(* * Copyright 2014, General Dynamics C4 Systems * * SPDX-License-Identifier: GPL-2.0-only *) theory Orphanage imports Refine.Refine begin text \ Proof that calling the kernel never leaves threads orphaned. More specifically, every active thread must be the current thread, or about to be switched to, or be in a scheduling queue. \ (*FIXME: arch_split: move up? *) context Arch begin requalify_facts switchToIdleThread_def switchToThread_def lemmas [crunch_def] = switchToIdleThread_def switchToThread_def context begin global_naming global requalify_facts Thread_H.switchToIdleThread_def Thread_H.switchToThread_def end end context begin interpretation Arch . (*FIXME: arch_split*) definition is_active_thread_state :: "thread_state \ bool" where "is_active_thread_state ts \ isRunning ts \ isRestart ts" definition is_active_tcb_ptr :: "machine_word \ kernel_state \ bool" where "is_active_tcb_ptr tcb_ptr s \ st_tcb_at' is_active_thread_state tcb_ptr s" lemma is_active_tcb_ptr_runnable': "is_active_tcb_ptr t s = st_tcb_at' runnable' t s" by (auto simp: is_active_tcb_ptr_def pred_tcb_at'_def obj_at'_def is_active_thread_state_def isRunning_def isRestart_def split: Structures_H.thread_state.split_asm) definition all_active_tcb_ptrs :: "kernel_state \ machine_word set" where "all_active_tcb_ptrs s \ { tcb_ptr. is_active_tcb_ptr tcb_ptr s }" definition all_queued_tcb_ptrs :: "kernel_state \ machine_word set" where "all_queued_tcb_ptrs s \ { tcb_ptr. \ priority. tcb_ptr : set ((ksReadyQueues s) priority) }" lemma st_tcb_at_neg': "(st_tcb_at' (\ ts. \ P ts) t s) = (tcb_at' t s \ \ st_tcb_at' P t s)" by (auto simp: pred_tcb_at'_def obj_at'_def) lemma st_tcb_at_neg2: "(\ st_tcb_at' P t s) = (st_tcb_at' (\ ts. \ P ts) t s \ \ tcb_at' t s)" by (auto simp: pred_tcb_at'_def obj_at'_def) lemma st_tcb_at_double_neg': "(st_tcb_at' (\ ts. \ P ts \ \ Q ts) t s) = ((st_tcb_at' (\ ts. \ P ts) t s) \ (st_tcb_at' (\ ts. \ Q ts) t s))" apply (auto simp: pred_tcb_at'_def obj_at'_def) done definition no_orphans :: " kernel_state \ bool" where "no_orphans s \ \ tcb_ptr. (tcb_ptr : all_active_tcb_ptrs s \ tcb_ptr = ksCurThread s \ tcb_ptr : all_queued_tcb_ptrs s \ ksSchedulerAction s = SwitchToThread tcb_ptr)" lemma no_orphans_disj: "no_orphans = (\ s. \ tcb_ptr. tcb_ptr = ksCurThread s \ tcb_ptr : all_queued_tcb_ptrs s \ \ typ_at' TCBT tcb_ptr s \ st_tcb_at' (\ state. \ is_active_thread_state state) tcb_ptr s \ ksSchedulerAction s = SwitchToThread tcb_ptr)" apply clarsimp apply (rule ext) apply (unfold no_orphans_def all_active_tcb_ptrs_def is_active_tcb_ptr_def st_tcb_at_neg' typ_at_tcb') apply (auto intro: pred_tcb_at') done lemma no_orphans_lift: assumes typ_at'_is_lifted: "\ tcb_ptr. \ \s. \ typ_at' TCBT tcb_ptr s\ f \ \_ s. \ typ_at' TCBT tcb_ptr s \" assumes ksCurThread_is_lifted: "\ tcb_ptr. \ \s. tcb_ptr = ksCurThread s \ f \ \_ s. tcb_ptr = ksCurThread s \" assumes st_tcb_at'_is_lifted: "\P p. \ \s. st_tcb_at' P p s\ f \ \_ s. st_tcb_at' P p s \" assumes ksReadyQueues_is_lifted: "\P. \ \s. P (ksReadyQueues s)\ f \ \_ s. P (ksReadyQueues s) \" assumes ksSchedulerAction_is_lifted: "\P. \ \s. P (ksSchedulerAction s)\ f \ \_ s. P (ksSchedulerAction s) \" shows "\ \s. no_orphans s \ f \ \_ s. no_orphans s \" apply (unfold no_orphans_disj all_active_tcb_ptrs_def all_queued_tcb_ptrs_def) apply (wp hoare_vcg_all_lift hoare_vcg_disj_lift) apply (rule ksCurThread_is_lifted) apply (wp hoare_vcg_disj_lift) apply (rule ksReadyQueues_is_lifted) apply (wp hoare_vcg_disj_lift) apply (rule typ_at'_is_lifted) apply (wp hoare_vcg_disj_lift) apply (rule st_tcb_at'_is_lifted) apply (rule ksSchedulerAction_is_lifted) apply simp done lemma st_tcb_at'_is_active_tcb_ptr_lift: assumes "\P P' t. \\s. P (st_tcb_at' P' t s)\ f \\rv s. P (st_tcb_at' P' t s)\" shows "\\s. P (is_active_tcb_ptr t s)\ f \\_ s. P (is_active_tcb_ptr t s)\" by (clarsimp simp: is_active_tcb_ptr_def) (rule assms) lemma st_tcb_at'_all_active_tcb_ptrs_lift: assumes "\P P' t. \\s. P (st_tcb_at' P' t s)\ f \\rv s. P (st_tcb_at' P' t s)\" shows "\\s. P (t \ all_active_tcb_ptrs s)\ f \\_ s. P (t \ all_active_tcb_ptrs s)\" by (clarsimp simp: all_active_tcb_ptrs_def) (rule st_tcb_at'_is_active_tcb_ptr_lift [OF assms]) lemma ksQ_all_queued_tcb_ptrs_lift: assumes "\P p. \\s. P (ksReadyQueues s p)\ f \\rv s. P (ksReadyQueues s p)\" shows "\\s. P (t \ all_queued_tcb_ptrs s)\ f \\_ s. P (t \ all_queued_tcb_ptrs s)\" apply (clarsimp simp: all_queued_tcb_ptrs_def) apply (rule_tac P=P in P_bool_lift) apply (wp hoare_ex_wp assms) apply (clarsimp) apply (wp hoare_vcg_all_lift assms) done definition almost_no_orphans :: "word32 \ kernel_state \ bool" where "almost_no_orphans tcb_ptr s \ \ ptr. ptr = tcb_ptr \ (ptr : all_active_tcb_ptrs s \ ptr = ksCurThread s \ ptr : all_queued_tcb_ptrs s \ ksSchedulerAction s = SwitchToThread ptr)" lemma no_orphans_strg_almost: "no_orphans s \ almost_no_orphans tcb_ptr s" unfolding no_orphans_def almost_no_orphans_def apply simp done lemma almost_no_orphans_disj: "almost_no_orphans tcb_ptr = (\ s. \ ptr. ptr = ksCurThread s \ ptr : all_queued_tcb_ptrs s \ \ typ_at' TCBT ptr s \ st_tcb_at' (\ thread_state. \ is_active_thread_state thread_state) ptr s \ ptr = tcb_ptr \ ksSchedulerAction s = SwitchToThread ptr)" apply clarsimp apply (rule ext) apply (unfold almost_no_orphans_def all_active_tcb_ptrs_def is_active_tcb_ptr_def st_tcb_at_neg' typ_at_tcb') apply (auto intro: pred_tcb_at') done (****************************************************************************************************) crunch ksCurThread [wp]: setVMRoot "\ s. P (ksCurThread s)" (wp: crunch_wps simp: crunch_simps) crunch ksReadyQueues [wp]: asUser "\s. P (ksReadyQueues s)" (wp: crunch_wps simp: crunch_simps) crunch no_orphans [wp]: getCurThread "no_orphans" crunch no_orphans [wp]: threadGet "no_orphans" crunch no_orphans [wp]: getNotification "no_orphans" lemma no_orphans_ksReadyQueuesL1Bitmap_update[simp]: "no_orphans (s\ ksReadyQueuesL1Bitmap := x \) = no_orphans s" unfolding no_orphans_def all_active_tcb_ptrs_def all_queued_tcb_ptrs_def is_active_tcb_ptr_def by auto lemma no_orphans_ksReadyQueuesL2Bitmap_update[simp]: "no_orphans (s\ ksReadyQueuesL2Bitmap := x \) = no_orphans s" unfolding no_orphans_def all_active_tcb_ptrs_def all_queued_tcb_ptrs_def is_active_tcb_ptr_def by auto crunch no_orphans [wp]: addToBitmap "no_orphans" crunch no_orphans [wp]: removeFromBitmap "no_orphans" lemma almost_no_orphans_ksReadyQueuesL1Bitmap_update[simp]: "almost_no_orphans t (s\ ksReadyQueuesL1Bitmap := x \) = almost_no_orphans t s" unfolding almost_no_orphans_def all_active_tcb_ptrs_def all_queued_tcb_ptrs_def is_active_tcb_ptr_def by auto lemma almost_no_orphans_ksReadyQueuesL2Bitmap_update[simp]: "almost_no_orphans t (s\ ksReadyQueuesL2Bitmap := x \) = almost_no_orphans t s" unfolding almost_no_orphans_def all_active_tcb_ptrs_def all_queued_tcb_ptrs_def is_active_tcb_ptr_def by auto crunch almost_no_orphans [wp]: addToBitmap "almost_no_orphans x" crunch almost_no_orphans [wp]: removeFromBitmap "almost_no_orphans x" lemma setCTE_no_orphans [wp]: "\ \s. no_orphans s \ setCTE p cte \ \rv s. no_orphans s \" apply (rule no_orphans_lift) apply (wp setCTE_typ_at' setCTE_pred_tcb_at')+ done lemma setCTE_almost_no_orphans [wp]: "\ \s. almost_no_orphans tcb_ptr s \ setCTE p cte \ \rv s. almost_no_orphans tcb_ptr s \" unfolding almost_no_orphans_disj all_queued_tcb_ptrs_def apply (wp hoare_vcg_all_lift hoare_vcg_disj_lift setCTE_typ_at' setCTE_pred_tcb_at') done crunch no_orphans [wp]: activateIdleThread "no_orphans" lemma asUser_no_orphans [wp]: "\ \s. no_orphans s \ asUser thread f \ \rv s. no_orphans s \" unfolding no_orphans_disj all_queued_tcb_ptrs_def apply (wp hoare_vcg_all_lift hoare_vcg_disj_lift) done lemma threadSet_no_orphans: "\tcb. \ is_active_thread_state (tcbState tcb) \ \ is_active_thread_state (tcbState (F tcb)) \ \ \s. no_orphans s \ threadSet F tptr \ \rv s. no_orphans s \" unfolding no_orphans_disj all_queued_tcb_ptrs_def apply (wp hoare_vcg_all_lift hoare_vcg_disj_lift threadSet_st_tcb_at2 | clarsimp)+ done lemma threadSet_almost_no_orphans: "\tcb. \ is_active_thread_state (tcbState tcb) \ \ is_active_thread_state (tcbState (F tcb)) \ \ \s. almost_no_orphans ptr s \ threadSet F tptr \ \rv s. almost_no_orphans ptr s \" unfolding almost_no_orphans_disj all_queued_tcb_ptrs_def apply (wp hoare_vcg_all_lift hoare_vcg_disj_lift threadSet_st_tcb_at2 | clarsimp)+ done lemma all_active_tcb_ptrs_queue [simp]: "all_active_tcb_ptrs (ksReadyQueues_update f s) = all_active_tcb_ptrs s" by (clarsimp simp: all_active_tcb_ptrs_def is_active_tcb_ptr_def) lemma setQueue_no_orphans_enq: "\ \s. no_orphans s \ set (ksReadyQueues s (d, prio)) \ set qs \ setQueue d prio qs \ \_ s. no_orphans s \" unfolding setQueue_def apply wp apply (clarsimp simp: no_orphans_def all_queued_tcb_ptrs_def split: if_split_asm) apply fastforce done lemma setQueue_almost_no_orphans_enq: "\ \s. almost_no_orphans tcb_ptr s \ set (ksReadyQueues s (d, prio)) \ set qs \ tcb_ptr \ set qs \ setQueue d prio qs \ \_ s. no_orphans s \" unfolding setQueue_def apply wp apply (clarsimp simp: no_orphans_def almost_no_orphans_def all_queued_tcb_ptrs_def split: if_split_asm) apply fastforce done lemma setQueue_almost_no_orphans_enq_lift: "\ \s. almost_no_orphans tcb_ptr s \ set (ksReadyQueues s (d, prio)) \ set qs \ setQueue d prio qs \ \_ s. almost_no_orphans tcb_ptr s \" unfolding setQueue_def apply wp apply (clarsimp simp: almost_no_orphans_def all_queued_tcb_ptrs_def split: if_split_asm) apply fastforce done lemma tcbSchedEnqueue_no_orphans[wp]: "\ \s. no_orphans s \ tcbSchedEnqueue tcb_ptr \ \rv s. no_orphans s \" unfolding tcbSchedEnqueue_def apply (wp setQueue_no_orphans_enq threadSet_no_orphans | clarsimp simp: unless_def)+ apply (wp getObject_tcb_wp | clarsimp simp: threadGet_def)+ apply (drule obj_at_ko_at') apply auto done lemma tcbSchedAppend_no_orphans[wp]: "\ \s. no_orphans s \ tcbSchedAppend tcb_ptr \ \rv s. no_orphans s \" unfolding tcbSchedAppend_def apply (wp setQueue_no_orphans_enq threadSet_no_orphans | clarsimp simp: unless_def)+ apply (wp getObject_tcb_wp | clarsimp simp: threadGet_def)+ apply (drule obj_at_ko_at') apply auto done lemma ko_at_obj_at': "ko_at' ko p s \ P ko \ obj_at' P p s" unfolding obj_at'_def apply clarsimp done lemma queued_in_queue: "\valid_queues' s; ko_at' tcb tcb_ptr s; tcbQueued tcb\ \ \ p. tcb_ptr \ set (ksReadyQueues s p)" unfolding valid_queues'_def apply (drule_tac x="tcbDomain tcb" in spec) apply (drule_tac x="tcbPriority tcb" in spec) apply (drule_tac x="tcb_ptr" in spec) apply (drule mp) apply (rule ko_at_obj_at') apply (auto simp: inQ_def) done lemma tcbSchedEnqueue_almost_no_orphans: "\ \s. almost_no_orphans tcb_ptr s \ valid_queues' s \ tcbSchedEnqueue tcb_ptr \ \rv s. no_orphans s \" unfolding tcbSchedEnqueue_def apply simp apply (wp setQueue_almost_no_orphans_enq[where tcb_ptr=tcb_ptr] threadSet_no_orphans | clarsimp)+ apply (wp getObject_tcb_wp | clarsimp simp: threadGet_def)+ apply normalise_obj_at' apply (rule_tac x=ko in exI) apply (clarsimp simp: subset_insertI) apply (unfold no_orphans_def almost_no_orphans_def) apply clarsimp apply (drule(2) queued_in_queue) apply (fastforce simp: all_queued_tcb_ptrs_def) done lemma tcbSchedEnqueue_almost_no_orphans_lift: "\ \s. almost_no_orphans ptr s \ tcbSchedEnqueue tcb_ptr \ \rv s. almost_no_orphans ptr s \" unfolding tcbSchedEnqueue_def apply (wp setQueue_almost_no_orphans_enq_lift threadSet_almost_no_orphans | clarsimp simp: unless_def)+ apply (wp getObject_tcb_wp | clarsimp simp: threadGet_def)+ apply (drule obj_at_ko_at') apply auto done lemma ssa_no_orphans: "\ \s. no_orphans s \ (\t. sch_act_not t s \ t : all_queued_tcb_ptrs s \ ksCurThread s = t) \ setSchedulerAction sa \ \rv s. no_orphans s \" unfolding setSchedulerAction_def no_orphans_disj all_queued_tcb_ptrs_def apply wp apply auto done lemma ssa_almost_no_orphans: "\ \s. almost_no_orphans tcb_ptr s \ (\t. sch_act_not t s \ t : all_queued_tcb_ptrs s \ ksCurThread s = t) \ setSchedulerAction (SwitchToThread tcb_ptr) \ \rv s. no_orphans s \" unfolding setSchedulerAction_def no_orphans_disj almost_no_orphans_disj all_queued_tcb_ptrs_def apply wp apply auto done lemma ssa_almost_no_orphans_lift [wp]: "\ \s. almost_no_orphans tcb_ptr s \ (\t. sch_act_not t s \ t : all_queued_tcb_ptrs s \ ksCurThread s = t) \ setSchedulerAction sa \ \rv s. almost_no_orphans tcb_ptr s \" unfolding setSchedulerAction_def almost_no_orphans_disj all_queued_tcb_ptrs_def apply wp apply auto done lemma tcbSchedEnqueue_inQueue [wp]: "\ \s. valid_queues' s \ tcbSchedEnqueue tcb_ptr \ \rv s. tcb_ptr \ all_queued_tcb_ptrs s \" unfolding tcbSchedEnqueue_def all_queued_tcb_ptrs_def apply (wp | clarsimp simp: unless_def)+ apply (rule_tac Q="\rv. \" in hoare_post_imp) apply fastforce apply (wp getObject_tcb_wp | clarsimp simp: threadGet_def)+ apply (fastforce simp: obj_at'_def valid_queues'_def inQ_def) done lemma tcbSchedAppend_inQueue [wp]: "\ \s. valid_queues' s \ tcbSchedAppend tcb_ptr \ \rv s. tcb_ptr \ all_queued_tcb_ptrs s \" unfolding tcbSchedAppend_def all_queued_tcb_ptrs_def apply (wp | clarsimp simp: unless_def)+ apply (rule_tac Q="\rv. \" in hoare_post_imp) apply fastforce apply (wp getObject_tcb_wp | clarsimp simp: threadGet_def)+ apply (fastforce simp: obj_at'_def valid_queues'_def inQ_def) done lemma rescheduleRequired_no_orphans [wp]: "\ \s. no_orphans s \ valid_queues' s \ rescheduleRequired \ \rv s. no_orphans s \" unfolding rescheduleRequired_def apply (wp tcbSchedEnqueue_no_orphans hoare_vcg_all_lift ssa_no_orphans | wpc | clarsimp)+ apply (wps tcbSchedEnqueue_nosch, wp static_imp_wp) apply (rename_tac word t p) apply (rule_tac P="word = t" in hoare_gen_asm) apply (wp hoare_disjI1 | clarsimp)+ done lemma rescheduleRequired_almost_no_orphans [wp]: "\ \s. almost_no_orphans tcb_ptr s \ valid_queues' s \ rescheduleRequired \ \rv s. almost_no_orphans tcb_ptr s \" unfolding rescheduleRequired_def apply (wp tcbSchedEnqueue_almost_no_orphans_lift hoare_vcg_all_lift | wpc | clarsimp)+ apply (wps tcbSchedEnqueue_nosch, wp static_imp_wp) apply (rename_tac word t p) apply (rule_tac P="word = t" in hoare_gen_asm) apply (wp hoare_disjI1 | clarsimp)+ done lemma setThreadState_current_no_orphans: "\ \s. no_orphans s \ ksCurThread s = tcb_ptr \ valid_queues' s \ setThreadState state tcb_ptr \ \rv s. no_orphans s \" unfolding setThreadState_def apply (wp | clarsimp)+ apply (rule_tac Q="\rv s. valid_queues' s \ no_orphans s" in hoare_post_imp) apply clarsimp apply (wp threadSet_valid_queues') apply (unfold no_orphans_disj all_queued_tcb_ptrs_def) apply (wp hoare_vcg_all_lift hoare_vcg_disj_lift threadSet_pred_tcb_at_state) apply (auto simp: inQ_def) done lemma setThreadState_isRestart_no_orphans: "\ \s. no_orphans s \ st_tcb_at' isRestart tcb_ptr s \ valid_queues' s\ setThreadState state tcb_ptr \ \rv s. no_orphans s \" unfolding setThreadState_def apply (wp | clarsimp)+ apply (rule_tac Q="\rv s. valid_queues' s \ no_orphans s" in hoare_post_imp) apply clarsimp apply (wp threadSet_valid_queues') apply (unfold no_orphans_disj all_queued_tcb_ptrs_def is_active_thread_state_def) apply (wp hoare_vcg_all_lift hoare_vcg_disj_lift threadSet_pred_tcb_at_state) apply (auto simp: st_tcb_at_double_neg' st_tcb_at_neg' inQ_def) done lemma setThreadState_almost_no_orphans [wp]: "\ \s. no_orphans s \ valid_queues' s\ setThreadState state tcb_ptr \ \rv s. almost_no_orphans tcb_ptr s \" unfolding setThreadState_def apply (wp | clarsimp)+ apply (rule_tac Q="\rv s. valid_queues' s \ almost_no_orphans tcb_ptr s" in hoare_post_imp) apply clarsimp apply (wp threadSet_valid_queues') apply (unfold no_orphans_disj almost_no_orphans_disj all_queued_tcb_ptrs_def) apply (wp hoare_vcg_all_lift hoare_vcg_disj_lift threadSet_pred_tcb_at_state) apply (auto simp: inQ_def) done lemma setThreadState_not_active_no_orphans: "\ is_active_thread_state state \ \ \s. no_orphans s \ valid_queues' s \ setThreadState state tcb_ptr \ \rv s. no_orphans s \" unfolding setThreadState_def apply (wp | clarsimp)+ apply (rule_tac Q="\rv s. valid_queues' s \ no_orphans s" in hoare_post_imp) apply clarsimp apply (wp threadSet_valid_queues') apply (unfold no_orphans_disj all_queued_tcb_ptrs_def is_active_thread_state_def) apply (wp hoare_vcg_all_lift hoare_vcg_disj_lift threadSet_pred_tcb_at_state) apply (auto simp: isRunning_def isRestart_def inQ_def) done lemma setThreadState_not_active_almost_no_orphans: "\ is_active_thread_state state \ \ \s. almost_no_orphans thread s \ valid_queues' s \ setThreadState state tcb_ptr \ \rv s. almost_no_orphans thread s \" unfolding setThreadState_def apply (wp | clarsimp)+ apply (rule_tac Q="\rv s. valid_queues' s \ almost_no_orphans thread s" in hoare_post_imp) apply clarsimp apply (wp threadSet_valid_queues') apply (unfold almost_no_orphans_disj all_queued_tcb_ptrs_def is_active_thread_state_def) apply (wp hoare_vcg_all_lift hoare_vcg_disj_lift threadSet_pred_tcb_at_state) apply (auto simp: isRunning_def isRestart_def inQ_def) done lemma activateThread_no_orphans [wp]: "\ \s. no_orphans s \ ct_in_state' activatable' s \ invs' s \ activateThread \ \rv s. no_orphans s \" unfolding activateThread_def apply (wp gts_wp' setThreadState_isRestart_no_orphans | wpc | clarsimp)+ apply (auto simp: ct_in_state'_def pred_tcb_at'_def obj_at'_def isRestart_def) done lemma setQueue_no_orphans_deq: "\ \s. \ tcb_ptr. no_orphans s \ \ is_active_tcb_ptr tcb_ptr s \ queue = [x\((ksReadyQueues s) (d, priority)). x \ tcb_ptr] \ setQueue d priority queue \ \rv s. no_orphans s \" unfolding setQueue_def apply (wp | clarsimp)+ apply (fastforce simp: no_orphans_def all_queued_tcb_ptrs_def all_active_tcb_ptrs_def is_active_tcb_ptr_def) done lemma setQueue_almost_no_orphans_deq [wp]: "\ \s. almost_no_orphans tcb_ptr s \ queue = [x\((ksReadyQueues s) (d, priority)). x \ tcb_ptr] \ setQueue d priority queue \ \rv s. almost_no_orphans tcb_ptr s \" unfolding setQueue_def apply (wp | clarsimp)+ apply (fastforce simp: almost_no_orphans_def all_queued_tcb_ptrs_def all_active_tcb_ptrs_def is_active_tcb_ptr_def) done lemma tcbSchedDequeue_almost_no_orphans [wp]: "\ \s. no_orphans s \ tcbSchedDequeue thread \ \rv s. almost_no_orphans thread s \" unfolding tcbSchedDequeue_def apply (wp threadSet_almost_no_orphans | simp cong: if_cong)+ apply (simp add:no_orphans_strg_almost cong: if_cong) done lemma tcbSchedDequeue_no_orphans [wp]: "\ \s. no_orphans s \ \ is_active_tcb_ptr tcb_ptr s \ tcbSchedDequeue tcb_ptr \ \rv s. no_orphans s \" unfolding tcbSchedDequeue_def apply (wp setQueue_no_orphans_deq threadSet_no_orphans | clarsimp)+ apply (wp getObject_tcb_wp | clarsimp simp: threadGet_def)+ apply (drule obj_at_ko_at') apply auto done lemma switchToIdleThread_no_orphans' [wp]: "\ \s. no_orphans s \ (is_active_tcb_ptr (ksCurThread s) s \ ksCurThread s \ all_queued_tcb_ptrs s) \ switchToIdleThread \ \rv s. no_orphans s \" unfolding switchToIdleThread_def setCurThread_def ARM_H.switchToIdleThread_def apply (simp add: no_orphans_disj all_queued_tcb_ptrs_def) apply (wp hoare_vcg_all_lift hoare_vcg_imp_lift hoare_vcg_disj_lift | clarsimp)+ apply (auto simp: no_orphans_disj all_queued_tcb_ptrs_def is_active_tcb_ptr_def st_tcb_at_neg' tcb_at_typ_at') done lemma ct_in_state_ksSched [simp]: "ct_in_state' activatable' (ksSchedulerAction_update f s) = ct_in_state' activatable' s" unfolding ct_in_state'_def apply auto done lemma no_orphans_ksIdle [simp]: "no_orphans (ksIdleThread_update f s) = no_orphans s" unfolding no_orphans_def all_active_tcb_ptrs_def all_queued_tcb_ptrs_def is_active_tcb_ptr_def apply auto done crunch no_orphans [wp]: "Arch.switchToThread" "no_orphans" (wp: no_orphans_lift ignore: ARM.clearExMonitor) crunch ksCurThread [wp]: "Arch.switchToThread" "\ s. P (ksCurThread s)" (ignore: ARM.clearExMonitor) crunch ksIdleThread [wp]: "Arch.switchToThread" "\ s. P (ksIdleThread s)" (ignore: ARM.clearExMonitor) lemma ArchThreadDecls_H_switchToThread_all_queued_tcb_ptrs [wp]: "\ \s. P (all_queued_tcb_ptrs s) \ Arch.switchToThread tcb_ptr \ \rv s. P (all_queued_tcb_ptrs s) \" unfolding ARM_H.switchToThread_def all_queued_tcb_ptrs_def apply (wp | clarsimp)+ done crunch ksSchedulerAction [wp]: "Arch.switchToThread" "\s. P (ksSchedulerAction s)" (ignore: ARM.clearExMonitor) lemma setCurThread_no_orphans [wp]: "\ \s. no_orphans s \ (is_active_tcb_ptr (ksCurThread s) s \ ksCurThread s : all_queued_tcb_ptrs s) \ setCurThread newThread \ \rv s. no_orphans s \" unfolding setCurThread_def apply (wp | clarsimp)+ apply (unfold no_orphans_def all_queued_tcb_ptrs_def all_active_tcb_ptrs_def is_active_tcb_ptr_def) apply auto done lemma tcbSchedDequeue_all_queued_tcb_ptrs: "\\s. x \ all_queued_tcb_ptrs s \ x \ t \ tcbSchedDequeue t \\_ s. x \ all_queued_tcb_ptrs s\" apply (rule_tac Q="(\s. x \ all_queued_tcb_ptrs s) and K (x \ t)" in hoare_pre_imp, clarsimp) apply (rule hoare_gen_asm) apply (clarsimp simp: tcbSchedDequeue_def all_queued_tcb_ptrs_def) apply (rule hoare_pre) apply (wp, clarsimp) apply (wp hoare_ex_wp)+ apply (rename_tac d p) apply (rule_tac Q="\_ s. x \ set (ksReadyQueues s (d, p))" in hoare_post_imp, clarsimp) apply (wp hoare_vcg_all_lift | simp)+ done lemma tcbSchedDequeue_all_active_tcb_ptrs[wp]: "\\s. P (t' \ all_active_tcb_ptrs s)\ tcbSchedDequeue t \\_ s. P (t' \ all_active_tcb_ptrs s)\" by (clarsimp simp: all_active_tcb_ptrs_def is_active_tcb_ptr_def) wp lemma setCurThread_almost_no_orphans: "\\s. almost_no_orphans t s \ (ksCurThread s \ t \ ksCurThread s \ all_active_tcb_ptrs s \ ksCurThread s \ all_queued_tcb_ptrs s)\ setCurThread t \\_. no_orphans\" unfolding setCurThread_def apply wp apply (fastforce simp: almost_no_orphans_def no_orphans_def all_queued_tcb_ptrs_def all_active_tcb_ptrs_def is_active_tcb_ptr_def) done lemmas ArchThreadDecls_H_switchToThread_all_active_tcb_ptrs[wp] = st_tcb_at'_all_active_tcb_ptrs_lift [OF Arch_switchToThread_pred_tcb'] lemmas ArchThreadDecls_H_switchToThread_all_queued_tcb_ptrs_lift[wp] = ksQ_all_queued_tcb_ptrs_lift [OF arch_switch_thread_ksQ] lemma ThreadDecls_H_switchToThread_no_orphans: "\ \s. no_orphans s \ st_tcb_at' runnable' tcb_ptr s \ (ksCurThread s \ all_active_tcb_ptrs s \ ksCurThread s \ all_queued_tcb_ptrs s)\ ThreadDecls_H.switchToThread tcb_ptr \ \rv s. no_orphans s \" unfolding Thread_H.switchToThread_def apply (wp setCurThread_almost_no_orphans tcbSchedDequeue_almost_no_orphans) apply (wps tcbSchedDequeue_ct') apply (wp tcbSchedDequeue_all_queued_tcb_ptrs hoare_convert_imp)+ apply (wps) apply (wp)+ apply (wps) apply (wp) apply (clarsimp) done lemma findM_failure': "\ \x S. \ \s. P S s \ f x \ \rv s. \ rv \ P (insert x S) s \ \ \ \ \s. P S s \ findM f xs \ \rv s. rv = None \ P (S \ set xs) s \" apply (induct xs arbitrary: S) apply (clarsimp, wp, clarsimp) apply clarsimp apply (rule hoare_seq_ext[rotated], assumption) apply (case_tac r) apply (clarsimp, wp, clarsimp) apply clarsimp apply (rule hoare_strengthen_post, assumption) apply clarsimp done lemmas findM_failure = findM_failure'[where S="{}", simplified] lemma tcbSchedEnqueue_inQueue_eq: "\ valid_queues' and K (tcb_ptr = tcb_ptr') \ tcbSchedEnqueue tcb_ptr \ \rv s. tcb_ptr' \ all_queued_tcb_ptrs s \" apply (rule hoare_gen_asm, simp) apply wp done lemma tcbSchedAppend_inQueue_eq: "\ valid_queues' and K (tcb_ptr = tcb_ptr') \ tcbSchedAppend tcb_ptr \ \rv s. tcb_ptr' \ all_queued_tcb_ptrs s \" apply (rule hoare_gen_asm, simp) apply wp done lemma findM_on_success: "\ \x. \ P x \ f x \ \rv s. rv \; \x y. \ P x \ f y \ \rv. P x \ \ \ \ \s. \x \ set xs. P x s \ findM f xs \ \rv s. \ y. rv = Some y \" apply (induct xs; clarsimp) apply wp+ apply (clarsimp simp: imp_conv_disj Bex_def) apply (wp hoare_vcg_disj_lift hoare_ex_wp | clarsimp | assumption)+ done crunch st_tcb' [wp]: switchToThread "\s. P' (st_tcb_at' P t s)" (ignore: ARM.clearExMonitor) lemma setQueue_deq_not_empty: "\ \s. (\tcb. tcb \ set (ksReadyQueues s p) \ st_tcb_at' P tcb s) \ (\tcb_ptr. \ st_tcb_at' P tcb_ptr s \ queue = [x\((ksReadyQueues s) (d, priority)). x \ tcb_ptr]) \ setQueue d priority queue \ \rv s. \tcb. tcb \ set (ksReadyQueues s p) \ st_tcb_at' P tcb s \" unfolding setQueue_def apply wp apply auto done lemma tcbSchedDequeue_not_empty: "\ \s. (\tcb. tcb \ set (ksReadyQueues s p) \ st_tcb_at' P tcb s) \ \ st_tcb_at' P thread s \ tcbSchedDequeue thread \ \rv s. \tcb. tcb \ set (ksReadyQueues s p) \ st_tcb_at' P tcb s \" unfolding tcbSchedDequeue_def apply wp apply (wp hoare_ex_wp threadSet_pred_tcb_no_state) apply clarsimp apply (wp setQueue_deq_not_empty) apply clarsimp apply (rule hoare_pre_post, assumption) apply (clarsimp simp: bitmap_fun_defs) apply wp apply clarsimp apply clarsimp apply (wp setQueue_deq_not_empty)+ apply (rule_tac Q="\rv s. \ st_tcb_at' P thread s" in hoare_post_imp) apply fastforce apply (wp weak_if_wp | clarsimp)+ done lemmas switchToThread_all_active_tcb_ptrs[wp] = st_tcb_at'_all_active_tcb_ptrs_lift [OF switchToThread_st_tcb'] (* ksSchedulerAction s = ChooseNewThread *) lemma chooseThread_no_orphans [wp]: notes hoare_TrueI[simp] shows "\\s. no_orphans s \ all_invs_but_ct_idle_or_in_cur_domain' s \ (is_active_tcb_ptr (ksCurThread s) s \ ksCurThread s \ all_queued_tcb_ptrs s)\ chooseThread \ \rv s. no_orphans s \" (is "\?PRE\ _ \_\") unfolding chooseThread_def Let_def numDomains_def curDomain_def apply (simp only: return_bind, simp) apply (rule hoare_seq_ext[where B="\rv s. ?PRE s \ rv = ksCurDomain s"]) apply (rule_tac B="\rv s. ?PRE 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), simp) (* we have a thread to switch to *) apply (clarsimp simp: bitmap_fun_defs) apply (wp assert_inv ThreadDecls_H_switchToThread_no_orphans) apply (clarsimp simp: all_invs_but_ct_idle_or_in_cur_domain'_def valid_state'_def valid_queues_def st_tcb_at'_def) apply (fastforce dest!: lookupBitmapPriority_obj_at' elim: obj_at'_weaken simp: all_active_tcb_ptrs_def) apply (simp add: bitmap_fun_defs | wp)+ done lemma tcbSchedAppend_in_ksQ: "\valid_queues' and tcb_at' t\ tcbSchedAppend 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 (rule hoare_vcg_ex_lift)+ apply (simp add: tcbSchedAppend_def unless_def) apply wpsimp 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 clarsimp apply normalise_obj_at' apply (drule(1) valid_queues'_ko_atD, simp+) done lemma hoare_neg_imps: "\P\ f \\ rv s. \ R rv s\ \ \P\ f \\r s. R r s \ Q r s\" by (auto simp: valid_def) lemma setCurThread_ct [wp]: "\ \ \ setCurThread tcb_ptr \ \rv s. ksCurThread s = tcb_ptr \" unfolding setCurThread_def apply (wp | clarsimp)+ done lemma ThreadDecls_H_switchToThread_ct [wp]: "\ \ \ switchToThread tcb_ptr \ \rv s. ksCurThread s = tcb_ptr \" unfolding switchToThread_def apply (wp | clarsimp)+ done crunch no_orphans [wp]: nextDomain no_orphans (wp: no_orphans_lift simp: Let_def) crunch ksQ [wp]: nextDomain "\s. P (ksReadyQueues s p)" (simp: Let_def) crunch st_tcb_at' [wp]: nextDomain "\s. P (st_tcb_at' P' p s)" (simp: Let_def) crunch ct' [wp]: nextDomain "\s. P (ksCurThread s)" (simp: Let_def) crunch sch_act_not [wp]: nextDomain "sch_act_not t" (simp: Let_def) lemma tcbSchedEnqueue_in_ksQ': "\valid_queues' and tcb_at' t and K (t = t')\ tcbSchedEnqueue t' \\r s. \domain priority. t \ set (ksReadyQueues s (domain, priority))\" apply (rule hoare_gen_asm) apply (wp tcbSchedEnqueue_in_ksQ | clarsimp)+ done lemma all_invs_but_ct_idle_or_in_cur_domain'_strg: "invs' s \ all_invs_but_ct_idle_or_in_cur_domain' s" by (clarsimp simp: invs'_to_invs_no_cicd'_def) lemma setSchedulerAction_cnt_sch_act_not[wp]: "\ \ \ setSchedulerAction ChooseNewThread \\rv s. sch_act_not x s\" by (rule hoare_pre, rule hoare_strengthen_post[OF setSchedulerAction_direct]) auto lemma obj_at'_static_fix: "\ obj_at' (\(ko::'a::pspace_storable). True) p s ; P \ \ obj_at' (\(ko::'a::pspace_storable). P) p s" by (erule obj_at'_weakenE, simp) lemma tcbSchedEnqueue_in_ksQ_aqtp[wp]: "\valid_queues' and tcb_at' t\ tcbSchedEnqueue t \\r s. t \ all_queued_tcb_ptrs s\" apply (clarsimp simp: all_queued_tcb_ptrs_def) apply (rule tcbSchedEnqueue_in_ksQ) done crunch ksReadyQueues[wp]: threadGet "\s. P (ksReadyQueues s)" lemma tcbSchedEnqueue_in_ksQ_already_queued: "\\s. valid_queues' s \ tcb_at' t s \ (\domain priority. t' \ set (ksReadyQueues s (domain, priority))) \ tcbSchedEnqueue t \\r s. \domain priority. t' \ set (ksReadyQueues s (domain, priority))\" apply (case_tac "t'=t", wpsimp wp: tcbSchedEnqueue_in_ksQ) apply (wpsimp simp: tcbSchedEnqueue_def unless_def) apply (rule_tac Q="\_ s. \domain priority. t' \ set (ksReadyQueues s (domain, priority))" in hoare_post_imp) apply metis apply wpsimp+ done lemma tcbSchedAppend_in_ksQ_already_queued: "\\s. valid_queues' s \ tcb_at' t s \ (\domain priority. t' \ set (ksReadyQueues s (domain, priority))) \ tcbSchedAppend t \\r s. \domain priority. t' \ set (ksReadyQueues s (domain, priority))\" apply (case_tac "t'=t", wpsimp wp: tcbSchedAppend_in_ksQ) apply (wpsimp simp: tcbSchedAppend_def unless_def) apply (rule_tac Q="\_ s. \domain priority. t' \ set (ksReadyQueues s (domain, priority))" in hoare_post_imp) apply metis apply wpsimp+ done lemma tcbSchedEnqueue_in_ksQ'': "\\s. valid_queues' s \ tcb_at' t s \ (t' \ t \ (\domain priority. t' \ set (ksReadyQueues s (domain, priority)))) \ tcbSchedEnqueue t \\r s. \domain priority. t' \ set (ksReadyQueues s (domain, priority))\" apply (case_tac "t'=t", wpsimp wp: tcbSchedEnqueue_in_ksQ) apply clarsimp apply (wpsimp simp: tcbSchedEnqueue_def unless_def) apply (rule_tac Q="\_ s. \domain priority. t' \ set (ksReadyQueues s (domain, priority))" in hoare_post_imp) apply metis apply wpsimp+ done lemma tcbSchedAppend_in_ksQ'': "\\s. valid_queues' s \ tcb_at' t s \ (t' \ t \ (\domain priority. t' \ set (ksReadyQueues s (domain, priority)))) \ tcbSchedAppend t \\r s. \domain priority. t' \ set (ksReadyQueues s (domain, priority))\" apply (case_tac "t'=t", wpsimp wp: tcbSchedAppend_in_ksQ) apply clarsimp apply (wpsimp simp: tcbSchedAppend_def unless_def) apply (rule_tac Q="\_ s. \domain priority. t' \ set (ksReadyQueues s (domain, priority))" in hoare_post_imp) apply metis apply wpsimp+ done crunches setSchedulerAction for pred_tcb_at': "\s. P (pred_tcb_at' proj Q t s)" and ct': "\s. P (ksCurThread s)" (wp_del: ssa_wp) lemmas ssa_st_tcb_at'_ksCurThread[wp] = hoare_lift_Pf2[where f=ksCurThread, OF setSchedulerAction_pred_tcb_at' setSchedulerAction_ct'] lemma ct_active_st_tcb_at': "ct_active' s = st_tcb_at' runnable' (ksCurThread s) s" apply (rule iffI) apply (drule ct_active_runnable') apply (simp add: ct_in_state'_def) apply (clarsimp simp: ct_in_state'_def) apply (erule pred_tcb'_weakenE) apply (case_tac st, auto) done lemma tcbSchedEnqueue_in_ksQ_already_queued_aqtp: "\\s. valid_queues' s \ tcb_at' t s \ t' \ all_queued_tcb_ptrs s \ tcbSchedEnqueue t \\r s. t' \ all_queued_tcb_ptrs s \" by (clarsimp simp: all_queued_tcb_ptrs_def tcbSchedEnqueue_in_ksQ_already_queued) (* FIXME move *) lemma invs_switchToThread_runnable': "\ invs' s ; ksSchedulerAction s = SwitchToThread t \ \ st_tcb_at' runnable' t s" by (simp add: invs'_def valid_state'_def) (* for shoving pred_tcb_at' through hoare_vcg_imp_lift for tcbs we know are there *) lemma not_pred_tcb_at'I: "\ pred_tcb_at' f (Not \ P) t s ; tcb_at' t s \ \ \ pred_tcb_at' f P t s" by (subst (asm) pred_tcb_at'_Not, blast) lemma in_all_active_tcb_ptrsD: "t \ all_active_tcb_ptrs s \ st_tcb_at' runnable' t s" unfolding all_active_tcb_ptrs_def is_active_tcb_ptr_def is_active_thread_state_def isRunning_def isRestart_def apply clarsimp apply (erule pred_tcb'_weakenE) apply (case_tac st; clarsimp) done lemma scheduleChooseNewThread_no_orphans: "\ invs' and no_orphans and (\s. ksSchedulerAction s = ChooseNewThread \ (st_tcb_at' runnable' (ksCurThread s) s \ (\d p. ksCurThread s \ set (ksReadyQueues s (d, p))))) \ scheduleChooseNewThread \\_. no_orphans \" unfolding scheduleChooseNewThread_def apply (wp add: ssa_no_orphans hoare_vcg_all_lift) apply (wp hoare_disjI1 chooseThread_nosch)+ apply (wp nextDomain_invs_no_cicd' hoare_vcg_imp_lift hoare_lift_Pf2 [OF ksQ_all_queued_tcb_ptrs_lift[OF nextDomain_ksQ] nextDomain_ct'] hoare_lift_Pf2 [OF st_tcb_at'_is_active_tcb_ptr_lift[OF nextDomain_st_tcb_at'] nextDomain_ct'] hoare_vcg_all_lift getDomainTime_wp)[2] apply (wpsimp simp: if_apply_def2 invs'_invs_no_cicd all_queued_tcb_ptrs_def is_active_tcb_ptr_runnable')+ done lemma schedule_no_orphans[wp]: notes ssa_wp[wp del] shows "\ \s. no_orphans s \ invs' s \ schedule \ \rv s. no_orphans s \" proof - have do_switch_to: "\candidate. \\s. no_orphans s \ ksSchedulerAction s = SwitchToThread candidate \ st_tcb_at' runnable' candidate s \ (st_tcb_at' runnable' (ksCurThread s) s \ (\d p. ksCurThread s \ set (ksReadyQueues s (d, p)))) \ do ThreadDecls_H.switchToThread candidate; setSchedulerAction ResumeCurrentThread od \\rv. no_orphans\" apply (wpsimp wp: scheduleChooseNewThread_no_orphans ssa_no_orphans hoare_vcg_all_lift ThreadDecls_H_switchToThread_no_orphans)+ apply (rule_tac Q="\_ s. (t = candidate \ ksCurThread s = candidate) \ (t \ candidate \ sch_act_not t s)" in hoare_post_imp) apply (wpsimp wp: stt_nosch static_imp_wp)+ apply (fastforce dest!: in_all_active_tcb_ptrsD simp: all_queued_tcb_ptrs_def comp_def) done have abort_switch_to_enq: "\candidate. \\s. no_orphans s \ invs' s \ valid_queues' s \ ksSchedulerAction s = SwitchToThread candidate \ (st_tcb_at' runnable' (ksCurThread s) s \ (\d p. ksCurThread s \ set (ksReadyQueues s (d, p)))) \ do tcbSchedEnqueue candidate; setSchedulerAction ChooseNewThread; scheduleChooseNewThread od \\rv. no_orphans\" apply (rule hoare_pre) apply (wp scheduleChooseNewThread_no_orphans ssa_no_orphans setSchedulerAction_direct) apply (wpsimp wp: hoare_vcg_imp_lift' hoare_vcg_ex_lift simp: is_active_tcb_ptr_runnable' all_queued_tcb_ptrs_def | rule hoare_lift_Pf2[where f=ksCurThread, OF setSchedulerAction_ksQ])+ apply (wp tcbSchedEnqueue_in_ksQ' tcbSchedEnqueue_no_orphans hoare_vcg_all_lift hoare_vcg_imp_lift' hoare_vcg_disj_lift) apply (wp hoare_lift_Pf2[where f=ksCurThread, OF tcbSchedEnqueue_pred_tcb_at'] hoare_lift_Pf2[where f=ksCurThread, OF tcbSchedEnqueue_in_ksQ_already_queued] tcbSchedEnqueue_no_orphans | strengthen not_pred_tcb_at'_strengthen | wp (once) hoare_vcg_imp_lift')+ apply (clarsimp) apply (frule invs_sch_act_wf', clarsimp simp: pred_tcb_at') apply (simp add: st_tcb_at_neg' tcb_at_invs') done have abort_switch_to_app: "\candidate. \\s. no_orphans s \ invs' s \ valid_queues' s \ ksSchedulerAction s = SwitchToThread candidate \ (st_tcb_at' runnable' (ksCurThread s) s \ (\d p. ksCurThread s \ set (ksReadyQueues s (d, p))) ) \ do tcbSchedAppend candidate; setSchedulerAction ChooseNewThread; scheduleChooseNewThread od \\rv. no_orphans\" apply (rule hoare_pre) apply (wp scheduleChooseNewThread_no_orphans ssa_no_orphans setSchedulerAction_direct) apply (wpsimp wp: hoare_vcg_imp_lift' hoare_vcg_ex_lift simp: is_active_tcb_ptr_runnable' all_queued_tcb_ptrs_def | rule hoare_lift_Pf2[where f=ksCurThread, OF setSchedulerAction_ksQ])+ apply (wp tcbSchedAppend_in_ksQ'' tcbSchedAppend_no_orphans hoare_vcg_all_lift hoare_vcg_imp_lift' hoare_vcg_disj_lift) apply (wp hoare_lift_Pf2[where f=ksCurThread, OF tcbSchedAppend_pred_tcb_at'] hoare_lift_Pf2[where f=ksCurThread, OF tcbSchedAppend_in_ksQ_already_queued] tcbSchedAppend_no_orphans | strengthen not_pred_tcb_at'_strengthen | wp (once) hoare_vcg_imp_lift')+ apply (clarsimp) apply (frule invs_sch_act_wf', clarsimp simp: pred_tcb_at') apply (simp add: st_tcb_at_neg' tcb_at_invs') done show ?thesis unfolding schedule_def supply if_weak_cong[cong] apply (wp, wpc) \ \action = ResumeCurrentThread\ apply (wp)[1] \ \action = ChooseNewThread\ apply (clarsimp simp: when_def scheduleChooseNewThread_def) apply (wp ssa_no_orphans hoare_vcg_all_lift) apply (wp hoare_disjI1 chooseThread_nosch) apply (wp nextDomain_invs_no_cicd' hoare_vcg_imp_lift hoare_lift_Pf2 [OF ksQ_all_queued_tcb_ptrs_lift [OF nextDomain_ksQ] nextDomain_ct'] hoare_lift_Pf2 [OF st_tcb_at'_is_active_tcb_ptr_lift [OF nextDomain_st_tcb_at'] nextDomain_ct'] hoare_vcg_all_lift getDomainTime_wp)[2] apply ((wp tcbSchedEnqueue_no_orphans tcbSchedEnqueue_in_ksQ' hoare_drop_imp | clarsimp simp: all_queued_tcb_ptrs_def | strengthen all_invs_but_ct_idle_or_in_cur_domain'_strg | wps tcbSchedEnqueue_ct')+)[1] apply ((wp tcbSchedEnqueue_no_orphans tcbSchedEnqueue_in_ksQ' hoare_drop_imp | clarsimp simp: all_queued_tcb_ptrs_def | strengthen all_invs_but_ct_idle_or_in_cur_domain'_strg | wps tcbSchedEnqueue_ct')+)[1] apply wp[1] \ \action = SwitchToThread candidate\ apply (clarsimp) apply (rename_tac candidate) apply (wpsimp wp: do_switch_to abort_switch_to_enq abort_switch_to_app) (* isHighestPrio *) apply (wp hoare_drop_imps) apply (wp add: tcbSchedEnqueue_no_orphans)+ apply (clarsimp simp: conj_comms cong: conj_cong imp_cong split del: if_split) apply (wp hoare_lift_Pf2[where f=ksCurThread, OF tcbSchedEnqueue_pred_tcb_at'] hoare_lift_Pf2[where f=ksCurThread, OF tcbSchedEnqueue_in_ksQ'] hoare_vcg_imp_lift' | strengthen not_pred_tcb_at'_strengthen)+ apply (clarsimp simp: comp_def) apply (frule invs_queues) apply (clarsimp simp: invs_valid_queues' tcb_at_invs' st_tcb_at_neg' is_active_tcb_ptr_runnable') apply (fastforce simp: all_invs_but_ct_idle_or_in_cur_domain'_strg invs_switchToThread_runnable') done qed lemma setNotification_no_orphans [wp]: "\ \s. no_orphans s \ setNotification p ntfn \ \_ s. no_orphans s \" apply (rule no_orphans_lift) apply (wp | clarsimp simp: setNotification_def updateObject_default_def)+ done crunches doMachineOp, setMessageInfo for no_orphans [wp]: "no_orphans" (wp: no_orphans_lift) crunches completeSignal for no_orphans [wp]: "no_orphans" (simp: crunch_simps wp: crunch_wps) lemma possibleSwitchTo_almost_no_orphans [wp]: "\ \s. almost_no_orphans target s \ valid_queues' s \ st_tcb_at' runnable' target s \ weak_sch_act_wf (ksSchedulerAction s) s \ possibleSwitchTo target \ \rv s. no_orphans s \" unfolding possibleSwitchTo_def by (wp rescheduleRequired_valid_queues'_weak tcbSchedEnqueue_almost_no_orphans ssa_almost_no_orphans static_imp_wp | wpc | clarsimp | wp (once) hoare_drop_imp)+ lemma possibleSwitchTo_almost_no_orphans': "\ \s. almost_no_orphans target s \ valid_queues' s \ st_tcb_at' runnable' target s \ sch_act_wf (ksSchedulerAction s) s \ possibleSwitchTo target \ \rv s. no_orphans s \" by wp (strengthen sch_act_wf_weak, assumption) lemma tcbSchedAppend_almost_no_orphans: "\ \s. almost_no_orphans thread s \ valid_queues' s \ tcbSchedAppend thread \ \_ s. no_orphans s \" unfolding tcbSchedAppend_def apply (wp setQueue_almost_no_orphans_enq[where tcb_ptr=thread] threadSet_no_orphans | clarsimp simp: unless_def | simp only: subset_insertI)+ apply (unfold threadGet_def) apply (wp getObject_tcb_wp | clarsimp)+ apply (drule obj_at_ko_at', clarsimp) apply (rule_tac x=ko in exI) apply (clarsimp simp: almost_no_orphans_def no_orphans_def) apply (drule queued_in_queue | simp)+ apply (auto simp: all_queued_tcb_ptrs_def) done lemma no_orphans_is_almost[simp]: "no_orphans s \ almost_no_orphans t s" by (clarsimp simp: no_orphans_def almost_no_orphans_def) crunches decDomainTime for no_orphans [wp]: no_orphans and valid_queues' [wp]: valid_queues' (wp: no_orphans_lift) lemma timerTick_no_orphans [wp]: "\ \s. no_orphans s \ invs' s \ timerTick \ \_ s. no_orphans s \" unfolding timerTick_def getDomainTime_def numDomains_def apply (rule hoare_pre) apply (wp hoare_drop_imps | clarsimp)+ apply (wp threadSet_valid_queues' tcbSchedAppend_almost_no_orphans threadSet_almost_no_orphans threadSet_no_orphans tcbSchedAppend_sch_act_wf | wpc | clarsimp | strengthen sch_act_wf_weak)+ apply (rule_tac Q="\rv s. no_orphans s \ valid_queues' s \ tcb_at' thread s \ sch_act_wf (ksSchedulerAction s) s" in hoare_post_imp) apply (clarsimp simp: inQ_def) apply (wp hoare_drop_imps | clarsimp)+ apply auto done lemma handleDoubleFault_no_orphans [wp]: "\ \s. no_orphans s \ valid_queues' s \ handleDoubleFault tptr ex1 ex2 \ \rv s. no_orphans s \" unfolding handleDoubleFault_def apply (wp setThreadState_not_active_no_orphans | clarsimp simp: is_active_thread_state_def isRestart_def isRunning_def)+ done crunches cteInsert, getThreadCallerSlot, getThreadReplySlot for st_tcb' [wp]: "st_tcb_at' (\st. P st) t" and no_orphans [wp]: "no_orphans" and almost_no_orphans [wp]: "almost_no_orphans tcb_ptr" (wp: crunch_wps) lemma setupCallerCap_no_orphans [wp]: "\ \s. no_orphans s \ valid_queues' s \ setupCallerCap sender receiver gr \ \rv s. no_orphans s \" unfolding setupCallerCap_def apply (wp setThreadState_not_active_no_orphans | clarsimp simp: is_active_thread_state_def isRestart_def isRunning_def)+ done lemma setupCallerCap_almost_no_orphans [wp]: "\ \s. almost_no_orphans tcb_ptr s \ valid_queues' s \ setupCallerCap sender receiver gr \ \rv s. almost_no_orphans tcb_ptr s \" unfolding setupCallerCap_def apply (wp setThreadState_not_active_almost_no_orphans | clarsimp simp: is_active_thread_state_def isRestart_def isRunning_def)+ done crunches doIPCTransfer, setMRs, setEndpoint for ksReadyQueues [wp]: "\s. P (ksReadyQueues s)" and no_orphans [wp]: "no_orphans" (wp: transferCapsToSlots_pres1 crunch_wps no_orphans_lift setObject_queues_unchanged_tcb updateObject_default_inv) lemma sendIPC_no_orphans [wp]: "\ \s. no_orphans s \ valid_queues' s \ valid_objs' s \ sch_act_wf (ksSchedulerAction s) s \ sendIPC blocking call badge canGrant canGrantReply thread epptr \ \rv s. no_orphans s \" unfolding sendIPC_def apply (wp hoare_drop_imps setThreadState_not_active_no_orphans sts_st_tcb' possibleSwitchTo_almost_no_orphans' | wpc | clarsimp simp: is_active_thread_state_def isRestart_def isRunning_def)+ apply (rule_tac Q="\rv. no_orphans and valid_queues' and valid_objs' and ko_at' rv epptr and (\s. sch_act_wf (ksSchedulerAction s) s)" in hoare_post_imp) apply (fastforce simp: valid_objs'_def valid_obj'_def valid_ep'_def obj_at'_def projectKOs) apply (wp get_ep_sp' | clarsimp)+ done lemma sendFaultIPC_no_orphans [wp]: "\ \s. no_orphans s \ valid_queues' s \ valid_objs' s \ sch_act_wf (ksSchedulerAction s) s \ sendFaultIPC tptr fault \ \rv s. no_orphans s \" unfolding sendFaultIPC_def apply (rule hoare_pre) apply (wp threadSet_valid_queues' threadSet_no_orphans threadSet_valid_objs' threadSet_sch_act | wpc | clarsimp)+ apply (rule_tac Q'="\handlerCap s. no_orphans s \ valid_queues' s \ valid_objs' s \ sch_act_wf (ksSchedulerAction s) s" in hoare_post_imp_R) apply (wp | clarsimp simp: inQ_def valid_tcb'_def tcb_cte_cases_def)+ done lemma sendIPC_valid_queues' [wp]: "\ \s. valid_queues' s \ valid_objs' s \ sch_act_wf (ksSchedulerAction s) s \ sendIPC blocking call badge canGrant canGrantReply thread epptr \ \rv s. valid_queues' s \" unfolding sendIPC_def apply (wpsimp wp: hoare_drop_imps) apply (wpsimp | wp (once) sts_st_tcb')+ apply (rule_tac Q="\rv. valid_queues' and valid_objs' and ko_at' rv epptr and (\s. sch_act_wf (ksSchedulerAction s) s)" in hoare_post_imp) apply (clarsimp) apply (wp get_ep_sp' | clarsimp)+ done lemma sendFaultIPC_valid_queues' [wp]: "\ \s. valid_queues' s \ valid_objs' s \ sch_act_wf (ksSchedulerAction s) s \ sendFaultIPC tptr fault \ \rv s. valid_queues' s \" unfolding sendFaultIPC_def apply (rule hoare_pre) apply (wp threadSet_valid_queues' threadSet_valid_objs' threadSet_sch_act | wpc | clarsimp)+ apply (rule_tac Q'="\handlerCap s. valid_queues' s \ valid_objs' s \ sch_act_wf (ksSchedulerAction s) s" in hoare_post_imp_R) apply (wp | clarsimp simp: inQ_def valid_tcb'_def tcb_cte_cases_def)+ done lemma handleFault_no_orphans [wp]: "\ \s. no_orphans s \ valid_queues' s \ valid_objs' s \ sch_act_wf (ksSchedulerAction s) s \ handleFault tptr ex1 \ \rv s. no_orphans s \" unfolding handleFault_def apply (rule hoare_pre) apply (wp | clarsimp)+ done lemma replyFromKernel_no_orphans [wp]: "\ \s. no_orphans s \ replyFromKernel thread r \ \rv s. no_orphans s \" apply (cases r, simp_all add: replyFromKernel_def) apply wp done crunch ksSchedulerAction [wp]: setMessageInfo "\s. P (ksSchedulerAction s)" crunch ksCurThread [wp]: createNewCaps "\ s. P (ksCurThread s)" crunch ksReadyQueues [wp]: createNewCaps "\ s. P (ksReadyQueues s)" crunch inv [wp]: alignError "P" lemma createObjects_no_orphans [wp]: "\ \s. no_orphans s \ pspace_aligned' s \ pspace_no_overlap' ptr sz s \ pspace_distinct' s \ n \ 0 \ range_cover ptr sz (objBitsKO val + gbits) n \ \ case_option False (is_active_thread_state \ tcbState) (projectKO_opt val) \ createObjects ptr n val gbits \ \rv s. no_orphans s \" apply (clarsimp simp: no_orphans_def all_active_tcb_ptrs_def is_active_tcb_ptr_def all_queued_tcb_ptrs_def) apply (simp only: imp_conv_disj pred_tcb_at'_def createObjects_def) apply (wp hoare_vcg_all_lift hoare_vcg_disj_lift createObjects_orig_obj_at2') apply clarsimp apply (erule(1) impE) apply clarsimp apply (drule_tac x = x in spec) apply (erule impE) apply (clarsimp simp:obj_at'_def projectKOs split: option.splits) apply simp done lemma copyGlobalMappings_no_orphans [wp]: "\ \s. no_orphans s \ copyGlobalMappings newPD \ \rv s. no_orphans s \" unfolding no_orphans_disj all_queued_tcb_ptrs_def apply (wp hoare_vcg_all_lift hoare_vcg_disj_lift) done lemma no_orphans_update_simps[simp]: "no_orphans (gsCNodes_update f s) = no_orphans s" "no_orphans (gsUserPages_update g s) = no_orphans s" "no_orphans (gsUntypedZeroRanges_update h s) = no_orphans s" by (simp_all add: no_orphans_def all_active_tcb_ptrs_def is_active_tcb_ptr_def all_queued_tcb_ptrs_def) crunch no_orphans [wp]: insertNewCap "no_orphans" (wp: hoare_drop_imps) lemma createNewCaps_no_orphans: "\ (\s. no_orphans s \ pspace_aligned' s \ pspace_distinct' s \ pspace_no_overlap' ptr sz s \ (tp = APIObjectType CapTableObject \ us > 0)) and K (range_cover ptr sz (APIType_capBits tp us) n \ 0 < n) \ createNewCaps tp ptr n us d \ \rv s. no_orphans s \" apply (clarsimp simp: createNewCaps_def toAPIType_def split del: if_split cong: option.case_cong) apply (cases tp, simp_all split del: if_split) apply (rename_tac apiobject_type) apply (case_tac apiobject_type, simp_all) apply (wp mapM_x_wp' threadSet_no_orphans | clarsimp simp: is_active_thread_state_def makeObject_tcb projectKO_opt_tcb isRunning_def isRestart_def APIType_capBits_def Arch_createNewCaps_def objBits_if_dev split del: if_split | simp add: objBits_simps | fastforce simp:pageBits_def pteBits_def archObjSize_def ptBits_def pdBits_def pdeBits_def objBits_simps)+ done lemma createObject_no_orphans: "\pspace_no_overlap' ptr sz and pspace_aligned' and pspace_distinct' and cte_wp_at' (\cte. cteCap cte = (capability.UntypedCap d ptr sz idx)) cref and K (range_cover ptr sz (APIType_capBits tp us) (Suc 0)) and no_orphans\ RetypeDecls_H.createObject tp ptr us d \\xa. no_orphans\" apply (simp only: createObject_def ARM_H.createObject_def placeNewObject_def2) apply (wpsimp wp: createObjects'_wp_subst threadSet_no_orphans createObjects_no_orphans[where sz = sz] simp: placeNewObject_def2 placeNewDataObject_def projectKO_opt_tcb cte_wp_at_ctes_of projectKO_opt_ep is_active_thread_state_def makeObject_tcb pageBits_def unless_def projectKO_opt_tcb isRunning_def isRestart_def APIType_capBits_def objBits_simps split_del: if_split) apply (clarsimp simp: toAPIType_def APIType_capBits_def objBits_simps archObjSize_def pteBits_def ptBits_def pdeBits_def pdBits_def split: object_type.split_asm apiobject_type.split_asm) done lemma createNewObjects_no_orphans: "\\s. no_orphans s \ invs' s \ pspace_no_overlap' ptr sz s \ (\slot\set slots. cte_wp_at' (\c. cteCap c = capability.NullCap) slot s) \ cte_wp_at' (\cte. cteCap cte = UntypedCap d (ptr && ~~ mask sz) sz idx) cref s \ caps_no_overlap'' ptr sz s \ range_cover ptr sz (APIType_capBits tp us) (length slots) \ (tp = APIObjectType ArchTypes_H.CapTableObject \ us > 0) \ caps_overlap_reserved' {ptr..ptr + of_nat (length slots) * 2 ^ APIType_capBits tp us - 1} s \ slots \ [] \ distinct slots \ ptr \ 0\ createNewObjects tp cref slots ptr us d \ \rv s. no_orphans s \" apply (rule hoare_name_pre_state) apply clarsimp apply (rule hoare_pre) apply (rule createNewObjects_wp_helper) apply simp+ apply (simp add:insertNewCaps_def) apply wp apply (rule_tac P = "length caps = length slots" in hoare_gen_asm) apply (wp zipWithM_x_inv) apply simp apply (wp createNewCaps_no_orphans[where sz = sz] | clarsimp)+ apply (rule hoare_strengthen_post[OF createNewCaps_ret_len]) apply simp apply (clarsimp simp:invs_pspace_aligned' invs_valid_pspace' invs_pspace_distinct') apply (intro conjI) apply (erule range_cover.range_cover_n_less[where 'a=32, folded word_bits_def]) apply (clarsimp simp:cte_wp_at_ctes_of) apply (simp add:invs'_def valid_state'_def) apply (simp add: invs_ksCurDomain_maxDomain') done lemma ksMachineState_ksPSpace_upd_comm: "ksPSpace_update g (ksMachineState_update f s) = ksMachineState_update f (ksPSpace_update g s)" by simp lemma deleteObjects_no_orphans [wp]: "\ (\s. no_orphans s \ pspace_distinct' s) and K (is_aligned ptr bits) \ deleteObjects ptr bits \ \rv s. no_orphans s \" apply (rule hoare_gen_asm) apply (unfold deleteObjects_def2 doMachineOp_def split_def) apply (wp hoare_drop_imps | clarsimp)+ apply (clarsimp simp: no_orphans_def all_active_tcb_ptrs_def all_queued_tcb_ptrs_def is_active_tcb_ptr_def ksMachineState_ksPSpace_upd_comm) apply (drule_tac x=tcb_ptr in spec) apply (clarsimp simp: pred_tcb_at'_def obj_at_delete'[simplified field_simps] cong: if_cong) done lemma no_orphans_ksWorkUnits [simp]: "no_orphans (ksWorkUnitsCompleted_update f s) = no_orphans s" unfolding no_orphans_def all_active_tcb_ptrs_def all_queued_tcb_ptrs_def is_active_tcb_ptr_def apply auto done lemma no_orphans_irq_state_independent[intro!, simp]: "no_orphans (s \ksMachineState := ksMachineState s \ irq_state := f (irq_state (ksMachineState s)) \ \) = no_orphans s" by (simp add: no_orphans_def all_active_tcb_ptrs_def all_queued_tcb_ptrs_def is_active_tcb_ptr_def) add_upd_simps "no_orphans (gsUntypedZeroRanges_update f s)" declare upd_simps[simp] crunch no_orphans[wp]: updateFreeIndex "no_orphans" lemma resetUntypedCap_no_orphans [wp]: "\ (\s. no_orphans s \ pspace_distinct' s \ valid_objs' s) and cte_wp_at' (isUntypedCap o cteCap) slot\ resetUntypedCap slot \ \rv s. no_orphans s \" apply (simp add: resetUntypedCap_def) apply (rule hoare_pre) apply (wp mapME_x_inv_wp preemptionPoint_inv getSlotCap_wp hoare_drop_imps | simp add: unless_def split del: if_split)+ apply (clarsimp simp: cte_wp_at_ctes_of split del: if_split) apply (frule(1) cte_wp_at_valid_objs_valid_cap'[OF ctes_of_cte_wpD]) apply (clarsimp simp: isCap_simps valid_cap_simps' capAligned_def) done lemma invokeUntyped_no_orphans [wp]: "\ \s. no_orphans s \ invs' s \ valid_untyped_inv' ui s \ ct_active' s \ invokeUntyped ui \ \reply s. no_orphans s \" apply (rule hoare_pre, rule hoare_strengthen_post) apply (rule invokeUntyped_invs''[where Q=no_orphans]) apply (wp createNewCaps_no_orphans)+ apply (clarsimp simp: valid_pspace'_def) apply (intro conjI, simp_all)[1] apply (wp | simp)+ apply (cases ui, auto simp: cte_wp_at_ctes_of)[1] done lemma setInterruptState_no_orphans [wp]: "\ \s. no_orphans s \ setInterruptState a \ \rv s. no_orphans s \" unfolding no_orphans_disj all_queued_tcb_ptrs_def apply (wp hoare_vcg_all_lift hoare_vcg_disj_lift | clarsimp)+ done crunch no_orphans [wp]: emptySlot "no_orphans" lemma mapM_x_match: "\I and V xs\ mapM_x m xs \\rv. Q\ \ \I and V xs\ mapM_x m xs \\rv. Q\" by assumption lemma cancelAllIPC_no_orphans [wp]: "\ \s. no_orphans s \ valid_queues' s \ valid_objs' s \ cancelAllIPC epptr \ \rv s. no_orphans s \" unfolding cancelAllIPC_def apply (wp sts_valid_objs' set_ep_valid_objs' sts_st_tcb' hoare_vcg_const_Ball_lift tcbSchedEnqueue_almost_no_orphans | wpc | rule mapM_x_match, rename_tac list, rule_tac V="\_. valid_queues' and valid_objs'" and I="no_orphans and (\s. \t\set list. tcb_at' t s)" in mapM_x_inv_wp2 | clarsimp simp: valid_tcb_state'_def)+ apply (rule_tac Q="\rv. no_orphans and valid_objs' and valid_queues' and ko_at' rv epptr" in hoare_post_imp) apply (fastforce simp: valid_obj'_def valid_ep'_def obj_at'_def projectKOs) apply (wp get_ep_sp' | clarsimp)+ done lemma cancelAllSignals_no_orphans [wp]: "\ \s. no_orphans s \ valid_queues' s \ valid_objs' s \ cancelAllSignals ntfn \ \rv s. no_orphans s \" unfolding cancelAllSignals_def apply (wp sts_valid_objs' set_ntfn_valid_objs' sts_st_tcb' hoare_vcg_const_Ball_lift tcbSchedEnqueue_almost_no_orphans | wpc | clarsimp simp: valid_tcb_state'_def)+ apply (rename_tac list) apply (rule_tac V="\_. valid_queues' and valid_objs'" and I="no_orphans and (\s. \t\set list. tcb_at' t s)" in mapM_x_inv_wp2) apply simp apply (wp sts_valid_objs' set_ntfn_valid_objs' sts_st_tcb' hoare_vcg_const_Ball_lift tcbSchedEnqueue_almost_no_orphans| clarsimp simp: valid_tcb_state'_def)+ apply (rule_tac Q="\rv. no_orphans and valid_objs' and valid_queues' and ko_at' rv ntfn" in hoare_post_imp) apply (fastforce simp: valid_obj'_def valid_ntfn'_def obj_at'_def projectKOs) apply (wp get_ntfn_sp' | clarsimp)+ done crunch no_orphans[wp]: setBoundNotification "no_orphans" lemma unbindNotification_no_orphans[wp]: "\\s. no_orphans s\ unbindNotification t \ \rv s. no_orphans s\" unfolding unbindNotification_def apply (rule hoare_seq_ext[OF _ gbn_sp']) apply (case_tac ntfnPtr, simp_all, wp, simp) apply (rule hoare_seq_ext[OF _ get_ntfn_sp']) apply (wp | simp)+ done lemma unbindMaybeNotification_no_orphans[wp]: "\\s. no_orphans s\ unbindMaybeNotification a \ \rv s. no_orphans s\" unfolding unbindMaybeNotification_def by (wp getNotification_wp | simp | wpc)+ lemma finaliseCapTrue_standin_no_orphans [wp]: "\ \s. no_orphans s \ valid_queues' s \ valid_objs' s \ finaliseCapTrue_standin cap final \ \rv s. no_orphans s \" unfolding finaliseCapTrue_standin_def apply (rule hoare_pre) apply (wp | clarsimp simp: Let_def | wpc)+ done lemma cteDeleteOne_no_orphans [wp]: "\ \s. no_orphans s \ valid_queues' s \ valid_objs' s \ cteDeleteOne slot \ \rv s. no_orphans s \" unfolding cteDeleteOne_def apply (wp assert_inv isFinalCapability_inv weak_if_wp | clarsimp simp: unless_def)+ done crunch valid_objs' [wp]: getThreadReplySlot "valid_objs'" lemma cancelSignal_no_orphans [wp]: "\ \s. no_orphans s \ valid_queues' s \ valid_objs' s \ cancelSignal t ntfn \ \rv s. no_orphans s \" unfolding cancelSignal_def Let_def apply (rule hoare_pre) apply (wp hoare_drop_imps setThreadState_not_active_no_orphans | wpc | clarsimp simp: is_active_thread_state_def isRestart_def isRunning_def)+ done lemma cancelIPC_no_orphans [wp]: "\ \s. no_orphans s \ valid_queues' s \ valid_objs' s \ cancelIPC t \ \rv s. no_orphans s \" unfolding cancelIPC_def Let_def apply (rule hoare_pre) apply (wp setThreadState_not_active_no_orphans hoare_drop_imps weak_if_wp threadSet_valid_queues' threadSet_valid_objs' threadSet_no_orphans | wpc | clarsimp simp: is_active_thread_state_def isRestart_def isRunning_def inQ_def valid_tcb'_def tcb_cte_cases_def)+ done lemma asUser_almost_no_orphans: "\almost_no_orphans t\ asUser a f \\_. almost_no_orphans t\" unfolding almost_no_orphans_disj all_queued_tcb_ptrs_def apply (wp hoare_vcg_all_lift hoare_vcg_disj_lift) done crunch almost_no_orphans[wp]: asUser "almost_no_orphans t" (simp: almost_no_orphans_disj all_queued_tcb_ptrs_def wp: hoare_vcg_all_lift hoare_vcg_disj_lift crunch_wps) lemma sendSignal_no_orphans [wp]: "\ \s. no_orphans s \ valid_queues' s \ valid_objs' s \ sch_act_wf (ksSchedulerAction s) s\ sendSignal ntfnptr badge \ \_ s. no_orphans s \" unfolding sendSignal_def apply (rule hoare_pre) apply (wp sts_st_tcb' gts_wp' getNotification_wp asUser_almost_no_orphans cancelIPC_weak_sch_act_wf | wpc | clarsimp simp: sch_act_wf_weak)+ done lemma handleInterrupt_no_orphans [wp]: "\ \s. no_orphans s \ invs' s \ handleInterrupt irq \ \rv s. no_orphans s \" unfolding handleInterrupt_def apply (rule hoare_pre) apply (wp hoare_drop_imps hoare_vcg_all_lift getIRQState_inv | wpc | clarsimp simp: invs'_def valid_state'_def maskIrqSignal_def handleReservedIRQ_def)+ done lemma updateRestartPC_no_orphans[wp]: "\ \s. no_orphans s \ invs' s \ updateRestartPC t \ \rv s. no_orphans s \" by (wpsimp simp: updateRestartPC_def asUser_no_orphans) lemma updateRestartPC_valid_queues'[wp]: "\ \s. valid_queues' s \ updateRestartPC t \ \rv s. valid_queues' s \" unfolding updateRestartPC_def apply (rule asUser_valid_queues') done lemma updateRestartPC_no_orphans_invs'_valid_queues'[wp]: "\\s. no_orphans s \ invs' s \ valid_queues' s \ updateRestartPC t \\rv s. no_orphans s \ valid_queues' s \" by (wpsimp simp: updateRestartPC_def asUser_no_orphans) lemma suspend_no_orphans [wp]: "\ \s. no_orphans s \ invs' s \ sch_act_simple s \ tcb_at' t s \ suspend t \ \rv s. no_orphans s \" unfolding suspend_def apply (wp | clarsimp simp: unless_def | rule conjI)+ apply (clarsimp simp: is_active_tcb_ptr_def is_active_thread_state_def st_tcb_at_neg2) apply (wp setThreadState_not_active_no_orphans hoare_disjI1 setThreadState_st_tcb | clarsimp simp: is_active_thread_state_def isRunning_def isRestart_def)+ apply (wp hoare_drop_imp)+ apply auto done lemma storeHWASID_no_orphans [wp]: "\ \s. no_orphans s \ storeHWASID asid hw_asid \ \reply s. no_orphans s \" unfolding no_orphans_disj all_queued_tcb_ptrs_def apply (wp hoare_vcg_all_lift hoare_vcg_disj_lift) done lemma invalidateHWASIDEntry_no_orphans [wp]: "\ \s. no_orphans s \ invalidateHWASIDEntry hwASID \ \reply s. no_orphans s \" unfolding no_orphans_disj all_queued_tcb_ptrs_def apply (wp hoare_vcg_all_lift hoare_vcg_disj_lift) done lemma invalidateASID_no_orphans [wp]: "\ \s. no_orphans s \ invalidateASID asid \ \reply s. no_orphans s \" unfolding no_orphans_disj all_queued_tcb_ptrs_def apply (wp hoare_vcg_all_lift hoare_vcg_disj_lift) done lemma findFreeHWASID_no_orphans [wp]: "\ \s. no_orphans s \ findFreeHWASID \ \reply s. no_orphans s \" unfolding no_orphans_disj all_queued_tcb_ptrs_def apply (wp hoare_vcg_all_lift hoare_vcg_disj_lift) done crunch ksCurThread [wp]: invalidateASIDEntry "\ s. P (ksCurThread s)" crunch ksReadyQueues[wp]: invalidateASIDEntry "\s. P (ksReadyQueues s)" lemma invalidateASIDEntry_no_orphans [wp]: "\ \s. no_orphans s \ invalidateASIDEntry asid \ \rv s. no_orphans s \" unfolding no_orphans_disj all_queued_tcb_ptrs_def apply (wp hoare_vcg_all_lift hoare_vcg_disj_lift) done crunch no_orphans [wp]: flushSpace "no_orphans" lemma deleteASIDPool_no_orphans [wp]: "\ \s. no_orphans s \ deleteASIDPool asid pool \ \rv s. no_orphans s \" unfolding deleteASIDPool_def apply (wp | clarsimp)+ apply (rule_tac Q="\rv s. no_orphans s" in hoare_post_imp) apply (clarsimp simp: no_orphans_def all_queued_tcb_ptrs_def all_active_tcb_ptrs_def is_active_tcb_ptr_def) apply (wp mapM_wp_inv getObject_inv loadObject_default_inv | clarsimp)+ done lemma storePTE_no_orphans [wp]: "\ \s. no_orphans s \ storePTE ptr val \ \rv s. no_orphans s \" unfolding no_orphans_disj all_queued_tcb_ptrs_def apply (wp hoare_vcg_all_lift hoare_vcg_disj_lift) done lemma storePDE_no_orphans [wp]: "\ \s. no_orphans s \ storePDE ptr val \ \rv s. no_orphans s \" unfolding no_orphans_disj all_queued_tcb_ptrs_def apply (wp hoare_vcg_all_lift hoare_vcg_disj_lift) done crunches unmapPage for no_orphans [wp]: "no_orphans" (wp: crunch_wps) lemma flushTable_no_orphans [wp]: "\ \s. no_orphans s \ flushTable pd asid vptr \ \reply s. no_orphans s \" unfolding flushTable_def apply (wp hoare_drop_imps | wpc | clarsimp)+ done crunches unmapPageTable, prepareThreadDelete for no_orphans [wp]: "no_orphans" lemma setASIDPool_no_orphans [wp]: "\ \s. no_orphans s \ setObject p (ap :: asidpool) \ \rv s. no_orphans s \" unfolding no_orphans_disj all_queued_tcb_ptrs_def apply (wp hoare_vcg_all_lift hoare_vcg_disj_lift) done lemma deleteASID_no_orphans [wp]: "\ \s. no_orphans s \ deleteASID asid pd \ \rv s. no_orphans s \" unfolding deleteASID_def apply (wp getObject_inv loadObject_default_inv | wpc | clarsimp)+ done lemma arch_finaliseCap_no_orphans [wp]: "\ \s. no_orphans s \ Arch.finaliseCap cap fin \ \rv s. no_orphans s \" unfolding ARM_H.finaliseCap_def apply (wpsimp simp: isCap_simps) apply (safe; wpsimp) done lemma deletingIRQHandler_no_orphans [wp]: "\ \s. no_orphans s \ invs' s \ deletingIRQHandler irq \ \rv s. no_orphans s \" unfolding deletingIRQHandler_def apply (wp, auto) done lemma finaliseCap_no_orphans [wp]: "\ \s. no_orphans s \ invs' s \ sch_act_simple s \ valid_cap' cap s \ finaliseCap cap final flag \ \rv s. no_orphans s \" apply (simp add: finaliseCap_def Let_def cong: if_cong split del: if_split) apply (rule hoare_pre) apply (wp | clarsimp simp: o_def | wpc)+ apply (auto simp: valid_cap'_def dest!: isCapDs) done crunches cteSwap, capSwapForDelete for no_orphans [wp]: "no_orphans" declare withoutPreemption_lift [wp del] lemma no_orphans_finalise_prop_stuff: "no_cte_prop no_orphans = no_orphans" "finalise_prop_stuff no_orphans" by (simp_all add: no_cte_prop_def finalise_prop_stuff_def setCTE_no_orphans, simp_all add: no_orphans_def all_active_tcb_ptrs_def is_active_tcb_ptr_def all_queued_tcb_ptrs_def) lemma finaliseSlot_no_orphans [wp]: "\ \s. no_orphans s \ invs' s \ sch_act_simple s \ (\ e \ ex_cte_cap_to' slot s) \ finaliseSlot slot e \ \rv s. no_orphans s \" unfolding finaliseSlot_def apply (rule validE_valid, rule hoare_pre, rule hoare_post_impErr, rule use_spec) apply (rule finaliseSlot_invs'[where p=slot and slot=slot and Pr=no_orphans]) apply (simp_all add: no_orphans_finalise_prop_stuff) apply (wp | simp)+ apply (auto dest: cte_wp_at_valid_objs_valid_cap') done lemma cteDelete_no_orphans [wp]: "\ no_orphans and invs' and sch_act_simple and K ex \ cteDelete ptr ex \ \rv s. no_orphans s \" apply (rule hoare_gen_asm) apply (clarsimp simp: cteDelete_def whenE_def split_def) apply (rule hoare_pre, wp) apply clarsimp done crunch no_orphans [wp]: cteMove "no_orphans" (wp: crunch_wps) lemma cteRevoke_no_orphans [wp]: "\ \s. no_orphans s \ invs' s \ sch_act_simple s \ cteRevoke ptr \ \rv s. no_orphans s \" apply (rule_tac Q="\rv s. no_orphans s \ invs' s \ sch_act_simple s" in hoare_strengthen_post) apply (wp cteRevoke_preservation cteDelete_invs' cteDelete_sch_act_simple)+ apply auto done lemma cancelBadgedSends_no_orphans [wp]: "\ \s. no_orphans s \ valid_queues' s \ cancelBadgedSends epptr badge \ \rv s. no_orphans s \" unfolding cancelBadgedSends_def apply (rule hoare_pre) apply (wp hoare_drop_imps | wpc | clarsimp)+ apply (wp filterM_preserved tcbSchedEnqueue_almost_no_orphans gts_wp' sts_st_tcb' hoare_drop_imps | clarsimp)+ done crunch no_orphans [wp]: invalidateTLBByASID "no_orphans" crunch no_orphans [wp]: handleFaultReply "no_orphans" crunch valid_queues' [wp]: handleFaultReply "valid_queues'" lemma doReplyTransfer_no_orphans[wp]: "\no_orphans and invs' and tcb_at' sender and tcb_at' receiver\ doReplyTransfer sender receiver slot grant \\rv. no_orphans\" unfolding doReplyTransfer_def apply (wp sts_st_tcb' setThreadState_not_active_no_orphans threadSet_no_orphans threadSet_valid_queues' threadSet_weak_sch_act_wf | wpc | clarsimp simp: is_active_thread_state_def isRunning_def isRestart_def | wp (once) hoare_drop_imps | strengthen sch_act_wf_weak invs_valid_queues')+ apply (rule_tac Q="\rv. invs' and no_orphans" in hoare_post_imp) apply (fastforce simp: inQ_def) apply (wp hoare_drop_imps | clarsimp)+ apply (clarsimp simp:invs'_def valid_state'_def valid_pspace'_def) done lemma cancelSignal_valid_queues' [wp]: "\ \s. valid_queues' s \ valid_objs' s \ cancelSignal t ntfn \ \rv s. valid_queues' s \" unfolding cancelSignal_def Let_def apply (rule hoare_pre) apply (wp hoare_drop_imps | wpc | clarsimp)+ done crunch no_orphans [wp]: setupReplyMaster "no_orphans" (wp: crunch_wps simp: crunch_simps) crunch valid_queues' [wp]: setupReplyMaster "valid_queues'" lemma restart_no_orphans [wp]: "\ \s. no_orphans s \ invs' s \ sch_act_simple s \ tcb_at' t s \ restart t \ \rv s. no_orphans s \" unfolding restart_def isStopped_def2 apply (wp tcbSchedEnqueue_almost_no_orphans sts_st_tcb' cancelIPC_weak_sch_act_wf | clarsimp simp: o_def if_apply_def2 | strengthen no_orphans_strg_almost | strengthen invs_valid_queues' | wp (once) hoare_drop_imps)+ apply auto done lemma readreg_no_orphans: "\ \s. no_orphans s \ invs' s \ sch_act_simple s \ tcb_at' src s \ invokeTCB (tcbinvocation.ReadRegisters src susp n arch) \ \rv s. no_orphans s \" unfolding invokeTCB_def performTransfer_def apply (wp | clarsimp)+ done lemma writereg_no_orphans: "\ \s. no_orphans s \ invs' s \ sch_act_simple s \ tcb_at' dest s \ ex_nonz_cap_to' dest s\ invokeTCB (tcbinvocation.WriteRegisters dest resume values arch) \ \rv s. no_orphans s \" unfolding invokeTCB_def performTransfer_def postModifyRegisters_def apply simp apply (rule hoare_pre) by (wp hoare_vcg_if_lift hoare_vcg_conj_lift restart_invs' static_imp_wp | strengthen invs_valid_queues' | clarsimp simp: invs'_def valid_state'_def dest!: global'_no_ex_cap )+ lemma copyreg_no_orphans: "\ \s. no_orphans s \ invs' s \ sch_act_simple s \ tcb_at' src s \ tcb_at' dest s \ ex_nonz_cap_to' src s \ ex_nonz_cap_to' dest s \ invokeTCB (tcbinvocation.CopyRegisters dest src susp resume frames ints arch) \ \rv s. no_orphans s \" unfolding invokeTCB_def performTransfer_def postModifyRegisters_def supply if_weak_cong[cong] apply simp apply (wp hoare_vcg_if_lift static_imp_wp) apply (wp hoare_vcg_imp_lift' mapM_x_wp' asUser_no_orphans | wpc | clarsimp split del: if_splits)+ apply (wp static_imp_wp hoare_vcg_conj_lift hoare_drop_imp mapM_x_wp' restart_invs' restart_no_orphans asUser_no_orphans suspend_nonz_cap_to_tcb | strengthen invs_valid_queues' | wpc | simp add: if_apply_def2)+ apply (fastforce simp: invs'_def valid_state'_def dest!: global'_no_ex_cap) done lemma settlsbase_no_orphans: "\ \s. no_orphans s \ invs' s \ invokeTCB (tcbinvocation.SetTLSBase src dest) \ \rv s. no_orphans s \" unfolding invokeTCB_def performTransfer_def apply simp apply (wp hoare_vcg_if_lift static_imp_wp) apply (wpsimp wp: hoare_vcg_imp_lift' mapM_x_wp' asUser_no_orphans)+ done lemma almost_no_orphans_no_orphans: "\ almost_no_orphans t s; \ is_active_tcb_ptr t s \ \ no_orphans s" by (auto simp: almost_no_orphans_def no_orphans_def all_active_tcb_ptrs_def) lemma almost_no_orphans_no_orphans': "\ almost_no_orphans t s; ksCurThread s = t\ \ no_orphans s" by (auto simp: almost_no_orphans_def no_orphans_def all_active_tcb_ptrs_def) lemma setPriority_no_orphans [wp]: "\ \s. no_orphans s \ invs' s \ tcb_at' tptr s \ setPriority tptr prio \ \rv s. no_orphans s \" unfolding setPriority_def apply wpsimp apply (rule_tac Q="\rv s. almost_no_orphans tptr s \ valid_queues' s \ weak_sch_act_wf (ksSchedulerAction s) s" in hoare_post_imp) apply clarsimp apply (clarsimp simp: is_active_tcb_ptr_runnable' pred_tcb_at'_def obj_at'_def almost_no_orphans_no_orphans elim!: almost_no_orphans_no_orphans') apply (wp threadSet_almost_no_orphans threadSet_valid_queues' | clarsimp simp: inQ_def)+ apply (wpsimp wp: threadSet_weak_sch_act_wf) apply (wp tcbSchedDequeue_almost_no_orphans| clarsimp)+ apply (rule_tac Q="\rv. obj_at' (Not \ tcbQueued) tptr and invs' and (\s. weak_sch_act_wf (ksSchedulerAction s) s)" in hoare_post_imp) apply (clarsimp simp: obj_at'_def inQ_def) apply (wp tcbSchedDequeue_not_queued | clarsimp)+ done lemma setMCPriority_no_orphans[wp]: "\no_orphans\ setMCPriority t prio \\rv. no_orphans\" unfolding setMCPriority_def apply (rule hoare_pre) apply (wp threadSet_no_orphans) by clarsimp+ lemma threadSet_ipcbuffer_invs: "is_aligned a msg_align_bits \ \invs' and tcb_at' t\ threadSet (tcbIPCBuffer_update (\_. a)) t \\rv. invs'\" apply (wp threadSet_invs_trivial, simp_all add: inQ_def cong: conj_cong) done lemma tc_no_orphans: "\ no_orphans and invs' and sch_act_simple and tcb_at' a and ex_nonz_cap_to' a and case_option \ (valid_cap' o fst) e' and K (case_option True (isCNodeCap o fst) e') and case_option \ (valid_cap' o fst) f' and K (case_option True (isValidVTableRoot o fst) f') and case_option \ (valid_cap') (case_option None (case_option None (Some o fst) o snd) g) and K (case_option True isArchObjectCap (case_option None (case_option None (Some o fst) o snd) g)) and K (case_option True (swp is_aligned 2 o fst) g) and K (case_option True (swp is_aligned msg_align_bits o fst) g) and K (case g of None \ True | Some x \ (case_option True (isArchObjectCap \ fst) \ snd) x) and K (valid_option_prio d \ valid_option_prio mcp) \ invokeTCB (tcbinvocation.ThreadControl a sl b' mcp d e' f' g) \ \rv s. no_orphans s \" apply (rule hoare_gen_asm) apply (rule hoare_gen_asm) apply (rule hoare_gen_asm) apply (simp add: invokeTCB_def getThreadCSpaceRoot getThreadVSpaceRoot getThreadBufferSlot_def split_def) apply (simp only: eq_commute[where a="a"]) apply (rule hoare_walk_assmsE) apply (clarsimp simp: pred_conj_def option.splits[where P="\x. x s" for s]) apply ((wp case_option_wp threadSet_no_orphans threadSet_invs_trivial threadSet_cap_to' hoare_vcg_all_lift static_imp_wp | clarsimp simp: inQ_def)+)[2] apply (rule hoare_walk_assmsE) apply (cases mcp; clarsimp simp: pred_conj_def option.splits[where P="\x. x s" for s]) apply ((wp case_option_wp threadSet_no_orphans threadSet_invs_trivial setMCPriority_invs' typ_at_lifts[OF setMCPriority_typ_at'] threadSet_cap_to' hoare_vcg_all_lift static_imp_wp | clarsimp simp: inQ_def)+)[3] apply ((simp only: simp_thms cong: conj_cong | wp cteDelete_deletes cteDelete_invs' cteDelete_sch_act_simple case_option_wp[where m'="return ()", OF setPriority_no_orphans return_inv,simplified] checkCap_inv[where P="valid_cap' c" for c] checkCap_inv[where P=sch_act_simple] checkCap_inv[where P=no_orphans] checkCap_inv[where P="tcb_at' a"] threadSet_cte_wp_at' hoare_vcg_all_lift_R hoare_vcg_all_lift threadSet_no_orphans hoare_vcg_const_imp_lift_R static_imp_wp hoare_drop_imp threadSet_ipcbuffer_invs | strengthen invs_valid_queues' | (simp add: locateSlotTCB_def locateSlotBasic_def objBits_def objBitsKO_def tcbIPCBufferSlot_def tcb_cte_cases_def, wp hoare_return_sp) | wpc | clarsimp)+) apply (fastforce simp: objBits_defs isCap_simps dest!: isValidVTableRootD) done lemma bindNotification_no_orphans[wp]: "\no_orphans\ bindNotification t ntfn \\_. no_orphans\" unfolding bindNotification_def by wp lemma invokeTCB_no_orphans [wp]: "\ \s. no_orphans s \ invs' s \ sch_act_simple s \ tcb_inv_wf' tinv s \ invokeTCB tinv \ \rv s. no_orphans s \" apply (case_tac tinv, simp_all) apply (clarsimp simp: invokeTCB_def) apply (wp, clarsimp) apply (clarsimp simp: invokeTCB_def) apply (wp, clarsimp) apply (wp tc_no_orphans) apply (clarsimp split: option.splits simp: msg_align_bits elim!:is_aligned_weaken) apply (rename_tac option) apply (case_tac option) apply ((wp | simp add: invokeTCB_def)+)[2] apply (wp writereg_no_orphans readreg_no_orphans copyreg_no_orphans settlsbase_no_orphans | clarsimp)+ done lemma invokeCNode_no_orphans [wp]: "\ \s. no_orphans s \ invs' s \ valid_cnode_inv' cinv s \ sch_act_simple s \ invokeCNode cinv \ \rv. no_orphans \" unfolding invokeCNode_def apply (rule hoare_pre) apply (wp hoare_drop_imps hoare_unless_wp | wpc | clarsimp split del: if_split)+ apply (simp add: invs_valid_queues') done lemma invokeIRQControl_no_orphans [wp]: "\ \s. no_orphans s \ performIRQControl i \ \rv s. no_orphans s \" apply (cases i, simp_all add: performIRQControl_def ARM_H.performIRQControl_def) apply (rename_tac archinv) apply (case_tac archinv) apply (wp | clarsimp)+ done lemma invokeIRQHandler_no_orphans [wp]: "\ \s. no_orphans s \ invs' s \ invokeIRQHandler i \ \reply s. no_orphans s \" apply (cases i, simp_all add: invokeIRQHandler_def) apply (wp | clarsimp | fastforce)+ done crunch no_orphans [wp]: setVMRootForFlush "no_orphans" (wp: crunch_wps) lemma performPageTableInvocation_no_orphans [wp]: "\ \s. no_orphans s \ performPageTableInvocation pti \ \reply s. no_orphans s \" apply (cases pti, simp_all add: performPageTableInvocation_def) apply (rule hoare_pre) apply (wp mapM_x_wp' | wpc | clarsimp)+ done lemma performPageInvocation_no_orphans [wp]: "\ \s. no_orphans s \ performPageInvocation pgi \ \reply s. no_orphans s \" apply (simp add: performPageInvocation_def cong: page_invocation.case_cong) apply (rule hoare_pre) apply (wp mapM_x_wp' mapM_wp' static_imp_wp | wpc | clarsimp simp: pdeCheckIfMapped_def pteCheckIfMapped_def)+ done lemma performASIDControlInvocation_no_orphans [wp]: notes blah[simp del] = atLeastAtMost_iff atLeastatMost_subset_iff atLeastLessThan_iff Int_atLeastAtMost atLeastatMost_empty_iff split_paired_Ex usableUntypedRange.simps shows "\ \s. no_orphans s \ invs' s \ valid_aci' aci s \ ct_active' s \ performASIDControlInvocation aci \ \reply s. no_orphans s \" apply (rule hoare_name_pre_state) apply (clarsimp simp:valid_aci'_def cte_wp_at_ctes_of split:asidcontrol_invocation.splits) apply (rename_tac s ptr_base p cref ptr null_cte ut_cte idx) proof - fix s ptr_base p cref ptr null_cte ut_cte idx assume no_orphans: "no_orphans s" and invs' : "invs' s" and cte : "ctes_of s p = Some null_cte" "cteCap null_cte = capability.NullCap" "ctes_of s cref = Some ut_cte" "cteCap ut_cte = capability.UntypedCap False ptr_base pageBits idx" and desc : "descendants_of' cref (ctes_of s) = {}" and misc : "p \ cref" "ex_cte_cap_wp_to' (\_. True) p s" "sch_act_simple s" "is_aligned ptr asid_low_bits" "(ptr :: word32) < 2 ^ asid_bits" "ct_active' s" have vc:"s \' UntypedCap False ptr_base pageBits idx" using cte misc invs' apply - apply (case_tac ut_cte) apply (rule ctes_of_valid_cap') apply simp apply fastforce done hence cover: "range_cover ptr_base pageBits pageBits (Suc 0)" apply - apply (rule range_cover_full) apply (simp add:valid_cap'_def capAligned_def) apply simp done have exclude: "cref \ {ptr_base..ptr_base + 2 ^ pageBits - 1}" apply (rule descendants_range_ex_cte'[where cte = "ut_cte"]) apply (rule empty_descendants_range_in'[OF desc]) apply (rule if_unsafe_then_capD'[where P = "\c. c = ut_cte"]) apply (clarsimp simp: cte_wp_at_ctes_of cte) apply (simp add:invs' invs_unsafe_then_cap') apply (simp add:cte invs')+ done show "\(=) s\performASIDControlInvocation (asidcontrol_invocation.MakePool ptr_base p cref ptr) \\reply. no_orphans\" apply (clarsimp simp: performASIDControlInvocation_def split: asidcontrol_invocation.splits) apply (wp static_imp_wp | clarsimp)+ apply (rule_tac Q="\rv s. no_orphans s" in hoare_post_imp) apply (clarsimp simp: no_orphans_def all_active_tcb_ptrs_def is_active_tcb_ptr_def all_queued_tcb_ptrs_def) apply (wp | clarsimp simp:placeNewObject_def2)+ apply (wp createObjects'_wp_subst)+ apply (wp static_imp_wp updateFreeIndex_pspace_no_overlap'[where sz= pageBits] getSlotCap_wp | simp)+ apply (strengthen invs_pspace_aligned' invs_pspace_distinct' invs_valid_pspace') apply (clarsimp simp:conj_comms) apply (wp deleteObjects_invs'[where idx = idx and d=False] hoare_ex_wp deleteObjects_cte_wp_at'[where idx = idx and d=False] hoare_vcg_const_imp_lift ) using invs' misc cte exclude no_orphans cover apply (clarsimp simp: is_active_thread_state_def makeObject_tcb valid_aci'_def cte_wp_at_ctes_of invs_pspace_aligned' invs_pspace_distinct' projectKO_opt_tcb isRunning_def isRestart_def conj_comms invs_valid_pspace' vc objBits_simps archObjSize_def range_cover.aligned) apply (intro conjI) apply (rule vc) apply (simp add:descendants_range'_def2) apply (rule empty_descendants_range_in'[OF desc]) apply clarsimp done qed lemma performASIDPoolInvocation_no_orphans [wp]: "\ \s. no_orphans s \ performASIDPoolInvocation api \ \reply s. no_orphans s \" apply (cases api, simp_all add: performASIDPoolInvocation_def) apply (wp getObject_inv loadObject_default_inv | clarsimp)+ done lemma performPageDirectoryInvocation_no_orphans [wp]: "\ \s. no_orphans s \ performPageDirectoryInvocation pdi \ \reply s. no_orphans s \" apply (cases pdi, simp_all add: performPageDirectoryInvocation_def) apply (wp | simp)+ done lemma arch_performInvocation_no_orphans [wp]: "\ \s. no_orphans s \ invs' s \ valid_arch_inv' i s \ ct_active' s \ Arch.performInvocation i \ \reply s. no_orphans s \" unfolding ARM_H.performInvocation_def performARMMMUInvocation_def apply (cases i, simp_all add: valid_arch_inv'_def) apply (wp | clarsimp)+ done lemma setDomain_no_orphans [wp]: "\no_orphans and valid_queues and valid_queues' and cur_tcb'\ setDomain tptr newdom \\_. no_orphans\" apply (simp add: setDomain_def when_def) apply (wp tcbSchedEnqueue_almost_no_orphans hoare_vcg_imp_lift threadSet_almost_no_orphans threadSet_valid_queues'_no_state threadSet_st_tcb_at2 hoare_vcg_disj_lift threadSet_no_orphans | clarsimp simp: st_tcb_at_neg2 not_obj_at')+ apply (auto simp: tcb_at_typ_at' st_tcb_at_neg' is_active_tcb_ptr_runnable' cur_tcb'_def obj_at'_def dest: pred_tcb_at') done crunch no_orphans[wp]: InterruptDecls_H.invokeIRQHandler no_orphans lemma performInvocation_no_orphans [wp]: "\ \s. no_orphans s \ invs' s \ valid_invocation' i s \ ct_active' s \ sch_act_simple s \ performInvocation block call i \ \reply s. no_orphans s \" apply (simp add: performInvocation_def cong: invocation.case_cong) apply (rule hoare_pre) apply (wp | wpc | clarsimp)+ apply auto done lemma getThreadState_restart [wp]: "\ \s. tcb_at' thread s \ getThreadState thread \ \rv s. rv = Structures_H.thread_state.Restart \ st_tcb_at' isRestart thread s \" apply (rule hoare_strengthen_post) apply (rule gts_st_tcb') apply (clarsimp simp add: pred_tcb_at'_def obj_at'_def isRestart_def) done lemma K_bind_hoareE [wp]: "\P\ f \Q\,\E\ \ \P\ K_bind f x \Q\,\E\" by simp crunch valid_queues' [wp]: replyFromKernel "valid_queues'" lemma handleInvocation_no_orphans [wp]: "\ \s. no_orphans s \ invs' s \ vs_valid_duplicates' (ksPSpace s) \ ct_active' s \ ksSchedulerAction s = ResumeCurrentThread \ handleInvocation isCall isBlocking \ \rv s. no_orphans s \" unfolding handleInvocation_def apply (rule hoare_pre) apply (wp syscall_valid' setThreadState_isRestart_no_orphans | wpc | clarsimp)+ apply (rule_tac Q="\state s. no_orphans s \ invs' s \ (state = Structures_H.thread_state.Restart \ st_tcb_at' isRestart thread s)" in hoare_post_imp) apply (wp | clarsimp)+ apply (wp setThreadState_current_no_orphans sts_invs_minor' ct_in_state'_set setThreadState_st_tcb hoare_vcg_all_lift | simp add: split_def split del: if_split)+ apply (wps setThreadState_ct') apply (wp sts_ksQ setThreadState_current_no_orphans sts_invs_minor' ct_in_state'_set setThreadState_st_tcb | simp add: split_def split del: if_split)+ apply (clarsimp simp: if_apply_def2) apply (frule(1) ct_not_ksQ) by (auto simp: ct_in_state'_def pred_tcb_at'_def obj_at'_def invs'_def cur_tcb'_def valid_state'_def valid_idle'_def) lemma receiveSignal_no_orphans [wp]: "\ \s. no_orphans s \ valid_queues' s \ receiveSignal thread cap isBlocking \ \rv s. no_orphans s \" unfolding receiveSignal_def apply (wp hoare_drop_imps setThreadState_not_active_no_orphans | wpc | clarsimp simp: is_active_thread_state_def isRunning_def isRestart_def doNBRecvFailedTransfer_def)+ done lemma receiveIPC_no_orphans [wp]: "\ \s. no_orphans s \ invs' s \ receiveIPC thread cap is_blocking \ \rv s. no_orphans s \" unfolding receiveIPC_def apply (rule hoare_pre) apply (wp setThreadState_not_active_no_orphans hoare_drop_imps hoare_vcg_all_lift sts_st_tcb' | wpc | clarsimp simp: is_active_thread_state_def isRunning_def isRestart_def doNBRecvFailedTransfer_def invs_valid_queues' | strengthen sch_act_wf_weak)+ done crunch valid_objs' [wp]: getThreadCallerSlot "valid_objs'" lemma deleteCallerCap_no_orphans [wp]: "\ \s. no_orphans s \ invs' s \ deleteCallerCap receiver \ \rv s. no_orphans s \" unfolding deleteCallerCap_def by wpsimp auto lemma remove_neg_strg: "(A \ B) \ ((x \ A) \ (\ x \ B))" by blast lemma handleRecv_no_orphans [wp]: notes if_cong[cong] shows "\ \s. no_orphans s \ invs' s \ handleRecv isBlocking \ \rv . no_orphans \" unfolding handleRecv_def apply (clarsimp simp: whenE_def split del: if_split | wp hoare_drop_imps getNotification_wp | wpc )+ (*takes a while*) apply (rule_tac Q'="\rv s. no_orphans s \ invs' s" in hoare_post_imp_R) apply (wp, fastforce) apply (rule_tac Q="\rv s. no_orphans s \ invs' s" in hoare_post_imp) apply (wp | clarsimp | fastforce)+ done crunches getThreadCallerSlot, handleHypervisorFault for invs' [wp]: "invs'" lemma handleReply_no_orphans [wp]: "\no_orphans and invs'\ handleReply \\_. no_orphans\" unfolding handleReply_def apply (rule hoare_pre) apply (wp hoare_drop_imps | wpc | clarsimp)+ apply (wp hoare_vcg_all_lift) apply (rule_tac Q="\rv s. no_orphans s \ invs' s \ tcb_at' thread s \ valid_cap' rv s" in hoare_post_imp) apply (wp hoare_drop_imps | clarsimp simp: valid_cap'_def | clarsimp simp: invs'_def cur_tcb'_def valid_state'_def)+ done lemma handleYield_no_orphans [wp]: "\ \s. no_orphans s \ invs' s \ handleYield \ \rv . no_orphans \" unfolding handleYield_def apply (wp tcbSchedAppend_almost_no_orphans) apply auto done lemma activatable_from_running': "ct_running' s \ ct_in_state' activatable' s" by (clarsimp simp: ct_in_state'_def elim!: pred_tcb'_weakenE) (* FIXME move *) lemma sts_tcb_at'_preserve: "\ st_tcb_at' P t and K (P st) \ setThreadState st t' \\_. st_tcb_at' P t \" by (wpsimp wp: sts_st_tcb') (* FIXME move *) (* e.g. if you set a non-runnable thread to Inactive, all runnable threads are still runnable *) lemma sts_tcb_at'_preserve': "\ st_tcb_at' P t and st_tcb_at' (\st. \ P st) t' and K (\ P st) \ setThreadState st t' \\_. st_tcb_at' P t \" by (wpsimp wp: sts_st_tcb' simp: st_tcb_at_neg') lemma handleEvent_no_orphans [wp]: "\ \s. invs' s \ vs_valid_duplicates' (ksPSpace s) \ (e \ Interrupt \ ct_running' s) \ ksSchedulerAction s = ResumeCurrentThread \ no_orphans s \ handleEvent e \ \rv s. no_orphans s \" apply (simp add: handleEvent_def handleSend_def handleCall_def cong: event.case_cong syscall.case_cong) apply (rule hoare_pre) apply (wp hv_inv' hoare_drop_imps | wpc | clarsimp simp: handleHypervisorFault_def)+ apply (auto simp: activatable_from_running' active_from_running') done theorem callKernel_no_orphans [wp]: "\ \s. invs' s \ vs_valid_duplicates' (ksPSpace s) \ (e \ Interrupt \ ct_running' s) \ ksSchedulerAction s = ResumeCurrentThread \ no_orphans s \ callKernel e \ \rv s. no_orphans s \" unfolding callKernel_def by (wpsimp wp: weak_if_wp schedule_invs' hoare_drop_imps) end end