| author | zhangx | 
| Wed, 03 Feb 2016 12:04:03 +0800 | |
| changeset 101 | c7db2ccba18a | 
| parent 81 | c495eb16beb6 | 
| permissions | -rw-r--r-- | 
| 81 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1 | theory CpsG | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2 | imports PIPDefs | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 5 | lemma Max_f_mono: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 6 | assumes seq: "A \<subseteq> B" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 7 |   and np: "A \<noteq> {}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 8 | and fnt: "finite B" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 9 | shows "Max (f ` A) \<le> Max (f ` B)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 10 | proof(rule Max_mono) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 11 | from seq show "f ` A \<subseteq> f ` B" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 12 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 13 |   from np show "f ` A \<noteq> {}" by auto
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 14 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 15 | from fnt and seq show "finite (f ` B)" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 16 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 17 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 18 | (* I am going to use this file as a start point to retrofiting | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 19 | PIPBasics.thy, which is originally called CpsG.ghy *) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 20 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 21 | locale valid_trace = | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 22 | fixes s | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 23 | assumes vt : "vt s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 24 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 25 | locale valid_trace_e = valid_trace + | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 26 | fixes e | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 27 | assumes vt_e: "vt (e#s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 28 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 29 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 30 | lemma pip_e: "PIP s e" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 31 | using vt_e by (cases, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 32 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 33 | end | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 34 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 35 | locale valid_trace_create = valid_trace_e + | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 36 | fixes th prio | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 37 | assumes is_create: "e = Create th prio" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 38 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 39 | locale valid_trace_exit = valid_trace_e + | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 40 | fixes th | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 41 | assumes is_exit: "e = Exit th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 42 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 43 | locale valid_trace_p = valid_trace_e + | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 44 | fixes th cs | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 45 | assumes is_p: "e = P th cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 46 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 47 | locale valid_trace_v = valid_trace_e + | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 48 | fixes th cs | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 49 | assumes is_v: "e = V th cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 50 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 51 | definition "rest = tl (wq s cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 52 | definition "wq' = (SOME q. distinct q \<and> set q = set rest)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 53 | end | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 54 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 55 | locale valid_trace_v_n = valid_trace_v + | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 56 | assumes rest_nnl: "rest \<noteq> []" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 57 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 58 | locale valid_trace_v_e = valid_trace_v + | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 59 | assumes rest_nil: "rest = []" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 60 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 61 | locale valid_trace_set= valid_trace_e + | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 62 | fixes th prio | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 63 | assumes is_set: "e = Set th prio" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 64 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 65 | context valid_trace | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 66 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 67 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 68 | lemma ind [consumes 0, case_names Nil Cons, induct type]: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 69 | assumes "PP []" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 70 | and "(\<And>s e. valid_trace_e s e \<Longrightarrow> | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 71 | PP s \<Longrightarrow> PIP s e \<Longrightarrow> PP (e # s))" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 72 | shows "PP s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 73 | proof(induct rule:vt.induct[OF vt, case_names Init Step]) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 74 | case Init | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 75 | from assms(1) show ?case . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 76 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 77 | case (Step s e) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 78 | show ?case | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 79 | proof(rule assms(2)) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 80 | show "valid_trace_e s e" using Step by (unfold_locales, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 81 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 82 | show "PP s" using Step by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 83 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 84 | show "PIP s e" using Step by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 85 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 86 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 87 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 88 | end | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 89 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 90 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 91 | lemma waiting_eq: "waiting s th cs = waiting (wq s) th cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 92 | by (unfold s_waiting_def cs_waiting_def wq_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 93 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 94 | lemma holding_eq: "holding (s::state) th cs = holding (wq s) th cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 95 | by (unfold s_holding_def wq_def cs_holding_def, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 96 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 97 | lemma runing_ready: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 98 | shows "runing s \<subseteq> readys s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 99 | unfolding runing_def readys_def | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 100 | by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 101 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 102 | lemma readys_threads: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 103 | shows "readys s \<subseteq> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 104 | unfolding readys_def | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 105 | by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 106 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 107 | lemma wq_v_neq [simp]: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 108 | "cs \<noteq> cs' \<Longrightarrow> wq (V thread cs#s) cs' = wq s cs'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 109 | by (auto simp:wq_def Let_def cp_def split:list.splits) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 110 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 111 | lemma runing_head: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 112 | assumes "th \<in> runing s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 113 | and "th \<in> set (wq_fun (schs s) cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 114 | shows "th = hd (wq_fun (schs s) cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 115 | using assms | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 116 | by (simp add:runing_def readys_def s_waiting_def wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 117 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 118 | context valid_trace | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 119 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 120 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 121 | lemma runing_wqE: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 122 | assumes "th \<in> runing s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 123 | and "th \<in> set (wq s cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 124 | obtains rest where "wq s cs = th#rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 125 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 126 | from assms(2) obtain th' rest where eq_wq: "wq s cs = th'#rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 127 | by (meson list.set_cases) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 128 | have "th' = th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 129 | proof(rule ccontr) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 130 | assume "th' \<noteq> th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 131 | hence "th \<noteq> hd (wq s cs)" using eq_wq by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 132 | with assms(2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 133 | have "waiting s th cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 134 | by (unfold s_waiting_def, fold wq_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 135 | with assms show False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 136 | by (unfold runing_def readys_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 137 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 138 | with eq_wq that show ?thesis by metis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 139 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 140 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 141 | end | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 142 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 143 | context valid_trace_p | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 144 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 145 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 146 | lemma wq_neq_simp [simp]: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 147 | assumes "cs' \<noteq> cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 148 | shows "wq (e#s) cs' = wq s cs'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 149 | using assms unfolding is_p wq_def | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 150 | by (auto simp:Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 151 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 152 | lemma runing_th_s: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 153 | shows "th \<in> runing s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 154 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 155 | from pip_e[unfolded is_p] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 156 | show ?thesis by (cases, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 157 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 158 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 159 | lemma th_not_waiting: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 160 | "\<not> waiting s th c" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 161 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 162 | have "th \<in> readys s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 163 | using runing_ready runing_th_s by blast | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 164 | thus ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 165 | by (unfold readys_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 166 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 167 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 168 | lemma waiting_neq_th: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 169 | assumes "waiting s t c" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 170 | shows "t \<noteq> th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 171 | using assms using th_not_waiting by blast | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 172 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 173 | lemma th_not_in_wq: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 174 | shows "th \<notin> set (wq s cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 175 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 176 | assume otherwise: "th \<in> set (wq s cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 177 | from runing_wqE[OF runing_th_s this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 178 | obtain rest where eq_wq: "wq s cs = th#rest" by blast | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 179 | with otherwise | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 180 | have "holding s th cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 181 | by (unfold s_holding_def, fold wq_def, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 182 | hence cs_th_RAG: "(Cs cs, Th th) \<in> RAG s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 183 | by (unfold s_RAG_def, fold holding_eq, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 184 | from pip_e[unfolded is_p] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 185 | show False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 186 | proof(cases) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 187 | case (thread_P) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 188 | with cs_th_RAG show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 189 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 190 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 191 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 192 | lemma wq_es_cs: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 193 | "wq (e#s) cs = wq s cs @ [th]" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 194 | by (unfold is_p wq_def, auto simp:Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 195 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 196 | lemma wq_distinct_kept: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 197 | assumes "distinct (wq s cs')" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 198 | shows "distinct (wq (e#s) cs')" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 199 | proof(cases "cs' = cs") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 200 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 201 | show ?thesis using True assms th_not_in_wq | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 202 | by (unfold True wq_es_cs, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 203 | qed (insert assms, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 204 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 205 | end | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 206 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 207 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 208 | context valid_trace_v | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 209 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 210 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 211 | lemma wq_neq_simp [simp]: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 212 | assumes "cs' \<noteq> cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 213 | shows "wq (e#s) cs' = wq s cs'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 214 | using assms unfolding is_v wq_def | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 215 | by (auto simp:Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 216 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 217 | lemma runing_th_s: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 218 | shows "th \<in> runing s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 219 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 220 | from pip_e[unfolded is_v] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 221 | show ?thesis by (cases, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 222 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 223 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 224 | lemma th_not_waiting: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 225 | "\<not> waiting s th c" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 226 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 227 | have "th \<in> readys s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 228 | using runing_ready runing_th_s by blast | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 229 | thus ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 230 | by (unfold readys_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 231 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 232 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 233 | lemma waiting_neq_th: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 234 | assumes "waiting s t c" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 235 | shows "t \<noteq> th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 236 | using assms using th_not_waiting by blast | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 237 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 238 | lemma wq_s_cs: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 239 | "wq s cs = th#rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 240 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 241 | from pip_e[unfolded is_v] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 242 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 243 | proof(cases) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 244 | case (thread_V) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 245 | from this(2) show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 246 | by (unfold rest_def s_holding_def, fold wq_def, | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 247 | metis empty_iff list.collapse list.set(1)) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 248 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 249 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 250 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 251 | lemma wq_es_cs: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 252 | "wq (e#s) cs = wq'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 253 | using wq_s_cs[unfolded wq_def] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 254 | by (auto simp:Let_def wq_def rest_def wq'_def is_v, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 255 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 256 | lemma wq_distinct_kept: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 257 | assumes "distinct (wq s cs')" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 258 | shows "distinct (wq (e#s) cs')" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 259 | proof(cases "cs' = cs") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 260 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 261 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 262 | proof(unfold True wq_es_cs wq'_def, rule someI2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 263 | show "distinct rest \<and> set rest = set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 264 | using assms[unfolded True wq_s_cs] by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 265 | qed simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 266 | qed (insert assms, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 267 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 268 | end | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 269 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 270 | context valid_trace | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 271 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 272 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 273 | lemma actor_inv: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 274 | assumes "PIP s e" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 275 | and "\<not> isCreate e" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 276 | shows "actor e \<in> runing s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 277 | using assms | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 278 | by (induct, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 279 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 280 | lemma isP_E: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 281 | assumes "isP e" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 282 | obtains cs where "e = P (actor e) cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 283 | using assms by (cases e, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 284 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 285 | lemma isV_E: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 286 | assumes "isV e" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 287 | obtains cs where "e = V (actor e) cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 288 | using assms by (cases e, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 289 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 290 | lemma wq_distinct: "distinct (wq s cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 291 | proof(induct rule:ind) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 292 | case (Cons s e) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 293 | interpret vt_e: valid_trace_e s e using Cons by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 294 | show ?case | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 295 | proof(cases e) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 296 | case (V th cs) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 297 | interpret vt_v: valid_trace_v s e th cs using V by (unfold_locales, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 298 | show ?thesis using Cons by (simp add: vt_v.wq_distinct_kept) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 299 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 300 | qed (unfold wq_def Let_def, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 301 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 302 | end | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 303 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 304 | context valid_trace_e | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 305 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 306 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 307 | text {*
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 308 |   The following lemma shows that only the @{text "P"}
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 309 | operation can add new thread into waiting queues. | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 310 | Such kind of lemmas are very obvious, but need to be checked formally. | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 311 | This is a kind of confirmation that our modelling is correct. | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 312 | *} | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 313 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 314 | lemma wq_in_inv: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 315 | assumes s_ni: "thread \<notin> set (wq s cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 316 | and s_i: "thread \<in> set (wq (e#s) cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 317 | shows "e = P thread cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 318 | proof(cases e) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 319 |   -- {* This is the only non-trivial case: *}
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 320 | case (V th cs1) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 321 | have False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 322 | proof(cases "cs1 = cs") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 323 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 324 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 325 | proof(cases "(wq s cs1)") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 326 | case (Cons w_hd w_tl) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 327 | have "set (wq (e#s) cs) \<subseteq> set (wq s cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 328 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 329 | have "(wq (e#s) cs) = (SOME q. distinct q \<and> set q = set w_tl)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 330 | using Cons V by (auto simp:wq_def Let_def True split:if_splits) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 331 | moreover have "set ... \<subseteq> set (wq s cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 332 | proof(rule someI2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 333 | show "distinct w_tl \<and> set w_tl = set w_tl" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 334 | by (metis distinct.simps(2) local.Cons wq_distinct) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 335 | qed (insert Cons True, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 336 | ultimately show ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 337 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 338 | with assms show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 339 | qed (insert assms V True, auto simp:wq_def Let_def split:if_splits) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 340 | qed (insert assms V, auto simp:wq_def Let_def split:if_splits) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 341 | thus ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 342 | qed (insert assms, auto simp:wq_def Let_def split:if_splits) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 343 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 344 | lemma wq_out_inv: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 345 | assumes s_in: "thread \<in> set (wq s cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 346 | and s_hd: "thread = hd (wq s cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 347 | and s_i: "thread \<noteq> hd (wq (e#s) cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 348 | shows "e = V thread cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 349 | proof(cases e) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 350 | -- {* There are only two non-trivial cases: *}
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 351 | case (V th cs1) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 352 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 353 | proof(cases "cs1 = cs") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 354 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 355 | have "PIP s (V th cs)" using pip_e[unfolded V[unfolded True]] . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 356 | thus ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 357 | proof(cases) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 358 | case (thread_V) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 359 | moreover have "th = thread" using thread_V(2) s_hd | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 360 | by (unfold s_holding_def wq_def, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 361 | ultimately show ?thesis using V True by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 362 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 363 | qed (insert assms V, auto simp:wq_def Let_def split:if_splits) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 364 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 365 | case (P th cs1) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 366 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 367 | proof(cases "cs1 = cs") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 368 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 369 | with P have "wq (e#s) cs = wq_fun (schs s) cs @ [th]" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 370 | by (auto simp:wq_def Let_def split:if_splits) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 371 | with s_i s_hd s_in have False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 372 | by (metis empty_iff hd_append2 list.set(1) wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 373 | thus ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 374 | qed (insert assms P, auto simp:wq_def Let_def split:if_splits) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 375 | qed (insert assms, auto simp:wq_def Let_def split:if_splits) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 376 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 377 | end | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 378 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 379 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 380 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 381 | context valid_trace | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 382 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 383 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 384 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 385 | text {* (* ddd *)
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 386 | The nature of the work is like this: since it starts from a very simple and basic | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 387 | model, even intuitively very `basic` and `obvious` properties need to derived from scratch. | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 388 | For instance, the fact | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 389 | that one thread can not be blocked by two critical resources at the same time | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 390 | is obvious, because only running threads can make new requests, if one is waiting for | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 391 | a critical resource and get blocked, it can not make another resource request and get | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 392 | blocked the second time (because it is not running). | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 393 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 394 | To derive this fact, one needs to prove by contraction and | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 395 |   reason about time (or @{text "moement"}). The reasoning is based on a generic theorem
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 396 |   named @{text "p_split"}, which is about status changing along the time axis. It says if 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 397 |   a condition @{text "Q"} is @{text "True"} at a state @{text "s"},
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 398 |   but it was @{text "False"} at the very beginning, then there must exits a moment @{text "t"} 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 399 |   in the history of @{text "s"} (notice that @{text "s"} itself is essentially the history 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 400 |   of events leading to it), such that @{text "Q"} switched 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 401 |   from being @{text "False"} to @{text "True"} and kept being @{text "True"}
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 402 |   till the last moment of @{text "s"}.
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 403 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 404 |   Suppose a thread @{text "th"} is blocked
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 405 |   on @{text "cs1"} and @{text "cs2"} in some state @{text "s"}, 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 406 | since no thread is blocked at the very beginning, by applying | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 407 |   @{text "p_split"} to these two blocking facts, there exist 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 408 |   two moments @{text "t1"} and @{text "t2"}  in @{text "s"}, such that 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 409 |   @{text "th"} got blocked on @{text "cs1"} and @{text "cs2"} 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 410 | and kept on blocked on them respectively ever since. | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 411 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 412 |   Without lost of generality, we assume @{text "t1"} is earlier than @{text "t2"}.
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 413 |   However, since @{text "th"} was blocked ever since memonent @{text "t1"}, so it was still
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 414 |   in blocked state at moment @{text "t2"} and could not
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 415 | make any request and get blocked the second time: Contradiction. | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 416 | *} | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 417 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 418 | lemma waiting_unique_pre: (* ddd *) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 419 | assumes h11: "thread \<in> set (wq s cs1)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 420 | and h12: "thread \<noteq> hd (wq s cs1)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 421 | assumes h21: "thread \<in> set (wq s cs2)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 422 | and h22: "thread \<noteq> hd (wq s cs2)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 423 | and neq12: "cs1 \<noteq> cs2" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 424 | shows "False" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 425 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 426 | let "?Q" = "\<lambda> cs s. thread \<in> set (wq s cs) \<and> thread \<noteq> hd (wq s cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 427 | from h11 and h12 have q1: "?Q cs1 s" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 428 | from h21 and h22 have q2: "?Q cs2 s" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 429 | have nq1: "\<not> ?Q cs1 []" by (simp add:wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 430 | have nq2: "\<not> ?Q cs2 []" by (simp add:wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 431 | from p_split [of "?Q cs1", OF q1 nq1] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 432 | obtain t1 where lt1: "t1 < length s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 433 | and np1: "\<not> ?Q cs1 (moment t1 s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 434 | and nn1: "(\<forall>i'>t1. ?Q cs1 (moment i' s))" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 435 | from p_split [of "?Q cs2", OF q2 nq2] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 436 | obtain t2 where lt2: "t2 < length s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 437 | and np2: "\<not> ?Q cs2 (moment t2 s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 438 | and nn2: "(\<forall>i'>t2. ?Q cs2 (moment i' s))" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 439 |   { fix s cs
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 440 | assume q: "?Q cs s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 441 | have "thread \<notin> runing s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 442 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 443 | assume "thread \<in> runing s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 444 | hence " \<forall>cs. \<not> (thread \<in> set (wq_fun (schs s) cs) \<and> | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 445 | thread \<noteq> hd (wq_fun (schs s) cs))" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 446 | by (unfold runing_def s_waiting_def readys_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 447 | from this[rule_format, of cs] q | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 448 | show False by (simp add: wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 449 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 450 | } note q_not_runing = this | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 451 |   { fix t1 t2 cs1 cs2
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 452 | assume lt1: "t1 < length s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 453 | and np1: "\<not> ?Q cs1 (moment t1 s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 454 | and nn1: "(\<forall>i'>t1. ?Q cs1 (moment i' s))" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 455 | and lt2: "t2 < length s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 456 | and np2: "\<not> ?Q cs2 (moment t2 s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 457 | and nn2: "(\<forall>i'>t2. ?Q cs2 (moment i' s))" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 458 | and lt12: "t1 < t2" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 459 | let ?t3 = "Suc t2" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 460 | from lt2 have le_t3: "?t3 \<le> length s" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 461 | from moment_plus [OF this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 462 | obtain e where eq_m: "moment ?t3 s = e#moment t2 s" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 463 | have "t2 < ?t3" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 464 | from nn2 [rule_format, OF this] and eq_m | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 465 | have h1: "thread \<in> set (wq (e#moment t2 s) cs2)" and | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 466 | h2: "thread \<noteq> hd (wq (e#moment t2 s) cs2)" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 467 | have "vt (e#moment t2 s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 468 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 469 | from vt_moment | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 470 | have "vt (moment ?t3 s)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 471 | with eq_m show ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 472 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 473 | then interpret vt_e: valid_trace_e "moment t2 s" "e" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 474 | by (unfold_locales, auto, cases, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 475 | have ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 476 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 477 | have "thread \<in> runing (moment t2 s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 478 | proof(cases "thread \<in> set (wq (moment t2 s) cs2)") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 479 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 480 | have "e = V thread cs2" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 481 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 482 | have eq_th: "thread = hd (wq (moment t2 s) cs2)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 483 | using True and np2 by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 484 | from vt_e.wq_out_inv[OF True this h2] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 485 | show ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 486 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 487 | thus ?thesis using vt_e.actor_inv[OF vt_e.pip_e] by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 488 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 489 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 490 | have "e = P thread cs2" using vt_e.wq_in_inv[OF False h1] . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 491 | with vt_e.actor_inv[OF vt_e.pip_e] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 492 | show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 493 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 494 | moreover have "thread \<notin> runing (moment t2 s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 495 | by (rule q_not_runing[OF nn1[rule_format, OF lt12]]) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 496 | ultimately show ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 497 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 498 | } note lt_case = this | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 499 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 500 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 501 |     { assume "t1 < t2"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 502 | from lt_case[OF lt1 np1 nn1 lt2 np2 nn2 this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 503 | have ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 504 |     } moreover {
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 505 | assume "t2 < t1" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 506 | from lt_case[OF lt2 np2 nn2 lt1 np1 nn1 this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 507 | have ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 508 |     } moreover {
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 509 | assume eq_12: "t1 = t2" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 510 | let ?t3 = "Suc t2" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 511 | from lt2 have le_t3: "?t3 \<le> length s" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 512 | from moment_plus [OF this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 513 | obtain e where eq_m: "moment ?t3 s = e#moment t2 s" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 514 | have lt_2: "t2 < ?t3" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 515 | from nn2 [rule_format, OF this] and eq_m | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 516 | have h1: "thread \<in> set (wq (e#moment t2 s) cs2)" and | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 517 | h2: "thread \<noteq> hd (wq (e#moment t2 s) cs2)" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 518 | from nn1[rule_format, OF lt_2[folded eq_12]] eq_m[folded eq_12] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 519 | have g1: "thread \<in> set (wq (e#moment t1 s) cs1)" and | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 520 | g2: "thread \<noteq> hd (wq (e#moment t1 s) cs1)" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 521 | have "vt (e#moment t2 s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 522 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 523 | from vt_moment | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 524 | have "vt (moment ?t3 s)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 525 | with eq_m show ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 526 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 527 | then interpret vt_e: valid_trace_e "moment t2 s" "e" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 528 | by (unfold_locales, auto, cases, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 529 | have "e = V thread cs2 \<or> e = P thread cs2" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 530 | proof(cases "thread \<in> set (wq (moment t2 s) cs2)") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 531 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 532 | have "e = V thread cs2" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 533 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 534 | have eq_th: "thread = hd (wq (moment t2 s) cs2)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 535 | using True and np2 by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 536 | from vt_e.wq_out_inv[OF True this h2] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 537 | show ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 538 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 539 | thus ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 540 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 541 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 542 | have "e = P thread cs2" using vt_e.wq_in_inv[OF False h1] . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 543 | thus ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 544 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 545 | moreover have "e = V thread cs1 \<or> e = P thread cs1" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 546 | proof(cases "thread \<in> set (wq (moment t1 s) cs1)") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 547 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 548 | have eq_th: "thread = hd (wq (moment t1 s) cs1)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 549 | using True and np1 by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 550 | from vt_e.wq_out_inv[folded eq_12, OF True this g2] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 551 | have "e = V thread cs1" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 552 | thus ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 553 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 554 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 555 | have "e = P thread cs1" using vt_e.wq_in_inv[folded eq_12, OF False g1] . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 556 | thus ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 557 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 558 | ultimately have ?thesis using neq12 by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 559 | } ultimately show ?thesis using nat_neq_iff by blast | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 560 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 561 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 562 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 563 | text {*
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 564 |   This lemma is a simple corrolary of @{text "waiting_unique_pre"}.
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 565 | *} | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 566 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 567 | lemma waiting_unique: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 568 | assumes "waiting s th cs1" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 569 | and "waiting s th cs2" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 570 | shows "cs1 = cs2" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 571 | using waiting_unique_pre assms | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 572 | unfolding wq_def s_waiting_def | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 573 | by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 574 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 575 | end | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 576 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 577 | (* not used *) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 578 | text {*
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 579 | Every thread can only be blocked on one critical resource, | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 580 | symmetrically, every critical resource can only be held by one thread. | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 581 | This fact is much more easier according to our definition. | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 582 | *} | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 583 | lemma held_unique: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 584 | assumes "holding (s::event list) th1 cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 585 | and "holding s th2 cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 586 | shows "th1 = th2" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 587 | by (insert assms, unfold s_holding_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 588 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 589 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 590 | lemma last_set_lt: "th \<in> threads s \<Longrightarrow> last_set th s < length s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 591 | apply (induct s, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 592 | by (case_tac a, auto split:if_splits) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 593 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 594 | lemma last_set_unique: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 595 | "\<lbrakk>last_set th1 s = last_set th2 s; th1 \<in> threads s; th2 \<in> threads s\<rbrakk> | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 596 | \<Longrightarrow> th1 = th2" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 597 | apply (induct s, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 598 | by (case_tac a, auto split:if_splits dest:last_set_lt) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 599 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 600 | lemma preced_unique : | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 601 | assumes pcd_eq: "preced th1 s = preced th2 s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 602 | and th_in1: "th1 \<in> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 603 | and th_in2: " th2 \<in> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 604 | shows "th1 = th2" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 605 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 606 | from pcd_eq have "last_set th1 s = last_set th2 s" by (simp add:preced_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 607 | from last_set_unique [OF this th_in1 th_in2] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 608 | show ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 609 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 610 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 611 | lemma preced_linorder: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 612 | assumes neq_12: "th1 \<noteq> th2" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 613 | and th_in1: "th1 \<in> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 614 | and th_in2: " th2 \<in> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 615 | shows "preced th1 s < preced th2 s \<or> preced th1 s > preced th2 s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 616 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 617 | from preced_unique [OF _ th_in1 th_in2] and neq_12 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 618 | have "preced th1 s \<noteq> preced th2 s" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 619 | thus ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 620 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 621 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 622 | (* An aux lemma used later *) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 623 | lemma unique_minus: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 624 | assumes unique: "\<And> a b c. \<lbrakk>(a, b) \<in> r; (a, c) \<in> r\<rbrakk> \<Longrightarrow> b = c" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 625 | and xy: "(x, y) \<in> r" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 626 | and xz: "(x, z) \<in> r^+" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 627 | and neq: "y \<noteq> z" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 628 | shows "(y, z) \<in> r^+" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 629 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 630 | from xz and neq show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 631 | proof(induct) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 632 | case (base ya) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 633 | have "(x, ya) \<in> r" by fact | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 634 | from unique [OF xy this] have "y = ya" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 635 | with base show ?case by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 636 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 637 | case (step ya z) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 638 | show ?case | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 639 | proof(cases "y = ya") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 640 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 641 | from step True show ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 642 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 643 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 644 | from step False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 645 | show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 646 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 647 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 648 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 649 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 650 | lemma unique_base: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 651 | assumes unique: "\<And> a b c. \<lbrakk>(a, b) \<in> r; (a, c) \<in> r\<rbrakk> \<Longrightarrow> b = c" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 652 | and xy: "(x, y) \<in> r" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 653 | and xz: "(x, z) \<in> r^+" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 654 | and neq_yz: "y \<noteq> z" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 655 | shows "(y, z) \<in> r^+" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 656 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 657 | from xz neq_yz show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 658 | proof(induct) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 659 | case (base ya) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 660 | from xy unique base show ?case by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 661 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 662 | case (step ya z) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 663 | show ?case | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 664 | proof(cases "y = ya") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 665 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 666 | from True step show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 667 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 668 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 669 | from False step | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 670 | have "(y, ya) \<in> r\<^sup>+" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 671 | with step show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 672 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 673 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 674 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 675 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 676 | lemma unique_chain: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 677 | assumes unique: "\<And> a b c. \<lbrakk>(a, b) \<in> r; (a, c) \<in> r\<rbrakk> \<Longrightarrow> b = c" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 678 | and xy: "(x, y) \<in> r^+" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 679 | and xz: "(x, z) \<in> r^+" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 680 | and neq_yz: "y \<noteq> z" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 681 | shows "(y, z) \<in> r^+ \<or> (z, y) \<in> r^+" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 682 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 683 | from xy xz neq_yz show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 684 | proof(induct) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 685 | case (base y) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 686 | have h1: "(x, y) \<in> r" and h2: "(x, z) \<in> r\<^sup>+" and h3: "y \<noteq> z" using base by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 687 | from unique_base [OF _ h1 h2 h3] and unique show ?case by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 688 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 689 | case (step y za) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 690 | show ?case | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 691 | proof(cases "y = z") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 692 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 693 | from True step show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 694 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 695 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 696 | from False step have "(y, z) \<in> r\<^sup>+ \<or> (z, y) \<in> r\<^sup>+" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 697 | thus ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 698 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 699 | assume "(z, y) \<in> r\<^sup>+" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 700 | with step have "(z, za) \<in> r\<^sup>+" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 701 | thus ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 702 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 703 | assume h: "(y, z) \<in> r\<^sup>+" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 704 | from step have yza: "(y, za) \<in> r" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 705 | from step have "za \<noteq> z" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 706 | from unique_minus [OF _ yza h this] and unique | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 707 | have "(za, z) \<in> r\<^sup>+" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 708 | thus ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 709 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 710 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 711 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 712 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 713 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 714 | text {*
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 715 |   The following three lemmas show that @{text "RAG"} does not change
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 716 |   by the happening of @{text "Set"}, @{text "Create"} and @{text "Exit"}
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 717 | events, respectively. | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 718 | *} | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 719 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 720 | lemma RAG_set_unchanged: "(RAG (Set th prio # s)) = RAG s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 721 | apply (unfold s_RAG_def s_waiting_def wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 722 | by (simp add:Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 723 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 724 | lemma RAG_create_unchanged: "(RAG (Create th prio # s)) = RAG s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 725 | apply (unfold s_RAG_def s_waiting_def wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 726 | by (simp add:Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 727 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 728 | lemma RAG_exit_unchanged: "(RAG (Exit th # s)) = RAG s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 729 | apply (unfold s_RAG_def s_waiting_def wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 730 | by (simp add:Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 731 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 732 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 733 | context valid_trace_v | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 734 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 735 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 736 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 737 | lemma distinct_rest: "distinct rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 738 | by (simp add: distinct_tl rest_def wq_distinct) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 739 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 740 | definition "wq' = (SOME q. distinct q \<and> set q = set rest)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 741 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 742 | lemma runing_th_s: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 743 | shows "th \<in> runing s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 744 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 745 | from pip_e[unfolded is_v] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 746 | show ?thesis by (cases, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 747 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 748 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 749 | lemma holding_cs_eq_th: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 750 | assumes "holding s t cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 751 | shows "t = th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 752 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 753 | from pip_e[unfolded is_v] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 754 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 755 | proof(cases) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 756 | case (thread_V) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 757 | from held_unique[OF this(2) assms] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 758 | show ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 759 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 760 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 761 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 762 | lemma th_not_waiting: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 763 | "\<not> waiting s th c" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 764 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 765 | have "th \<in> readys s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 766 | using runing_ready runing_th_s by blast | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 767 | thus ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 768 | by (unfold readys_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 769 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 770 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 771 | lemma waiting_neq_th: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 772 | assumes "waiting s t c" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 773 | shows "t \<noteq> th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 774 | using assms using th_not_waiting by blast | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 775 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 776 | lemma wq_s_cs: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 777 | "wq s cs = th#rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 778 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 779 | from pip_e[unfolded is_v] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 780 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 781 | proof(cases) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 782 | case (thread_V) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 783 | from this(2) show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 784 | by (unfold rest_def s_holding_def, fold wq_def, | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 785 | metis empty_iff list.collapse list.set(1)) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 786 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 787 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 788 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 789 | lemma wq_es_cs: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 790 | "wq (e#s) cs = wq'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 791 | using wq_s_cs[unfolded wq_def] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 792 | by (auto simp:Let_def wq_def rest_def wq'_def is_v, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 793 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 794 | lemma distinct_wq': "distinct wq'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 795 | by (metis (mono_tags, lifting) distinct_rest some_eq_ex wq'_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 796 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 797 | lemma th'_in_inv: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 798 | assumes "th' \<in> set wq'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 799 | shows "th' \<in> set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 800 | using assms | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 801 | by (metis (mono_tags, lifting) distinct.simps(2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 802 | rest_def some_eq_ex wq'_def wq_distinct wq_s_cs) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 803 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 804 | lemma neq_t_th: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 805 | assumes "waiting (e#s) t c" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 806 | shows "t \<noteq> th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 807 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 808 | assume otherwise: "t = th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 809 | show False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 810 | proof(cases "c = cs") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 811 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 812 | have "t \<in> set wq'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 813 | using assms[unfolded True s_waiting_def, folded wq_def, unfolded wq_es_cs] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 814 | by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 815 | from th'_in_inv[OF this] have "t \<in> set rest" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 816 | with wq_s_cs[folded otherwise] wq_distinct[of cs] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 817 | show ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 818 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 819 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 820 | have "wq (e#s) c = wq s c" using False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 821 | by (unfold is_v, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 822 | hence "waiting s t c" using assms | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 823 | by (simp add: cs_waiting_def waiting_eq) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 824 | hence "t \<notin> readys s" by (unfold readys_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 825 | hence "t \<notin> runing s" using runing_ready by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 826 | with runing_th_s[folded otherwise] show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 827 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 828 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 829 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 830 | lemma waiting_esI1: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 831 | assumes "waiting s t c" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 832 | and "c \<noteq> cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 833 | shows "waiting (e#s) t c" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 834 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 835 | have "wq (e#s) c = wq s c" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 836 | using assms(2) is_v by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 837 | with assms(1) show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 838 | using cs_waiting_def waiting_eq by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 839 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 840 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 841 | lemma holding_esI2: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 842 | assumes "c \<noteq> cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 843 | and "holding s t c" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 844 | shows "holding (e#s) t c" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 845 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 846 | from assms(1) have "wq (e#s) c = wq s c" using is_v by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 847 | from assms(2)[unfolded s_holding_def, folded wq_def, | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 848 | folded this, unfolded wq_def, folded s_holding_def] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 849 | show ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 850 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 851 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 852 | lemma holding_esI1: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 853 | assumes "holding s t c" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 854 | and "t \<noteq> th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 855 | shows "holding (e#s) t c" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 856 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 857 | have "c \<noteq> cs" using assms using holding_cs_eq_th by blast | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 858 | from holding_esI2[OF this assms(1)] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 859 | show ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 860 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 861 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 862 | end | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 863 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 864 | context valid_trace_v_n | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 865 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 866 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 867 | lemma neq_wq': "wq' \<noteq> []" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 868 | proof (unfold wq'_def, rule someI2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 869 | show "distinct rest \<and> set rest = set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 870 | by (simp add: distinct_rest) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 871 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 872 | fix x | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 873 | assume " distinct x \<and> set x = set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 874 | thus "x \<noteq> []" using rest_nnl by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 875 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 876 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 877 | definition "taker = hd wq'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 878 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 879 | definition "rest' = tl wq'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 880 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 881 | lemma eq_wq': "wq' = taker # rest'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 882 | by (simp add: neq_wq' rest'_def taker_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 883 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 884 | lemma next_th_taker: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 885 | shows "next_th s th cs taker" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 886 | using rest_nnl taker_def wq'_def wq_s_cs | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 887 | by (auto simp:next_th_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 888 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 889 | lemma taker_unique: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 890 | assumes "next_th s th cs taker'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 891 | shows "taker' = taker" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 892 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 893 | from assms | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 894 | obtain rest' where | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 895 | h: "wq s cs = th # rest'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 896 | "taker' = hd (SOME q. distinct q \<and> set q = set rest')" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 897 | by (unfold next_th_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 898 | with wq_s_cs have "rest' = rest" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 899 | thus ?thesis using h(2) taker_def wq'_def by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 900 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 901 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 902 | lemma waiting_set_eq: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 903 |   "{(Th th', Cs cs) |th'. next_th s th cs th'} = {(Th taker, Cs cs)}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 904 | by (smt all_not_in_conv bot.extremum insertI1 insert_subset | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 905 | mem_Collect_eq next_th_taker subsetI subset_antisym taker_def taker_unique) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 906 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 907 | lemma holding_set_eq: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 908 |   "{(Cs cs, Th th') |th'.  next_th s th cs th'} = {(Cs cs, Th taker)}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 909 | using next_th_taker taker_def waiting_set_eq | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 910 | by fastforce | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 911 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 912 | lemma holding_taker: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 913 | shows "holding (e#s) taker cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 914 | by (unfold s_holding_def, fold wq_def, unfold wq_es_cs, | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 915 | auto simp:neq_wq' taker_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 916 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 917 | lemma waiting_esI2: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 918 | assumes "waiting s t cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 919 | and "t \<noteq> taker" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 920 | shows "waiting (e#s) t cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 921 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 922 | have "t \<in> set wq'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 923 | proof(unfold wq'_def, rule someI2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 924 | show "distinct rest \<and> set rest = set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 925 | by (simp add: distinct_rest) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 926 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 927 | fix x | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 928 | assume "distinct x \<and> set x = set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 929 | moreover have "t \<in> set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 930 | using assms(1) cs_waiting_def waiting_eq wq_s_cs by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 931 | ultimately show "t \<in> set x" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 932 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 933 | moreover have "t \<noteq> hd wq'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 934 | using assms(2) taker_def by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 935 | ultimately show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 936 | by (unfold s_waiting_def, fold wq_def, unfold wq_es_cs, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 937 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 938 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 939 | lemma waiting_esE: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 940 | assumes "waiting (e#s) t c" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 941 | obtains "c \<noteq> cs" "waiting s t c" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 942 | | "c = cs" "t \<noteq> taker" "waiting s t cs" "t \<in> set rest'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 943 | proof(cases "c = cs") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 944 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 945 | hence "wq (e#s) c = wq s c" using is_v by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 946 | with assms have "waiting s t c" using cs_waiting_def waiting_eq by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 947 | from that(1)[OF False this] show ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 948 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 949 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 950 | from assms[unfolded s_waiting_def True, folded wq_def, unfolded wq_es_cs] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 951 | have "t \<noteq> hd wq'" "t \<in> set wq'" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 952 | hence "t \<noteq> taker" by (simp add: taker_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 953 | moreover hence "t \<noteq> th" using assms neq_t_th by blast | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 954 | moreover have "t \<in> set rest" by (simp add: `t \<in> set wq'` th'_in_inv) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 955 | ultimately have "waiting s t cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 956 | by (metis cs_waiting_def list.distinct(2) list.sel(1) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 957 | list.set_sel(2) rest_def waiting_eq wq_s_cs) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 958 | show ?thesis using that(2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 959 | using True `t \<in> set wq'` `t \<noteq> taker` `waiting s t cs` eq_wq' by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 960 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 961 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 962 | lemma holding_esI1: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 963 | assumes "c = cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 964 | and "t = taker" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 965 | shows "holding (e#s) t c" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 966 | by (unfold assms, simp add: holding_taker) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 967 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 968 | lemma holding_esE: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 969 | assumes "holding (e#s) t c" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 970 | obtains "c = cs" "t = taker" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 971 | | "c \<noteq> cs" "holding s t c" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 972 | proof(cases "c = cs") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 973 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 974 | from assms[unfolded True, unfolded s_holding_def, | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 975 | folded wq_def, unfolded wq_es_cs] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 976 | have "t = taker" by (simp add: taker_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 977 | from that(1)[OF True this] show ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 978 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 979 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 980 | hence "wq (e#s) c = wq s c" using is_v by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 981 | from assms[unfolded s_holding_def, folded wq_def, | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 982 | unfolded this, unfolded wq_def, folded s_holding_def] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 983 | have "holding s t c" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 984 | from that(2)[OF False this] show ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 985 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 986 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 987 | end | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 988 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 989 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 990 | context valid_trace_v_n | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 991 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 992 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 993 | lemma nil_wq': "wq' = []" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 994 | proof (unfold wq'_def, rule someI2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 995 | show "distinct rest \<and> set rest = set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 996 | by (simp add: distinct_rest) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 997 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 998 | fix x | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 999 | assume " distinct x \<and> set x = set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1000 | thus "x = []" using rest_nil by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1001 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1002 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1003 | lemma no_taker: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1004 | assumes "next_th s th cs taker" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1005 | shows "False" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1006 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1007 | from assms[unfolded next_th_def] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1008 | obtain rest' where "wq s cs = th # rest'" "rest' \<noteq> []" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1009 | by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1010 | thus ?thesis using rest_def rest_nil by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1011 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1012 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1013 | lemma waiting_set_eq: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1014 |   "{(Th th', Cs cs) |th'. next_th s th cs th'} = {}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1015 | using no_taker by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1016 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1017 | lemma holding_set_eq: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1018 |   "{(Cs cs, Th th') |th'.  next_th s th cs th'} = {}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1019 | using no_taker by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1020 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1021 | lemma no_holding: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1022 | assumes "holding (e#s) taker cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1023 | shows False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1024 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1025 | from wq_es_cs[unfolded nil_wq'] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1026 | have " wq (e # s) cs = []" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1027 | from assms[unfolded s_holding_def, folded wq_def, unfolded this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1028 | show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1029 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1030 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1031 | lemma no_waiting: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1032 | assumes "waiting (e#s) t cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1033 | shows False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1034 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1035 | from wq_es_cs[unfolded nil_wq'] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1036 | have " wq (e # s) cs = []" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1037 | from assms[unfolded s_waiting_def, folded wq_def, unfolded this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1038 | show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1039 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1040 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1041 | lemma waiting_esI2: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1042 | assumes "waiting s t c" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1043 | shows "waiting (e#s) t c" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1044 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1045 | have "c \<noteq> cs" using assms | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1046 | using cs_waiting_def rest_nil waiting_eq wq_s_cs by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1047 | from waiting_esI1[OF assms this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1048 | show ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1049 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1050 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1051 | lemma waiting_esE: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1052 | assumes "waiting (e#s) t c" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1053 | obtains "c \<noteq> cs" "waiting s t c" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1054 | proof(cases "c = cs") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1055 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1056 | hence "wq (e#s) c = wq s c" using is_v by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1057 | with assms have "waiting s t c" using cs_waiting_def waiting_eq by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1058 | from that(1)[OF False this] show ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1059 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1060 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1061 | from no_waiting[OF assms[unfolded True]] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1062 | show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1063 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1064 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1065 | lemma holding_esE: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1066 | assumes "holding (e#s) t c" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1067 | obtains "c \<noteq> cs" "holding s t c" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1068 | proof(cases "c = cs") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1069 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1070 | from no_holding[OF assms[unfolded True]] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1071 | show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1072 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1073 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1074 | hence "wq (e#s) c = wq s c" using is_v by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1075 | from assms[unfolded s_holding_def, folded wq_def, | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1076 | unfolded this, unfolded wq_def, folded s_holding_def] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1077 | have "holding s t c" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1078 | from that[OF False this] show ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1079 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1080 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1081 | end (* ccc *) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1082 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1083 | lemma rel_eqI: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1084 | assumes "\<And> x y. (x,y) \<in> A \<Longrightarrow> (x,y) \<in> B" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1085 | and "\<And> x y. (x,y) \<in> B \<Longrightarrow> (x, y) \<in> A" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1086 | shows "A = B" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1087 | using assms by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1088 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1089 | lemma in_RAG_E: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1090 | assumes "(n1, n2) \<in> RAG (s::state)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1091 | obtains (waiting) th cs where "n1 = Th th" "n2 = Cs cs" "waiting s th cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1092 | | (holding) th cs where "n1 = Cs cs" "n2 = Th th" "holding s th cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1093 | using assms[unfolded s_RAG_def, folded waiting_eq holding_eq] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1094 | by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1095 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1096 | context valid_trace_v | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1097 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1098 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1099 | lemma RAG_es: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1100 | "RAG (e # s) = | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1101 |    RAG s - {(Cs cs, Th th)} -
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1102 |      {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1103 |      {(Cs cs, Th th') |th'.  next_th s th cs th'}" (is "?L = ?R")
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1104 | proof(rule rel_eqI) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1105 | fix n1 n2 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1106 | assume "(n1, n2) \<in> ?L" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1107 | thus "(n1, n2) \<in> ?R" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1108 | proof(cases rule:in_RAG_E) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1109 | case (waiting th' cs') | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1110 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1111 | proof(cases "rest = []") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1112 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1113 | interpret h_n: valid_trace_v_n s e th cs | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1114 | by (unfold_locales, insert False, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1115 | from waiting(3) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1116 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1117 | proof(cases rule:h_n.waiting_esE) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1118 | case 1 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1119 | with waiting(1,2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1120 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1121 | by (unfold h_n.waiting_set_eq h_n.holding_set_eq s_RAG_def, | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1122 | fold waiting_eq, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1123 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1124 | case 2 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1125 | with waiting(1,2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1126 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1127 | by (unfold h_n.waiting_set_eq h_n.holding_set_eq s_RAG_def, | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1128 | fold waiting_eq, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1129 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1130 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1131 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1132 | interpret h_e: valid_trace_v_e s e th cs | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1133 | by (unfold_locales, insert True, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1134 | from waiting(3) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1135 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1136 | proof(cases rule:h_e.waiting_esE) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1137 | case 1 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1138 | with waiting(1,2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1139 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1140 | by (unfold h_e.waiting_set_eq h_e.holding_set_eq s_RAG_def, | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1141 | fold waiting_eq, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1142 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1143 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1144 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1145 | case (holding th' cs') | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1146 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1147 | proof(cases "rest = []") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1148 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1149 | interpret h_n: valid_trace_v_n s e th cs | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1150 | by (unfold_locales, insert False, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1151 | from holding(3) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1152 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1153 | proof(cases rule:h_n.holding_esE) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1154 | case 1 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1155 | with holding(1,2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1156 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1157 | by (unfold h_n.waiting_set_eq h_n.holding_set_eq s_RAG_def, | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1158 | fold waiting_eq, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1159 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1160 | case 2 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1161 | with holding(1,2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1162 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1163 | by (unfold h_n.waiting_set_eq h_n.holding_set_eq s_RAG_def, | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1164 | fold holding_eq, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1165 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1166 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1167 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1168 | interpret h_e: valid_trace_v_e s e th cs | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1169 | by (unfold_locales, insert True, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1170 | from holding(3) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1171 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1172 | proof(cases rule:h_e.holding_esE) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1173 | case 1 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1174 | with holding(1,2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1175 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1176 | by (unfold h_e.waiting_set_eq h_e.holding_set_eq s_RAG_def, | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1177 | fold holding_eq, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1178 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1179 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1180 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1181 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1182 | fix n1 n2 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1183 | assume h: "(n1, n2) \<in> ?R" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1184 | show "(n1, n2) \<in> ?L" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1185 | proof(cases "rest = []") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1186 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1187 | interpret h_n: valid_trace_v_n s e th cs | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1188 | by (unfold_locales, insert False, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1189 | from h[unfolded h_n.waiting_set_eq h_n.holding_set_eq] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1190 | have "((n1, n2) \<in> RAG s \<and> (n1 \<noteq> Cs cs \<or> n2 \<noteq> Th th) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1191 | \<and> (n1 \<noteq> Th h_n.taker \<or> n2 \<noteq> Cs cs)) \<or> | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1192 | (n2 = Th h_n.taker \<and> n1 = Cs cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1193 | by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1194 | thus ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1195 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1196 | assume "n2 = Th h_n.taker \<and> n1 = Cs cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1197 | with h_n.holding_taker | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1198 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1199 | by (unfold s_RAG_def, fold holding_eq, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1200 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1201 | assume h: "(n1, n2) \<in> RAG s \<and> | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1202 | (n1 \<noteq> Cs cs \<or> n2 \<noteq> Th th) \<and> (n1 \<noteq> Th h_n.taker \<or> n2 \<noteq> Cs cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1203 | hence "(n1, n2) \<in> RAG s" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1204 | thus ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1205 | proof(cases rule:in_RAG_E) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1206 | case (waiting th' cs') | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1207 | from h and this(1,2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1208 | have "th' \<noteq> h_n.taker \<or> cs' \<noteq> cs" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1209 | hence "waiting (e#s) th' cs'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1210 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1211 | assume "cs' \<noteq> cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1212 | from waiting_esI1[OF waiting(3) this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1213 | show ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1214 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1215 | assume neq_th': "th' \<noteq> h_n.taker" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1216 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1217 | proof(cases "cs' = cs") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1218 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1219 | from waiting_esI1[OF waiting(3) this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1220 | show ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1221 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1222 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1223 | from h_n.waiting_esI2[OF waiting(3)[unfolded True] neq_th', folded True] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1224 | show ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1225 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1226 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1227 | thus ?thesis using waiting(1,2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1228 | by (unfold s_RAG_def, fold waiting_eq, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1229 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1230 | case (holding th' cs') | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1231 | from h this(1,2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1232 | have "cs' \<noteq> cs \<or> th' \<noteq> th" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1233 | hence "holding (e#s) th' cs'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1234 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1235 | assume "cs' \<noteq> cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1236 | from holding_esI2[OF this holding(3)] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1237 | show ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1238 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1239 | assume "th' \<noteq> th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1240 | from holding_esI1[OF holding(3) this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1241 | show ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1242 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1243 | thus ?thesis using holding(1,2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1244 | by (unfold s_RAG_def, fold holding_eq, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1245 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1246 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1247 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1248 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1249 | interpret h_e: valid_trace_v_e s e th cs | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1250 | by (unfold_locales, insert True, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1251 | from h[unfolded h_e.waiting_set_eq h_e.holding_set_eq] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1252 | have h_s: "(n1, n2) \<in> RAG s" "(n1, n2) \<noteq> (Cs cs, Th th)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1253 | by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1254 | from h_s(1) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1255 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1256 | proof(cases rule:in_RAG_E) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1257 | case (waiting th' cs') | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1258 | from h_e.waiting_esI2[OF this(3)] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1259 | show ?thesis using waiting(1,2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1260 | by (unfold s_RAG_def, fold waiting_eq, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1261 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1262 | case (holding th' cs') | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1263 | with h_s(2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1264 | have "cs' \<noteq> cs \<or> th' \<noteq> th" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1265 | thus ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1266 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1267 | assume neq_cs: "cs' \<noteq> cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1268 | from holding_esI2[OF this holding(3)] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1269 | show ?thesis using holding(1,2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1270 | by (unfold s_RAG_def, fold holding_eq, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1271 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1272 | assume "th' \<noteq> th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1273 | from holding_esI1[OF holding(3) this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1274 | show ?thesis using holding(1,2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1275 | by (unfold s_RAG_def, fold holding_eq, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1276 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1277 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1278 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1279 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1280 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1281 | end | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1282 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1283 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1284 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1285 | context valid_trace | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1286 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1287 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1288 | lemma finite_threads: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1289 | shows "finite (threads s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1290 | using vt by (induct) (auto elim: step.cases) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1291 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1292 | lemma cp_eq_cpreced: "cp s th = cpreced (wq s) s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1293 | unfolding cp_def wq_def | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1294 | apply(induct s rule: schs.induct) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1295 | apply(simp add: Let_def cpreced_initial) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1296 | apply(simp add: Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1297 | apply(simp add: Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1298 | apply(simp add: Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1299 | apply(subst (2) schs.simps) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1300 | apply(simp add: Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1301 | apply(subst (2) schs.simps) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1302 | apply(simp add: Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1303 | done | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1304 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1305 | lemma RAG_target_th: "(Th th, x) \<in> RAG (s::state) \<Longrightarrow> \<exists> cs. x = Cs cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1306 | by (unfold s_RAG_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1307 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1308 | lemma wq_threads: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1309 | assumes h: "th \<in> set (wq s cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1310 | shows "th \<in> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1311 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1312 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1313 | lemma wq_threads: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1314 | assumes h: "th \<in> set (wq s cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1315 | shows "th \<in> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1316 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1317 | from vt and h show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1318 | proof(induct arbitrary: th cs) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1319 | case (vt_cons s e) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1320 | interpret vt_s: valid_trace s | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1321 | using vt_cons(1) by (unfold_locales, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1322 | assume ih: "\<And>th cs. th \<in> set (wq s cs) \<Longrightarrow> th \<in> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1323 | and stp: "step s e" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1324 | and vt: "vt s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1325 | and h: "th \<in> set (wq (e # s) cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1326 | show ?case | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1327 | proof(cases e) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1328 | case (Create th' prio) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1329 | with ih h show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1330 | by (auto simp:wq_def Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1331 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1332 | case (Exit th') | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1333 | with stp ih h show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1334 | apply (auto simp:wq_def Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1335 | apply (ind_cases "step s (Exit th')") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1336 | apply (auto simp:runing_def readys_def s_holding_def s_waiting_def holdents_def | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1337 | s_RAG_def s_holding_def cs_holding_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1338 | done | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1339 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1340 | case (V th' cs') | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1341 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1342 | proof(cases "cs' = cs") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1343 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1344 | with h | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1345 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1346 | apply(unfold wq_def V, auto simp:Let_def V split:prod.splits, fold wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1347 | by (drule_tac ih, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1348 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1349 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1350 | from h | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1351 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1352 | proof(unfold V wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1353 | assume th_in: "th \<in> set (wq_fun (schs (V th' cs' # s)) cs)" (is "th \<in> set ?l") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1354 | show "th \<in> threads (V th' cs' # s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1355 | proof(cases "cs = cs'") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1356 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1357 | hence "?l = wq_fun (schs s) cs" by (simp add:Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1358 | with th_in have " th \<in> set (wq s cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1359 | by (fold wq_def, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1360 | from ih [OF this] show ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1361 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1362 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1363 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1364 | proof(cases "wq_fun (schs s) cs'") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1365 | case Nil | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1366 | with h V show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1367 | apply (auto simp:wq_def Let_def split:if_splits) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1368 | by (fold wq_def, drule_tac ih, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1369 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1370 | case (Cons a rest) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1371 | assume eq_wq: "wq_fun (schs s) cs' = a # rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1372 | with h V show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1373 | apply (auto simp:Let_def wq_def split:if_splits) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1374 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1375 | assume th_in: "th \<in> set (SOME q. distinct q \<and> set q = set rest)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1376 | have "set (SOME q. distinct q \<and> set q = set rest) = set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1377 | proof(rule someI2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1378 | from vt_s.wq_distinct[of cs'] and eq_wq[folded wq_def] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1379 | show "distinct rest \<and> set rest = set rest" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1380 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1381 | show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1382 | by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1383 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1384 | with eq_wq th_in have "th \<in> set (wq_fun (schs s) cs')" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1385 | from ih[OF this[folded wq_def]] show "th \<in> threads s" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1386 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1387 | assume th_in: "th \<in> set (wq_fun (schs s) cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1388 | from ih[OF this[folded wq_def]] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1389 | show "th \<in> threads s" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1390 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1391 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1392 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1393 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1394 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1395 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1396 | case (P th' cs') | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1397 | from h stp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1398 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1399 | apply (unfold P wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1400 | apply (auto simp:Let_def split:if_splits, fold wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1401 | apply (auto intro:ih) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1402 | apply(ind_cases "step s (P th' cs')") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1403 | by (unfold runing_def readys_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1404 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1405 | case (Set thread prio) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1406 | with ih h show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1407 | by (auto simp:wq_def Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1408 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1409 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1410 | case vt_nil | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1411 | thus ?case by (auto simp:wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1412 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1413 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1414 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1415 | lemma dm_RAG_threads: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1416 | assumes in_dom: "(Th th) \<in> Domain (RAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1417 | shows "th \<in> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1418 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1419 | from in_dom obtain n where "(Th th, n) \<in> RAG s" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1420 | moreover from RAG_target_th[OF this] obtain cs where "n = Cs cs" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1421 | ultimately have "(Th th, Cs cs) \<in> RAG s" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1422 | hence "th \<in> set (wq s cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1423 | by (unfold s_RAG_def, auto simp:cs_waiting_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1424 | from wq_threads [OF this] show ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1425 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1426 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1427 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1428 | lemma cp_le: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1429 | assumes th_in: "th \<in> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1430 | shows "cp s th \<le> Max ((\<lambda> th. (preced th s)) ` threads s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1431 | proof(unfold cp_eq_cpreced cpreced_def cs_dependants_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1432 |   show "Max ((\<lambda>th. preced th s) ` ({th} \<union> {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+}))
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1433 | \<le> Max ((\<lambda>th. preced th s) ` threads s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1434 | (is "Max (?f ` ?A) \<le> Max (?f ` ?B)") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1435 | proof(rule Max_f_mono) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1436 |     show "{th} \<union> {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<noteq> {}" by simp
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1437 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1438 | from finite_threads | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1439 | show "finite (threads s)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1440 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1441 | from th_in | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1442 |     show "{th} \<union> {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<subseteq> threads s"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1443 | apply (auto simp:Domain_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1444 | apply (rule_tac dm_RAG_threads) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1445 | apply (unfold trancl_domain [of "RAG s", symmetric]) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1446 | by (unfold cs_RAG_def s_RAG_def, auto simp:Domain_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1447 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1448 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1449 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1450 | lemma le_cp: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1451 | shows "preced th s \<le> cp s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1452 | proof(unfold cp_eq_cpreced preced_def cpreced_def, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1453 | show "Prc (priority th s) (last_set th s) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1454 | \<le> Max (insert (Prc (priority th s) (last_set th s)) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1455 | ((\<lambda>th. Prc (priority th s) (last_set th s)) ` dependants (wq s) th))" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1456 | (is "?l \<le> Max (insert ?l ?A)") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1457 |   proof(cases "?A = {}")
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1458 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1459 | have "finite ?A" (is "finite (?f ` ?B)") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1460 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1461 | have "finite ?B" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1462 | proof- | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1463 |         have "finite {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1464 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1465 | let ?F = "\<lambda> (x, y). the_th x" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1466 |           have "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1467 | apply (auto simp:image_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1468 | by (rule_tac x = "(Th x, Th th)" in bexI, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1469 | moreover have "finite \<dots>" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1470 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1471 | from finite_RAG have "finite (RAG s)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1472 | hence "finite ((RAG (wq s))\<^sup>+)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1473 | apply (unfold finite_trancl) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1474 | by (auto simp: s_RAG_def cs_RAG_def wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1475 | thus ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1476 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1477 | ultimately show ?thesis by (auto intro:finite_subset) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1478 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1479 | thus ?thesis by (simp add:cs_dependants_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1480 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1481 | thus ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1482 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1483 | from Max_insert [OF this False, of ?l] show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1484 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1485 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1486 | thus ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1487 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1488 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1489 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1490 | lemma max_cp_eq: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1491 | shows "Max ((cp s) ` threads s) = Max ((\<lambda> th. (preced th s)) ` threads s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1492 | (is "?l = ?r") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1493 | proof(cases "threads s = {}")
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1494 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1495 | thus ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1496 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1497 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1498 | have "?l \<in> ((cp s) ` threads s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1499 | proof(rule Max_in) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1500 | from finite_threads | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1501 | show "finite (cp s ` threads s)" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1502 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1503 |     from False show "cp s ` threads s \<noteq> {}" by auto
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1504 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1505 | then obtain th | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1506 | where th_in: "th \<in> threads s" and eq_l: "?l = cp s th" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1507 | have "\<dots> \<le> ?r" by (rule cp_le[OF th_in]) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1508 | moreover have "?r \<le> cp s th" (is "Max (?f ` ?A) \<le> cp s th") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1509 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1510 | have "?r \<in> (?f ` ?A)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1511 | proof(rule Max_in) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1512 | from finite_threads | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1513 | show " finite ((\<lambda>th. preced th s) ` threads s)" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1514 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1515 |       from False show " (\<lambda>th. preced th s) ` threads s \<noteq> {}" by auto
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1516 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1517 | then obtain th' where | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1518 | th_in': "th' \<in> ?A " and eq_r: "?r = ?f th'" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1519 | from le_cp [of th'] eq_r | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1520 | have "?r \<le> cp s th'" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1521 | moreover have "\<dots> \<le> cp s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1522 | proof(fold eq_l) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1523 | show " cp s th' \<le> Max (cp s ` threads s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1524 | proof(rule Max_ge) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1525 | from th_in' show "cp s th' \<in> cp s ` threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1526 | by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1527 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1528 | from finite_threads | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1529 | show "finite (cp s ` threads s)" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1530 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1531 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1532 | ultimately show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1533 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1534 | ultimately show ?thesis using eq_l by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1535 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1536 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1537 | lemma max_cp_eq_the_preced: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1538 | shows "Max ((cp s) ` threads s) = Max (the_preced s ` threads s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1539 | using max_cp_eq using the_preced_def by presburger | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1540 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1541 | end | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1542 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1543 | lemma preced_v [simp]: "preced th' (V th cs#s) = preced th' s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1544 | by (unfold preced_def, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1545 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1546 | lemma the_preced_v[simp]: "the_preced (V th cs#s) = the_preced s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1547 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1548 | fix th' | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1549 | show "the_preced (V th cs # s) th' = the_preced s th'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1550 | by (unfold the_preced_def preced_def, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1551 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1552 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1553 | lemma step_RAG_v: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1554 | assumes vt: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1555 | "vt (V th cs#s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1556 | shows " | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1557 | RAG (V th cs # s) = | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1558 |   RAG s - {(Cs cs, Th th)} -
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1559 |   {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1560 |   {(Cs cs, Th th') |th'.  next_th s th cs th'}" (is "?L = ?R")
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1561 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1562 | interpret vt_v: valid_trace_v s "V th cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1563 | using assms step_back_vt by (unfold_locales, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1564 | show ?thesis using vt_v.RAG_es . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1565 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1566 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1567 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1568 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1569 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1570 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1571 | text {* (* ddd *) 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1572 |   The following @{text "step_RAG_v"} lemma charaterizes how @{text "RAG"} is changed
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1573 |   with the happening of @{text "V"}-events:
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1574 | *} | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1575 | lemma step_RAG_v: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1576 | assumes vt: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1577 | "vt (V th cs#s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1578 | shows " | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1579 | RAG (V th cs # s) = | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1580 |   RAG s - {(Cs cs, Th th)} -
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1581 |   {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1582 |   {(Cs cs, Th th') |th'.  next_th s th cs th'}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1583 | apply (insert vt, unfold s_RAG_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1584 | apply (auto split:if_splits list.splits simp:Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1585 | apply (auto elim: step_v_waiting_mono step_v_hold_inv | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1586 | step_v_release step_v_wait_inv | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1587 | step_v_get_hold step_v_release_inv) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1588 | apply (erule_tac step_v_not_wait, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1589 | done | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1590 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1591 | text {* 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1592 |   The following @{text "step_RAG_p"} lemma charaterizes how @{text "RAG"} is changed
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1593 |   with the happening of @{text "P"}-events:
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1594 | *} | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1595 | lemma step_RAG_p: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1596 | "vt (P th cs#s) \<Longrightarrow> | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1597 |   RAG (P th cs # s) =  (if (wq s cs = []) then RAG s \<union> {(Cs cs, Th th)}
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1598 |                                              else RAG s \<union> {(Th th, Cs cs)})"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1599 | apply(simp only: s_RAG_def wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1600 | apply (auto split:list.splits prod.splits simp:Let_def wq_def cs_waiting_def cs_holding_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1601 | apply(case_tac "csa = cs", auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1602 | apply(fold wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1603 | apply(drule_tac step_back_step) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1604 | apply(ind_cases " step s (P (hd (wq s cs)) cs)") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1605 | apply(simp add:s_RAG_def wq_def cs_holding_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1606 | apply(auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1607 | done | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1608 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1609 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1610 | lemma RAG_target_th: "(Th th, x) \<in> RAG (s::state) \<Longrightarrow> \<exists> cs. x = Cs cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1611 | by (unfold s_RAG_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1612 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1613 | context valid_trace | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1614 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1615 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1616 | text {*
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1617 |   The following lemma shows that @{text "RAG"} is acyclic.
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1618 |   The overall structure is by induction on the formation of @{text "vt s"}
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1619 |   and then case analysis on event @{text "e"}, where the non-trivial cases 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1620 |   for those for @{text "V"} and @{text "P"} events.
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1621 | *} | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1622 | lemma acyclic_RAG: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1623 | shows "acyclic (RAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1624 | using vt | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1625 | proof(induct) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1626 | case (vt_cons s e) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1627 | interpret vt_s: valid_trace s using vt_cons(1) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1628 | by (unfold_locales, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1629 | assume ih: "acyclic (RAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1630 | and stp: "step s e" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1631 | and vt: "vt s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1632 | show ?case | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1633 | proof(cases e) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1634 | case (Create th prio) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1635 | with ih | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1636 | show ?thesis by (simp add:RAG_create_unchanged) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1637 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1638 | case (Exit th) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1639 | with ih show ?thesis by (simp add:RAG_exit_unchanged) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1640 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1641 | case (V th cs) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1642 | from V vt stp have vtt: "vt (V th cs#s)" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1643 | from step_RAG_v [OF this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1644 | have eq_de: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1645 | "RAG (e # s) = | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1646 |       RAG s - {(Cs cs, Th th)} - {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1647 |       {(Cs cs, Th th') |th'. next_th s th cs th'}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1648 | (is "?L = (?A - ?B - ?C) \<union> ?D") by (simp add:V) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1649 | from ih have ac: "acyclic (?A - ?B - ?C)" by (auto elim:acyclic_subset) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1650 | from step_back_step [OF vtt] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1651 | have "step s (V th cs)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1652 | thus ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1653 | proof(cases) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1654 | assume "holding s th cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1655 | hence th_in: "th \<in> set (wq s cs)" and | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1656 | eq_hd: "th = hd (wq s cs)" unfolding s_holding_def wq_def by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1657 | then obtain rest where | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1658 | eq_wq: "wq s cs = th#rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1659 | by (cases "wq s cs", auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1660 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1661 | proof(cases "rest = []") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1662 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1663 | let ?th' = "hd (SOME q. distinct q \<and> set q = set rest)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1664 |         from eq_wq False have eq_D: "?D = {(Cs cs, Th ?th')}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1665 | by (unfold next_th_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1666 | let ?E = "(?A - ?B - ?C)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1667 | have "(Th ?th', Cs cs) \<notin> ?E\<^sup>*" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1668 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1669 | assume "(Th ?th', Cs cs) \<in> ?E\<^sup>*" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1670 | hence " (Th ?th', Cs cs) \<in> ?E\<^sup>+" by (simp add: rtrancl_eq_or_trancl) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1671 | from tranclD [OF this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1672 | obtain x where th'_e: "(Th ?th', x) \<in> ?E" by blast | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1673 | hence th_d: "(Th ?th', x) \<in> ?A" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1674 | from RAG_target_th [OF this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1675 | obtain cs' where eq_x: "x = Cs cs'" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1676 | with th_d have "(Th ?th', Cs cs') \<in> ?A" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1677 | hence wt_th': "waiting s ?th' cs'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1678 | unfolding s_RAG_def s_waiting_def cs_waiting_def wq_def by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1679 | hence "cs' = cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1680 | proof(rule vt_s.waiting_unique) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1681 | from eq_wq vt_s.wq_distinct[of cs] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1682 | show "waiting s ?th' cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1683 | apply (unfold s_waiting_def wq_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1684 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1685 | assume hd_in: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1686 | and eq_wq: "wq_fun (schs s) cs = th # rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1687 | have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1688 | proof(rule someI2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1689 | from vt_s.wq_distinct[of cs] and eq_wq | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1690 | show "distinct rest \<and> set rest = set rest" unfolding wq_def by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1691 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1692 | fix x assume "distinct x \<and> set x = set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1693 | with False show "x \<noteq> []" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1694 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1695 | hence "hd (SOME q. distinct q \<and> set q = set rest) \<in> | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1696 | set (SOME q. distinct q \<and> set q = set rest)" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1697 | moreover have "\<dots> = set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1698 | proof(rule someI2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1699 | from vt_s.wq_distinct[of cs] and eq_wq | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1700 | show "distinct rest \<and> set rest = set rest" unfolding wq_def by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1701 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1702 | show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1703 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1704 | moreover note hd_in | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1705 | ultimately show "hd (SOME q. distinct q \<and> set q = set rest) = th" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1706 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1707 | assume hd_in: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1708 | and eq_wq: "wq s cs = hd (SOME q. distinct q \<and> set q = set rest) # rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1709 | have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1710 | proof(rule someI2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1711 | from vt_s.wq_distinct[of cs] and eq_wq | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1712 | show "distinct rest \<and> set rest = set rest" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1713 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1714 | fix x assume "distinct x \<and> set x = set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1715 | with False show "x \<noteq> []" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1716 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1717 | hence "hd (SOME q. distinct q \<and> set q = set rest) \<in> | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1718 | set (SOME q. distinct q \<and> set q = set rest)" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1719 | moreover have "\<dots> = set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1720 | proof(rule someI2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1721 | from vt_s.wq_distinct[of cs] and eq_wq | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1722 | show "distinct rest \<and> set rest = set rest" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1723 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1724 | show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1725 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1726 | moreover note hd_in | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1727 | ultimately show False by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1728 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1729 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1730 | with th'_e eq_x have "(Th ?th', Cs cs) \<in> ?E" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1731 | with False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1732 | show "False" by (auto simp: next_th_def eq_wq) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1733 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1734 | with acyclic_insert[symmetric] and ac | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1735 | and eq_de eq_D show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1736 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1737 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1738 | with eq_wq | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1739 |         have eq_D: "?D = {}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1740 | by (unfold next_th_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1741 | with eq_de ac | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1742 | show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1743 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1744 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1745 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1746 | case (P th cs) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1747 | from P vt stp have vtt: "vt (P th cs#s)" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1748 | from step_RAG_p [OF this] P | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1749 | have "RAG (e # s) = | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1750 |       (if wq s cs = [] then RAG s \<union> {(Cs cs, Th th)} else 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1751 |       RAG s \<union> {(Th th, Cs cs)})" (is "?L = ?R")
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1752 | by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1753 | moreover have "acyclic ?R" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1754 | proof(cases "wq s cs = []") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1755 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1756 |       hence eq_r: "?R =  RAG s \<union> {(Cs cs, Th th)}" by simp
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1757 | have "(Th th, Cs cs) \<notin> (RAG s)\<^sup>*" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1758 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1759 | assume "(Th th, Cs cs) \<in> (RAG s)\<^sup>*" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1760 | hence "(Th th, Cs cs) \<in> (RAG s)\<^sup>+" by (simp add: rtrancl_eq_or_trancl) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1761 | from tranclD2 [OF this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1762 | obtain x where "(x, Cs cs) \<in> RAG s" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1763 | with True show False by (auto simp:s_RAG_def cs_waiting_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1764 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1765 | with acyclic_insert ih eq_r show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1766 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1767 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1768 |       hence eq_r: "?R =  RAG s \<union> {(Th th, Cs cs)}" by simp
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1769 | have "(Cs cs, Th th) \<notin> (RAG s)\<^sup>*" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1770 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1771 | assume "(Cs cs, Th th) \<in> (RAG s)\<^sup>*" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1772 | hence "(Cs cs, Th th) \<in> (RAG s)\<^sup>+" by (simp add: rtrancl_eq_or_trancl) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1773 | moreover from step_back_step [OF vtt] have "step s (P th cs)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1774 | ultimately show False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1775 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1776 | show " \<lbrakk>(Cs cs, Th th) \<in> (RAG s)\<^sup>+; step s (P th cs)\<rbrakk> \<Longrightarrow> False" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1777 | by (ind_cases "step s (P th cs)", simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1778 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1779 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1780 | with acyclic_insert ih eq_r show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1781 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1782 | ultimately show ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1783 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1784 | case (Set thread prio) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1785 | with ih | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1786 | thm RAG_set_unchanged | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1787 | show ?thesis by (simp add:RAG_set_unchanged) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1788 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1789 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1790 | case vt_nil | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1791 | show "acyclic (RAG ([]::state))" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1792 | by (auto simp: s_RAG_def cs_waiting_def | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1793 | cs_holding_def wq_def acyclic_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1794 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1795 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1796 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1797 | lemma finite_RAG: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1798 | shows "finite (RAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1799 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1800 | from vt show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1801 | proof(induct) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1802 | case (vt_cons s e) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1803 | interpret vt_s: valid_trace s using vt_cons(1) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1804 | by (unfold_locales, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1805 | assume ih: "finite (RAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1806 | and stp: "step s e" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1807 | and vt: "vt s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1808 | show ?case | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1809 | proof(cases e) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1810 | case (Create th prio) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1811 | with ih | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1812 | show ?thesis by (simp add:RAG_create_unchanged) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1813 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1814 | case (Exit th) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1815 | with ih show ?thesis by (simp add:RAG_exit_unchanged) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1816 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1817 | case (V th cs) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1818 | from V vt stp have vtt: "vt (V th cs#s)" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1819 | from step_RAG_v [OF this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1820 | have eq_de: "RAG (e # s) = | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1821 |                    RAG s - {(Cs cs, Th th)} - {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1822 |                       {(Cs cs, Th th') |th'. next_th s th cs th'}
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1823 | " | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1824 | (is "?L = (?A - ?B - ?C) \<union> ?D") by (simp add:V) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1825 | moreover from ih have ac: "finite (?A - ?B - ?C)" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1826 | moreover have "finite ?D" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1827 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1828 |         have "?D = {} \<or> (\<exists> a. ?D = {a})" 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1829 | by (unfold next_th_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1830 | thus ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1831 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1832 |           assume h: "?D = {}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1833 | show ?thesis by (unfold h, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1834 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1835 |           assume "\<exists> a. ?D = {a}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1836 | thus ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1837 | by (metis finite.simps) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1838 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1839 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1840 | ultimately show ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1841 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1842 | case (P th cs) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1843 | from P vt stp have vtt: "vt (P th cs#s)" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1844 | from step_RAG_p [OF this] P | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1845 | have "RAG (e # s) = | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1846 |               (if wq s cs = [] then RAG s \<union> {(Cs cs, Th th)} else 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1847 |                                     RAG s \<union> {(Th th, Cs cs)})" (is "?L = ?R")
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1848 | by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1849 | moreover have "finite ?R" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1850 | proof(cases "wq s cs = []") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1851 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1852 |         hence eq_r: "?R =  RAG s \<union> {(Cs cs, Th th)}" by simp
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1853 | with True and ih show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1854 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1855 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1856 |         hence "?R = RAG s \<union> {(Th th, Cs cs)}" by simp
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1857 | with False and ih show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1858 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1859 | ultimately show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1860 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1861 | case (Set thread prio) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1862 | with ih | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1863 | show ?thesis by (simp add:RAG_set_unchanged) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1864 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1865 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1866 | case vt_nil | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1867 | show "finite (RAG ([]::state))" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1868 | by (auto simp: s_RAG_def cs_waiting_def | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1869 | cs_holding_def wq_def acyclic_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1870 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1871 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1872 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1873 | text {* Several useful lemmas *}
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1874 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1875 | lemma wf_dep_converse: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1876 | shows "wf ((RAG s)^-1)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1877 | proof(rule finite_acyclic_wf_converse) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1878 | from finite_RAG | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1879 | show "finite (RAG s)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1880 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1881 | from acyclic_RAG | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1882 | show "acyclic (RAG s)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1883 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1884 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1885 | end | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1886 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1887 | lemma hd_np_in: "x \<in> set l \<Longrightarrow> hd l \<in> set l" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1888 | by (induct l, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1889 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1890 | lemma th_chasing: "(Th th, Cs cs) \<in> RAG (s::state) \<Longrightarrow> \<exists> th'. (Cs cs, Th th') \<in> RAG s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1891 | by (auto simp:s_RAG_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1892 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1893 | context valid_trace | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1894 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1895 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1896 | lemma wq_threads: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1897 | assumes h: "th \<in> set (wq s cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1898 | shows "th \<in> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1899 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1900 | from vt and h show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1901 | proof(induct arbitrary: th cs) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1902 | case (vt_cons s e) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1903 | interpret vt_s: valid_trace s | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1904 | using vt_cons(1) by (unfold_locales, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1905 | assume ih: "\<And>th cs. th \<in> set (wq s cs) \<Longrightarrow> th \<in> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1906 | and stp: "step s e" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1907 | and vt: "vt s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1908 | and h: "th \<in> set (wq (e # s) cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1909 | show ?case | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1910 | proof(cases e) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1911 | case (Create th' prio) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1912 | with ih h show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1913 | by (auto simp:wq_def Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1914 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1915 | case (Exit th') | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1916 | with stp ih h show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1917 | apply (auto simp:wq_def Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1918 | apply (ind_cases "step s (Exit th')") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1919 | apply (auto simp:runing_def readys_def s_holding_def s_waiting_def holdents_def | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1920 | s_RAG_def s_holding_def cs_holding_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1921 | done | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1922 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1923 | case (V th' cs') | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1924 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1925 | proof(cases "cs' = cs") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1926 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1927 | with h | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1928 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1929 | apply(unfold wq_def V, auto simp:Let_def V split:prod.splits, fold wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1930 | by (drule_tac ih, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1931 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1932 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1933 | from h | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1934 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1935 | proof(unfold V wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1936 | assume th_in: "th \<in> set (wq_fun (schs (V th' cs' # s)) cs)" (is "th \<in> set ?l") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1937 | show "th \<in> threads (V th' cs' # s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1938 | proof(cases "cs = cs'") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1939 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1940 | hence "?l = wq_fun (schs s) cs" by (simp add:Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1941 | with th_in have " th \<in> set (wq s cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1942 | by (fold wq_def, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1943 | from ih [OF this] show ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1944 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1945 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1946 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1947 | proof(cases "wq_fun (schs s) cs'") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1948 | case Nil | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1949 | with h V show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1950 | apply (auto simp:wq_def Let_def split:if_splits) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1951 | by (fold wq_def, drule_tac ih, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1952 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1953 | case (Cons a rest) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1954 | assume eq_wq: "wq_fun (schs s) cs' = a # rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1955 | with h V show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1956 | apply (auto simp:Let_def wq_def split:if_splits) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1957 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1958 | assume th_in: "th \<in> set (SOME q. distinct q \<and> set q = set rest)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1959 | have "set (SOME q. distinct q \<and> set q = set rest) = set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1960 | proof(rule someI2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1961 | from vt_s.wq_distinct[of cs'] and eq_wq[folded wq_def] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1962 | show "distinct rest \<and> set rest = set rest" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1963 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1964 | show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1965 | by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1966 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1967 | with eq_wq th_in have "th \<in> set (wq_fun (schs s) cs')" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1968 | from ih[OF this[folded wq_def]] show "th \<in> threads s" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1969 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1970 | assume th_in: "th \<in> set (wq_fun (schs s) cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1971 | from ih[OF this[folded wq_def]] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1972 | show "th \<in> threads s" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1973 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1974 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1975 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1976 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1977 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1978 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1979 | case (P th' cs') | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1980 | from h stp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1981 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1982 | apply (unfold P wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1983 | apply (auto simp:Let_def split:if_splits, fold wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1984 | apply (auto intro:ih) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1985 | apply(ind_cases "step s (P th' cs')") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1986 | by (unfold runing_def readys_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1987 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1988 | case (Set thread prio) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1989 | with ih h show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1990 | by (auto simp:wq_def Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1991 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1992 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1993 | case vt_nil | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1994 | thus ?case by (auto simp:wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1995 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1996 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1997 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1998 | lemma range_in: "\<lbrakk>(Th th) \<in> Range (RAG (s::state))\<rbrakk> \<Longrightarrow> th \<in> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 1999 | apply(unfold s_RAG_def cs_waiting_def cs_holding_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2000 | by (auto intro:wq_threads) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2001 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2002 | lemma readys_v_eq: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2003 | assumes neq_th: "th \<noteq> thread" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2004 | and eq_wq: "wq s cs = thread#rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2005 | and not_in: "th \<notin> set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2006 | shows "(th \<in> readys (V thread cs#s)) = (th \<in> readys s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2007 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2008 | from assms show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2009 | apply (auto simp:readys_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2010 | apply(simp add:s_waiting_def[folded wq_def]) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2011 | apply (erule_tac x = csa in allE) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2012 | apply (simp add:s_waiting_def wq_def Let_def split:if_splits) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2013 | apply (case_tac "csa = cs", simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2014 | apply (erule_tac x = cs in allE) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2015 | apply(auto simp add: s_waiting_def[folded wq_def] Let_def split: list.splits) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2016 | apply(auto simp add: wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2017 | apply (auto simp:s_waiting_def wq_def Let_def split:list.splits) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2018 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2019 | assume th_nin: "th \<notin> set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2020 | and th_in: "th \<in> set (SOME q. distinct q \<and> set q = set rest)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2021 | and eq_wq: "wq_fun (schs s) cs = thread # rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2022 | have "set (SOME q. distinct q \<and> set q = set rest) = set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2023 | proof(rule someI2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2024 | from wq_distinct[of cs, unfolded wq_def] and eq_wq[unfolded wq_def] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2025 | show "distinct rest \<and> set rest = set rest" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2026 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2027 | show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2028 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2029 | with th_nin th_in show False by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2030 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2031 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2032 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2033 | text {* \noindent
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2034 |   The following lemmas shows that: starting from any node in @{text "RAG"}, 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2035 | by chasing out-going edges, it is always possible to reach a node representing a ready | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2036 |   thread. In this lemma, it is the @{text "th'"}.
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2037 | *} | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2038 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2039 | lemma chain_building: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2040 | shows "node \<in> Domain (RAG s) \<longrightarrow> (\<exists> th'. th' \<in> readys s \<and> (node, Th th') \<in> (RAG s)^+)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2041 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2042 | from wf_dep_converse | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2043 | have h: "wf ((RAG s)\<inverse>)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2044 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2045 | proof(induct rule:wf_induct [OF h]) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2046 | fix x | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2047 | assume ih [rule_format]: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2048 | "\<forall>y. (y, x) \<in> (RAG s)\<inverse> \<longrightarrow> | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2049 | y \<in> Domain (RAG s) \<longrightarrow> (\<exists>th'. th' \<in> readys s \<and> (y, Th th') \<in> (RAG s)\<^sup>+)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2050 | show "x \<in> Domain (RAG s) \<longrightarrow> (\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (RAG s)\<^sup>+)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2051 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2052 | assume x_d: "x \<in> Domain (RAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2053 | show "\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (RAG s)\<^sup>+" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2054 | proof(cases x) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2055 | case (Th th) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2056 | from x_d Th obtain cs where x_in: "(Th th, Cs cs) \<in> RAG s" by (auto simp:s_RAG_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2057 | with Th have x_in_r: "(Cs cs, x) \<in> (RAG s)^-1" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2058 | from th_chasing [OF x_in] obtain th' where "(Cs cs, Th th') \<in> RAG s" by blast | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2059 | hence "Cs cs \<in> Domain (RAG s)" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2060 | from ih [OF x_in_r this] obtain th' | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2061 | where th'_ready: " th' \<in> readys s" and cs_in: "(Cs cs, Th th') \<in> (RAG s)\<^sup>+" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2062 | have "(x, Th th') \<in> (RAG s)\<^sup>+" using Th x_in cs_in by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2063 | with th'_ready show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2064 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2065 | case (Cs cs) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2066 | from x_d Cs obtain th' where th'_d: "(Th th', x) \<in> (RAG s)^-1" by (auto simp:s_RAG_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2067 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2068 | proof(cases "th' \<in> readys s") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2069 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2070 | from True and th'_d show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2071 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2072 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2073 | from th'_d and range_in have "th' \<in> threads s" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2074 | with False have "Th th' \<in> Domain (RAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2075 | by (auto simp:readys_def wq_def s_waiting_def s_RAG_def cs_waiting_def Domain_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2076 | from ih [OF th'_d this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2077 | obtain th'' where | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2078 | th''_r: "th'' \<in> readys s" and | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2079 | th''_in: "(Th th', Th th'') \<in> (RAG s)\<^sup>+" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2080 | from th'_d and th''_in | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2081 | have "(x, Th th'') \<in> (RAG s)\<^sup>+" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2082 | with th''_r show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2083 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2084 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2085 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2086 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2087 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2088 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2089 | text {* \noindent
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2090 |   The following is just an instance of @{text "chain_building"}.
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2091 | *} | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2092 | lemma th_chain_to_ready: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2093 | assumes th_in: "th \<in> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2094 | shows "th \<in> readys s \<or> (\<exists> th'. th' \<in> readys s \<and> (Th th, Th th') \<in> (RAG s)^+)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2095 | proof(cases "th \<in> readys s") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2096 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2097 | thus ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2098 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2099 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2100 | from False and th_in have "Th th \<in> Domain (RAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2101 | by (auto simp:readys_def s_waiting_def s_RAG_def wq_def cs_waiting_def Domain_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2102 | from chain_building [rule_format, OF this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2103 | show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2104 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2105 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2106 | end | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2107 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2108 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2109 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2110 | lemma holding_unique: "\<lbrakk>holding (s::state) th1 cs; holding s th2 cs\<rbrakk> \<Longrightarrow> th1 = th2" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2111 | by (unfold s_holding_def cs_holding_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2112 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2113 | context valid_trace | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2114 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2115 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2116 | lemma unique_RAG: "\<lbrakk>(n, n1) \<in> RAG s; (n, n2) \<in> RAG s\<rbrakk> \<Longrightarrow> n1 = n2" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2117 | apply(unfold s_RAG_def, auto, fold waiting_eq holding_eq) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2118 | by(auto elim:waiting_unique holding_unique) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2119 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2120 | end | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2121 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2122 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2123 | lemma trancl_split: "(a, b) \<in> r^+ \<Longrightarrow> \<exists> c. (a, c) \<in> r" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2124 | by (induct rule:trancl_induct, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2125 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2126 | context valid_trace | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2127 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2128 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2129 | lemma dchain_unique: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2130 | assumes th1_d: "(n, Th th1) \<in> (RAG s)^+" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2131 | and th1_r: "th1 \<in> readys s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2132 | and th2_d: "(n, Th th2) \<in> (RAG s)^+" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2133 | and th2_r: "th2 \<in> readys s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2134 | shows "th1 = th2" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2135 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2136 |   { assume neq: "th1 \<noteq> th2"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2137 | hence "Th th1 \<noteq> Th th2" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2138 | from unique_chain [OF _ th1_d th2_d this] and unique_RAG | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2139 | have "(Th th1, Th th2) \<in> (RAG s)\<^sup>+ \<or> (Th th2, Th th1) \<in> (RAG s)\<^sup>+" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2140 | hence "False" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2141 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2142 | assume "(Th th1, Th th2) \<in> (RAG s)\<^sup>+" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2143 | from trancl_split [OF this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2144 | obtain n where dd: "(Th th1, n) \<in> RAG s" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2145 | then obtain cs where eq_n: "n = Cs cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2146 | by (auto simp:s_RAG_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2147 | from dd eq_n have "th1 \<notin> readys s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2148 | by (auto simp:readys_def s_RAG_def wq_def s_waiting_def cs_waiting_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2149 | with th1_r show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2150 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2151 | assume "(Th th2, Th th1) \<in> (RAG s)\<^sup>+" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2152 | from trancl_split [OF this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2153 | obtain n where dd: "(Th th2, n) \<in> RAG s" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2154 | then obtain cs where eq_n: "n = Cs cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2155 | by (auto simp:s_RAG_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2156 | from dd eq_n have "th2 \<notin> readys s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2157 | by (auto simp:readys_def wq_def s_RAG_def s_waiting_def cs_waiting_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2158 | with th2_r show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2159 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2160 | } thus ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2161 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2162 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2163 | end | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2164 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2165 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2166 | lemma step_holdents_p_add: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2167 | assumes vt: "vt (P th cs#s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2168 | and "wq s cs = []" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2169 |   shows "holdents (P th cs#s) th = holdents s th \<union> {cs}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2170 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2171 | from assms show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2172 | unfolding holdents_test step_RAG_p[OF vt] by (auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2173 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2174 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2175 | lemma step_holdents_p_eq: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2176 | assumes vt: "vt (P th cs#s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2177 | and "wq s cs \<noteq> []" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2178 | shows "holdents (P th cs#s) th = holdents s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2179 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2180 | from assms show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2181 | unfolding holdents_test step_RAG_p[OF vt] by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2182 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2183 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2184 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2185 | lemma (in valid_trace) finite_holding : | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2186 | shows "finite (holdents s th)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2187 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2188 | let ?F = "\<lambda> (x, y). the_cs x" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2189 | from finite_RAG | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2190 | have "finite (RAG s)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2191 | hence "finite (?F `(RAG s))" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2192 |   moreover have "{cs . (Cs cs, Th th) \<in> RAG s} \<subseteq> \<dots>" 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2193 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2194 |     { have h: "\<And> a A f. a \<in> A \<Longrightarrow> f a \<in> f ` A" by auto
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2195 | fix x assume "(Cs x, Th th) \<in> RAG s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2196 | hence "?F (Cs x, Th th) \<in> ?F `(RAG s)" by (rule h) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2197 | moreover have "?F (Cs x, Th th) = x" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2198 | ultimately have "x \<in> (\<lambda>(x, y). the_cs x) ` RAG s" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2199 | } thus ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2200 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2201 | ultimately show ?thesis by (unfold holdents_test, auto intro:finite_subset) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2202 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2203 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2204 | lemma cntCS_v_dec: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2205 | assumes vtv: "vt (V thread cs#s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2206 | shows "(cntCS (V thread cs#s) thread + 1) = cntCS s thread" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2207 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2208 | from vtv interpret vt_s: valid_trace s | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2209 | by (cases, unfold_locales, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2210 | from vtv interpret vt_v: valid_trace "V thread cs#s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2211 | by (unfold_locales, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2212 | from step_back_step[OF vtv] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2213 | have cs_in: "cs \<in> holdents s thread" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2214 | apply (cases, unfold holdents_test s_RAG_def, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2215 | by (unfold cs_holding_def s_holding_def wq_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2216 | moreover have cs_not_in: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2217 |     "(holdents (V thread cs#s) thread) = holdents s thread - {cs}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2218 | apply (insert vt_s.wq_distinct[of cs]) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2219 | apply (unfold holdents_test, unfold step_RAG_v[OF vtv], | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2220 | auto simp:next_th_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2221 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2222 | fix rest | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2223 | assume dst: "distinct (rest::thread list)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2224 | and ne: "rest \<noteq> []" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2225 | and hd_ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2226 | moreover have "set (SOME q. distinct q \<and> set q = set rest) = set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2227 | proof(rule someI2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2228 | from dst show "distinct rest \<and> set rest = set rest" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2229 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2230 | show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2231 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2232 | ultimately have "hd (SOME q. distinct q \<and> set q = set rest) \<notin> | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2233 | set (SOME q. distinct q \<and> set q = set rest)" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2234 | moreover have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2235 | proof(rule someI2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2236 | from dst show "distinct rest \<and> set rest = set rest" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2237 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2238 | fix x assume " distinct x \<and> set x = set rest" with ne | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2239 | show "x \<noteq> []" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2240 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2241 | ultimately | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2242 | show "(Cs cs, Th (hd (SOME q. distinct q \<and> set q = set rest))) \<in> RAG s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2243 | by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2244 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2245 | fix rest | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2246 | assume dst: "distinct (rest::thread list)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2247 | and ne: "rest \<noteq> []" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2248 | and hd_ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2249 | moreover have "set (SOME q. distinct q \<and> set q = set rest) = set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2250 | proof(rule someI2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2251 | from dst show "distinct rest \<and> set rest = set rest" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2252 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2253 | show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2254 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2255 | ultimately have "hd (SOME q. distinct q \<and> set q = set rest) \<notin> | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2256 | set (SOME q. distinct q \<and> set q = set rest)" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2257 | moreover have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2258 | proof(rule someI2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2259 | from dst show "distinct rest \<and> set rest = set rest" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2260 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2261 | fix x assume " distinct x \<and> set x = set rest" with ne | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2262 | show "x \<noteq> []" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2263 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2264 | ultimately show "False" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2265 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2266 | ultimately | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2267 | have "holdents s thread = insert cs (holdents (V thread cs#s) thread)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2268 | by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2269 | moreover have "card \<dots> = | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2270 |                     Suc (card ((holdents (V thread cs#s) thread) - {cs}))"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2271 | proof(rule card_insert) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2272 | from vt_v.finite_holding | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2273 | show " finite (holdents (V thread cs # s) thread)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2274 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2275 | moreover from cs_not_in | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2276 | have "cs \<notin> (holdents (V thread cs#s) thread)" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2277 | ultimately show ?thesis by (simp add:cntCS_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2278 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2279 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2280 | lemma count_rec1 [simp]: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2281 | assumes "Q e" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2282 | shows "count Q (e#es) = Suc (count Q es)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2283 | using assms | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2284 | by (unfold count_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2285 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2286 | lemma count_rec2 [simp]: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2287 | assumes "\<not>Q e" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2288 | shows "count Q (e#es) = (count Q es)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2289 | using assms | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2290 | by (unfold count_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2291 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2292 | lemma count_rec3 [simp]: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2293 | shows "count Q [] = 0" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2294 | by (unfold count_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2295 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2296 | lemma cntP_diff_inv: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2297 | assumes "cntP (e#s) th \<noteq> cntP s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2298 | shows "isP e \<and> actor e = th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2299 | proof(cases e) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2300 | case (P th' pty) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2301 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2302 | by (cases "(\<lambda>e. \<exists>cs. e = P th cs) (P th' pty)", | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2303 | insert assms P, auto simp:cntP_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2304 | qed (insert assms, auto simp:cntP_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2305 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2306 | lemma cntV_diff_inv: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2307 | assumes "cntV (e#s) th \<noteq> cntV s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2308 | shows "isV e \<and> actor e = th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2309 | proof(cases e) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2310 | case (V th' pty) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2311 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2312 | by (cases "(\<lambda>e. \<exists>cs. e = V th cs) (V th' pty)", | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2313 | insert assms V, auto simp:cntV_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2314 | qed (insert assms, auto simp:cntV_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2315 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2316 | context valid_trace | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2317 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2318 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2319 | text {* (* ddd *) \noindent
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2320 |   The relationship between @{text "cntP"}, @{text "cntV"} and @{text "cntCS"} 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2321 | of one particular thread. t | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2322 | *} | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2323 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2324 | lemma cnp_cnv_cncs: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2325 | shows "cntP s th = cntV s th + (if (th \<in> readys s \<or> th \<notin> threads s) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2326 | then cntCS s th else cntCS s th + 1)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2327 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2328 | from vt show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2329 | proof(induct arbitrary:th) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2330 | case (vt_cons s e) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2331 | interpret vt_s: valid_trace s using vt_cons(1) by (unfold_locales, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2332 | assume vt: "vt s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2333 | and ih: "\<And>th. cntP s th = cntV s th + | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2334 | (if (th \<in> readys s \<or> th \<notin> threads s) then cntCS s th else cntCS s th + 1)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2335 | and stp: "step s e" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2336 | from stp show ?case | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2337 | proof(cases) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2338 | case (thread_create thread prio) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2339 | assume eq_e: "e = Create thread prio" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2340 | and not_in: "thread \<notin> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2341 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2342 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2343 |         { fix cs 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2344 | assume "thread \<in> set (wq s cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2345 | from vt_s.wq_threads [OF this] have "thread \<in> threads s" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2346 | with not_in have "False" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2347 |         } with eq_e have eq_readys: "readys (e#s) = readys s \<union> {thread}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2348 | by (auto simp:readys_def threads.simps s_waiting_def | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2349 | wq_def cs_waiting_def Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2350 | from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2351 | from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2352 | have eq_cncs: "cntCS (e#s) th = cntCS s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2353 | unfolding cntCS_def holdents_test | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2354 | by (simp add:RAG_create_unchanged eq_e) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2355 |         { assume "th \<noteq> thread"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2356 | with eq_readys eq_e | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2357 | have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2358 | (th \<in> readys (s) \<or> th \<notin> threads (s))" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2359 | by (simp add:threads.simps) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2360 | with eq_cnp eq_cnv eq_cncs ih not_in | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2361 | have ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2362 |         } moreover {
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2363 | assume eq_th: "th = thread" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2364 | with not_in ih have " cntP s th = cntV s th + cntCS s th" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2365 | moreover from eq_th and eq_readys have "th \<in> readys (e#s)" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2366 | moreover note eq_cnp eq_cnv eq_cncs | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2367 | ultimately have ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2368 | } ultimately show ?thesis by blast | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2369 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2370 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2371 | case (thread_exit thread) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2372 | assume eq_e: "e = Exit thread" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2373 | and is_runing: "thread \<in> runing s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2374 |       and no_hold: "holdents s thread = {}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2375 | from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2376 | from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2377 | have eq_cncs: "cntCS (e#s) th = cntCS s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2378 | unfolding cntCS_def holdents_test | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2379 | by (simp add:RAG_exit_unchanged eq_e) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2380 |       { assume "th \<noteq> thread"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2381 | with eq_e | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2382 | have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2383 | (th \<in> readys (s) \<or> th \<notin> threads (s))" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2384 | apply (simp add:threads.simps readys_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2385 | apply (subst s_waiting_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2386 | apply (simp add:Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2387 | apply (subst s_waiting_def, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2388 | done | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2389 | with eq_cnp eq_cnv eq_cncs ih | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2390 | have ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2391 |       } moreover {
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2392 | assume eq_th: "th = thread" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2393 | with ih is_runing have " cntP s th = cntV s th + cntCS s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2394 | by (simp add:runing_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2395 | moreover from eq_th eq_e have "th \<notin> threads (e#s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2396 | by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2397 | moreover note eq_cnp eq_cnv eq_cncs | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2398 | ultimately have ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2399 | } ultimately show ?thesis by blast | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2400 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2401 | case (thread_P thread cs) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2402 | assume eq_e: "e = P thread cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2403 | and is_runing: "thread \<in> runing s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2404 | and no_dep: "(Cs cs, Th thread) \<notin> (RAG s)\<^sup>+" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2405 | from thread_P vt stp ih have vtp: "vt (P thread cs#s)" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2406 | then interpret vt_p: valid_trace "(P thread cs#s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2407 | by (unfold_locales, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2408 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2409 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2410 |         { have hh: "\<And> A B C. (B = C) \<Longrightarrow> (A \<and> B) = (A \<and> C)" by blast
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2411 | assume neq_th: "th \<noteq> thread" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2412 | with eq_e | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2413 | have eq_readys: "(th \<in> readys (e#s)) = (th \<in> readys (s))" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2414 | apply (simp add:readys_def s_waiting_def wq_def Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2415 | apply (rule_tac hh) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2416 | apply (intro iffI allI, clarify) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2417 | apply (erule_tac x = csa in allE, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2418 | apply (subgoal_tac "wq_fun (schs s) cs \<noteq> []", auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2419 | apply (erule_tac x = cs in allE, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2420 | by (case_tac "(wq_fun (schs s) cs)", auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2421 | moreover from neq_th eq_e have "cntCS (e # s) th = cntCS s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2422 | apply (simp add:cntCS_def holdents_test) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2423 | by (unfold step_RAG_p [OF vtp], auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2424 | moreover from eq_e neq_th have "cntP (e # s) th = cntP s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2425 | by (simp add:cntP_def count_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2426 | moreover from eq_e neq_th have "cntV (e#s) th = cntV s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2427 | by (simp add:cntV_def count_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2428 | moreover from eq_e neq_th have "threads (e#s) = threads s" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2429 | moreover note ih [of th] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2430 | ultimately have ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2431 |         } moreover {
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2432 | assume eq_th: "th = thread" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2433 | have ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2434 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2435 | from eq_e eq_th have eq_cnp: "cntP (e # s) th = 1 + (cntP s th)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2436 | by (simp add:cntP_def count_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2437 | from eq_e eq_th have eq_cnv: "cntV (e#s) th = cntV s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2438 | by (simp add:cntV_def count_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2439 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2440 | proof (cases "wq s cs = []") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2441 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2442 | with is_runing | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2443 | have "th \<in> readys (e#s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2444 | apply (unfold eq_e wq_def, unfold readys_def s_RAG_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2445 | apply (simp add: wq_def[symmetric] runing_def eq_th s_waiting_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2446 | by (auto simp:readys_def wq_def Let_def s_waiting_def wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2447 | moreover have "cntCS (e # s) th = 1 + cntCS s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2448 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2449 |                 have "card {csa. csa = cs \<or> (Cs csa, Th thread) \<in> RAG s} =
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2450 |                   Suc (card {cs. (Cs cs, Th thread) \<in> RAG s})" (is "card ?L = Suc (card ?R)")
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2451 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2452 | have "?L = insert cs ?R" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2453 |                   moreover have "card \<dots> = Suc (card (?R - {cs}))" 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2454 | proof(rule card_insert) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2455 | from vt_s.finite_holding [of thread] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2456 |                     show " finite {cs. (Cs cs, Th thread) \<in> RAG s}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2457 | by (unfold holdents_test, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2458 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2459 |                   moreover have "?R - {cs} = ?R"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2460 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2461 | have "cs \<notin> ?R" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2462 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2463 |                       assume "cs \<in> {cs. (Cs cs, Th thread) \<in> RAG s}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2464 | with no_dep show False by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2465 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2466 | thus ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2467 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2468 | ultimately show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2469 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2470 | thus ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2471 | apply (unfold eq_e eq_th cntCS_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2472 | apply (simp add: holdents_test) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2473 | by (unfold step_RAG_p [OF vtp], auto simp:True) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2474 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2475 | moreover from is_runing have "th \<in> readys s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2476 | by (simp add:runing_def eq_th) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2477 | moreover note eq_cnp eq_cnv ih [of th] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2478 | ultimately show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2479 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2480 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2481 | have eq_wq: "wq (e#s) cs = wq s cs @ [th]" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2482 | by (unfold eq_th eq_e wq_def, auto simp:Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2483 | have "th \<notin> readys (e#s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2484 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2485 | assume "th \<in> readys (e#s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2486 | hence "\<forall>cs. \<not> waiting (e # s) th cs" by (simp add:readys_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2487 | from this[rule_format, of cs] have " \<not> waiting (e # s) th cs" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2488 | hence "th \<in> set (wq (e#s) cs) \<Longrightarrow> th = hd (wq (e#s) cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2489 | by (simp add:s_waiting_def wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2490 | moreover from eq_wq have "th \<in> set (wq (e#s) cs)" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2491 | ultimately have "th = hd (wq (e#s) cs)" by blast | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2492 | with eq_wq have "th = hd (wq s cs @ [th])" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2493 | hence "th = hd (wq s cs)" using False by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2494 | with False eq_wq vt_p.wq_distinct [of cs] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2495 | show False by (fold eq_e, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2496 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2497 | moreover from is_runing have "th \<in> threads (e#s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2498 | by (unfold eq_e, auto simp:runing_def readys_def eq_th) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2499 | moreover have "cntCS (e # s) th = cntCS s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2500 | apply (unfold cntCS_def holdents_test eq_e step_RAG_p[OF vtp]) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2501 | by (auto simp:False) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2502 | moreover note eq_cnp eq_cnv ih[of th] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2503 | moreover from is_runing have "th \<in> readys s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2504 | by (simp add:runing_def eq_th) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2505 | ultimately show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2506 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2507 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2508 | } ultimately show ?thesis by blast | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2509 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2510 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2511 | case (thread_V thread cs) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2512 | from assms vt stp ih thread_V have vtv: "vt (V thread cs # s)" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2513 | then interpret vt_v: valid_trace "(V thread cs # s)" by (unfold_locales, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2514 | assume eq_e: "e = V thread cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2515 | and is_runing: "thread \<in> runing s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2516 | and hold: "holding s thread cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2517 | from hold obtain rest | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2518 | where eq_wq: "wq s cs = thread # rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2519 | by (case_tac "wq s cs", auto simp: wq_def s_holding_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2520 | have eq_threads: "threads (e#s) = threads s" by (simp add: eq_e) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2521 | have eq_set: "set (SOME q. distinct q \<and> set q = set rest) = set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2522 | proof(rule someI2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2523 | from vt_v.wq_distinct[of cs] and eq_wq | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2524 | show "distinct rest \<and> set rest = set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2525 | by (metis distinct.simps(2) vt_s.wq_distinct) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2526 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2527 | show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2528 | by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2529 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2530 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2531 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2532 |         { assume eq_th: "th = thread"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2533 | from eq_th have eq_cnp: "cntP (e # s) th = cntP s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2534 | by (unfold eq_e, simp add:cntP_def count_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2535 | moreover from eq_th have eq_cnv: "cntV (e#s) th = 1 + cntV s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2536 | by (unfold eq_e, simp add:cntV_def count_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2537 | moreover from cntCS_v_dec [OF vtv] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2538 | have "cntCS (e # s) thread + 1 = cntCS s thread" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2539 | by (simp add:eq_e) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2540 | moreover from is_runing have rd_before: "thread \<in> readys s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2541 | by (unfold runing_def, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2542 | moreover have "thread \<in> readys (e # s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2543 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2544 | from is_runing | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2545 | have "thread \<in> threads (e#s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2546 | by (unfold eq_e, auto simp:runing_def readys_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2547 | moreover have "\<forall> cs1. \<not> waiting (e#s) thread cs1" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2548 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2549 | fix cs1 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2550 |               { assume eq_cs: "cs1 = cs" 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2551 | have "\<not> waiting (e # s) thread cs1" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2552 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2553 | from eq_wq | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2554 | have "thread \<notin> set (wq (e#s) cs1)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2555 | apply(unfold eq_e wq_def eq_cs s_holding_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2556 | apply (auto simp:Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2557 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2558 | assume "thread \<in> set (SOME q. distinct q \<and> set q = set rest)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2559 | with eq_set have "thread \<in> set rest" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2560 | with vt_v.wq_distinct[of cs] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2561 | and eq_wq show False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2562 | by (metis distinct.simps(2) vt_s.wq_distinct) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2563 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2564 | thus ?thesis by (simp add:wq_def s_waiting_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2565 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2566 |               } moreover {
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2567 | assume neq_cs: "cs1 \<noteq> cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2568 | have "\<not> waiting (e # s) thread cs1" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2569 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2570 | from wq_v_neq [OF neq_cs[symmetric]] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2571 | have "wq (V thread cs # s) cs1 = wq s cs1" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2572 | moreover have "\<not> waiting s thread cs1" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2573 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2574 | from runing_ready and is_runing | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2575 | have "thread \<in> readys s" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2576 | thus ?thesis by (simp add:readys_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2577 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2578 | ultimately show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2579 | by (auto simp:wq_def s_waiting_def eq_e) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2580 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2581 | } ultimately show "\<not> waiting (e # s) thread cs1" by blast | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2582 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2583 | ultimately show ?thesis by (simp add:readys_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2584 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2585 | moreover note eq_th ih | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2586 | ultimately have ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2587 |         } moreover {
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2588 | assume neq_th: "th \<noteq> thread" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2589 | from neq_th eq_e have eq_cnp: "cntP (e # s) th = cntP s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2590 | by (simp add:cntP_def count_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2591 | from neq_th eq_e have eq_cnv: "cntV (e # s) th = cntV s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2592 | by (simp add:cntV_def count_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2593 | have ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2594 | proof(cases "th \<in> set rest") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2595 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2596 | have "(th \<in> readys (e # s)) = (th \<in> readys s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2597 | apply (insert step_back_vt[OF vtv]) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2598 | by (simp add: False eq_e eq_wq neq_th vt_s.readys_v_eq) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2599 | moreover have "cntCS (e#s) th = cntCS s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2600 | apply (insert neq_th, unfold eq_e cntCS_def holdents_test step_RAG_v[OF vtv], auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2601 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2602 |                 have "{csa. (Cs csa, Th th) \<in> RAG s \<or> csa = cs \<and> next_th s thread cs th} =
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2603 |                       {cs. (Cs cs, Th th) \<in> RAG s}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2604 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2605 | from False eq_wq | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2606 | have " next_th s thread cs th \<Longrightarrow> (Cs cs, Th th) \<in> RAG s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2607 | apply (unfold next_th_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2608 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2609 | assume ne: "rest \<noteq> []" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2610 | and ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2611 | and eq_wq: "wq s cs = thread # rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2612 | from eq_set ni have "hd (SOME q. distinct q \<and> set q = set rest) \<notin> | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2613 | set (SOME q. distinct q \<and> set q = set rest) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2614 | " by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2615 | moreover have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2616 | proof(rule someI2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2617 | from vt_s.wq_distinct[ of cs] and eq_wq | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2618 | show "distinct rest \<and> set rest = set rest" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2619 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2620 | fix x assume "distinct x \<and> set x = set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2621 | with ne show "x \<noteq> []" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2622 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2623 | ultimately show | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2624 | "(Cs cs, Th (hd (SOME q. distinct q \<and> set q = set rest))) \<in> RAG s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2625 | by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2626 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2627 | thus ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2628 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2629 |                 thus "card {csa. (Cs csa, Th th) \<in> RAG s \<or> csa = cs \<and> next_th s thread cs th} =
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2630 |                              card {cs. (Cs cs, Th th) \<in> RAG s}" by simp 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2631 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2632 | moreover note ih eq_cnp eq_cnv eq_threads | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2633 | ultimately show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2634 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2635 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2636 | assume th_in: "th \<in> set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2637 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2638 | proof(cases "next_th s thread cs th") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2639 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2640 | with eq_wq and th_in have | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2641 | neq_hd: "th \<noteq> hd (SOME q. distinct q \<and> set q = set rest)" (is "th \<noteq> hd ?rest") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2642 | by (auto simp:next_th_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2643 | have "(th \<in> readys (e # s)) = (th \<in> readys s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2644 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2645 | from eq_wq and th_in | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2646 | have "\<not> th \<in> readys s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2647 | apply (auto simp:readys_def s_waiting_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2648 | apply (rule_tac x = cs in exI, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2649 | by (insert vt_s.wq_distinct[of cs], auto simp add: wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2650 | moreover | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2651 | from eq_wq and th_in and neq_hd | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2652 | have "\<not> (th \<in> readys (e # s))" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2653 | apply (auto simp:readys_def s_waiting_def eq_e wq_def Let_def split:list.splits) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2654 | by (rule_tac x = cs in exI, auto simp:eq_set) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2655 | ultimately show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2656 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2657 | moreover have "cntCS (e#s) th = cntCS s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2658 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2659 | from eq_wq and th_in and neq_hd | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2660 | have "(holdents (e # s) th) = (holdents s th)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2661 | apply (unfold eq_e step_RAG_v[OF vtv], | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2662 | auto simp:next_th_def eq_set s_RAG_def holdents_test wq_def | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2663 | Let_def cs_holding_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2664 | by (insert vt_s.wq_distinct[of cs], auto simp:wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2665 | thus ?thesis by (simp add:cntCS_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2666 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2667 | moreover note ih eq_cnp eq_cnv eq_threads | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2668 | ultimately show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2669 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2670 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2671 | let ?rest = " (SOME q. distinct q \<and> set q = set rest)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2672 | let ?t = "hd ?rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2673 | from True eq_wq th_in neq_th | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2674 | have "th \<in> readys (e # s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2675 | apply (auto simp:eq_e readys_def s_waiting_def wq_def | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2676 | Let_def next_th_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2677 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2678 | assume eq_wq: "wq_fun (schs s) cs = thread # rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2679 | and t_in: "?t \<in> set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2680 | show "?t \<in> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2681 | proof(rule vt_s.wq_threads) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2682 | from eq_wq and t_in | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2683 | show "?t \<in> set (wq s cs)" by (auto simp:wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2684 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2685 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2686 | fix csa | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2687 | assume eq_wq: "wq_fun (schs s) cs = thread # rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2688 | and t_in: "?t \<in> set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2689 | and neq_cs: "csa \<noteq> cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2690 | and t_in': "?t \<in> set (wq_fun (schs s) csa)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2691 | show "?t = hd (wq_fun (schs s) csa)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2692 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2693 |                   { assume neq_hd': "?t \<noteq> hd (wq_fun (schs s) csa)"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2694 | from vt_s.wq_distinct[of cs] and | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2695 | eq_wq[folded wq_def] and t_in eq_wq | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2696 | have "?t \<noteq> thread" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2697 | with eq_wq and t_in | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2698 | have w1: "waiting s ?t cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2699 | by (auto simp:s_waiting_def wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2700 | from t_in' neq_hd' | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2701 | have w2: "waiting s ?t csa" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2702 | by (auto simp:s_waiting_def wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2703 | from vt_s.waiting_unique[OF w1 w2] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2704 | and neq_cs have "False" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2705 | } thus ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2706 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2707 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2708 | moreover have "cntP s th = cntV s th + cntCS s th + 1" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2709 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2710 | have "th \<notin> readys s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2711 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2712 | from True eq_wq neq_th th_in | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2713 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2714 | apply (unfold readys_def s_waiting_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2715 | by (rule_tac x = cs in exI, auto simp add: wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2716 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2717 | moreover have "th \<in> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2718 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2719 | from th_in eq_wq | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2720 | have "th \<in> set (wq s cs)" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2721 | from vt_s.wq_threads [OF this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2722 | show ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2723 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2724 | ultimately show ?thesis using ih by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2725 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2726 | moreover from True neq_th have "cntCS (e # s) th = 1 + cntCS s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2727 | apply (unfold cntCS_def holdents_test eq_e step_RAG_v[OF vtv], auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2728 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2729 |                 show "card {csa. (Cs csa, Th th) \<in> RAG s \<or> csa = cs} =
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2730 |                                Suc (card {cs. (Cs cs, Th th) \<in> RAG s})"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2731 | (is "card ?A = Suc (card ?B)") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2732 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2733 | have "?A = insert cs ?B" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2734 | hence "card ?A = card (insert cs ?B)" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2735 | also have "\<dots> = Suc (card ?B)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2736 | proof(rule card_insert_disjoint) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2737 | have "?B \<subseteq> ((\<lambda> (x, y). the_cs x) ` RAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2738 | apply (auto simp:image_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2739 | by (rule_tac x = "(Cs x, Th th)" in bexI, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2740 | with vt_s.finite_RAG | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2741 |                     show "finite {cs. (Cs cs, Th th) \<in> RAG s}" by (auto intro:finite_subset)
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2742 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2743 |                     show "cs \<notin> {cs. (Cs cs, Th th) \<in> RAG s}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2744 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2745 |                       assume "cs \<in> {cs. (Cs cs, Th th) \<in> RAG s}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2746 | hence "(Cs cs, Th th) \<in> RAG s" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2747 | with True neq_th eq_wq show False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2748 | by (auto simp:next_th_def s_RAG_def cs_holding_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2749 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2750 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2751 | finally show ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2752 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2753 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2754 | moreover note eq_cnp eq_cnv | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2755 | ultimately show ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2756 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2757 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2758 | } ultimately show ?thesis by blast | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2759 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2760 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2761 | case (thread_set thread prio) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2762 | assume eq_e: "e = Set thread prio" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2763 | and is_runing: "thread \<in> runing s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2764 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2765 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2766 | from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2767 | from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2768 | have eq_cncs: "cntCS (e#s) th = cntCS s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2769 | unfolding cntCS_def holdents_test | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2770 | by (simp add:RAG_set_unchanged eq_e) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2771 | from eq_e have eq_readys: "readys (e#s) = readys s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2772 | by (simp add:readys_def cs_waiting_def s_waiting_def wq_def, | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2773 | auto simp:Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2774 |         { assume "th \<noteq> thread"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2775 | with eq_readys eq_e | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2776 | have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2777 | (th \<in> readys (s) \<or> th \<notin> threads (s))" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2778 | by (simp add:threads.simps) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2779 | with eq_cnp eq_cnv eq_cncs ih is_runing | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2780 | have ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2781 |         } moreover {
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2782 | assume eq_th: "th = thread" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2783 | with is_runing ih have " cntP s th = cntV s th + cntCS s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2784 | by (unfold runing_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2785 | moreover from eq_th and eq_readys is_runing have "th \<in> readys (e#s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2786 | by (simp add:runing_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2787 | moreover note eq_cnp eq_cnv eq_cncs | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2788 | ultimately have ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2789 | } ultimately show ?thesis by blast | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2790 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2791 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2792 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2793 | case vt_nil | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2794 | show ?case | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2795 | by (unfold cntP_def cntV_def cntCS_def, | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2796 | auto simp:count_def holdents_test s_RAG_def wq_def cs_holding_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2797 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2798 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2799 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2800 | lemma not_thread_cncs: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2801 | assumes not_in: "th \<notin> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2802 | shows "cntCS s th = 0" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2803 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2804 | from vt not_in show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2805 | proof(induct arbitrary:th) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2806 | case (vt_cons s e th) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2807 | interpret vt_s: valid_trace s using vt_cons(1) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2808 | by (unfold_locales, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2809 | assume vt: "vt s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2810 | and ih: "\<And>th. th \<notin> threads s \<Longrightarrow> cntCS s th = 0" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2811 | and stp: "step s e" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2812 | and not_in: "th \<notin> threads (e # s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2813 | from stp show ?case | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2814 | proof(cases) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2815 | case (thread_create thread prio) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2816 | assume eq_e: "e = Create thread prio" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2817 | and not_in': "thread \<notin> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2818 | have "cntCS (e # s) th = cntCS s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2819 | apply (unfold eq_e cntCS_def holdents_test) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2820 | by (simp add:RAG_create_unchanged) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2821 | moreover have "th \<notin> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2822 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2823 | from not_in eq_e show ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2824 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2825 | moreover note ih ultimately show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2826 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2827 | case (thread_exit thread) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2828 | assume eq_e: "e = Exit thread" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2829 |       and nh: "holdents s thread = {}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2830 | have eq_cns: "cntCS (e # s) th = cntCS s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2831 | apply (unfold eq_e cntCS_def holdents_test) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2832 | by (simp add:RAG_exit_unchanged) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2833 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2834 | proof(cases "th = thread") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2835 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2836 | have "cntCS s th = 0" by (unfold cntCS_def, auto simp:nh True) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2837 | with eq_cns show ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2838 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2839 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2840 | with not_in and eq_e | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2841 | have "th \<notin> threads s" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2842 | from ih[OF this] and eq_cns show ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2843 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2844 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2845 | case (thread_P thread cs) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2846 | assume eq_e: "e = P thread cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2847 | and is_runing: "thread \<in> runing s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2848 | from assms thread_P ih vt stp thread_P have vtp: "vt (P thread cs#s)" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2849 | have neq_th: "th \<noteq> thread" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2850 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2851 | from not_in eq_e have "th \<notin> threads s" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2852 | moreover from is_runing have "thread \<in> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2853 | by (simp add:runing_def readys_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2854 | ultimately show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2855 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2856 | hence "cntCS (e # s) th = cntCS s th " | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2857 | apply (unfold cntCS_def holdents_test eq_e) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2858 | by (unfold step_RAG_p[OF vtp], auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2859 | moreover have "cntCS s th = 0" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2860 | proof(rule ih) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2861 | from not_in eq_e show "th \<notin> threads s" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2862 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2863 | ultimately show ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2864 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2865 | case (thread_V thread cs) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2866 | assume eq_e: "e = V thread cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2867 | and is_runing: "thread \<in> runing s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2868 | and hold: "holding s thread cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2869 | have neq_th: "th \<noteq> thread" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2870 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2871 | from not_in eq_e have "th \<notin> threads s" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2872 | moreover from is_runing have "thread \<in> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2873 | by (simp add:runing_def readys_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2874 | ultimately show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2875 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2876 | from assms thread_V vt stp ih | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2877 | have vtv: "vt (V thread cs#s)" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2878 | then interpret vt_v: valid_trace "(V thread cs#s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2879 | by (unfold_locales, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2880 | from hold obtain rest | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2881 | where eq_wq: "wq s cs = thread # rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2882 | by (case_tac "wq s cs", auto simp: wq_def s_holding_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2883 | from not_in eq_e eq_wq | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2884 | have "\<not> next_th s thread cs th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2885 | apply (auto simp:next_th_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2886 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2887 | assume ne: "rest \<noteq> []" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2888 | and ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> threads s" (is "?t \<notin> threads s") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2889 | have "?t \<in> set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2890 | proof(rule someI2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2891 | from vt_v.wq_distinct[of cs] and eq_wq | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2892 | show "distinct rest \<and> set rest = set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2893 | by (metis distinct.simps(2) vt_s.wq_distinct) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2894 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2895 | fix x assume "distinct x \<and> set x = set rest" with ne | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2896 | show "hd x \<in> set rest" by (cases x, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2897 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2898 | with eq_wq have "?t \<in> set (wq s cs)" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2899 | from vt_s.wq_threads[OF this] and ni | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2900 | show False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2901 | using `hd (SOME q. distinct q \<and> set q = set rest) \<in> set (wq s cs)` | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2902 | ni vt_s.wq_threads by blast | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2903 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2904 | moreover note neq_th eq_wq | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2905 | ultimately have "cntCS (e # s) th = cntCS s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2906 | by (unfold eq_e cntCS_def holdents_test step_RAG_v[OF vtv], auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2907 | moreover have "cntCS s th = 0" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2908 | proof(rule ih) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2909 | from not_in eq_e show "th \<notin> threads s" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2910 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2911 | ultimately show ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2912 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2913 | case (thread_set thread prio) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2914 | print_facts | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2915 | assume eq_e: "e = Set thread prio" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2916 | and is_runing: "thread \<in> runing s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2917 | from not_in and eq_e have "th \<notin> threads s" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2918 | from ih [OF this] and eq_e | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2919 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2920 | apply (unfold eq_e cntCS_def holdents_test) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2921 | by (simp add:RAG_set_unchanged) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2922 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2923 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2924 | case vt_nil | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2925 | show ?case | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2926 | by (unfold cntCS_def, | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2927 | auto simp:count_def holdents_test s_RAG_def wq_def cs_holding_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2928 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2929 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2930 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2931 | end | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2932 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2933 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2934 | context valid_trace | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2935 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2936 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2937 | lemma dm_RAG_threads: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2938 | assumes in_dom: "(Th th) \<in> Domain (RAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2939 | shows "th \<in> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2940 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2941 | from in_dom obtain n where "(Th th, n) \<in> RAG s" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2942 | moreover from RAG_target_th[OF this] obtain cs where "n = Cs cs" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2943 | ultimately have "(Th th, Cs cs) \<in> RAG s" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2944 | hence "th \<in> set (wq s cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2945 | by (unfold s_RAG_def, auto simp:cs_waiting_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2946 | from wq_threads [OF this] show ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2947 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2948 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2949 | end | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2950 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2951 | lemma cp_eq_cpreced: "cp s th = cpreced (wq s) s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2952 | unfolding cp_def wq_def | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2953 | apply(induct s rule: schs.induct) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2954 | thm cpreced_initial | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2955 | apply(simp add: Let_def cpreced_initial) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2956 | apply(simp add: Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2957 | apply(simp add: Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2958 | apply(simp add: Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2959 | apply(subst (2) schs.simps) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2960 | apply(simp add: Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2961 | apply(subst (2) schs.simps) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2962 | apply(simp add: Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2963 | done | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2964 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2965 | context valid_trace | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2966 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2967 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2968 | lemma runing_unique: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2969 | assumes runing_1: "th1 \<in> runing s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2970 | and runing_2: "th2 \<in> runing s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2971 | shows "th1 = th2" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2972 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2973 | from runing_1 and runing_2 have "cp s th1 = cp s th2" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2974 | unfolding runing_def | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2975 | apply(simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2976 | done | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2977 |   hence eq_max: "Max ((\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1)) =
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2978 |                  Max ((\<lambda>th. preced th s) ` ({th2} \<union> dependants (wq s) th2))"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2979 | (is "Max (?f ` ?A) = Max (?f ` ?B)") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2980 | unfolding cp_eq_cpreced | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2981 | unfolding cpreced_def . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2982 | obtain th1' where th1_in: "th1' \<in> ?A" and eq_f_th1: "?f th1' = Max (?f ` ?A)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2983 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2984 | have h1: "finite (?f ` ?A)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2985 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2986 | have "finite ?A" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2987 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2988 | have "finite (dependants (wq s) th1)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2989 | proof- | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2990 |           have "finite {th'. (Th th', Th th1) \<in> (RAG (wq s))\<^sup>+}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2991 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2992 | let ?F = "\<lambda> (x, y). the_th x" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2993 |             have "{th'. (Th th', Th th1) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2994 | apply (auto simp:image_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2995 | by (rule_tac x = "(Th x, Th th1)" in bexI, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2996 | moreover have "finite \<dots>" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2997 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2998 | from finite_RAG have "finite (RAG s)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 2999 | hence "finite ((RAG (wq s))\<^sup>+)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3000 | apply (unfold finite_trancl) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3001 | by (auto simp: s_RAG_def cs_RAG_def wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3002 | thus ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3003 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3004 | ultimately show ?thesis by (auto intro:finite_subset) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3005 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3006 | thus ?thesis by (simp add:cs_dependants_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3007 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3008 | thus ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3009 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3010 | thus ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3011 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3012 |     moreover have h2: "(?f ` ?A) \<noteq> {}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3013 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3014 |       have "?A \<noteq> {}" by simp
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3015 | thus ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3016 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3017 | from Max_in [OF h1 h2] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3018 | have "Max (?f ` ?A) \<in> (?f ` ?A)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3019 | thus ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3020 | thm cpreced_def | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3021 | unfolding cpreced_def[symmetric] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3022 | unfolding cp_eq_cpreced[symmetric] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3023 | unfolding cpreced_def | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3024 | using that[intro] by (auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3025 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3026 | obtain th2' where th2_in: "th2' \<in> ?B" and eq_f_th2: "?f th2' = Max (?f ` ?B)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3027 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3028 | have h1: "finite (?f ` ?B)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3029 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3030 | have "finite ?B" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3031 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3032 | have "finite (dependants (wq s) th2)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3033 | proof- | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3034 |           have "finite {th'. (Th th', Th th2) \<in> (RAG (wq s))\<^sup>+}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3035 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3036 | let ?F = "\<lambda> (x, y). the_th x" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3037 |             have "{th'. (Th th', Th th2) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3038 | apply (auto simp:image_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3039 | by (rule_tac x = "(Th x, Th th2)" in bexI, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3040 | moreover have "finite \<dots>" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3041 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3042 | from finite_RAG have "finite (RAG s)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3043 | hence "finite ((RAG (wq s))\<^sup>+)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3044 | apply (unfold finite_trancl) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3045 | by (auto simp: s_RAG_def cs_RAG_def wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3046 | thus ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3047 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3048 | ultimately show ?thesis by (auto intro:finite_subset) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3049 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3050 | thus ?thesis by (simp add:cs_dependants_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3051 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3052 | thus ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3053 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3054 | thus ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3055 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3056 |     moreover have h2: "(?f ` ?B) \<noteq> {}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3057 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3058 |       have "?B \<noteq> {}" by simp
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3059 | thus ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3060 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3061 | from Max_in [OF h1 h2] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3062 | have "Max (?f ` ?B) \<in> (?f ` ?B)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3063 | thus ?thesis by (auto intro:that) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3064 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3065 | from eq_f_th1 eq_f_th2 eq_max | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3066 | have eq_preced: "preced th1' s = preced th2' s" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3067 | hence eq_th12: "th1' = th2'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3068 | proof (rule preced_unique) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3069 | from th1_in have "th1' = th1 \<or> (th1' \<in> dependants (wq s) th1)" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3070 | thus "th1' \<in> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3071 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3072 | assume "th1' \<in> dependants (wq s) th1" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3073 | hence "(Th th1') \<in> Domain ((RAG s)^+)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3074 | apply (unfold cs_dependants_def cs_RAG_def s_RAG_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3075 | by (auto simp:Domain_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3076 | hence "(Th th1') \<in> Domain (RAG s)" by (simp add:trancl_domain) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3077 | from dm_RAG_threads[OF this] show ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3078 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3079 | assume "th1' = th1" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3080 | with runing_1 show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3081 | by (unfold runing_def readys_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3082 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3083 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3084 | from th2_in have "th2' = th2 \<or> (th2' \<in> dependants (wq s) th2)" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3085 | thus "th2' \<in> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3086 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3087 | assume "th2' \<in> dependants (wq s) th2" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3088 | hence "(Th th2') \<in> Domain ((RAG s)^+)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3089 | apply (unfold cs_dependants_def cs_RAG_def s_RAG_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3090 | by (auto simp:Domain_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3091 | hence "(Th th2') \<in> Domain (RAG s)" by (simp add:trancl_domain) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3092 | from dm_RAG_threads[OF this] show ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3093 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3094 | assume "th2' = th2" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3095 | with runing_2 show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3096 | by (unfold runing_def readys_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3097 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3098 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3099 | from th1_in have "th1' = th1 \<or> th1' \<in> dependants (wq s) th1" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3100 | thus ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3101 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3102 | assume eq_th': "th1' = th1" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3103 | from th2_in have "th2' = th2 \<or> th2' \<in> dependants (wq s) th2" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3104 | thus ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3105 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3106 | assume "th2' = th2" thus ?thesis using eq_th' eq_th12 by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3107 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3108 | assume "th2' \<in> dependants (wq s) th2" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3109 | with eq_th12 eq_th' have "th1 \<in> dependants (wq s) th2" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3110 | hence "(Th th1, Th th2) \<in> (RAG s)^+" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3111 | by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3112 | hence "Th th1 \<in> Domain ((RAG s)^+)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3113 | apply (unfold cs_dependants_def cs_RAG_def s_RAG_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3114 | by (auto simp:Domain_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3115 | hence "Th th1 \<in> Domain (RAG s)" by (simp add:trancl_domain) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3116 | then obtain n where d: "(Th th1, n) \<in> RAG s" by (auto simp:Domain_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3117 | from RAG_target_th [OF this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3118 | obtain cs' where "n = Cs cs'" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3119 | with d have "(Th th1, Cs cs') \<in> RAG s" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3120 | with runing_1 have "False" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3121 | apply (unfold runing_def readys_def s_RAG_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3122 | by (auto simp:waiting_eq) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3123 | thus ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3124 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3125 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3126 | assume th1'_in: "th1' \<in> dependants (wq s) th1" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3127 | from th2_in have "th2' = th2 \<or> th2' \<in> dependants (wq s) th2" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3128 | thus ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3129 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3130 | assume "th2' = th2" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3131 | with th1'_in eq_th12 have "th2 \<in> dependants (wq s) th1" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3132 | hence "(Th th2, Th th1) \<in> (RAG s)^+" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3133 | by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3134 | hence "Th th2 \<in> Domain ((RAG s)^+)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3135 | apply (unfold cs_dependants_def cs_RAG_def s_RAG_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3136 | by (auto simp:Domain_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3137 | hence "Th th2 \<in> Domain (RAG s)" by (simp add:trancl_domain) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3138 | then obtain n where d: "(Th th2, n) \<in> RAG s" by (auto simp:Domain_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3139 | from RAG_target_th [OF this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3140 | obtain cs' where "n = Cs cs'" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3141 | with d have "(Th th2, Cs cs') \<in> RAG s" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3142 | with runing_2 have "False" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3143 | apply (unfold runing_def readys_def s_RAG_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3144 | by (auto simp:waiting_eq) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3145 | thus ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3146 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3147 | assume "th2' \<in> dependants (wq s) th2" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3148 | with eq_th12 have "th1' \<in> dependants (wq s) th2" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3149 | hence h1: "(Th th1', Th th2) \<in> (RAG s)^+" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3150 | by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3151 | from th1'_in have h2: "(Th th1', Th th1) \<in> (RAG s)^+" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3152 | by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3153 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3154 | proof(rule dchain_unique[OF h1 _ h2, symmetric]) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3155 | from runing_1 show "th1 \<in> readys s" by (simp add:runing_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3156 | from runing_2 show "th2 \<in> readys s" by (simp add:runing_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3157 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3158 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3159 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3160 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3161 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3162 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3163 | lemma "card (runing s) \<le> 1" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3164 | apply(subgoal_tac "finite (runing s)") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3165 | prefer 2 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3166 | apply (metis finite_nat_set_iff_bounded lessI runing_unique) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3167 | apply(rule ccontr) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3168 | apply(simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3169 | apply(case_tac "Suc (Suc 0) \<le> card (runing s)") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3170 | apply(subst (asm) card_le_Suc_iff) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3171 | apply(simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3172 | apply(auto)[1] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3173 | apply (metis insertCI runing_unique) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3174 | apply(auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3175 | done | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3176 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3177 | end | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3178 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3179 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3180 | lemma create_pre: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3181 | assumes stp: "step s e" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3182 | and not_in: "th \<notin> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3183 | and is_in: "th \<in> threads (e#s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3184 | obtains prio where "e = Create th prio" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3185 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3186 | from assms | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3187 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3188 | proof(cases) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3189 | case (thread_create thread prio) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3190 | with is_in not_in have "e = Create th prio" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3191 | from that[OF this] show ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3192 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3193 | case (thread_exit thread) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3194 | with assms show ?thesis by (auto intro!:that) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3195 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3196 | case (thread_P thread) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3197 | with assms show ?thesis by (auto intro!:that) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3198 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3199 | case (thread_V thread) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3200 | with assms show ?thesis by (auto intro!:that) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3201 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3202 | case (thread_set thread) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3203 | with assms show ?thesis by (auto intro!:that) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3204 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3205 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3206 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3207 | context valid_trace | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3208 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3209 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3210 | lemma cnp_cnv_eq: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3211 | assumes "th \<notin> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3212 | shows "cntP s th = cntV s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3213 | using assms | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3214 | using cnp_cnv_cncs not_thread_cncs by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3215 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3216 | end | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3217 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3218 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3219 | lemma eq_RAG: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3220 | "RAG (wq s) = RAG s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3221 | by (unfold cs_RAG_def s_RAG_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3222 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3223 | context valid_trace | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3224 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3225 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3226 | lemma count_eq_dependants: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3227 | assumes eq_pv: "cntP s th = cntV s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3228 |   shows "dependants (wq s) th = {}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3229 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3230 | from cnp_cnv_cncs and eq_pv | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3231 | have "cntCS s th = 0" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3232 | by (auto split:if_splits) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3233 |   moreover have "finite {cs. (Cs cs, Th th) \<in> RAG s}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3234 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3235 | from finite_holding[of th] show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3236 | by (simp add:holdents_test) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3237 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3238 |   ultimately have h: "{cs. (Cs cs, Th th) \<in> RAG s} = {}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3239 | by (unfold cntCS_def holdents_test cs_dependants_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3240 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3241 | proof(unfold cs_dependants_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3242 |     { assume "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<noteq> {}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3243 | then obtain th' where "(Th th', Th th) \<in> (RAG (wq s))\<^sup>+" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3244 | hence "False" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3245 | proof(cases) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3246 | assume "(Th th', Th th) \<in> RAG (wq s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3247 | thus "False" by (auto simp:cs_RAG_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3248 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3249 | fix c | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3250 | assume "(c, Th th) \<in> RAG (wq s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3251 | with h and eq_RAG show "False" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3252 | by (cases c, auto simp:cs_RAG_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3253 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3254 |     } thus "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} = {}" by auto
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3255 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3256 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3257 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3258 | lemma dependants_threads: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3259 | shows "dependants (wq s) th \<subseteq> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3260 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3261 |   { fix th th'
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3262 |     assume h: "th \<in> {th'a. (Th th'a, Th th') \<in> (RAG (wq s))\<^sup>+}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3263 | have "Th th \<in> Domain (RAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3264 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3265 | from h obtain th' where "(Th th, Th th') \<in> (RAG (wq s))\<^sup>+" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3266 | hence "(Th th) \<in> Domain ( (RAG (wq s))\<^sup>+)" by (auto simp:Domain_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3267 | with trancl_domain have "(Th th) \<in> Domain (RAG (wq s))" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3268 | thus ?thesis using eq_RAG by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3269 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3270 | from dm_RAG_threads[OF this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3271 | have "th \<in> threads s" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3272 | } note hh = this | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3273 | fix th1 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3274 | assume "th1 \<in> dependants (wq s) th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3275 |   hence "th1 \<in> {th'a. (Th th'a, Th th) \<in> (RAG (wq s))\<^sup>+}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3276 | by (unfold cs_dependants_def, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3277 | from hh [OF this] show "th1 \<in> threads s" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3278 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3279 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3280 | lemma finite_threads: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3281 | shows "finite (threads s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3282 | using vt by (induct) (auto elim: step.cases) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3283 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3284 | end | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3285 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3286 | lemma Max_f_mono: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3287 | assumes seq: "A \<subseteq> B" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3288 |   and np: "A \<noteq> {}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3289 | and fnt: "finite B" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3290 | shows "Max (f ` A) \<le> Max (f ` B)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3291 | proof(rule Max_mono) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3292 | from seq show "f ` A \<subseteq> f ` B" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3293 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3294 |   from np show "f ` A \<noteq> {}" by auto
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3295 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3296 | from fnt and seq show "finite (f ` B)" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3297 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3298 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3299 | context valid_trace | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3300 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3301 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3302 | lemma cp_le: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3303 | assumes th_in: "th \<in> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3304 | shows "cp s th \<le> Max ((\<lambda> th. (preced th s)) ` threads s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3305 | proof(unfold cp_eq_cpreced cpreced_def cs_dependants_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3306 |   show "Max ((\<lambda>th. preced th s) ` ({th} \<union> {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+}))
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3307 | \<le> Max ((\<lambda>th. preced th s) ` threads s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3308 | (is "Max (?f ` ?A) \<le> Max (?f ` ?B)") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3309 | proof(rule Max_f_mono) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3310 |     show "{th} \<union> {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<noteq> {}" by simp
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3311 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3312 | from finite_threads | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3313 | show "finite (threads s)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3314 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3315 | from th_in | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3316 |     show "{th} \<union> {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<subseteq> threads s"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3317 | apply (auto simp:Domain_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3318 | apply (rule_tac dm_RAG_threads) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3319 | apply (unfold trancl_domain [of "RAG s", symmetric]) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3320 | by (unfold cs_RAG_def s_RAG_def, auto simp:Domain_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3321 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3322 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3323 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3324 | lemma le_cp: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3325 | shows "preced th s \<le> cp s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3326 | proof(unfold cp_eq_cpreced preced_def cpreced_def, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3327 | show "Prc (priority th s) (last_set th s) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3328 | \<le> Max (insert (Prc (priority th s) (last_set th s)) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3329 | ((\<lambda>th. Prc (priority th s) (last_set th s)) ` dependants (wq s) th))" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3330 | (is "?l \<le> Max (insert ?l ?A)") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3331 |   proof(cases "?A = {}")
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3332 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3333 | have "finite ?A" (is "finite (?f ` ?B)") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3334 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3335 | have "finite ?B" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3336 | proof- | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3337 |         have "finite {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3338 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3339 | let ?F = "\<lambda> (x, y). the_th x" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3340 |           have "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3341 | apply (auto simp:image_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3342 | by (rule_tac x = "(Th x, Th th)" in bexI, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3343 | moreover have "finite \<dots>" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3344 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3345 | from finite_RAG have "finite (RAG s)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3346 | hence "finite ((RAG (wq s))\<^sup>+)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3347 | apply (unfold finite_trancl) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3348 | by (auto simp: s_RAG_def cs_RAG_def wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3349 | thus ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3350 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3351 | ultimately show ?thesis by (auto intro:finite_subset) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3352 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3353 | thus ?thesis by (simp add:cs_dependants_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3354 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3355 | thus ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3356 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3357 | from Max_insert [OF this False, of ?l] show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3358 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3359 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3360 | thus ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3361 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3362 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3363 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3364 | lemma max_cp_eq: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3365 | shows "Max ((cp s) ` threads s) = Max ((\<lambda> th. (preced th s)) ` threads s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3366 | (is "?l = ?r") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3367 | proof(cases "threads s = {}")
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3368 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3369 | thus ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3370 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3371 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3372 | have "?l \<in> ((cp s) ` threads s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3373 | proof(rule Max_in) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3374 | from finite_threads | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3375 | show "finite (cp s ` threads s)" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3376 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3377 |     from False show "cp s ` threads s \<noteq> {}" by auto
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3378 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3379 | then obtain th | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3380 | where th_in: "th \<in> threads s" and eq_l: "?l = cp s th" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3381 | have "\<dots> \<le> ?r" by (rule cp_le[OF th_in]) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3382 | moreover have "?r \<le> cp s th" (is "Max (?f ` ?A) \<le> cp s th") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3383 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3384 | have "?r \<in> (?f ` ?A)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3385 | proof(rule Max_in) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3386 | from finite_threads | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3387 | show " finite ((\<lambda>th. preced th s) ` threads s)" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3388 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3389 |       from False show " (\<lambda>th. preced th s) ` threads s \<noteq> {}" by auto
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3390 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3391 | then obtain th' where | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3392 | th_in': "th' \<in> ?A " and eq_r: "?r = ?f th'" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3393 | from le_cp [of th'] eq_r | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3394 | have "?r \<le> cp s th'" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3395 | moreover have "\<dots> \<le> cp s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3396 | proof(fold eq_l) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3397 | show " cp s th' \<le> Max (cp s ` threads s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3398 | proof(rule Max_ge) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3399 | from th_in' show "cp s th' \<in> cp s ` threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3400 | by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3401 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3402 | from finite_threads | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3403 | show "finite (cp s ` threads s)" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3404 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3405 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3406 | ultimately show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3407 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3408 | ultimately show ?thesis using eq_l by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3409 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3410 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3411 | lemma max_cp_readys_threads_pre: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3412 |   assumes np: "threads s \<noteq> {}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3413 | shows "Max (cp s ` readys s) = Max (cp s ` threads s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3414 | proof(unfold max_cp_eq) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3415 | show "Max (cp s ` readys s) = Max ((\<lambda>th. preced th s) ` threads s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3416 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3417 | let ?p = "Max ((\<lambda>th. preced th s) ` threads s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3418 | let ?f = "(\<lambda>th. preced th s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3419 | have "?p \<in> ((\<lambda>th. preced th s) ` threads s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3420 | proof(rule Max_in) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3421 | from finite_threads show "finite (?f ` threads s)" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3422 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3423 |       from np show "?f ` threads s \<noteq> {}" by simp
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3424 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3425 | then obtain tm where tm_max: "?f tm = ?p" and tm_in: "tm \<in> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3426 | by (auto simp:Image_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3427 | from th_chain_to_ready [OF tm_in] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3428 | have "tm \<in> readys s \<or> (\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (RAG s)\<^sup>+)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3429 | thus ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3430 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3431 | assume "\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (RAG s)\<^sup>+ " | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3432 | then obtain th' where th'_in: "th' \<in> readys s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3433 | and tm_chain:"(Th tm, Th th') \<in> (RAG s)\<^sup>+" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3434 | have "cp s th' = ?f tm" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3435 | proof(subst cp_eq_cpreced, subst cpreced_def, rule Max_eqI) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3436 | from dependants_threads finite_threads | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3437 |         show "finite ((\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th'))" 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3438 | by (auto intro:finite_subset) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3439 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3440 |         fix p assume p_in: "p \<in> (\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th')"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3441 | from tm_max have " preced tm s = Max ((\<lambda>th. preced th s) ` threads s)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3442 | moreover have "p \<le> \<dots>" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3443 | proof(rule Max_ge) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3444 | from finite_threads | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3445 | show "finite ((\<lambda>th. preced th s) ` threads s)" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3446 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3447 | from p_in and th'_in and dependants_threads[of th'] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3448 | show "p \<in> (\<lambda>th. preced th s) ` threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3449 | by (auto simp:readys_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3450 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3451 | ultimately show "p \<le> preced tm s" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3452 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3453 |         show "preced tm s \<in> (\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th')"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3454 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3455 | from tm_chain | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3456 | have "tm \<in> dependants (wq s) th'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3457 | by (unfold cs_dependants_def s_RAG_def cs_RAG_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3458 | thus ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3459 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3460 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3461 | with tm_max | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3462 | have h: "cp s th' = Max ((\<lambda>th. preced th s) ` threads s)" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3463 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3464 | proof (fold h, rule Max_eqI) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3465 | fix q | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3466 | assume "q \<in> cp s ` readys s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3467 | then obtain th1 where th1_in: "th1 \<in> readys s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3468 | and eq_q: "q = cp s th1" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3469 | show "q \<le> cp s th'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3470 | apply (unfold h eq_q) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3471 | apply (unfold cp_eq_cpreced cpreced_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3472 | apply (rule Max_mono) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3473 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3474 | from dependants_threads [of th1] th1_in | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3475 |           show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) \<subseteq> 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3476 | (\<lambda>th. preced th s) ` threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3477 | by (auto simp:readys_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3478 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3479 |           show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) \<noteq> {}" by simp
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3480 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3481 | from finite_threads | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3482 | show " finite ((\<lambda>th. preced th s) ` threads s)" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3483 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3484 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3485 | from finite_threads | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3486 | show "finite (cp s ` readys s)" by (auto simp:readys_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3487 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3488 | from th'_in | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3489 | show "cp s th' \<in> cp s ` readys s" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3490 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3491 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3492 | assume tm_ready: "tm \<in> readys s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3493 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3494 | proof(fold tm_max) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3495 | have cp_eq_p: "cp s tm = preced tm s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3496 | proof(unfold cp_eq_cpreced cpreced_def, rule Max_eqI) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3497 | fix y | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3498 |           assume hy: "y \<in> (\<lambda>th. preced th s) ` ({tm} \<union> dependants (wq s) tm)"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3499 | show "y \<le> preced tm s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3500 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3501 |             { fix y'
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3502 | assume hy' : "y' \<in> ((\<lambda>th. preced th s) ` dependants (wq s) tm)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3503 | have "y' \<le> preced tm s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3504 | proof(unfold tm_max, rule Max_ge) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3505 | from hy' dependants_threads[of tm] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3506 | show "y' \<in> (\<lambda>th. preced th s) ` threads s" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3507 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3508 | from finite_threads | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3509 | show "finite ((\<lambda>th. preced th s) ` threads s)" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3510 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3511 | } with hy show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3512 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3513 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3514 | from dependants_threads[of tm] finite_threads | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3515 |           show "finite ((\<lambda>th. preced th s) ` ({tm} \<union> dependants (wq s) tm))"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3516 | by (auto intro:finite_subset) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3517 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3518 |           show "preced tm s \<in> (\<lambda>th. preced th s) ` ({tm} \<union> dependants (wq s) tm)"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3519 | by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3520 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3521 | moreover have "Max (cp s ` readys s) = cp s tm" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3522 | proof(rule Max_eqI) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3523 | from tm_ready show "cp s tm \<in> cp s ` readys s" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3524 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3525 | from finite_threads | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3526 | show "finite (cp s ` readys s)" by (auto simp:readys_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3527 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3528 | fix y assume "y \<in> cp s ` readys s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3529 | then obtain th1 where th1_readys: "th1 \<in> readys s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3530 | and h: "y = cp s th1" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3531 | show "y \<le> cp s tm" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3532 | apply(unfold cp_eq_p h) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3533 | apply(unfold cp_eq_cpreced cpreced_def tm_max, rule Max_mono) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3534 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3535 | from finite_threads | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3536 | show "finite ((\<lambda>th. preced th s) ` threads s)" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3537 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3538 |             show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) \<noteq> {}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3539 | by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3540 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3541 | from dependants_threads[of th1] th1_readys | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3542 |             show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3543 | \<subseteq> (\<lambda>th. preced th s) ` threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3544 | by (auto simp:readys_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3545 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3546 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3547 | ultimately show " Max (cp s ` readys s) = preced tm s" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3548 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3549 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3550 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3551 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3552 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3553 | text {* (* ccc *) \noindent
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3554 | Since the current precedence of the threads in ready queue will always be boosted, | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3555 | there must be one inside it has the maximum precedence of the whole system. | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3556 | *} | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3557 | lemma max_cp_readys_threads: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3558 | shows "Max (cp s ` readys s) = Max (cp s ` threads s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3559 | proof(cases "threads s = {}")
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3560 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3561 | thus ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3562 | by (auto simp:readys_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3563 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3564 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3565 | show ?thesis by (rule max_cp_readys_threads_pre[OF False]) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3566 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3567 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3568 | end | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3569 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3570 | lemma eq_holding: "holding (wq s) th cs = holding s th cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3571 | apply (unfold s_holding_def cs_holding_def wq_def, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3572 | done | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3573 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3574 | lemma f_image_eq: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3575 | assumes h: "\<And> a. a \<in> A \<Longrightarrow> f a = g a" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3576 | shows "f ` A = g ` A" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3577 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3578 | show "f ` A \<subseteq> g ` A" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3579 | by(rule image_subsetI, auto intro:h) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3580 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3581 | show "g ` A \<subseteq> f ` A" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3582 | by (rule image_subsetI, auto intro:h[symmetric]) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3583 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3584 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3585 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3586 | definition detached :: "state \<Rightarrow> thread \<Rightarrow> bool" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3587 | where "detached s th \<equiv> (\<not>(\<exists> cs. holding s th cs)) \<and> (\<not>(\<exists>cs. waiting s th cs))" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3588 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3589 | lemma detached_test: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3590 | shows "detached s th = (Th th \<notin> Field (RAG s))" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3591 | apply(simp add: detached_def Field_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3592 | apply(simp add: s_RAG_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3593 | apply(simp add: s_holding_abv s_waiting_abv) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3594 | apply(simp add: Domain_iff Range_iff) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3595 | apply(simp add: wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3596 | apply(auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3597 | done | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3598 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3599 | context valid_trace | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3600 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3601 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3602 | lemma detached_intro: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3603 | assumes eq_pv: "cntP s th = cntV s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3604 | shows "detached s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3605 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3606 | from cnp_cnv_cncs | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3607 | have eq_cnt: "cntP s th = | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3608 | cntV s th + (if th \<in> readys s \<or> th \<notin> threads s then cntCS s th else cntCS s th + 1)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3609 | hence cncs_zero: "cntCS s th = 0" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3610 | by (auto simp:eq_pv split:if_splits) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3611 | with eq_cnt | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3612 | have "th \<in> readys s \<or> th \<notin> threads s" by (auto simp:eq_pv) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3613 | thus ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3614 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3615 | assume "th \<notin> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3616 | with range_in dm_RAG_threads | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3617 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3618 | by (auto simp add: detached_def s_RAG_def s_waiting_abv s_holding_abv wq_def Domain_iff Range_iff) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3619 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3620 | assume "th \<in> readys s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3621 | moreover have "Th th \<notin> Range (RAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3622 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3623 | from card_0_eq [OF finite_holding] and cncs_zero | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3624 |       have "holdents s th = {}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3625 | by (simp add:cntCS_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3626 | thus ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3627 | apply(auto simp:holdents_test) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3628 | apply(case_tac a) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3629 | apply(auto simp:holdents_test s_RAG_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3630 | done | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3631 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3632 | ultimately show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3633 | by (auto simp add: detached_def s_RAG_def s_waiting_abv s_holding_abv wq_def readys_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3634 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3635 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3636 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3637 | lemma detached_elim: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3638 | assumes dtc: "detached s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3639 | shows "cntP s th = cntV s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3640 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3641 | from cnp_cnv_cncs | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3642 | have eq_pv: " cntP s th = | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3643 | cntV s th + (if th \<in> readys s \<or> th \<notin> threads s then cntCS s th else cntCS s th + 1)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3644 | have cncs_z: "cntCS s th = 0" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3645 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3646 |     from dtc have "holdents s th = {}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3647 | unfolding detached_def holdents_test s_RAG_def | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3648 | by (simp add: s_waiting_abv wq_def s_holding_abv Domain_iff Range_iff) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3649 | thus ?thesis by (auto simp:cntCS_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3650 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3651 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3652 | proof(cases "th \<in> threads s") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3653 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3654 | with dtc | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3655 | have "th \<in> readys s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3656 | by (unfold readys_def detached_def Field_def Domain_def Range_def, | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3657 | auto simp:waiting_eq s_RAG_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3658 | with cncs_z and eq_pv show ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3659 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3660 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3661 | with cncs_z and eq_pv show ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3662 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3663 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3664 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3665 | lemma detached_eq: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3666 | shows "(detached s th) = (cntP s th = cntV s th)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3667 | by (insert vt, auto intro:detached_intro detached_elim) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3668 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3669 | end | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3670 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3671 | text {* 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3672 | The lemmas in this .thy file are all obvious lemmas, however, they still needs to be derived | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3673 | from the concise and miniature model of PIP given in PrioGDef.thy. | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3674 | *} | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3675 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3676 | lemma eq_dependants: "dependants (wq s) = dependants s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3677 | by (simp add: s_dependants_abv wq_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3678 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3679 | lemma next_th_unique: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3680 | assumes nt1: "next_th s th cs th1" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3681 | and nt2: "next_th s th cs th2" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3682 | shows "th1 = th2" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3683 | using assms by (unfold next_th_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3684 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3685 | lemma birth_time_lt: "s \<noteq> [] \<Longrightarrow> last_set th s < length s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3686 | apply (induct s, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3687 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3688 | fix a s | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3689 | assume ih: "s \<noteq> [] \<Longrightarrow> last_set th s < length s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3690 | and eq_as: "a # s \<noteq> []" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3691 | show "last_set th (a # s) < length (a # s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3692 | proof(cases "s \<noteq> []") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3693 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3694 | from False show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3695 | by (cases a, auto simp:last_set.simps) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3696 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3697 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3698 | from ih [OF True] show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3699 | by (cases a, auto simp:last_set.simps) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3700 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3701 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3702 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3703 | lemma th_in_ne: "th \<in> threads s \<Longrightarrow> s \<noteq> []" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3704 | by (induct s, auto simp:threads.simps) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3705 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3706 | lemma preced_tm_lt: "th \<in> threads s \<Longrightarrow> preced th s = Prc x y \<Longrightarrow> y < length s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3707 | apply (drule_tac th_in_ne) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3708 | by (unfold preced_def, auto intro: birth_time_lt) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3709 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3710 | lemma inj_the_preced: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3711 | "inj_on (the_preced s) (threads s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3712 | by (metis inj_onI preced_unique the_preced_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3713 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3714 | lemma tRAG_alt_def: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3715 |   "tRAG s = {(Th th1, Th th2) | th1 th2. 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3716 | \<exists> cs. (Th th1, Cs cs) \<in> RAG s \<and> (Cs cs, Th th2) \<in> RAG s}" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3717 | by (auto simp:tRAG_def RAG_split wRAG_def hRAG_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3718 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3719 | lemma tRAG_Field: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3720 | "Field (tRAG s) \<subseteq> Field (RAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3721 | by (unfold tRAG_alt_def Field_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3722 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3723 | lemma tRAG_ancestorsE: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3724 | assumes "x \<in> ancestors (tRAG s) u" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3725 | obtains th where "x = Th th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3726 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3727 | from assms have "(u, x) \<in> (tRAG s)^+" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3728 | by (unfold ancestors_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3729 | from tranclE[OF this] obtain c where "(c, x) \<in> tRAG s" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3730 | then obtain th where "x = Th th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3731 | by (unfold tRAG_alt_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3732 | from that[OF this] show ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3733 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3734 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3735 | lemma tRAG_mono: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3736 | assumes "RAG s' \<subseteq> RAG s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3737 | shows "tRAG s' \<subseteq> tRAG s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3738 | using assms | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3739 | by (unfold tRAG_alt_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3740 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3741 | lemma holding_next_thI: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3742 | assumes "holding s th cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3743 | and "length (wq s cs) > 1" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3744 | obtains th' where "next_th s th cs th'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3745 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3746 | from assms(1)[folded eq_holding, unfolded cs_holding_def] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3747 | have " th \<in> set (wq s cs) \<and> th = hd (wq s cs)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3748 | then obtain rest where h1: "wq s cs = th#rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3749 | by (cases "wq s cs", auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3750 | with assms(2) have h2: "rest \<noteq> []" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3751 | let ?th' = "hd (SOME q. distinct q \<and> set q = set rest)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3752 | have "next_th s th cs ?th'" using h1(1) h2 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3753 | by (unfold next_th_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3754 | from that[OF this] show ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3755 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3756 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3757 | lemma RAG_tRAG_transfer: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3758 | assumes "vt s'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3759 |   assumes "RAG s = RAG s' \<union> {(Th th, Cs cs)}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3760 | and "(Cs cs, Th th'') \<in> RAG s'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3761 |   shows "tRAG s = tRAG s' \<union> {(Th th, Th th'')}" (is "?L = ?R")
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3762 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3763 | interpret vt_s': valid_trace "s'" using assms(1) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3764 | by (unfold_locales, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3765 | interpret rtree: rtree "RAG s'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3766 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3767 | show "single_valued (RAG s')" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3768 | apply (intro_locales) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3769 | by (unfold single_valued_def, | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3770 | auto intro:vt_s'.unique_RAG) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3771 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3772 | show "acyclic (RAG s')" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3773 | by (rule vt_s'.acyclic_RAG) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3774 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3775 |   { fix n1 n2
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3776 | assume "(n1, n2) \<in> ?L" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3777 | from this[unfolded tRAG_alt_def] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3778 | obtain th1 th2 cs' where | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3779 | h: "n1 = Th th1" "n2 = Th th2" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3780 | "(Th th1, Cs cs') \<in> RAG s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3781 | "(Cs cs', Th th2) \<in> RAG s" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3782 | from h(4) and assms(2) have cs_in: "(Cs cs', Th th2) \<in> RAG s'" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3783 | from h(3) and assms(2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3784 | have "(Th th1, Cs cs') = (Th th, Cs cs) \<or> | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3785 | (Th th1, Cs cs') \<in> RAG s'" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3786 | hence "(n1, n2) \<in> ?R" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3787 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3788 | assume h1: "(Th th1, Cs cs') = (Th th, Cs cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3789 | hence eq_th1: "th1 = th" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3790 | moreover have "th2 = th''" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3791 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3792 | from h1 have "cs' = cs" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3793 | from assms(3) cs_in[unfolded this] rtree.sgv | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3794 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3795 | by (unfold single_valued_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3796 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3797 | ultimately show ?thesis using h(1,2) by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3798 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3799 | assume "(Th th1, Cs cs') \<in> RAG s'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3800 | with cs_in have "(Th th1, Th th2) \<in> tRAG s'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3801 | by (unfold tRAG_alt_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3802 | from this[folded h(1, 2)] show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3803 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3804 |   } moreover {
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3805 | fix n1 n2 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3806 | assume "(n1, n2) \<in> ?R" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3807 | hence "(n1, n2) \<in>tRAG s' \<or> (n1, n2) = (Th th, Th th'')" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3808 | hence "(n1, n2) \<in> ?L" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3809 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3810 | assume "(n1, n2) \<in> tRAG s'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3811 | moreover have "... \<subseteq> ?L" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3812 | proof(rule tRAG_mono) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3813 | show "RAG s' \<subseteq> RAG s" by (unfold assms(2), auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3814 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3815 | ultimately show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3816 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3817 | assume eq_n: "(n1, n2) = (Th th, Th th'')" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3818 | from assms(2, 3) have "(Cs cs, Th th'') \<in> RAG s" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3819 | moreover have "(Th th, Cs cs) \<in> RAG s" using assms(2) by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3820 | ultimately show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3821 | by (unfold eq_n tRAG_alt_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3822 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3823 | } ultimately show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3824 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3825 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3826 | context valid_trace | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3827 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3828 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3829 | lemmas RAG_tRAG_transfer = RAG_tRAG_transfer[OF vt] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3830 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3831 | end | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3832 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3833 | lemma cp_alt_def: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3834 | "cp s th = | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3835 |            Max ((the_preced s) ` {th'. Th th' \<in> (subtree (RAG s) (Th th))})"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3836 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3837 |   have "Max (the_preced s ` ({th} \<union> dependants (wq s) th)) =
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3838 |         Max (the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th)})" 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3839 | (is "Max (_ ` ?L) = Max (_ ` ?R)") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3840 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3841 | have "?L = ?R" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3842 | by (auto dest:rtranclD simp:cs_dependants_def cs_RAG_def s_RAG_def subtree_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3843 | thus ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3844 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3845 | thus ?thesis by (unfold cp_eq_cpreced cpreced_def, fold the_preced_def, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3846 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3847 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3848 | lemma cp_gen_alt_def: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3849 | "cp_gen s = (Max \<circ> (\<lambda>x. (the_preced s \<circ> the_thread) ` subtree (tRAG s) x))" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3850 | by (auto simp:cp_gen_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3851 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3852 | lemma tRAG_nodeE: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3853 | assumes "(n1, n2) \<in> tRAG s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3854 | obtains th1 th2 where "n1 = Th th1" "n2 = Th th2" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3855 | using assms | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3856 | by (auto simp: tRAG_def wRAG_def hRAG_def tRAG_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3857 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3858 | lemma subtree_nodeE: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3859 | assumes "n \<in> subtree (tRAG s) (Th th)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3860 | obtains th1 where "n = Th th1" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3861 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3862 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3863 | proof(rule subtreeE[OF assms]) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3864 | assume "n = Th th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3865 | from that[OF this] show ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3866 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3867 | assume "Th th \<in> ancestors (tRAG s) n" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3868 | hence "(n, Th th) \<in> (tRAG s)^+" by (auto simp:ancestors_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3869 | hence "\<exists> th1. n = Th th1" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3870 | proof(induct) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3871 | case (base y) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3872 | from tRAG_nodeE[OF this] show ?case by metis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3873 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3874 | case (step y z) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3875 | thus ?case by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3876 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3877 | with that show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3878 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3879 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3880 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3881 | lemma tRAG_star_RAG: "(tRAG s)^* \<subseteq> (RAG s)^*" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3882 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3883 | have "(wRAG s O hRAG s)^* \<subseteq> (RAG s O RAG s)^*" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3884 | by (rule rtrancl_mono, auto simp:RAG_split) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3885 | also have "... \<subseteq> ((RAG s)^*)^*" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3886 | by (rule rtrancl_mono, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3887 | also have "... = (RAG s)^*" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3888 | finally show ?thesis by (unfold tRAG_def, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3889 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3890 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3891 | lemma tRAG_subtree_RAG: "subtree (tRAG s) x \<subseteq> subtree (RAG s) x" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3892 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3893 |   { fix a
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3894 | assume "a \<in> subtree (tRAG s) x" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3895 | hence "(a, x) \<in> (tRAG s)^*" by (auto simp:subtree_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3896 | with tRAG_star_RAG[of s] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3897 | have "(a, x) \<in> (RAG s)^*" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3898 | hence "a \<in> subtree (RAG s) x" by (auto simp:subtree_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3899 | } thus ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3900 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3901 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3902 | lemma tRAG_trancl_eq: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3903 |    "{th'. (Th th', Th th)  \<in> (tRAG s)^+} = 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3904 |     {th'. (Th th', Th th)  \<in> (RAG s)^+}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3905 | (is "?L = ?R") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3906 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3907 |   { fix th'
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3908 | assume "th' \<in> ?L" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3909 | hence "(Th th', Th th) \<in> (tRAG s)^+" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3910 | from tranclD[OF this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3911 | obtain z where h: "(Th th', z) \<in> tRAG s" "(z, Th th) \<in> (tRAG s)\<^sup>*" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3912 | from tRAG_subtree_RAG[of s] and this(2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3913 | have "(z, Th th) \<in> (RAG s)^*" by (meson subsetCE tRAG_star_RAG) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3914 | moreover from h(1) have "(Th th', z) \<in> (RAG s)^+" using tRAG_alt_def by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3915 | ultimately have "th' \<in> ?R" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3916 | } moreover | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3917 |   { fix th'
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3918 | assume "th' \<in> ?R" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3919 | hence "(Th th', Th th) \<in> (RAG s)^+" by (auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3920 | from plus_rpath[OF this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3921 | obtain xs where rp: "rpath (RAG s) (Th th') xs (Th th)" "xs \<noteq> []" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3922 | hence "(Th th', Th th) \<in> (tRAG s)^+" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3923 | proof(induct xs arbitrary:th' th rule:length_induct) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3924 | case (1 xs th' th) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3925 | then obtain x1 xs1 where Cons1: "xs = x1#xs1" by (cases xs, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3926 | show ?case | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3927 | proof(cases "xs1") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3928 | case Nil | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3929 | from 1(2)[unfolded Cons1 Nil] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3930 | have rp: "rpath (RAG s) (Th th') [x1] (Th th)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3931 | hence "(Th th', x1) \<in> (RAG s)" by (cases, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3932 | then obtain cs where "x1 = Cs cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3933 | by (unfold s_RAG_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3934 | from rpath_nnl_lastE[OF rp[unfolded this]] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3935 | show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3936 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3937 | case (Cons x2 xs2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3938 | from 1(2)[unfolded Cons1[unfolded this]] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3939 | have rp: "rpath (RAG s) (Th th') (x1 # x2 # xs2) (Th th)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3940 | from rpath_edges_on[OF this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3941 | have eds: "edges_on (Th th' # x1 # x2 # xs2) \<subseteq> RAG s" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3942 | have "(Th th', x1) \<in> edges_on (Th th' # x1 # x2 # xs2)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3943 | by (simp add: edges_on_unfold) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3944 | with eds have rg1: "(Th th', x1) \<in> RAG s" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3945 | then obtain cs1 where eq_x1: "x1 = Cs cs1" by (unfold s_RAG_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3946 | have "(x1, x2) \<in> edges_on (Th th' # x1 # x2 # xs2)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3947 | by (simp add: edges_on_unfold) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3948 | from this eds | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3949 | have rg2: "(x1, x2) \<in> RAG s" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3950 | from this[unfolded eq_x1] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3951 | obtain th1 where eq_x2: "x2 = Th th1" by (unfold s_RAG_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3952 | from rg1[unfolded eq_x1] rg2[unfolded eq_x1 eq_x2] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3953 | have rt1: "(Th th', Th th1) \<in> tRAG s" by (unfold tRAG_alt_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3954 | from rp have "rpath (RAG s) x2 xs2 (Th th)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3955 | by (elim rpath_ConsE, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3956 | from this[unfolded eq_x2] have rp': "rpath (RAG s) (Th th1) xs2 (Th th)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3957 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3958 | proof(cases "xs2 = []") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3959 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3960 | from rpath_nilE[OF rp'[unfolded this]] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3961 | have "th1 = th" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3962 | from rt1[unfolded this] show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3963 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3964 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3965 | from 1(1)[rule_format, OF _ rp' this, unfolded Cons1 Cons] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3966 | have "(Th th1, Th th) \<in> (tRAG s)\<^sup>+" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3967 | with rt1 show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3968 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3969 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3970 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3971 | hence "th' \<in> ?L" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3972 | } ultimately show ?thesis by blast | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3973 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3974 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3975 | lemma tRAG_trancl_eq_Th: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3976 |    "{Th th' | th'. (Th th', Th th)  \<in> (tRAG s)^+} = 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3977 |     {Th th' | th'. (Th th', Th th)  \<in> (RAG s)^+}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3978 | using tRAG_trancl_eq by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3979 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3980 | lemma dependants_alt_def: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3981 |   "dependants s th = {th'. (Th th', Th th) \<in> (tRAG s)^+}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3982 | by (metis eq_RAG s_dependants_def tRAG_trancl_eq) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3983 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3984 | context valid_trace | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3985 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3986 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3987 | lemma count_eq_tRAG_plus: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3988 | assumes "cntP s th = cntV s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3989 |   shows "{th'. (Th th', Th th) \<in> (tRAG s)^+} = {}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3990 | using assms count_eq_dependants dependants_alt_def eq_dependants by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3991 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3992 | lemma count_eq_RAG_plus: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3993 | assumes "cntP s th = cntV s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3994 |   shows "{th'. (Th th', Th th) \<in> (RAG s)^+} = {}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3995 | using assms count_eq_dependants cs_dependants_def eq_RAG by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3996 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3997 | lemma count_eq_RAG_plus_Th: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3998 | assumes "cntP s th = cntV s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 3999 |   shows "{Th th' | th'. (Th th', Th th) \<in> (RAG s)^+} = {}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4000 | using count_eq_RAG_plus[OF assms] by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4001 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4002 | lemma count_eq_tRAG_plus_Th: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4003 | assumes "cntP s th = cntV s th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4004 |   shows "{Th th' | th'. (Th th', Th th) \<in> (tRAG s)^+} = {}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4005 | using count_eq_tRAG_plus[OF assms] by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4006 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4007 | end | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4008 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4009 | lemma tRAG_subtree_eq: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4010 |    "(subtree (tRAG s) (Th th)) = {Th th' | th'. Th th'  \<in> (subtree (RAG s) (Th th))}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4011 | (is "?L = ?R") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4012 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4013 |   { fix n 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4014 | assume h: "n \<in> ?L" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4015 | hence "n \<in> ?R" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4016 | by (smt mem_Collect_eq subsetCE subtree_def subtree_nodeE tRAG_subtree_RAG) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4017 |   } moreover {
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4018 | fix n | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4019 | assume "n \<in> ?R" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4020 | then obtain th' where h: "n = Th th'" "(Th th', Th th) \<in> (RAG s)^*" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4021 | by (auto simp:subtree_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4022 | from rtranclD[OF this(2)] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4023 | have "n \<in> ?L" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4024 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4025 | assume "Th th' \<noteq> Th th \<and> (Th th', Th th) \<in> (RAG s)\<^sup>+" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4026 |       with h have "n \<in> {Th th' | th'. (Th th', Th th)  \<in> (RAG s)^+}" by auto
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4027 | thus ?thesis using subtree_def tRAG_trancl_eq by fastforce | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4028 | qed (insert h, auto simp:subtree_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4029 | } ultimately show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4030 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4031 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4032 | lemma threads_set_eq: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4033 | "the_thread ` (subtree (tRAG s) (Th th)) = | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4034 |                   {th'. Th th' \<in> (subtree (RAG s) (Th th))}" (is "?L = ?R")
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4035 | by (auto intro:rev_image_eqI simp:tRAG_subtree_eq) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4036 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4037 | lemma cp_alt_def1: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4038 | "cp s th = Max ((the_preced s o the_thread) ` (subtree (tRAG s) (Th th)))" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4039 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4040 | have "(the_preced s ` the_thread ` subtree (tRAG s) (Th th)) = | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4041 | ((the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th))" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4042 | by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4043 | thus ?thesis by (unfold cp_alt_def, fold threads_set_eq, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4044 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4045 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4046 | lemma cp_gen_def_cond: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4047 | assumes "x = Th th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4048 | shows "cp s th = cp_gen s (Th th)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4049 | by (unfold cp_alt_def1 cp_gen_def, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4050 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4051 | lemma cp_gen_over_set: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4052 | assumes "\<forall> x \<in> A. \<exists> th. x = Th th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4053 | shows "cp_gen s ` A = (cp s \<circ> the_thread) ` A" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4054 | proof(rule f_image_eq) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4055 | fix a | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4056 | assume "a \<in> A" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4057 | from assms[rule_format, OF this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4058 | obtain th where eq_a: "a = Th th" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4059 | show "cp_gen s a = (cp s \<circ> the_thread) a" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4060 | by (unfold eq_a, simp, unfold cp_gen_def_cond[OF refl[of "Th th"]], simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4061 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4062 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4063 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4064 | context valid_trace | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4065 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4066 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4067 | lemma RAG_threads: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4068 | assumes "(Th th) \<in> Field (RAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4069 | shows "th \<in> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4070 | using assms | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4071 | by (metis Field_def UnE dm_RAG_threads range_in vt) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4072 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4073 | lemma subtree_tRAG_thread: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4074 | assumes "th \<in> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4075 | shows "subtree (tRAG s) (Th th) \<subseteq> Th ` threads s" (is "?L \<subseteq> ?R") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4076 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4077 |   have "?L = {Th th' |th'. Th th' \<in> subtree (RAG s) (Th th)}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4078 | by (unfold tRAG_subtree_eq, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4079 | also have "... \<subseteq> ?R" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4080 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4081 | fix x | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4082 |     assume "x \<in> {Th th' |th'. Th th' \<in> subtree (RAG s) (Th th)}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4083 | then obtain th' where h: "x = Th th'" "Th th' \<in> subtree (RAG s) (Th th)" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4084 | from this(2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4085 | show "x \<in> ?R" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4086 | proof(cases rule:subtreeE) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4087 | case 1 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4088 | thus ?thesis by (simp add: assms h(1)) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4089 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4090 | case 2 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4091 | thus ?thesis by (metis ancestors_Field dm_RAG_threads h(1) image_eqI) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4092 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4093 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4094 | finally show ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4095 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4096 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4097 | lemma readys_root: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4098 | assumes "th \<in> readys s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4099 | shows "root (RAG s) (Th th)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4100 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4101 |   { fix x
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4102 | assume "x \<in> ancestors (RAG s) (Th th)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4103 | hence h: "(Th th, x) \<in> (RAG s)^+" by (auto simp:ancestors_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4104 | from tranclD[OF this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4105 | obtain z where "(Th th, z) \<in> RAG s" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4106 | with assms(1) have False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4107 | apply (case_tac z, auto simp:readys_def s_RAG_def s_waiting_def cs_waiting_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4108 | by (fold wq_def, blast) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4109 | } thus ?thesis by (unfold root_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4110 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4111 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4112 | lemma readys_in_no_subtree: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4113 | assumes "th \<in> readys s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4114 | and "th' \<noteq> th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4115 | shows "Th th \<notin> subtree (RAG s) (Th th')" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4116 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4117 | assume "Th th \<in> subtree (RAG s) (Th th')" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4118 | thus False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4119 | proof(cases rule:subtreeE) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4120 | case 1 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4121 | with assms show ?thesis by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4122 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4123 | case 2 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4124 | with readys_root[OF assms(1)] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4125 | show ?thesis by (auto simp:root_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4126 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4127 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4128 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4129 | lemma not_in_thread_isolated: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4130 | assumes "th \<notin> threads s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4131 | shows "(Th th) \<notin> Field (RAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4132 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4133 | assume "(Th th) \<in> Field (RAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4134 | with dm_RAG_threads and range_in assms | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4135 | show False by (unfold Field_def, blast) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4136 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4137 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4138 | lemma wf_RAG: "wf (RAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4139 | proof(rule finite_acyclic_wf) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4140 | from finite_RAG show "finite (RAG s)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4141 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4142 | from acyclic_RAG show "acyclic (RAG s)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4143 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4144 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4145 | lemma sgv_wRAG: "single_valued (wRAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4146 | using waiting_unique | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4147 | by (unfold single_valued_def wRAG_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4148 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4149 | lemma sgv_hRAG: "single_valued (hRAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4150 | using holding_unique | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4151 | by (unfold single_valued_def hRAG_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4152 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4153 | lemma sgv_tRAG: "single_valued (tRAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4154 | by (unfold tRAG_def, rule single_valued_relcomp, | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4155 | insert sgv_wRAG sgv_hRAG, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4156 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4157 | lemma acyclic_tRAG: "acyclic (tRAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4158 | proof(unfold tRAG_def, rule acyclic_compose) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4159 | show "acyclic (RAG s)" using acyclic_RAG . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4160 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4161 | show "wRAG s \<subseteq> RAG s" unfolding RAG_split by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4162 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4163 | show "hRAG s \<subseteq> RAG s" unfolding RAG_split by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4164 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4165 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4166 | lemma sgv_RAG: "single_valued (RAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4167 | using unique_RAG by (auto simp:single_valued_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4168 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4169 | lemma rtree_RAG: "rtree (RAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4170 | using sgv_RAG acyclic_RAG | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4171 | by (unfold rtree_def rtree_axioms_def sgv_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4172 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4173 | end | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4174 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4175 | sublocale valid_trace < rtree_RAG: rtree "RAG s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4176 | proof | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4177 | show "single_valued (RAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4178 | apply (intro_locales) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4179 | by (unfold single_valued_def, | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4180 | auto intro:unique_RAG) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4181 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4182 | show "acyclic (RAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4183 | by (rule acyclic_RAG) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4184 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4185 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4186 | sublocale valid_trace < rtree_s: rtree "tRAG s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4187 | proof(unfold_locales) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4188 | from sgv_tRAG show "single_valued (tRAG s)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4189 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4190 | from acyclic_tRAG show "acyclic (tRAG s)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4191 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4192 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4193 | sublocale valid_trace < fsbtRAGs : fsubtree "RAG s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4194 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4195 | show "fsubtree (RAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4196 | proof(intro_locales) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4197 | show "fbranch (RAG s)" using finite_fbranchI[OF finite_RAG] . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4198 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4199 | show "fsubtree_axioms (RAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4200 | proof(unfold fsubtree_axioms_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4201 | from wf_RAG show "wf (RAG s)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4202 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4203 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4204 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4205 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4206 | sublocale valid_trace < fsbttRAGs: fsubtree "tRAG s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4207 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4208 | have "fsubtree (tRAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4209 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4210 | have "fbranch (tRAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4211 | proof(unfold tRAG_def, rule fbranch_compose) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4212 | show "fbranch (wRAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4213 | proof(rule finite_fbranchI) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4214 | from finite_RAG show "finite (wRAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4215 | by (unfold RAG_split, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4216 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4217 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4218 | show "fbranch (hRAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4219 | proof(rule finite_fbranchI) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4220 | from finite_RAG | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4221 | show "finite (hRAG s)" by (unfold RAG_split, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4222 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4223 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4224 | moreover have "wf (tRAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4225 | proof(rule wf_subset) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4226 | show "wf (RAG s O RAG s)" using wf_RAG | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4227 | by (fold wf_comp_self, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4228 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4229 | show "tRAG s \<subseteq> (RAG s O RAG s)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4230 | by (unfold tRAG_alt_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4231 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4232 | ultimately show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4233 | by (unfold fsubtree_def fsubtree_axioms_def,auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4234 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4235 | from this[folded tRAG_def] show "fsubtree (tRAG s)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4236 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4237 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4238 | lemma Max_UNION: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4239 | assumes "finite A" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4240 |   and "A \<noteq> {}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4241 | and "\<forall> M \<in> f ` A. finite M" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4242 |   and "\<forall> M \<in> f ` A. M \<noteq> {}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4243 | shows "Max (\<Union>x\<in> A. f x) = Max (Max ` f ` A)" (is "?L = ?R") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4244 | using assms[simp] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4245 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4246 | have "?L = Max (\<Union>(f ` A))" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4247 | by (fold Union_image_eq, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4248 | also have "... = ?R" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4249 | by (subst Max_Union, simp+) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4250 | finally show ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4251 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4252 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4253 | lemma max_Max_eq: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4254 | assumes "finite A" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4255 |     and "A \<noteq> {}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4256 | and "x = y" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4257 |   shows "max x (Max A) = Max ({y} \<union> A)" (is "?L = ?R")
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4258 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4259 | have "?R = Max (insert y A)" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4260 | also from assms have "... = ?L" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4261 | by (subst Max.insert, simp+) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4262 | finally show ?thesis by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4263 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4264 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4265 | context valid_trace | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4266 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4267 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4268 | (* ddd *) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4269 | lemma cp_gen_rec: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4270 | assumes "x = Th th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4271 |   shows "cp_gen s x = Max ({the_preced s th} \<union> (cp_gen s) ` children (tRAG s) x)"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4272 | proof(cases "children (tRAG s) x = {}")
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4273 | case True | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4274 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4275 | by (unfold True cp_gen_def subtree_children, simp add:assms) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4276 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4277 | case False | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4278 |   hence [simp]: "children (tRAG s) x \<noteq> {}" by auto
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4279 | note fsbttRAGs.finite_subtree[simp] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4280 | have [simp]: "finite (children (tRAG s) x)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4281 | by (intro rev_finite_subset[OF fsbttRAGs.finite_subtree], | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4282 | rule children_subtree) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4283 |   { fix r x
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4284 |     have "subtree r x \<noteq> {}" by (auto simp:subtree_def)
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4285 | } note this[simp] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4286 |   have [simp]: "\<exists>x\<in>children (tRAG s) x. subtree (tRAG s) x \<noteq> {}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4287 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4288 | from False obtain q where "q \<in> children (tRAG s) x" by blast | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4289 |     moreover have "subtree (tRAG s) q \<noteq> {}" by simp
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4290 | ultimately show ?thesis by blast | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4291 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4292 | have h: "Max ((the_preced s \<circ> the_thread) ` | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4293 |                 ({x} \<union> \<Union>(subtree (tRAG s) ` children (tRAG s) x))) =
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4294 |         Max ({the_preced s th} \<union> cp_gen s ` children (tRAG s) x)"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4295 | (is "?L = ?R") | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4296 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4297 | let "Max (?f ` (?A \<union> \<Union> (?g ` ?B)))" = ?L | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4298 | let "Max (_ \<union> (?h ` ?B))" = ?R | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4299 | let ?L1 = "?f ` \<Union>(?g ` ?B)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4300 | have eq_Max_L1: "Max ?L1 = Max (?h ` ?B)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4301 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4302 | have "?L1 = ?f ` (\<Union> x \<in> ?B.(?g x))" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4303 | also have "... = (\<Union> x \<in> ?B. ?f ` (?g x))" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4304 | finally have "Max ?L1 = Max ..." by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4305 | also have "... = Max (Max ` (\<lambda>x. ?f ` subtree (tRAG s) x) ` ?B)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4306 | by (subst Max_UNION, simp+) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4307 | also have "... = Max (cp_gen s ` children (tRAG s) x)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4308 | by (unfold image_comp cp_gen_alt_def, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4309 | finally show ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4310 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4311 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4312 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4313 | have "?L = Max (?f ` ?A \<union> ?L1)" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4314 | also have "... = max (the_preced s (the_thread x)) (Max ?L1)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4315 | by (subst Max_Un, simp+) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4316 | also have "... = max (?f x) (Max (?h ` ?B))" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4317 | by (unfold eq_Max_L1, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4318 | also have "... =?R" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4319 | by (rule max_Max_eq, (simp)+, unfold assms, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4320 | finally show ?thesis . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4321 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4322 | qed thus ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4323 | by (fold h subtree_children, unfold cp_gen_def, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4324 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4325 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4326 | lemma cp_rec: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4327 |   "cp s th = Max ({the_preced s th} \<union> 
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4328 | (cp s o the_thread) ` children (tRAG s) (Th th))" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4329 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4330 | have "Th th = Th th" by simp | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4331 | note h = cp_gen_def_cond[OF this] cp_gen_rec[OF this] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4332 | show ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4333 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4334 | have "cp_gen s ` children (tRAG s) (Th th) = | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4335 | (cp s \<circ> the_thread) ` children (tRAG s) (Th th)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4336 | proof(rule cp_gen_over_set) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4337 | show " \<forall>x\<in>children (tRAG s) (Th th). \<exists>th. x = Th th" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4338 | by (unfold tRAG_alt_def, auto simp:children_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4339 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4340 | thus ?thesis by (subst (1) h(1), unfold h(2), simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4341 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4342 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4343 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4344 | end | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4345 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4346 | (* keep *) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4347 | lemma next_th_holding: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4348 | assumes vt: "vt s" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4349 | and nxt: "next_th s th cs th'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4350 | shows "holding (wq s) th cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4351 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4352 | from nxt[unfolded next_th_def] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4353 | obtain rest where h: "wq s cs = th # rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4354 | "rest \<noteq> []" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4355 | "th' = hd (SOME q. distinct q \<and> set q = set rest)" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4356 | thus ?thesis | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4357 | by (unfold cs_holding_def, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4358 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4359 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4360 | context valid_trace | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4361 | begin | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4362 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4363 | lemma next_th_waiting: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4364 | assumes nxt: "next_th s th cs th'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4365 | shows "waiting (wq s) th' cs" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4366 | proof - | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4367 | from nxt[unfolded next_th_def] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4368 | obtain rest where h: "wq s cs = th # rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4369 | "rest \<noteq> []" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4370 | "th' = hd (SOME q. distinct q \<and> set q = set rest)" by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4371 | from wq_distinct[of cs, unfolded h] | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4372 | have dst: "distinct (th # rest)" . | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4373 | have in_rest: "th' \<in> set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4374 | proof(unfold h, rule someI2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4375 | show "distinct rest \<and> set rest = set rest" using dst by auto | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4376 | next | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4377 | fix x assume "distinct x \<and> set x = set rest" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4378 | with h(2) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4379 | show "hd x \<in> set (rest)" by (cases x, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4380 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4381 | hence "th' \<in> set (wq s cs)" by (unfold h(1), auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4382 | moreover have "th' \<noteq> hd (wq s cs)" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4383 | by (unfold h(1), insert in_rest dst, auto) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4384 | ultimately show ?thesis by (auto simp:cs_waiting_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4385 | qed | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4386 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4387 | lemma next_th_RAG: | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4388 | assumes nxt: "next_th (s::event list) th cs th'" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4389 |   shows "{(Cs cs, Th th), (Th th', Cs cs)} \<subseteq> RAG s"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4390 | using vt assms next_th_holding next_th_waiting | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4391 | by (unfold s_RAG_def, simp) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4392 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4393 | end | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4394 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4395 | -- {* A useless definition *}
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4396 | definition cps:: "state \<Rightarrow> (thread \<times> precedence) set" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4397 | where "cps s = {(th, cp s th) | th . th \<in> threads s}"
 | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4398 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4399 | lemma "wq (V th cs # s) cs1 = ttt" | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4400 | apply (unfold wq_def, auto simp:Let_def) | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4401 | |
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4402 | end | 
| 
c495eb16beb6
CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
 zhangx parents: diff
changeset | 4403 |