Merged back ExtGG.thy and PrioG.thy.
--- a/Correctness.thy Thu Jan 28 07:43:05 2016 +0800
+++ b/Correctness.thy Thu Jan 28 16:33:49 2016 +0800
@@ -474,42 +474,37 @@
section {* The `blocking thread` *}
text {*
- The purpose of PIP is to ensure that the most
- urgent thread @{term th} is not blocked unreasonably.
- Therefore, a clear picture of the blocking thread is essential
- to assure people that the purpose is fulfilled.
-
- In this section, we are going to derive a series of lemmas
- with finally give rise to a picture of the blocking thread.
- By `blocking thread`, we mean a thread in running state but
- different from thread @{term th}.
+ The purpose of PIP is to ensure that the most urgent thread @{term
+ th} is not blocked unreasonably. Therefore, below, we will derive
+ properties of the blocking thread. By blocking thread, we mean a
+ thread in running state t @ s, but is different from thread @{term
+ th}.
+
+ The first lemmas shows that the @{term cp}-value of the blocking
+ thread @{text th'} equals to the highest precedence in the whole
+ system.
+
*}
-text {*
- The following lemmas shows that the @{term cp}-value
- of the blocking thread @{text th'} equals to the highest
- precedence in the whole system.
-*}
lemma runing_preced_inversion:
- assumes runing': "th' \<in> runing (t@s)"
- shows "cp (t@s) th' = preced th s" (is "?L = ?R")
+ assumes runing': "th' \<in> runing (t @ s)"
+ shows "cp (t @ s) th' = preced th s"
proof -
- have "?L = Max (cp (t @ s) ` readys (t @ s))" using assms
- by (unfold runing_def, auto)
- also have "\<dots> = ?R"
- by (metis th_cp_max th_cp_preced vat_t.max_cp_readys_threads)
+ have "cp (t @ s) th' = Max (cp (t @ s) ` readys (t @ s))"
+ using assms by (unfold runing_def, auto)
+ also have "\<dots> = preced th s"
+ by (metis th_cp_max th_cp_preced vat_t.max_cp_readys_threads)
finally show ?thesis .
qed
text {*
- The following lemma shows how the counters for @{term "P"} and
- @{term "V"} operations relate to the running threads in the states
- @{term s} and @{term "t @ s"}. The lemma shows that if a thread's
- @{term "P"}-count equals its @{term "V"}-count (which means it no
- longer has any resource in its possession), it cannot be a running
- thread.
+ The next lemma shows how the counters for @{term "P"} and @{term
+ "V"} operations relate to the running threads in the states @{term
+ s} and @{term "t @ s"}: if a thread's @{term "P"}-count equals its
+ @{term "V"}-count (which means it no longer has any resource in its
+ possession), it cannot be a running thread.
The proof is by contraction with the assumption @{text "th' \<noteq> th"}.
The key is the use of @{thm count_eq_dependants} to derive the
@@ -521,7 +516,7 @@
runing_preced_inversion}, its @{term cp}-value equals to the
precedence of @{term th}.
- Combining the above two resukts we have that @{text th'} and @{term
+ Combining the above two results we have that @{text th'} and @{term
th} have the same precedence. By uniqueness of precedences, we have
@{text "th' = th"}, which is in contradiction with the assumption
@{text "th' \<noteq> th"}.
@@ -530,13 +525,13 @@
lemma eq_pv_blocked: (* ddd *)
assumes neq_th': "th' \<noteq> th"
- and eq_pv: "cntP (t@s) th' = cntV (t@s) th'"
- shows "th' \<notin> runing (t@s)"
+ and eq_pv: "cntP (t @ s) th' = cntV (t @ s) th'"
+ shows "th' \<notin> runing (t @ s)"
proof
- assume otherwise: "th' \<in> runing (t@s)"
+ assume otherwise: "th' \<in> runing (t @ s)"
show False
proof -
- have th'_in: "th' \<in> threads (t@s)"
+ have th'_in: "th' \<in> threads (t @ s)"
using otherwise readys_threads runing_def by auto
have "th' = th"
proof(rule preced_unique)
@@ -550,13 +545,12 @@
-- {* Since the counts of @{term th'} are balanced, the subtree
of it contains only itself, so, its @{term cp}-value
equals its @{term preced}-value: *}
- have "?L = cp (t@s) th'"
+ have "?L = cp (t @ s) th'"
by (unfold cp_eq_cpreced cpreced_def count_eq_dependants[OF eq_pv], simp)
-- {* Since @{term "th'"} is running, by @{thm runing_preced_inversion},
its @{term cp}-value equals @{term "preced th s"},
which equals to @{term "?R"} by simplification: *}
also have "... = ?R"
- thm runing_preced_inversion
using runing_preced_inversion[OF otherwise] by simp
finally show ?thesis .
qed
@@ -574,8 +568,8 @@
lemma eq_pv_persist: (* ddd *)
assumes neq_th': "th' \<noteq> th"
and eq_pv: "cntP s th' = cntV s th'"
- shows "cntP (t@s) th' = cntV (t@s) th'"
-proof(induction rule:ind) -- {* The proof goes by induction. *}
+ shows "cntP (t @ s) th' = cntV (t @ s) th'"
+proof(induction rule: ind)
-- {* The nontrivial case is for the @{term Cons}: *}
case (Cons e t)
-- {* All results derived so far hold for both @{term s} and @{term "t@s"}: *}
@@ -624,22 +618,28 @@
qed (auto simp:eq_pv)
text {*
- By combining @{thm eq_pv_blocked} and @{thm eq_pv_persist},
- it can be derived easily that @{term th'} can not be running in the future:
+
+ By combining @{thm eq_pv_blocked} and @{thm eq_pv_persist}, it can
+ be derived easily that @{term th'} can not be running in the future:
+
*}
+
lemma eq_pv_blocked_persist:
assumes neq_th': "th' \<noteq> th"
and eq_pv: "cntP s th' = cntV s th'"
- shows "th' \<notin> runing (t@s)"
+ shows "th' \<notin> runing (t @ s)"
using assms
by (simp add: eq_pv_blocked eq_pv_persist)
text {*
- The following lemma shows the blocking thread @{term th'}
- must hold some resource in the very beginning.
+
+ The following lemma shows the blocking thread @{term th'} must hold
+ some resource in the very beginning.
+
*}
+
lemma runing_cntP_cntV_inv: (* ddd *)
- assumes is_runing: "th' \<in> runing (t@s)"
+ assumes is_runing: "th' \<in> runing (t @ s)"
and neq_th': "th' \<noteq> th"
shows "cntP s th' > cntV s th'"
using assms
@@ -665,11 +665,13 @@
text {*
- The following lemmas shows the blocking thread @{text th'} must be live
- at the very beginning, i.e. the moment (or state) @{term s}.
+ The following lemmas shows the blocking thread @{text th'} must be
+ live at the very beginning, i.e. the moment (or state) @{term s}.
The proof is a simple combination of the results above:
+
*}
+
lemma runing_threads_inv:
assumes runing': "th' \<in> runing (t@s)"
and neq_th': "th' \<noteq> th"
@@ -687,9 +689,12 @@
qed
text {*
- The following lemma summarizes several foregoing
- lemmas to give an overall picture of the blocking thread @{text "th'"}:
+
+ The following lemma summarises the above lemmas to give an overall
+ characterisationof the blocking thread @{text "th'"}:
+
*}
+
lemma runing_inversion: (* ddd, one of the main lemmas to present *)
assumes runing': "th' \<in> runing (t@s)"
and neq_th: "th' \<noteq> th"
@@ -707,18 +712,23 @@
show "cp (t@s) th' = preced th s" .
qed
+
section {* The existence of `blocking thread` *}
text {*
- Suppose @{term th} is not running, it is first shown that
- there is a path in RAG leading from node @{term th} to another thread @{text "th'"}
- in the @{term readys}-set (So @{text "th'"} is an ancestor of @{term th}}).
+
+ Suppose @{term th} is not running, it is first shown that there is a
+ path in RAG leading from node @{term th} to another thread @{text
+ "th'"} in the @{term readys}-set (So @{text "th'"} is an ancestor of
+ @{term th}}).
- Now, since @{term readys}-set is non-empty, there must be
- one in it which holds the highest @{term cp}-value, which, by definition,
- is the @{term runing}-thread. However, we are going to show more: this running thread
- is exactly @{term "th'"}.
- *}
+ Now, since @{term readys}-set is non-empty, there must be one in it
+ which holds the highest @{term cp}-value, which, by definition, is
+ the @{term runing}-thread. However, we are going to show more: this
+ running thread is exactly @{term "th'"}.
+
+*}
+
lemma th_blockedE: (* ddd, the other main lemma to be presented: *)
assumes "th \<notin> runing (t@s)"
obtains th' where "Th th' \<in> ancestors (RAG (t @ s)) (Th th)"
@@ -780,11 +790,15 @@
qed
text {*
- Now it is easy to see there is always a thread to run by case analysis
- on whether thread @{term th} is running: if the answer is Yes, the
- the running thread is obviously @{term th} itself; otherwise, the running
- thread is the @{text th'} given by lemma @{thm th_blockedE}.
+
+ Now it is easy to see there is always a thread to run by case
+ analysis on whether thread @{term th} is running: if the answer is
+ yes, the the running thread is obviously @{term th} itself;
+ otherwise, the running thread is the @{text th'} given by lemma
+ @{thm th_blockedE}.
+
*}
+
lemma live: "runing (t@s) \<noteq> {}"
proof(cases "th \<in> runing (t@s)")
case True thus ?thesis by auto
--- a/ExtGG.thy Thu Jan 28 07:43:05 2016 +0800
+++ b/ExtGG.thy Thu Jan 28 16:33:49 2016 +0800
@@ -1,3 +1,4 @@
+<<<<<<< local
section {*
This file contains lemmas used to guide the recalculation of current precedence
after every system call (or system operation)
@@ -706,3 +707,927 @@
end
+=======
+theory ExtGG
+imports PrioG CpsG
+begin
+
+text {*
+ The following two auxiliary lemmas are used to reason about @{term Max}.
+*}
+lemma image_Max_eqI:
+ assumes "finite B"
+ and "b \<in> B"
+ and "\<forall> x \<in> B. f x \<le> f b"
+ shows "Max (f ` B) = f b"
+ using assms
+ using Max_eqI by blast
+
+lemma image_Max_subset:
+ assumes "finite A"
+ and "B \<subseteq> A"
+ and "a \<in> B"
+ and "Max (f ` A) = f a"
+ shows "Max (f ` B) = f a"
+proof(rule image_Max_eqI)
+ show "finite B"
+ using assms(1) assms(2) finite_subset by auto
+next
+ show "a \<in> B" using assms by simp
+next
+ show "\<forall>x\<in>B. f x \<le> f a"
+ by (metis Max_ge assms(1) assms(2) assms(4)
+ finite_imageI image_eqI subsetCE)
+qed
+
+text {*
+ The following locale @{text "highest_gen"} sets the basic context for our
+ investigation: supposing thread @{text th} holds the highest @{term cp}-value
+ in state @{text s}, which means the task for @{text th} is the
+ most urgent. We want to show that
+ @{text th} is treated correctly by PIP, which means
+ @{text th} will not be blocked unreasonably by other less urgent
+ threads.
+*}
+locale highest_gen =
+ fixes s th prio tm
+ assumes vt_s: "vt s"
+ and threads_s: "th \<in> threads s"
+ and highest: "preced th s = Max ((cp s)`threads s)"
+ -- {* The internal structure of @{term th}'s precedence is exposed:*}
+ and preced_th: "preced th s = Prc prio tm"
+
+-- {* @{term s} is a valid trace, so it will inherit all results derived for
+ a valid trace: *}
+sublocale highest_gen < vat_s: valid_trace "s"
+ by (unfold_locales, insert vt_s, simp)
+
+context highest_gen
+begin
+
+text {*
+ @{term tm} is the time when the precedence of @{term th} is set, so
+ @{term tm} must be a valid moment index into @{term s}.
+*}
+lemma lt_tm: "tm < length s"
+ by (insert preced_tm_lt[OF threads_s preced_th], simp)
+
+text {*
+ Since @{term th} holds the highest precedence and @{text "cp"}
+ is the highest precedence of all threads in the sub-tree of
+ @{text "th"} and @{text th} is among these threads,
+ its @{term cp} must equal to its precedence:
+*}
+lemma eq_cp_s_th: "cp s th = preced th s" (is "?L = ?R")
+proof -
+ have "?L \<le> ?R"
+ by (unfold highest, rule Max_ge,
+ auto simp:threads_s finite_threads)
+ moreover have "?R \<le> ?L"
+ by (unfold vat_s.cp_rec, rule Max_ge,
+ auto simp:the_preced_def vat_s.fsbttRAGs.finite_children)
+ ultimately show ?thesis by auto
+qed
+
+(* ccc *)
+lemma highest_cp_preced: "cp s th = Max ((\<lambda> th'. preced th' s) ` threads s)"
+ by (fold max_cp_eq, unfold eq_cp_s_th, insert highest, simp)
+
+lemma highest_preced_thread: "preced th s = Max ((\<lambda> th'. preced th' s) ` threads s)"
+ by (fold eq_cp_s_th, unfold highest_cp_preced, simp)
+
+lemma highest': "cp s th = Max (cp s ` threads s)"
+proof -
+ from highest_cp_preced max_cp_eq[symmetric]
+ show ?thesis by simp
+qed
+
+end
+
+locale extend_highest_gen = highest_gen +
+ fixes t
+ assumes vt_t: "vt (t@s)"
+ and create_low: "Create th' prio' \<in> set t \<Longrightarrow> prio' \<le> prio"
+ and set_diff_low: "Set th' prio' \<in> set t \<Longrightarrow> th' \<noteq> th \<and> prio' \<le> prio"
+ and exit_diff: "Exit th' \<in> set t \<Longrightarrow> th' \<noteq> th"
+
+sublocale extend_highest_gen < vat_t: valid_trace "t@s"
+ by (unfold_locales, insert vt_t, simp)
+
+lemma step_back_vt_app:
+ assumes vt_ts: "vt (t@s)"
+ shows "vt s"
+proof -
+ from vt_ts show ?thesis
+ proof(induct t)
+ case Nil
+ from Nil show ?case by auto
+ next
+ case (Cons e t)
+ assume ih: " vt (t @ s) \<Longrightarrow> vt s"
+ and vt_et: "vt ((e # t) @ s)"
+ show ?case
+ proof(rule ih)
+ show "vt (t @ s)"
+ proof(rule step_back_vt)
+ from vt_et show "vt (e # t @ s)" by simp
+ qed
+ qed
+ qed
+qed
+
+
+locale red_extend_highest_gen = extend_highest_gen +
+ fixes i::nat
+
+sublocale red_extend_highest_gen < red_moment: extend_highest_gen "s" "th" "prio" "tm" "(moment i t)"
+ apply (insert extend_highest_gen_axioms, subst (asm) (1) moment_restm_s [of i t, symmetric])
+ apply (unfold extend_highest_gen_def extend_highest_gen_axioms_def, clarsimp)
+ by (unfold highest_gen_def, auto dest:step_back_vt_app)
+
+
+context extend_highest_gen
+begin
+
+ lemma ind [consumes 0, case_names Nil Cons, induct type]:
+ assumes
+ h0: "R []"
+ and h2: "\<And> e t. \<lbrakk>vt (t@s); step (t@s) e;
+ extend_highest_gen s th prio tm t;
+ extend_highest_gen s th prio tm (e#t); R t\<rbrakk> \<Longrightarrow> R (e#t)"
+ shows "R t"
+proof -
+ from vt_t extend_highest_gen_axioms show ?thesis
+ proof(induct t)
+ from h0 show "R []" .
+ next
+ case (Cons e t')
+ assume ih: "\<lbrakk>vt (t' @ s); extend_highest_gen s th prio tm t'\<rbrakk> \<Longrightarrow> R t'"
+ and vt_e: "vt ((e # t') @ s)"
+ and et: "extend_highest_gen s th prio tm (e # t')"
+ from vt_e and step_back_step have stp: "step (t'@s) e" by auto
+ from vt_e and step_back_vt have vt_ts: "vt (t'@s)" by auto
+ show ?case
+ proof(rule h2 [OF vt_ts stp _ _ _ ])
+ show "R t'"
+ proof(rule ih)
+ from et show ext': "extend_highest_gen s th prio tm t'"
+ by (unfold extend_highest_gen_def extend_highest_gen_axioms_def, auto dest:step_back_vt)
+ next
+ from vt_ts show "vt (t' @ s)" .
+ qed
+ next
+ from et show "extend_highest_gen s th prio tm (e # t')" .
+ next
+ from et show ext': "extend_highest_gen s th prio tm t'"
+ by (unfold extend_highest_gen_def extend_highest_gen_axioms_def, auto dest:step_back_vt)
+ qed
+ qed
+qed
+
+
+lemma th_kept: "th \<in> threads (t @ s) \<and>
+ preced th (t@s) = preced th s" (is "?Q t")
+proof -
+ show ?thesis
+ proof(induct rule:ind)
+ case Nil
+ from threads_s
+ show ?case
+ by auto
+ next
+ case (Cons e t)
+ interpret h_e: extend_highest_gen _ _ _ _ "(e # t)" using Cons by auto
+ interpret h_t: extend_highest_gen _ _ _ _ t using Cons by auto
+ show ?case
+ proof(cases e)
+ case (Create thread prio)
+ show ?thesis
+ proof -
+ from Cons and Create have "step (t@s) (Create thread prio)" by auto
+ hence "th \<noteq> thread"
+ proof(cases)
+ case thread_create
+ with Cons show ?thesis by auto
+ qed
+ hence "preced th ((e # t) @ s) = preced th (t @ s)"
+ by (unfold Create, auto simp:preced_def)
+ moreover note Cons
+ ultimately show ?thesis
+ by (auto simp:Create)
+ qed
+ next
+ case (Exit thread)
+ from h_e.exit_diff and Exit
+ have neq_th: "thread \<noteq> th" by auto
+ with Cons
+ show ?thesis
+ by (unfold Exit, auto simp:preced_def)
+ next
+ case (P thread cs)
+ with Cons
+ show ?thesis
+ by (auto simp:P preced_def)
+ next
+ case (V thread cs)
+ with Cons
+ show ?thesis
+ by (auto simp:V preced_def)
+ next
+ case (Set thread prio')
+ show ?thesis
+ proof -
+ from h_e.set_diff_low and Set
+ have "th \<noteq> thread" by auto
+ hence "preced th ((e # t) @ s) = preced th (t @ s)"
+ by (unfold Set, auto simp:preced_def)
+ moreover note Cons
+ ultimately show ?thesis
+ by (auto simp:Set)
+ qed
+ qed
+ qed
+qed
+
+text {*
+ According to @{thm th_kept}, thread @{text "th"} has its living status
+ and precedence kept along the way of @{text "t"}. The following lemma
+ shows that this preserved precedence of @{text "th"} remains as the highest
+ along the way of @{text "t"}.
+
+ The proof goes by induction over @{text "t"} using the specialized
+ induction rule @{thm ind}, followed by case analysis of each possible
+ operations of PIP. All cases follow the same pattern rendered by the
+ generalized introduction rule @{thm "image_Max_eqI"}.
+
+ The very essence is to show that precedences, no matter whether they are newly introduced
+ or modified, are always lower than the one held by @{term "th"},
+ which by @{thm th_kept} is preserved along the way.
+*}
+lemma max_kept: "Max (the_preced (t @ s) ` (threads (t@s))) = preced th s"
+proof(induct rule:ind)
+ case Nil
+ from highest_preced_thread
+ show ?case
+ by (unfold the_preced_def, simp)
+next
+ case (Cons e t)
+ interpret h_e: extend_highest_gen _ _ _ _ "(e # t)" using Cons by auto
+ interpret h_t: extend_highest_gen _ _ _ _ t using Cons by auto
+ show ?case
+ proof(cases e)
+ case (Create thread prio')
+ show ?thesis (is "Max (?f ` ?A) = ?t")
+ proof -
+ -- {* The following is the common pattern of each branch of the case analysis. *}
+ -- {* The major part is to show that @{text "th"} holds the highest precedence: *}
+ have "Max (?f ` ?A) = ?f th"
+ proof(rule image_Max_eqI)
+ show "finite ?A" using h_e.finite_threads by auto
+ next
+ show "th \<in> ?A" using h_e.th_kept by auto
+ next
+ show "\<forall>x\<in>?A. ?f x \<le> ?f th"
+ proof
+ fix x
+ assume "x \<in> ?A"
+ hence "x = thread \<or> x \<in> threads (t@s)" by (auto simp:Create)
+ thus "?f x \<le> ?f th"
+ proof
+ assume "x = thread"
+ thus ?thesis
+ apply (simp add:Create the_preced_def preced_def, fold preced_def)
+ using Create h_e.create_low h_t.th_kept lt_tm preced_leI2 preced_th by force
+ next
+ assume h: "x \<in> threads (t @ s)"
+ from Cons(2)[unfolded Create]
+ have "x \<noteq> thread" using h by (cases, auto)
+ hence "?f x = the_preced (t@s) x"
+ by (simp add:Create the_preced_def preced_def)
+ hence "?f x \<le> Max (the_preced (t@s) ` threads (t@s))"
+ by (simp add: h_t.finite_threads h)
+ also have "... = ?f th"
+ by (metis Cons.hyps(5) h_e.th_kept the_preced_def)
+ finally show ?thesis .
+ qed
+ qed
+ qed
+ -- {* The minor part is to show that the precedence of @{text "th"}
+ equals to preserved one, given by the foregoing lemma @{thm th_kept} *}
+ also have "... = ?t" using h_e.th_kept the_preced_def by auto
+ -- {* Then it follows trivially that the precedence preserved
+ for @{term "th"} remains the maximum of all living threads along the way. *}
+ finally show ?thesis .
+ qed
+ next
+ case (Exit thread)
+ show ?thesis (is "Max (?f ` ?A) = ?t")
+ proof -
+ have "Max (?f ` ?A) = ?f th"
+ proof(rule image_Max_eqI)
+ show "finite ?A" using h_e.finite_threads by auto
+ next
+ show "th \<in> ?A" using h_e.th_kept by auto
+ next
+ show "\<forall>x\<in>?A. ?f x \<le> ?f th"
+ proof
+ fix x
+ assume "x \<in> ?A"
+ hence "x \<in> threads (t@s)" by (simp add: Exit)
+ hence "?f x \<le> Max (?f ` threads (t@s))"
+ by (simp add: h_t.finite_threads)
+ also have "... \<le> ?f th"
+ apply (simp add:Exit the_preced_def preced_def, fold preced_def)
+ using Cons.hyps(5) h_t.th_kept the_preced_def by auto
+ finally show "?f x \<le> ?f th" .
+ qed
+ qed
+ also have "... = ?t" using h_e.th_kept the_preced_def by auto
+ finally show ?thesis .
+ qed
+ next
+ case (P thread cs)
+ with Cons
+ show ?thesis by (auto simp:preced_def the_preced_def)
+ next
+ case (V thread cs)
+ with Cons
+ show ?thesis by (auto simp:preced_def the_preced_def)
+ next
+ case (Set thread prio')
+ show ?thesis (is "Max (?f ` ?A) = ?t")
+ proof -
+ have "Max (?f ` ?A) = ?f th"
+ proof(rule image_Max_eqI)
+ show "finite ?A" using h_e.finite_threads by auto
+ next
+ show "th \<in> ?A" using h_e.th_kept by auto
+ next
+ show "\<forall>x\<in>?A. ?f x \<le> ?f th"
+ proof
+ fix x
+ assume h: "x \<in> ?A"
+ show "?f x \<le> ?f th"
+ proof(cases "x = thread")
+ case True
+ moreover have "the_preced (Set thread prio' # t @ s) thread \<le> the_preced (t @ s) th"
+ proof -
+ have "the_preced (t @ s) th = Prc prio tm"
+ using h_t.th_kept preced_th by (simp add:the_preced_def)
+ moreover have "prio' \<le> prio" using Set h_e.set_diff_low by auto
+ ultimately show ?thesis by (insert lt_tm, auto simp:the_preced_def preced_def)
+ qed
+ ultimately show ?thesis
+ by (unfold Set, simp add:the_preced_def preced_def)
+ next
+ case False
+ then have "?f x = the_preced (t@s) x"
+ by (simp add:the_preced_def preced_def Set)
+ also have "... \<le> Max (the_preced (t@s) ` threads (t@s))"
+ using Set h h_t.finite_threads by auto
+ also have "... = ?f th" by (metis Cons.hyps(5) h_e.th_kept the_preced_def)
+ finally show ?thesis .
+ qed
+ qed
+ qed
+ also have "... = ?t" using h_e.th_kept the_preced_def by auto
+ finally show ?thesis .
+ qed
+ qed
+qed
+
+lemma max_preced: "preced th (t@s) = Max (the_preced (t@s) ` (threads (t@s)))"
+ by (insert th_kept max_kept, auto)
+
+text {*
+ The reason behind the following lemma is that:
+ Since @{term "cp"} is defined as the maximum precedence
+ of those threads contained in the sub-tree of node @{term "Th th"}
+ in @{term "RAG (t@s)"}, and all these threads are living threads, and
+ @{term "th"} is also among them, the maximum precedence of
+ them all must be the one for @{text "th"}.
+*}
+lemma th_cp_max_preced:
+ "cp (t@s) th = Max (the_preced (t@s) ` (threads (t@s)))" (is "?L = ?R")
+proof -
+ let ?f = "the_preced (t@s)"
+ have "?L = ?f th"
+ proof(unfold cp_alt_def, rule image_Max_eqI)
+ show "finite {th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}"
+ proof -
+ have "{th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)} =
+ the_thread ` {n . n \<in> subtree (RAG (t @ s)) (Th th) \<and>
+ (\<exists> th'. n = Th th')}"
+ by (smt Collect_cong Setcompr_eq_image mem_Collect_eq the_thread.simps)
+ moreover have "finite ..." by (simp add: vat_t.fsbtRAGs.finite_subtree)
+ ultimately show ?thesis by simp
+ qed
+ next
+ show "th \<in> {th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}"
+ by (auto simp:subtree_def)
+ next
+ show "\<forall>x\<in>{th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}.
+ the_preced (t @ s) x \<le> the_preced (t @ s) th"
+ proof
+ fix th'
+ assume "th' \<in> {th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}"
+ hence "Th th' \<in> subtree (RAG (t @ s)) (Th th)" by auto
+ moreover have "... \<subseteq> Field (RAG (t @ s)) \<union> {Th th}"
+ by (meson subtree_Field)
+ ultimately have "Th th' \<in> ..." by auto
+ hence "th' \<in> threads (t@s)"
+ proof
+ assume "Th th' \<in> {Th th}"
+ thus ?thesis using th_kept by auto
+ next
+ assume "Th th' \<in> Field (RAG (t @ s))"
+ thus ?thesis using vat_t.not_in_thread_isolated by blast
+ qed
+ thus "the_preced (t @ s) th' \<le> the_preced (t @ s) th"
+ by (metis Max_ge finite_imageI finite_threads image_eqI
+ max_kept th_kept the_preced_def)
+ qed
+ qed
+ also have "... = ?R" by (simp add: max_preced the_preced_def)
+ finally show ?thesis .
+qed
+
+lemma th_cp_max: "cp (t@s) th = Max (cp (t@s) ` threads (t@s))"
+ using max_cp_eq th_cp_max_preced the_preced_def vt_t by presburger
+
+lemma th_cp_preced: "cp (t@s) th = preced th s"
+ by (fold max_kept, unfold th_cp_max_preced, simp)
+
+lemma preced_less:
+ assumes th'_in: "th' \<in> threads s"
+ and neq_th': "th' \<noteq> th"
+ shows "preced th' s < preced th s"
+ using assms
+by (metis Max.coboundedI finite_imageI highest not_le order.trans
+ preced_linorder rev_image_eqI threads_s vat_s.finite_threads
+ vat_s.le_cp)
+
+text {*
+ Counting of the number of @{term "P"} and @{term "V"} operations
+ is the cornerstone of a large number of the following proofs.
+ The reason is that this counting is quite easy to calculate and
+ convenient to use in the reasoning.
+
+ The following lemma shows that the counting controls whether
+ a thread is running or not.
+*}
+
+lemma pv_blocked_pre:
+ assumes th'_in: "th' \<in> threads (t@s)"
+ and neq_th': "th' \<noteq> th"
+ and eq_pv: "cntP (t@s) th' = cntV (t@s) th'"
+ shows "th' \<notin> runing (t@s)"
+proof
+ assume otherwise: "th' \<in> runing (t@s)"
+ show False
+ proof -
+ have "th' = th"
+ proof(rule preced_unique)
+ show "preced th' (t @ s) = preced th (t @ s)" (is "?L = ?R")
+ proof -
+ have "?L = cp (t@s) th'"
+ by (unfold cp_eq_cpreced cpreced_def count_eq_dependants[OF eq_pv], simp)
+ also have "... = cp (t @ s) th" using otherwise
+ by (metis (mono_tags, lifting) mem_Collect_eq
+ runing_def th_cp_max vat_t.max_cp_readys_threads)
+ also have "... = ?R" by (metis th_cp_preced th_kept)
+ finally show ?thesis .
+ qed
+ qed (auto simp: th'_in th_kept)
+ moreover have "th' \<noteq> th" using neq_th' .
+ ultimately show ?thesis by simp
+ qed
+qed
+
+lemmas pv_blocked = pv_blocked_pre[folded detached_eq]
+
+lemma runing_precond_pre:
+ fixes th'
+ assumes th'_in: "th' \<in> threads s"
+ and eq_pv: "cntP s th' = cntV s th'"
+ and neq_th': "th' \<noteq> th"
+ shows "th' \<in> threads (t@s) \<and>
+ cntP (t@s) th' = cntV (t@s) th'"
+proof(induct rule:ind)
+ case (Cons e t)
+ interpret vat_t: extend_highest_gen s th prio tm t using Cons by simp
+ interpret vat_e: extend_highest_gen s th prio tm "(e # t)" using Cons by simp
+ show ?case
+ proof(cases e)
+ case (P thread cs)
+ show ?thesis
+ proof -
+ have "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
+ proof -
+ have "thread \<noteq> th'"
+ proof -
+ have "step (t@s) (P thread cs)" using Cons P by auto
+ thus ?thesis
+ proof(cases)
+ assume "thread \<in> runing (t@s)"
+ moreover have "th' \<notin> runing (t@s)" using Cons(5)
+ by (metis neq_th' vat_t.pv_blocked_pre)
+ ultimately show ?thesis by auto
+ qed
+ qed with Cons show ?thesis
+ by (unfold P, simp add:cntP_def cntV_def count_def)
+ qed
+ moreover have "th' \<in> threads ((e # t) @ s)" using Cons by (unfold P, simp)
+ ultimately show ?thesis by auto
+ qed
+ next
+ case (V thread cs)
+ show ?thesis
+ proof -
+ have "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
+ proof -
+ have "thread \<noteq> th'"
+ proof -
+ have "step (t@s) (V thread cs)" using Cons V by auto
+ thus ?thesis
+ proof(cases)
+ assume "thread \<in> runing (t@s)"
+ moreover have "th' \<notin> runing (t@s)" using Cons(5)
+ by (metis neq_th' vat_t.pv_blocked_pre)
+ ultimately show ?thesis by auto
+ qed
+ qed with Cons show ?thesis
+ by (unfold V, simp add:cntP_def cntV_def count_def)
+ qed
+ moreover have "th' \<in> threads ((e # t) @ s)" using Cons by (unfold V, simp)
+ ultimately show ?thesis by auto
+ qed
+ next
+ case (Create thread prio')
+ show ?thesis
+ proof -
+ have "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
+ proof -
+ have "thread \<noteq> th'"
+ proof -
+ have "step (t@s) (Create thread prio')" using Cons Create by auto
+ thus ?thesis using Cons(5) by (cases, auto)
+ qed with Cons show ?thesis
+ by (unfold Create, simp add:cntP_def cntV_def count_def)
+ qed
+ moreover have "th' \<in> threads ((e # t) @ s)" using Cons by (unfold Create, simp)
+ ultimately show ?thesis by auto
+ qed
+ next
+ case (Exit thread)
+ show ?thesis
+ proof -
+ have neq_thread: "thread \<noteq> th'"
+ proof -
+ have "step (t@s) (Exit thread)" using Cons Exit by auto
+ thus ?thesis apply (cases) using Cons(5)
+ by (metis neq_th' vat_t.pv_blocked_pre)
+ qed
+ hence "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'" using Cons
+ by (unfold Exit, simp add:cntP_def cntV_def count_def)
+ moreover have "th' \<in> threads ((e # t) @ s)" using Cons neq_thread
+ by (unfold Exit, simp)
+ ultimately show ?thesis by auto
+ qed
+ next
+ case (Set thread prio')
+ with Cons
+ show ?thesis
+ by (auto simp:cntP_def cntV_def count_def)
+ qed
+next
+ case Nil
+ with assms
+ show ?case by auto
+qed
+
+text {* Changing counting balance to detachedness *}
+lemmas runing_precond_pre_dtc = runing_precond_pre
+ [folded vat_t.detached_eq vat_s.detached_eq]
+
+lemma runing_precond:
+ fixes th'
+ assumes th'_in: "th' \<in> threads s"
+ and neq_th': "th' \<noteq> th"
+ and is_runing: "th' \<in> runing (t@s)"
+ shows "cntP s th' > cntV s th'"
+ using assms
+proof -
+ have "cntP s th' \<noteq> cntV s th'"
+ by (metis is_runing neq_th' pv_blocked_pre runing_precond_pre th'_in)
+ moreover have "cntV s th' \<le> cntP s th'" using vat_s.cnp_cnv_cncs by auto
+ ultimately show ?thesis by auto
+qed
+
+lemma moment_blocked_pre:
+ assumes neq_th': "th' \<noteq> th"
+ and th'_in: "th' \<in> threads ((moment i t)@s)"
+ and eq_pv: "cntP ((moment i t)@s) th' = cntV ((moment i t)@s) th'"
+ shows "cntP ((moment (i+j) t)@s) th' = cntV ((moment (i+j) t)@s) th' \<and>
+ th' \<in> threads ((moment (i+j) t)@s)"
+proof -
+ interpret h_i: red_extend_highest_gen _ _ _ _ _ i
+ by (unfold_locales)
+ interpret h_j: red_extend_highest_gen _ _ _ _ _ "i+j"
+ by (unfold_locales)
+ interpret h: extend_highest_gen "((moment i t)@s)" th prio tm "moment j (restm i t)"
+ proof(unfold_locales)
+ show "vt (moment i t @ s)" by (metis h_i.vt_t)
+ next
+ show "th \<in> threads (moment i t @ s)" by (metis h_i.th_kept)
+ next
+ show "preced th (moment i t @ s) =
+ Max (cp (moment i t @ s) ` threads (moment i t @ s))"
+ by (metis h_i.th_cp_max h_i.th_cp_preced h_i.th_kept)
+ next
+ show "preced th (moment i t @ s) = Prc prio tm" by (metis h_i.th_kept preced_th)
+ next
+ show "vt (moment j (restm i t) @ moment i t @ s)"
+ using moment_plus_split by (metis add.commute append_assoc h_j.vt_t)
+ next
+ fix th' prio'
+ assume "Create th' prio' \<in> set (moment j (restm i t))"
+ thus "prio' \<le> prio" using assms
+ by (metis Un_iff add.commute h_j.create_low moment_plus_split set_append)
+ next
+ fix th' prio'
+ assume "Set th' prio' \<in> set (moment j (restm i t))"
+ thus "th' \<noteq> th \<and> prio' \<le> prio"
+ by (metis Un_iff add.commute h_j.set_diff_low moment_plus_split set_append)
+ next
+ fix th'
+ assume "Exit th' \<in> set (moment j (restm i t))"
+ thus "th' \<noteq> th"
+ by (metis Un_iff add.commute h_j.exit_diff moment_plus_split set_append)
+ qed
+ show ?thesis
+ by (metis add.commute append_assoc eq_pv h.runing_precond_pre
+ moment_plus_split neq_th' th'_in)
+qed
+
+lemma moment_blocked_eqpv:
+ assumes neq_th': "th' \<noteq> th"
+ and th'_in: "th' \<in> threads ((moment i t)@s)"
+ and eq_pv: "cntP ((moment i t)@s) th' = cntV ((moment i t)@s) th'"
+ and le_ij: "i \<le> j"
+ shows "cntP ((moment j t)@s) th' = cntV ((moment j t)@s) th' \<and>
+ th' \<in> threads ((moment j t)@s) \<and>
+ th' \<notin> runing ((moment j t)@s)"
+proof -
+ from moment_blocked_pre [OF neq_th' th'_in eq_pv, of "j-i"] and le_ij
+ have h1: "cntP ((moment j t)@s) th' = cntV ((moment j t)@s) th'"
+ and h2: "th' \<in> threads ((moment j t)@s)" by auto
+ moreover have "th' \<notin> runing ((moment j t)@s)"
+ proof -
+ interpret h: red_extend_highest_gen _ _ _ _ _ j by (unfold_locales)
+ show ?thesis
+ using h.pv_blocked_pre h1 h2 neq_th' by auto
+ qed
+ ultimately show ?thesis by auto
+qed
+
+(* The foregoing two lemmas are preparation for this one, but
+ in long run can be combined. Maybe I am wrong.
+*)
+lemma moment_blocked:
+ assumes neq_th': "th' \<noteq> th"
+ and th'_in: "th' \<in> threads ((moment i t)@s)"
+ and dtc: "detached (moment i t @ s) th'"
+ and le_ij: "i \<le> j"
+ shows "detached (moment j t @ s) th' \<and>
+ th' \<in> threads ((moment j t)@s) \<and>
+ th' \<notin> runing ((moment j t)@s)"
+proof -
+ interpret h_i: red_extend_highest_gen _ _ _ _ _ i by (unfold_locales)
+ interpret h_j: red_extend_highest_gen _ _ _ _ _ j by (unfold_locales)
+ have cnt_i: "cntP (moment i t @ s) th' = cntV (moment i t @ s) th'"
+ by (metis dtc h_i.detached_elim)
+ from moment_blocked_eqpv[OF neq_th' th'_in cnt_i le_ij]
+ show ?thesis by (metis h_j.detached_intro)
+qed
+
+lemma runing_preced_inversion:
+ assumes runing': "th' \<in> runing (t@s)"
+ shows "cp (t@s) th' = preced th s" (is "?L = ?R")
+proof -
+ have "?L = Max (cp (t @ s) ` readys (t @ s))" using assms
+ by (unfold runing_def, auto)
+ also have "\<dots> = ?R"
+ by (metis th_cp_max th_cp_preced vat_t.max_cp_readys_threads)
+ finally show ?thesis .
+qed
+
+text {*
+ The situation when @{term "th"} is blocked is analyzed by the following lemmas.
+*}
+
+text {*
+ The following lemmas shows the running thread @{text "th'"}, if it is different from
+ @{term th}, must be live at the very beginning. By the term {\em the very beginning},
+ we mean the moment where the formal investigation starts, i.e. the moment (or state)
+ @{term s}.
+*}
+
+lemma runing_inversion_0:
+ assumes neq_th': "th' \<noteq> th"
+ and runing': "th' \<in> runing (t@s)"
+ shows "th' \<in> threads s"
+proof -
+ -- {* The proof is by contradiction: *}
+ { assume otherwise: "\<not> ?thesis"
+ have "th' \<notin> runing (t @ s)"
+ proof -
+ -- {* Since @{term "th'"} is running at time @{term "t@s"}, so it exists that time. *}
+ have th'_in: "th' \<in> threads (t@s)" using runing' by (simp add:runing_def readys_def)
+ -- {* However, @{text "th'"} does not exist at very beginning. *}
+ have th'_notin: "th' \<notin> threads (moment 0 t @ s)" using otherwise
+ by (metis append.simps(1) moment_zero)
+ -- {* Therefore, there must be a moment during @{text "t"}, when
+ @{text "th'"} came into being. *}
+ -- {* Let us suppose the moment being @{text "i"}: *}
+ from p_split_gen[OF th'_in th'_notin]
+ obtain i where lt_its: "i < length t"
+ and le_i: "0 \<le> i"
+ and pre: " th' \<notin> threads (moment i t @ s)" (is "th' \<notin> threads ?pre")
+ and post: "(\<forall>i'>i. th' \<in> threads (moment i' t @ s))" by (auto)
+ interpret h_i: red_extend_highest_gen _ _ _ _ _ i by (unfold_locales)
+ interpret h_i': red_extend_highest_gen _ _ _ _ _ "(Suc i)" by (unfold_locales)
+ from lt_its have "Suc i \<le> length t" by auto
+ -- {* Let us also suppose the event which makes this change is @{text e}: *}
+ from moment_head[OF this] obtain e where
+ eq_me: "moment (Suc i) t = e # moment i t" by blast
+ hence "vt (e # (moment i t @ s))" by (metis append_Cons h_i'.vt_t)
+ hence "PIP (moment i t @ s) e" by (cases, simp)
+ -- {* It can be derived that this event @{text "e"}, which
+ gives birth to @{term "th'"} must be a @{term "Create"}: *}
+ from create_pre[OF this, of th']
+ obtain prio where eq_e: "e = Create th' prio"
+ by (metis append_Cons eq_me lessI post pre)
+ have h1: "th' \<in> threads (moment (Suc i) t @ s)" using post by auto
+ have h2: "cntP (moment (Suc i) t @ s) th' = cntV (moment (Suc i) t@ s) th'"
+ proof -
+ have "cntP (moment i t@s) th' = cntV (moment i t@s) th'"
+ by (metis h_i.cnp_cnv_eq pre)
+ thus ?thesis by (simp add:eq_me eq_e cntP_def cntV_def count_def)
+ qed
+ show ?thesis
+ using moment_blocked_eqpv [OF neq_th' h1 h2, of "length t"] lt_its moment_ge
+ by auto
+ qed
+ with `th' \<in> runing (t@s)`
+ have False by simp
+ } thus ?thesis by auto
+qed
+
+text {*
+ The second lemma says, if the running thread @{text th'} is different from
+ @{term th}, then this @{text th'} must in the possession of some resources
+ at the very beginning.
+
+ To ease the reasoning of resource possession of one particular thread,
+ we used two auxiliary functions @{term cntV} and @{term cntP},
+ which are the counters of @{term P}-operations and
+ @{term V}-operations respectively.
+ If the number of @{term V}-operation is less than the number of
+ @{term "P"}-operations, the thread must have some unreleased resource.
+*}
+
+lemma runing_inversion_1: (* ddd *)
+ assumes neq_th': "th' \<noteq> th"
+ and runing': "th' \<in> runing (t@s)"
+ -- {* thread @{term "th'"} is a live on in state @{term "s"} and
+ it has some unreleased resource. *}
+ shows "th' \<in> threads s \<and> cntV s th' < cntP s th'"
+proof -
+ -- {* The proof is a simple composition of @{thm runing_inversion_0} and
+ @{thm runing_precond}: *}
+ -- {* By applying @{thm runing_inversion_0} to assumptions,
+ it can be shown that @{term th'} is live in state @{term s}: *}
+ have "th' \<in> threads s" using runing_inversion_0[OF assms(1,2)] .
+ -- {* Then the thesis is derived easily by applying @{thm runing_precond}: *}
+ with runing_precond [OF this neq_th' runing'] show ?thesis by simp
+qed
+
+text {*
+ The following lemma is just a rephrasing of @{thm runing_inversion_1}:
+*}
+lemma runing_inversion_2:
+ assumes runing': "th' \<in> runing (t@s)"
+ shows "th' = th \<or> (th' \<noteq> th \<and> th' \<in> threads s \<and> cntV s th' < cntP s th')"
+proof -
+ from runing_inversion_1[OF _ runing']
+ show ?thesis by auto
+qed
+
+lemma runing_inversion_3:
+ assumes runing': "th' \<in> runing (t@s)"
+ and neq_th: "th' \<noteq> th"
+ shows "th' \<in> threads s \<and> (cntV s th' < cntP s th' \<and> cp (t@s) th' = preced th s)"
+ by (metis neq_th runing' runing_inversion_2 runing_preced_inversion)
+
+lemma runing_inversion_4:
+ assumes runing': "th' \<in> runing (t@s)"
+ and neq_th: "th' \<noteq> th"
+ shows "th' \<in> threads s"
+ and "\<not>detached s th'"
+ and "cp (t@s) th' = preced th s"
+ apply (metis neq_th runing' runing_inversion_2)
+ apply (metis neq_th pv_blocked runing' runing_inversion_2 runing_precond_pre_dtc)
+ by (metis neq_th runing' runing_inversion_3)
+
+
+text {*
+ Suppose @{term th} is not running, it is first shown that
+ there is a path in RAG leading from node @{term th} to another thread @{text "th'"}
+ in the @{term readys}-set (So @{text "th'"} is an ancestor of @{term th}}).
+
+ Now, since @{term readys}-set is non-empty, there must be
+ one in it which holds the highest @{term cp}-value, which, by definition,
+ is the @{term runing}-thread. However, we are going to show more: this running thread
+ is exactly @{term "th'"}.
+ *}
+lemma th_blockedE: (* ddd *)
+ assumes "th \<notin> runing (t@s)"
+ obtains th' where "Th th' \<in> ancestors (RAG (t @ s)) (Th th)"
+ "th' \<in> runing (t@s)"
+proof -
+ -- {* According to @{thm vat_t.th_chain_to_ready}, either
+ @{term "th"} is in @{term "readys"} or there is path leading from it to
+ one thread in @{term "readys"}. *}
+ have "th \<in> readys (t @ s) \<or> (\<exists>th'. th' \<in> readys (t @ s) \<and> (Th th, Th th') \<in> (RAG (t @ s))\<^sup>+)"
+ using th_kept vat_t.th_chain_to_ready by auto
+ -- {* However, @{term th} can not be in @{term readys}, because otherwise, since
+ @{term th} holds the highest @{term cp}-value, it must be @{term "runing"}. *}
+ moreover have "th \<notin> readys (t@s)"
+ using assms runing_def th_cp_max vat_t.max_cp_readys_threads by auto
+ -- {* So, there must be a path from @{term th} to another thread @{text "th'"} in
+ term @{term readys}: *}
+ ultimately obtain th' where th'_in: "th' \<in> readys (t@s)"
+ and dp: "(Th th, Th th') \<in> (RAG (t @ s))\<^sup>+" by auto
+ -- {* We are going to show that this @{term th'} is running. *}
+ have "th' \<in> runing (t@s)"
+ proof -
+ -- {* We only need to show that this @{term th'} holds the highest @{term cp}-value: *}
+ have "cp (t@s) th' = Max (cp (t@s) ` readys (t@s))" (is "?L = ?R")
+ proof -
+ have "?L = Max ((the_preced (t @ s) \<circ> the_thread) ` subtree (tRAG (t @ s)) (Th th'))"
+ by (unfold cp_alt_def1, simp)
+ also have "... = (the_preced (t @ s) \<circ> the_thread) (Th th)"
+ proof(rule image_Max_subset)
+ show "finite (Th ` (threads (t@s)))" by (simp add: vat_t.finite_threads)
+ next
+ show "subtree (tRAG (t @ s)) (Th th') \<subseteq> Th ` threads (t @ s)"
+ by (metis Range.intros dp trancl_range vat_t.range_in vat_t.subtree_tRAG_thread)
+ next
+ show "Th th \<in> subtree (tRAG (t @ s)) (Th th')" using dp
+ by (unfold tRAG_subtree_eq, auto simp:subtree_def)
+ next
+ show "Max ((the_preced (t @ s) \<circ> the_thread) ` Th ` threads (t @ s)) =
+ (the_preced (t @ s) \<circ> the_thread) (Th th)" (is "Max ?L = _")
+ proof -
+ have "?L = the_preced (t @ s) ` threads (t @ s)"
+ by (unfold image_comp, rule image_cong, auto)
+ thus ?thesis using max_preced the_preced_def by auto
+ qed
+ qed
+ also have "... = ?R"
+ using th_cp_max th_cp_preced th_kept
+ the_preced_def vat_t.max_cp_readys_threads by auto
+ finally show ?thesis .
+ qed
+ -- {* Now, since @{term th'} holds the highest @{term cp}
+ and we have already show it is in @{term readys},
+ it is @{term runing} by definition. *}
+ with `th' \<in> readys (t@s)` show ?thesis by (simp add: runing_def)
+ qed
+ -- {* It is easy to show @{term th'} is an ancestor of @{term th}: *}
+ moreover have "Th th' \<in> ancestors (RAG (t @ s)) (Th th)"
+ using `(Th th, Th th') \<in> (RAG (t @ s))\<^sup>+` by (auto simp:ancestors_def)
+ ultimately show ?thesis using that by metis
+qed
+
+text {*
+ Now it is easy to see there is always a thread to run by case analysis
+ on whether thread @{term th} is running: if the answer is Yes, the
+ the running thread is obviously @{term th} itself; otherwise, the running
+ thread is the @{text th'} given by lemma @{thm th_blockedE}.
+*}
+lemma live: "runing (t@s) \<noteq> {}"
+proof(cases "th \<in> runing (t@s)")
+ case True thus ?thesis by auto
+next
+ case False
+ thus ?thesis using th_blockedE by auto
+qed
+
+end
+end
+
+
+
+>>>>>>> other
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/ExtGG.ty Thu Jan 28 16:33:49 2016 +0800
@@ -0,0 +1,922 @@
+theory ExtGG
+imports PrioG CpsG
+begin
+
+text {*
+ The following two auxiliary lemmas are used to reason about @{term Max}.
+*}
+lemma image_Max_eqI:
+ assumes "finite B"
+ and "b \<in> B"
+ and "\<forall> x \<in> B. f x \<le> f b"
+ shows "Max (f ` B) = f b"
+ using assms
+ using Max_eqI by blast
+
+lemma image_Max_subset:
+ assumes "finite A"
+ and "B \<subseteq> A"
+ and "a \<in> B"
+ and "Max (f ` A) = f a"
+ shows "Max (f ` B) = f a"
+proof(rule image_Max_eqI)
+ show "finite B"
+ using assms(1) assms(2) finite_subset by auto
+next
+ show "a \<in> B" using assms by simp
+next
+ show "\<forall>x\<in>B. f x \<le> f a"
+ by (metis Max_ge assms(1) assms(2) assms(4)
+ finite_imageI image_eqI subsetCE)
+qed
+
+text {*
+ The following locale @{text "highest_gen"} sets the basic context for our
+ investigation: supposing thread @{text th} holds the highest @{term cp}-value
+ in state @{text s}, which means the task for @{text th} is the
+ most urgent. We want to show that
+ @{text th} is treated correctly by PIP, which means
+ @{text th} will not be blocked unreasonably by other less urgent
+ threads.
+*}
+locale highest_gen =
+ fixes s th prio tm
+ assumes vt_s: "vt s"
+ and threads_s: "th \<in> threads s"
+ and highest: "preced th s = Max ((cp s)`threads s)"
+ -- {* The internal structure of @{term th}'s precedence is exposed:*}
+ and preced_th: "preced th s = Prc prio tm"
+
+-- {* @{term s} is a valid trace, so it will inherit all results derived for
+ a valid trace: *}
+sublocale highest_gen < vat_s: valid_trace "s"
+ by (unfold_locales, insert vt_s, simp)
+
+context highest_gen
+begin
+
+text {*
+ @{term tm} is the time when the precedence of @{term th} is set, so
+ @{term tm} must be a valid moment index into @{term s}.
+*}
+lemma lt_tm: "tm < length s"
+ by (insert preced_tm_lt[OF threads_s preced_th], simp)
+
+text {*
+ Since @{term th} holds the highest precedence and @{text "cp"}
+ is the highest precedence of all threads in the sub-tree of
+ @{text "th"} and @{text th} is among these threads,
+ its @{term cp} must equal to its precedence:
+*}
+lemma eq_cp_s_th: "cp s th = preced th s" (is "?L = ?R")
+proof -
+ have "?L \<le> ?R"
+ by (unfold highest, rule Max_ge,
+ auto simp:threads_s finite_threads)
+ moreover have "?R \<le> ?L"
+ by (unfold vat_s.cp_rec, rule Max_ge,
+ auto simp:the_preced_def vat_s.fsbttRAGs.finite_children)
+ ultimately show ?thesis by auto
+qed
+
+(* ccc *)
+lemma highest_cp_preced: "cp s th = Max ((\<lambda> th'. preced th' s) ` threads s)"
+ by (fold max_cp_eq, unfold eq_cp_s_th, insert highest, simp)
+
+lemma highest_preced_thread: "preced th s = Max ((\<lambda> th'. preced th' s) ` threads s)"
+ by (fold eq_cp_s_th, unfold highest_cp_preced, simp)
+
+lemma highest': "cp s th = Max (cp s ` threads s)"
+proof -
+ from highest_cp_preced max_cp_eq[symmetric]
+ show ?thesis by simp
+qed
+
+end
+
+locale extend_highest_gen = highest_gen +
+ fixes t
+ assumes vt_t: "vt (t@s)"
+ and create_low: "Create th' prio' \<in> set t \<Longrightarrow> prio' \<le> prio"
+ and set_diff_low: "Set th' prio' \<in> set t \<Longrightarrow> th' \<noteq> th \<and> prio' \<le> prio"
+ and exit_diff: "Exit th' \<in> set t \<Longrightarrow> th' \<noteq> th"
+
+sublocale extend_highest_gen < vat_t: valid_trace "t@s"
+ by (unfold_locales, insert vt_t, simp)
+
+lemma step_back_vt_app:
+ assumes vt_ts: "vt (t@s)"
+ shows "vt s"
+proof -
+ from vt_ts show ?thesis
+ proof(induct t)
+ case Nil
+ from Nil show ?case by auto
+ next
+ case (Cons e t)
+ assume ih: " vt (t @ s) \<Longrightarrow> vt s"
+ and vt_et: "vt ((e # t) @ s)"
+ show ?case
+ proof(rule ih)
+ show "vt (t @ s)"
+ proof(rule step_back_vt)
+ from vt_et show "vt (e # t @ s)" by simp
+ qed
+ qed
+ qed
+qed
+
+
+locale red_extend_highest_gen = extend_highest_gen +
+ fixes i::nat
+
+sublocale red_extend_highest_gen < red_moment: extend_highest_gen "s" "th" "prio" "tm" "(moment i t)"
+ apply (insert extend_highest_gen_axioms, subst (asm) (1) moment_restm_s [of i t, symmetric])
+ apply (unfold extend_highest_gen_def extend_highest_gen_axioms_def, clarsimp)
+ by (unfold highest_gen_def, auto dest:step_back_vt_app)
+
+
+context extend_highest_gen
+begin
+
+ lemma ind [consumes 0, case_names Nil Cons, induct type]:
+ assumes
+ h0: "R []"
+ and h2: "\<And> e t. \<lbrakk>vt (t@s); step (t@s) e;
+ extend_highest_gen s th prio tm t;
+ extend_highest_gen s th prio tm (e#t); R t\<rbrakk> \<Longrightarrow> R (e#t)"
+ shows "R t"
+proof -
+ from vt_t extend_highest_gen_axioms show ?thesis
+ proof(induct t)
+ from h0 show "R []" .
+ next
+ case (Cons e t')
+ assume ih: "\<lbrakk>vt (t' @ s); extend_highest_gen s th prio tm t'\<rbrakk> \<Longrightarrow> R t'"
+ and vt_e: "vt ((e # t') @ s)"
+ and et: "extend_highest_gen s th prio tm (e # t')"
+ from vt_e and step_back_step have stp: "step (t'@s) e" by auto
+ from vt_e and step_back_vt have vt_ts: "vt (t'@s)" by auto
+ show ?case
+ proof(rule h2 [OF vt_ts stp _ _ _ ])
+ show "R t'"
+ proof(rule ih)
+ from et show ext': "extend_highest_gen s th prio tm t'"
+ by (unfold extend_highest_gen_def extend_highest_gen_axioms_def, auto dest:step_back_vt)
+ next
+ from vt_ts show "vt (t' @ s)" .
+ qed
+ next
+ from et show "extend_highest_gen s th prio tm (e # t')" .
+ next
+ from et show ext': "extend_highest_gen s th prio tm t'"
+ by (unfold extend_highest_gen_def extend_highest_gen_axioms_def, auto dest:step_back_vt)
+ qed
+ qed
+qed
+
+
+lemma th_kept: "th \<in> threads (t @ s) \<and>
+ preced th (t@s) = preced th s" (is "?Q t")
+proof -
+ show ?thesis
+ proof(induct rule:ind)
+ case Nil
+ from threads_s
+ show ?case
+ by auto
+ next
+ case (Cons e t)
+ interpret h_e: extend_highest_gen _ _ _ _ "(e # t)" using Cons by auto
+ interpret h_t: extend_highest_gen _ _ _ _ t using Cons by auto
+ show ?case
+ proof(cases e)
+ case (Create thread prio)
+ show ?thesis
+ proof -
+ from Cons and Create have "step (t@s) (Create thread prio)" by auto
+ hence "th \<noteq> thread"
+ proof(cases)
+ case thread_create
+ with Cons show ?thesis by auto
+ qed
+ hence "preced th ((e # t) @ s) = preced th (t @ s)"
+ by (unfold Create, auto simp:preced_def)
+ moreover note Cons
+ ultimately show ?thesis
+ by (auto simp:Create)
+ qed
+ next
+ case (Exit thread)
+ from h_e.exit_diff and Exit
+ have neq_th: "thread \<noteq> th" by auto
+ with Cons
+ show ?thesis
+ by (unfold Exit, auto simp:preced_def)
+ next
+ case (P thread cs)
+ with Cons
+ show ?thesis
+ by (auto simp:P preced_def)
+ next
+ case (V thread cs)
+ with Cons
+ show ?thesis
+ by (auto simp:V preced_def)
+ next
+ case (Set thread prio')
+ show ?thesis
+ proof -
+ from h_e.set_diff_low and Set
+ have "th \<noteq> thread" by auto
+ hence "preced th ((e # t) @ s) = preced th (t @ s)"
+ by (unfold Set, auto simp:preced_def)
+ moreover note Cons
+ ultimately show ?thesis
+ by (auto simp:Set)
+ qed
+ qed
+ qed
+qed
+
+text {*
+ According to @{thm th_kept}, thread @{text "th"} has its living status
+ and precedence kept along the way of @{text "t"}. The following lemma
+ shows that this preserved precedence of @{text "th"} remains as the highest
+ along the way of @{text "t"}.
+
+ The proof goes by induction over @{text "t"} using the specialized
+ induction rule @{thm ind}, followed by case analysis of each possible
+ operations of PIP. All cases follow the same pattern rendered by the
+ generalized introduction rule @{thm "image_Max_eqI"}.
+
+ The very essence is to show that precedences, no matter whether they are newly introduced
+ or modified, are always lower than the one held by @{term "th"},
+ which by @{thm th_kept} is preserved along the way.
+*}
+lemma max_kept: "Max (the_preced (t @ s) ` (threads (t@s))) = preced th s"
+proof(induct rule:ind)
+ case Nil
+ from highest_preced_thread
+ show ?case
+ by (unfold the_preced_def, simp)
+next
+ case (Cons e t)
+ interpret h_e: extend_highest_gen _ _ _ _ "(e # t)" using Cons by auto
+ interpret h_t: extend_highest_gen _ _ _ _ t using Cons by auto
+ show ?case
+ proof(cases e)
+ case (Create thread prio')
+ show ?thesis (is "Max (?f ` ?A) = ?t")
+ proof -
+ -- {* The following is the common pattern of each branch of the case analysis. *}
+ -- {* The major part is to show that @{text "th"} holds the highest precedence: *}
+ have "Max (?f ` ?A) = ?f th"
+ proof(rule image_Max_eqI)
+ show "finite ?A" using h_e.finite_threads by auto
+ next
+ show "th \<in> ?A" using h_e.th_kept by auto
+ next
+ show "\<forall>x\<in>?A. ?f x \<le> ?f th"
+ proof
+ fix x
+ assume "x \<in> ?A"
+ hence "x = thread \<or> x \<in> threads (t@s)" by (auto simp:Create)
+ thus "?f x \<le> ?f th"
+ proof
+ assume "x = thread"
+ thus ?thesis
+ apply (simp add:Create the_preced_def preced_def, fold preced_def)
+ using Create h_e.create_low h_t.th_kept lt_tm preced_leI2 preced_th by force
+ next
+ assume h: "x \<in> threads (t @ s)"
+ from Cons(2)[unfolded Create]
+ have "x \<noteq> thread" using h by (cases, auto)
+ hence "?f x = the_preced (t@s) x"
+ by (simp add:Create the_preced_def preced_def)
+ hence "?f x \<le> Max (the_preced (t@s) ` threads (t@s))"
+ by (simp add: h_t.finite_threads h)
+ also have "... = ?f th"
+ by (metis Cons.hyps(5) h_e.th_kept the_preced_def)
+ finally show ?thesis .
+ qed
+ qed
+ qed
+ -- {* The minor part is to show that the precedence of @{text "th"}
+ equals to preserved one, given by the foregoing lemma @{thm th_kept} *}
+ also have "... = ?t" using h_e.th_kept the_preced_def by auto
+ -- {* Then it follows trivially that the precedence preserved
+ for @{term "th"} remains the maximum of all living threads along the way. *}
+ finally show ?thesis .
+ qed
+ next
+ case (Exit thread)
+ show ?thesis (is "Max (?f ` ?A) = ?t")
+ proof -
+ have "Max (?f ` ?A) = ?f th"
+ proof(rule image_Max_eqI)
+ show "finite ?A" using h_e.finite_threads by auto
+ next
+ show "th \<in> ?A" using h_e.th_kept by auto
+ next
+ show "\<forall>x\<in>?A. ?f x \<le> ?f th"
+ proof
+ fix x
+ assume "x \<in> ?A"
+ hence "x \<in> threads (t@s)" by (simp add: Exit)
+ hence "?f x \<le> Max (?f ` threads (t@s))"
+ by (simp add: h_t.finite_threads)
+ also have "... \<le> ?f th"
+ apply (simp add:Exit the_preced_def preced_def, fold preced_def)
+ using Cons.hyps(5) h_t.th_kept the_preced_def by auto
+ finally show "?f x \<le> ?f th" .
+ qed
+ qed
+ also have "... = ?t" using h_e.th_kept the_preced_def by auto
+ finally show ?thesis .
+ qed
+ next
+ case (P thread cs)
+ with Cons
+ show ?thesis by (auto simp:preced_def the_preced_def)
+ next
+ case (V thread cs)
+ with Cons
+ show ?thesis by (auto simp:preced_def the_preced_def)
+ next
+ case (Set thread prio')
+ show ?thesis (is "Max (?f ` ?A) = ?t")
+ proof -
+ have "Max (?f ` ?A) = ?f th"
+ proof(rule image_Max_eqI)
+ show "finite ?A" using h_e.finite_threads by auto
+ next
+ show "th \<in> ?A" using h_e.th_kept by auto
+ next
+ show "\<forall>x\<in>?A. ?f x \<le> ?f th"
+ proof
+ fix x
+ assume h: "x \<in> ?A"
+ show "?f x \<le> ?f th"
+ proof(cases "x = thread")
+ case True
+ moreover have "the_preced (Set thread prio' # t @ s) thread \<le> the_preced (t @ s) th"
+ proof -
+ have "the_preced (t @ s) th = Prc prio tm"
+ using h_t.th_kept preced_th by (simp add:the_preced_def)
+ moreover have "prio' \<le> prio" using Set h_e.set_diff_low by auto
+ ultimately show ?thesis by (insert lt_tm, auto simp:the_preced_def preced_def)
+ qed
+ ultimately show ?thesis
+ by (unfold Set, simp add:the_preced_def preced_def)
+ next
+ case False
+ then have "?f x = the_preced (t@s) x"
+ by (simp add:the_preced_def preced_def Set)
+ also have "... \<le> Max (the_preced (t@s) ` threads (t@s))"
+ using Set h h_t.finite_threads by auto
+ also have "... = ?f th" by (metis Cons.hyps(5) h_e.th_kept the_preced_def)
+ finally show ?thesis .
+ qed
+ qed
+ qed
+ also have "... = ?t" using h_e.th_kept the_preced_def by auto
+ finally show ?thesis .
+ qed
+ qed
+qed
+
+lemma max_preced: "preced th (t@s) = Max (the_preced (t@s) ` (threads (t@s)))"
+ by (insert th_kept max_kept, auto)
+
+text {*
+ The reason behind the following lemma is that:
+ Since @{term "cp"} is defined as the maximum precedence
+ of those threads contained in the sub-tree of node @{term "Th th"}
+ in @{term "RAG (t@s)"}, and all these threads are living threads, and
+ @{term "th"} is also among them, the maximum precedence of
+ them all must be the one for @{text "th"}.
+*}
+lemma th_cp_max_preced:
+ "cp (t@s) th = Max (the_preced (t@s) ` (threads (t@s)))" (is "?L = ?R")
+proof -
+ let ?f = "the_preced (t@s)"
+ have "?L = ?f th"
+ proof(unfold cp_alt_def, rule image_Max_eqI)
+ show "finite {th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}"
+ proof -
+ have "{th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)} =
+ the_thread ` {n . n \<in> subtree (RAG (t @ s)) (Th th) \<and>
+ (\<exists> th'. n = Th th')}"
+ by (smt Collect_cong Setcompr_eq_image mem_Collect_eq the_thread.simps)
+ moreover have "finite ..." by (simp add: vat_t.fsbtRAGs.finite_subtree)
+ ultimately show ?thesis by simp
+ qed
+ next
+ show "th \<in> {th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}"
+ by (auto simp:subtree_def)
+ next
+ show "\<forall>x\<in>{th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}.
+ the_preced (t @ s) x \<le> the_preced (t @ s) th"
+ proof
+ fix th'
+ assume "th' \<in> {th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}"
+ hence "Th th' \<in> subtree (RAG (t @ s)) (Th th)" by auto
+ moreover have "... \<subseteq> Field (RAG (t @ s)) \<union> {Th th}"
+ by (meson subtree_Field)
+ ultimately have "Th th' \<in> ..." by auto
+ hence "th' \<in> threads (t@s)"
+ proof
+ assume "Th th' \<in> {Th th}"
+ thus ?thesis using th_kept by auto
+ next
+ assume "Th th' \<in> Field (RAG (t @ s))"
+ thus ?thesis using vat_t.not_in_thread_isolated by blast
+ qed
+ thus "the_preced (t @ s) th' \<le> the_preced (t @ s) th"
+ by (metis Max_ge finite_imageI finite_threads image_eqI
+ max_kept th_kept the_preced_def)
+ qed
+ qed
+ also have "... = ?R" by (simp add: max_preced the_preced_def)
+ finally show ?thesis .
+qed
+
+lemma th_cp_max: "cp (t@s) th = Max (cp (t@s) ` threads (t@s))"
+ using max_cp_eq th_cp_max_preced the_preced_def vt_t by presburger
+
+lemma th_cp_preced: "cp (t@s) th = preced th s"
+ by (fold max_kept, unfold th_cp_max_preced, simp)
+
+lemma preced_less:
+ assumes th'_in: "th' \<in> threads s"
+ and neq_th': "th' \<noteq> th"
+ shows "preced th' s < preced th s"
+ using assms
+by (metis Max.coboundedI finite_imageI highest not_le order.trans
+ preced_linorder rev_image_eqI threads_s vat_s.finite_threads
+ vat_s.le_cp)
+
+text {*
+ Counting of the number of @{term "P"} and @{term "V"} operations
+ is the cornerstone of a large number of the following proofs.
+ The reason is that this counting is quite easy to calculate and
+ convenient to use in the reasoning.
+
+ The following lemma shows that the counting controls whether
+ a thread is running or not.
+*}
+
+lemma pv_blocked_pre:
+ assumes th'_in: "th' \<in> threads (t@s)"
+ and neq_th': "th' \<noteq> th"
+ and eq_pv: "cntP (t@s) th' = cntV (t@s) th'"
+ shows "th' \<notin> runing (t@s)"
+proof
+ assume otherwise: "th' \<in> runing (t@s)"
+ show False
+ proof -
+ have "th' = th"
+ proof(rule preced_unique)
+ show "preced th' (t @ s) = preced th (t @ s)" (is "?L = ?R")
+ proof -
+ have "?L = cp (t@s) th'"
+ by (unfold cp_eq_cpreced cpreced_def count_eq_dependants[OF eq_pv], simp)
+ also have "... = cp (t @ s) th" using otherwise
+ by (metis (mono_tags, lifting) mem_Collect_eq
+ runing_def th_cp_max vat_t.max_cp_readys_threads)
+ also have "... = ?R" by (metis th_cp_preced th_kept)
+ finally show ?thesis .
+ qed
+ qed (auto simp: th'_in th_kept)
+ moreover have "th' \<noteq> th" using neq_th' .
+ ultimately show ?thesis by simp
+ qed
+qed
+
+lemmas pv_blocked = pv_blocked_pre[folded detached_eq]
+
+lemma runing_precond_pre:
+ fixes th'
+ assumes th'_in: "th' \<in> threads s"
+ and eq_pv: "cntP s th' = cntV s th'"
+ and neq_th': "th' \<noteq> th"
+ shows "th' \<in> threads (t@s) \<and>
+ cntP (t@s) th' = cntV (t@s) th'"
+proof(induct rule:ind)
+ case (Cons e t)
+ interpret vat_t: extend_highest_gen s th prio tm t using Cons by simp
+ interpret vat_e: extend_highest_gen s th prio tm "(e # t)" using Cons by simp
+ show ?case
+ proof(cases e)
+ case (P thread cs)
+ show ?thesis
+ proof -
+ have "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
+ proof -
+ have "thread \<noteq> th'"
+ proof -
+ have "step (t@s) (P thread cs)" using Cons P by auto
+ thus ?thesis
+ proof(cases)
+ assume "thread \<in> runing (t@s)"
+ moreover have "th' \<notin> runing (t@s)" using Cons(5)
+ by (metis neq_th' vat_t.pv_blocked_pre)
+ ultimately show ?thesis by auto
+ qed
+ qed with Cons show ?thesis
+ by (unfold P, simp add:cntP_def cntV_def count_def)
+ qed
+ moreover have "th' \<in> threads ((e # t) @ s)" using Cons by (unfold P, simp)
+ ultimately show ?thesis by auto
+ qed
+ next
+ case (V thread cs)
+ show ?thesis
+ proof -
+ have "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
+ proof -
+ have "thread \<noteq> th'"
+ proof -
+ have "step (t@s) (V thread cs)" using Cons V by auto
+ thus ?thesis
+ proof(cases)
+ assume "thread \<in> runing (t@s)"
+ moreover have "th' \<notin> runing (t@s)" using Cons(5)
+ by (metis neq_th' vat_t.pv_blocked_pre)
+ ultimately show ?thesis by auto
+ qed
+ qed with Cons show ?thesis
+ by (unfold V, simp add:cntP_def cntV_def count_def)
+ qed
+ moreover have "th' \<in> threads ((e # t) @ s)" using Cons by (unfold V, simp)
+ ultimately show ?thesis by auto
+ qed
+ next
+ case (Create thread prio')
+ show ?thesis
+ proof -
+ have "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
+ proof -
+ have "thread \<noteq> th'"
+ proof -
+ have "step (t@s) (Create thread prio')" using Cons Create by auto
+ thus ?thesis using Cons(5) by (cases, auto)
+ qed with Cons show ?thesis
+ by (unfold Create, simp add:cntP_def cntV_def count_def)
+ qed
+ moreover have "th' \<in> threads ((e # t) @ s)" using Cons by (unfold Create, simp)
+ ultimately show ?thesis by auto
+ qed
+ next
+ case (Exit thread)
+ show ?thesis
+ proof -
+ have neq_thread: "thread \<noteq> th'"
+ proof -
+ have "step (t@s) (Exit thread)" using Cons Exit by auto
+ thus ?thesis apply (cases) using Cons(5)
+ by (metis neq_th' vat_t.pv_blocked_pre)
+ qed
+ hence "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'" using Cons
+ by (unfold Exit, simp add:cntP_def cntV_def count_def)
+ moreover have "th' \<in> threads ((e # t) @ s)" using Cons neq_thread
+ by (unfold Exit, simp)
+ ultimately show ?thesis by auto
+ qed
+ next
+ case (Set thread prio')
+ with Cons
+ show ?thesis
+ by (auto simp:cntP_def cntV_def count_def)
+ qed
+next
+ case Nil
+ with assms
+ show ?case by auto
+qed
+
+text {* Changing counting balance to detachedness *}
+lemmas runing_precond_pre_dtc = runing_precond_pre
+ [folded vat_t.detached_eq vat_s.detached_eq]
+
+lemma runing_precond:
+ fixes th'
+ assumes th'_in: "th' \<in> threads s"
+ and neq_th': "th' \<noteq> th"
+ and is_runing: "th' \<in> runing (t@s)"
+ shows "cntP s th' > cntV s th'"
+ using assms
+proof -
+ have "cntP s th' \<noteq> cntV s th'"
+ by (metis is_runing neq_th' pv_blocked_pre runing_precond_pre th'_in)
+ moreover have "cntV s th' \<le> cntP s th'" using vat_s.cnp_cnv_cncs by auto
+ ultimately show ?thesis by auto
+qed
+
+lemma moment_blocked_pre:
+ assumes neq_th': "th' \<noteq> th"
+ and th'_in: "th' \<in> threads ((moment i t)@s)"
+ and eq_pv: "cntP ((moment i t)@s) th' = cntV ((moment i t)@s) th'"
+ shows "cntP ((moment (i+j) t)@s) th' = cntV ((moment (i+j) t)@s) th' \<and>
+ th' \<in> threads ((moment (i+j) t)@s)"
+proof -
+ interpret h_i: red_extend_highest_gen _ _ _ _ _ i
+ by (unfold_locales)
+ interpret h_j: red_extend_highest_gen _ _ _ _ _ "i+j"
+ by (unfold_locales)
+ interpret h: extend_highest_gen "((moment i t)@s)" th prio tm "moment j (restm i t)"
+ proof(unfold_locales)
+ show "vt (moment i t @ s)" by (metis h_i.vt_t)
+ next
+ show "th \<in> threads (moment i t @ s)" by (metis h_i.th_kept)
+ next
+ show "preced th (moment i t @ s) =
+ Max (cp (moment i t @ s) ` threads (moment i t @ s))"
+ by (metis h_i.th_cp_max h_i.th_cp_preced h_i.th_kept)
+ next
+ show "preced th (moment i t @ s) = Prc prio tm" by (metis h_i.th_kept preced_th)
+ next
+ show "vt (moment j (restm i t) @ moment i t @ s)"
+ using moment_plus_split by (metis add.commute append_assoc h_j.vt_t)
+ next
+ fix th' prio'
+ assume "Create th' prio' \<in> set (moment j (restm i t))"
+ thus "prio' \<le> prio" using assms
+ by (metis Un_iff add.commute h_j.create_low moment_plus_split set_append)
+ next
+ fix th' prio'
+ assume "Set th' prio' \<in> set (moment j (restm i t))"
+ thus "th' \<noteq> th \<and> prio' \<le> prio"
+ by (metis Un_iff add.commute h_j.set_diff_low moment_plus_split set_append)
+ next
+ fix th'
+ assume "Exit th' \<in> set (moment j (restm i t))"
+ thus "th' \<noteq> th"
+ by (metis Un_iff add.commute h_j.exit_diff moment_plus_split set_append)
+ qed
+ show ?thesis
+ by (metis add.commute append_assoc eq_pv h.runing_precond_pre
+ moment_plus_split neq_th' th'_in)
+qed
+
+lemma moment_blocked_eqpv:
+ assumes neq_th': "th' \<noteq> th"
+ and th'_in: "th' \<in> threads ((moment i t)@s)"
+ and eq_pv: "cntP ((moment i t)@s) th' = cntV ((moment i t)@s) th'"
+ and le_ij: "i \<le> j"
+ shows "cntP ((moment j t)@s) th' = cntV ((moment j t)@s) th' \<and>
+ th' \<in> threads ((moment j t)@s) \<and>
+ th' \<notin> runing ((moment j t)@s)"
+proof -
+ from moment_blocked_pre [OF neq_th' th'_in eq_pv, of "j-i"] and le_ij
+ have h1: "cntP ((moment j t)@s) th' = cntV ((moment j t)@s) th'"
+ and h2: "th' \<in> threads ((moment j t)@s)" by auto
+ moreover have "th' \<notin> runing ((moment j t)@s)"
+ proof -
+ interpret h: red_extend_highest_gen _ _ _ _ _ j by (unfold_locales)
+ show ?thesis
+ using h.pv_blocked_pre h1 h2 neq_th' by auto
+ qed
+ ultimately show ?thesis by auto
+qed
+
+(* The foregoing two lemmas are preparation for this one, but
+ in long run can be combined. Maybe I am wrong.
+*)
+lemma moment_blocked:
+ assumes neq_th': "th' \<noteq> th"
+ and th'_in: "th' \<in> threads ((moment i t)@s)"
+ and dtc: "detached (moment i t @ s) th'"
+ and le_ij: "i \<le> j"
+ shows "detached (moment j t @ s) th' \<and>
+ th' \<in> threads ((moment j t)@s) \<and>
+ th' \<notin> runing ((moment j t)@s)"
+proof -
+ interpret h_i: red_extend_highest_gen _ _ _ _ _ i by (unfold_locales)
+ interpret h_j: red_extend_highest_gen _ _ _ _ _ j by (unfold_locales)
+ have cnt_i: "cntP (moment i t @ s) th' = cntV (moment i t @ s) th'"
+ by (metis dtc h_i.detached_elim)
+ from moment_blocked_eqpv[OF neq_th' th'_in cnt_i le_ij]
+ show ?thesis by (metis h_j.detached_intro)
+qed
+
+lemma runing_preced_inversion:
+ assumes runing': "th' \<in> runing (t@s)"
+ shows "cp (t@s) th' = preced th s" (is "?L = ?R")
+proof -
+ have "?L = Max (cp (t @ s) ` readys (t @ s))" using assms
+ by (unfold runing_def, auto)
+ also have "\<dots> = ?R"
+ by (metis th_cp_max th_cp_preced vat_t.max_cp_readys_threads)
+ finally show ?thesis .
+qed
+
+text {*
+ The situation when @{term "th"} is blocked is analyzed by the following lemmas.
+*}
+
+text {*
+ The following lemmas shows the running thread @{text "th'"}, if it is different from
+ @{term th}, must be live at the very beginning. By the term {\em the very beginning},
+ we mean the moment where the formal investigation starts, i.e. the moment (or state)
+ @{term s}.
+*}
+
+lemma runing_inversion_0:
+ assumes neq_th': "th' \<noteq> th"
+ and runing': "th' \<in> runing (t@s)"
+ shows "th' \<in> threads s"
+proof -
+ -- {* The proof is by contradiction: *}
+ { assume otherwise: "\<not> ?thesis"
+ have "th' \<notin> runing (t @ s)"
+ proof -
+ -- {* Since @{term "th'"} is running at time @{term "t@s"}, so it exists that time. *}
+ have th'_in: "th' \<in> threads (t@s)" using runing' by (simp add:runing_def readys_def)
+ -- {* However, @{text "th'"} does not exist at very beginning. *}
+ have th'_notin: "th' \<notin> threads (moment 0 t @ s)" using otherwise
+ by (metis append.simps(1) moment_zero)
+ -- {* Therefore, there must be a moment during @{text "t"}, when
+ @{text "th'"} came into being. *}
+ -- {* Let us suppose the moment being @{text "i"}: *}
+ from p_split_gen[OF th'_in th'_notin]
+ obtain i where lt_its: "i < length t"
+ and le_i: "0 \<le> i"
+ and pre: " th' \<notin> threads (moment i t @ s)" (is "th' \<notin> threads ?pre")
+ and post: "(\<forall>i'>i. th' \<in> threads (moment i' t @ s))" by (auto)
+ interpret h_i: red_extend_highest_gen _ _ _ _ _ i by (unfold_locales)
+ interpret h_i': red_extend_highest_gen _ _ _ _ _ "(Suc i)" by (unfold_locales)
+ from lt_its have "Suc i \<le> length t" by auto
+ -- {* Let us also suppose the event which makes this change is @{text e}: *}
+ from moment_head[OF this] obtain e where
+ eq_me: "moment (Suc i) t = e # moment i t" by blast
+ hence "vt (e # (moment i t @ s))" by (metis append_Cons h_i'.vt_t)
+ hence "PIP (moment i t @ s) e" by (cases, simp)
+ -- {* It can be derived that this event @{text "e"}, which
+ gives birth to @{term "th'"} must be a @{term "Create"}: *}
+ from create_pre[OF this, of th']
+ obtain prio where eq_e: "e = Create th' prio"
+ by (metis append_Cons eq_me lessI post pre)
+ have h1: "th' \<in> threads (moment (Suc i) t @ s)" using post by auto
+ have h2: "cntP (moment (Suc i) t @ s) th' = cntV (moment (Suc i) t@ s) th'"
+ proof -
+ have "cntP (moment i t@s) th' = cntV (moment i t@s) th'"
+ by (metis h_i.cnp_cnv_eq pre)
+ thus ?thesis by (simp add:eq_me eq_e cntP_def cntV_def count_def)
+ qed
+ show ?thesis
+ using moment_blocked_eqpv [OF neq_th' h1 h2, of "length t"] lt_its moment_ge
+ by auto
+ qed
+ with `th' \<in> runing (t@s)`
+ have False by simp
+ } thus ?thesis by auto
+qed
+
+text {*
+ The second lemma says, if the running thread @{text th'} is different from
+ @{term th}, then this @{text th'} must in the possession of some resources
+ at the very beginning.
+
+ To ease the reasoning of resource possession of one particular thread,
+ we used two auxiliary functions @{term cntV} and @{term cntP},
+ which are the counters of @{term P}-operations and
+ @{term V}-operations respectively.
+ If the number of @{term V}-operation is less than the number of
+ @{term "P"}-operations, the thread must have some unreleased resource.
+*}
+
+lemma runing_inversion_1: (* ddd *)
+ assumes neq_th': "th' \<noteq> th"
+ and runing': "th' \<in> runing (t@s)"
+ -- {* thread @{term "th'"} is a live on in state @{term "s"} and
+ it has some unreleased resource. *}
+ shows "th' \<in> threads s \<and> cntV s th' < cntP s th'"
+proof -
+ -- {* The proof is a simple composition of @{thm runing_inversion_0} and
+ @{thm runing_precond}: *}
+ -- {* By applying @{thm runing_inversion_0} to assumptions,
+ it can be shown that @{term th'} is live in state @{term s}: *}
+ have "th' \<in> threads s" using runing_inversion_0[OF assms(1,2)] .
+ -- {* Then the thesis is derived easily by applying @{thm runing_precond}: *}
+ with runing_precond [OF this neq_th' runing'] show ?thesis by simp
+qed
+
+text {*
+ The following lemma is just a rephrasing of @{thm runing_inversion_1}:
+*}
+lemma runing_inversion_2:
+ assumes runing': "th' \<in> runing (t@s)"
+ shows "th' = th \<or> (th' \<noteq> th \<and> th' \<in> threads s \<and> cntV s th' < cntP s th')"
+proof -
+ from runing_inversion_1[OF _ runing']
+ show ?thesis by auto
+qed
+
+lemma runing_inversion_3:
+ assumes runing': "th' \<in> runing (t@s)"
+ and neq_th: "th' \<noteq> th"
+ shows "th' \<in> threads s \<and> (cntV s th' < cntP s th' \<and> cp (t@s) th' = preced th s)"
+ by (metis neq_th runing' runing_inversion_2 runing_preced_inversion)
+
+lemma runing_inversion_4:
+ assumes runing': "th' \<in> runing (t@s)"
+ and neq_th: "th' \<noteq> th"
+ shows "th' \<in> threads s"
+ and "\<not>detached s th'"
+ and "cp (t@s) th' = preced th s"
+ apply (metis neq_th runing' runing_inversion_2)
+ apply (metis neq_th pv_blocked runing' runing_inversion_2 runing_precond_pre_dtc)
+ by (metis neq_th runing' runing_inversion_3)
+
+
+text {*
+ Suppose @{term th} is not running, it is first shown that
+ there is a path in RAG leading from node @{term th} to another thread @{text "th'"}
+ in the @{term readys}-set (So @{text "th'"} is an ancestor of @{term th}}).
+
+ Now, since @{term readys}-set is non-empty, there must be
+ one in it which holds the highest @{term cp}-value, which, by definition,
+ is the @{term runing}-thread. However, we are going to show more: this running thread
+ is exactly @{term "th'"}.
+ *}
+lemma th_blockedE: (* ddd *)
+ assumes "th \<notin> runing (t@s)"
+ obtains th' where "Th th' \<in> ancestors (RAG (t @ s)) (Th th)"
+ "th' \<in> runing (t@s)"
+proof -
+ -- {* According to @{thm vat_t.th_chain_to_ready}, either
+ @{term "th"} is in @{term "readys"} or there is path leading from it to
+ one thread in @{term "readys"}. *}
+ have "th \<in> readys (t @ s) \<or> (\<exists>th'. th' \<in> readys (t @ s) \<and> (Th th, Th th') \<in> (RAG (t @ s))\<^sup>+)"
+ using th_kept vat_t.th_chain_to_ready by auto
+ -- {* However, @{term th} can not be in @{term readys}, because otherwise, since
+ @{term th} holds the highest @{term cp}-value, it must be @{term "runing"}. *}
+ moreover have "th \<notin> readys (t@s)"
+ using assms runing_def th_cp_max vat_t.max_cp_readys_threads by auto
+ -- {* So, there must be a path from @{term th} to another thread @{text "th'"} in
+ term @{term readys}: *}
+ ultimately obtain th' where th'_in: "th' \<in> readys (t@s)"
+ and dp: "(Th th, Th th') \<in> (RAG (t @ s))\<^sup>+" by auto
+ -- {* We are going to show that this @{term th'} is running. *}
+ have "th' \<in> runing (t@s)"
+ proof -
+ -- {* We only need to show that this @{term th'} holds the highest @{term cp}-value: *}
+ have "cp (t@s) th' = Max (cp (t@s) ` readys (t@s))" (is "?L = ?R")
+ proof -
+ have "?L = Max ((the_preced (t @ s) \<circ> the_thread) ` subtree (tRAG (t @ s)) (Th th'))"
+ by (unfold cp_alt_def1, simp)
+ also have "... = (the_preced (t @ s) \<circ> the_thread) (Th th)"
+ proof(rule image_Max_subset)
+ show "finite (Th ` (threads (t@s)))" by (simp add: vat_t.finite_threads)
+ next
+ show "subtree (tRAG (t @ s)) (Th th') \<subseteq> Th ` threads (t @ s)"
+ by (metis Range.intros dp trancl_range vat_t.range_in vat_t.subtree_tRAG_thread)
+ next
+ show "Th th \<in> subtree (tRAG (t @ s)) (Th th')" using dp
+ by (unfold tRAG_subtree_eq, auto simp:subtree_def)
+ next
+ show "Max ((the_preced (t @ s) \<circ> the_thread) ` Th ` threads (t @ s)) =
+ (the_preced (t @ s) \<circ> the_thread) (Th th)" (is "Max ?L = _")
+ proof -
+ have "?L = the_preced (t @ s) ` threads (t @ s)"
+ by (unfold image_comp, rule image_cong, auto)
+ thus ?thesis using max_preced the_preced_def by auto
+ qed
+ qed
+ also have "... = ?R"
+ using th_cp_max th_cp_preced th_kept
+ the_preced_def vat_t.max_cp_readys_threads by auto
+ finally show ?thesis .
+ qed
+ -- {* Now, since @{term th'} holds the highest @{term cp}
+ and we have already show it is in @{term readys},
+ it is @{term runing} by definition. *}
+ with `th' \<in> readys (t@s)` show ?thesis by (simp add: runing_def)
+ qed
+ -- {* It is easy to show @{term th'} is an ancestor of @{term th}: *}
+ moreover have "Th th' \<in> ancestors (RAG (t @ s)) (Th th)"
+ using `(Th th, Th th') \<in> (RAG (t @ s))\<^sup>+` by (auto simp:ancestors_def)
+ ultimately show ?thesis using that by metis
+qed
+
+text {*
+ Now it is easy to see there is always a thread to run by case analysis
+ on whether thread @{term th} is running: if the answer is Yes, the
+ the running thread is obviously @{term th} itself; otherwise, the running
+ thread is the @{text th'} given by lemma @{thm th_blockedE}.
+*}
+lemma live: "runing (t@s) \<noteq> {}"
+proof(cases "th \<in> runing (t@s)")
+ case True thus ?thesis by auto
+next
+ case False
+ thus ?thesis using th_blockedE by auto
+qed
+
+end
+end
+
+
+
--- a/Journal/Paper.thy Thu Jan 28 07:43:05 2016 +0800
+++ b/Journal/Paper.thy Thu Jan 28 16:33:49 2016 +0800
@@ -784,51 +784,56 @@
begin
(*>*)
text {*
+
Sha et al.~state their first correctness criterion for PIP in terms
of the number of low-priority threads \cite[Theorem 3]{Sha90}: if
there are @{text n} low-priority threads, then a blocked job with
high priority can only be blocked a maximum of @{text n} times.
- Their second correctness criterion is given
- in terms of the number of critical resources \cite[Theorem 6]{Sha90}: if there are
- @{text m} critical resources, then a blocked job with high priority
- can only be blocked a maximum of @{text m} times. Both results on their own, strictly speaking, do
- \emph{not} prevent indefinite, or unbounded, Priority Inversion,
- because if a low-priority thread does not give up its critical
- resource (the one the high-priority thread is waiting for), then the
- high-priority thread can never run. The argument of Sha et al.~is
- that \emph{if} threads release locked resources in a finite amount
- of time, then indefinite Priority Inversion cannot occur---the high-priority
- thread is guaranteed to run eventually. The assumption is that
- programmers must ensure that threads are programmed in this way. However, even
- taking this assumption into account, the correctness properties of
- Sha et al.~are
- \emph{not} true for their version of PIP---despite being ``proved''. As Yodaiken
- \cite{Yodaiken02} and Moylan et al.~\cite{deinheritance} pointed out: If a low-priority thread possesses
- locks to two resources for which two high-priority threads are
- waiting for, then lowering the priority prematurely after giving up
- only one lock, can cause indefinite Priority Inversion for one of the
- high-priority threads, invalidating their two bounds.
+ Their second correctness criterion is given in terms of the number
+ of critical resources \cite[Theorem 6]{Sha90}: if there are @{text
+ m} critical resources, then a blocked job with high priority can
+ only be blocked a maximum of @{text m} times. Both results on their
+ own, strictly speaking, do \emph{not} prevent indefinite, or
+ unbounded, Priority Inversion, because if a low-priority thread does
+ not give up its critical resource (the one the high-priority thread
+ is waiting for), then the high-priority thread can never run. The
+ argument of Sha et al.~is that \emph{if} threads release locked
+ resources in a finite amount of time, then indefinite Priority
+ Inversion cannot occur---the high-priority thread is guaranteed to
+ run eventually. The assumption is that programmers must ensure that
+ threads are programmed in this way. However, even taking this
+ assumption into account, the correctness properties of Sha et
+ al.~are \emph{not} true for their version of PIP---despite being
+ ``proved''. As Yodaiken \cite{Yodaiken02} and Moylan et
+ al.~\cite{deinheritance} pointed out: If a low-priority thread
+ possesses locks to two resources for which two high-priority threads
+ are waiting for, then lowering the priority prematurely after giving
+ up only one lock, can cause indefinite Priority Inversion for one of
+ the high-priority threads, invalidating their two bounds.
Even when fixed, their proof idea does not seem to go through for
us, because of the way we have set up our formal model of PIP. One
- reason is that we allow critical sections, which start with a @{text P}-event
- and finish with a corresponding @{text V}-event, to arbitrarily overlap
- (something Sha et al.~explicitly exclude). Therefore we have
- designed a different correctness criterion for PIP. The idea behind
- our criterion is as follows: for all states @{text s}, we know the
- corresponding thread @{text th} with the highest precedence; we show
- that in every future state (denoted by @{text "s' @ s"}) in which
- @{text th} is still alive, either @{text th} is running or it is
- blocked by a thread that was alive in the state @{text s} and was waiting
- for or in the possession of a lock in @{text s}. Since in @{text s}, as in
- every state, the set of alive threads is finite, @{text th} can only
- be blocked a finite number of times. This is independent of how many
- threads of lower priority are created in @{text "s'"}. We will actually prove a
+ reason is that we allow critical sections, which start with a @{text
+ P}-event and finish with a corresponding @{text V}-event, to
+ arbitrarily overlap (something Sha et al.~explicitly exclude).
+ Therefore we have designed a different correctness criterion for
+ PIP. The idea behind our criterion is as follows: for all states
+ @{text s}, we know the corresponding thread @{text th} with the
+ highest precedence; we show that in every future state (denoted by
+ @{text "s' @ s"}) in which @{text th} is still alive, either @{text
+ th} is running or it is blocked by a thread that was alive in the
+ state @{text s} and was waiting for or in the possession of a lock
+ in @{text s}. Since in @{text s}, as in every state, the set of
+ alive threads is finite, @{text th} can only be blocked a finite
+ number of times. This is independent of how many threads of lower
+ priority are created in @{text "s'"}. We will actually prove a
stronger statement where we also provide the current precedence of
- the blocking thread. However, this correctness criterion hinges upon
- a number of assumptions about the states @{text s} and @{text "s' @
- s"}, the thread @{text th} and the events happening in @{text
- s'}. We list them next:
+ the blocking thread.
+
+ However, this correctness criterion hinges upon a number of
+ assumptions about the states @{text s} and @{text "s' @ s"}, the
+ thread @{text th} and the events happening in @{text s'}. We list
+ them next:
\begin{quote}
{\bf Assumptions on the states {\boldmath@{text s}} and
@@ -872,9 +877,9 @@
\end{isabelle}
\end{quote}
- \noindent
- The locale mechanism of Isabelle helps us to manage conveniently such assumptions~\cite{Haftmann08}.
- Under these assumptions we shall prove the following correctness property:
+ \noindent The locale mechanism of Isabelle helps us to manage
+ conveniently such assumptions~\cite{Haftmann08}. Under these
+ assumptions we shall prove the following correctness property:
\begin{theorem}\label{mainthm}
Given the assumptions about states @{text "s"} and @{text "s' @ s"},
@@ -884,29 +889,28 @@
@{term "cp (s' @ s) th' = prec th s"}.
\end{theorem}
- \noindent
- This theorem ensures that the thread @{text th}, which has the
- highest precedence in the state @{text s}, can only be blocked in
- the state @{text "s' @ s"} by a thread @{text th'} that already
- existed in @{text s} and requested or had a lock on at least
- one resource---that means the thread was not \emph{detached} in @{text s}.
- As we shall see shortly, that means there are only finitely
- many threads that can block @{text th} in this way and then they
+ \noindent This theorem ensures that the thread @{text th}, which has
+ the highest precedence in the state @{text s}, can only be blocked
+ in the state @{text "s' @ s"} by a thread @{text th'} that already
+ existed in @{text s} and requested or had a lock on at least one
+ resource---that means the thread was not \emph{detached} in @{text
+ s}. As we shall see shortly, that means there are only finitely
+ many threads that can block @{text th} in this way and then they
need to run with the same precedence as @{text th}.
Like in the argument by Sha et al.~our finite bound does not
guarantee absence of indefinite Priority Inversion. For this we
further have to assume that every thread gives up its resources
after a finite amount of time. We found that this assumption is
- awkward to formalise in our model. There are mainly two reasons for this:
- First, we do not specify what ``running'' the code of a thread
+ awkward to formalise in our model. There are mainly two reasons for
+ this: First, we do not specify what ``running'' the code of a thread
means, for example by giving an operational semantics for machine
instructions. Therefore we cannot characterise what are ``good''
programs that contain for every looking request also a corresponding
unlocking request for a resource. Second, we would need to specify a
kind of global clock that tracks the time how long a thread locks a
resource. But this seems difficult, because how do we conveniently
- distinguish between a thread that ``just'' lock a resource for a
+ distinguish between a thread that ``just'' locks a resource for a
very long time and one that locks it forever. Therefore we decided
to leave out this property and let the programmer assume the
responsibility to program threads in such a benign manner (in
@@ -915,6 +919,25 @@
However, we are able to combine their two separate bounds into a
single theorem improving their bound.
+ In what follows we will describe properties of PIP that allow us to
+ prove Theorem~\ref{mainthm} and, when instructive, briefly describe
+ our argument. Recall we want to prove that in state @{term "s' @ s"}
+ either @{term th} is either running or blocked by a thread @{term
+ "th'"} (@{term "th \<noteq> th'"}) which was alive in state @{term s}. We
+ can show that
+
+
+
+ \begin{lemma}
+ If @{thm (prem 2) eq_pv_blocked}
+ then @{thm (concl) eq_pv_blocked}
+ \end{lemma}
+
+ \begin{lemma}
+ If @{thm (prem 2) eq_pv_persist}
+ then @{thm (concl) eq_pv_persist}
+ \end{lemma}
+
\subsection*{\bf OUTLINE}
Since @{term "th"} is the most urgent thread, if it is somehow
@@ -962,13 +985,7 @@
Lemmas we want to describe:
- \begin{lemma}
- @{thm eq_pv_persist}
- \end{lemma}
-
- \begin{lemma}
- @{thm eq_pv_blocked}
- \end{lemma}
+
\begin{lemma}
@{thm runing_cntP_cntV_inv}
--- a/PrioG.thy Thu Jan 28 07:43:05 2016 +0800
+++ b/PrioG.thy Thu Jan 28 16:33:49 2016 +0800
@@ -1,3 +1,4 @@
+<<<<<<< local
theory Correctness
imports PIPBasics
begin
@@ -796,3 +797,817 @@
end
end
+=======
+theory Correctness
+imports PIPBasics
+begin
+
+
+text {*
+ The following two auxiliary lemmas are used to reason about @{term Max}.
+*}
+lemma image_Max_eqI:
+ assumes "finite B"
+ and "b \<in> B"
+ and "\<forall> x \<in> B. f x \<le> f b"
+ shows "Max (f ` B) = f b"
+ using assms
+ using Max_eqI by blast
+
+lemma image_Max_subset:
+ assumes "finite A"
+ and "B \<subseteq> A"
+ and "a \<in> B"
+ and "Max (f ` A) = f a"
+ shows "Max (f ` B) = f a"
+proof(rule image_Max_eqI)
+ show "finite B"
+ using assms(1) assms(2) finite_subset by auto
+next
+ show "a \<in> B" using assms by simp
+next
+ show "\<forall>x\<in>B. f x \<le> f a"
+ by (metis Max_ge assms(1) assms(2) assms(4)
+ finite_imageI image_eqI subsetCE)
+qed
+
+text {*
+ The following locale @{text "highest_gen"} sets the basic context for our
+ investigation: supposing thread @{text th} holds the highest @{term cp}-value
+ in state @{text s}, which means the task for @{text th} is the
+ most urgent. We want to show that
+ @{text th} is treated correctly by PIP, which means
+ @{text th} will not be blocked unreasonably by other less urgent
+ threads.
+*}
+locale highest_gen =
+ fixes s th prio tm
+ assumes vt_s: "vt s"
+ and threads_s: "th \<in> threads s"
+ and highest: "preced th s = Max ((cp s)`threads s)"
+ -- {* The internal structure of @{term th}'s precedence is exposed:*}
+ and preced_th: "preced th s = Prc prio tm"
+
+-- {* @{term s} is a valid trace, so it will inherit all results derived for
+ a valid trace: *}
+sublocale highest_gen < vat_s: valid_trace "s"
+ by (unfold_locales, insert vt_s, simp)
+
+context highest_gen
+begin
+
+text {*
+ @{term tm} is the time when the precedence of @{term th} is set, so
+ @{term tm} must be a valid moment index into @{term s}.
+*}
+lemma lt_tm: "tm < length s"
+ by (insert preced_tm_lt[OF threads_s preced_th], simp)
+
+text {*
+ Since @{term th} holds the highest precedence and @{text "cp"}
+ is the highest precedence of all threads in the sub-tree of
+ @{text "th"} and @{text th} is among these threads,
+ its @{term cp} must equal to its precedence:
+*}
+lemma eq_cp_s_th: "cp s th = preced th s" (is "?L = ?R")
+proof -
+ have "?L \<le> ?R"
+ by (unfold highest, rule Max_ge,
+ auto simp:threads_s finite_threads)
+ moreover have "?R \<le> ?L"
+ by (unfold vat_s.cp_rec, rule Max_ge,
+ auto simp:the_preced_def vat_s.fsbttRAGs.finite_children)
+ ultimately show ?thesis by auto
+qed
+
+lemma highest_cp_preced: "cp s th = Max (the_preced s ` threads s)"
+ using eq_cp_s_th highest max_cp_eq the_preced_def by presburger
+
+
+lemma highest_preced_thread: "preced th s = Max (the_preced s ` threads s)"
+ by (fold eq_cp_s_th, unfold highest_cp_preced, simp)
+
+lemma highest': "cp s th = Max (cp s ` threads s)"
+ by (simp add: eq_cp_s_th highest)
+
+end
+
+locale extend_highest_gen = highest_gen +
+ fixes t
+ assumes vt_t: "vt (t@s)"
+ and create_low: "Create th' prio' \<in> set t \<Longrightarrow> prio' \<le> prio"
+ and set_diff_low: "Set th' prio' \<in> set t \<Longrightarrow> th' \<noteq> th \<and> prio' \<le> prio"
+ and exit_diff: "Exit th' \<in> set t \<Longrightarrow> th' \<noteq> th"
+
+sublocale extend_highest_gen < vat_t: valid_trace "t@s"
+ by (unfold_locales, insert vt_t, simp)
+
+lemma step_back_vt_app:
+ assumes vt_ts: "vt (t@s)"
+ shows "vt s"
+proof -
+ from vt_ts show ?thesis
+ proof(induct t)
+ case Nil
+ from Nil show ?case by auto
+ next
+ case (Cons e t)
+ assume ih: " vt (t @ s) \<Longrightarrow> vt s"
+ and vt_et: "vt ((e # t) @ s)"
+ show ?case
+ proof(rule ih)
+ show "vt (t @ s)"
+ proof(rule step_back_vt)
+ from vt_et show "vt (e # t @ s)" by simp
+ qed
+ qed
+ qed
+qed
+
+(* locale red_extend_highest_gen = extend_highest_gen +
+ fixes i::nat
+*)
+
+(*
+sublocale red_extend_highest_gen < red_moment: extend_highest_gen "s" "th" "prio" "tm" "(moment i t)"
+ apply (insert extend_highest_gen_axioms, subst (asm) (1) moment_restm_s [of i t, symmetric])
+ apply (unfold extend_highest_gen_def extend_highest_gen_axioms_def, clarsimp)
+ by (unfold highest_gen_def, auto dest:step_back_vt_app)
+*)
+
+context extend_highest_gen
+begin
+
+ lemma ind [consumes 0, case_names Nil Cons, induct type]:
+ assumes
+ h0: "R []"
+ and h2: "\<And> e t. \<lbrakk>vt (t@s); step (t@s) e;
+ extend_highest_gen s th prio tm t;
+ extend_highest_gen s th prio tm (e#t); R t\<rbrakk> \<Longrightarrow> R (e#t)"
+ shows "R t"
+proof -
+ from vt_t extend_highest_gen_axioms show ?thesis
+ proof(induct t)
+ from h0 show "R []" .
+ next
+ case (Cons e t')
+ assume ih: "\<lbrakk>vt (t' @ s); extend_highest_gen s th prio tm t'\<rbrakk> \<Longrightarrow> R t'"
+ and vt_e: "vt ((e # t') @ s)"
+ and et: "extend_highest_gen s th prio tm (e # t')"
+ from vt_e and step_back_step have stp: "step (t'@s) e" by auto
+ from vt_e and step_back_vt have vt_ts: "vt (t'@s)" by auto
+ show ?case
+ proof(rule h2 [OF vt_ts stp _ _ _ ])
+ show "R t'"
+ proof(rule ih)
+ from et show ext': "extend_highest_gen s th prio tm t'"
+ by (unfold extend_highest_gen_def extend_highest_gen_axioms_def, auto dest:step_back_vt)
+ next
+ from vt_ts show "vt (t' @ s)" .
+ qed
+ next
+ from et show "extend_highest_gen s th prio tm (e # t')" .
+ next
+ from et show ext': "extend_highest_gen s th prio tm t'"
+ by (unfold extend_highest_gen_def extend_highest_gen_axioms_def, auto dest:step_back_vt)
+ qed
+ qed
+qed
+
+
+lemma th_kept: "th \<in> threads (t @ s) \<and>
+ preced th (t@s) = preced th s" (is "?Q t")
+proof -
+ show ?thesis
+ proof(induct rule:ind)
+ case Nil
+ from threads_s
+ show ?case
+ by auto
+ next
+ case (Cons e t)
+ interpret h_e: extend_highest_gen _ _ _ _ "(e # t)" using Cons by auto
+ interpret h_t: extend_highest_gen _ _ _ _ t using Cons by auto
+ show ?case
+ proof(cases e)
+ case (Create thread prio)
+ show ?thesis
+ proof -
+ from Cons and Create have "step (t@s) (Create thread prio)" by auto
+ hence "th \<noteq> thread"
+ proof(cases)
+ case thread_create
+ with Cons show ?thesis by auto
+ qed
+ hence "preced th ((e # t) @ s) = preced th (t @ s)"
+ by (unfold Create, auto simp:preced_def)
+ moreover note Cons
+ ultimately show ?thesis
+ by (auto simp:Create)
+ qed
+ next
+ case (Exit thread)
+ from h_e.exit_diff and Exit
+ have neq_th: "thread \<noteq> th" by auto
+ with Cons
+ show ?thesis
+ by (unfold Exit, auto simp:preced_def)
+ next
+ case (P thread cs)
+ with Cons
+ show ?thesis
+ by (auto simp:P preced_def)
+ next
+ case (V thread cs)
+ with Cons
+ show ?thesis
+ by (auto simp:V preced_def)
+ next
+ case (Set thread prio')
+ show ?thesis
+ proof -
+ from h_e.set_diff_low and Set
+ have "th \<noteq> thread" by auto
+ hence "preced th ((e # t) @ s) = preced th (t @ s)"
+ by (unfold Set, auto simp:preced_def)
+ moreover note Cons
+ ultimately show ?thesis
+ by (auto simp:Set)
+ qed
+ qed
+ qed
+qed
+
+text {*
+ According to @{thm th_kept}, thread @{text "th"} has its living status
+ and precedence kept along the way of @{text "t"}. The following lemma
+ shows that this preserved precedence of @{text "th"} remains as the highest
+ along the way of @{text "t"}.
+
+ The proof goes by induction over @{text "t"} using the specialized
+ induction rule @{thm ind}, followed by case analysis of each possible
+ operations of PIP. All cases follow the same pattern rendered by the
+ generalized introduction rule @{thm "image_Max_eqI"}.
+
+ The very essence is to show that precedences, no matter whether they
+ are newly introduced or modified, are always lower than the one held
+ by @{term "th"}, which by @{thm th_kept} is preserved along the way.
+*}
+lemma max_kept: "Max (the_preced (t @ s) ` (threads (t@s))) = preced th s"
+proof(induct rule:ind)
+ case Nil
+ from highest_preced_thread
+ show ?case by simp
+next
+ case (Cons e t)
+ interpret h_e: extend_highest_gen _ _ _ _ "(e # t)" using Cons by auto
+ interpret h_t: extend_highest_gen _ _ _ _ t using Cons by auto
+ show ?case
+ proof(cases e)
+ case (Create thread prio')
+ show ?thesis (is "Max (?f ` ?A) = ?t")
+ proof -
+ -- {* The following is the common pattern of each branch of the case analysis. *}
+ -- {* The major part is to show that @{text "th"} holds the highest precedence: *}
+ have "Max (?f ` ?A) = ?f th"
+ proof(rule image_Max_eqI)
+ show "finite ?A" using h_e.finite_threads by auto
+ next
+ show "th \<in> ?A" using h_e.th_kept by auto
+ next
+ show "\<forall>x\<in>?A. ?f x \<le> ?f th"
+ proof
+ fix x
+ assume "x \<in> ?A"
+ hence "x = thread \<or> x \<in> threads (t@s)" by (auto simp:Create)
+ thus "?f x \<le> ?f th"
+ proof
+ assume "x = thread"
+ thus ?thesis
+ apply (simp add:Create the_preced_def preced_def, fold preced_def)
+ using Create h_e.create_low h_t.th_kept lt_tm preced_leI2
+ preced_th by force
+ next
+ assume h: "x \<in> threads (t @ s)"
+ from Cons(2)[unfolded Create]
+ have "x \<noteq> thread" using h by (cases, auto)
+ hence "?f x = the_preced (t@s) x"
+ by (simp add:Create the_preced_def preced_def)
+ hence "?f x \<le> Max (the_preced (t@s) ` threads (t@s))"
+ by (simp add: h_t.finite_threads h)
+ also have "... = ?f th"
+ by (metis Cons.hyps(5) h_e.th_kept the_preced_def)
+ finally show ?thesis .
+ qed
+ qed
+ qed
+ -- {* The minor part is to show that the precedence of @{text "th"}
+ equals to preserved one, given by the foregoing lemma @{thm th_kept} *}
+ also have "... = ?t" using h_e.th_kept the_preced_def by auto
+ -- {* Then it follows trivially that the precedence preserved
+ for @{term "th"} remains the maximum of all living threads along the way. *}
+ finally show ?thesis .
+ qed
+ next
+ case (Exit thread)
+ show ?thesis (is "Max (?f ` ?A) = ?t")
+ proof -
+ have "Max (?f ` ?A) = ?f th"
+ proof(rule image_Max_eqI)
+ show "finite ?A" using h_e.finite_threads by auto
+ next
+ show "th \<in> ?A" using h_e.th_kept by auto
+ next
+ show "\<forall>x\<in>?A. ?f x \<le> ?f th"
+ proof
+ fix x
+ assume "x \<in> ?A"
+ hence "x \<in> threads (t@s)" by (simp add: Exit)
+ hence "?f x \<le> Max (?f ` threads (t@s))"
+ by (simp add: h_t.finite_threads)
+ also have "... \<le> ?f th"
+ apply (simp add:Exit the_preced_def preced_def, fold preced_def)
+ using Cons.hyps(5) h_t.th_kept the_preced_def by auto
+ finally show "?f x \<le> ?f th" .
+ qed
+ qed
+ also have "... = ?t" using h_e.th_kept the_preced_def by auto
+ finally show ?thesis .
+ qed
+ next
+ case (P thread cs)
+ with Cons
+ show ?thesis by (auto simp:preced_def the_preced_def)
+ next
+ case (V thread cs)
+ with Cons
+ show ?thesis by (auto simp:preced_def the_preced_def)
+ next
+ case (Set thread prio')
+ show ?thesis (is "Max (?f ` ?A) = ?t")
+ proof -
+ have "Max (?f ` ?A) = ?f th"
+ proof(rule image_Max_eqI)
+ show "finite ?A" using h_e.finite_threads by auto
+ next
+ show "th \<in> ?A" using h_e.th_kept by auto
+ next
+ show "\<forall>x\<in>?A. ?f x \<le> ?f th"
+ proof
+ fix x
+ assume h: "x \<in> ?A"
+ show "?f x \<le> ?f th"
+ proof(cases "x = thread")
+ case True
+ moreover have "the_preced (Set thread prio' # t @ s) thread \<le> the_preced (t @ s) th"
+ proof -
+ have "the_preced (t @ s) th = Prc prio tm"
+ using h_t.th_kept preced_th by (simp add:the_preced_def)
+ moreover have "prio' \<le> prio" using Set h_e.set_diff_low by auto
+ ultimately show ?thesis by (insert lt_tm, auto simp:the_preced_def preced_def)
+ qed
+ ultimately show ?thesis
+ by (unfold Set, simp add:the_preced_def preced_def)
+ next
+ case False
+ then have "?f x = the_preced (t@s) x"
+ by (simp add:the_preced_def preced_def Set)
+ also have "... \<le> Max (the_preced (t@s) ` threads (t@s))"
+ using Set h h_t.finite_threads by auto
+ also have "... = ?f th" by (metis Cons.hyps(5) h_e.th_kept the_preced_def)
+ finally show ?thesis .
+ qed
+ qed
+ qed
+ also have "... = ?t" using h_e.th_kept the_preced_def by auto
+ finally show ?thesis .
+ qed
+ qed
+qed
+
+lemma max_preced: "preced th (t@s) = Max (the_preced (t@s) ` (threads (t@s)))"
+ by (insert th_kept max_kept, auto)
+
+text {*
+ The reason behind the following lemma is that:
+ Since @{term "cp"} is defined as the maximum precedence
+ of those threads contained in the sub-tree of node @{term "Th th"}
+ in @{term "RAG (t@s)"}, and all these threads are living threads, and
+ @{term "th"} is also among them, the maximum precedence of
+ them all must be the one for @{text "th"}.
+*}
+lemma th_cp_max_preced:
+ "cp (t@s) th = Max (the_preced (t@s) ` (threads (t@s)))" (is "?L = ?R")
+proof -
+ let ?f = "the_preced (t@s)"
+ have "?L = ?f th"
+ proof(unfold cp_alt_def, rule image_Max_eqI)
+ show "finite {th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}"
+ proof -
+ have "{th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)} =
+ the_thread ` {n . n \<in> subtree (RAG (t @ s)) (Th th) \<and>
+ (\<exists> th'. n = Th th')}"
+ by (smt Collect_cong Setcompr_eq_image mem_Collect_eq the_thread.simps)
+ moreover have "finite ..." by (simp add: vat_t.fsbtRAGs.finite_subtree)
+ ultimately show ?thesis by simp
+ qed
+ next
+ show "th \<in> {th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}"
+ by (auto simp:subtree_def)
+ next
+ show "\<forall>x\<in>{th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}.
+ the_preced (t @ s) x \<le> the_preced (t @ s) th"
+ proof
+ fix th'
+ assume "th' \<in> {th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}"
+ hence "Th th' \<in> subtree (RAG (t @ s)) (Th th)" by auto
+ moreover have "... \<subseteq> Field (RAG (t @ s)) \<union> {Th th}"
+ by (meson subtree_Field)
+ ultimately have "Th th' \<in> ..." by auto
+ hence "th' \<in> threads (t@s)"
+ proof
+ assume "Th th' \<in> {Th th}"
+ thus ?thesis using th_kept by auto
+ next
+ assume "Th th' \<in> Field (RAG (t @ s))"
+ thus ?thesis using vat_t.not_in_thread_isolated by blast
+ qed
+ thus "the_preced (t @ s) th' \<le> the_preced (t @ s) th"
+ by (metis Max_ge finite_imageI finite_threads image_eqI
+ max_kept th_kept the_preced_def)
+ qed
+ qed
+ also have "... = ?R" by (simp add: max_preced the_preced_def)
+ finally show ?thesis .
+qed
+
+lemma th_cp_max[simp]: "Max (cp (t@s) ` threads (t@s)) = cp (t@s) th"
+ using max_cp_eq th_cp_max_preced the_preced_def vt_t by presburger
+
+lemma [simp]: "Max (cp (t@s) ` threads (t@s)) = Max (the_preced (t@s) ` threads (t@s))"
+ by (simp add: th_cp_max_preced)
+
+lemma [simp]: "Max (the_preced (t@s) ` threads (t@s)) = the_preced (t@s) th"
+ using max_kept th_kept the_preced_def by auto
+
+lemma [simp]: "the_preced (t@s) th = preced th (t@s)"
+ using the_preced_def by auto
+
+lemma [simp]: "preced th (t@s) = preced th s"
+ by (simp add: th_kept)
+
+lemma [simp]: "cp s th = preced th s"
+ by (simp add: eq_cp_s_th)
+
+lemma th_cp_preced [simp]: "cp (t@s) th = preced th s"
+ by (fold max_kept, unfold th_cp_max_preced, simp)
+
+lemma preced_less:
+ assumes th'_in: "th' \<in> threads s"
+ and neq_th': "th' \<noteq> th"
+ shows "preced th' s < preced th s"
+ using assms
+by (metis Max.coboundedI finite_imageI highest not_le order.trans
+ preced_linorder rev_image_eqI threads_s vat_s.finite_threads
+ vat_s.le_cp)
+
+section {* The `blocking thread` *}
+
+text {*
+
+ The purpose of PIP is to ensure that the most urgent thread @{term
+ th} is not blocked unreasonably. Therefore, below, we will derive
+ properties of the blocking thread. By blocking thread, we mean a
+ thread in running state t @ s, but is different from thread @{term
+ th}.
+
+ The first lemmas shows that the @{term cp}-value of the blocking
+ thread @{text th'} equals to the highest precedence in the whole
+ system.
+
+*}
+
+lemma runing_preced_inversion:
+ assumes runing': "th' \<in> runing (t @ s)"
+ shows "cp (t @ s) th' = preced th s"
+proof -
+ have "cp (t @ s) th' = Max (cp (t @ s) ` readys (t @ s))"
+ using assms by (unfold runing_def, auto)
+ also have "\<dots> = preced th s"
+ by (metis th_cp_max th_cp_preced vat_t.max_cp_readys_threads)
+ finally show ?thesis .
+qed
+
+text {*
+
+ The next lemma shows how the counters for @{term "P"} and @{term
+ "V"} operations relate to the running threads in the states @{term
+ s} and @{term "t @ s"}: if a thread's @{term "P"}-count equals its
+ @{term "V"}-count (which means it no longer has any resource in its
+ possession), it cannot be a running thread.
+
+ The proof is by contraction with the assumption @{text "th' \<noteq> th"}.
+ The key is the use of @{thm count_eq_dependants} to derive the
+ emptiness of @{text th'}s @{term dependants}-set from the balance of
+ its @{term P} and @{term V} counts. From this, it can be shown
+ @{text th'}s @{term cp}-value equals to its own precedence.
+
+ On the other hand, since @{text th'} is running, by @{thm
+ runing_preced_inversion}, its @{term cp}-value equals to the
+ precedence of @{term th}.
+
+ Combining the above two results we have that @{text th'} and @{term
+ th} have the same precedence. By uniqueness of precedences, we have
+ @{text "th' = th"}, which is in contradiction with the assumption
+ @{text "th' \<noteq> th"}.
+
+*}
+
+lemma eq_pv_blocked: (* ddd *)
+ assumes neq_th': "th' \<noteq> th"
+ and eq_pv: "cntP (t @ s) th' = cntV (t @ s) th'"
+ shows "th' \<notin> runing (t @ s)"
+proof
+ assume otherwise: "th' \<in> runing (t @ s)"
+ show False
+ proof -
+ have th'_in: "th' \<in> threads (t @ s)"
+ using otherwise readys_threads runing_def by auto
+ have "th' = th"
+ proof(rule preced_unique)
+ -- {* The proof goes like this:
+ it is first shown that the @{term preced}-value of @{term th'}
+ equals to that of @{term th}, then by uniqueness
+ of @{term preced}-values (given by lemma @{thm preced_unique}),
+ @{term th'} equals to @{term th}: *}
+ show "preced th' (t @ s) = preced th (t @ s)" (is "?L = ?R")
+ proof -
+ -- {* Since the counts of @{term th'} are balanced, the subtree
+ of it contains only itself, so, its @{term cp}-value
+ equals its @{term preced}-value: *}
+ have "?L = cp (t @ s) th'"
+ by (unfold cp_eq_cpreced cpreced_def count_eq_dependants[OF eq_pv], simp)
+ -- {* Since @{term "th'"} is running, by @{thm runing_preced_inversion},
+ its @{term cp}-value equals @{term "preced th s"},
+ which equals to @{term "?R"} by simplification: *}
+ also have "... = ?R"
+ using runing_preced_inversion[OF otherwise] by simp
+ finally show ?thesis .
+ qed
+ qed (auto simp: th'_in th_kept)
+ with `th' \<noteq> th` show ?thesis by simp
+ qed
+qed
+
+text {*
+ The following lemma is the extrapolation of @{thm eq_pv_blocked}.
+ It says if a thread, different from @{term th},
+ does not hold any resource at the very beginning,
+ it will keep hand-emptied in the future @{term "t@s"}.
+*}
+lemma eq_pv_persist: (* ddd *)
+ assumes neq_th': "th' \<noteq> th"
+ and eq_pv: "cntP s th' = cntV s th'"
+ shows "cntP (t @ s) th' = cntV (t @ s) th'"
+proof(induction rule: ind)
+ -- {* The nontrivial case is for the @{term Cons}: *}
+ case (Cons e t)
+ -- {* All results derived so far hold for both @{term s} and @{term "t@s"}: *}
+ interpret vat_t: extend_highest_gen s th prio tm t using Cons by simp
+ interpret vat_e: extend_highest_gen s th prio tm "(e # t)" using Cons by simp
+ show ?case
+ proof -
+ -- {* It can be proved that @{term cntP}-value of @{term th'} does not change
+ by the happening of event @{term e}: *}
+ have "cntP ((e#t)@s) th' = cntP (t@s) th'"
+ proof(rule ccontr) -- {* Proof by contradiction. *}
+ -- {* Suppose @{term cntP}-value of @{term th'} is changed by @{term e}: *}
+ assume otherwise: "cntP ((e # t) @ s) th' \<noteq> cntP (t @ s) th'"
+ -- {* Then the actor of @{term e} must be @{term th'} and @{term e}
+ must be a @{term P}-event: *}
+ hence "isP e" "actor e = th'" by (auto simp:cntP_diff_inv)
+ with vat_t.actor_inv[OF Cons(2)]
+ -- {* According to @{thm actor_inv}, @{term th'} must be running at
+ the moment @{term "t@s"}: *}
+ have "th' \<in> runing (t@s)" by (cases e, auto)
+ -- {* However, an application of @{thm eq_pv_blocked} to induction hypothesis
+ shows @{term th'} can not be running at moment @{term "t@s"}: *}
+ moreover have "th' \<notin> runing (t@s)"
+ using vat_t.eq_pv_blocked[OF neq_th' Cons(5)] .
+ -- {* Contradiction is finally derived: *}
+ ultimately show False by simp
+ qed
+ -- {* It can also be proved that @{term cntV}-value of @{term th'} does not change
+ by the happening of event @{term e}: *}
+ -- {* The proof follows exactly the same pattern as the case for @{term cntP}-value: *}
+ moreover have "cntV ((e#t)@s) th' = cntV (t@s) th'"
+ proof(rule ccontr) -- {* Proof by contradiction. *}
+ assume otherwise: "cntV ((e # t) @ s) th' \<noteq> cntV (t @ s) th'"
+ hence "isV e" "actor e = th'" by (auto simp:cntV_diff_inv)
+ with vat_t.actor_inv[OF Cons(2)]
+ have "th' \<in> runing (t@s)" by (cases e, auto)
+ moreover have "th' \<notin> runing (t@s)"
+ using vat_t.eq_pv_blocked[OF neq_th' Cons(5)] .
+ ultimately show False by simp
+ qed
+ -- {* Finally, it can be shown that the @{term cntP} and @{term cntV}
+ value for @{term th'} are still in balance, so @{term th'}
+ is still hand-emptied after the execution of event @{term e}: *}
+ ultimately show ?thesis using Cons(5) by metis
+ qed
+qed (auto simp:eq_pv)
+
+text {*
+
+ By combining @{thm eq_pv_blocked} and @{thm eq_pv_persist}, it can
+ be derived easily that @{term th'} can not be running in the future:
+
+*}
+
+lemma eq_pv_blocked_persist:
+ assumes neq_th': "th' \<noteq> th"
+ and eq_pv: "cntP s th' = cntV s th'"
+ shows "th' \<notin> runing (t @ s)"
+ using assms
+ by (simp add: eq_pv_blocked eq_pv_persist)
+
+text {*
+
+ The following lemma shows the blocking thread @{term th'} must hold
+ some resource in the very beginning.
+
+*}
+
+lemma runing_cntP_cntV_inv: (* ddd *)
+ assumes is_runing: "th' \<in> runing (t @ s)"
+ and neq_th': "th' \<noteq> th"
+ shows "cntP s th' > cntV s th'"
+ using assms
+proof -
+ -- {* First, it can be shown that the number of @{term P} and
+ @{term V} operations can not be equal for thred @{term th'} *}
+ have "cntP s th' \<noteq> cntV s th'"
+ proof
+ -- {* The proof goes by contradiction, suppose otherwise: *}
+ assume otherwise: "cntP s th' = cntV s th'"
+ -- {* By applying @{thm eq_pv_blocked_persist} to this: *}
+ from eq_pv_blocked_persist[OF neq_th' otherwise]
+ -- {* we have that @{term th'} can not be running at moment @{term "t@s"}: *}
+ have "th' \<notin> runing (t@s)" .
+ -- {* This is obvious in contradiction with assumption @{thm is_runing} *}
+ thus False using is_runing by simp
+ qed
+ -- {* However, the number of @{term V} is always less or equal to @{term P}: *}
+ moreover have "cntV s th' \<le> cntP s th'" using vat_s.cnp_cnv_cncs by auto
+ -- {* Thesis is finally derived by combining the these two results: *}
+ ultimately show ?thesis by auto
+qed
+
+
+text {*
+
+ The following lemmas shows the blocking thread @{text th'} must be
+ live at the very beginning, i.e. the moment (or state) @{term s}.
+ The proof is a simple combination of the results above:
+
+*}
+
+lemma runing_threads_inv:
+ assumes runing': "th' \<in> runing (t@s)"
+ and neq_th': "th' \<noteq> th"
+ shows "th' \<in> threads s"
+proof(rule ccontr) -- {* Proof by contradiction: *}
+ assume otherwise: "th' \<notin> threads s"
+ have "th' \<notin> runing (t @ s)"
+ proof -
+ from vat_s.cnp_cnv_eq[OF otherwise]
+ have "cntP s th' = cntV s th'" .
+ from eq_pv_blocked_persist[OF neq_th' this]
+ show ?thesis .
+ qed
+ with runing' show False by simp
+qed
+
+text {*
+
+ The following lemma summarises the above lemmas to give an overall
+ characterisationof the blocking thread @{text "th'"}:
+
+*}
+
+lemma runing_inversion: (* ddd, one of the main lemmas to present *)
+ assumes runing': "th' \<in> runing (t@s)"
+ and neq_th: "th' \<noteq> th"
+ shows "th' \<in> threads s"
+ and "\<not>detached s th'"
+ and "cp (t@s) th' = preced th s"
+proof -
+ from runing_threads_inv[OF assms]
+ show "th' \<in> threads s" .
+next
+ from runing_cntP_cntV_inv[OF runing' neq_th]
+ show "\<not>detached s th'" using vat_s.detached_eq by simp
+next
+ from runing_preced_inversion[OF runing']
+ show "cp (t@s) th' = preced th s" .
+qed
+
+
+section {* The existence of `blocking thread` *}
+
+text {*
+
+ Suppose @{term th} is not running, it is first shown that there is a
+ path in RAG leading from node @{term th} to another thread @{text
+ "th'"} in the @{term readys}-set (So @{text "th'"} is an ancestor of
+ @{term th}}).
+
+ Now, since @{term readys}-set is non-empty, there must be one in it
+ which holds the highest @{term cp}-value, which, by definition, is
+ the @{term runing}-thread. However, we are going to show more: this
+ running thread is exactly @{term "th'"}.
+
+*}
+
+lemma th_blockedE: (* ddd, the other main lemma to be presented: *)
+ assumes "th \<notin> runing (t@s)"
+ obtains th' where "Th th' \<in> ancestors (RAG (t @ s)) (Th th)"
+ "th' \<in> runing (t@s)"
+proof -
+ -- {* According to @{thm vat_t.th_chain_to_ready}, either
+ @{term "th"} is in @{term "readys"} or there is path leading from it to
+ one thread in @{term "readys"}. *}
+ have "th \<in> readys (t @ s) \<or> (\<exists>th'. th' \<in> readys (t @ s) \<and> (Th th, Th th') \<in> (RAG (t @ s))\<^sup>+)"
+ using th_kept vat_t.th_chain_to_ready by auto
+ -- {* However, @{term th} can not be in @{term readys}, because otherwise, since
+ @{term th} holds the highest @{term cp}-value, it must be @{term "runing"}. *}
+ moreover have "th \<notin> readys (t@s)"
+ using assms runing_def th_cp_max vat_t.max_cp_readys_threads by auto
+ -- {* So, there must be a path from @{term th} to another thread @{text "th'"} in
+ term @{term readys}: *}
+ ultimately obtain th' where th'_in: "th' \<in> readys (t@s)"
+ and dp: "(Th th, Th th') \<in> (RAG (t @ s))\<^sup>+" by auto
+ -- {* We are going to show that this @{term th'} is running. *}
+ have "th' \<in> runing (t@s)"
+ proof -
+ -- {* We only need to show that this @{term th'} holds the highest @{term cp}-value: *}
+ have "cp (t@s) th' = Max (cp (t@s) ` readys (t@s))" (is "?L = ?R")
+ proof -
+ have "?L = Max ((the_preced (t @ s) \<circ> the_thread) ` subtree (tRAG (t @ s)) (Th th'))"
+ by (unfold cp_alt_def1, simp)
+ also have "... = (the_preced (t @ s) \<circ> the_thread) (Th th)"
+ proof(rule image_Max_subset)
+ show "finite (Th ` (threads (t@s)))" by (simp add: vat_t.finite_threads)
+ next
+ show "subtree (tRAG (t @ s)) (Th th') \<subseteq> Th ` threads (t @ s)"
+ by (metis Range.intros dp trancl_range vat_t.range_in vat_t.subtree_tRAG_thread)
+ next
+ show "Th th \<in> subtree (tRAG (t @ s)) (Th th')" using dp
+ by (unfold tRAG_subtree_eq, auto simp:subtree_def)
+ next
+ show "Max ((the_preced (t @ s) \<circ> the_thread) ` Th ` threads (t @ s)) =
+ (the_preced (t @ s) \<circ> the_thread) (Th th)" (is "Max ?L = _")
+ proof -
+ have "?L = the_preced (t @ s) ` threads (t @ s)"
+ by (unfold image_comp, rule image_cong, auto)
+ thus ?thesis using max_preced the_preced_def by auto
+ qed
+ qed
+ also have "... = ?R"
+ using th_cp_max th_cp_preced th_kept
+ the_preced_def vat_t.max_cp_readys_threads by auto
+ finally show ?thesis .
+ qed
+ -- {* Now, since @{term th'} holds the highest @{term cp}
+ and we have already show it is in @{term readys},
+ it is @{term runing} by definition. *}
+ with `th' \<in> readys (t@s)` show ?thesis by (simp add: runing_def)
+ qed
+ -- {* It is easy to show @{term th'} is an ancestor of @{term th}: *}
+ moreover have "Th th' \<in> ancestors (RAG (t @ s)) (Th th)"
+ using `(Th th, Th th') \<in> (RAG (t @ s))\<^sup>+` by (auto simp:ancestors_def)
+ ultimately show ?thesis using that by metis
+qed
+
+text {*
+
+ Now it is easy to see there is always a thread to run by case
+ analysis on whether thread @{term th} is running: if the answer is
+ yes, the the running thread is obviously @{term th} itself;
+ otherwise, the running thread is the @{text th'} given by lemma
+ @{thm th_blockedE}.
+
+*}
+
+lemma live: "runing (t@s) \<noteq> {}"
+proof(cases "th \<in> runing (t@s)")
+ case True thus ?thesis by auto
+next
+ case False
+ thus ?thesis using th_blockedE by auto
+qed
+
+
+end
+end
+>>>>>>> other
Binary file journal.pdf has changed
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/log Thu Jan 28 16:33:49 2016 +0800
@@ -0,0 +1,447 @@
+修改集: 86:2106021bae53
+标签: tip
+用户: zhangx
+日期: Thu Jan 28 07:46:05 2016 +0800
+摘要: Added PrioG.thy again
+
+修改集: 85:61a4429e7d4d
+父亲: 84:c0a4e840aefe
+父亲: 33:7f87232d9424
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Wed Jan 27 13:50:02 2016 +0000
+摘要: merged
+
+修改集: 84:c0a4e840aefe
+父亲: 77:b6ea51cd2e88
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Wed Jan 27 13:47:08 2016 +0000
+摘要: some small changes to Correctness and Paper
+
+修改集: 83:d239aa953315
+用户: zhangx
+日期: Thu Jan 28 07:43:05 2016 +0800
+摘要: Added PrioG.thy as a parallel copy of Correctness.thy
+
+修改集: 82:cfd644dfc3b4
+用户: zhangx
+日期: Wed Jan 27 23:34:23 2016 +0800
+摘要: The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
+
+修改集: 81:c495eb16beb6
+用户: zhangx
+日期: Wed Jan 27 19:28:42 2016 +0800
+摘要: CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
+
+修改集: 80:17305a85493d
+用户: zhangx
+日期: Wed Jan 27 19:26:56 2016 +0800
+摘要: CpsG.thy retrofiting almost completed. An important mile stone.
+
+修改集: 79:8067efcb43da
+用户: zhangx
+日期: Sun Jan 17 22:18:35 2016 +0800
+摘要: Still improving CpsG.thy
+
+修改集: 78:df0334468335
+父亲: 75:d37703e0c5c4
+父亲: 77:b6ea51cd2e88
+用户: zhangx
+日期: Sat Jan 16 11:02:17 2016 +0800
+摘要: Merged with 77
+
+修改集: 77:b6ea51cd2e88
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Fri Jan 15 02:05:29 2016 +0000
+摘要: some small changes to the paper
+
+修改集: 76:2aa37de77f31
+父亲: 74:83ba2d8c859a
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Thu Jan 14 03:29:22 2016 +0000
+摘要: updated paper
+
+修改集: 75:d37703e0c5c4
+用户: zhangx
+日期: Sat Jan 16 10:59:03 2016 +0800
+摘要: CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
+
+修改集: 74:83ba2d8c859a
+用户: zhangx
+日期: Thu Jan 14 00:55:54 2016 +0800
+摘要: Moment.thy further simplified.
+
+修改集: 73:b0054fb0d1ce
+用户: zhangx
+日期: Wed Jan 13 23:39:59 2016 +0800
+摘要: Moment.thy further improved.
+
+修改集: 72:3fa70b12c117
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Wed Jan 13 15:22:14 2016 +0000
+摘要: another simplification
+
+修改集: 71:04caf0ccb3ae
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Wed Jan 13 15:16:59 2016 +0000
+摘要: some small change
+
+修改集: 70:92ca2410b3d9
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Wed Jan 13 14:20:58 2016 +0000
+摘要: further simplificaton of Moment.thy
+
+修改集: 69:1dc801552dfd
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Wed Jan 13 13:20:45 2016 +0000
+摘要: simplified Moment.thy
+
+修改集: 68:db196b066b97
+用户: zhangx
+日期: Tue Jan 12 08:35:36 2016 +0800
+摘要: Before retrofiting PIPBasics.thy
+
+修改集: 67:25fd656667a7
+用户: zhangx
+日期: Sat Jan 09 22:19:27 2016 +0800
+摘要: Correctness simplified a great deal.
+
+修改集: 66:2af87bb52fca
+用户: zhangx
+日期: Thu Jan 07 22:10:06 2016 +0800
+摘要: Some small improvements in Correctness.thy.
+
+修改集: 65:633b1fc8631b
+用户: zhangx
+日期: Thu Jan 07 08:33:13 2016 +0800
+摘要: Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
+
+修改集: 64:b4bcd1edbb6d
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Wed Jan 06 16:34:26 2016 +0000
+摘要: renamed files
+
+修改集: 63:b620a2a0806a
+用户: zhangx
+日期: Wed Jan 06 20:46:14 2016 +0800
+摘要: ExtGG.thy finished, but more comments are needed.
+
+修改集: 62:031d2ae9c9b8
+用户: zhangx
+日期: Tue Dec 22 23:13:31 2015 +0800
+摘要: In the middle of retrofiting ExtGG.thy.
+
+修改集: 61:f8194fd6214f
+用户: zhangx
+日期: Fri Dec 18 22:47:32 2015 +0800
+摘要: CpsG.thy has been cleaned up.
+
+修改集: 60:f98a95f3deae
+用户: zhangx
+日期: Fri Dec 18 19:13:19 2015 +0800
+摘要: Main proofs in CpsG.thy completed.
+
+修改集: 59:0a069a667301
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Tue Dec 15 15:10:40 2015 +0000
+摘要: removed some fixes about which Isabelle complains
+
+修改集: 58:ad57323fd4d6
+用户: zhangx
+日期: Tue Dec 15 21:45:46 2015 +0800
+摘要: Extended RTree.thy
+
+修改集: 57:f1b39d77db00
+用户: xingyuan zhang <xingyuanzhang@126.com>
+日期: Thu Dec 03 14:34:29 2015 +0800
+摘要: Added generic theory "RTree.thy"
+
+修改集: 56:0fd478e14e87
+用户: xingyuan zhang <xingyuanzhang@126.com>
+日期: Thu Dec 03 14:34:00 2015 +0800
+摘要: Before switching to generic theory of relational trees.
+
+修改集: 55:b85cfbd58f59
+用户: xingyuan zhang <xingyuanzhang@126.com>
+日期: Fri Oct 30 20:40:11 2015 +0800
+摘要: Comments for Set-operation finished
+
+修改集: 54:fee01b2858a2
+父亲: 50:8142e80f5d58
+父亲: 53:78adeef368c1
+用户: xingyuan zhang <xingyuanzhang@126.com>
+日期: Sat Oct 17 16:14:30 2015 +0800
+摘要: Merge with tip
+
+修改集: 53:78adeef368c1
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Tue Oct 06 14:22:34 2015 +0100
+摘要: test
+
+修改集: 52:d462d449505f
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Tue Oct 06 14:13:52 2015 +0100
+摘要: another test
+
+修改集: 51:38ad30559775
+父亲: 49:8679d75b1d76
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Tue Oct 06 14:11:28 2015 +0100
+摘要: test
+
+修改集: 50:8142e80f5d58
+用户: xingyuan zhang <xingyuanzhang@126.com>
+日期: Sat Oct 17 16:10:33 2015 +0800
+摘要: Finished comments on PrioGDef.thy
+
+修改集: 49:8679d75b1d76
+用户: xingyuan zhang <xingyuanzhang@126.com>
+日期: Tue Oct 06 18:52:04 2015 +0800
+摘要: A little more change.
+
+修改集: 48:c0f14399c12f
+用户: xingyuan zhang <xingyuanzhang@126.com>
+日期: Tue Oct 06 13:08:00 2015 +0800
+摘要: Some changes in the PrioGDef.thy.
+
+修改集: 47:2e6c8d530216
+用户: xingyuan zhang <xingyuanzhang@126.com>
+日期: Tue Oct 06 11:26:18 2015 +0800
+摘要: Just a test change
+
+修改集: 46:331137d43625
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Sun Oct 04 23:02:57 2015 +0100
+摘要: added one more reference to an incorrect specification
+
+修改集: 45:fc83f79009bd
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Wed Sep 09 11:24:19 2015 +0100
+摘要: updated for Isabelle 2015
+
+修改集: 44:f676a68935a0
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Tue Jul 15 17:25:53 2014 +0200
+摘要: updated teh theories to newer Isabelle version
+
+修改集: 43:45e1d324c493
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Thu Jun 12 10:14:50 2014 +0100
+摘要: a few additions
+
+修改集: 42:0069bca6dd51
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Tue Jun 10 10:44:48 2014 +0100
+摘要: added scheduling book
+
+修改集: 41:66ed924aaa5c
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Mon Jun 09 16:01:28 2014 +0100
+摘要: added another book that makes the error, some more proofs
+
+修改集: 40:0781a2fc93f1
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Tue Jun 03 15:00:12 2014 +0100
+摘要: added a library about graphs
+
+修改集: 39:7ea6b019ce24
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Mon Jun 02 14:58:42 2014 +0100
+摘要: updated
+
+修改集: 38:c89013dca1aa
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Fri May 30 07:56:39 2014 +0100
+摘要: finished proof of acyclity
+
+修改集: 37:c820ac0f3088
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Sat May 24 12:39:12 2014 +0100
+摘要: simplified RAG_acyclic proof
+
+修改集: 36:af38526275f8
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Fri May 23 15:19:32 2014 +0100
+摘要: added a test theory for polishing teh proofs
+
+修改集: 35:92f61f6a0fe7
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Thu May 22 17:40:39 2014 +0100
+摘要: added a bit more text to the paper and separated a theory about Max
+
+修改集: 34:313acffe63b6
+父亲: 32:9b9f2117561f
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Tue May 20 12:49:21 2014 +0100
+摘要: updated ROOT file
+
+修改集: 33:7f87232d9424
+父亲: 29:408ff78ce28f
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Wed May 14 11:52:53 2014 +0100
+摘要: test
+
+修改集: 32:9b9f2117561f
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Thu May 15 16:02:44 2014 +0100
+摘要: simplified the cp_rec proof
+
+修改集: 31:e861aff29655
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Tue May 06 14:36:40 2014 +0100
+摘要: made some modifications.
+
+修改集: 30:8f026b608378
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Wed Mar 12 10:08:20 2014 +0000
+摘要: added paper
+
+修改集: 29:408ff78ce28f
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Tue Mar 04 16:47:54 2014 +0000
+摘要: updated readme
+
+修改集: 28:7fa738a9615a
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Tue Mar 04 16:38:38 2014 +0000
+摘要: updated
+
+修改集: 27:6b1141c5e24c
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Tue Mar 04 15:49:36 2014 +0000
+摘要: cleaned up
+
+修改集: 26:da7a6ccfa7a9
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Tue Mar 04 15:30:24 2014 +0000
+摘要: updated
+
+修改集: 25:a9c0eeb00cc3
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Tue Mar 04 15:27:59 2014 +0000
+摘要: added two more references
+
+修改集: 24:6f50e6a8c6e0
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Tue Mar 04 09:40:40 2014 +0000
+摘要: some additions
+
+修改集: 23:24e6884d9258
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Tue Mar 04 08:45:11 2014 +0000
+摘要: made some small chages
+
+修改集: 22:9f0b78fcc894
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Mon Mar 03 16:22:48 2014 +0000
+摘要: updated
+
+修改集: 21:55d1591b17f0
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Fri Feb 28 12:49:58 2014 +0000
+摘要: added llncs to journal
+
+修改集: 20:b56616fd88dd
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Tue Feb 25 20:01:47 2014 +0000
+摘要: added
+
+修改集: 19:3cc70bd49588
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Thu Jun 20 23:28:26 2013 -0400
+摘要: added paper
+
+修改集: 18:598409a21f4c
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Thu Jun 20 13:50:01 2013 -0400
+摘要: added nasa talk
+
+修改集: 17:105715a0a807
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Sat Dec 22 14:50:29 2012 +0000
+摘要: updated
+
+修改集: 16:9764023f719e
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Sat Dec 22 01:58:45 2012 +0000
+摘要: added
+
+修改集: 15:9e664c268e25
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Fri Dec 21 23:32:58 2012 +0000
+摘要: added
+
+修改集: 14:1bf194825a4e
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Fri Dec 21 18:06:00 2012 +0000
+摘要: more one the implementation
+
+修改集: 13:735e36c64a71
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Fri Dec 21 13:30:14 2012 +0000
+摘要: added explanation of the code
+
+修改集: 12:85116bc854c0
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Fri Dec 21 00:24:30 2012 +0000
+摘要: updated
+
+修改集: 11:8e02fb168350
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Thu Dec 20 14:54:06 2012 +0000
+摘要: added
+
+修改集: 10:242a781135ba
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Thu Dec 20 12:23:44 2012 +0000
+摘要: added two papers about PIP on multiprocs
+
+修改集: 9:a8e8ec87a933
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Thu Dec 20 10:53:31 2012 +0000
+摘要: added original Sha paper
+
+修改集: 8:5ba3d79622da
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Wed Dec 19 23:46:36 2012 +0000
+摘要: added a paragraph about RAGS
+
+修改集: 7:0514be2ad83e
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Wed Dec 19 12:51:06 2012 +0000
+摘要: started code explanation
+
+修改集: 6:7f2493296c39
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Mon Dec 17 12:34:24 2012 +0000
+摘要: updated
+
+修改集: 5:0f2d4b78f839
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Mon Dec 10 21:27:22 2012 +0000
+摘要: updated
+
+修改集: 4:9d667d545e32
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Thu Dec 06 16:30:57 2012 +0000
+摘要: added
+
+修改集: 3:51019d035a79
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Thu Dec 06 15:49:20 2012 +0000
+摘要: made everything working
+
+修改集: 2:a04084de4946
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Thu Dec 06 15:12:49 2012 +0000
+摘要: added
+
+修改集: 1:c4783e4ef43f
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Thu Dec 06 15:11:51 2012 +0000
+摘要: added
+
+修改集: 0:110247f9d47e
+用户: Christian Urban <christian dot urban at kcl dot ac dot uk>
+日期: Thu Dec 06 15:11:21 2012 +0000
+摘要: added
+