diff -r 12e9aa68d5db -r 4190df6f4488 prio/ExtS.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/prio/ExtS.thy Tue Jan 24 00:20:09 2012 +0000 @@ -0,0 +1,1019 @@ +theory ExtS +imports Prio +begin + +locale highest_set = + fixes s' th prio fixes s + defines s_def : "s \ (Set th prio#s')" + assumes vt_s: "vt step s" + and highest: "preced th s = Max ((cp s)`threads s)" + +context highest_set +begin + + +lemma vt_s': "vt step s'" + by (insert vt_s, unfold s_def, drule_tac step_back_vt, simp) + +lemma step_set: "step s' (Set th prio)" + by (insert vt_s, unfold s_def, drule_tac step_back_step, simp) + +lemma step_set_elim: + "\\th \ runing s'\ \ Q\ \ Q" + by (insert step_set, ind_cases "step s' (Set th prio)", auto) + + +lemma threads_s: "th \ threads s" + by (rule step_set_elim, unfold runing_def readys_def, auto simp:s_def) + +lemma same_depend: "depend s = depend s'" + by (insert depend_set_unchanged, unfold s_def, simp) + +lemma same_dependents: + "dependents (wq s) th = dependents (wq s') th" + apply (unfold cs_dependents_def) + by (unfold eq_depend same_depend, simp) + +lemma eq_cp_s_th: "cp s th = preced th s" +proof - + from highest and max_cp_eq[OF vt_s] + have is_max: "preced th s = Max ((\th. preced th s) ` threads s)" by simp + have sbs: "({th} \ dependents (wq s) th) \ threads s" + proof - + from threads_s and dependents_threads[OF vt_s, of th] + show ?thesis by auto + qed + show ?thesis + proof(unfold cp_eq_cpreced cpreced_def, rule Max_eqI) + show "preced th s \ (\th. preced th s) ` ({th} \ dependents (wq s) th)" by simp + next + fix y + assume "y \ (\th. preced th s) ` ({th} \ dependents (wq s) th)" + then obtain th1 where th1_in: "th1 \ ({th} \ dependents (wq s) th)" + and eq_y: "y = preced th1 s" by auto + show "y \ preced th s" + proof(unfold is_max, rule Max_ge) + from finite_threads[OF vt_s] + show "finite ((\th. preced th s) ` threads s)" by simp + next + from sbs th1_in and eq_y + show "y \ (\th. preced th s) ` threads s" by auto + qed + next + from sbs and finite_threads[OF vt_s] + show "finite ((\th. preced th s) ` ({th} \ dependents (wq s) th))" + by (auto intro:finite_subset) + qed +qed + +lemma highest_cp_preced: "cp s th = Max ((\ th'. preced th' s) ` threads s)" + by (fold max_cp_eq[OF vt_s], unfold eq_cp_s_th, insert highest, simp) + +lemma highest_preced_thread: "preced th s = Max ((\ th'. preced th' s) ` threads s)" + by (fold eq_cp_s_th, unfold highest_cp_preced, simp) + +lemma is_ready: "th \ readys s" +proof - + have "\cs. \ waiting s th cs" + apply (rule step_set_elim, unfold s_def, insert depend_set_unchanged[of th prio s']) + apply (unfold s_depend_def, unfold runing_def readys_def) + apply (auto, fold s_def) + apply (erule_tac x = cs in allE, auto simp:waiting_eq) + proof - + fix cs + assume h: + "{(Th t, Cs c) |t c. waiting (wq s) t c} \ {(Cs c, Th t) |c t. holding (wq s) t c} = + {(Th t, Cs c) |t c. waiting (wq s') t c} \ {(Cs c, Th t) |c t. holding (wq s') t c}" + (is "?L = ?R") + and wt: "waiting (wq s) th cs" and nwt: "\ waiting (wq s') th cs" + from wt have "(Th th, Cs cs) \ ?L" by auto + with h have "(Th th, Cs cs) \ ?R" by simp + hence "waiting (wq s') th cs" by auto with nwt + show False by auto + qed + with threads_s show ?thesis + by (unfold readys_def, auto) +qed + +lemma highest': "cp s th = Max (cp s ` threads s)" +proof - + from highest_cp_preced max_cp_eq[OF vt_s, symmetric] + show ?thesis by simp +qed + +lemma is_runing: "th \ runing s" +proof - + have "Max (cp s ` threads s) = Max (cp s ` readys s)" + proof - + have " Max (cp s ` readys s) = cp s th" + proof(rule Max_eqI) + from finite_threads[OF vt_s] readys_threads finite_subset + have "finite (readys s)" by blast + thus "finite (cp s ` readys s)" by auto + next + from is_ready show "cp s th \ cp s ` readys s" by auto + next + fix y + assume "y \ cp s ` readys s" + then obtain th1 where + eq_y: "y = cp s th1" and th1_in: "th1 \ readys s" by auto + show "y \ cp s th" + proof - + have "y \ Max (cp s ` threads s)" + proof(rule Max_ge) + from eq_y and th1_in + show "y \ cp s ` threads s" + by (auto simp:readys_def) + next + from finite_threads[OF vt_s] + show "finite (cp s ` threads s)" by auto + qed + with highest' show ?thesis by auto + qed + qed + with highest' show ?thesis by auto + qed + thus ?thesis + by (unfold runing_def, insert highest' is_ready, auto) +qed + +end + +locale extend_highest_set = highest_set + + fixes t + assumes vt_t: "vt step (t@s)" + and create_low: "Create th' prio' \ set t \ prio' \ prio" + and set_diff_low: "Set th' prio' \ set t \ th' \ th \ prio' \ prio" + and exit_diff: "Exit th' \ set t \ th' \ th" + +lemma step_back_vt_app: + assumes vt_ts: "vt cs (t@s)" + shows "vt cs 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 cs (t @ s) \ vt cs s" + and vt_et: "vt cs ((e # t) @ s)" + show ?case + proof(rule ih) + show "vt cs (t @ s)" + proof(rule step_back_vt) + from vt_et show "vt cs (e # t @ s)" by simp + qed + qed + qed +qed + +context extend_highest_set +begin + +lemma red_moment: + "extend_highest_set s' th prio (moment i t)" + apply (insert extend_highest_set_axioms, subst (asm) (1) moment_restm_s [of i t, symmetric]) + apply (unfold extend_highest_set_def extend_highest_set_axioms_def, clarsimp) + by (unfold highest_set_def, auto dest:step_back_vt_app) + +lemma ind [consumes 0, case_names Nil Cons, induct type]: + assumes + h0: "R []" + and h2: "\ e t. \vt step (t@s); step (t@s) e; + extend_highest_set s' th prio t; + extend_highest_set s' th prio (e#t); R t\ \ R (e#t)" + shows "R t" +proof - + from vt_t extend_highest_set_axioms show ?thesis + proof(induct t) + from h0 show "R []" . + next + case (Cons e t') + assume ih: "\vt step (t' @ s); extend_highest_set s' th prio t'\ \ R t'" + and vt_e: "vt step ((e # t') @ s)" + and et: "extend_highest_set s' th prio (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 step (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_set s' th prio t'" + by (unfold extend_highest_set_def extend_highest_set_axioms_def, auto dest:step_back_vt) + next + from vt_ts show "vt step (t' @ s)" . + qed + next + from et show "extend_highest_set s' th prio (e # t')" . + next + from et show ext': "extend_highest_set s' th prio t'" + by (unfold extend_highest_set_def extend_highest_set_axioms_def, auto dest:step_back_vt) + qed + qed +qed + +lemma th_kept: "th \ threads (t @ s) \ + preced th (t@s) = preced th s" (is "?Q t") +proof - + show ?thesis + proof(induct rule:ind) + case Nil + from threads_s + show "th \ threads ([] @ s) \ preced th ([] @ s) = preced th s" + by auto + next + case (Cons e t) + show ?case + proof(cases e) + case (Create thread prio) + assume eq_e: " e = Create thread prio" + show ?thesis + proof - + from Cons and eq_e have "step (t@s) (Create thread prio)" by auto + hence "th \ thread" + proof(cases) + assume "thread \ threads (t @ s)" + with Cons show ?thesis by auto + qed + hence "preced th ((e # t) @ s) = preced th (t @ s)" + by (unfold eq_e, auto simp:preced_def) + moreover note Cons + ultimately show ?thesis + by (auto simp:eq_e) + qed + next + case (Exit thread) + assume eq_e: "e = Exit thread" + from Cons have "extend_highest_set s' th prio (e # t)" by auto + from extend_highest_set.exit_diff [OF this] and eq_e + have neq_th: "thread \ th" by auto + with Cons + show ?thesis + by (unfold eq_e, auto simp:preced_def) + next + case (P thread cs) + assume eq_e: "e = P thread cs" + with Cons + show ?thesis + by (auto simp:eq_e preced_def) + next + case (V thread cs) + assume eq_e: "e = V thread cs" + with Cons + show ?thesis + by (auto simp:eq_e preced_def) + next + case (Set thread prio') + assume eq_e: " e = Set thread prio'" + show ?thesis + proof - + from Cons have "extend_highest_set s' th prio (e # t)" by auto + from extend_highest_set.set_diff_low[OF this] and eq_e + have "th \ thread" by auto + hence "preced th ((e # t) @ s) = preced th (t @ s)" + by (unfold eq_e, auto simp:preced_def) + moreover note Cons + ultimately show ?thesis + by (auto simp:eq_e) + qed + qed + qed +qed + +lemma max_kept: "Max ((\ th'. preced th' (t @ s)) ` (threads (t@s))) = preced th s" +proof(induct rule:ind) + case Nil + from highest_preced_thread + show "Max ((\th'. preced th' ([] @ s)) ` threads ([] @ s)) = preced th s" + by simp +next + case (Cons e t) + show ?case + proof(cases e) + case (Create thread prio') + assume eq_e: " e = Create thread prio'" + from Cons and eq_e have stp: "step (t@s) (Create thread prio')" by auto + hence neq_thread: "thread \ th" + proof(cases) + assume "thread \ threads (t @ s)" + moreover have "th \ threads (t@s)" + proof - + from Cons have "extend_highest_set s' th prio t" by auto + from extend_highest_set.th_kept[OF this] show ?thesis by (simp add:s_def) + qed + ultimately show ?thesis by auto + qed + from Cons have "extend_highest_set s' th prio t" by auto + from extend_highest_set.th_kept[OF this] + have h': " th \ threads (t @ s) \ preced th (t @ s) = preced th s" + by (auto simp:s_def) + from stp + have thread_ts: "thread \ threads (t @ s)" + by (cases, auto) + show ?thesis (is "Max (?f ` ?A) = ?t") + proof - + have "Max (?f ` ?A) = Max(insert (?f thread) (?f ` (threads (t@s))))" + by (unfold eq_e, simp) + moreover have "\ = max (?f thread) (Max (?f ` (threads (t@s))))" + proof(rule Max_insert) + from Cons have "vt step (t @ s)" by auto + from finite_threads[OF this] + show "finite (?f ` (threads (t@s)))" by simp + next + from h' show "(?f ` (threads (t@s))) \ {}" by auto + qed + moreover have "(Max (?f ` (threads (t@s)))) = ?t" + proof - + have "(\th'. preced th' ((e # t) @ s)) ` threads (t @ s) = + (\th'. preced th' (t @ s)) ` threads (t @ s)" (is "?f1 ` ?B = ?f2 ` ?B") + proof - + { fix th' + assume "th' \ ?B" + with thread_ts eq_e + have "?f1 th' = ?f2 th'" by (auto simp:preced_def) + } thus ?thesis + apply (auto simp:Image_def) + proof - + fix th' + assume h: "\th'. th' \ threads (t @ s) \ + preced th' (e # t @ s) = preced th' (t @ s)" + and h1: "th' \ threads (t @ s)" + show "preced th' (t @ s) \ (\th'. preced th' (e # t @ s)) ` threads (t @ s)" + proof - + from h1 have "?f1 th' \ ?f1 ` ?B" by auto + moreover from h[OF h1] have "?f1 th' = ?f2 th'" by simp + ultimately show ?thesis by simp + qed + qed + qed + with Cons show ?thesis by auto + qed + moreover have "?f thread < ?t" + proof - + from Cons have " extend_highest_set s' th prio (e # t)" by auto + from extend_highest_set.create_low[OF this] and eq_e + have "prio' \ prio" by auto + thus ?thesis + by (unfold eq_e, auto simp:preced_def s_def precedence_less_def) + qed + ultimately show ?thesis by (auto simp:max_def) + qed +next + case (Exit thread) + assume eq_e: "e = Exit thread" + from Cons have vt_e: "vt step (e#(t @ s))" by auto + from Cons and eq_e have stp: "step (t@s) (Exit thread)" by auto + from stp have thread_ts: "thread \ threads (t @ s)" + by(cases, unfold runing_def readys_def, auto) + from Cons have "extend_highest_set s' th prio (e # t)" by auto + from extend_highest_set.exit_diff[OF this] and eq_e + have neq_thread: "thread \ th" by auto + from Cons have "extend_highest_set s' th prio t" by auto + from extend_highest_set.th_kept[OF this, folded s_def] + have h': "th \ threads (t @ s) \ preced th (t @ s) = preced th s" . + show ?thesis (is "Max (?f ` ?A) = ?t") + proof - + have "threads (t@s) = insert thread ?A" + by (insert stp thread_ts, unfold eq_e, auto) + hence "Max (?f ` (threads (t@s))) = Max (?f ` \)" by simp + also from this have "\ = Max (insert (?f thread) (?f ` ?A))" by simp + also have "\ = max (?f thread) (Max (?f ` ?A))" + proof(rule Max_insert) + from finite_threads [OF vt_e] + show "finite (?f ` ?A)" by simp + next + from Cons have "extend_highest_set s' th prio (e # t)" by auto + from extend_highest_set.th_kept[OF this] + show "?f ` ?A \ {}" by (auto simp:s_def) + qed + finally have "Max (?f ` (threads (t@s))) = max (?f thread) (Max (?f ` ?A))" . + moreover have "Max (?f ` (threads (t@s))) = ?t" + proof - + from Cons show ?thesis + by (unfold eq_e, auto simp:preced_def) + qed + ultimately have "max (?f thread) (Max (?f ` ?A)) = ?t" by simp + moreover have "?f thread < ?t" + proof(unfold eq_e, simp add:preced_def, fold preced_def) + show "preced thread (t @ s) < ?t" + proof - + have "preced thread (t @ s) \ ?t" + proof - + from Cons + have "?t = Max ((\th'. preced th' (t @ s)) ` threads (t @ s))" + (is "?t = Max (?g ` ?B)") by simp + moreover have "?g thread \ \" + proof(rule Max_ge) + have "vt step (t@s)" by fact + from finite_threads [OF this] + show "finite (?g ` ?B)" by simp + next + from thread_ts + show "?g thread \ (?g ` ?B)" by auto + qed + ultimately show ?thesis by auto + qed + moreover have "preced thread (t @ s) \ ?t" + proof + assume "preced thread (t @ s) = preced th s" + with h' have "preced thread (t @ s) = preced th (t@s)" by simp + from preced_unique [OF this] have "thread = th" + proof + from h' show "th \ threads (t @ s)" by simp + next + from thread_ts show "thread \ threads (t @ s)" . + qed(simp) + with neq_thread show "False" by simp + qed + ultimately show ?thesis by auto + qed + qed + ultimately show ?thesis + by (auto simp:max_def split:if_splits) + qed + next + case (P thread cs) + with Cons + show ?thesis by (auto simp:preced_def) + next + case (V thread cs) + with Cons + show ?thesis by (auto simp:preced_def) + next + case (Set thread prio') + show ?thesis (is "Max (?f ` ?A) = ?t") + proof - + let ?B = "threads (t@s)" + from Cons have "extend_highest_set s' th prio (e # t)" by auto + from extend_highest_set.set_diff_low[OF this] and Set + have neq_thread: "thread \ th" and le_p: "prio' \ prio" by auto + from Set have "Max (?f ` ?A) = Max (?f ` ?B)" by simp + also have "\ = ?t" + proof(rule Max_eqI) + fix y + assume y_in: "y \ ?f ` ?B" + then obtain th1 where + th1_in: "th1 \ ?B" and eq_y: "y = ?f th1" by auto + show "y \ ?t" + proof(cases "th1 = thread") + case True + with neq_thread le_p eq_y s_def Set + show ?thesis + by (auto simp:preced_def precedence_le_def) + next + case False + with Set eq_y + have "y = preced th1 (t@s)" + by (simp add:preced_def) + moreover have "\ \ ?t" + proof - + from Cons + have "?t = Max ((\ th'. preced th' (t@s)) ` (threads (t@s)))" + by auto + moreover have "preced th1 (t@s) \ \" + proof(rule Max_ge) + from th1_in + show "preced th1 (t @ s) \ (\th'. preced th' (t @ s)) ` threads (t @ s)" + by simp + next + show "finite ((\th'. preced th' (t @ s)) ` threads (t @ s))" + proof - + from Cons have "vt step (t @ s)" by auto + from finite_threads[OF this] show ?thesis by auto + qed + qed + ultimately show ?thesis by auto + qed + ultimately show ?thesis by auto + qed + next + from Cons and finite_threads + show "finite (?f ` ?B)" by auto + next + from Cons have "extend_highest_set s' th prio t" by auto + from extend_highest_set.th_kept [OF this, folded s_def] + have h: "th \ threads (t @ s) \ preced th (t @ s) = preced th s" . + show "?t \ (?f ` ?B)" + proof - + from neq_thread Set h + have "?t = ?f th" by (auto simp:preced_def) + with h show ?thesis by auto + qed + qed + finally show ?thesis . + qed + qed +qed + +lemma max_preced: "preced th (t@s) = Max ((\ th'. preced th' (t @ s)) ` (threads (t@s)))" + by (insert th_kept max_kept, auto) + +lemma th_cp_max_preced: "cp (t@s) th = Max ((\ th'. preced th' (t @ s)) ` (threads (t@s)))" + (is "?L = ?R") +proof - + have "?L = cpreced (t@s) (wq (t@s)) th" + by (unfold cp_eq_cpreced, simp) + also have "\ = ?R" + proof(unfold cpreced_def) + show "Max ((\th. preced th (t @ s)) ` ({th} \ dependents (wq (t @ s)) th)) = + Max ((\th'. preced th' (t @ s)) ` threads (t @ s))" + (is "Max (?f ` ({th} \ ?A)) = Max (?f ` ?B)") + proof(cases "?A = {}") + case False + have "Max (?f ` ({th} \ ?A)) = Max (insert (?f th) (?f ` ?A))" by simp + moreover have "\ = max (?f th) (Max (?f ` ?A))" + proof(rule Max_insert) + show "finite (?f ` ?A)" + proof - + from dependents_threads[OF vt_t] + have "?A \ threads (t@s)" . + moreover from finite_threads[OF vt_t] have "finite \" . + ultimately show ?thesis + by (auto simp:finite_subset) + qed + next + from False show "(?f ` ?A) \ {}" by simp + qed + moreover have "\ = Max (?f ` ?B)" + proof - + from max_preced have "?f th = Max (?f ` ?B)" . + moreover have "Max (?f ` ?A) \ \" + proof(rule Max_mono) + from False show "(?f ` ?A) \ {}" by simp + next + show "?f ` ?A \ ?f ` ?B" + proof - + have "?A \ ?B" by (rule dependents_threads[OF vt_t]) + thus ?thesis by auto + qed + next + from finite_threads[OF vt_t] + show "finite (?f ` ?B)" by simp + qed + ultimately show ?thesis + by (auto simp:max_def) + qed + ultimately show ?thesis by auto + next + case True + with max_preced show ?thesis by auto + qed + qed + finally show ?thesis . +qed + +lemma th_cp_max: "cp (t@s) th = Max (cp (t@s) ` threads (t@s))" + by (unfold max_cp_eq[OF vt_t] th_cp_max_preced, simp) + +lemma th_cp_preced: "cp (t@s) th = preced th s" + by (fold max_kept, unfold th_cp_max_preced, simp) + +lemma preced_less': + fixes th' + assumes th'_in: "th' \ threads s" + and neq_th': "th' \ th" + shows "preced th' s < preced th s" +proof - + have "preced th' s \ Max ((\th'. preced th' s) ` threads s)" + proof(rule Max_ge) + from finite_threads [OF vt_s] + show "finite ((\th'. preced th' s) ` threads s)" by simp + next + from th'_in show "preced th' s \ (\th'. preced th' s) ` threads s" + by simp + qed + moreover have "preced th' s \ preced th s" + proof + assume "preced th' s = preced th s" + from preced_unique[OF this th'_in] neq_th' is_ready + show "False" by (auto simp:readys_def) + qed + ultimately show ?thesis using highest_preced_thread + by auto +qed + +lemma pv_blocked: + fixes th' + assumes th'_in: "th' \ threads (t@s)" + and neq_th': "th' \ th" + and eq_pv: "cntP (t@s) th' = cntV (t@s) th'" + shows "th' \ runing (t@s)" +proof + assume "th' \ runing (t@s)" + hence "cp (t@s) th' = Max (cp (t@s) ` readys (t@s))" + by (auto simp:runing_def) + with max_cp_readys_threads [OF vt_t] + have "cp (t @ s) th' = Max (cp (t@s) ` threads (t@s))" + by auto + moreover from th_cp_max have "cp (t @ s) th = \" by simp + ultimately have "cp (t @ s) th' = cp (t @ s) th" by simp + moreover from th_cp_preced and th_kept have "\ = preced th (t @ s)" + by simp + finally have h: "cp (t @ s) th' = preced th (t @ s)" . + show False + proof - + have "dependents (wq (t @ s)) th' = {}" + by (rule count_eq_dependents [OF vt_t eq_pv]) + moreover have "preced th' (t @ s) \ preced th (t @ s)" + proof + assume "preced th' (t @ s) = preced th (t @ s)" + hence "th' = th" + proof(rule preced_unique) + from th_kept show "th \ threads (t @ s)" by simp + next + from th'_in show "th' \ threads (t @ s)" by simp + qed + with assms show False by simp + qed + ultimately show ?thesis + by (insert h, unfold cp_eq_cpreced cpreced_def, simp) + qed +qed + +lemma runing_precond_pre: + fixes th' + assumes th'_in: "th' \ threads s" + and eq_pv: "cntP s th' = cntV s th'" + and neq_th': "th' \ th" + shows "th' \ threads (t@s) \ + cntP (t@s) th' = cntV (t@s) th'" +proof - + show ?thesis + proof(induct rule:ind) + case (Cons e t) + from Cons + have in_thread: "th' \ threads (t @ s)" + and not_holding: "cntP (t @ s) th' = cntV (t @ s) th'" by auto + have "extend_highest_set s' th prio t" by fact + from extend_highest_set.pv_blocked + [OF this, folded s_def, OF in_thread neq_th' not_holding] + have not_runing: "th' \ runing (t @ s)" . + show ?case + proof(cases e) + case (V thread cs) + from Cons and V have vt_v: "vt step (V thread cs#(t@s))" by auto + + show ?thesis + proof - + from Cons and V have "step (t@s) (V thread cs)" by auto + hence neq_th': "thread \ th'" + proof(cases) + assume "thread \ runing (t@s)" + moreover have "th' \ runing (t@s)" by fact + ultimately show ?thesis by auto + qed + with not_holding have cnt_eq: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'" + by (unfold V, simp add:cntP_def cntV_def count_def) + moreover from in_thread + have in_thread': "th' \ threads ((e # t) @ s)" by (unfold V, simp) + ultimately show ?thesis by auto + qed + next + case (P thread cs) + from Cons and P have "step (t@s) (P thread cs)" by auto + hence neq_th': "thread \ th'" + proof(cases) + assume "thread \ runing (t@s)" + moreover note not_runing + ultimately show ?thesis by auto + qed + with Cons and P have eq_cnt: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'" + by (auto simp:cntP_def cntV_def count_def) + moreover from Cons and P have in_thread': "th' \ threads ((e # t) @ s)" + by auto + ultimately show ?thesis by auto + next + case (Create thread prio') + from Cons and Create have "step (t@s) (Create thread prio')" by auto + hence neq_th': "thread \ th'" + proof(cases) + assume "thread \ threads (t @ s)" + moreover have "th' \ threads (t@s)" by fact + ultimately show ?thesis by auto + qed + with Cons and Create + have eq_cnt: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'" + by (auto simp:cntP_def cntV_def count_def) + moreover from Cons and Create + have in_thread': "th' \ threads ((e # t) @ s)" by auto + ultimately show ?thesis by auto + next + case (Exit thread) + from Cons and Exit have "step (t@s) (Exit thread)" by auto + hence neq_th': "thread \ th'" + proof(cases) + assume "thread \ runing (t @ s)" + moreover note not_runing + ultimately show ?thesis by auto + qed + with Cons and Exit + have eq_cnt: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'" + by (auto simp:cntP_def cntV_def count_def) + moreover from Cons and Exit and neq_th' + have in_thread': "th' \ threads ((e # t) @ s)" + by auto + ultimately show ?thesis by auto + 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 +qed + +(* +lemma runing_precond: + fixes th' + assumes th'_in: "th' \ threads s" + and eq_pv: "cntP s th' = cntV s th'" + and neq_th': "th' \ th" + shows "th' \ runing (t@s)" +proof - + from runing_precond_pre[OF th'_in eq_pv neq_th'] + have h1: "th' \ threads (t @ s)" and h2: "cntP (t @ s) th' = cntV (t @ s) th'" by auto + from pv_blocked[OF h1 neq_th' h2] + show ?thesis . +qed +*) + +lemma runing_precond: + fixes th' + assumes th'_in: "th' \ threads s" + and neq_th': "th' \ th" + and is_runing: "th' \ runing (t@s)" + shows "cntP s th' > cntV s th'" +proof - + have "cntP s th' \ cntV s th'" + proof + assume eq_pv: "cntP s th' = cntV s th'" + from runing_precond_pre[OF th'_in eq_pv neq_th'] + have h1: "th' \ threads (t @ s)" + and h2: "cntP (t @ s) th' = cntV (t @ s) th'" by auto + from pv_blocked[OF h1 neq_th' h2] have " th' \ runing (t @ s)" . + with is_runing show "False" by simp + qed + moreover from cnp_cnv_cncs[OF vt_s, of th'] + have "cntV s th' \ cntP s th'" by auto + ultimately show ?thesis by auto +qed + +lemma moment_blocked_pre: + assumes neq_th': "th' \ th" + and th'_in: "th' \ 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' \ + th' \ threads ((moment (i+j) t)@s)" +proof(induct j) + case (Suc k) + show ?case + proof - + { assume True: "Suc (i+k) \ length t" + from moment_head [OF this] + obtain e where + eq_me: "moment (Suc(i+k)) t = e#(moment (i+k) t)" + by blast + from red_moment[of "Suc(i+k)"] + and eq_me have "extend_highest_set s' th prio (e # moment (i + k) t)" by simp + hence vt_e: "vt step (e#(moment (i + k) t)@s)" + by (unfold extend_highest_set_def extend_highest_set_axioms_def + highest_set_def s_def, auto) + have not_runing': "th' \ runing (moment (i + k) t @ s)" + proof(unfold s_def) + show "th' \ runing (moment (i + k) t @ Set th prio # s')" + proof(rule extend_highest_set.pv_blocked) + from Suc show "th' \ threads (moment (i + k) t @ Set th prio # s')" + by (simp add:s_def) + next + from neq_th' show "th' \ th" . + next + from red_moment show "extend_highest_set s' th prio (moment (i + k) t)" . + next + from Suc show "cntP (moment (i + k) t @ Set th prio # s') th' = + cntV (moment (i + k) t @ Set th prio # s') th'" + by (auto simp:s_def) + qed + qed + from step_back_step[OF vt_e] + have "step ((moment (i + k) t)@s) e" . + hence "cntP (e#(moment (i + k) t)@s) th' = cntV (e#(moment (i + k) t)@s) th' \ + th' \ threads (e#(moment (i + k) t)@s) + " + proof(cases) + case (thread_create thread prio) + with Suc show ?thesis by (auto simp:cntP_def cntV_def count_def) + next + case (thread_exit thread) + moreover have "thread \ th'" + proof - + have "thread \ runing (moment (i + k) t @ s)" by fact + moreover note not_runing' + ultimately show ?thesis by auto + qed + moreover note Suc + ultimately show ?thesis by (auto simp:cntP_def cntV_def count_def) + next + case (thread_P thread cs) + moreover have "thread \ th'" + proof - + have "thread \ runing (moment (i + k) t @ s)" by fact + moreover note not_runing' + ultimately show ?thesis by auto + qed + moreover note Suc + ultimately show ?thesis by (auto simp:cntP_def cntV_def count_def) + next + case (thread_V thread cs) + moreover have "thread \ th'" + proof - + have "thread \ runing (moment (i + k) t @ s)" by fact + moreover note not_runing' + ultimately show ?thesis by auto + qed + moreover note Suc + ultimately show ?thesis by (auto simp:cntP_def cntV_def count_def) + next + case (thread_set thread prio') + with Suc show ?thesis + by (auto simp:cntP_def cntV_def count_def) + qed + with eq_me have ?thesis using eq_me by auto + } note h = this + show ?thesis + proof(cases "Suc (i+k) \ length t") + case True + from h [OF this] show ?thesis . + next + case False + with moment_ge + have eq_m: "moment (i + Suc k) t = moment (i+k) t" by auto + with Suc show ?thesis by auto + qed + qed +next + case 0 + from assms show ?case by auto +qed + +lemma moment_blocked: + assumes neq_th': "th' \ th" + and th'_in: "th' \ threads ((moment i t)@s)" + and eq_pv: "cntP ((moment i t)@s) th' = cntV ((moment i t)@s) th'" + and le_ij: "i \ j" + shows "cntP ((moment j t)@s) th' = cntV ((moment j t)@s) th' \ + th' \ threads ((moment j t)@s) \ + th' \ 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' \ threads ((moment j t)@s)" by auto + with extend_highest_set.pv_blocked [OF red_moment [of j], folded s_def, OF h2 neq_th' h1] + show ?thesis by auto +qed + +lemma runing_inversion_1: + assumes neq_th': "th' \ th" + and runing': "th' \ runing (t@s)" + shows "th' \ threads s \ cntV s th' < cntP s th'" +proof(cases "th' \ threads s") + case True + with runing_precond [OF this neq_th' runing'] show ?thesis by simp +next + case False + let ?Q = "\ t. th' \ threads (t@s)" + let ?q = "moment 0 t" + from moment_eq and False have not_thread: "\ ?Q ?q" by simp + from runing' have "th' \ threads (t@s)" by (simp add:runing_def readys_def) + from p_split_gen [of ?Q, OF this not_thread] + obtain i where lt_its: "i < length t" + and le_i: "0 \ i" + and pre: " th' \ threads (moment i t @ s)" (is "th' \ threads ?pre") + and post: "(\i'>i. th' \ threads (moment i' t @ s))" by auto + from lt_its have "Suc i \ length t" by auto + from moment_head[OF this] obtain e where + eq_me: "moment (Suc i) t = e # moment i t" by blast + from red_moment[of "Suc i"] and eq_me + have "extend_highest_set s' th prio (e # moment i t)" by simp + hence vt_e: "vt step (e#(moment i t)@s)" + by (unfold extend_highest_set_def extend_highest_set_axioms_def + highest_set_def s_def, auto) + from step_back_step[OF this] have stp_i: "step (moment i t @ s) e" . + from post[rule_format, of "Suc i"] and eq_me + have not_in': "th' \ threads (e # moment i t@s)" by auto + from create_pre[OF stp_i pre this] + obtain prio where eq_e: "e = Create th' prio" . + have "cntP (moment i t@s) th' = cntV (moment i t@s) th'" + proof(rule cnp_cnv_eq) + from step_back_vt [OF vt_e] + show "vt step (moment i t @ s)" . + next + from eq_e and stp_i + have "step (moment i t @ s) (Create th' prio)" by simp + thus "th' \ threads (moment i t @ s)" by (cases, simp) + qed + with eq_e + have "cntP ((e#moment i t)@s) th' = cntV ((e#moment i t)@s) th'" + by (simp add:cntP_def cntV_def count_def) + with eq_me[symmetric] + have h1: "cntP (moment (Suc i) t @ s) th' = cntV (moment (Suc i) t@ s) th'" + by simp + from eq_e have "th' \ threads ((e#moment i t)@s)" by simp + with eq_me [symmetric] + have h2: "th' \ threads (moment (Suc i) t @ s)" by simp + from moment_blocked [OF neq_th' h2 h1, of "length t"] and lt_its + and moment_ge + have "th' \ runing (t @ s)" by auto + with runing' + show ?thesis by auto +qed + +lemma runing_inversion_2: + assumes runing': "th' \ runing (t@s)" + shows "th' = th \ (th' \ th \ th' \ threads s \ cntV s th' < cntP s th')" +proof - + from runing_inversion_1[OF _ runing'] + show ?thesis by auto +qed + +lemma live: "runing (t@s) \ {}" +proof(cases "th \ runing (t@s)") + case True thus ?thesis by auto +next + case False + then have not_ready: "th \ readys (t@s)" + apply (unfold runing_def, + insert th_cp_max max_cp_readys_threads[OF vt_t, symmetric]) + by auto + from th_kept have "th \ threads (t@s)" by auto + from th_chain_to_ready[OF vt_t this] and not_ready + obtain th' where th'_in: "th' \ readys (t@s)" + and dp: "(Th th, Th th') \ (depend (t @ s))\<^sup>+" by auto + have "th' \ runing (t@s)" + proof - + have "cp (t @ s) th' = Max (cp (t @ s) ` readys (t @ s))" + proof - + have " Max ((\th. preced th (t @ s)) ` ({th'} \ dependents (wq (t @ s)) th')) = + preced th (t@s)" + proof(rule Max_eqI) + fix y + assume "y \ (\th. preced th (t @ s)) ` ({th'} \ dependents (wq (t @ s)) th')" + then obtain th1 where + h1: "th1 = th' \ th1 \ dependents (wq (t @ s)) th'" + and eq_y: "y = preced th1 (t@s)" by auto + show "y \ preced th (t @ s)" + proof - + from max_preced + have "preced th (t @ s) = Max ((\th'. preced th' (t @ s)) ` threads (t @ s))" . + moreover have "y \ \" + proof(rule Max_ge) + from h1 + have "th1 \ threads (t@s)" + proof + assume "th1 = th'" + with th'_in show ?thesis by (simp add:readys_def) + next + assume "th1 \ dependents (wq (t @ s)) th'" + with dependents_threads [OF vt_t] + show "th1 \ threads (t @ s)" by auto + qed + with eq_y show " y \ (\th'. preced th' (t @ s)) ` threads (t @ s)" by simp + next + from finite_threads[OF vt_t] + show "finite ((\th'. preced th' (t @ s)) ` threads (t @ s))" by simp + qed + ultimately show ?thesis by auto + qed + next + from finite_threads[OF vt_t] dependents_threads [OF vt_t, of th'] + show "finite ((\th. preced th (t @ s)) ` ({th'} \ dependents (wq (t @ s)) th'))" + by (auto intro:finite_subset) + next + from dp + have "th \ dependents (wq (t @ s)) th'" + by (unfold cs_dependents_def, auto simp:eq_depend) + thus "preced th (t @ s) \ + (\th. preced th (t @ s)) ` ({th'} \ dependents (wq (t @ s)) th')" + by auto + qed + moreover have "\ = Max (cp (t @ s) ` readys (t @ s))" + proof - + from max_preced and max_cp_eq[OF vt_t, symmetric] + have "preced th (t @ s) = Max (cp (t @ s) ` threads (t @ s))" by simp + with max_cp_readys_threads[OF vt_t] show ?thesis by simp + qed + ultimately show ?thesis by (unfold cp_eq_cpreced cpreced_def, simp) + qed + with th'_in show ?thesis by (auto simp:runing_def) + qed + thus ?thesis by auto +qed + +end + +end +