(* * Copyright 2014, General Dynamics C4 Systems * * This software may be distributed and modified according to the terms of * the GNU General Public License version 2. Note that NO WARRANTY is provided. * See "LICENSE_GPLv2.txt" for details. * * @TAG(GD_GPL) *) theory Schedule_R imports VSpace_R begin context begin interpretation Arch . (*FIXME: arch_split*) (* FIXME REMOVE and rename other uses, also there is a duplicate of this in KHeap_R too *) declare doMachineOp_cte_wp_at'[wp] declare static_imp_wp[wp_split del] (* FIXME: move *) lemma corres_gets_pre_lhs: "(\x. corres r (P x) P' (g x) g') \ corres r (\s. P (f s) s) P' (gets f >>= (\x. g x)) g'" by (simp add: corres_underlying_gets_pre_lhs) (* FIXME: move *) lemma corres_if_lhs: assumes "P \ corres r A Q f f'" assumes "\P \ corres r B Q g f'" shows "corres r (\s. (P \ A s) \ (\P \ B s)) Q (if P then f else g) f'" by (simp add: assms) (* Levity: added (20090713 10:04:12) *) declare sts_rel_idle [simp] (* FIXME: move to NonDetMonadVCG *) lemma return_wp_exs_valid [wp]: "\ P x \ return x \\ P \" by (simp add: exs_valid_def return_def) (* FIXME: move to NonDetMonadVCG *) lemma get_exs_valid [wp]: "\op = s\ get \\\r. op = s\" by (simp add: get_def exs_valid_def) lemma countLeadingZeros_word_clz[simp]: "countLeadingZeros w = word_clz w" unfolding countLeadingZeros_def word_clz_def by (simp add: to_bl_upt) lemma wordLog2_word_log2[simp]: "wordLog2 = word_log2" apply (rule ext) unfolding wordLog2_def word_log2_def by (simp add: word_size wordBits_def) lemma invs_no_cicd'_queues: "invs_no_cicd' s \ valid_queues s" unfolding invs_no_cicd'_def by simp lemma gsa_wf [wp]: "\invs'\ getSchedulerAction \sch_act_wf\" by (simp add: getSchedulerAction_def invs'_def valid_state'_def | wp)+ lemma corres_if2: "\ G = G'; G \ corres r P P' a c; \ G' \ corres r Q Q' b d \ \ corres r (if G then P else Q) (if G' then P' else Q') (if G then a else b) (if G' then c else d)" by simp lemma findM_awesome': assumes x: "\x xs. suffix (x # xs) xs' \ corres (\a b. if b then (\a'. a = Some a' \ r a' (Some x)) else a = None) P (P' (x # xs)) ((f >>= (\x. return (Some x))) OR (return None)) (g x)" assumes y: "corres r P (P' []) f (return None)" assumes z: "\x xs. suffix (x # xs) xs' \ \P' (x # xs)\ g x \\rv s. \ rv \ P' xs s\" assumes p: "suffix xs xs'" shows "corres r P (P' xs) f (findM g xs)" proof - have P: "f = do x \ (do x \ f; return (Some x) od) OR return None; if x \ None then return (the x) else f od" apply (rule ext) apply (auto simp add: bind_def alternative_def return_def split_def prod_eq_iff) done have Q: "\P\ (do x \ f; return (Some x) od) OR return None \\rv. if rv \ None then \ else P\" by (wp alternative_wp | simp)+ show ?thesis using p apply (induct xs) apply (simp add: y del: dc_simp) apply (simp only: findM.simps) apply (subst P) apply (rule corres_guard_imp) apply (rule corres_split [OF _ x]) apply (rule corres_if2) apply (case_tac ra, clarsimp+)[1] apply (rule corres_trivial, clarsimp) apply (case_tac ra, simp_all)[1] apply (erule(1) meta_mp [OF _ suffix_ConsD]) apply assumption apply (rule Q) apply (rule hoare_post_imp [OF _ z]) apply simp+ done qed lemmas findM_awesome = findM_awesome' [OF _ _ _ suffix_refl] lemma corres_rhs_disj_division: "\ P \ Q; P \ corres r R S x y; Q \ corres r T U x y \ \ corres r (R and T) (\s. (P \ S s) \ (Q \ U s)) x y" apply (rule corres_guard_imp) apply (erule corres_disj_division) apply simp+ done lemma findM_alternative_awesome: assumes x: "\x. corres (\a b. if b then (\a'. a = Some a') else a = None) P (P' and (\s. x \ fn s)) ((f >>= (\x. return (Some x))) OR (return None)) (g x)" assumes z: "\x xs. \\s. P' s \ x \ fn s \ set xs \ fn s\ g x \\rv s. \ rv \ P' s \ set xs \ fn s\" assumes on_none: "corres dc P P' f g'" shows "corres dc P (P' and (\s. set xs \ fn s)) f (findM g xs >>= (\x. when (x = None) g'))" proof - have P: "f = do x \ (do x \ f; return (Some x) od) OR return None; if x \ None then return (the x) else f od" apply (rule ext) apply (auto simp add: bind_def alternative_def return_def split_def prod_eq_iff) done have Q: "\P\ (do x \ f; return (Some x) od) OR return None \\rv. if rv \ None then \ else P\" by (wp alternative_wp | simp)+ have R: "\P x g xs. (do x \ if P then return (Some x) else findM g xs; when (x = None) g' od) = (if P then return () else (findM g xs >>= (\x . when (x = None) g')))" by (simp add: when_def) show ?thesis apply (induct xs) apply (simp add: when_def) apply (fold dc_def, rule on_none) apply (simp only: findM.simps bind_assoc) apply (subst P) apply (rule corres_guard_imp) apply (rule corres_split [OF _ x]) apply (subst R) apply (rule corres_if2) apply (case_tac xa, simp_all)[1] apply (rule corres_trivial, simp) apply assumption apply (rule Q) apply (rule hoare_post_imp [OF _ z]) apply simp+ done qed lemma awesome_case1: assumes x: "corres op = P P' (return False) (g x)" shows "corres (\a b. if b then (\a'. a = Some a' \ r a' (Some x)) else a = None) P P' ((f >>= (\x. return (Some x))) OR (return None)) (g x)" proof - have P: "return None = liftM (\x. None) (return False)" by (simp add: liftM_def) show ?thesis apply (rule corres_alternate2) apply (subst P, simp only: corres_liftM_simp, simp) apply (subst corres_cong [OF refl refl refl refl]) defer apply (rule x) apply (simp add: return_def) done qed lemma awesome_case2: assumes x: "corres (\a b. r a (Some x) \ b) P P' f (g x)" shows "corres (\a b. if b then (\a'. a = Some a' \ r a' (Some x)) else a = None) P P' ((f >>= (\x. return (Some x))) OR (return None)) (g x)" apply (rule corres_alternate1) apply (fold liftM_def) apply (simp add: o_def) apply (rule corres_rel_imp [OF x]) apply simp done lemma bind_select_corres: assumes x: "corres (\rvs rv'. \rv\rvs. r rv rv') P P' m m'" shows "corres r P P' (m >>= select) m'" apply (insert x) apply (clarsimp simp: corres_underlying_def bind_def select_def split_def) done lemma bind_select_fail_corres: assumes x: "corres (\rvs rv'. \rv\rvs. r rv rv') P P' m m'" shows "corres r P P' (m >>= (\x. select x \ fail)) m'" apply (insert x) apply (clarsimp simp: corres_underlying_def bind_def select_def fail_def alternative_def split_def) done lemma st_tcb_at_coerce_abstract: assumes t: "st_tcb_at' P t c" assumes sr: "(a, c) \ state_relation" shows "st_tcb_at (\st. \st'. thread_state_relation st st' \ P st') t a" using assms apply (clarsimp simp: state_relation_def pred_tcb_at'_def obj_at'_def projectKOs objBits_simps) apply (erule(1) pspace_dom_relatedE) apply (erule(1) obj_relation_cutsE, simp_all) apply (clarsimp simp: st_tcb_at_def obj_at_def other_obj_relation_def tcb_relation_def split: Structures_A.kernel_object.split_asm if_split_asm ARM_A.arch_kernel_obj.split_asm)+ apply fastforce done lemma runnable_coerce_abstract: "\ runnable' st'; thread_state_relation st st' \ \ runnable st" by (case_tac st, simp_all) (* Levity: added (20090721 10:56:29) *) declare objBitsT_koTypeOf [simp] lemma arch_switch_thread_corres: "corres dc (valid_arch_state and valid_objs and valid_asid_map and valid_arch_objs and pspace_aligned and pspace_distinct and valid_vs_lookup and valid_global_objs and unique_table_refs o caps_of_state and st_tcb_at runnable t) (valid_arch_state' and valid_pspace' and st_tcb_at' runnable' t) (arch_switch_to_thread t) (Arch.switchToThread t)" apply (simp add: arch_switch_to_thread_def ARM_H.switchToThread_def) apply (rule corres_guard_imp) apply (rule corres_split' [OF set_vm_root_corres]) apply (rule corres_machine_op[OF corres_rel_imp]) apply (rule corres_underlying_trivial) apply (simp add: ARM.clearExMonitor_def | wp)+ apply clarsimp apply (erule st_tcb_at_tcb_at) apply (clarsimp simp: valid_pspace'_def) done lemma tcbSchedAppend_corres: notes trans_state_update'[symmetric, simp del] shows "corres dc (is_etcb_at t) (tcb_at' t and Invariants_H.valid_queues and valid_queues') (tcb_sched_action (tcb_sched_append) t) (tcbSchedAppend t)" apply (simp only: tcbSchedAppend_def tcb_sched_action_def) apply (rule corres_symb_exec_r [OF _ _ threadGet_inv, where Q'="\rv. tcb_at' t and Invariants_H.valid_queues and valid_queues' and obj_at' (\obj. tcbQueued obj = rv) t"]) defer apply (wp threadGet_obj_at', simp, simp) apply (rule no_fail_pre, wp, simp) apply (case_tac queued) apply (simp add: unless_def when_def) apply (rule corres_no_failI) apply (rule no_fail_pre, wp) apply (clarsimp simp: in_monad ethread_get_def gets_the_def bind_assoc assert_opt_def exec_gets is_etcb_at_def get_etcb_def get_tcb_queue_def set_tcb_queue_def simpler_modify_def) apply (subgoal_tac "tcb_sched_append t (ready_queues a (tcb_domain y) (tcb_priority y)) = (ready_queues a (tcb_domain y) (tcb_priority y))") apply (simp add: state_relation_def ready_queues_relation_def) apply (clarsimp simp: tcb_sched_append_def state_relation_def valid_queues'_def ready_queues_relation_def ekheap_relation_def etcb_relation_def obj_at'_def inQ_def projectKO_eq project_inject) apply (drule_tac x=t in bspec,clarsimp) apply clarsimp apply (clarsimp simp: unless_def when_def cong: if_cong) apply (rule stronger_corres_guard_imp) apply (rule corres_split[where r'="op =", OF _ ethreadget_corres]) apply (rule corres_split[where r'="op =", OF _ ethreadget_corres]) apply (rule corres_split[where r'="op ="]) apply (rule corres_split_noop_rhs2) apply (rule corres_split_noop_rhs2) apply (rule threadSet_corres_noop, simp_all add: tcb_relation_def exst_same_def)[1] apply (rule addToBitmap_if_null_corres_noop) apply wp apply (simp add: tcb_sched_append_def) apply (intro conjI impI) apply (rule corres_guard_imp) apply (rule setQueue_corres) prefer 3 apply (rule_tac P=\ and Q="K (t \ set queuea)" in corres_assume_pre) apply (wp getQueue_corres getObject_tcb_wp | simp add: etcb_relation_def threadGet_def)+ apply (fastforce simp: valid_queues_def valid_queues_no_bitmap_def obj_at'_def inQ_def projectKO_eq project_inject) done crunch valid_pspace'[wp]: tcbSchedEnqueue valid_pspace' (simp: unless_def) crunch valid_pspace'[wp]: tcbSchedAppend valid_pspace' (simp: unless_def) crunch valid_pspace'[wp]: tcbSchedDequeue valid_pspace' crunch valid_arch_state'[wp]: tcbSchedEnqueue valid_arch_state' (simp: unless_def) crunch valid_arch_state'[wp]: tcbSchedAppend valid_arch_state' (simp: unless_def) crunch valid_arch_state'[wp]: tcbSchedDequeue valid_arch_state' crunch pred_tcb_at'[wp]: tcbSchedAppend, tcbSchedDequeue "pred_tcb_at' proj P t" (wp: threadSet_pred_tcb_no_state simp: unless_def tcb_to_itcb'_def ignore: getObject setObject) crunch state_refs_of'[wp]: setQueue "\s. P (state_refs_of' s)" lemma removeFromBitmap_valid_queues_no_bitmap[wp]: " \ valid_queues_no_bitmap \ removeFromBitmap d p \\_. valid_queues_no_bitmap \" unfolding bitmapQ_defs bitmap_fun_defs valid_queues_no_bitmap_def by (wp, clarsimp) lemma removeFromBitmap_valid_queues_no_bitmap_except[wp]: " \ valid_queues_no_bitmap_except t \ removeFromBitmap d p \\_. valid_queues_no_bitmap_except t \" unfolding bitmapQ_defs bitmap_fun_defs valid_queues_no_bitmap_except_def by (wp, clarsimp) lemma removeFromBitmap_bitmapQ: "\ \s. True \ removeFromBitmap d p \\_ s. \ bitmapQ d p s \" unfolding bitmapQ_defs bitmap_fun_defs apply wp apply (clarsimp simp: bitmapQ_def ucast_and_mask[symmetric] is_up unat_ucast_upcast) apply (subst (asm) complement_nth_w2p, simp_all) apply (fastforce intro!: order_less_le_trans[OF word_unat_mask_lt] simp: word_size wordRadix_def') done lemma removeFromBitmap_valid_bitmapQ[wp]: " \ valid_bitmapQ_except d p and bitmapQ_no_L2_orphans and bitmapQ_no_L1_orphans and (\s. ksReadyQueues s (d,p) = []) \ removeFromBitmap d p \\_. valid_bitmapQ \" proof - have "\ valid_bitmapQ_except d p and bitmapQ_no_L2_orphans and bitmapQ_no_L1_orphans and (\s. ksReadyQueues s (d,p) = []) \ removeFromBitmap d p \\_. valid_bitmapQ_except d p and bitmapQ_no_L2_orphans and bitmapQ_no_L1_orphans and (\s. \ bitmapQ d p s \ ksReadyQueues s (d,p) = []) \" by (rule hoare_pre) (wp removeFromBitmap_valid_queues_no_bitmap_except removeFromBitmap_valid_bitmapQ_except removeFromBitmap_bitmapQ, simp) thus ?thesis by - (erule hoare_strengthen_post; fastforce elim: valid_bitmap_valid_bitmapQ_exceptE) qed lemma threadSet_bitmapQ_no_L1_orphans[wp]: "\ bitmapQ_no_L1_orphans \ threadSet f t \\_. bitmapQ_no_L1_orphans \" unfolding bitmapQ_defs threadSet_def by (wp hoare_vcg_all_lift hoare_vcg_imp_lift) simp lemma threadSet_bitmapQ_no_L2_orphans[wp]: "\ bitmapQ_no_L2_orphans \ threadSet f t \\_. bitmapQ_no_L2_orphans \" unfolding bitmapQ_defs threadSet_def by (wp hoare_vcg_all_lift hoare_vcg_imp_lift) simp lemma threadSet_valid_bitmapQ[wp]: "\ valid_bitmapQ \ threadSet f t \\_. valid_bitmapQ \" unfolding bitmapQ_defs threadSet_def by (wp hoare_vcg_all_lift hoare_vcg_imp_lift hoare_vcg_conj_lift | wps)+ lemma setQueue_valid_queues_no_bitmap_except: "\ valid_queues_no_bitmap_except t and (\s. \t \ set ts. obj_at' (inQ d p and runnable' \ tcbState) t s) and K (t \ set ts) and K (distinct ts) and K (d \ maxDomain \ p \ maxPriority)\ setQueue d p ts \\_. valid_queues_no_bitmap_except t \" unfolding setQueue_def valid_queues_def by (wp, fastforce simp: valid_queues_no_bitmap_except_def) (* this should be the actual weakest precondition to establish valid_queues under tagging a thread as not queued *) lemma threadSet_valid_queues_dequeue_wp: "\ valid_queues_no_bitmap_except t and valid_bitmapQ and bitmapQ_no_L2_orphans and bitmapQ_no_L1_orphans and (\s. \d p. t \ set (ksReadyQueues s (d,p))) \ threadSet (tcbQueued_update (\_. False)) t \\rv. valid_queues \" unfolding threadSet_def apply (rule hoare_seq_ext[OF _ getObject_tcb_sp]) apply (rule hoare_pre) apply (simp add: valid_queues_def valid_queues_no_bitmap_except_def valid_queues_no_bitmap_def) apply (wp setObject_queues_unchanged_tcb hoare_Ball_helper hoare_vcg_all_lift setObject_tcb_strongest) apply (clarsimp simp: valid_queues_no_bitmap_except_def obj_at'_def valid_queues_no_bitmap_def) done (* FIXME move *) lemmas obj_at'_conjI = obj_at_conj' lemma obj_at'_disj: "\ obj_at' (\x. P x s) t s \ obj_at' (\x. Q x s) t s \ \ obj_at' (\x. P x s \ Q x s) t s" by (auto simp: obj_at'_def) (* FIXME move, also not_obj_at' is too type constrained *) lemma obj_at'_imp: "\ obj_at' (\(_::'a::pspace_storable). True) t s ; obj_at' (\(x::'a). P x s) t s \ obj_at' (\x. Q x s) t s \ \ obj_at' (\x. P x s \ Q x s) t s" by (auto simp: obj_at'_def) lemma setQueue_valid_queues_no_bitmap_except_dequeue_wp: "\d p ts t. \ \s. valid_queues_no_bitmap_except t s \ (\t' \ set ts. obj_at' (inQ d p and runnable' \ tcbState) t' s) \ t \ set ts \ distinct ts \ p \ maxPriority \ d \ maxDomain \ setQueue d p ts \\rv. valid_queues_no_bitmap_except t \" unfolding setQueue_def valid_queues_no_bitmap_except_def null_def by wp force definition (* if t is in a queue, it should be tagged with right priority and domain *) "correct_queue t s \ \d p. t \ set(ksReadyQueues s (d, p)) \ (obj_at' (\tcb. tcbQueued tcb \ tcbDomain tcb = d \ tcbPriority tcb = p) t s)" lemma valid_queues_no_bitmap_correct_queueI[intro]: "valid_queues_no_bitmap s \ correct_queue t s" unfolding correct_queue_def valid_queues_no_bitmap_def by (fastforce simp: obj_at'_def inQ_def) lemma tcbSchedDequeue_valid_queues_weak: "\ valid_queues_no_bitmap_except t and valid_bitmapQ and bitmapQ_no_L2_orphans and bitmapQ_no_L1_orphans and correct_queue t and obj_at' (\tcb. tcbDomain tcb \ maxDomain \ tcbPriority tcb \ maxPriority) t \ tcbSchedDequeue t \\_. Invariants_H.valid_queues\" proof - show ?thesis unfolding tcbSchedDequeue_def null_def valid_queues_def apply (rule hoare_pre) apply wp (* stops on threadSet *) apply (rule hoare_post_eq[OF _ threadSet_valid_queues_dequeue_wp], simp add: valid_queues_def) apply (wp hoare_vcg_if_lift hoare_vcg_conj_lift hoare_vcg_imp_lift) apply (wp hoare_vcg_imp_lift setQueue_valid_queues_no_bitmap_except_dequeue_wp setQueue_valid_bitmapQ threadGet_const_tcb_at) (* wp done *) apply (normalise_obj_at') apply (clarsimp simp: correct_queue_def) apply (normalise_obj_at') apply (fastforce simp add: valid_queues_no_bitmap_except_def valid_queues_no_bitmap_def elim: obj_at'_weaken)+ done qed lemma tcbSchedDequeue_valid_queues: "\Invariants_H.valid_queues and obj_at' (\tcb. tcbDomain tcb \ maxDomain) t and obj_at' (\tcb. tcbPriority tcb \ maxPriority) t\ tcbSchedDequeue t \\_. Invariants_H.valid_queues\" apply (rule hoare_pre, rule tcbSchedDequeue_valid_queues_weak) apply (fastforce simp: valid_queues_def valid_queues_no_bitmap_def obj_at'_def inQ_def) done lemma tcbSchedAppend_valid_queues'[wp]: (* most of this is identical to tcbSchedEnqueue_valid_queues' in TcbAcc_R *) "\valid_queues' and tcb_at' t\ tcbSchedAppend t \\_. valid_queues'\" apply (simp add: tcbSchedAppend_def) apply (rule hoare_pre) apply (rule_tac B="\rv. valid_queues' and obj_at' (\obj. tcbQueued obj = rv) t" in hoare_seq_ext) apply (rename_tac queued) apply (case_tac queued; simp_all add: unless_def when_def) apply (wp threadSet_valid_queues' setQueue_valid_queues' | simp)+ apply (subst conj_commute, wp) apply (rule hoare_pre_post, assumption) apply (clarsimp simp: addToBitmap_def modifyReadyQueuesL1Bitmap_def modifyReadyQueuesL2Bitmap_def getReadyQueuesL1Bitmap_def getReadyQueuesL2Bitmap_def) apply wp apply fastforce apply wp apply (subst conj_commute) apply (clarsimp simp: imp_conj_taut) apply (rule_tac Q="\rv. valid_queues' and obj_at' (\obj. \ tcbQueued obj) t and obj_at' (\obj. tcbPriority obj = prio) t and obj_at' (\obj. tcbDomain obj = tdom) t and (\s. t \ set (ksReadyQueues s (tdom, prio)))" in hoare_post_imp) apply (clarsimp simp: valid_queues'_def obj_at'_def projectKOs inQ_def) apply (wp setQueue_valid_queues' | simp | simp add: setQueue_def)+ apply (wp getObject_tcb_wp | simp add: threadGet_def)+ apply (clarsimp simp: obj_at'_def inQ_def projectKOs valid_queues'_def) apply (wp getObject_tcb_wp | simp add: threadGet_def)+ apply (clarsimp simp: obj_at'_def) done crunch norq[wp]: threadSet "\s. P (ksReadyQueues s)" (simp: updateObject_default_def ignore: setObject getObject) lemma removeFromBitmap_valid_queues'[wp]: "\ valid_queues' \ removeFromBitmap d p \\_. valid_queues' \" unfolding valid_queues'_def bitmap_fun_defs by (wp, simp) lemma threadSet_valid_queues'_dequeue: (* threadSet_valid_queues' is too weak for dequeue *) "\\s. (\d p t'. obj_at' (inQ d p) t' s \ t' \ t \ t' \ set (ksReadyQueues s (d, p))) \ obj_at' (inQ d p) t s \ threadSet (tcbQueued_update (\_. False)) t \\rv. valid_queues' \" unfolding valid_queues'_def apply (rule hoare_pre) apply (wp hoare_vcg_all_lift) apply (simp only: imp_conv_disj not_obj_at') apply (wp hoare_vcg_all_lift hoare_vcg_disj_lift) apply (simp add: not_obj_at') apply (clarsimp simp: typ_at_tcb') apply normalise_obj_at' apply (fastforce elim: obj_at'_weaken simp: inQ_def) done lemma setQueue_ksReadyQueues_lift: "\ \s. P (s\ksReadyQueues := (ksReadyQueues s)((d, p) := ts)\) ts \ setQueue d p ts \ \_ s. P s (ksReadyQueues s (d,p))\" unfolding setQueue_def by (wp, clarsimp simp: fun_upd_def) lemma tcbSchedDequeue_valid_queues'[wp]: "\valid_queues' and tcb_at' t\ tcbSchedDequeue t \\_. valid_queues'\" unfolding tcbSchedDequeue_def apply (rule_tac B="\rv. valid_queues' and obj_at' (\obj. tcbQueued obj = rv) t" in hoare_seq_ext) prefer 2 apply (wp threadGet_const_tcb_at) apply (fastforce simp: obj_at'_def) apply clarsimp apply (rename_tac queued) apply (case_tac queued, simp_all) apply wp apply (rule_tac d=tdom and p=prio in threadSet_valid_queues'_dequeue) apply (rule hoare_pre_post, assumption) apply (wp | clarsimp simp: bitmap_fun_defs)+ apply (wp hoare_vcg_all_lift setQueue_ksReadyQueues_lift) apply clarsimp apply (wp threadGet_obj_at' threadGet_const_tcb_at) apply clarsimp apply (rule context_conjI, clarsimp simp: obj_at'_def) apply (clarsimp simp: valid_queues'_def obj_at'_def projectKOs inQ_def|wp)+ done crunch tcb_at'[wp]: tcbSchedEnqueue "tcb_at' t" (simp: unless_def) crunch tcb_at'[wp]: tcbSchedAppend "tcb_at' t" (simp: unless_def) crunch tcb_at'[wp]: tcbSchedDequeue "tcb_at' t" crunch nosch[wp]: tcbSchedEnqueue "\s. P (ksSchedulerAction s)" (simp: unless_def) crunch nosch[wp]: tcbSchedAppend "\s. P (ksSchedulerAction s)" (simp: unless_def) crunch nosch[wp]: tcbSchedDequeue "\s. P (ksSchedulerAction s)" crunch state_refs_of'[wp]: tcbSchedEnqueue "\s. P (state_refs_of' s)" (wp: refl simp: crunch_simps unless_def) crunch state_refs_of'[wp]: tcbSchedAppend "\s. P (state_refs_of' s)" (wp: refl simp: crunch_simps unless_def) crunch state_refs_of'[wp]: tcbSchedDequeue "\s. P (state_refs_of' s)" (wp: refl simp: crunch_simps) crunch cap_to'[wp]: tcbSchedEnqueue "ex_nonz_cap_to' p" (simp: unless_def) crunch cap_to'[wp]: tcbSchedAppend "ex_nonz_cap_to' p" (simp: unless_def) crunch cap_to'[wp]: tcbSchedDequeue "ex_nonz_cap_to' p" crunch iflive'[wp]: setQueue if_live_then_nonz_cap' lemma tcbSchedAppend_iflive'[wp]: "\if_live_then_nonz_cap' and ex_nonz_cap_to' tcb\ tcbSchedAppend tcb \\_. if_live_then_nonz_cap'\" apply (simp add: tcbSchedAppend_def unless_def) apply (wp threadSet_iflive' hoare_drop_imps | simp add: crunch_simps)+ done lemma tcbSchedDequeue_iflive'[wp]: "\if_live_then_nonz_cap'\ tcbSchedDequeue tcb \\_. if_live_then_nonz_cap'\" apply (simp add: tcbSchedDequeue_def) apply (wp threadSet_iflive' hoare_when_weak_wp | simp add: crunch_simps)+ apply ((wp | clarsimp simp: bitmap_fun_defs)+)[1] (* deal with removeFromBitmap *) apply (wp threadSet_iflive' hoare_when_weak_wp | simp add: crunch_simps)+ apply (rule_tac Q="\rv. \" in hoare_post_imp, fastforce) apply (wp | simp add: crunch_simps)+ done crunch ifunsafe'[wp]: tcbSchedEnqueue if_unsafe_then_cap' (simp: unless_def) crunch ifunsafe'[wp]: tcbSchedAppend if_unsafe_then_cap' (simp: unless_def) crunch ifunsafe'[wp]: tcbSchedDequeue if_unsafe_then_cap' crunch idle'[wp]: tcbSchedEnqueue valid_idle' (simp: crunch_simps unless_def) crunch idle'[wp]: tcbSchedAppend valid_idle' (simp: crunch_simps unless_def) crunch idle'[wp]: tcbSchedDequeue valid_idle' (simp: crunch_simps) crunch global_refs'[wp]: tcbSchedEnqueue valid_global_refs' (wp: threadSet_global_refs simp: unless_def ignore: getObject setObject) crunch global_refs'[wp]: tcbSchedAppend valid_global_refs' (wp: threadSet_global_refs simp: unless_def) crunch global_refs'[wp]: tcbSchedDequeue valid_global_refs' (wp: threadSet_global_refs) crunch irq_node'[wp]: tcbSchedEnqueue "\s. P (irq_node' s)" (simp: unless_def) crunch irq_node'[wp]: tcbSchedAppend "\s. P (irq_node' s)" (simp: unless_def) crunch irq_node'[wp]: tcbSchedDequeue "\s. P (irq_node' s)" crunch typ_at'[wp]: tcbSchedEnqueue "\s. P (typ_at' T p s)" (simp: unless_def) crunch typ_at'[wp]: tcbSchedAppend "\s. P (typ_at' T p s)" (simp: unless_def) crunch typ_at'[wp]: tcbSchedDequeue "\s. P (typ_at' T p s)" crunch ctes_of[wp]: tcbSchedEnqueue "\s. P (ctes_of s)" (simp: unless_def) crunch ctes_of[wp]: tcbSchedAppend "\s. P (ctes_of s)" (simp: unless_def) crunch ctes_of[wp]: tcbSchedDequeue "\s. P (ctes_of s)" crunch ksInterrupt[wp]: tcbSchedEnqueue "\s. P (ksInterruptState s)" (simp: unless_def) crunch ksInterrupt[wp]: tcbSchedAppend "\s. P (ksInterruptState s)" (simp: unless_def) crunch ksInterrupt[wp]: tcbSchedDequeue "\s. P (ksInterruptState s)" crunch irq_states[wp]: tcbSchedEnqueue valid_irq_states' (simp: unless_def) crunch irq_states[wp]: tcbSchedAppend valid_irq_states' (simp: unless_def) crunch irq_states[wp]: tcbSchedDequeue valid_irq_states' crunch ct'[wp]: tcbSchedEnqueue "\s. P (ksCurThread s)" (simp: unless_def) crunch ct'[wp]: tcbSchedAppend "\s. P (ksCurThread s)" (simp: unless_def) crunch ct'[wp]: tcbSchedDequeue "\s. P (ksCurThread s)" crunch pde_mappings'[wp]: tcbSchedEnqueue "valid_pde_mappings'" (simp: unless_def) crunch pde_mappings'[wp]: tcbSchedAppend "valid_pde_mappings'" (simp: unless_def) crunch pde_mappings'[wp]: tcbSchedDequeue "valid_pde_mappings'" lemma tcbSchedEnqueue_vms'[wp]: "\valid_machine_state'\ tcbSchedEnqueue t \\_. valid_machine_state'\" apply (simp add: valid_machine_state'_def pointerInUserData_def pointerInDeviceData_def) apply (wp hoare_vcg_all_lift hoare_vcg_disj_lift tcbSchedEnqueue_ksMachine) done crunch ksCurDomain[wp]: tcbSchedEnqueue "\s. P (ksCurDomain s)" (simp: unless_def) lemma tcbSchedEnqueue_tcb_in_cur_domain'[wp]: "\tcb_in_cur_domain' t'\ tcbSchedEnqueue t \\_. tcb_in_cur_domain' t' \" apply (rule tcb_in_cur_domain'_lift) apply wp apply (clarsimp simp: tcbSchedEnqueue_def) apply wp apply (case_tac queued, simp_all add: unless_def when_def) apply (wp | simp)+ done lemma ct_idle_or_in_cur_domain'_lift2: "\ \t. \tcb_in_cur_domain' t\ f \\_. tcb_in_cur_domain' t\; \P. \\s. P (ksCurThread s) \ f \\_ s. P (ksCurThread s) \; \P. \\s. P (ksIdleThread s) \ f \\_ s. P (ksIdleThread s) \; \P. \\s. P (ksSchedulerAction s) \ f \\_ s. P (ksSchedulerAction s) \\ \ \ ct_idle_or_in_cur_domain'\ f \\_. ct_idle_or_in_cur_domain' \" apply (unfold ct_idle_or_in_cur_domain'_def) apply (rule hoare_lift_Pf2[where f=ksCurThread]) apply (rule hoare_lift_Pf2[where f=ksSchedulerAction]) apply (wp static_imp_wp hoare_vcg_disj_lift ) apply simp+ done lemma tcbSchedEnqueue_invs'[wp]: "\invs' and st_tcb_at' runnable' t and (\s. ksSchedulerAction s = ResumeCurrentThread \ ksCurThread s \ t)\ tcbSchedEnqueue t \\_. invs'\" apply (simp add: invs'_def valid_state'_def) apply (rule hoare_pre) apply (wp tcbSchedEnqueue_ct_not_inQ valid_irq_node_lift irqs_masked_lift hoare_vcg_disj_lift valid_irq_handlers_lift' cur_tcb_lift ct_idle_or_in_cur_domain'_lift2 untyped_ranges_zero_lift | simp add: cteCaps_of_def o_def | auto elim!: st_tcb_ex_cap'' valid_objs'_maxDomain valid_objs'_maxPriority split: thread_state.split_asm simp: valid_pspace'_def)+ done lemma tcb_at'_has_tcbPriority: "tcb_at' t s \ \p. obj_at' (\tcb. tcbPriority tcb = p) t s" by (clarsimp simp add: obj_at'_def) lemma tcb_at'_has_tcbDomain: "tcb_at' t s \ \p. obj_at' (\tcb. tcbDomain tcb = p) t s" by (clarsimp simp add: obj_at'_def) lemma tcbSchedEnqueue_in_ksQ: "\valid_queues' and tcb_at' t\ tcbSchedEnqueue t \\r s. \domain priority. t \ set (ksReadyQueues s (domain, priority))\" apply (rule_tac Q="\s. \d p. valid_queues' s \ obj_at' (\tcb. tcbPriority tcb = p) t s \ obj_at' (\tcb. tcbDomain tcb = d) t s" in hoare_pre_imp) apply (clarsimp simp: tcb_at'_has_tcbPriority tcb_at'_has_tcbDomain) apply (wp hoare_vcg_ex_lift) apply (simp add: tcbSchedEnqueue_def unless_def) apply (wp) apply clarsimp apply wp apply (rule_tac Q="\rv s. tdom = d \ rv = p \ obj_at' (\tcb. tcbPriority tcb = p) t s \ obj_at' (\tcb. tcbDomain tcb = d) t s" in hoare_post_imp, clarsimp) apply (wp, wp threadGet_const) apply (rule_tac Q="\rv s. obj_at' (\tcb. tcbPriority tcb = p) t s \ obj_at' (\tcb. tcbDomain tcb = d) t s \ obj_at' (\tcb. tcbQueued tcb = rv) t s \ (rv \ t \ set (ksReadyQueues s (d, p)))" in hoare_post_imp) apply (clarsimp simp: o_def elim!: obj_at'_weakenE) apply (wp threadGet_obj_at' hoare_vcg_imp_lift threadGet_const) apply (case_tac "obj_at' (Not \ tcbQueued) t s") apply (clarsimp) apply (clarsimp simp: valid_queues'_def) apply (drule_tac x=d in spec) apply (drule_tac x=p in spec) apply (drule_tac x=t in spec) apply (subgoal_tac "obj_at' (inQ d p) t s", clarsimp) apply (clarsimp simp: obj_at'_def inQ_def)+ done crunch ksMachine[wp]: tcbSchedAppend "\s. P (ksMachineState s)" (simp: unless_def) lemma tcbSchedAppend_vms'[wp]: "\valid_machine_state'\ tcbSchedAppend t \\_. valid_machine_state'\" apply (simp add: valid_machine_state'_def pointerInUserData_def pointerInDeviceData_def) apply (wp hoare_vcg_all_lift hoare_vcg_disj_lift tcbSchedAppend_ksMachine) done crunch pspace_domain_valid[wp]: tcbSchedAppend "pspace_domain_valid" (simp: unless_def) lemma tcbSchedAppend_ct_not_inQ: "\ ct_not_inQ and (\s. t \ ksCurThread s) \ tcbSchedAppend t \ \_. ct_not_inQ \" (is "\?PRE\ _ \_\") proof - have ts: "\?PRE\ threadSet (tcbQueued_update (\_. True)) t \\_. ct_not_inQ\" apply (simp add: ct_not_inQ_def) apply (rule hoare_weaken_pre) apply (wps setObject_ct_inv) apply (wp static_imp_wp, clarsimp simp: comp_def) done have sq: "\p q d. \ct_not_inQ\ setQueue d p q \\_. ct_not_inQ\" apply (simp add: ct_not_inQ_def setQueue_def) apply (wp) apply (clarsimp) done show ?thesis apply (simp add: tcbSchedAppend_def unless_def) apply (wp ts sq | clarsimp)+ apply (rule_tac Q="\_. ?PRE" in hoare_post_imp, clarsimp) apply (wp) done qed crunch ksCurDomain[wp]: tcbSchedAppend "\s. P (ksCurDomain s)" (simp: unless_def) crunch ksIdleThread[wp]: tcbSchedAppend "\s. P (ksIdleThread s)" (simp: unless_def) crunch ksDomSchedule[wp]: tcbSchedAppend "\s. P (ksDomSchedule s)" (simp: unless_def) lemma tcbSchedAppend_tcbDomain[wp]: "\ obj_at' (\tcb. P (tcbDomain tcb)) t' \ tcbSchedAppend t \ \_. obj_at' (\tcb. P (tcbDomain tcb)) t' \" apply (clarsimp simp: tcbSchedAppend_def) apply wp apply (case_tac queued, simp_all add: unless_def when_def) apply (wp | simp)+ done lemma tcbSchedAppend_tcbPriority[wp]: "\ obj_at' (\tcb. P (tcbPriority tcb)) t' \ tcbSchedAppend t \ \_. obj_at' (\tcb. P (tcbPriority tcb)) t' \" apply (clarsimp simp: tcbSchedAppend_def) apply wp apply (case_tac queued, simp_all add: unless_def when_def) apply (wp | simp)+ done lemma tcbSchedAppend_tcb_in_cur_domain'[wp]: "\tcb_in_cur_domain' t'\ tcbSchedAppend t \\_. tcb_in_cur_domain' t' \" apply (rule tcb_in_cur_domain'_lift) apply wp done crunch ksDomScheduleIdx[wp]: tcbSchedAppend "\s. P (ksDomScheduleIdx s)" (simp: unless_def) crunch gsUntypedZeroRanges[wp]: tcbSchedAppend, tcbSchedDequeue "\s. P (gsUntypedZeroRanges s)" (simp: unless_def) lemma tcbSchedAppend_invs'[wp]: "\invs' and st_tcb_at' runnable' t and (\s. t \ ksCurThread s) \ tcbSchedAppend t \\_. invs'\" apply (simp add: invs'_def valid_state'_def) apply (rule hoare_pre) apply (wp tcbSchedAppend_ct_not_inQ sch_act_wf_lift valid_irq_node_lift irqs_masked_lift valid_irq_handlers_lift' cur_tcb_lift ct_idle_or_in_cur_domain'_lift2 untyped_ranges_zero_lift | simp add: cteCaps_of_def o_def | fastforce elim!: st_tcb_ex_cap'' split: thread_state.split_asm)+ done lemma tcbSchedAppend_invs_but_ct_not_inQ': "\invs' and st_tcb_at' runnable' t and tcb_in_cur_domain' t \ tcbSchedAppend t \\_. all_invs_but_ct_not_inQ'\" apply (simp add: invs'_def valid_state'_def) apply (rule hoare_pre) apply (wp sch_act_wf_lift valid_irq_node_lift irqs_masked_lift valid_irq_handlers_lift' cur_tcb_lift ct_idle_or_in_cur_domain'_lift2 untyped_ranges_zero_lift | simp add: cteCaps_of_def o_def | fastforce elim!: st_tcb_ex_cap'' split: thread_state.split_asm)+ done crunch ksMachine[wp]: tcbSchedDequeue "\s. P (ksMachineState s)" (simp: unless_def) lemma tcbSchedDequeue_vms'[wp]: "\valid_machine_state'\ tcbSchedDequeue t \\_. valid_machine_state'\" apply (simp add: valid_machine_state'_def pointerInUserData_def pointerInDeviceData_def) apply (wp hoare_vcg_all_lift hoare_vcg_disj_lift tcbSchedDequeue_ksMachine) done crunch pspace_domain_valid[wp]: tcbSchedDequeue "pspace_domain_valid" crunch ksCurDomain[wp]: tcbSchedDequeue "\s. P (ksCurDomain s)" (simp: unless_def) crunch ksIdleThread[wp]: tcbSchedDequeue "\s. P (ksIdleThread s)" (simp: unless_def) crunch ksDomSchedule[wp]: tcbSchedDequeue "\s. P (ksDomSchedule s)" (simp: unless_def) crunch ksDomScheduleIdx[wp]: tcbSchedDequeue "\s. P (ksDomScheduleIdx s)" (simp: unless_def) lemma tcbSchedDequeue_tcb_in_cur_domain'[wp]: "\tcb_in_cur_domain' t'\ tcbSchedDequeue t \\_. tcb_in_cur_domain' t' \" apply (rule tcb_in_cur_domain'_lift) apply wp apply (clarsimp simp: tcbSchedDequeue_def) apply (wp hoare_when_weak_wp | simp)+ done lemma tcbSchedDequeue_tcbDomain[wp]: "\ obj_at' (\tcb. P (tcbDomain tcb)) t' \ tcbSchedDequeue t \ \_. obj_at' (\tcb. P (tcbDomain tcb)) t' \" apply (clarsimp simp: tcbSchedDequeue_def) apply (wp hoare_when_weak_wp | simp)+ done lemma tcbSchedDequeue_tcbPriority[wp]: "\ obj_at' (\tcb. P (tcbPriority tcb)) t' \ tcbSchedDequeue t \ \_. obj_at' (\tcb. P (tcbPriority tcb)) t' \" apply (clarsimp simp: tcbSchedDequeue_def) apply (wp hoare_when_weak_wp | simp)+ done lemma tcbSchedDequeue_invs'[wp]: "\invs' and tcb_at' t\ tcbSchedDequeue t \\_. invs'\" unfolding invs'_def valid_state'_def apply (rule hoare_pre) apply (wp tcbSchedDequeue_ct_not_inQ sch_act_wf_lift valid_irq_node_lift irqs_masked_lift valid_irq_handlers_lift' cur_tcb_lift ct_idle_or_in_cur_domain'_lift2 tcbSchedDequeue_valid_queues untyped_ranges_zero_lift | simp add: cteCaps_of_def o_def)+ apply (fastforce elim: valid_objs'_maxDomain valid_objs'_maxPriority simp: valid_pspace'_def)+ done lemma no_fail_isRunnable[wp]: "no_fail (tcb_at' t) (isRunnable t)" apply (simp add: isRunnable_def isBlocked_def) apply (rule no_fail_pre, wp) apply (clarsimp simp: pred_tcb_at'_def) done lemma corres_when_r: "\ G \ corres r P P' f g; \ G \ corres r Q Q' f (return ()) \ \ corres r (P and Q) (\s. (G \ P' s) \ (\G \ Q' s)) f (when G g)" apply (cases G, simp_all add: when_def) apply (rule corres_guard_imp, simp+)+ done lemma cur_thread_update_corres: "corres dc \ \ (modify (cur_thread_update (\_. t))) (setCurThread t)" apply (unfold setCurThread_def) apply (rule corres_modify) apply (simp add: state_relation_def swp_def) done lemma arch_switch_thread_tcb_at' [wp]: "\tcb_at' t\ Arch.switchToThread t \\_. tcb_at' t\" by (unfold ARM_H.switchToThread_def, wp typ_at_lift_tcb') crunch typ_at'[wp]: "switchToThread" "\s. P (typ_at' T p s)" (ignore: clearExMonitor) lemma Arch_switchToThread_pred_tcb'[wp]: "\\s. P (pred_tcb_at' proj P' t' s)\ Arch.switchToThread t \\rv s. P (pred_tcb_at' proj P' t' s)\" proof - have pos: "\P t t'. \pred_tcb_at' proj P t'\ Arch.switchToThread t \\rv. pred_tcb_at' proj P t'\" apply (simp add: pred_tcb_at'_def ARM_H.switchToThread_def) apply (rule hoare_seq_ext)+ apply (rule doMachineOp_obj_at) apply (rule setVMRoot_obj_at) done show ?thesis apply (rule P_bool_lift [OF pos]) by (rule lift_neg_pred_tcb_at' [OF ArchThreadDecls_H_ARM_H_switchToThread_typ_at' pos]) qed crunch ksQ[wp]: storeWordUser "\s. P (ksReadyQueues s p)" crunch ksQ[wp]: setVMRoot "\s. P (ksReadyQueues s)" (wp: crunch_wps simp: crunch_simps) crunch ksIdleThread[wp]: storeWordUser "\s. P (ksIdleThread s)" crunch ksIdleThread[wp]: asUser "\s. P (ksIdleThread s)" (wp: crunch_wps simp: crunch_simps) crunch ksQ[wp]: asUser "\s. P (ksReadyQueues s p)" (wp: crunch_wps simp: crunch_simps) lemma arch_switch_thread_ksQ[wp]: "\\s. P (ksReadyQueues s p)\ Arch.switchToThread t \\_ s. P (ksReadyQueues s p)\" apply (simp add: ARM_H.switchToThread_def) apply (wp) done crunch valid_queues[wp]: "Arch.switchToThread" "Invariants_H.valid_queues" (wp: crunch_wps simp: crunch_simps ignore: clearExMonitor) lemma switch_thread_corres: "corres dc (valid_arch_state and valid_objs and valid_asid_map and valid_arch_objs and pspace_aligned and pspace_distinct and valid_vs_lookup and valid_global_objs and unique_table_refs o caps_of_state and st_tcb_at runnable t and valid_etcbs) (valid_arch_state' and valid_pspace' and Invariants_H.valid_queues and st_tcb_at' runnable' t and cur_tcb') (switch_to_thread t) (switchToThread t)" (is "corres _ ?PA ?PH _ _") proof - have mainpart: "corres dc (?PA) (?PH) (do y \ arch_switch_to_thread t; y \ (tcb_sched_action tcb_sched_dequeue t); modify (cur_thread_update (\_. t)) od) (do y \ Arch.switchToThread t; y \ tcbSchedDequeue t; setCurThread t od)" apply (rule corres_guard_imp) apply (rule corres_split [OF _ arch_switch_thread_corres]) apply (rule corres_split[OF cur_thread_update_corres tcbSchedDequeue_corres]) apply (wp|clarsimp simp: tcb_at_is_etcb_at st_tcb_at_tcb_at)+ done show ?thesis apply - apply (simp add: switch_to_thread_def Thread_H.switchToThread_def K_bind_def) apply (rule corres_symb_exec_l [where Q = "\ s rv. (?PA and op = rv) s", OF corres_symb_exec_l [OF mainpart]]) apply (auto intro: no_fail_pre [OF no_fail_assert] no_fail_pre [OF no_fail_get] dest: st_tcb_at_tcb_at [THEN get_tcb_at] | simp add: assert_def | wp)+ done qed lemma allActiveTCBs_corres: assumes "\threads. corres r (P threads) P' (m threads) m'" shows "corres r (\s. P {x. getActiveTCB x s \ None} s) P' (do threads \ allActiveTCBs; m threads od) m'" apply (simp add: allActiveTCBs_def) apply (simp add: bind_def get_def) apply (insert assms) apply (simp add: corres_underlying_def) apply force done lemma typ_at'_typ_at'_mask: "\s. \ typ_at' t (P s) s \ \ typ_at' t (P s && ~~mask (objBitsT t)) s" apply (rule split_state_strg [where P = "typ_at' t", THEN mp]) apply (frule typ_at_aligned') apply (clarsimp dest!: is_aligned_neg_mask_eq) done lemma arch_switch_idle_thread_corres: "corres dc \ (valid_arch_state' and pspace_aligned') arch_switch_to_idle_thread Arch.switchToIdleThread" apply (simp add: arch_switch_to_idle_thread_def ARM_H.switchToIdleThread_def) done lemma switch_idle_thread_corres: "corres dc invs invs_no_cicd' switch_to_idle_thread switchToIdleThread" apply (simp add: switch_to_idle_thread_def Thread_H.switchToIdleThread_def) apply (rule corres_guard_imp) apply (rule corres_split [OF _ git_corres]) apply (rule corres_split [OF _ arch_switch_idle_thread_corres]) apply (unfold setCurThread_def) apply (rule corres_trivial, rule corres_modify) apply (simp add: state_relation_def cdt_relation_def) apply (wp, simp+) apply (simp add: all_invs_but_ct_idle_or_in_cur_domain'_def valid_state'_def valid_pspace'_def) done lemma gq_sp: "\P\ getQueue d p \\rv. P and (\s. ksReadyQueues s (d, p) = rv)\" by (unfold getQueue_def, rule gets_sp) lemma gq_se: "\\s'. (s, s') \ state_relation \ True\ getQueue d p \\rv s'. (s, s') \ state_relation\" by (simp add: getQueue_def) lemma setQueue_no_change_ct[wp]: "\ct_in_state' st\ setQueue d p q \\rv. ct_in_state' st\" apply (simp add: setQueue_def) apply wp apply (simp add: ct_in_state'_def pred_tcb_at'_def) done lemma sch_act_wf: "sch_act_wf sa s = ((\t. sa = SwitchToThread t \ st_tcb_at' runnable' t s \ tcb_in_cur_domain' t s) \ (sa = ResumeCurrentThread \ ct_in_state' activatable' s))" by (case_tac sa, simp_all add: ) declare gq_wp[wp] declare setQueue_obj_at[wp] lemma ready_tcb': "\ t \ set (ksReadyQueues s (d, p)); invs' s \ \ obj_at' (inQ d p) t s" apply (clarsimp simp: invs'_def valid_state'_def valid_pspace'_def Invariants_H.valid_queues_def valid_queues_no_bitmap_def) apply (drule_tac x=d in spec) apply (drule_tac x=p in spec) apply (clarsimp) apply (drule(1) bspec) apply (erule obj_at'_weakenE) apply (simp) done lemma ready_tcb_valid_queues: "\ t \ set (ksReadyQueues s (d, p)); valid_queues s \ \ obj_at' (inQ d p) t s" apply (clarsimp simp: invs'_def valid_state'_def valid_pspace'_def Invariants_H.valid_queues_def valid_queues_no_bitmap_def) apply (drule_tac x=d in spec) apply (drule_tac x=p in spec) apply (clarsimp) apply (drule(1) bspec) apply (erule obj_at'_weakenE) apply (simp) done lemma queued_vs_state[wp]: "\ct_in_state' st\ threadSet (tcbQueued_update x) t \\rv. ct_in_state' st\" apply (rule hoare_vcg_precond_imp) apply (rule hoare_post_imp[where Q="\rv s. \t. t = ksCurThread s \ st_tcb_at' st t s"]) apply (simp add: ct_in_state'_def) apply (unfold pred_tcb_at'_def) apply (wp hoare_ex_wp threadSet_ct) apply (simp add: ct_in_state'_def pred_tcb_at'_def) done lemma threadSet_timeslice_invs: "\invs' and tcb_at' t\ threadSet (tcbTimeSlice_update b) t \\rv. invs'\" by (wp threadSet_invs_trivial, simp_all add: inQ_def cong: conj_cong) lemma setCurThread_invs_no_cicd': "\invs_no_cicd' and st_tcb_at' activatable' t and obj_at' (\x. \ tcbQueued x) t and tcb_in_cur_domain' t\ setCurThread t \\rv. invs'\" proof - have obj_at'_ct: "\f P addr s. obj_at' P addr (ksCurThread_update f s) = obj_at' P addr s" by (fastforce intro: obj_at'_pspaceI) have valid_pspace'_ct: "\s f. valid_pspace' (ksCurThread_update f s) = valid_pspace' s" by (rule valid_pspace'_ksPSpace, simp) have vms'_ct: "\s f. valid_machine_state' (ksCurThread_update f s) = valid_machine_state' s" by (simp add: valid_machine_state'_def) have ct_not_inQ_ct: "\s t . \ ct_not_inQ s; obj_at' (\x. \ tcbQueued x) t s\ \ ct_not_inQ (s\ ksCurThread := t \)" apply (simp add: ct_not_inQ_def o_def) done have tcb_in_cur_domain_ct: "\s f t. tcb_in_cur_domain' t (ksCurThread_update f s)= tcb_in_cur_domain' t s" by (fastforce simp: tcb_in_cur_domain'_def) show ?thesis apply (simp add: setCurThread_def) apply wp apply (clarsimp simp add: all_invs_but_ct_idle_or_in_cur_domain'_def invs'_def cur_tcb'_def valid_state'_def obj_at'_ct valid_pspace'_ct vms'_ct Invariants_H.valid_queues_def sch_act_wf ct_in_state'_def state_refs_of'_def ps_clear_def valid_irq_node'_def valid_queues'_def ct_not_inQ_ct tcb_in_cur_domain_ct ct_idle_or_in_cur_domain'_def bitmapQ_defs valid_queues_no_bitmap_def cong: option.case_cong) done qed (* Don't use this rule when considering the idle thread. The invariant ct_idle_or_in_cur_domain' says that either "tcb_in_cur_domain' t" or "t = ksIdleThread s". Use setCurThread_invs_idle_thread instead. *) lemma setCurThread_invs: "\invs' and st_tcb_at' activatable' t and obj_at' (\x. \ tcbQueued x) t and tcb_in_cur_domain' t\ setCurThread t \\rv. invs'\" by (rule hoare_pre, rule setCurThread_invs_no_cicd') (simp add: invs'_to_invs_no_cicd'_def) lemma valid_queues_not_runnable_not_queued: fixes s assumes vq: "Invariants_H.valid_queues s" and vq': "valid_queues' s" and st: "st_tcb_at' (Not \ runnable') t s" shows "obj_at' (Not \ tcbQueued) t s" proof (rule ccontr) assume "\ obj_at' (Not \ tcbQueued) t s" moreover from st have "typ_at' TCBT t s" by (rule pred_tcb_at' [THEN tcb_at_typ_at' [THEN iffD1]]) ultimately have "obj_at' tcbQueued t s" by (clarsimp simp: not_obj_at' comp_def) moreover from st [THEN pred_tcb_at', THEN tcb_at'_has_tcbPriority] obtain p where tp: "obj_at' (\tcb. tcbPriority tcb = p) t s" by clarsimp moreover from st [THEN pred_tcb_at', THEN tcb_at'_has_tcbDomain] obtain d where td: "obj_at' (\tcb. tcbDomain tcb = d) t s" by clarsimp ultimately have "t \ set (ksReadyQueues s (d, p))" using vq' unfolding valid_queues'_def apply - apply (drule_tac x=d in spec) apply (drule_tac x=p in spec) apply (drule_tac x=t in spec) apply (erule impE) apply (fastforce simp add: inQ_def obj_at'_def) apply (assumption) done with vq have "st_tcb_at' runnable' t s" unfolding Invariants_H.valid_queues_def valid_queues_no_bitmap_def apply - apply clarsimp apply (drule_tac x=d in spec) apply (drule_tac x=p in spec) apply (clarsimp simp add: st_tcb_at'_def) apply (drule(1) bspec) apply (erule obj_at'_weakenE) apply (clarsimp) done with st show False apply - apply (drule(1) pred_tcb_at_conj') apply (clarsimp) done qed (* * The idle thread is not part of any ready queues. *) lemma idle'_not_tcbQueued': assumes vq: "Invariants_H.valid_queues s" and vq': "valid_queues' s" and idle: "valid_idle' s" shows "obj_at' (Not \ tcbQueued) (ksIdleThread s) s" proof - from idle have stidle: "st_tcb_at' (Not \ runnable') (ksIdleThread s) s" by (clarsimp simp add: valid_idle'_def pred_tcb_at'_def obj_at'_def projectKOs) with vq vq' show ?thesis by (rule valid_queues_not_runnable_not_queued) qed lemma idle'_not_tcbQueued: assumes "invs' s" shows "obj_at' (Not \ tcbQueued) (ksIdleThread s) s" by (insert assms) (clarsimp simp: invs'_def valid_state'_def elim!: idle'_not_tcbQueued') lemma setCurThread_invs_no_cicd'_idle_thread: "\invs_no_cicd' and (\s. t = ksIdleThread s) \ setCurThread t \\rv. invs'\" proof - have vms'_ct: "\s f. valid_machine_state' (ksCurThread_update f s) = valid_machine_state' s" by (simp add: valid_machine_state'_def) have ct_not_inQ_ct: "\s t . \ ct_not_inQ s; obj_at' (\x. \ tcbQueued x) t s\ \ ct_not_inQ (s\ ksCurThread := t \)" apply (simp add: ct_not_inQ_def o_def) done have idle'_activatable': "\ s t. st_tcb_at' idle' t s \ st_tcb_at' activatable' t s" apply (clarsimp simp: st_tcb_at'_def o_def obj_at'_def) done show ?thesis apply (simp add: setCurThread_def) apply wp apply (clarsimp simp add: vms'_ct ct_not_inQ_ct idle'_activatable' idle'_not_tcbQueued'[simplified o_def] invs'_def cur_tcb'_def valid_state'_def valid_idle'_def sch_act_wf ct_in_state'_def state_refs_of'_def ps_clear_def valid_irq_node'_def ct_idle_or_in_cur_domain'_def tcb_in_cur_domain'_def valid_queues_def bitmapQ_defs valid_queues_no_bitmap_def valid_queues'_def all_invs_but_ct_idle_or_in_cur_domain'_def pred_tcb_at'_def cong: option.case_cong) apply (clarsimp simp: obj_at'_def projectKOs) done qed lemma setCurThread_invs_idle_thread: "\invs' and (\s. t = ksIdleThread s) \ setCurThread t \\rv. invs'\" by (rule hoare_pre, rule setCurThread_invs_no_cicd'_idle_thread) (clarsimp simp: invs'_to_invs_no_cicd'_def all_invs_but_ct_idle_or_in_cur_domain'_def) lemma clearExMonitor_invs'[wp]: "\invs'\ doMachineOp ARM.clearExMonitor \\rv. invs'\" apply (wp dmo_invs' no_irq) apply (simp add: no_irq_clearExMonitor) apply (clarsimp simp: ARM.clearExMonitor_def machine_op_lift_def in_monad select_f_def) done lemma Arch_switchToThread_invs[wp]: "\invs' and tcb_at' t\ Arch.switchToThread t \\rv. invs'\" apply (simp add: ARM_H.switchToThread_def) apply (wp; auto) done lemma Arch_switchToThread_tcb'[wp]: "\tcb_at' t\ Arch.switchToThread t \\rv. tcb_at' t\" apply (simp add: ARM_H.switchToThread_def) apply (wp doMachineOp_obj_at hoare_drop_imps)+ done crunch ksCurDomain[wp]: "Arch.switchToThread" "\s. P (ksCurDomain s)" (simp: whenE_def) lemma Arch_swichToThread_tcbDomain_triv[wp]: "\ obj_at' (\tcb. P (tcbDomain tcb)) t' \ Arch.switchToThread t \ \_. obj_at' (\tcb. P (tcbDomain tcb)) t' \" apply (clarsimp simp: ARM_H.switchToThread_def storeWordUser_def) apply (wp hoare_drop_imp | simp)+ done lemma Arch_swichToThread_tcbPriority_triv[wp]: "\ obj_at' (\tcb. P (tcbPriority tcb)) t' \ Arch.switchToThread t \ \_. obj_at' (\tcb. P (tcbPriority tcb)) t' \" apply (clarsimp simp: ARM_H.switchToThread_def storeWordUser_def) apply (wp hoare_drop_imp | simp)+ done lemma Arch_switchToThread_tcb_in_cur_domain'[wp]: "\tcb_in_cur_domain' t'\ Arch.switchToThread t \\_. tcb_in_cur_domain' t' \" apply (rule tcb_in_cur_domain'_lift) apply wp done lemma tcbSchedDequeue_not_tcbQueued: "\ tcb_at' t \ tcbSchedDequeue t \ \_. obj_at' (\x. \ tcbQueued x) t \" apply (simp add: tcbSchedDequeue_def) apply (wp|clarsimp)+ apply (rule_tac Q="\queued. obj_at' (\x. tcbQueued x = queued) t" in hoare_post_imp) apply (clarsimp simp: obj_at'_def) apply (wp threadGet_obj_at') apply (simp) done lemma asUser_obj_at[wp]: "\obj_at' (P \ tcbState) t\ asUser t' f \\rv. obj_at' (P \ tcbState) t\" apply (simp add: asUser_def threadGet_stateAssert_gets_asUser) apply (wp|wpc)+ apply (simp add: asUser_fetch_def obj_at'_def) done lemma Arch_switchToThread_obj_at[wp]: "\obj_at' (P \ tcbState) t\ Arch.switchToThread t \\rv. obj_at' (P \ tcbState) t\" apply (simp add: ARM_H.switchToThread_def ) apply (rule hoare_seq_ext)+ apply (rule doMachineOp_obj_at) apply (rule setVMRoot_obj_at) done declare doMachineOp_obj_at[wp] lemma clearExMonitor_invs_no_cicd'[wp]: "\invs_no_cicd'\ doMachineOp ARM.clearExMonitor \\rv. invs_no_cicd'\" apply (wp dmo_invs_no_cicd' no_irq) apply (simp add: no_irq_clearExMonitor) apply (clarsimp simp: ARM.clearExMonitor_def machine_op_lift_def in_monad select_f_def) done crunch valid_arch_state'[wp]: asUser "valid_arch_state'" (wp: crunch_wps simp: crunch_simps) crunch valid_irq_states'[wp]: asUser "valid_irq_states'" (wp: crunch_wps simp: crunch_simps) crunch valid_machine_state'[wp]: asUser "valid_machine_state'" (wp: crunch_wps simp: crunch_simps) crunch valid_queues'[wp]: asUser "valid_queues'" (wp: crunch_wps simp: crunch_simps) lemma asUser_valid_irq_node'[wp]: "\\s. valid_irq_node' (irq_node' s) s\ asUser t (setRegister f r) \\_ s. valid_irq_node' (irq_node' s) s\" apply (rule_tac valid_irq_node_lift) apply (simp add: asUser_def) apply (rule hoare_seq_ext) defer apply (wp threadGet_irq_node'|wpc)+ apply clarsimp done crunch irq_masked'_helper: asUser "\s. P (intStateIRQTable (ksInterruptState s))" (wp: crunch_wps simp: crunch_simps) lemma asUser_irq_masked'[wp]: "\irqs_masked'\ asUser t (setRegister f r) \\_ . irqs_masked'\" apply (rule irqs_masked_lift) apply (rule asUser_irq_masked'_helper) done lemma asUser_ct_not_inQ[wp]: "\ct_not_inQ\ asUser t (setRegister f r) \\_ . ct_not_inQ\" apply (clarsimp simp: submonad_asUser.fn_is_sm submonad_fn_def) apply (rule hoare_seq_ext)+ prefer 4 apply (rule stateAssert_sp) prefer 3 apply (rule gets_inv) defer apply (rule select_f_inv) apply (case_tac x; simp) apply (clarsimp simp: projectKOs asUser_replace_def obj_at'_def fun_upd_def split: option.split kernel_object.split) apply (clarsimp simp: ct_not_inQ_def obj_at'_def projectKOs objBitsKO_def ps_clear_def dom_def) apply (rule conjI; clarsimp; blast) done crunch valid_pde_mappings'[wp]: asUser "valid_pde_mappings'" (wp: crunch_wps simp: crunch_simps) crunch pspace_domain_valid[wp]: asUser "pspace_domain_valid" (wp: crunch_wps simp: crunch_simps) crunch valid_dom_schedule'[wp]: asUser "valid_dom_schedule'" (wp: crunch_wps simp: crunch_simps) crunch gsUntypedZeroRanges[wp]: asUser "\s. P (gsUntypedZeroRanges s)" (wp: crunch_wps simp: unless_def) crunch ctes_of[wp]: asUser "\s. P (ctes_of s)" (wp: crunch_wps simp: unless_def) lemmas asUser_cteCaps_of[wp] = cteCaps_of_ctes_of_lift[OF asUser_ctes_of] lemma asUser_utr[wp]: "\untyped_ranges_zero'\ asUser f t \\_. untyped_ranges_zero'\" apply (simp add: cteCaps_of_def) apply (rule hoare_pre, wp untyped_ranges_zero_lift) apply (simp add: o_def) done lemma asUser_invs_no_cicd'[wp]: "\invs_no_cicd'\ asUser t (setRegister f r) \\rv. invs_no_cicd'\" apply (simp add: invs_no_cicd'_def) apply (rule hoare_pre) apply (wpc | wp asUser_global_refs' asUser_irq_handlers')+ apply (rule hoare_post_imp[rotated]) apply (rule hoare_vcg_conj_lift) apply (rule asUser_valid_dom_schedule') apply (rule asUser_utr) apply clarsimp apply clarsimp done lemma Arch_switchToThread_invs_no_cicd': "\invs_no_cicd'\ Arch.switchToThread t \\rv. invs_no_cicd'\" apply (simp add: ARM_H.switchToThread_def) by (wp|rule setVMRoot_invs_no_cicd')+ lemma tcbSchedDequeue_invs_no_cicd'[wp]: "\invs_no_cicd' and tcb_at' t\ tcbSchedDequeue t \\_. invs_no_cicd'\" unfolding all_invs_but_ct_idle_or_in_cur_domain'_def valid_state'_def apply (rule hoare_pre) apply (wp tcbSchedDequeue_ct_not_inQ sch_act_wf_lift valid_irq_node_lift irqs_masked_lift valid_irq_handlers_lift' cur_tcb_lift ct_idle_or_in_cur_domain'_lift2 tcbSchedDequeue_valid_queues_weak untyped_ranges_zero_lift | simp add: cteCaps_of_def o_def)+ apply clarsimp apply (fastforce simp: valid_pspace'_def valid_queues_def elim: valid_objs'_maxDomain valid_objs'_maxPriority intro: obj_at'_conjI) done lemma switchToThread_invs_no_cicd': "\invs_no_cicd' and st_tcb_at' runnable' t and tcb_in_cur_domain' t \ ThreadDecls_H.switchToThread t \\rv. invs' \" apply (simp add: Thread_H.switchToThread_def ) apply (wp setCurThread_invs_no_cicd' Arch_switchToThread_invs_no_cicd' Arch_switchToThread_pred_tcb' tcbSchedDequeue_not_tcbQueued) apply (clarsimp elim!: pred_tcb'_weakenE)+ done lemma switchToThread_invs[wp]: "\invs' and st_tcb_at' runnable' t and tcb_in_cur_domain' t \ switchToThread t \\rv. invs' \" apply (simp add: Thread_H.switchToThread_def ) apply (wp threadSet_timeslice_invs setCurThread_invs Arch_switchToThread_invs dmo_invs' doMachineOp_obj_at tcbSchedDequeue_not_tcbQueued) by (clarsimp elim!: pred_tcb'_weakenE)+ lemma setCurThread_ct_in_state: "\obj_at' (P \ tcbState) t\ setCurThread t \\rv. ct_in_state' P\" proof - have obj_at'_ct: "\P addr f s. obj_at' P addr (ksCurThread_update f s) = obj_at' P addr s" by (fastforce intro: obj_at'_pspaceI) show ?thesis apply (simp add: setCurThread_def) apply wp apply (simp add: ct_in_state'_def pred_tcb_at'_def obj_at'_ct o_def) done qed lemma switchToThread_ct_in_state[wp]: "\obj_at' (P \ tcbState) t\ switchToThread t \\rv. ct_in_state' P\" proof - have P: "\f x. tcbState (tcbTimeSlice_update f x) = tcbState x" by (case_tac x, simp) show ?thesis apply (simp add: Thread_H.switchToThread_def tcbSchedEnqueue_def unless_def) apply (wp setCurThread_ct_in_state Arch_switchToThread_obj_at | simp add: P o_def cong: if_cong)+ done qed lemma setCurThread_obj_at[wp]: "\obj_at' P addr\ setCurThread t \\rv. obj_at' P addr\" apply (simp add: setCurThread_def) apply wp apply (fastforce intro: obj_at'_pspaceI) done lemma switchToThread_tcb'[wp]: "\tcb_at' t\ switchToThread t \\rv. tcb_at' t\" apply (simp add: Thread_H.switchToThread_def when_def) apply wp done crunch cap_to'[wp]: setQueue "ex_nonz_cap_to' p" lemma dmo_cap_to'[wp]: "\ex_nonz_cap_to' p\ doMachineOp m \\rv. ex_nonz_cap_to' p\" by (wp ex_nonz_cap_to_pres') lemma sct_cap_to'[wp]: "\ex_nonz_cap_to' p\ setCurThread t \\rv. ex_nonz_cap_to' p\" apply (simp add: setCurThread_def) apply (wp ex_nonz_cap_to_pres') apply (clarsimp elim!: cte_wp_at'_pspaceI) done crunch cap_to'[wp]: "Arch.switchToThread" "ex_nonz_cap_to' p" (simp: crunch_simps ignore: ARM.clearExMonitor) crunch cap_to'[wp]: switchToThread "ex_nonz_cap_to' p" (simp: crunch_simps ignore: ARM.clearExMonitor) lemma no_longer_inQ[simp]: "\ inQ d p (tcbQueued_update (\x. False) tcb)" by (simp add: inQ_def) lemma rq_distinct: "invs' s \ distinct (ksReadyQueues s (d, p))" by (clarsimp simp: invs'_def valid_state'_def Invariants_H.valid_queues_def valid_queues_no_bitmap_def) lemma iflive_inQ_nonz_cap_strg: "if_live_then_nonz_cap' s \ obj_at' (inQ d prio) t s \ ex_nonz_cap_to' t s" by (clarsimp simp: obj_at'_real_def projectKOs inQ_def elim!: if_live_then_nonz_capE' ko_wp_at'_weakenE) lemmas iflive_inQ_nonz_cap[elim] = mp [OF iflive_inQ_nonz_cap_strg, OF conjI[rotated]] crunch ksRQ[wp]: threadSet "\s. P (ksReadyQueues s)" (ignore: setObject getObject wp: updateObject_default_inv) declare Cons_eq_tails[simp] lemma invs_ksReadyQueues_update_triv: assumes a: "\d p. set (f (ksReadyQueues s) (d,p)) = set (ksReadyQueues s (d,p))" assumes b: "\d p. distinct (f (ksReadyQueues s) (d,p)) = distinct (ksReadyQueues s (d,p))" assumes c: "\d p. (d > maxDomain \ p > maxPriority \ (f (ksReadyQueues s) (d,p) = [])) = (d > maxDomain \ p > maxPriority \ (ksReadyQueues s (d,p) = []))" shows "invs' (ksReadyQueues_update f s) = invs' s" proof - have orphans: "bitmapQ_no_L1_orphans (ksReadyQueues_update f s) = bitmapQ_no_L1_orphans s" "bitmapQ_no_L2_orphans (ksReadyQueues_update f s) = bitmapQ_no_L2_orphans s" using a b c unfolding bitmapQ_no_L2_orphans_def bitmapQ_no_L1_orphans_def by auto have bitmapQ: "valid_bitmapQ (ksReadyQueues_update f s) = valid_bitmapQ s" using a b c unfolding valid_bitmapQ_def by (simp add: orphans bitmapQ_def set_empty[symmetric] del: set_empty) show ?thesis unfolding invs'_def valid_state'_def ct_not_inQ_def valid_irq_node'_def cur_tcb'_def Invariants_H.valid_queues_def valid_queues'_def valid_queues_no_bitmap_def ct_idle_or_in_cur_domain'_def tcb_in_cur_domain'_def using a b c by (clarsimp simp: cur_tcb'_def orphans bitmapQ) qed lemma tcbSchedDequeue_ksReadyQueues_eq: "\\s. obj_at' (inQ d p) t s \ filter (op \ t) (ksReadyQueues s (d, p)) = ts\ tcbSchedDequeue t \\rv s. ksReadyQueues s (d, p) = ts\" apply (simp add: tcbSchedDequeue_def threadGet_def liftM_def) apply (wp getObject_tcb_wp|clarsimp)+ apply (clarsimp simp: obj_at'_def projectKOs inQ_def split del: if_split) apply (simp add: eq_commute) apply (rule_tac x=obj in exI, clarsimp)+ done lemma hd_ksReadyQueues_runnable: "\invs' s; ksReadyQueues s (d, p) = a # as\ \ st_tcb_at' runnable' a s" apply(simp add: invs'_def valid_state'_def, rule valid_queues_running, clarsimp) apply(rule suffix_Cons_mem) apply(simp add: suffix_def) apply(rule exI) apply(rule eq_Nil_appendI) by auto lemma chooseThread_invs_fragment: "\invs' and (\s. ksCurDomain s = d)\ do x \ getQueue d r; (case x of [] \ return False | thread # x \ do y \ ThreadDecls_H.switchToThread thread; return True od) od \\_. invs'\" apply (rule seq_ext) apply (rule gq_sp) apply (case_tac x) apply (simp) apply (wp) apply (clarsimp) apply (simp) apply (wp) apply (clarsimp simp: invs'_def valid_state'_def , rule conjI) apply (rule valid_queues_running) apply (rule suffix_Cons_mem) apply (simp add: suffix_def) apply (rule exI) apply (rule eq_Nil_appendI) apply simp+ apply (clarsimp simp add: tcb_in_cur_domain'_def valid_queues_def valid_queues_no_bitmap_def) apply (drule_tac x="ksCurDomain s" in spec) apply (drule_tac x=r in spec) apply (fastforce simp: obj_at'_def inQ_def) done crunch ksCurDomain[wp]: "ThreadDecls_H.switchToThread" "\s. P (ksCurDomain s)" lemma chooseThread_ksCurDomain_fragment: "\\s. P (ksCurDomain s)\ do x \ getQueue d r; (case x of [] \ return False | thread # x \ do y \ ThreadDecls_H.switchToThread thread; return True od) od \\_ s. P (ksCurDomain s)\" apply (wp | wpc | clarsimp)+ done (* FIXME move *) lemma obj_tcb_at': "obj_at' (\tcb::tcb. P tcb) t s \ tcb_at' t s" by (clarsimp simp: obj_at'_def) lemma valid_queues_queued_runnable: fixes s assumes vq: "Invariants_H.valid_queues s" and vq': "valid_queues' s" and oa: "obj_at' tcbQueued t s" shows "st_tcb_at' runnable' t s" proof - from oa [THEN obj_tcb_at', THEN tcb_at'_has_tcbPriority] obtain p where tp: "obj_at' (\tcb. tcbPriority tcb = p) t s" by clarsimp from oa [THEN obj_tcb_at', THEN tcb_at'_has_tcbDomain] obtain d where td: "obj_at' (\tcb. tcbDomain tcb = d) t s" by clarsimp with oa tp have "obj_at' (inQ d p) t s" by (fastforce simp add: inQ_def obj_at'_def) with vq' have "t \ set (ksReadyQueues s (d, p))" unfolding valid_queues'_def by (fastforce) with vq show ?thesis unfolding Invariants_H.valid_queues_def valid_queues_no_bitmap_def apply - apply clarify apply (drule_tac x=d in spec) apply (drule_tac x=p in spec) apply (clarsimp simp: pred_tcb_at'_def) apply (drule(1) bspec) apply (erule obj_at'_weakenE) apply (simp) done qed lemma tcb_at_typ_at': "tcb_at' p s = typ_at' TCBT p s" unfolding typ_at'_def apply rule apply (clarsimp simp add: obj_at'_def ko_wp_at'_def projectKOs) apply (clarsimp simp add: obj_at'_def ko_wp_at'_def projectKOs) apply (case_tac ko, simp_all) done lemma invs'_not_runnable_not_queued: fixes s assumes inv: "invs' s" and st: "st_tcb_at' (Not \ runnable') t s" shows "obj_at' (Not \ tcbQueued) t s" apply (insert assms) apply (rule valid_queues_not_runnable_not_queued) apply (clarsimp simp add: invs'_def valid_state'_def)+ done lemma valid_queues_not_tcbQueued_not_ksQ: fixes s assumes vq: "Invariants_H.valid_queues s" and notq: "obj_at' (Not \ tcbQueued) t s" shows "\d p. t \ set (ksReadyQueues s (d, p))" proof (rule ccontr, simp , erule exE, erule exE) fix d p assume "t \ set (ksReadyQueues s (d, p))" with vq have "obj_at' (inQ d p) t s" unfolding Invariants_H.valid_queues_def valid_queues_no_bitmap_def apply clarify apply (drule_tac x=d in spec) apply (drule_tac x=p in spec) apply (clarsimp) apply (drule(1) bspec) apply (erule obj_at'_weakenE) apply (simp) done hence "obj_at' tcbQueued t s" apply (rule obj_at'_weakenE) apply (simp only: inQ_def) done with notq show "False" by (clarsimp simp: obj_at'_def) qed lemma not_tcbQueued_not_ksQ: fixes s assumes "invs' s" and "obj_at' (Not \ tcbQueued) t s" shows "\d p. t \ set (ksReadyQueues s (d, p))" apply (insert assms) apply (clarsimp simp add: invs'_def valid_state'_def) apply (drule(1) valid_queues_not_tcbQueued_not_ksQ) apply (clarsimp) done lemma ct_not_ksQ: "\ invs' s; ksSchedulerAction s = ResumeCurrentThread \ \ \p. ksCurThread s \ set (ksReadyQueues s p)" apply (clarsimp simp: invs'_def valid_state'_def ct_not_inQ_def) apply (frule(1) valid_queues_not_tcbQueued_not_ksQ) apply (fastforce) done crunch nosch[wp]: getCurThread "\s. P (ksSchedulerAction s)" lemma setThreadState_rct: "\\s. (runnable' st \ ksCurThread s \ t) \ ksSchedulerAction s = ResumeCurrentThread\ setThreadState st t \\_ s. ksSchedulerAction s = ResumeCurrentThread\" apply (simp add: setThreadState_def) apply (rule hoare_pre_disj') apply (rule hoare_seq_ext [OF _ hoare_vcg_conj_lift [OF threadSet_tcbState_st_tcb_at' [where P=runnable'] threadSet_nosch]]) apply (rule hoare_seq_ext [OF _ hoare_vcg_conj_lift [OF isRunnable_const isRunnable_inv]]) apply (clarsimp simp: when_def) apply (case_tac x) apply (clarsimp, wp)[1] apply (clarsimp) apply (rule hoare_seq_ext [OF _ hoare_vcg_conj_lift [OF threadSet_ct threadSet_nosch]]) apply (rule hoare_seq_ext [OF _ isRunnable_inv]) apply (rule hoare_seq_ext [OF _ hoare_vcg_conj_lift [OF gct_wp getCurThread_nosch]]) apply (rename_tac ct) apply (case_tac "ct\t") apply (clarsimp simp: when_def) apply (wp)[1] apply (clarsimp) done lemma switchToIdleThread_invs'[wp]: "\invs'\ switchToIdleThread \\rv. invs'\" apply (clarsimp simp: Thread_H.switchToIdleThread_def ARM_H.switchToIdleThread_def) apply (wp_trace setCurThread_invs_idle_thread) apply clarsimp done lemma bind_dummy_ret_val: "do y \ a; b od = do a; b od" by simp lemma valid_bitmapQ_bitmapQ_simp: "\ valid_bitmapQ s \ \ bitmapQ d p s = (ksReadyQueues s (d, p) \ [])" unfolding valid_bitmapQ_def by simp lemma invs_bitmapQ_tcb_at_cur_domain'_hd: "\ invs' s; bitmapQ (ksCurDomain s) p s; st_tcb_at' runnable' (hd (ksReadyQueues s (ksCurDomain s, p))) s\ \ tcb_in_cur_domain' (hd (ksReadyQueues s (ksCurDomain s, p))) s" apply (rule_tac xs="tl (ksReadyQueues s (ksCurDomain s, p))" in invs_queues_tcb_in_cur_domain'[rotated], assumption, rule refl) apply (drule invs_queues) apply (clarsimp simp: valid_queues_def) apply (simp add: valid_bitmapQ_bitmapQ_simp) done lemma shiftr_le_0: "unat (w::'a::len word) < 2 ^ n \ w >> n = (0::'a::len word)" by (rule word_unat.Rep_eqD) (simp add: shiftr_div_2n') lemma prioToL1Index_l1IndexToPrio_or_id: "\ unat (w'::priority) < 2 ^ wordRadix ; w < size w' \ \ prioToL1Index ((l1IndexToPrio w) || w') = w" unfolding l1IndexToPrio_def prioToL1Index_def apply (simp add: shiftr_over_or_dist shiftr_le_0 wordRadix_def') apply (subst shiftl_shiftr_id, simp, simp add: word_size) apply (rule word_of_nat_less) apply simp apply (subst unat_of_nat_eq, simp_all add: word_size) done lemma word_exists_nth: "(w::'a::len word) \ 0 \ \i. w !! i" using word_log2_nth_same by blast lemma bitmapQ_no_L1_orphansD: "\ bitmapQ_no_L1_orphans s ; ksReadyQueuesL1Bitmap s d !! i \ \ ksReadyQueuesL2Bitmap s (d, i) \ 0 \ i < numPriorities div wordBits" unfolding bitmapQ_no_L1_orphans_def by simp lemma l1IndexToPrio_wordRadix_mask[simp]: "l1IndexToPrio i && mask wordRadix = 0" unfolding l1IndexToPrio_def by (simp add: wordRadix_def') definition (* convenience for use with proofs, used for chooseThread proofs *) "lookupBitmapPriority d s \ l1IndexToPrio (word_log2 (ksReadyQueuesL1Bitmap s d)) || of_nat (word_log2 (ksReadyQueuesL2Bitmap s (d, word_log2 (ksReadyQueuesL1Bitmap s d))))" lemma bitmapQ_lookupBitmapPriority_simp: (* neater unfold, actual unfold is really ugly *) "\ ksReadyQueuesL1Bitmap s d \ 0 ; valid_bitmapQ s ; bitmapQ_no_L1_orphans s \ \ bitmapQ d (lookupBitmapPriority d s) s = (ksReadyQueuesL1Bitmap s d !! word_log2 (ksReadyQueuesL1Bitmap s d) \ ksReadyQueuesL2Bitmap s (d, word_log2 (ksReadyQueuesL1Bitmap s d)) !! word_log2 (ksReadyQueuesL2Bitmap s (d, word_log2 (ksReadyQueuesL1Bitmap s d))))" unfolding bitmapQ_def lookupBitmapPriority_def apply (drule word_log2_nth_same, clarsimp) apply (drule (1) bitmapQ_no_L1_orphansD, clarsimp) apply (drule word_log2_nth_same, clarsimp) apply (frule test_bit_size[where n="word_log2 (ksReadyQueuesL2Bitmap _ _)"]) apply (clarsimp simp: numPriorities_def wordBits_def word_size) apply (subst prioToL1Index_l1IndexToPrio_or_id) apply (simp add: wordRadix_def' unat_of_nat word_size) apply (simp add: wordRadix_def' unat_of_nat word_size) apply (subst prioToL1Index_l1IndexToPrio_or_id) apply (simp add: wordRadix_def' unat_of_nat word_size) apply (simp add: wordRadix_def' unat_of_nat word_size) apply (simp add: word_ao_dist) apply (subst less_mask_eq) apply (fastforce intro: word_of_nat_less simp: wordRadix_def' unat_of_nat word_size)+ done lemma bitmapQ_from_bitmap_lookup: "\ ksReadyQueuesL1Bitmap s d \ 0 ; valid_bitmapQ s ; bitmapQ_no_L1_orphans s \ \ bitmapQ d (lookupBitmapPriority d s) s" apply (simp add: bitmapQ_lookupBitmapPriority_simp) apply (drule word_log2_nth_same) apply (drule (1) bitmapQ_no_L1_orphansD) apply (fastforce dest!: word_log2_nth_same simp: word_ao_dist lookupBitmapPriority_def word_size numPriorities_def wordBits_def) done lemma lookupBitmapPriority_obj_at': "\ksReadyQueuesL1Bitmap s (ksCurDomain s) \ 0; valid_queues_no_bitmap s; valid_bitmapQ s; bitmapQ_no_L1_orphans s\ \ obj_at' (inQ (ksCurDomain s) (lookupBitmapPriority (ksCurDomain s) s) and runnable' \ tcbState) (hd (ksReadyQueues s (ksCurDomain s, lookupBitmapPriority (ksCurDomain s) s))) s" apply (drule (2) bitmapQ_from_bitmap_lookup) apply (simp add: valid_bitmapQ_bitmapQ_simp) apply (case_tac "ksReadyQueues s (ksCurDomain s, lookupBitmapPriority (ksCurDomain s) s)") (* FIXME use a split rule?*) apply simp apply (clarsimp, rename_tac t ts) apply (drule cons_set_intro) apply (drule (2) valid_queues_no_bitmap_objD) done lemma bitmapL1_zero_ksReadyQueues: "\ valid_bitmapQ s ; bitmapQ_no_L1_orphans s \ \ (ksReadyQueuesL1Bitmap s d = 0) = (\p. ksReadyQueues s (d,p) = [])" apply (cases "ksReadyQueuesL1Bitmap s d = 0") apply (force simp add: bitmapQ_def valid_bitmapQ_def) apply (fastforce dest: bitmapQ_from_bitmap_lookup simp: valid_bitmapQ_bitmapQ_simp) done lemma prioToL1Index_le_mask: "\ prioToL1Index p = prioToL1Index p' ; p && mask wordRadix \ p' && mask wordRadix \ \ p \ p'" unfolding prioToL1Index_def apply (simp add: wordRadix_def word_le_nat_alt[symmetric]) apply (drule shiftr_eq_neg_mask_eq) apply (metis add.commute word_and_le2 word_plus_and_or_coroll2 word_plus_mono_left) done lemma prioToL1Index_le_index: "\ prioToL1Index p \ prioToL1Index p' ; prioToL1Index p \ prioToL1Index p' \ \ p \ p'" unfolding prioToL1Index_def apply (simp add: wordRadix_def word_le_nat_alt[symmetric]) apply (erule (1) le_shiftr') done lemma bitmapL1_highest_lookup: "\ valid_bitmapQ s ; bitmapQ_no_L1_orphans s ; bitmapQ d p' s \ \ p' \ lookupBitmapPriority d s" apply (subgoal_tac "ksReadyQueuesL1Bitmap s d \ 0") prefer 2 apply (clarsimp simp add: bitmapQ_def) apply (case_tac "prioToL1Index (lookupBitmapPriority d s) = prioToL1Index p'") apply (rule prioToL1Index_le_mask, simp) apply (frule (2) bitmapQ_from_bitmap_lookup) apply (clarsimp simp: bitmapQ_lookupBitmapPriority_simp) apply (clarsimp simp: bitmapQ_def lookupBitmapPriority_def) apply (subst mask_or_not_mask[where n=wordRadix and x=p', symmetric]) apply (subst word_bw_comms(2)) (* || commute *) apply (simp add: word_ao_dist mask_AND_NOT_mask mask_twice) apply (subst less_mask_eq[where x="of_nat _"]) apply (subst word_less_nat_alt) apply (subst unat_of_nat_eq) apply (rule order_less_le_trans[OF word_log2_max]) apply (simp add: word_size) apply (rule order_less_le_trans[OF word_log2_max]) apply (simp add: word_size wordRadix_def') apply (subst word_le_nat_alt) apply (subst unat_of_nat_eq) apply (rule order_less_le_trans[OF word_log2_max], simp add: word_size) apply (rule word_log2_highest) apply (subst (asm) prioToL1Index_l1IndexToPrio_or_id) apply (subst unat_of_nat_eq) apply (rule order_less_le_trans[OF word_log2_max], simp add: word_size) apply (rule order_less_le_trans[OF word_log2_max], simp add: word_size wordRadix_def') apply (simp add: word_size) apply (drule (1) bitmapQ_no_L1_orphansD[where d=d and i="word_log2 _"]) apply (simp add: numPriorities_def wordBits_def word_size) apply simp apply (rule prioToL1Index_le_index[rotated], simp) apply (frule (2) bitmapQ_from_bitmap_lookup) apply (clarsimp simp: bitmapQ_lookupBitmapPriority_simp) apply (clarsimp simp: bitmapQ_def lookupBitmapPriority_def) apply (subst prioToL1Index_l1IndexToPrio_or_id) apply (subst unat_of_nat_eq) apply (rule order_less_le_trans[OF word_log2_max], simp add: word_size) apply (rule order_less_le_trans[OF word_log2_max], simp add: word_size wordRadix_def') apply (fastforce dest: bitmapQ_no_L1_orphansD simp: wordBits_def numPriorities_def word_size) apply (erule word_log2_highest) done lemma bitmapQ_ksReadyQueuesI: "\ bitmapQ d p s ; valid_bitmapQ s \ \ ksReadyQueues s (d, p) \ []" unfolding valid_bitmapQ_def by simp lemma getReadyQueuesL2Bitmap_inv[wp]: "\ P \ getReadyQueuesL2Bitmap d i \\_. P\" unfolding getReadyQueuesL2Bitmap_def by wp lemma switchToThread_lookupBitmapPriority_wp: "\\s. invs_no_cicd' s \ bitmapQ (ksCurDomain s) (lookupBitmapPriority (ksCurDomain s) s) s \ t = hd (ksReadyQueues s (ksCurDomain s, lookupBitmapPriority (ksCurDomain s) s)) \ ThreadDecls_H.switchToThread t \\rv. invs'\" proof - have switchToThread_pre: "\s p t.\ valid_queues s ; bitmapQ (ksCurDomain s) p s ; t = hd (ksReadyQueues s (ksCurDomain s,p)) \ \ st_tcb_at' runnable' t s \ tcb_in_cur_domain' t s" unfolding valid_queues_def apply (clarsimp dest!: bitmapQ_ksReadyQueuesI) apply (case_tac "ksReadyQueues s (ksCurDomain s, p)", simp) apply (rename_tac t ts) apply (drule_tac t=t and p=p and d="ksCurDomain s" in valid_queues_no_bitmap_objD) apply simp apply (fastforce elim: obj_at'_weaken simp: inQ_def tcb_in_cur_domain'_def st_tcb_at'_def) done thus ?thesis by (wp switchToThread_invs_no_cicd') (fastforce dest: invs_no_cicd'_queues) qed lemma switchToIdleThread_invs_no_cicd': "\invs_no_cicd'\ switchToIdleThread \\rv. invs'\" apply (clarsimp simp: Thread_H.switchToIdleThread_def ARM_H.switchToIdleThread_def) apply (wp setCurThread_invs_no_cicd'_idle_thread storeWordUser_invs_no_cicd') apply (clarsimp simp: all_invs_but_ct_idle_or_in_cur_domain'_def valid_idle'_def) done crunch obj_at'[wp]: "Arch.switchToIdleThread" "\s. obj_at' P t s" lemma switchToIdleThread_activatable[wp]: "\invs'\ switchToIdleThread \\rv. ct_in_state' activatable'\" apply (simp add: Thread_H.switchToIdleThread_def ARM_H.switchToIdleThread_def) apply (wp setCurThread_ct_in_state) apply (clarsimp simp: invs'_def valid_state'_def valid_idle'_def pred_tcb_at'_def obj_at'_def) done declare static_imp_conj_wp[wp_split del] lemma valid_queues_obj_at'_imp: "\ t \ set (ksReadyQueues s (d, p)); Invariants_H.valid_queues s; \obj. runnable' obj \ P obj \ \ obj_at' (P \ tcbState) t s" apply (clarsimp simp: Invariants_H.valid_queues_def valid_queues_no_bitmap_def o_def) apply(elim allE conjE) apply(rule obj_at'_weaken) apply(erule(1) ballE) apply(erule(1) notE) apply(clarsimp simp: obj_at'_def inQ_def) done lemma setCurThread_const: "\\_. P t \ setCurThread t \\_ s. P (ksCurThread s) \" by (simp add: setCurThread_def | wp)+ crunch it[wp]: switchToIdleThread "\s. P (ksIdleThread s)" crunch it[wp]: switchToThread "\s. P (ksIdleThread s)" (ignore: clearExMonitor) lemma switchToIdleThread_curr_is_idle: "\\\ switchToIdleThread \\rv s. ksCurThread s = ksIdleThread s\" apply (rule hoare_weaken_pre) apply (wps switchToIdleThread_it) apply (simp add: Thread_H.switchToIdleThread_def) apply (wp setCurThread_const) apply (simp) done lemma chooseThread_it[wp]: "\\s. P (ksIdleThread s)\ chooseThread \\_ s. P (ksIdleThread s)\" by (simp add: chooseThread_def curDomain_def numDomains_def bitmap_fun_defs) (wp, clarsimp) lemma valid_queues_ct_update[simp]: "Invariants_H.valid_queues (s\ksCurThread := t\) = Invariants_H.valid_queues s" by (simp add: Invariants_H.valid_queues_def bitmapQ_defs valid_queues_no_bitmap_def) lemma valid_queues'_ct_update[simp]: "valid_queues' (s\ksCurThread := t\) = valid_queues' s" by (simp add: valid_queues'_def) lemma valid_machine_state'_ct_update[simp]: "valid_machine_state' (s\ksCurThread := t\) = valid_machine_state' s" by (simp add: valid_machine_state'_def) lemma valid_irq_node'_ct_update[simp]: "valid_irq_node' w (s\ksCurThread := t\) = valid_irq_node' w s" by (simp add: valid_irq_node'_def) lemma switchToThread_ct_not_queued: "\invs' and tcb_at' t\ switchToThread t \\rv s. obj_at' (Not \ tcbQueued) (ksCurThread s) s\" (is "\_\ _ \\_. ?POST\") apply (simp add: Thread_H.switchToThread_def) apply (wp) apply (simp add: switchToThread_def setCurThread_def) apply (wp tcbSchedDequeue_not_tcbQueued | simp )+ done lemma valid_irq_node'_ksCurThread: "valid_irq_node' w (s\ksCurThread := t\) = valid_irq_node' w s" unfolding valid_irq_node'_def by simp lemma threadGet_inv [wp]: "\P\ threadGet f t \\rv. P\" apply (simp add: threadGet_def) apply (wp | simp)+ done lemma getThreadState_ct_in_state: "\\s. t = ksCurThread s \ tcb_at' t s\ getThreadState t \\rv s. P rv \ ct_in_state' P s\" apply (rule hoare_post_imp [where Q="\rv s. t = ksCurThread s \ ((\ P rv) \ st_tcb_at' P t s)"]) apply (clarsimp simp add: ct_in_state'_def) apply (rule hoare_vcg_precond_imp) apply (wp hoare_vcg_disj_lift) apply (clarsimp simp add: pred_tcb_at'_def obj_at'_def) done lemma gsa_wf_invs: "\invs'\ getSchedulerAction \\sa. invs' and sch_act_wf sa\" by wp (clarsimp simp add: invs'_def valid_state'_def) (* Helper for schedule_corres *) lemma gsa_wf_invs_and_P: "\invs' and P\ getSchedulerAction \\sa. invs' and P and (\s. ksSchedulerAction s = sa) and sch_act_wf sa\" apply ( wp | simp ) by auto lemma gets_the_simp: "f s = Some y \ gets_the f s = ({(y, s)}, False)" by (simp add: gets_the_def gets_def assert_opt_def get_def bind_def return_def fail_def split: option.splits) lemma corres_split_sched_act: "\sched_act_relation act act'; corres r P P' f1 g1; \t. corres r (Q t) (Q' t) (f2 t) (g2 t); corres r R R' f3 g3\ \ corres r (case act of resume_cur_thread \ P | switch_thread t \ Q t | choose_new_thread \ R) (case act' of ResumeCurrentThread \ P' | SwitchToThread t \ Q' t | ChooseThread \ R') (case act of resume_cur_thread \ f1 | switch_thread t \ f2 t | choose_new_thread \ f3) (case act' of ResumeCurrentThread \ g1 | ChooseNewThread \ g3 | SwitchToThread t \ g2 t)" apply (cases act) apply (rule corres_guard_imp, force+)+ done lemma get_sa_corres': "corres sched_act_relation P P' (gets scheduler_action) getSchedulerAction" by (clarsimp simp: getSchedulerAction_def state_relation_def) lemma corres_assert_ret: "corres dc (\s. P) \ (assert P) (return ())" apply (rule corres_no_failI) apply simp apply (simp add: assert_def return_def fail_def) done lemma corres_assert_assume_l: "corres dc P Q (f ()) g \ corres dc (P and (\s. P')) Q (assert P' >>= f) g" by (force simp: corres_underlying_def assert_def return_def bind_def fail_def) lemma corres_assert_assume_r: "corres dc P Q f (g ()) \ corres dc P (Q and (\s. Q')) f (assert Q' >>= g)" by (force simp: corres_underlying_def assert_def return_def bind_def fail_def) lemma cur_tcb'_ksReadyQueuesL1Bitmap_upd[simp]: "cur_tcb' (s\ksReadyQueuesL1Bitmap := x\) = cur_tcb' s" unfolding cur_tcb'_def by simp lemma cur_tcb'_ksReadyQueuesL2Bitmap_upd[simp]: "cur_tcb' (s\ksReadyQueuesL2Bitmap := x\) = cur_tcb' s" unfolding cur_tcb'_def by simp crunch cur[wp]: tcbSchedEnqueue cur_tcb' (simp: unless_def) (* FIXME: move *) lemma corres_noop3: assumes x: "\s s'. \P s; P' s'; (s, s') \ sr\ \ \op = s\ f \\\r. op = s\" assumes y: "\s s'. \P s; P' s'; (s, s') \ sr\ \ \op = s'\ g \\r. op = s'\" assumes z: "nf' \ no_fail P' g" shows "corres_underlying sr nf nf' dc P P' f g" apply (clarsimp simp: corres_underlying_def) apply (rule conjI) apply clarsimp apply (rule use_exs_valid) apply (rule exs_hoare_post_imp) prefer 2 apply (rule x) apply assumption+ apply simp_all apply (subgoal_tac "ba = b") apply simp apply (rule sym) apply (rule use_valid[OF _ y], assumption+) apply simp apply (insert z) apply (clarsimp simp: no_fail_def) done lemma corres_symb_exec_l': assumes z: "\rv. corres_underlying sr nf nf' r (Q' rv) P' (x rv) y" assumes x: "\s. P s \ \op = s\ m \\\r. op = s\" assumes y: "\Q\ m \Q'\" shows "corres_underlying sr nf nf' r (P and Q) P' (m >>= (\rv. x rv)) y" apply (rule corres_guard_imp) apply (subst gets_bind_ign[symmetric], rule corres_split) apply (rule z) apply (rule corres_noop3) apply (erule x) apply (rule gets_wp) apply (rule non_fail_gets) apply (rule y) apply (rule gets_wp) apply simp+ done lemma corres_symb_exec_r': assumes z: "\rv. corres_underlying sr nf nf' r P (P'' rv) x (y rv)" assumes y: "\P'\ m \P''\" assumes x: "\s. Q' s \ \op = s\ m \\r. op = s\" assumes nf: "nf' \ no_fail R' m" shows "corres_underlying sr nf nf' r P (P' and Q' and R') x (m >>= (\rv. y rv))" apply (rule corres_guard_imp) apply (subst gets_bind_ign[symmetric], rule corres_split) apply (rule z) apply (rule_tac P'="a' and a''" for a' a'' in corres_noop3) apply (simp add: simpler_gets_def exs_valid_def) apply clarsimp apply (erule x) apply (rule no_fail_pre) apply (erule nf) apply clarsimp apply assumption apply (rule gets_wp) apply (rule y) apply simp+ done lemma corres_case_list: "\list = list'; corres r P P' f1 g1; \x xs. corres r (Q x xs) (Q' x xs) (f2 x xs) (g2 x xs)\ \ corres r (case list of [] \ P | x # xs \ Q x xs) (case list' of [] \ P' | x # xs \ Q' x xs) (case list of [] \ f1 | x # xs \ f2 x xs) (case list' of [] \ g1 | x # xs \ g2 x xs)" apply (cases list) apply (rule corres_guard_imp, force+)+ done lemma findM_corres: "\\x. x \ set xs \ corres op = P P' (f x) (f' x); \x. x \ set xs \ \P\ f x \\r. P\; \x. x \ set xs \ \P'\ f' x \\r. P'\\ \ corres op = P P' (findM f xs) (findM f' xs)" apply (induct xs) apply simp apply simp apply (rule corres_guard_imp) apply (rule corres_split[where r'="op ="]) apply (rule corres_if[where P=P and P'=P']) apply simp apply simp apply force apply force apply (atomize, (erule_tac x=a in allE)+, force)+ done lemma thread_get_exs_valid[wp]: "tcb_at t s \ \op = s\ thread_get f t \\\r. op = s\" apply (clarsimp simp: get_thread_state_def assert_opt_def fail_def thread_get_def gets_the_def exs_valid_def gets_def get_def bind_def return_def split: option.splits) apply (erule get_tcb_at) done lemma gts_exs_valid[wp]: "tcb_at t s \ \op = s\ get_thread_state t \\\r. op = s\" apply (clarsimp simp: get_thread_state_def assert_opt_def fail_def thread_get_def gets_the_def exs_valid_def gets_def get_def bind_def return_def split: option.splits) apply (erule get_tcb_at) done lemma guarded_switch_to_corres: "corres dc (valid_arch_state and valid_objs and valid_asid_map and valid_arch_objs and pspace_aligned and pspace_distinct and valid_vs_lookup and valid_global_objs and unique_table_refs o caps_of_state and st_tcb_at runnable t and valid_etcbs) (valid_arch_state' and valid_pspace' and Invariants_H.valid_queues and st_tcb_at' runnable' t and cur_tcb') (guarded_switch_to t) (switchToThread t)" apply (simp add: guarded_switch_to_def) apply (rule corres_guard_imp) apply (rule corres_symb_exec_l'[OF _ gts_exs_valid]) apply (rule corres_assert_assume_l) apply (rule switch_thread_corres) apply (force simp: st_tcb_at_tcb_at) apply (wp gts_st_tcb_at) apply (force simp: st_tcb_at_tcb_at)+ done lemma findM_gets_outside: assumes terminate_False: "(\s s' x. (False,s') \ fst (g x s) \ s = s')" shows "findM (\x. gets (\ks. f ks x) >>= g) l = do f' \ gets (\ks. f ks); findM (\x. (g (f' x))) l od" apply (rule ext) apply (induct l) apply (clarsimp simp add: bind_def gets_def get_def return_def) apply simp apply (clarsimp simp add: bind_def gets_def get_def return_def cong: if_cong split: if_split_asm) apply safe apply (clarsimp split: if_split_asm) apply force apply (frule terminate_False) apply simp apply (rule_tac x="(False,ba)" in bexI,clarsimp+) apply (clarsimp split: if_split_asm)+ apply force apply (frule terminate_False) apply clarsimp apply (rule_tac x="(False,ba)" in bexI,clarsimp+) apply (clarsimp split: if_split_asm)+ apply (frule terminate_False) apply clarsimp apply force apply (clarsimp split: if_split_asm)+ apply (frule terminate_False) apply clarsimp apply (rule bexI [rotated], assumption) apply auto done lemma corres_gets_queues: "corres ready_queues_relation \ \ (gets ready_queues) (gets ksReadyQueues)" apply (simp add: state_relation_def) done lemma Max_lt_set: "\y\A. Max B < (y::word8) \ \y\A. y \ B" apply clarsimp apply (drule_tac x=y in bspec,simp) apply (case_tac "B = {}") apply simp+ apply blast done abbreviation "enumPrio \ [0.e.maxPriority]" lemma enumPrio_word_div: fixes v :: "8 word" assumes vlt: "unat v \ unat maxPriority" shows "\xs ys. enumPrio = xs @ [v] @ ys \ (\x\set xs. x < v) \ (\y\set ys. v < y)" apply (subst upto_enum_word) apply (subst upt_add_eq_append'[where j="unat v"]) apply simp apply (rule le_SucI) apply (rule vlt) apply (simp only: upt_conv_Cons vlt[simplified less_Suc_eq_le[symmetric]]) apply (intro exI conjI) apply fastforce apply clarsimp apply (drule of_nat_mono_maybe[rotated, where 'a="8"]) apply (fastforce simp: vlt) apply simp apply (clarsimp simp: Suc_le_eq) apply (erule disjE) apply (drule of_nat_mono_maybe[rotated, where 'a="8"]) apply (simp add: maxPriority_def numPriorities_def) apply (clarsimp simp: unat_of_nat_eq) apply (erule conjE) apply (drule_tac y="unat v" and x="x" in of_nat_mono_maybe[rotated, where 'a="8"]) apply (simp add: maxPriority_def numPriorities_def)+ done lemma rev_enumPrio_word_div: "unat v \ unat maxPriority \ \xs ys. rev enumPrio = ys @ [(v::word8)] @ xs \ (\x\set xs. x < v) \ (\y\set ys. v < y)" apply (cut_tac v=v in enumPrio_word_div) apply clarsimp apply clarsimp apply (rule_tac x="rev xs" in exI) apply (rule_tac x="rev ys" in exI) apply simp done lemma findM_ignore_head: "\y\ set ys. f y = return False \ findM f (ys @ xs) = findM f xs" apply (induct ys,simp+) done lemma curDomain_corres: "corres (op =) \ \ (gets cur_domain) (curDomain)" by (simp add: curDomain_def state_relation_def) lemma valid_queues_non_empty: "\s d p. \ valid_queues s; ksReadyQueues s (d, p) \ [] \ \ unat (Max {prio. ksReadyQueues s (d, prio) \ []}) \ unat maxPriority" apply (subst word_le_nat_alt[symmetric]) apply (subst Max_le_iff) apply simp apply blast apply (clarsimp simp: valid_queues_def valid_queues_no_bitmap_def) apply (drule spec) apply (drule_tac x="a" in spec) apply fastforce done lemma lookupBitmapPriority_Max_eqI: "\ valid_bitmapQ s ; bitmapQ_no_L1_orphans s ; ksReadyQueuesL1Bitmap s d \ 0 \ \ lookupBitmapPriority d s = (Max {prio. ksReadyQueues s (d, prio) \ []})" apply (rule Max_eqI[simplified eq_commute]; simp) apply (fastforce simp: bitmapL1_highest_lookup valid_bitmapQ_bitmapQ_simp) apply (metis valid_bitmapQ_bitmapQ_simp bitmapQ_from_bitmap_lookup) done lemma corres_gets_queues_getReadyQueuesL1Bitmap: "corres (\qs l1. ((l1 = 0) = (\p. qs p = []))) \ valid_queues (gets (\s. ready_queues s d)) (getReadyQueuesL1Bitmap d)" unfolding state_relation_def valid_queues_def getReadyQueuesL1Bitmap_def by (clarsimp simp: bitmapL1_zero_ksReadyQueues ready_queues_relation_def) lemma guarded_switch_to_chooseThread_fragment_corres: "corres dc (P and st_tcb_at runnable t and invs and valid_sched) (P' and st_tcb_at' runnable' t and invs_no_cicd') (guarded_switch_to t) (do runnable \ isRunnable t; y \ assert runnable; ThreadDecls_H.switchToThread t od)" unfolding guarded_switch_to_def isRunnable_def apply simp apply (rule corres_guard_imp) apply (rule corres_split[OF _ gts_corres]) apply (rule corres_assert_assume_l) apply (rule corres_assert_assume_r) apply (rule switch_thread_corres) apply (wp gts_st_tcb_at) apply (clarsimp simp: st_tcb_at_tcb_at invs_def valid_state_def valid_pspace_def valid_sched_def invs_valid_vs_lookup invs_unique_refs) apply (auto elim!: pred_tcb'_weakenE split: thread_state.splits simp: pred_tcb_at' runnable'_def all_invs_but_ct_idle_or_in_cur_domain'_def) done lemma bitmap_lookup_queue_is_max_non_empty: "\ valid_queues s'; (s, s') \ state_relation; invs s; ksReadyQueuesL1Bitmap s' (ksCurDomain s') \ 0 \ \ ksReadyQueues s' (ksCurDomain s', lookupBitmapPriority (ksCurDomain s') s') = max_non_empty_queue (ready_queues s (cur_domain s))" unfolding all_invs_but_ct_idle_or_in_cur_domain'_def valid_queues_def by (clarsimp simp add: max_non_empty_queue_def lookupBitmapPriority_Max_eqI state_relation_def ready_queues_relation_def) lemma ksReadyQueuesL1Bitmap_return_wp: "\\s. P (ksReadyQueuesL1Bitmap s d) s \ getReadyQueuesL1Bitmap d \\rv s. P rv s\" unfolding getReadyQueuesL1Bitmap_def by wp lemma ksReadyQueuesL2Bitmap_return_wp: "\\s. P (ksReadyQueuesL2Bitmap s (d,i)) s \ getReadyQueuesL2Bitmap d i \\rv s. P rv s\" unfolding getReadyQueuesL2Bitmap_def by wp lemma ksReadyQueuesL1Bitmap_st_tcb_at': "\ ksReadyQueuesL1Bitmap s (ksCurDomain s) \ 0 ; valid_queues s \ \ st_tcb_at' runnable' (hd (ksReadyQueues s (ksCurDomain s, lookupBitmapPriority (ksCurDomain s) s))) s" apply (drule bitmapQ_from_bitmap_lookup; clarsimp simp: valid_queues_def) apply (clarsimp simp add: valid_bitmapQ_bitmapQ_simp) apply (case_tac "ksReadyQueues s (ksCurDomain s, lookupBitmapPriority (ksCurDomain s) s)") apply simp apply (simp add: valid_queues_no_bitmap_def) apply (erule_tac x="ksCurDomain s" in allE) apply (erule_tac x="lookupBitmapPriority (ksCurDomain s) s" in allE) apply (clarsimp simp: st_tcb_at'_def) apply (erule obj_at'_weaken) apply simp done lemma chooseThread_corres: "corres dc (invs and valid_sched) (invs_no_cicd') choose_thread chooseThread" (is "corres _ ?PREI ?PREH _ _") proof - show ?thesis unfolding choose_thread_def chooseThread_def numDomains_def apply (simp only: numDomains_def return_bind Let_def) apply (simp cong: if_cong) (* clean up if 1 < numDomains *) apply (subst if_swap[where P="_ \ 0"]) (* put switchToIdleThread on first branch*) apply (rule corres_name_pre) apply (rule corres_guard_imp) apply (rule corres_split[OF _ curDomain_corres]) apply clarsimp apply (rule corres_split[OF _ corres_gets_queues_getReadyQueuesL1Bitmap]) apply (erule corres_if2[OF sym]) apply (rule switch_idle_thread_corres) apply (rule corres_symb_exec_r) apply (rule corres_symb_exec_r) apply (rule_tac P="\s. ?PREI s \ queues = ready_queues s (cur_domain s) \ st_tcb_at runnable (hd (max_non_empty_queue queues)) s" and P'="\s. (?PREH s \ st_tcb_at' runnable' (hd queue) s) \ l1 = ksReadyQueuesL1Bitmap s (ksCurDomain s) \ l1 \ 0 \ queue = ksReadyQueues s (ksCurDomain s, lookupBitmapPriority (ksCurDomain s) s)" and F="hd queue = hd (max_non_empty_queue queues)" in corres_req) apply (fastforce dest!: invs_no_cicd'_queues simp: bitmap_lookup_queue_is_max_non_empty) apply clarsimp apply (rule corres_guard_imp) apply (rule_tac P=\ and P'=\ in guarded_switch_to_chooseThread_fragment_corres) apply (wp ksReadyQueuesL2Bitmap_return_wp |clarsimp simp: getQueue_def getReadyQueuesL2Bitmap_def)+ apply (clarsimp simp: if_apply_def2) apply (wp hoare_vcg_conj_lift hoare_vcg_imp_lift ksReadyQueuesL1Bitmap_return_wp) apply (simp add: curDomain_def, wp)+ apply (clarsimp simp: valid_sched_def DetSchedInvs_AI.valid_queues_def max_non_empty_queue_def) apply (erule_tac x="cur_domain sa" in allE) apply (erule_tac x="Max {prio. ready_queues sa (cur_domain sa) prio \ []}" in allE) apply (case_tac "ready_queues sa (cur_domain sa) (Max {prio. ready_queues sa (cur_domain sa) prio \ []})") apply (clarsimp) apply (subgoal_tac "ready_queues sa (cur_domain sa) (Max {prio. ready_queues sa (cur_domain sa) prio \ []}) \ []") apply (fastforce elim!: setcomp_Max_has_prop)+ apply (clarsimp dest!: invs_no_cicd'_queues) apply (fold lookupBitmapPriority_def) apply (fastforce intro: ksReadyQueuesL1Bitmap_st_tcb_at') done qed lemma thread_get_comm: "do x \ thread_get f p; y \ gets g; k x y od = do y \ gets g; x \ thread_get f p; k x y od" apply (rule ext) apply (clarsimp simp add: gets_the_def assert_opt_def bind_def gets_def get_def return_def thread_get_def fail_def split: option.splits) done lemma schact_bind_inside: "do x \ f; (case act of resume_cur_thread \ f1 x | switch_thread t \ f2 t x | choose_new_thread \ f3 x) od = (case act of resume_cur_thread \ (do x \ f; f1 x od) | switch_thread t \ (do x \ f; f2 t x od) | choose_new_thread \ (do x \ f; f3 x od))" apply (case_tac act,simp_all) done interpretation tcb_sched_action_extended: is_extended' "tcb_sched_action f a" by (unfold_locales) lemma domain_time_corres: "corres (op =) \ \ (gets domain_time) getDomainTime" by (simp add: getDomainTime_def state_relation_def) lemma next_domain_corres: "corres dc \ \ next_domain nextDomain" apply (simp add: next_domain_def nextDomain_def) apply (rule corres_modify) apply (simp add: state_relation_def Let_def dschLength_def dschDomain_def) done lemma valid_queues'_ksCurDomain[simp]: "valid_queues' (ksCurDomain_update f s) = valid_queues' s" by (simp add: valid_queues'_def) lemma valid_queues'_ksDomScheduleIdx[simp]: "valid_queues' (ksDomScheduleIdx_update f s) = valid_queues' s" by (simp add: valid_queues'_def) lemma valid_queues'_ksDomSchedule[simp]: "valid_queues' (ksDomSchedule_update f s) = valid_queues' s" by (simp add: valid_queues'_def) lemma valid_queues'_ksDomainTime[simp]: "valid_queues' (ksDomainTime_update f s) = valid_queues' s" by (simp add: valid_queues'_def) lemma valid_queues'_ksWorkUnitsCompleted[simp]: "valid_queues' (ksWorkUnitsCompleted_update f s) = valid_queues' s" by (simp add: valid_queues'_def) lemma valid_queues_ksCurDomain[simp]: "Invariants_H.valid_queues (ksCurDomain_update f s) = Invariants_H.valid_queues s" by (simp add: Invariants_H.valid_queues_def valid_queues_no_bitmap_def bitmapQ_defs) lemma valid_queues_ksDomScheduleIdx[simp]: "Invariants_H.valid_queues (ksDomScheduleIdx_update f s) = Invariants_H.valid_queues s" by (simp add: Invariants_H.valid_queues_def valid_queues_no_bitmap_def bitmapQ_defs) lemma valid_queues_ksDomSchedule[simp]: "Invariants_H.valid_queues (ksDomSchedule_update f s) = Invariants_H.valid_queues s" by (simp add: Invariants_H.valid_queues_def valid_queues_no_bitmap_def bitmapQ_defs) lemma valid_queues_ksDomainTime[simp]: "Invariants_H.valid_queues (ksDomainTime_update f s) = Invariants_H.valid_queues s" by (simp add: Invariants_H.valid_queues_def valid_queues_no_bitmap_def bitmapQ_defs) lemma valid_queues_ksWorkUnitsCompleted[simp]: "Invariants_H.valid_queues (ksWorkUnitsCompleted_update f s) = Invariants_H.valid_queues s" by (simp add: Invariants_H.valid_queues_def valid_queues_no_bitmap_def bitmapQ_defs) lemma valid_irq_node'_ksCurDomain[simp]: "valid_irq_node' w (ksCurDomain_update f s) = valid_irq_node' w s" by (simp add: valid_irq_node'_def) lemma valid_irq_node'_ksDomScheduleIdx[simp]: "valid_irq_node' w (ksDomScheduleIdx_update f s) = valid_irq_node' w s" by (simp add: valid_irq_node'_def) lemma valid_irq_node'_ksDomSchedule[simp]: "valid_irq_node' w (ksDomSchedule_update f s) = valid_irq_node' w s" by (simp add: valid_irq_node'_def) lemma valid_irq_node'_ksDomainTime[simp]: "valid_irq_node' w (ksDomainTime_update f s) = valid_irq_node' w s" by (simp add: valid_irq_node'_def) lemma valid_irq_node'_ksWorkUnitsCompleted[simp]: "valid_irq_node' w (ksWorkUnitsCompleted_update f s) = valid_irq_node' w s" by (simp add: valid_irq_node'_def) lemma sch_act_wf_ksCurDomain [simp]: "sa = ChooseNewThread \ sch_act_wf sa (ksCurDomain_update f s) = sch_act_wf sa s" apply (cases sa) apply (simp_all add: ct_in_state'_def tcb_in_cur_domain'_def) done lemma next_domain_valid_sched[wp]: "\ valid_sched and (\s. scheduler_action s = choose_new_thread)\ next_domain \ \_. valid_sched \" apply (simp add: next_domain_def Let_def) apply (wp, simp add: valid_sched_def valid_sched_action_2_def ct_not_in_q_2_def) apply (simp add:valid_blocked_2_def) done lemma nextDomain_invs_no_cicd': "\ invs' and (\s. ksSchedulerAction s = ChooseNewThread)\ nextDomain \ \_. invs_no_cicd' \" apply (simp add: nextDomain_def Let_def dschLength_def dschDomain_def) apply wp apply (clarsimp simp: invs'_def valid_state'_def valid_machine_state'_def ct_not_inQ_def cur_tcb'_def ct_idle_or_in_cur_domain'_def dschDomain_def all_invs_but_ct_idle_or_in_cur_domain'_def) done lemma schedule_ChooseNewThread_fragment_corres: "corres dc (invs and valid_sched and (\s. scheduler_action s = choose_new_thread)) (invs' and (\s. ksSchedulerAction s = ChooseNewThread)) (do _ \ when (domainTime = 0) next_domain; choose_thread od) (do _ \ when (domainTime = 0) nextDomain; chooseThread od)" apply (subst bind_dummy_ret_val) apply (subst bind_dummy_ret_val) apply (rule corres_guard_imp) apply (rule corres_split[OF _ corres_when]) apply (simp add: K_bind_def) apply (rule chooseThread_corres) apply simp apply (rule next_domain_corres) apply (wp nextDomain_invs_no_cicd') apply (clarsimp simp: valid_sched_def invs'_def valid_state'_def all_invs_but_ct_idle_or_in_cur_domain'_def)+ done lemma schedule_corres: "corres dc (invs and valid_sched and valid_list) invs' (Schedule_A.schedule) ThreadDecls_H.schedule" apply (clarsimp simp: Schedule_A.schedule_def Thread_H.schedule_def) apply (subst thread_get_test) apply (subst thread_get_comm) apply (subst schact_bind_inside) apply (rule corres_guard_imp) apply (rule corres_split[OF _ gct_corres[THEN corres_rel_imp[where r="\x y. y = x"],simplified]]) apply (rule corres_guard_imp) apply (rule corres_split[OF _ get_sa_corres']) apply (rule corres_split_sched_act,assumption) apply (rule_tac P="tcb_at cur" in corres_symb_exec_l') apply (rule_tac corres_symb_exec_l') apply simp apply (rule corres_assert_ret) apply (wp gets_exs_valid | simp)+ apply (rule thread_get_wp') apply simp apply (rule corres_split[OF _ thread_get_isRunnable_corres]) apply (rule corres_split[OF _ corres_when]) apply (rule corres_split[OF _ guarded_switch_to_corres]) apply (rule set_sa_corres) apply (wp | simp)+ apply (rule tcbSchedEnqueue_corres) apply (wp thread_get_wp' | simp)+ apply (rule corres_split[OF _ thread_get_isRunnable_corres]) apply (rule corres_split[OF _ corres_when]) apply (rule corres_split[OF _ domain_time_corres], clarsimp, fold dc_def) apply (rule corres_split[OF _ schedule_ChooseNewThread_fragment_corres, simplified bind_assoc]) apply (rule set_sa_corres) apply (wp | simp)+ apply (wp | simp add: getDomainTime_def)+ apply (rule tcbSchedEnqueue_corres, simp) apply (simp_all only: cong: if_cong Deterministic_A.scheduler_action.case_cong Structures_H.scheduler_action.case_cong) apply ((wp thread_get_wp' hoare_vcg_conj_lift hoare_drop_imps | clarsimp)+) thm tcbSchedEnqueue_invs' apply (rule conjI,simp) apply (clarsimp split:Deterministic_A.scheduler_action.splits simp: invs_psp_aligned invs_distinct invs_valid_objs invs_arch_state invs_arch_objs) apply (intro impI conjI allI tcb_at_invs | (fastforce simp: invs_def cur_tcb_def valid_arch_caps_def valid_sched_def valid_sched_action_def is_activatable_def st_tcb_at_def obj_at_def valid_state_def only_idle_def valid_etcbs_def weak_valid_sched_action_def not_cur_thread_def tcb_at_invs ))+ apply (cut_tac s = s in valid_blocked_valid_blocked_except) prefer 2 apply (simp add:valid_sched_def) apply (simp add:valid_sched_def) apply simp by (fastforce simp: invs'_def cur_tcb'_def valid_state'_def st_tcb_at'_def sch_act_wf_def valid_pspace'_def valid_objs'_maxDomain valid_objs'_maxPriority comp_def split: scheduler_action.splits) lemma ssa_all_invs_but_ct_not_inQ': "\all_invs_but_ct_not_inQ' and sch_act_wf sa and (\s. sa = ResumeCurrentThread \ ksCurThread s = ksIdleThread s \ tcb_in_cur_domain' (ksCurThread s) s)\ setSchedulerAction sa \\rv. all_invs_but_ct_not_inQ'\" proof - have obj_at'_sa: "\P addr f s. obj_at' P addr (ksSchedulerAction_update f s) = obj_at' P addr s" by (fastforce intro: obj_at'_pspaceI) have valid_pspace'_sa: "\f s. valid_pspace' (ksSchedulerAction_update f s) = valid_pspace' s" by (rule valid_pspace'_ksPSpace, simp) have iflive_sa: "\f s. if_live_then_nonz_cap' (ksSchedulerAction_update f s) = if_live_then_nonz_cap' s" by (fastforce intro: if_live_then_nonz_cap'_pspaceI) have ifunsafe_sa[simp]: "\f s. if_unsafe_then_cap' (ksSchedulerAction_update f s) = if_unsafe_then_cap' s" by fastforce have idle_sa[simp]: "\f s. valid_idle' (ksSchedulerAction_update f s) = valid_idle' s" by fastforce show ?thesis apply (simp add: setSchedulerAction_def) apply wp apply (clarsimp simp add: invs'_def valid_state'_def cur_tcb'_def obj_at'_sa valid_pspace'_sa Invariants_H.valid_queues_def state_refs_of'_def iflive_sa ps_clear_def valid_irq_node'_def valid_queues'_def tcb_in_cur_domain'_def ct_idle_or_in_cur_domain'_def bitmapQ_defs valid_queues_no_bitmap_def cong: option.case_cong) done qed lemma ssa_ct_not_inQ: "\\s. sa = ResumeCurrentThread \ obj_at' (Not \ tcbQueued) (ksCurThread s) s\ setSchedulerAction sa \\rv. ct_not_inQ\" by (simp add: setSchedulerAction_def ct_not_inQ_def, wp, clarsimp) lemma ssa_all_invs_but_ct_not_inQ''[simplified]: "\\s. (all_invs_but_ct_not_inQ' s \ sch_act_wf sa s) \ (sa = ResumeCurrentThread \ ksCurThread s = ksIdleThread s \ tcb_in_cur_domain' (ksCurThread s) s) \ (sa = ResumeCurrentThread \ obj_at' (Not \ tcbQueued) (ksCurThread s) s)\ setSchedulerAction sa \\rv. invs'\" apply (simp only: all_invs_but_not_ct_inQ_check' [symmetric]) apply (rule hoare_elim_pred_conj) apply (wp hoare_vcg_conj_lift [OF ssa_all_invs_but_ct_not_inQ' ssa_ct_not_inQ]) apply (clarsimp) done lemma ssa_invs': "\invs' and sch_act_wf sa and (\s. sa = ResumeCurrentThread \ ksCurThread s = ksIdleThread s \ tcb_in_cur_domain' (ksCurThread s) s) and (\s. sa = ResumeCurrentThread \ obj_at' (Not \ tcbQueued) (ksCurThread s) s)\ setSchedulerAction sa \\rv. invs'\" apply (wp ssa_all_invs_but_ct_not_inQ'') apply (clarsimp simp add: invs'_def valid_state'_def) done lemma getDomainTime_wp[wp]: "\\s. P (ksDomainTime s) s \ getDomainTime \ P \" unfolding getDomainTime_def by wp lemma switchToThread_ct_not_queued_2: "\invs_no_cicd' and tcb_at' t\ switchToThread t \\rv s. obj_at' (Not \ tcbQueued) (ksCurThread s) s\" (is "\_\ _ \\_. ?POST\") apply (simp add: Thread_H.switchToThread_def) apply (wp) apply (simp add: ARM_H.switchToThread_def setCurThread_def) apply (wp tcbSchedDequeue_not_tcbQueued | simp )+ done lemma hd_ksReadyQueues_runnable_2: "\Invariants_H.valid_queues s; ksReadyQueues s (d, p) = a # as\ \ st_tcb_at' runnable' a s" apply( rule valid_queues_running) apply(rule suffix_Cons_mem) apply(simp add: suffix_def) apply(rule exI) apply(rule eq_Nil_appendI) by auto lemma setCurThread_obj_at': "\ obj_at' P t \ setCurThread t \\rv s. obj_at' P (ksCurThread s) s \" proof - have obj_at'_ct: "\P addr f s. obj_at' P addr (ksCurThread_update f s) = obj_at' P addr s" by (fastforce intro: obj_at'_pspaceI) show ?thesis apply (simp add: setCurThread_def) apply wp apply (simp add: ct_in_state'_def st_tcb_at'_def obj_at'_ct) done qed lemma switchToIdleThread_ct_not_queued_no_cicd': "\ invs_no_cicd' \ switchToIdleThread \\rv s. obj_at' (Not \ tcbQueued) (ksCurThread s) s \" apply (simp add: Thread_H.switchToIdleThread_def) apply (wp setCurThread_obj_at') apply (rule idle'_not_tcbQueued') apply (simp add: invs_no_cicd'_def)+ done lemma switchToIdleThread_activatable_2[wp]: "\invs_no_cicd'\ switchToIdleThread \\rv. ct_in_state' activatable'\" apply (simp add: Thread_H.switchToIdleThread_def ARM_H.switchToIdleThread_def) apply (wp setCurThread_ct_in_state) apply (clarsimp simp: all_invs_but_ct_idle_or_in_cur_domain'_def valid_state'_def valid_idle'_def pred_tcb_at'_def obj_at'_def) done lemma switchToThread_tcb_in_cur_domain': "\tcb_in_cur_domain' thread\ ThreadDecls_H.switchToThread thread \\y s. tcb_in_cur_domain' (ksCurThread s) s\" apply (simp add: Thread_H.switchToThread_def) apply (rule hoare_pre) apply (wp) apply (simp add: ARM_H.switchToThread_def setCurThread_def) apply (wp tcbSchedDequeue_not_tcbQueued | simp )+ apply (simp add:tcb_in_cur_domain'_def) apply (wp tcbSchedDequeue_tcbDomain | wps)+ apply (clarsimp simp:tcb_in_cur_domain'_def) done lemma chooseThread_invs_no_cicd'_posts: (* generic version *) "\ invs_no_cicd' \ chooseThread \\rv s. obj_at' (Not \ tcbQueued) (ksCurThread s) s \ ct_in_state' activatable' s \ (ksCurThread s = ksIdleThread s \ tcb_in_cur_domain' (ksCurThread s) s) \" (is "\_\ _ \\_. ?POST\") proof - note switchToThread_invs[wp del] note switchToThread_invs_no_cicd'[wp del] note switchToThread_lookupBitmapPriority_wp[wp] note assert_wp[wp del] show ?thesis unfolding chooseThread_def Let_def numDomains_def curDomain_def apply (simp only: return_bind, simp) apply (rule hoare_seq_ext[where B="\rv s. invs_no_cicd' s \ rv = ksCurDomain s"]) apply (rule_tac B="\rv s. invs_no_cicd' s \ curdom = ksCurDomain s \ rv = ksReadyQueuesL1Bitmap s curdom" in hoare_seq_ext) apply (rename_tac l1) apply (case_tac "l1 = 0") (* switch to idle thread *) apply simp apply (rule hoare_pre) apply (wp_once switchToIdleThread_ct_not_queued_no_cicd') apply (wp_once) apply ((wp hoare_disjI1 switchToIdleThread_curr_is_idle)+)[1] apply simp (* we have a thread to switch to *) apply (clarsimp simp: bitmap_fun_defs) apply (wp assert_inv switchToThread_ct_not_queued_2 assert_inv hoare_disjI2 switchToThread_tcb_in_cur_domain') apply clarsimp apply (clarsimp dest!: invs_no_cicd'_queues simp: valid_queues_def lookupBitmapPriority_def[symmetric]) apply (drule (3) lookupBitmapPriority_obj_at') apply normalise_obj_at' apply (fastforce simp: tcb_in_cur_domain'_def inQ_def elim: obj_at'_weaken) apply (wp | simp add: bitmap_fun_defs curDomain_def)+ done qed lemma chooseThread_activatable_2: "\invs_no_cicd'\ chooseThread \\rv. ct_in_state' activatable'\" apply (rule hoare_pre, rule hoare_strengthen_post) apply (rule chooseThread_invs_no_cicd'_posts) apply simp+ done lemma chooseThread_activatable: "\invs'\ chooseThread \\rv. ct_in_state' activatable'\" apply (rule hoare_pre, rule chooseThread_activatable_2) apply (simp add: invs'_to_invs_no_cicd'_def) done lemma chooseThread_ct_not_queued_2: "\ invs_no_cicd'\ chooseThread \\rv s. obj_at' (Not \ tcbQueued) (ksCurThread s) s\" (is "\_\ _ \\_. ?POST\") apply (rule hoare_pre, rule hoare_strengthen_post) apply (rule chooseThread_invs_no_cicd'_posts) apply simp+ done lemma chooseThread_invs_no_cicd': "\ invs_no_cicd' \ chooseThread \\rv. invs' \" proof - note switchToThread_invs[wp del] note switchToThread_invs_no_cicd'[wp del] note switchToThread_lookupBitmapPriority_wp[wp] note assert_wp[wp del] (* FIXME this is almost identical to the chooseThread_invs_no_cicd'_posts proof, can generalise? *) show ?thesis unfolding chooseThread_def Let_def numDomains_def curDomain_def apply (simp only: return_bind, simp) apply (rule hoare_seq_ext[where B="\rv s. invs_no_cicd' s \ rv = ksCurDomain s"]) apply (rule_tac B="\rv s. invs_no_cicd' s \ curdom = ksCurDomain s \ rv = ksReadyQueuesL1Bitmap s curdom" in hoare_seq_ext) apply (rename_tac l1) apply (case_tac "l1 = 0") (* switch to idle thread *) apply (simp, wp_once switchToIdleThread_invs_no_cicd', simp) (* we have a thread to switch to *) apply (clarsimp simp: bitmap_fun_defs) apply (wp assert_inv) apply (clarsimp dest!: invs_no_cicd'_queues simp: valid_queues_def) apply (fastforce elim: bitmapQ_from_bitmap_lookup simp: lookupBitmapPriority_def) apply (wp | simp add: bitmap_fun_defs curDomain_def)+ done qed lemma chooseThread_in_cur_domain': "\ invs_no_cicd' \ chooseThread \\rv s. ksCurThread s = ksIdleThread s \ tcb_in_cur_domain' (ksCurThread s) s\" apply (rule hoare_pre, rule hoare_strengthen_post) apply (rule chooseThread_invs_no_cicd'_posts, simp_all) done lemma schedule_ChooseNewThread_fragment_invs': "\ invs' and (\s. ksSchedulerAction s = ChooseNewThread) \ do _ \ when (domainTime = 0) nextDomain; chooseThread od \ \_ s. invs' s \ ct_in_state' activatable' s \ obj_at' (Not \ tcbQueued) (ksCurThread s) s \ (ksCurThread s = ksIdleThread s \ tcb_in_cur_domain' (ksCurThread s) s) \" apply (rule hoare_seq_ext) apply (wp chooseThread_ct_not_queued_2 chooseThread_activatable_2 chooseThread_invs_no_cicd') apply (wp chooseThread_in_cur_domain' nextDomain_invs_no_cicd') apply (simp add:nextDomain_def) apply (clarsimp simp: invs'_def all_invs_but_ct_idle_or_in_cur_domain'_def Let_def valid_state'_def) done lemma schedule_invs': "\invs'\ ThreadDecls_H.schedule \\rv. invs'\" apply (simp add: schedule_def) apply (rule_tac hoare_seq_ext, rename_tac t) apply (wp, wpc) -- "action = ResumeCurrentThread" apply (wp)[1] -- "action = ChooseNewThread" apply (rule_tac hoare_seq_ext, rename_tac r) apply (rule hoare_seq_ext, simp add: K_bind_def) apply (rule hoare_seq_ext) apply (rule seq_ext[OF schedule_ChooseNewThread_fragment_invs' _, simplified bind_assoc]) apply (wp ssa_invs' chooseThread_invs_no_cicd') apply clarsimp apply (wp)[3] -- "action = SwitchToThread" apply (rename_tac word) apply (rule_tac hoare_seq_ext, rename_tac r) apply (wp ssa_invs') apply (clarsimp) apply wp apply (rule_tac Q="\_. (\s. tcb_in_cur_domain' (ksCurThread s) s) and (\s. obj_at' (Not \ tcbQueued) (ksCurThread s) s)" in hoare_post_imp) apply simp apply (wp switchToThread_tcb_in_cur_domain' switchToThread_ct_not_queued) apply (rule_tac Q="\_. (\s. st_tcb_at' activatable' word s) and invs' and (\s. tcb_in_cur_domain' word s)" in hoare_post_imp) apply (clarsimp simp:st_tcb_at'_def valid_state'_def obj_at'_def) apply (wp) apply (frule invs_sch_act_wf') apply (auto elim!: obj_at'_weakenE simp: pred_tcb_at'_def) done lemma setCurThread_nosch: "\\s. P (ksSchedulerAction s)\ setCurThread t \\rv s. P (ksSchedulerAction s)\" apply (simp add: setCurThread_def) apply wp apply simp done lemma stt_nosch: "\\s. P (ksSchedulerAction s)\ switchToThread t \\rv s. P (ksSchedulerAction s)\" apply (simp add: Thread_H.switchToThread_def ARM_H.switchToThread_def storeWordUser_def) apply (wp setCurThread_nosch hoare_drop_imp |simp)+ done lemma stit_nosch[wp]: "\\s. P (ksSchedulerAction s)\ switchToIdleThread \\rv s. P (ksSchedulerAction s)\" apply (simp add: Thread_H.switchToIdleThread_def ARM_H.switchToIdleThread_def storeWordUser_def) apply (wp setCurThread_nosch | simp add: getIdleThread_def)+ done lemma chooseThread_nosch: "\\s. P (ksSchedulerAction s)\ chooseThread \\rv s. P (ksSchedulerAction s)\" unfolding chooseThread_def Let_def numDomains_def curDomain_def apply (simp only: return_bind, simp) apply (wp findM_inv | simp)+ apply (case_tac queue) apply (wp stt_nosch | simp add: curDomain_def bitmap_fun_defs)+ done lemma schedule_sch: "\\\ schedule \\rv s. ksSchedulerAction s = ResumeCurrentThread\" by (wp setSchedulerAction_direct | wpc| simp add: schedule_def)+ lemma schedule_sch_act_simple: "\\\ schedule \\rv. sch_act_simple\" apply (rule hoare_strengthen_post [OF schedule_sch]) apply (simp add: sch_act_simple_def) done lemma ssa_ct: "\ct_in_state' P\ setSchedulerAction sa \\rv. ct_in_state' P\" proof - have obj_at'_sa: "\P addr f s. obj_at' P addr (ksSchedulerAction_update f s) = obj_at' P addr s" by (fastforce intro: obj_at'_pspaceI) show ?thesis apply (unfold setSchedulerAction_def) apply wp apply (clarsimp simp add: ct_in_state'_def pred_tcb_at'_def obj_at'_sa) done qed lemma schedule_ct_activatable'[wp]: "\invs'\ ThreadDecls_H.schedule \\_. ct_in_state' activatable'\" apply (simp add: schedule_def) apply (rule_tac hoare_seq_ext, rename_tac t) apply (wp, wpc) -- "action = ResumeCurrentThread" apply (wp)[1] -- "action = ChooseNewThread" apply (rule_tac hoare_seq_ext, rename_tac r) apply (rule hoare_seq_ext, simp) apply (rule hoare_seq_ext) apply (rule seq_ext[OF schedule_ChooseNewThread_fragment_invs' _, simplified bind_assoc]) apply (wp ssa_invs') apply (clarsimp simp: ct_in_state'_def, simp) apply (wp)[3] -- "action = SwitchToThread" apply (rename_tac word) apply (rule_tac hoare_seq_ext, rename_tac r) apply (wp ssa_invs') apply (clarsimp simp: ct_in_state'_def, simp) apply (wp switchToThread_ct_not_queued sts_ct_in_state_neq') apply (rule_tac Q="\_. st_tcb_at' activatable' word and invs'" in hoare_post_imp) apply (fastforce simp: st_tcb_at'_def elim!: obj_at'_weakenE) apply (clarsimp simp: ) apply (wp) apply (frule invs_sch_act_wf') apply (auto elim!: obj_at'_weakenE simp: st_tcb_at'_def ) done lemma threadSet_sch_act_sane: "\sch_act_sane\ threadSet f t \\_. sch_act_sane\" by (wp sch_act_sane_lift) lemma rescheduleRequired_sch_act_sane[wp]: "\\\ rescheduleRequired \\rv. sch_act_sane\" apply (simp add: rescheduleRequired_def sch_act_sane_def setSchedulerAction_def) by (wp | wpc | clarsimp)+ lemma tcbSchedDequeue_sch_act_sane[wp]: "\sch_act_sane\ tcbSchedDequeue t \\_. sch_act_sane\" by (wp sch_act_sane_lift) lemma sts_sch_act_sane: "\sch_act_sane\ setThreadState st t \\_. sch_act_sane\" apply (simp add: setThreadState_def) apply (wp hoare_drop_imps | simp add: threadSet_sch_act_sane sane_update)+ done lemma sbn_sch_act_sane: "\sch_act_sane\ setBoundNotification ntfn t \\_. sch_act_sane\" apply (simp add: setBoundNotification_def) apply (wp | simp add: threadSet_sch_act_sane sane_update)+ done lemma possibleSwitchTo_corres: "corres dc (valid_etcbs and weak_valid_sched_action and cur_tcb and st_tcb_at runnable t) (Invariants_H.valid_queues and valid_queues' and (\s. weak_sch_act_wf (ksSchedulerAction s) s) and cur_tcb' and tcb_at' t and st_tcb_at' runnable' t and valid_objs') (possible_switch_to t b) (possibleSwitchTo t b)" apply (simp add: possible_switch_to_def possibleSwitchTo_def cong: if_cong) apply (rule corres_guard_imp) apply (rule corres_split[OF _ gct_corres]) apply simp apply (rule corres_split[OF _ curDomain_corres]) apply (rule corres_split[OF _ ethreadget_corres[where r="op ="]]) apply (rule corres_split[OF _ ethreadget_corres[where r="op ="]]) apply (rule corres_split[OF _ ethreadget_corres[where r="op ="]]) apply (rule corres_split[OF _ get_sa_corres]) apply (rule corres_if) apply (simp+)[2] apply (rule tcbSchedEnqueue_corres) apply (rule corres_split[where r'=dc]) apply (case_tac action,simp_all)[1] apply (rule rescheduleRequired_corres) apply (rule corres_if) apply (case_tac action,simp_all)[1] apply (rule set_sa_corres) apply simp apply (rule tcbSchedEnqueue_corres) apply (wp add: set_scheduler_action_wp del: ssa_lift | simp)+ apply (rule_tac Q="\r. st_tcb_at' runnable' t and tcb_in_cur_domain' t" in valid_prove_more) apply (rule ssa_lift) apply (clarsimp simp: valid_queues'_def weak_sch_act_wf_def Invariants_H.valid_queues_def tcb_in_cur_domain'_def st_tcb_at'_def valid_queues_no_bitmap_def bitmapQ_defs) apply (wp threadGet_wp | simp add: etcb_relation_def curDomain_def)+ apply (clarsimp simp: is_etcb_at_def valid_sched_action_def weak_valid_sched_action_def) apply (clarsimp simp: valid_sched_def etcb_at_def cur_tcb_def split: option.splits) apply (frule st_tcb_at_tcb_at) apply (frule(1) tcb_at_ekheap_dom[where x=t]) apply clarsimp apply (frule_tac x="cur_thread s" in tcb_at_ekheap_dom, simp) apply (clarsimp simp: is_etcb_at_def valid_sched_action_def weak_valid_sched_action_def) apply (clarsimp simp: cur_tcb'_def) apply (frule (1) valid_objs'_maxDomain) apply (frule (1) valid_objs'_maxDomain [where t=t]) apply (frule (1) valid_objs'_maxPriority) apply (frule (1) valid_objs'_maxPriority [where t=t]) apply (clarsimp simp: cur_tcb'_def obj_at'_def projectKOs objBits_simps) apply (rule_tac x=obje in exI) apply (auto simp: tcb_in_cur_domain'_def obj_at'_def projectKOs objBits_simps) done lemma attemptSwitchTo_corres: "corres dc (valid_etcbs and weak_valid_sched_action and cur_tcb and st_tcb_at runnable t) (Invariants_H.valid_queues and valid_queues' and (\s. weak_sch_act_wf (ksSchedulerAction s) s) and cur_tcb' and tcb_at' t and st_tcb_at' runnable' t and valid_objs') (attempt_switch_to t) (attemptSwitchTo t)" using possibleSwitchTo_corres apply (simp add: attempt_switch_to_def attemptSwitchTo_def) done lemma switchIfRequiredTo_corres: "corres dc (valid_etcbs and weak_valid_sched_action and cur_tcb and st_tcb_at runnable t) (Invariants_H.valid_queues and valid_queues' and (\s. weak_sch_act_wf (ksSchedulerAction s) s) and cur_tcb' and tcb_at' t and st_tcb_at' runnable' t and valid_objs') (switch_if_required_to t) (switchIfRequiredTo t)" using possibleSwitchTo_corres apply (simp add: switch_if_required_to_def switchIfRequiredTo_def) done end end