(* * Copyright 2014, General Dynamics C4 Systems * * SPDX-License-Identifier: GPL-2.0-only *) theory DetSchedSchedule_AI imports ArchDetSchedDomainTime_AI begin crunch ct_not_in_q[wp]: do_machine_op "ct_not_in_q" lemma set_tcb_queue_wp[wp]: "\\s. P (ready_queues_update (\_ t' p. if t' = t \ p = prio then queue else ready_queues s t' p) s)\ set_tcb_queue t prio queue \\_. P\" apply (simp add: set_tcb_queue_def) apply wp done lemma get_tcb_queue_wp[wp]: "\\s. P (ready_queues s t p) s\ get_tcb_queue t p \P\" apply (simp add: get_tcb_queue_def) apply wp done crunch valid_etcbs[wp]: tcb_sched_action "valid_etcbs" lemma enqueue_distinct[intro!]: "distinct queue \ distinct (tcb_sched_enqueue thread queue)" apply (simp add: tcb_sched_enqueue_def) done lemma is_activatable_trivs[iff]: "is_activatable_2 ct (switch_thread t) kh" "is_activatable_2 ct choose_new_thread kh" by (simp_all add: is_activatable_def) lemma weak_valid_sched_action_trivs[iff]: "weak_valid_sched_action_2 resume_cur_thread ekh kh" "weak_valid_sched_action_2 choose_new_thread ekh kh" by (simp_all add: weak_valid_sched_action_def) lemma switch_in_cur_domain_trivs[iff]: "switch_in_cur_domain_2 resume_cur_thread ekh cdom" "switch_in_cur_domain_2 choose_new_thread ekh cdom" by (simp_all add: switch_in_cur_domain_def) lemma ct_in_cur_domain_2_trivs[iff]: "ct_in_cur_domain_2 thread thread' (switch_thread t) cdom ekh" "ct_in_cur_domain_2 thread thread' choose_new_thread cdom ekh" by (simp_all add: ct_in_cur_domain_2_def) lemma valid_sched_action_triv[iff]: "valid_sched_action_2 choose_new_thread ekh kh ct cdom" by (simp add: valid_sched_action_def) lemma simple_sched_action_trivs[iff]: "simple_sched_action_2 resume_cur_thread" "simple_sched_action_2 choose_new_thread" by (simp_all add: simple_sched_action_def) lemma scheduler_act_not_trivs[iff]: "scheduler_act_not_2 resume_cur_thread t" "scheduler_act_not_2 choose_new_thread t" by (simp_all add: scheduler_act_not_def) lemma tcb_sched_action_enqueue_valid_queues[wp]: "\valid_queues and st_tcb_at runnable thread\ tcb_sched_action tcb_sched_enqueue thread \\_. valid_queues\" apply (simp add: tcb_sched_action_def, wp) apply simp apply (clarsimp simp: valid_queues_def2 etcb_at_def tcb_sched_enqueue_def is_etcb_at_def split: option.split) done lemma tcb_sched_action_append_valid_queues[wp]: "\valid_queues and st_tcb_at runnable thread\ tcb_sched_action tcb_sched_append thread \\_. valid_queues\" apply (simp add: tcb_sched_action_def, wp) apply (clarsimp simp: valid_queues_def2 etcb_at_def tcb_sched_append_def is_etcb_at_def split: option.split) done lemma tcb_sched_action_dequeue_valid_queues[wp]: "\valid_queues\ tcb_sched_action tcb_sched_dequeue thread \\_. valid_queues\" apply (simp add: tcb_sched_action_def, wp) apply (clarsimp simp: valid_queues_def2 etcb_at_def tcb_sched_dequeue_def split: option.splits) apply blast done lemma tcb_sched_action_enqueue_ct_not_in_q[wp]: "\ct_not_in_q and not_cur_thread thread\ tcb_sched_action tcb_sched_enqueue thread \\_. ct_not_in_q\" apply (simp add: tcb_sched_action_def, wp) apply (clarsimp simp: ct_not_in_q_def not_cur_thread_def etcb_at_def tcb_sched_enqueue_def not_queued_def split: option.splits) done lemma tcb_sched_action_append_ct_not_in_q[wp]: "\ct_not_in_q and not_cur_thread thread\ tcb_sched_action tcb_sched_append thread \\_. ct_not_in_q\" apply (simp add: tcb_sched_action_def, wp) apply (clarsimp simp: ct_not_in_q_def not_cur_thread_def etcb_at_def tcb_sched_append_def not_queued_def split: option.splits) done lemma tcb_sched_action_dequeue_ct_not_in_q[wp]: "\ct_not_in_q\ tcb_sched_action tcb_sched_dequeue thread \\_. ct_not_in_q\" apply (simp add: tcb_sched_action_def, wp) apply (fastforce simp: ct_not_in_q_def etcb_at_def tcb_sched_dequeue_def not_queued_def split: option.splits) done crunch is_activatable[wp]: tcb_sched_action "is_activatable t" crunch weak_valid_sched_action[wp]: tcb_sched_action "weak_valid_sched_action" crunch valid_sched_action[wp]: tcb_sched_action valid_sched_action crunch ct_in_cur_domain[wp]: tcb_sched_action ct_in_cur_domain lemma tcb_sched_action_enqueue_valid_blocked: "\valid_blocked_except thread\ tcb_sched_action tcb_sched_enqueue thread \\_. valid_blocked\" apply (simp add: tcb_sched_action_def, wp) apply (clarsimp simp: etcb_at_def tcb_sched_enqueue_def not_queued_def valid_blocked_def valid_blocked_except_def split: option.splits) apply (rule conjI) apply clarsimp apply (case_tac "t=thread") apply force apply (erule_tac x=t in allE) apply clarsimp apply force apply clarsimp apply (case_tac "t=thread") apply force apply (erule_tac x=t in allE) apply clarsimp apply force done lemma tcb_sched_action_append_valid_blocked: "\valid_blocked_except thread\ tcb_sched_action tcb_sched_append thread \\_. valid_blocked\" apply (simp add: tcb_sched_action_def, wp) apply (clarsimp simp: etcb_at_def tcb_sched_append_def not_queued_def valid_blocked_def valid_blocked_except_def split: option.splits) apply (rule conjI) apply clarsimp apply (case_tac "t=thread") apply force apply (erule_tac x=t in allE) apply clarsimp apply force apply clarsimp apply (case_tac "t=thread") apply force apply (erule_tac x=t in allE) apply clarsimp apply force done lemma tcb_sched_action_dequeue_valid_blocked: "\valid_blocked and st_tcb_at (\st. \ active st) thread\ tcb_sched_action tcb_sched_dequeue thread \\_. valid_blocked\" apply (simp add: tcb_sched_action_def, wp) apply (clarsimp simp: etcb_at_def tcb_sched_dequeue_def not_queued_def valid_blocked_def split: option.splits) apply (case_tac "t=thread") apply simp apply (force simp: st_tcb_at_def obj_at_def) apply (erule_tac x=t in allE) apply force done lemma tcb_sched_action_dequeue_valid_blocked_except: "\valid_blocked\ tcb_sched_action tcb_sched_dequeue thread \\_. valid_blocked_except thread\" apply (simp add: tcb_sched_action_def, wp) apply (clarsimp simp: etcb_at_def tcb_sched_dequeue_def not_queued_def valid_blocked_def valid_blocked_except_def split: option.splits) apply (case_tac "t=thread") apply simp apply (erule_tac x=t in allE) apply force done abbreviation valid_sched_except_blocked_2 where "valid_sched_except_blocked_2 queues ekh sa cdom kh ct it \ valid_etcbs_2 ekh kh \ valid_queues_2 queues ekh kh \ ct_not_in_q_2 queues sa ct \ valid_sched_action_2 sa ekh kh ct cdom \ ct_in_cur_domain_2 ct it sa cdom ekh \ valid_idle_etcb_2 ekh" abbreviation valid_sched_except_blocked :: "det_ext state \ bool" where "valid_sched_except_blocked s \ valid_sched_except_blocked_2 (ready_queues s) (ekheap s) (scheduler_action s) (cur_domain s) (kheap s) (cur_thread s) (idle_thread s)" declare valid_idle_etcb_lift[wp] lemma tcb_sched_action_enqueue_valid_sched[wp]: "\valid_sched_except_blocked and st_tcb_at runnable thread and not_cur_thread thread and valid_blocked_except thread\ tcb_sched_action tcb_sched_enqueue thread \\_. valid_sched\" by (simp add: valid_sched_def | wp tcb_sched_action_enqueue_valid_blocked)+ lemma tcb_sched_action_append_valid_sched[wp]: "\valid_sched_except_blocked and st_tcb_at runnable thread and not_cur_thread thread and valid_blocked_except thread\ tcb_sched_action tcb_sched_append thread \\_. valid_sched\" by (simp add: valid_sched_def | wp tcb_sched_action_append_valid_blocked)+ lemma tcb_sched_action_dequeue_valid_sched_not_runnable: "\valid_sched and st_tcb_at (\st. \ active st) thread\ tcb_sched_action tcb_sched_dequeue thread \\_. valid_sched\" by (simp add: valid_sched_def | wp tcb_sched_action_dequeue_valid_blocked)+ lemma tcb_sched_action_dequeue_valid_sched_except_blocked: "\valid_sched\ tcb_sched_action tcb_sched_dequeue thread \\_. valid_sched_except_blocked\" by (simp add: valid_sched_def | wp tcb_sched_action_dequeue_valid_blocked_except)+ crunch valid_etcbs[wp]: set_scheduler_action "valid_etcbs" crunch valid_queues[wp]: set_scheduler_action "valid_queues" lemma set_scheduler_action_rct_ct_not_in_q[wp]: "\ct_not_queued\ set_scheduler_action resume_cur_thread \\_. ct_not_in_q\" apply (simp add: set_scheduler_action_def, wp) apply (fastforce simp: ct_not_in_q_def not_queued_def) done lemma set_scheduler_action_rct_is_activatable[wp]: "\st_tcb_at activatable t\ set_scheduler_action resume_cur_thread \\_. is_activatable t\" apply (simp add: set_scheduler_action_def, wp) apply (simp add: is_activatable_def) done lemma set_scheduler_action_rct_weak_valid_sched_action[wp]: "\\\ set_scheduler_action resume_cur_thread \\_. weak_valid_sched_action\" apply (simp add: set_scheduler_action_def, wp) apply (simp add: weak_valid_sched_action_def) done lemma set_scheduler_action_wp: "\\s. Q (scheduler_action_update (\_. a) s)\ set_scheduler_action a \\_.Q\" apply (simp add: set_scheduler_action_def) apply wp done lemma set_scheduler_action_rct_valid_sched_action[wp]: "\\s. st_tcb_at activatable (cur_thread s) s\ set_scheduler_action resume_cur_thread \\_. valid_sched_action\" apply (simp add: valid_sched_action_def | wp set_scheduler_action_wp)+ apply (simp add: is_activatable_def) done lemma set_scheduler_action_rct_ct_in_cur_domain[wp]: "\\s. in_cur_domain (cur_thread s) s \ cur_thread s = idle_thread s\ set_scheduler_action resume_cur_thread \\_. ct_in_cur_domain\" apply (simp add: valid_sched_action_def | wp set_scheduler_action_wp)+ apply (clarsimp simp: in_cur_domain_def ct_in_cur_domain_2_def) done lemma set_scheduler_action_rct_valid_blocked: "\valid_blocked and simple_sched_action\ set_scheduler_action resume_cur_thread \\_. valid_blocked\" apply (simp add: valid_blocked_def simple_sched_action_def split: scheduler_action.splits | wp set_scheduler_action_wp)+ done crunch etcb_at[wp]: set_scheduler_action "etcb_at P t" lemma set_scheduler_action_rct_valid_sched: "\valid_sched and ct_not_queued and (\s. st_tcb_at activatable (cur_thread s) s) and (\s. in_cur_domain (cur_thread s) s \ cur_thread s = idle_thread s) and simple_sched_action\ set_scheduler_action resume_cur_thread \\_.valid_sched\" by (simp add: valid_sched_def | wp set_scheduler_action_rct_valid_blocked)+ lemma set_scheduler_action_cnt_ct_not_in_q[wp]: "\\\ set_scheduler_action choose_new_thread \\_. ct_not_in_q\" apply (simp add: set_scheduler_action_def, wp) apply (simp add: ct_not_in_q_def) done lemma set_scheduler_action_cnt_is_activatable[wp]: "\\\ set_scheduler_action choose_new_thread \\_. is_activatable t\" apply (simp add: set_scheduler_action_def, wp) apply (simp add: is_activatable_def) done lemma set_scheduler_action_cnt_ct_in_cur_domain[wp]: "\\\ set_scheduler_action choose_new_thread \\_. ct_in_cur_domain\" apply (simp add: set_scheduler_action_def, wp) apply (simp add: ct_in_cur_domain_def) done lemma set_scheduler_action_cnt_weak_valid_sched_action[wp]: "\\\ set_scheduler_action choose_new_thread \\_. weak_valid_sched_action\" apply (simp add: set_scheduler_action_def, wp) apply (simp add: weak_valid_sched_action_def) done lemma set_scheduler_action_cnt_valid_sched_action[wp]: "\\\ set_scheduler_action choose_new_thread \\_. valid_sched_action\" apply (simp add: valid_sched_action_def, wp set_scheduler_action_wp) apply (simp add: is_activatable_def) done lemma set_scheduler_action_cnt_valid_blocked_weak: "\valid_blocked and simple_sched_action\ set_scheduler_action choose_new_thread \\_. valid_blocked\" apply (simp add: valid_blocked_def, wp set_scheduler_action_wp) apply (simp add: simple_sched_action_def split: scheduler_action.splits) done lemma set_scheduler_action_cnt_valid_blocked: "\valid_blocked and (\s. \t. scheduler_action s = switch_thread t \ (\d p. t \ set (ready_queues s d p)))\ set_scheduler_action choose_new_thread \\_. valid_blocked\" apply (simp add: valid_blocked_def, wp set_scheduler_action_wp) apply clarsimp apply (erule_tac x=t in allE) apply (erule impCE) apply force apply (erule_tac x=st in allE) apply (force simp: not_queued_def) done lemma set_scheduler_action_cnt_valid_blocked_except: "\valid_blocked_except target and (\s. \t. scheduler_action s = switch_thread t \ (\d p. t \ set (ready_queues s d p)))\ set_scheduler_action choose_new_thread \\_. valid_blocked_except target\" apply (simp add: valid_blocked_except_def, wp set_scheduler_action_wp) apply clarsimp apply (erule_tac x=t in allE) apply (erule impCE) apply force apply (force simp: not_queued_def) done lemma set_scheduler_action_cnt_weak_valid_sched: "\valid_sched and simple_sched_action\ set_scheduler_action choose_new_thread \\_. valid_sched\" by (simp add: valid_sched_def simple_sched_action_def split: scheduler_action.splits | wp set_scheduler_action_cnt_valid_blocked)+ crunch valid_etcbs[wp]: reschedule_required "valid_etcbs" lemma reschedule_required_valid_queues[wp]: "\valid_queues and weak_valid_sched_action\ reschedule_required \\_. valid_queues\" apply (simp add: reschedule_required_def | wp | wpc)+ apply (simp add: weak_valid_sched_action_def) done lemma reschedule_required_ct_not_in_q[wp]: "\\\ reschedule_required \\_. ct_not_in_q\" by (simp add: reschedule_required_def, wp) lemma reschedule_required_is_activatable[wp]: "\\\ reschedule_required \\_. is_activatable t\" by (simp add: reschedule_required_def, wp) lemma reschedule_required_weak_valid_sched_action[wp]: "\\\ reschedule_required \\_. weak_valid_sched_action\" by (simp add: reschedule_required_def, wp) lemma tcb_sched_action_enqueue_valid_blocked_except: "\valid_blocked_except thread\ tcb_sched_action tcb_sched_enqueue thread' \\_. valid_blocked_except thread \" apply (simp add: tcb_sched_action_def, wp) apply (clarsimp simp: etcb_at_def tcb_sched_enqueue_def not_queued_def valid_blocked_def valid_blocked_except_def split: option.splits) apply (rule conjI) apply clarsimp apply (case_tac "t=thread") apply force apply (erule_tac x=t in allE) apply clarsimp apply force apply clarsimp apply (case_tac "t=thread") apply force apply (erule_tac x=t in allE) apply clarsimp apply force done lemma reschedule_required_valid_sched_action[wp]: "\\\ reschedule_required \\_. valid_sched_action\" apply (simp add: valid_sched_action_def reschedule_required_def) apply (wp set_scheduler_action_wp | simp)+ done lemma reschedule_required_ct_in_cur_domain[wp]: "\\\ reschedule_required \\_. ct_in_cur_domain\" by (simp add: reschedule_required_def, wp) lemma reschedule_required_valid_blocked: "\valid_blocked\ reschedule_required \\_. valid_blocked\" apply (simp add: reschedule_required_def | wp set_scheduler_action_cnt_valid_blocked tcb_sched_action_enqueue_valid_blocked hoare_vcg_all_lift | wpc)+ apply (simp add: tcb_sched_action_def) apply wp+ apply clarsimp apply (rule conjI) apply clarsimp apply (rule conjI) apply (force simp: etcb_at_def tcb_sched_enqueue_def valid_blocked_def valid_blocked_except_def split: option.splits) apply clarsimp done lemma reschedule_required_valid_blocked_except: "\valid_blocked_except target\ reschedule_required \\_. valid_blocked_except target\" apply (simp add: reschedule_required_def | wp set_scheduler_action_cnt_valid_blocked_except tcb_sched_action_enqueue_valid_blocked_except hoare_vcg_all_lift | wpc)+ apply (simp add: tcb_sched_action_def) apply wp+ apply clarsimp apply (rule conjI) apply clarsimp apply (rule conjI) apply (force simp: etcb_at_def tcb_sched_enqueue_def valid_blocked_except_def split: option.splits) apply clarsimp done crunch etcb_at[wp]: reschedule_required "etcb_at P t" lemma reschedule_required_valid_sched: "\valid_etcbs and valid_queues and weak_valid_sched_action and valid_blocked and valid_idle_etcb\ reschedule_required \\_. valid_sched\" by (simp add: valid_sched_def | wp reschedule_required_valid_blocked)+ lemma reschedule_required_valid_sched_cur_thread: "\(\s. target = cur_thread s) and valid_etcbs and valid_queues and weak_valid_sched_action and valid_blocked_except target and valid_idle_etcb\ reschedule_required \\_. valid_sched\" apply (simp add: valid_sched_def | wp reschedule_required_valid_blocked)+ apply (clarsimp simp: valid_blocked_except_def valid_blocked_def) done lemma get_tcb_st_tcb_at: "get_tcb t s = Some y \ st_tcb_at \ t s" apply (simp add: get_tcb_def pred_tcb_at_def obj_at_def split: option.splits kernel_object.splits) done lemma st_tcb_at_kh_if_split: "st_tcb_at_kh P ptr (\t. if t = ref then x else kh t) = (if ptr = ref then (\tcb. x = Some (TCB tcb) \ P (tcb_state tcb)) else st_tcb_at_kh P ptr kh)" by (fastforce simp: st_tcb_at_kh_def obj_at_kh_def obj_at_def) lemma set_thread_state_valid_etcbs[wp]: "\valid_etcbs\ set_thread_state ref ts \\_. valid_etcbs\" apply (simp add: set_thread_state_def) apply (wp | wpc | simp add: set_thread_state_ext_def set_object_def get_object_def)+ apply (fastforce simp: valid_etcbs_def st_tcb_at_kh_if_split dest: get_tcb_st_tcb_at) done lemma set_bound_notification_valid_etcbs[wp]: "\valid_etcbs\ set_bound_notification ref ntfn \\_. valid_etcbs\" apply (simp add: set_bound_notification_def) apply (wp | wpc | simp add: set_object_def get_object_def)+ apply (fastforce simp: valid_etcbs_def st_tcb_at_kh_if_split dest: get_tcb_st_tcb_at) done lemma set_thread_state_runnable_valid_queues: "\valid_queues and K (runnable ts)\ set_thread_state ref ts \\_. valid_queues\" apply (simp add: set_thread_state_def) apply (wp | wpc | simp add: set_thread_state_ext_def set_object_def get_object_def)+ apply (clarsimp simp: valid_queues_def st_tcb_at_kh_if_split) done lemma set_bound_notification_valid_queues: "\valid_queues\ set_bound_notification ref ntfn \\_. valid_queues\" apply (simp add: set_bound_notification_def) apply (wpsimp wp: set_object_wp) apply (clarsimp simp: valid_queues_def st_tcb_at_kh_if_split) apply (drule_tac x=d in spec) apply (drule_tac x=p in spec) apply clarsimp apply (drule (1) bspec, clarsimp simp: st_tcb_def2) done lemma set_thread_state_ct_not_in_q[wp]: "\ct_not_in_q\ set_thread_state ref ts \\_. ct_not_in_q\" apply (simp add: set_thread_state_def) apply (simp add: set_thread_state_ext_def set_object_def get_object_def | wp gts_wp)+ done lemma set_bound_notification_ct_not_in_q[wp]: "\ct_not_in_q\ set_bound_notification ref ntfn \\_. ct_not_in_q\" apply (simp add: set_bound_notification_def) apply (simp add: set_object_def get_object_def | wp)+ done lemma set_thread_state_cur_is_activatable[wp]: "\\s. is_activatable (cur_thread s) s\ set_thread_state ref ts \\_ (s::det_state). is_activatable (cur_thread s) s\" apply (simp add: set_thread_state_def) apply (simp add: set_thread_state_ext_def set_object_def get_object_def | wp set_scheduler_action_wp gts_wp)+ apply (clarsimp simp: is_activatable_def st_tcb_at_kh_if_split pred_tcb_at_def obj_at_def) done lemma set_bound_notification_cur_is_activatable[wp]: "\\s. is_activatable (cur_thread s) s\ set_bound_notification ref ntfn \\_ (s::det_state). is_activatable (cur_thread s) s\" apply (simp add: set_bound_notification_def) apply (simp add: set_object_def get_object_def | wp set_scheduler_action_wp)+ apply (clarsimp simp: is_activatable_def st_tcb_at_kh_if_split pred_tcb_at_def obj_at_def get_tcb_def) done lemma set_thread_state_runnable_weak_valid_sched_action: "\weak_valid_sched_action and (\s. runnable ts)\ set_thread_state ref ts \\_. weak_valid_sched_action\" apply (simp add: set_thread_state_def) apply (simp add: set_thread_state_ext_def set_object_def get_object_def | wp gts_wp)+ apply (clarsimp simp: weak_valid_sched_action_def st_tcb_at_kh_if_split) done lemma set_thread_state_switch_in_cur_domain[wp]: "\switch_in_cur_domain\ set_thread_state ref ts \\_. switch_in_cur_domain\" apply (simp add: set_thread_state_def) apply (simp add: set_thread_state_ext_def set_object_def get_object_def | wp set_scheduler_action_wp gts_wp)+ done lemma set_bound_notification_switch_in_cur_domain[wp]: "\switch_in_cur_domain\ set_bound_notification ref ts \\_. switch_in_cur_domain\" apply (simp add: set_bound_notification_def) apply (simp add: set_object_def get_object_def | wp set_scheduler_action_wp gbn_wp)+ done lemma set_bound_notification_weak_valid_sched_action: "\weak_valid_sched_action\ set_bound_notification ref ntfnptr \\_. weak_valid_sched_action\" apply (simp add: set_bound_notification_def) apply (simp add: set_object_def get_object_def | wp)+ apply (clarsimp simp: weak_valid_sched_action_def st_tcb_at_kh_if_split st_tcb_def2) done lemma set_thread_state_runnable_valid_sched_action: "\valid_sched_action and (\s. runnable ts)\ set_thread_state ref ts \\_. valid_sched_action\" apply (simp add: valid_sched_action_def | wp set_thread_state_runnable_weak_valid_sched_action)+ done lemma set_thread_state_cur_ct_in_cur_domain[wp]: "\ct_in_cur_domain\ set_thread_state ref ts \\_. ct_in_cur_domain\" apply (simp add: set_thread_state_def) apply (simp add: set_thread_state_ext_def set_object_def get_object_def | wp set_scheduler_action_wp gts_wp)+ done lemma set_bound_notification_cur_ct_in_cur_domain[wp]: "\ct_in_cur_domain\ set_bound_notification ref ts \\_. ct_in_cur_domain\" apply (simp add: set_bound_notification_def) apply (simp add: set_object_def get_object_def | wp set_scheduler_action_wp gbn_wp)+ done crunch etcb_at[wp]: set_thread_state, set_bound_notification, get_bound_notification "etcb_at P t" lemma set_thread_state_runnable_valid_sched_except_blocked: "\valid_sched and (\s. runnable ts)\ set_thread_state ref ts \\_. valid_sched_except_blocked\" apply (simp add: valid_sched_def | wp set_thread_state_runnable_valid_queues set_thread_state_runnable_valid_sched_action)+ done lemma set_bound_notification_valid_sched_action: "\valid_sched_action\ set_bound_notification ref ntfn \\_. valid_sched_action\" apply (simp add: valid_sched_action_def | wp set_bound_notification_weak_valid_sched_action)+ done lemma set_bound_notification_valid_blocked[wp]: "\valid_blocked\ set_bound_notification t ntfn \\_. valid_blocked\" apply (simp add: set_bound_notification_def) apply (simp add: set_object_def get_object_def | wp)+ apply (clarsimp simp: valid_blocked_def) apply (drule_tac x=ta in spec, clarsimp) apply (drule_tac x=st in spec) apply (simp only: st_tcb_at_kh_if_split) apply (simp add: st_tcb_def2 split: if_split_asm) done lemma set_bound_notification_valid_sched: "\valid_sched\ set_bound_notification ref ntfn \\_. valid_sched\" apply (simp add: valid_sched_def | wp set_bound_notification_valid_queues set_bound_notification_valid_sched_action)+ done lemma valid_blocked_valid_blocked_except[simp]: "valid_blocked s \ valid_blocked_except t s" by (simp add: valid_blocked_def valid_blocked_except_def) lemma set_thread_state_valid_blocked_except: "\valid_blocked\ set_thread_state ref ts \\_. valid_blocked_except ref\" apply (simp add: set_thread_state_def) apply (simp add: set_thread_state_ext_def set_object_def get_object_def | wp)+ apply (rule hoare_strengthen_post) apply (rule set_scheduler_action_cnt_valid_blocked_weak) apply simp apply (wp gts_wp)+ apply (clarsimp simp: valid_blocked_def valid_blocked_except_def st_tcb_at_def obj_at_def) done (* as user schedule invariants *) lemma as_user_valid_etcbs[wp]: "\valid_etcbs\ as_user ptr s \\_. valid_etcbs\" apply (simp add: as_user_def set_object_def get_object_def | wpc | wp)+ by (fastforce simp: valid_etcbs_def st_tcb_at_kh_if_split dest: get_tcb_st_tcb_at) lemma as_user_valid_queues[wp]: "\valid_queues\ as_user ptr s \\_. valid_queues\" apply (simp add: as_user_def set_object_def get_object_def | wpc | wp)+ apply (clarsimp simp: valid_queues_def st_tcb_at_kh_if_split st_tcb_at_def obj_at_def) apply (drule_tac x=d in spec) apply (drule_tac x=p in spec) apply clarsimp apply (drule (1) bspec, clarsimp simp: st_tcb_def2) apply (drule get_tcb_SomeD, auto) done lemma as_user_weak_valid_sched_action[wp]: "\weak_valid_sched_action\ as_user ptr s \\_. weak_valid_sched_action\" apply (simp add: as_user_def set_object_def get_object_def | wpc | wp)+ apply (clarsimp simp: weak_valid_sched_action_def st_tcb_at_kh_if_split st_tcb_at_def obj_at_def) by (drule get_tcb_SomeD, auto) lemma as_user_is_activatable[wp]: "\is_activatable t\ as_user ptr s \\_. is_activatable t\" apply (simp add: as_user_def set_object_def get_object_def | wpc | wp)+ apply (clarsimp simp: is_activatable_def st_tcb_at_kh_if_split st_tcb_at_def obj_at_def) by (drule get_tcb_SomeD, auto) lemma as_user_valid_sched_action[wp]: "\valid_sched_action\ as_user ptr s \\_. valid_sched_action\" apply (simp add: as_user_def set_object_def get_object_def | wpc | wp)+ apply (clarsimp simp: valid_sched_action_def st_tcb_at_def obj_at_def) apply (rule conjI) apply (clarsimp simp: is_activatable_def st_tcb_at_def obj_at_def st_tcb_at_kh_if_split) apply (drule get_tcb_SomeD, clarsimp) apply (clarsimp simp: weak_valid_sched_action_def st_tcb_at_def obj_at_def st_tcb_at_kh_if_split) apply (drule get_tcb_SomeD, clarsimp) done crunch ct_in_cur_domain[wp]: as_user ct_in_cur_domain (wp: ct_in_cur_domain_lift) lemma as_user_valid_idle_etcb[wp]: "\valid_idle_etcb\ as_user ptr s \\_. valid_idle_etcb\" by (simp add: as_user_def set_object_def get_object_def | wpc | wp)+ lemma as_user_valid_blocked[wp]: "\valid_blocked\ as_user ptr s \\_. valid_blocked\" apply (simp add: as_user_def set_object_def get_object_def | wpc | wp)+ apply (clarsimp simp: valid_blocked_def st_tcb_at_kh_if_split st_tcb_at_def obj_at_def) apply (drule_tac x=ptr in spec) by (drule get_tcb_SomeD, auto) definition ct_in_q where "ct_in_q s \ st_tcb_at runnable (cur_thread s) s \ (\d p. cur_thread s \ set (ready_queues s d p))" locale DetSchedSchedule_AI = assumes arch_switch_to_idle_thread_valid_etcbs'[wp]: "\valid_etcbs\ arch_switch_to_idle_thread \\_. valid_etcbs\" assumes arch_switch_to_thread_valid_etcbs'[wp]: "\t. \valid_etcbs\ arch_switch_to_thread t \\_. valid_etcbs\" assumes arch_switch_to_idle_thread_valid_queues'[wp]: "\valid_queues\ arch_switch_to_idle_thread \\_. valid_queues\" assumes arch_switch_to_thread_valid_queues'[wp]: "\t. \valid_queues\ arch_switch_to_thread t \\_. valid_queues\" assumes arch_switch_to_idle_thread_weak_valid_sched_action'[wp]: "\weak_valid_sched_action\ arch_switch_to_idle_thread \\_. weak_valid_sched_action\" assumes arch_switch_to_thread_weak_valid_sched_action'[wp]: "\t. \weak_valid_sched_action\ arch_switch_to_thread t \\_. weak_valid_sched_action\" assumes switch_to_idle_thread_ct_not_in_q[wp]: "\valid_queues and valid_idle\ switch_to_idle_thread \\_. ct_not_in_q\" assumes switch_to_idle_thread_valid_sched_action[wp]: "\valid_sched_action and valid_idle\ switch_to_idle_thread \\_. valid_sched_action\" assumes switch_to_idle_thread_ct_in_cur_domain[wp]: "\\\ switch_to_idle_thread \\_. ct_in_cur_domain\" assumes arch_switch_to_thread_ct_not_in_q'[wp]: "\t. \ct_not_in_q\ arch_switch_to_thread t \\_. ct_not_in_q\" assumes arch_switch_to_thread_is_activatable'[wp]: "\t t'. \is_activatable t'\ arch_switch_to_thread t \\_. is_activatable t'\" assumes arch_switch_to_thread_valid_sched_action'[wp]: "\t. \valid_sched_action\ arch_switch_to_thread t \\_. valid_sched_action\" assumes arch_switch_to_thread_valid_sched'[wp]: "\t. \valid_sched\ arch_switch_to_thread t \\_. valid_sched\" assumes arch_switch_to_thread_ct_in_cur_domain_2[wp]: "\t' t. \\s. ct_in_cur_domain_2 t' (idle_thread s) (scheduler_action s) (cur_domain s) (ekheap s)\ arch_switch_to_thread t \\_ s. ct_in_cur_domain_2 t' (idle_thread s) (scheduler_action s) (cur_domain s) (ekheap s)\" assumes arch_switch_to_thread_valid_blocked[wp]: "\t. \valid_blocked and ct_in_q\ arch_switch_to_thread t \\_. valid_blocked and ct_in_q\" assumes arch_switch_to_thread_etcb_at'[wp]: "\P t' t. \etcb_at P t'\ arch_switch_to_thread t \\_. etcb_at P t'\" assumes arch_switch_to_idle_thread_valid_idle[wp]: "\valid_idle :: det_ext state \ bool\ arch_switch_to_idle_thread \\_. valid_idle\" assumes switch_to_idle_thread_ct_not_queued[wp]: "\valid_queues and valid_etcbs and valid_idle\ switch_to_idle_thread \\rv s. not_queued (cur_thread s) s\" assumes switch_to_idle_thread_valid_blocked[wp]: "\valid_blocked and ct_in_q\ switch_to_idle_thread \\rv. valid_blocked\" assumes arch_switch_to_thread_exst'[wp]: "\P t. \\s. P (exst s :: det_ext)\ arch_switch_to_thread t \\rv s. P (exst s)\" assumes arch_switch_to_thread_scheduler_action'[wp]: "\P t. \\s. P (scheduler_action (s :: det_ext state))\ arch_switch_to_thread t \\rv s. P (scheduler_action (s :: det_ext state))\" assumes stit_activatable': "\valid_idle :: det_ext state \ bool\ switch_to_idle_thread \\rv . ct_in_state activatable\" assumes arch_switch_to_idle_thread_etcb_at'[wp]: "\P t. \etcb_at P t\ arch_switch_to_idle_thread \\_. etcb_at P t\" assumes switch_to_idle_thread_cur_thread_idle_thread [wp]: "\\ :: det_ext state \ bool\ switch_to_idle_thread \\_ s. cur_thread s = idle_thread s\" assumes arch_switch_to_idle_thread_scheduler_action'[wp]: "\P. \\s. P (scheduler_action s)\ arch_switch_to_idle_thread \\_ s. P (scheduler_action s)\" assumes arch_finalise_cap_ct_not_in_q'[wp]: "\acap final. \ct_not_in_q\ arch_finalise_cap acap final \\_. ct_not_in_q\" assumes arch_finalise_cap_valid_etcbs'[wp]: "\acap final. \valid_etcbs\ arch_finalise_cap acap final \\_. valid_etcbs\" assumes arch_finalise_cap_simple_sched_action'[wp]: "\acap final. \simple_sched_action\ arch_finalise_cap acap final \\_. simple_sched_action\" assumes arch_finalise_cap_valid_sched'[wp]: "\acap final. \valid_sched\ arch_finalise_cap acap final \\_. valid_sched\" assumes arch_post_cap_deletion_valid_idle[wp]: "\target. \valid_idle :: det_ext state \ bool\ arch_post_cap_deletion target \\_. valid_idle\" assumes handle_arch_fault_reply_valid_sched'[wp]: "\f t x y. \valid_sched\ handle_arch_fault_reply f t x y \\_. valid_sched\" assumes activate_thread_valid_sched: "\valid_sched\ activate_thread \\_. valid_sched\" assumes arch_perform_invocation_valid_sched[wp]: "\i. \invs and valid_sched and ct_active and valid_arch_inv i\ arch_perform_invocation i \\_.valid_sched\" assumes arch_invoke_irq_control_valid_sched'[wp]: "\i. \valid_sched\ arch_invoke_irq_control i \\_. valid_sched\" assumes handle_vm_fault_valid_sched'[wp]: "\t f. \valid_sched\ handle_vm_fault t f \\_. valid_sched\" assumes handle_vm_fault_not_queued'[wp]: "\t' t f. \not_queued t'\ handle_vm_fault t f \\_. not_queued t'\" assumes handle_vm_fault_scheduler_act_not'[wp]: "\t' t f. \scheduler_act_not t'\ handle_vm_fault t f \\_. scheduler_act_not t'\" assumes handle_arch_fault_reply_not_queued'[wp]: "\t' f t x y. \not_queued t'\ handle_arch_fault_reply f t x y \\_. not_queued t'\" assumes handle_arch_fault_reply_scheduler_act_not'[wp]: "\t' f t x y. \scheduler_act_not t'\ handle_arch_fault_reply f t x y \\_. scheduler_act_not t'\" assumes handle_arch_fault_reply_cur'[wp]: "\f t x y. \cur_tcb :: det_ext state \ bool\ handle_arch_fault_reply f t x y \\_. cur_tcb\" assumes hvmf_st_tcb_at[wp]: "\P t' t w. \st_tcb_at P t' :: det_ext state \ bool \ handle_vm_fault t w \\rv. st_tcb_at P t' \" assumes handle_vm_fault_st_tcb_cur_thread[wp]: "\P t f. \ \s :: det_ext state. st_tcb_at P (cur_thread s) s \ handle_vm_fault t f \\_ s. st_tcb_at P (cur_thread s) s \" assumes arch_activate_idle_thread_valid_list'[wp]: "\t. \valid_list\ arch_activate_idle_thread t \\_. valid_list\" assumes arch_switch_to_thread_valid_list'[wp]: "\t. \valid_list\ arch_switch_to_thread t \\_. valid_list\" assumes arch_switch_to_idle_thread_valid_list'[wp]: "\valid_list\ arch_switch_to_idle_thread \\_. valid_list\" assumes prepare_thread_delete_idel_thread[wp] : "\t. prepare_thread_delete t \\(s:: det_ext state). P (idle_thread s)\" assumes prepare_thread_delete_ct_not_in_q'[wp]: "\t. \ct_not_in_q\ prepare_thread_delete t \\_. ct_not_in_q\" assumes prepare_thread_delete_valid_etcbs'[wp]: "\t. \valid_etcbs\ prepare_thread_delete t \\_. valid_etcbs\" assumes prepare_thread_delete_simple_sched_action'[wp]: "\t. \simple_sched_action\ prepare_thread_delete t \\_. simple_sched_action\" assumes prepare_thread_delete_valid_sched'[wp]: "\t. \valid_sched\ prepare_thread_delete t \\_. valid_sched\" assumes make_fault_arch_msg_not_cur_thread[wp] : "\ft t t'. make_arch_fault_msg ft t \not_cur_thread t'\" assumes make_fault_arch_msg_valid_sched[wp] : "\ft t. make_arch_fault_msg ft t \valid_sched\" assumes make_fault_arch_msg_scheduler_action[wp] : "\P ft t. make_arch_fault_msg ft t \\s. P (scheduler_action s)\" assumes make_fault_arch_msg_ready_queues[wp] : "\P ft t. make_arch_fault_msg ft t \\s. P (ready_queues s)\" assumes make_fault_arch_msg_valid_etcbs[wp] : "\ft t. make_arch_fault_msg ft t \valid_etcbs\" assumes arch_get_sanitise_register_info_not_cur_thread[wp] : "\ft t'. arch_get_sanitise_register_info ft \not_cur_thread t'\" assumes arch_get_sanitise_register_info_valid_sched[wp] : "\ft. arch_get_sanitise_register_info ft \valid_sched\" assumes arch_get_sanitise_register_info_scheduler_action[wp] : "\P ft. arch_get_sanitise_register_info ft \\s. P (scheduler_action s)\" assumes arch_get_sanitise_register_info_ready_queues[wp] : "\P ft. arch_get_sanitise_register_info ft \\s. P (ready_queues s)\" assumes arch_get_sanitise_register_info_valid_etcbs[wp] : "\ft. arch_get_sanitise_register_info ft \valid_etcbs\" assumes arch_get_sanitise_register_info_cur'[wp]: "\f. \cur_tcb :: det_ext state \ bool\ arch_get_sanitise_register_info f \\_. cur_tcb\" assumes arch_post_modify_registers_not_cur_thread[wp] : "\c ft t'. arch_post_modify_registers c ft \not_cur_thread t'\" assumes arch_post_modify_registers_valid_sched[wp] : "\c ft. arch_post_modify_registers c ft \valid_sched\" assumes arch_post_modify_registers_scheduler_action[wp] : "\P c ft. arch_post_modify_registers c ft \\s. P (scheduler_action s)\" assumes arch_post_modify_registers_ready_queues[wp] : "\P c ft. arch_post_modify_registers c ft \\s. P (ready_queues s)\" assumes arch_post_modify_registers_valid_etcbs[wp] : "\c ft. arch_post_modify_registers c ft \valid_etcbs\" assumes arch_post_modify_registers_cur'[wp]: "\c f. \cur_tcb :: det_ext state \ bool\ arch_post_modify_registers c f \\_. cur_tcb\" assumes arch_post_modify_registers_not_idle_thread[wp]: "\c t. \\s::det_ext state. t \ idle_thread s\ arch_post_modify_registers c t \\_ s. t \ idle_thread s\" assumes arch_post_cap_deletion_valid_etcbs[wp] : "\c. arch_post_cap_deletion c \valid_etcbs\" assumes arch_post_cap_deletion_valid_sched[wp] : "\c. arch_post_cap_deletion c \valid_sched\" assumes arch_post_cap_deletion_ct_not_in_q[wp] : "\c. arch_post_cap_deletion c \ct_not_in_q\" assumes arch_post_cap_deletion_simple_sched_action[wp] : "\c. arch_post_cap_deletion c \simple_sched_action\" assumes arch_post_cap_deletion_not_cur_thread[wp] : "\c t. arch_post_cap_deletion c \not_cur_thread t\" assumes arch_post_cap_deletion_sched_act_not[wp] : "\c t. arch_post_cap_deletion c \scheduler_act_not t\" assumes arch_post_cap_deletion_not_queued[wp] : "\c t. arch_post_cap_deletion c \not_queued t\" assumes arch_post_cap_deletion_is_etcb_at[wp] : "\c t. arch_post_cap_deletion c \is_etcb_at t\" assumes arch_post_cap_deletion_weak_valid_sched_action[wp] : "\c. arch_post_cap_deletion c \weak_valid_sched_action\" assumes arch_finalise_cap_idle_thread[wp] : "\P b t. arch_finalise_cap t b \\ (s:: det_ext state). P (idle_thread s)\" assumes arch_invoke_irq_handler_valid_sched[wp]: "\i. arch_invoke_irq_handler i \valid_sched\" assumes arch_mask_irq_signal_valid_sched[wp]: "\irq. arch_mask_irq_signal irq \valid_sched\" context DetSchedSchedule_AI begin crunch valid_etcbs[wp]: switch_to_idle_thread, switch_to_thread valid_etcbs (simp: whenE_def) crunch valid_queues[wp]: switch_to_idle_thread, switch_to_thread valid_queues (simp: whenE_def ignore: set_tcb_queue tcb_sched_action) crunch weak_valid_sched_action[wp]: switch_to_idle_thread, switch_to_thread "weak_valid_sched_action" (simp: whenE_def) end lemma tcb_sched_action_dequeue_ct_not_in_q_2_ct_upd: "\valid_queues\ tcb_sched_action tcb_sched_dequeue thread \\r s. ct_not_in_q_2 (ready_queues s) (scheduler_action s) thread\" apply (simp add: tcb_sched_action_def unless_def set_tcb_queue_def) apply wp apply (fastforce simp: etcb_at_def ct_not_in_q_def valid_queues_def tcb_sched_dequeue_def not_queued_def split: option.split) done lemma tcb_sched_action_dequeue_valid_sched_action_2_ct_upd: "\valid_sched_action and (\s. is_activatable_2 thread (scheduler_action s) (kheap s))\ tcb_sched_action tcb_sched_dequeue thread \\r s. valid_sched_action_2 (scheduler_action s) (ekheap s) (kheap s) thread (cur_domain s)\" apply (simp add: tcb_sched_action_def unless_def set_tcb_queue_def) apply wp apply (clarsimp simp: etcb_at_def valid_sched_action_def split: option.split) done context DetSchedSchedule_AI begin lemma as_user_valid_sched[wp]: "\valid_sched\ as_user tptr f \\rv. valid_sched\" apply (simp add: as_user_def set_object_def get_object_def) apply (wp | wpc)+ apply clarsimp apply (fastforce simp: valid_sched_def valid_etcbs_def valid_queues_def valid_sched_action_def is_activatable_def weak_valid_sched_action_def st_tcb_at_kh_if_split st_tcb_def2 valid_blocked_def) done lemma switch_to_thread_ct_not_queued[wp]: "\valid_queues\ switch_to_thread t \\rv s. not_queued (cur_thread s) s\" apply (simp add: switch_to_thread_def) including no_pre apply wp prefer 4 apply (rule get_wp) prefer 3 apply (rule assert_inv) prefer 2 apply (rule arch_switch_to_thread_valid_queues') apply (simp add: tcb_sched_action_def tcb_sched_dequeue_def | wp)+ apply (clarsimp simp add: valid_queues_def etcb_at_def not_queued_def split: option.splits) done end lemma ct_not_in_q_def2: "ct_not_in_q_2 queues sa ct = (sa = resume_cur_thread \ (\d p. ct \ set (queues d p)))" by (fastforce simp add: ct_not_in_q_def not_queued_def) context DetSchedSchedule_AI begin lemma switch_to_thread_ct_not_in_q[wp]: "\valid_queues\ switch_to_thread t \\_. ct_not_in_q\" apply (simp add: ct_not_in_q_def2 not_queued_def[symmetric]) apply (wp hoare_drop_imps | simp)+ done lemma switch_to_thread_valid_sched_action[wp]: "\valid_sched_action and is_activatable t\ switch_to_thread t \\_. valid_sched_action\" by (wpsimp wp: tcb_sched_action_dequeue_valid_sched_action_2_ct_upd simp: switch_to_thread_def) end lemma tcb_sched_action_dequeue_ct_in_cur_domain': "\\s. ct_in_cur_domain_2 thread (idle_thread s) (scheduler_action s) (cur_domain s) (ekheap s)\ tcb_sched_action tcb_sched_dequeue thread \\_ s. ct_in_cur_domain (s\cur_thread := thread\)\" apply (simp add: tcb_sched_action_def) apply wp apply (simp add: etcb_at_def split: option.split) done context DetSchedSchedule_AI begin crunch ct_in_cur_domain_2[wp]: as_user "\s. ct_in_cur_domain_2 thread (idle_thread s) (scheduler_action s) (cur_domain s) (ekheap s)" (simp: whenE_def get_object_def) lemma switch_to_thread_ct_in_cur_domain[wp]: "\\s. ct_in_cur_domain_2 thread (idle_thread s) (scheduler_action s) (cur_domain s) (ekheap s)\ switch_to_thread thread \\_. ct_in_cur_domain\" apply (simp add: switch_to_thread_def) apply (wpsimp wp: tcb_sched_action_dequeue_ct_in_cur_domain') done end lemma ct_in_q_valid_blocked_ct_upd: "\ct_in_q s; valid_blocked s\ \ valid_blocked (s\cur_thread := thread\)" apply (clarsimp simp: valid_blocked_def ct_in_q_def) apply (erule_tac x=t in allE) apply clarsimp apply (case_tac "t=cur_thread s") apply (simp add: not_queued_def) apply (case_tac st, (force simp: st_tcb_at_def obj_at_def)+) done lemma tcb_sched_action_dequeue_valid_blocked': "\valid_blocked and ct_in_q\ tcb_sched_action tcb_sched_dequeue thread \\_ s. valid_blocked (s\cur_thread := thread\)\" apply (simp add: tcb_sched_action_def, wp) apply (clarsimp simp: etcb_at_def tcb_sched_dequeue_def not_queued_def valid_blocked_def split: option.splits) apply (case_tac "t = cur_thread s") apply (simp add: ct_in_q_def) apply clarsimp apply (rule ccontr) apply clarsimp apply (erule impE) apply (case_tac st, (force simp: st_tcb_at_def obj_at_def)+)[1] apply force apply (erule_tac x=t in allE) apply force done lemma ct_in_q_arch_state_upd[simp]: "ct_in_q (s\arch_state := f\) = ct_in_q s" by (simp add: ct_in_q_def) lemma ct_in_q_machine_state_upd[simp]: "ct_in_q (s\machine_state := f\) = ct_in_q s" by (simp add: ct_in_q_def) context DetSchedSchedule_AI begin lemma as_user_valid_blocked[wp]: "\valid_blocked and ct_in_q\ as_user tp S \\_. valid_blocked and ct_in_q\" apply (simp add: as_user_def set_object_def get_object_def | wpc | wp)+ apply (clarsimp simp: valid_blocked_def st_tcb_at_kh_if_split st_tcb_at_def obj_at_def) apply (drule_tac x=tp in spec) apply (drule get_tcb_SomeD, simp) apply (simp add: ct_in_q_def) apply (case_tac "tp = cur_thread s"; clarsimp simp: st_tcb_at_def obj_at_def) done end crunch ct_in_q[wp]: do_machine_op ct_in_q lemma do_machine_op_valid_blocked[wp]: "\valid_blocked and ct_in_q\ do_machine_op f \\_. valid_blocked and ct_in_q\" by (wpsimp | auto)+ context DetSchedSchedule_AI begin lemma switch_to_thread_valid_blocked[wp]: "\valid_blocked and ct_in_q\ switch_to_thread thread \\_. valid_blocked\" apply (simp add: switch_to_thread_def) apply (wp|wpc)+ prefer 4 apply (rule get_wp) prefer 3 apply (rule assert_inv) prefer 2 apply (rule arch_switch_to_thread_valid_blocked) by (wp tcb_sched_action_dequeue_ct_in_cur_domain' tcb_sched_action_dequeue_valid_blocked') crunch etcb_at[wp]: switch_to_thread "etcb_at P t" lemma switch_to_thread_valid_sched: "\is_activatable t and in_cur_domain t and valid_sched_action and valid_etcbs and valid_queues and valid_blocked and ct_in_q and valid_idle_etcb\ switch_to_thread t \\_. valid_sched\" apply (simp add: valid_sched_def | wp | simp add: ct_in_cur_domain_2_def)+ done crunch valid_idle[wp]: switch_to_idle_thread "valid_idle :: det_ext state \ bool" crunch scheduler_action[wp]: switch_to_thread "\s. P (scheduler_action (s :: det_ext state))" end crunch valid_etcbs[wp]: update_cdt_list "valid_etcbs" crunch valid_queues[wp]: update_cdt_list "valid_queues" crunch ct_not_in_q[wp]: update_cdt_list "ct_not_in_q" crunch weak_valid_sched_action[wp]: update_cdt_list "weak_valid_sched_action" crunch valid_sched_action[wp]: update_cdt_list "valid_sched_action" crunch valid_blocked[wp]: update_cdt_list valid_blocked crunch valid_sched[wp]: update_cdt_list "valid_sched" crunch valid_etcbs[wp]: set_cdt, set_cap valid_etcbs (wp: valid_etcbs_lift set_cap_typ_at) crunch valid_queues[wp]: set_cdt, set_cap valid_queues (wp: valid_queues_lift) crunch ct_not_in_q[wp]: set_cdt, set_cap ct_not_in_q (wp: ct_not_in_q_lift) crunch weak_valid_sched_action[wp]: set_cdt, set_cap weak_valid_sched_action (wp: weak_valid_sched_action_lift) crunch valid_sched_action[wp]: set_cdt, set_cap valid_sched_action (wp: valid_sched_action_lift) crunch valid_blocked[wp]: set_cdt, set_cap valid_blocked (wp: valid_blocked_lift) crunch valid_sched[wp]: set_cdt, set_cap valid_sched (wp: valid_sched_lift set_cap_typ_at) crunch ct_not_in_q[wp]: cap_insert ct_not_in_q (wp: hoare_drop_imps) crunch weak_valid_sched_action[wp]: cap_insert weak_valid_sched_action (wp: hoare_drop_imps) lemma valid_queues_trivial[simp]: "valid_queues_2 (\_ _. []) kh ekh" by (simp add: valid_queues_def) lemma ct_not_in_q_trivial[simp]: "ct_not_in_q_2 (\_ _. []) sa ct" by (simp add: ct_not_in_q_def not_queued_def) lemma st_tcb_at_get_lift: "get_tcb thread s = Some y \ test (tcb_state y) \ st_tcb_at test thread s" by (simp add: ct_in_state_def st_tcb_def2) lemmas det_ext_simps[simp] = select_switch_det_ext_ext_def unwrap_ext_det_ext_ext_def lemma guarded_switch_to_lift: "\P\ switch_to_thread thread \Q\ \ \P\ guarded_switch_to thread \Q\" by (wpsimp simp: guarded_switch_to_def get_thread_state_def thread_get_def) lemma tcb_sched_action_lift: "(\f s. P s \ P (ready_queues_update f s)) \ \P\ tcb_sched_action a b \\_. P\" by (wp | simp add: tcb_sched_action_def etcb_at_def)+ definition next_thread where "next_thread queues \ (hd (max_non_empty_queue queues))" lemma next_thread_update: assumes a: "P(next_thread queues)" assumes b: "P t" shows "P(next_thread (queues((p :: word8) := t # (queues prio))))" proof - have imp_comp[simp]: "\P Q. {x. P x \ Q x} = {x. \ P x} \ {x. Q x}" by blast { assume c: "{prio. queues prio \ []} = {}" from a b c have ?thesis by (simp add: next_thread_def max_non_empty_queue_def) } moreover { assume d: "{prio. queues prio \ []} \ {}" from a b have ?thesis apply (clarsimp simp: next_thread_def max_non_empty_queue_def Max_insert[OF finite_word d]) apply (clarsimp simp add: max_def) done } ultimately show ?thesis by blast qed lemma enqueue_thread_queued: "\\\ tcb_sched_action tcb_sched_enqueue thread \\_ s. (\d prio. thread \ set (ready_queues s d prio))\" apply (simp add: tcb_sched_action_def) apply wp apply (fastforce simp: etcb_at_def tcb_sched_enqueue_def split: option.splits) done lemma enqueue_nonempty: "\\\ tcb_sched_action tcb_sched_enqueue thread \\_ s. (\d prio. ready_queues s d prio \ [])\" by (rule hoare_post_imp[OF _ enqueue_thread_queued], metis empty_iff empty_set) lemma next_thread_queued: "queues p \ [] \ \p. next_thread queues \ set (queues p)" apply (clarsimp simp: next_thread_def max_non_empty_queue_def) apply (rule_tac x="Max {prio. queues prio \ []}" in exI) apply (rule Max_prop) apply simp apply blast done context DetSchedSchedule_AI begin crunch etcb_at[wp]: switch_to_idle_thread "etcb_at P t" lemma switch_to_idle_thread_valid_sched: "\valid_sched_action and valid_idle and valid_queues and valid_etcbs and valid_blocked and ct_in_q and valid_idle_etcb\ switch_to_idle_thread \\_. valid_sched\" by (simp add: valid_sched_def | wp)+ crunch etcb_at[wp]: choose_thread "etcb_at P t" (wp: crunch_wps) lemma choose_thread_valid_sched[wp]: "\valid_sched_action and valid_idle and valid_etcbs and valid_queues and valid_blocked and ct_in_q and valid_idle_etcb\ choose_thread \\_. valid_sched\" apply (simp add: choose_thread_def) apply (wp guarded_switch_to_lift switch_to_idle_thread_valid_sched switch_to_thread_valid_sched) apply (clarsimp simp: valid_queues_def next_thread_def is_activatable_2_def dest!: next_thread_queued) apply (fastforce simp: st_tcb_def2 in_cur_domain_def etcb_at_def split: option.splits) done end lemma tcb_sched_enqueue_in_cur_domain: "\in_cur_domain w\ tcb_sched_action tcb_sched_enqueue t \\_. in_cur_domain w\" apply (simp add: tcb_sched_action_def in_cur_domain_def | wp)+ done crunch valid_etcbs: next_domain valid_etcbs (simp: Let_def) crunch valid_queues: next_domain valid_queues (simp: Let_def) crunch valid_blocked: next_domain valid_blocked (simp: Let_def) crunch ct_in_q: next_domain ct_in_q (simp: Let_def ct_in_q_def) crunch ct_not_in_q: next_domain ct_not_in_q (simp: Let_def) lemma next_domain_valid_sched_action: "\\s. scheduler_action s = choose_new_thread\ next_domain \\_. valid_sched_action\" apply (simp add: next_domain_def thread_set_domain_def ethread_set_def set_eobject_def) apply wp apply (clarsimp simp: Let_def valid_sched_action_def weak_valid_sched_action_def get_etcb_def etcb_at_def) done lemma tcb_sched_action_dequeue_in_cur_domain: "\in_cur_domain thread\ tcb_sched_action tcb_sched_dequeue thread \\_. in_cur_domain thread\" apply (simp add: tcb_sched_action_def) apply wp apply (simp add: etcb_at_def split: option.split) done context DetSchedSchedule_AI begin lemma switch_to_thread_cur_in_cur_domain[wp]: "\in_cur_domain t\ switch_to_thread t \\_ s. in_cur_domain (cur_thread s) s\" apply (simp add: switch_to_thread_def | wp tcb_sched_action_dequeue_in_cur_domain)+ done end lemma tcb_sched_enqueue_cur_ct_in_q: "\\s. cur = cur_thread s\ tcb_sched_action tcb_sched_enqueue cur \\_. ct_in_q\" apply (simp add: tcb_sched_action_def | wp)+ apply (force simp: etcb_at_def ct_in_q_def tcb_sched_enqueue_def split: option.splits) done lemma tcb_sched_enqueue_ct_in_q: "\ ct_in_q \ tcb_sched_action tcb_sched_enqueue cur \\_. ct_in_q\" apply (simp add: tcb_sched_action_def | wp)+ apply (force simp: etcb_at_def ct_in_q_def tcb_sched_enqueue_def split: option.splits) done lemma tcb_sched_append_ct_in_q: "\ ct_in_q \ tcb_sched_action tcb_sched_append cur \\_. ct_in_q\" apply (simp add: tcb_sched_action_def | wp)+ apply (force simp: etcb_at_def ct_in_q_def tcb_sched_append_def split: option.splits) done context DetSchedSchedule_AI begin crunch scheduler_action[wp]: switch_to_idle_thread, next_domain "\s. P (scheduler_action s)" (simp: Let_def) end lemma set_scheduler_action_rct_switch_thread_valid_blocked: "\valid_blocked and (\s. scheduler_action s = switch_thread (cur_thread s))\ set_scheduler_action resume_cur_thread \\_. valid_blocked\" apply (simp add: valid_blocked_def, wp set_scheduler_action_wp) apply clarsimp done lemma set_scheduler_action_rct_switch_thread_valid_sched: "\valid_sched and ct_not_queued and (\s. st_tcb_at activatable (cur_thread s) s) and (\s. in_cur_domain (cur_thread s) s \ cur_thread s = idle_thread s) and (\s. scheduler_action s = switch_thread (cur_thread s))\ set_scheduler_action resume_cur_thread \\_.valid_sched\" by (simp add: valid_sched_def | wp set_scheduler_action_rct_switch_thread_valid_blocked | force)+ context DetSchedSchedule_AI begin lemma switch_to_thread_sched_act_is_cur: "\\s. scheduler_action s = switch_thread word\ switch_to_thread word \\rv s. scheduler_action s = switch_thread (cur_thread s)\" by (simp add: switch_to_thread_def | wp)+ end crunch valid_idle_etcb[wp]: next_domain "valid_idle_etcb" (simp: Let_def) context DetSchedSchedule_AI begin lemma choose_thread_ct_not_queued: "\ valid_queues and valid_idle and valid_etcbs \ choose_thread \\_. ct_not_queued \" by (wpsimp simp: choose_thread_def wp: guarded_switch_to_lift) lemma choose_thread_ct_activatable: "\ valid_queues and valid_idle \ choose_thread \\_ s. st_tcb_at activatable (cur_thread s) s \" apply (wpsimp simp: choose_thread_def wp: guarded_switch_to_lift stit_activatable'[simplified ct_in_state_def] stt_activatable[simplified ct_in_state_def]) apply (fastforce dest!: next_thread_queued simp: next_thread_def valid_queues_def is_activatable_def in_cur_domain_def) done lemma choose_thread_cur_dom_or_idle: "\ valid_queues \ choose_thread \\_ s. (in_cur_domain (cur_thread s) s \ cur_thread s = idle_thread s) \" apply (wpsimp simp: choose_thread_def wp: guarded_switch_to_lift) apply (rule hoare_disjI2) (* idle thread *) apply wp apply (rule hoare_disjI1) (* in current domain *) apply (wp guarded_switch_to_lift)+ apply clarsimp apply (fastforce dest!: next_thread_queued split: option.splits simp: etcb_at_def next_thread_def valid_queues_def in_cur_domain_def) done crunch sched_act[wp]: choose_thread "\s. P (scheduler_action s)" (wp: guarded_switch_to_lift) lemma valid_sched_action_from_choose_thread: "scheduler_action s = choose_new_thread \ valid_sched_action s" unfolding valid_sched_action_def by simp crunch sched_act[wp]: tcb_sched_action "in_cur_domain t" crunch ct_in_q[wp]: set_scheduler_action ct_in_q (simp: ct_in_q_def) lemma ct_in_q_scheduler_action_update_id[simp]: "ct_in_q (scheduler_action_update f s) = ct_in_q s" by (simp add: ct_in_q_def) lemma set_scheduler_action_cnt_simple[wp]: "\\\ set_scheduler_action choose_new_thread \\_. simple_sched_action \" by (wpsimp wp: set_scheduler_action_wp) lemma set_scheduler_action_obvious[wp]: "\\\ set_scheduler_action a \\_ s. scheduler_action s = a\" by (wpsimp wp: set_scheduler_action_wp) lemma set_scheduler_action_cnt_valid_blocked': "\valid_blocked and (\s. scheduler_action s = switch_thread t \ (\d p. t \ set (ready_queues s d p)))\ set_scheduler_action choose_new_thread \\_. valid_blocked\" apply (simp add: valid_blocked_def, wp set_scheduler_action_wp) apply clarsimp apply (erule_tac x=ta in allE) apply (erule impCE) apply force apply (erule_tac x=st in allE) apply (force simp: not_queued_def) done lemma set_scheduler_action_cnt_valid_sched: "\valid_sched and (\s. scheduler_action s = switch_thread t \ (\d p. t \ set (ready_queues s d p)))\ set_scheduler_action choose_new_thread \\_. valid_sched\" by (wpsimp simp: valid_sched_def wp: set_scheduler_action_cnt_valid_blocked'[where t=t]) lemma append_thread_queued: "\\\ tcb_sched_action tcb_sched_append thread \\_ s. (\d prio. thread \ set (ready_queues s d prio))\" apply (simp add: tcb_sched_action_def) apply wp apply (fastforce simp: etcb_at_def tcb_sched_append_def split: option.splits) done (* having is_highest_prio match gets_wp makes it very hard to stop and drop imps etc. *) definition "wrap_is_highest_prio cur_dom target_prio \ Nondet_Monad.gets (is_highest_prio cur_dom target_prio)" lemma schedule_choose_new_thread_valid_sched: "\ valid_idle and valid_etcbs and valid_idle_etcb and valid_queues and valid_blocked and (\s. scheduler_action s = choose_new_thread) and ct_in_q \ schedule_choose_new_thread \\_. valid_sched \" unfolding schedule_choose_new_thread_def apply (wpsimp wp_del: when_wp wp: set_scheduler_action_rct_valid_sched choose_thread_ct_not_queued choose_thread_ct_activatable choose_thread_cur_dom_or_idle hoare_vcg_disj_lift)+ apply (wpsimp wp: next_domain_valid_sched_action next_domain_valid_etcbs next_domain_valid_queues next_domain_valid_blocked next_domain_ct_in_q)+ done lemma schedule_valid_sched: "\valid_sched and valid_idle\ schedule \\_. valid_sched\" unfolding schedule_def supply ethread_get_wp[wp del] apply (wp, wpc) (* resume_cur_thread *) apply wp prefer 2 (* choose new thread *) apply (wp schedule_choose_new_thread_valid_sched tcb_sched_action_enqueue_valid_blocked tcb_sched_enqueue_cur_ct_in_q) (* switch_thread candidate *) apply (rename_tac candidate) apply (wp del: when_wp add: set_scheduler_action_rct_valid_sched schedule_choose_new_thread_valid_sched) apply (rule hoare_vcg_conj_lift) apply (rule_tac t=candidate in set_scheduler_action_cnt_valid_blocked') apply (wp tcb_sched_action_enqueue_valid_blocked tcb_sched_enqueue_ct_in_q enqueue_thread_queued)+ apply (wp schedule_choose_new_thread_valid_sched)+ apply (rule hoare_vcg_conj_lift) apply (rule_tac t=candidate in set_scheduler_action_cnt_valid_blocked') apply (wp tcb_sched_action_append_valid_blocked tcb_sched_append_ct_in_q append_thread_queued set_scheduler_action_rct_switch_thread_valid_sched guarded_switch_to_lift switch_to_thread_valid_sched stt_activatable[simplified ct_in_state_def] hoare_disjI1[OF switch_to_thread_cur_in_cur_domain] switch_to_thread_sched_act_is_cur)+ (* discard result of fastfail calculation *) apply (wpsimp wp: hoare_drop_imp)+ apply (strengthen valid_blocked_valid_blocked_except) apply (wp tcb_sched_action_enqueue_valid_blocked tcb_sched_enqueue_cur_ct_in_q gts_wp)+ apply (clarsimp simp: valid_sched_def valid_sched_action_def weak_valid_sched_action_def switch_in_cur_domain_def) apply (safe ; (fastforce elim: st_tcb_at_opeqI | fastforce simp: valid_sched_def valid_blocked_def valid_blocked_except_def valid_sched_action_def weak_valid_sched_action_def switch_in_cur_domain_def ct_in_q_def ct_not_in_q_def st_tcb_at_def obj_at_def)?) done crunch ct_not_in_q[wp]: as_user ct_not_in_q (wp: ct_not_in_q_lift) crunches update_restart_pc for ct_not_in_q[wp]: "\s. ct_not_in_q s" crunch ct_not_in_q[wp]: finalise_cap ct_not_in_q (wp: crunch_wps hoare_drop_imps unless_wp select_inv mapM_wp subset_refl if_fun_split simp: crunch_simps ignore: tcb_sched_action) end lemma set_simple_ko_valid_etcbs[wp]: "\valid_etcbs\ set_simple_ko f ptr ep \\rv. valid_etcbs\" by (wp hoare_drop_imps valid_etcbs_lift | simp add: set_simple_ko_def)+ lemma set_simple_ko_valid_queues[wp]: "\valid_queues\ set_simple_ko f ptr ep \\rv. valid_queues\" by (wp hoare_drop_imps valid_queues_lift | simp add: set_simple_ko_def)+ lemma set_simple_ko_weak_valid_sched_action[wp]: "\weak_valid_sched_action\ set_simple_ko f ptr ep \\rv. weak_valid_sched_action\" by (wp hoare_drop_imps weak_valid_sched_action_lift | simp add: set_simple_ko_def)+ lemma set_simple_ko_switch_in_cur_domain[wp]: "\switch_in_cur_domain\ set_simple_ko f ptr ep \\rv. switch_in_cur_domain\" by (wp hoare_drop_imps switch_in_cur_domain_lift | simp add: set_simple_ko_def)+ lemma set_simple_ko_ct_in_cur_domain[wp]: "\ct_in_cur_domain\ set_simple_ko f ptr ep \\_. ct_in_cur_domain\" by (wp hoare_drop_imps ct_in_cur_domain_lift | simp add: set_simple_ko_def)+ lemma set_simple_ko_valid_sched[wp]: "\valid_sched\ set_simple_ko f ptr ep \\rv. valid_sched\" by (wp hoare_drop_imps valid_sched_lift | simp add: set_simple_ko_def)+ lemma set_simple_ko_valid_blocked[wp]: "\valid_blocked\ set_simple_ko f ptr ep \\rv. valid_blocked\" by (wp hoare_drop_imps valid_blocked_lift | simp add: set_simple_ko_def)+ crunch etcb_at[wp]: set_simple_ko "etcb_at P t" (wp: crunch_wps simp: crunch_simps) lemma cancel_all_ipc_valid_sched[wp]: "\valid_sched\ cancel_all_ipc epptr \\rv. valid_sched\" apply (simp add: cancel_all_ipc_def) apply (wp mapM_x_wp' set_thread_state_runnable_valid_queues sts_st_tcb_at' hoare_drop_imps set_thread_state_runnable_weak_valid_sched_action hoare_vcg_all_lift set_thread_state_valid_blocked_except tcb_sched_action_enqueue_valid_blocked reschedule_required_valid_sched | wpc | rule subset_refl | simp)+ apply (simp_all add: valid_sched_def valid_sched_action_def) done lemma cancel_all_signals_valid_sched[wp]: "\valid_sched\ cancel_all_signals ntfnptr \\rv. valid_sched\" apply (simp add: cancel_all_signals_def) apply (wp mapM_x_wp' set_thread_state_runnable_valid_queues sts_st_tcb_at' hoare_drop_imps set_thread_state_runnable_weak_valid_sched_action hoare_vcg_all_lift set_thread_state_valid_blocked_except tcb_sched_action_enqueue_valid_blocked reschedule_required_valid_sched | wpc | rule subset_refl | simp)+ apply (simp_all add: valid_sched_def valid_sched_action_def) done lemma thread_set_valid_etcbs[wp]: "\valid_etcbs\ thread_set f tptr \\rv. valid_etcbs\" by (wp hoare_drop_imps valid_etcbs_lift | simp add: thread_set_def)+ lemma thread_set_valid_queues: "(\x. tcb_state (f x) = tcb_state x) \ \valid_queues\ thread_set f tptr \\rv. valid_queues\" by (wp hoare_drop_imps valid_queues_lift thread_set_no_change_tcb_state | simp add: thread_set_def)+ lemma thread_set_weak_valid_sched_action: "(\x. tcb_state (f x) = tcb_state x) \ \weak_valid_sched_action\ thread_set f tptr \\rv. weak_valid_sched_action\" by (wp hoare_drop_imps weak_valid_sched_action_lift thread_set_no_change_tcb_state | simp add: thread_set_def)+ lemma thread_set_not_state_valid_sched: "(\x. tcb_state (f x) = tcb_state x) \ \valid_sched\ thread_set f tptr \\rv. valid_sched\" by (wp hoare_drop_imps valid_sched_lift thread_set_no_change_tcb_state | simp add: thread_set_def)+ lemma unbind_notification_valid_sched[wp]: "\valid_sched\ unbind_notification ntfnptr \\rv. valid_sched\" apply (simp add: unbind_notification_def) apply (rule hoare_seq_ext[OF _ gbn_sp]) apply (case_tac ntfnptra, simp, wp, simp) apply (clarsimp) apply (rule hoare_seq_ext[OF _ get_simple_ko_sp]) apply (wp set_bound_notification_valid_sched, clarsimp) done crunches update_restart_pc for valid_etcbs[wp]: "valid_etcbs" context DetSchedSchedule_AI begin crunch valid_etcbs[wp]: finalise_cap valid_etcbs (wp: hoare_drop_imps unless_wp select_inv mapM_x_wp mapM_wp subset_refl if_fun_split simp: crunch_simps ignore: set_object) crunch valid_sched[wp]: cap_swap_for_delete, empty_slot, cap_delete_one valid_sched (simp: unless_def) lemma reply_cancel_ipc_valid_sched[wp]: "\valid_sched\ reply_cancel_ipc tptr \\rv. valid_sched\" apply (simp add: reply_cancel_ipc_def) apply (wp hoare_drop_imps thread_set_not_state_valid_sched | simp)+ done end lemma st_tcb_at_not: "st_tcb_at (\st. \ P st) t s = (\ st_tcb_at P t s \ tcb_at t s)" apply (clarsimp simp: not_pred_tcb) apply (fastforce simp: st_tcb_at_tcb_at) done lemma set_thread_state_not_runnable_valid_queues: "\valid_queues and ct_not_in_q and (\s. st_tcb_at (\ts. \ runnable ts) ref s)\ set_thread_state ref ts \\_. valid_queues\" apply (simp add: set_thread_state_def) apply (simp add: set_thread_state_ext_def set_object_def get_object_def | wp)+ apply (fastforce simp: valid_queues_def st_tcb_at_kh_if_split ct_not_in_q_def st_tcb_at_not) done lemma set_thread_state_not_runnable_weak_valid_sched_action: "\weak_valid_sched_action and (\s. st_tcb_at (\ts. \ runnable ts) ref s)\ set_thread_state ref ts \\_. weak_valid_sched_action\" apply (simp add: set_thread_state_def) apply (simp add: set_thread_state_ext_def set_object_def get_object_def | wp gts_wp)+ apply (clarsimp simp: weak_valid_sched_action_def st_tcb_at_kh_if_split st_tcb_at_not) done lemma set_thread_state_not_runnable_valid_sched_action: "\valid_sched_action and (\s. st_tcb_at (\ts. \ runnable ts) ref s)\ set_thread_state ref ts \\_. valid_sched_action\" apply (simp add: valid_sched_action_def | wp set_thread_state_not_runnable_weak_valid_sched_action)+ done lemma set_thread_state_not_runnable_valid_blocked: "\valid_blocked and K (\ runnable ts)\ set_thread_state ref ts \\_. valid_blocked\" apply (simp add: set_thread_state_def) apply (simp add: set_thread_state_ext_def set_object_def get_object_def | wp)+ apply (rule hoare_strengthen_post) apply (rule set_scheduler_action_cnt_valid_blocked_weak) apply simp apply (wp gts_wp)+ apply (clarsimp simp: valid_blocked_def valid_blocked_except_def st_tcb_at_def obj_at_def get_tcb_def) apply (case_tac ts, simp_all) done lemma set_thread_state_not_runnable_valid_sched: "\valid_sched and (\s. st_tcb_at (\ts. \ runnable ts) ref s) and K (\ runnable ts)\ set_thread_state ref ts \\_. valid_sched\" apply (simp add: valid_sched_def | wp set_thread_state_not_runnable_valid_queues set_thread_state_not_runnable_valid_sched_action set_thread_state_not_runnable_valid_blocked)+ done lemma blocked_cancel_ipc_valid_sched[wp]: "\valid_sched and (\s. st_tcb_at (\ts. \ runnable ts) tptr s)\ blocked_cancel_ipc ts tptr \\rv. valid_sched\" apply (simp add: blocked_cancel_ipc_def) apply (wp set_thread_state_not_runnable_valid_sched) apply simp done lemma cancel_signal_valid_sched[wp]: "\valid_sched and (\s. st_tcb_at (\ts. \ runnable ts) tptr s)\ cancel_signal tptr ntfnptr \\rv. valid_sched\" apply (simp add: cancel_signal_def) apply (wp set_thread_state_not_runnable_valid_sched hoare_drop_imps | wpc | simp)+ done lemma (in DetSchedSchedule_AI) cancel_ipc_valid_sched[wp]: "\valid_sched\ cancel_ipc tptr \\rv. valid_sched\" apply (simp add: cancel_ipc_def get_thread_state_def thread_get_def) apply (wp | wpc)+ apply (fastforce intro: st_tcb_at_get_lift) done (* valid_queues with thread not runnable *) lemma tcb_sched_action_dequeue_strong_valid_sched: "\valid_etcbs and ct_not_in_q and valid_sched_action and ct_in_cur_domain and valid_blocked and st_tcb_at (\st. \ runnable st) thread and (\es. \d p. (\t\set (ready_queues es d p). is_etcb_at t es \ etcb_at (\t. tcb_priority t = p \ tcb_domain t = d) t es \ (t \ thread \ st_tcb_at runnable t es)) \ distinct (ready_queues es d p)) and valid_idle_etcb\ tcb_sched_action tcb_sched_dequeue thread \\_. valid_sched\" apply (simp add: tcb_sched_action_def unless_def set_tcb_queue_def) apply wp apply (clarsimp simp: etcb_at_def valid_sched_def split: option.split) apply (rule conjI) apply (fastforce simp: etcb_at_def is_etcb_at_def valid_queues_def2 tcb_sched_dequeue_def) apply (rule conjI) apply (fastforce simp: ct_not_in_q_def not_queued_def tcb_sched_dequeue_def) apply (clarsimp simp: valid_blocked_def tcb_sched_dequeue_def) apply (case_tac "t=thread") apply simp apply (force simp: st_tcb_at_def obj_at_def) apply (erule_tac x=t in allE) apply (erule impE) apply (clarsimp simp: not_queued_def split: if_split_asm) apply (erule_tac x=d in allE) apply force apply force done lemma set_scheduler_action_simple_sched_action[wp]: "\K $ simple_sched_action_2 action\ set_scheduler_action action \\rv. simple_sched_action\" by (simp add: set_scheduler_action_def | wp)+ lemma reschedule_required_simple_sched_action[wp]: "\simple_sched_action\ reschedule_required \\rv. simple_sched_action\" by (simp add: reschedule_required_def | wp)+ crunch simple_sched_action[wp]: tcb_sched_action, update_cdt_list simple_sched_action context DetSchedSchedule_AI begin crunches update_restart_pc for simple_sched_action[wp]: "simple_sched_action" and valid_sched[wp]: "valid_sched" (simp: crunch_simps ignore: set_object) crunch simple_sched_action[wp]: finalise_cap simple_sched_action (wp: hoare_drop_imps mapM_x_wp mapM_wp subset_refl simp: unless_def if_fun_split) lemma suspend_valid_sched[wp]: notes seq_ext_inv = seq_ext[where A=I and B="\_. I" for I] shows "\valid_sched and simple_sched_action\ suspend t \\rv. valid_sched\" apply (simp add: suspend_def) apply (rule seq_ext_inv) apply wpsimp apply (rule seq_ext_inv) apply wp apply (rule seq_ext_inv) apply wpsimp apply (wp tcb_sched_action_dequeue_strong_valid_sched | simp)+ apply (simp add: set_thread_state_def) apply (wp gts_wp | wpc | simp add: set_thread_state_def set_thread_state_ext_def reschedule_required_def set_scheduler_action_def tcb_sched_action_def set_object_def get_object_def)+ apply (simp only: st_tcb_at_kh_simp[symmetric]) apply (clarsimp simp: valid_sched_def valid_queues_def st_tcb_at_kh_if_split valid_sched_action_def simple_sched_action_def is_activatable_def valid_blocked_def split: scheduler_action.splits)+ done crunch valid_sched[wp]: finalise_cap valid_sched (wp: crunch_wps simp: crunch_simps) end crunch simple_sched_action[wp]: cap_swap_for_delete, cap_insert simple_sched_action (wp: mapM_wp subset_refl hoare_drop_imps) context DetSchedSchedule_AI begin lemma rec_del_valid_sched'[wp]: "\valid_sched and simple_sched_action\ rec_del call \\rv. valid_sched and simple_sched_action\" apply (rule rec_del_preservation) apply (wp preemption_point_inv' | simp)+ done lemma rec_del_valid_sched[wp]: "\valid_sched and simple_sched_action\ rec_del call \\rv. valid_sched\" apply (rule hoare_strengthen_post) apply (rule rec_del_valid_sched') apply simp done lemma rec_del_simple_sched_action[wp]: "\simple_sched_action\ rec_del call \\rv. simple_sched_action\" by (wp rec_del_preservation preemption_point_inv' | simp)+ end lemma ethread_set_valid_etcbs[wp]: "\valid_etcbs\ ethread_set f tptr \\_. valid_etcbs\" apply (simp add: ethread_set_def set_eobject_def | wp)+ apply (clarsimp simp: valid_etcbs_def is_etcb_at_def get_etcb_def) done crunch ct_not_in_q[wp]: ethread_set "ct_not_in_q" crunch not_cur_thread[wp]: ethread_set, tcb_sched_action "not_cur_thread t" crunch cur_is_activatable[wp]: ethread_set "\s. is_activatable (cur_thread s) s" lemma ethread_set_not_queued_valid_queues: "\valid_queues and not_queued tptr\ ethread_set f tptr \\_. valid_queues\" apply (simp add: ethread_set_def set_eobject_def not_queued_def | wp)+ apply (clarsimp simp: valid_queues_def is_etcb_at_def etcb_at_def) done lemma ethread_set_weak_valid_sched_action[wp]: "\weak_valid_sched_action\ ethread_set f tptr \\_. weak_valid_sched_action\" apply (simp add: ethread_set_def set_eobject_def | wp)+ apply (clarsimp simp: weak_valid_sched_action_def is_etcb_at_def etcb_at_def get_etcb_def) done lemma ethread_set_not_domain_switch_in_cur_domain: "(\x. tcb_domain (f x) = tcb_domain x) \ \switch_in_cur_domain\ ethread_set f tptr \\_. switch_in_cur_domain\" apply (simp add: ethread_set_def set_eobject_def | wp)+ apply (clarsimp simp: switch_in_cur_domain_def in_cur_domain_def etcb_at_def get_etcb_def) done lemma ethread_set_not_domain_ct_in_cur_domain: "(\x. tcb_domain (f x) = tcb_domain x) \ \ct_in_cur_domain\ ethread_set f tptr \\_. ct_in_cur_domain\" apply (simp add: ethread_set_def set_eobject_def | wp)+ apply (clarsimp simp: ct_in_cur_domain_def in_cur_domain_def etcb_at_def get_etcb_def) done lemma ethread_set_not_domain_valid_blocked: "(\x. tcb_domain (f x) = tcb_domain x) \ \valid_blocked\ ethread_set f tptr \\_. valid_blocked\" apply (simp add: ethread_set_def set_eobject_def | wp)+ done lemma ethread_set_not_domain_valid_blocked_except: "(\x. tcb_domain (f x) = tcb_domain x) \ \valid_blocked_except ref\ ethread_set f tptr \\_. valid_blocked_except ref\" apply (simp add: ethread_set_def set_eobject_def | wp)+ done lemma ethread_set_not_domain_valid_sched_action: "(\x. tcb_domain (f x) = tcb_domain x) \ \valid_sched_action\ ethread_set f tptr \\_. valid_sched_action\" apply (simp add: valid_sched_action_def | wp ethread_set_not_domain_switch_in_cur_domain)+ done lemma ethread_set_valid_idle_etcb: "(\x. tcb_domain (f x) = tcb_domain x) \ \valid_idle_etcb\ ethread_set f tptr \\_. valid_idle_etcb\" apply(simp add: ethread_set_def valid_idle_etcb_def set_eobject_def) apply wp apply(simp add: etcb_at_def get_etcb_def split: option.splits) done lemma ethread_set_not_queued_valid_sched: "(\x. tcb_domain (f x) = tcb_domain x) \ \valid_sched and not_queued tptr\ ethread_set f tptr \\_. valid_sched\" apply (simp add: valid_sched_def valid_sched_action_def | wp ethread_set_not_queued_valid_queues ethread_set_not_domain_switch_in_cur_domain ethread_set_not_domain_ct_in_cur_domain ethread_set_not_domain_valid_blocked ethread_set_valid_idle_etcb)+ done lemma tcb_dequeue_not_queued: "\valid_queues\ tcb_sched_action tcb_sched_dequeue tptr \\_. not_queued tptr\" apply (simp add: tcb_sched_action_def | wp)+ apply (clarsimp simp: etcb_at_def tcb_sched_dequeue_def not_queued_def valid_queues_def split: option.splits) done crunch ct_in_state[wp]: set_eobject, set_tcb_queue "ct_in_state P" (simp: ct_in_state_def pred_tcb_at_def obj_at_def) crunch ct_in_state[wp]: tcb_sched_action, ethread_set "ct_in_state P" crunch not_pred_tcb[wp]: set_eobject, set_tcb_queue "\s. \ pred_tcb_at proj P t s" (simp: ct_in_state_def pred_tcb_at_def obj_at_def) crunch not_pred_tcb[wp]: tcb_sched_action, ethread_set "\s. \ pred_tcb_at proj P t s" lemma ct_in_state_def2: "ct_in_state test s = st_tcb_at test (cur_thread s) s" by (simp add: ct_in_state_def) lemma set_mcpriority_valid_sched[wp]: "\valid_sched\ set_mcpriority tptr prio \\_. valid_sched\" by (simp add: set_mcpriority_def thread_set_not_state_valid_sched) lemma set_nonmember_if_cong: "(a \ set (if P then x else y)) = (if P then a \ set x else a \ set y)" by auto lemma reschedule_preserves_valid_sched: "\ valid_sched \ reschedule_required \ \rv. valid_sched \" unfolding reschedule_required_def set_scheduler_action_def tcb_sched_action_def apply (rule hoare_pre) apply (wp|wpc)+ apply clarsimp apply (rule conjI) apply (clarsimp simp: valid_sched_2_def ct_not_in_q_2_def valid_blocked_2_def) apply (rule conjI) defer apply (clarsimp simp: valid_sched_2_def ct_not_in_q_2_def valid_blocked_2_def) apply (clarsimp simp only: etcb_at_def) apply (case_tac "ekheap s x"; simp) apply (clarsimp simp: valid_sched_2_def) apply (rule conjI) apply (clarsimp simp: valid_queues_2_def valid_sched_action_2_def tcb_sched_enqueue_def weak_valid_sched_action_2_def etcb_at_conj_is_etcb_at) apply (rule conjI) apply (clarsimp simp: ct_not_in_q_2_def) apply (clarsimp simp: valid_blocked_2_def) apply (clarsimp simp: not_queued_def) apply (erule_tac x=t in allE; simp) by (clarsimp simp: set_nonmember_if_cong tcb_sched_enqueue_def split: if_split_asm; blast) lemma set_scheduler_action_swt_weak_valid_sched_except: "\valid_sched_except_blocked and valid_blocked_except t and st_tcb_at runnable t and in_cur_domain t and simple_sched_action\ set_scheduler_action (switch_thread t) \\_.valid_sched\" apply (simp add: set_scheduler_action_def) apply wp apply (force simp: valid_sched_def ct_not_in_q_def valid_sched_action_def weak_valid_sched_action_def in_cur_domain_def ct_in_cur_domain_def switch_in_cur_domain_def valid_blocked_except_def simple_sched_action_def valid_blocked_2_def is_activatable_2_def valid_idle_etcb_def split: scheduler_action.splits) done lemma set_scheduler_action_swt_weak_valid_sched: "\valid_sched and st_tcb_at runnable t and in_cur_domain t and simple_sched_action\ set_scheduler_action (switch_thread t) \\_.valid_sched\" apply (simp add: set_scheduler_action_def) apply wp apply (clarsimp simp: valid_sched_def ct_not_in_q_def valid_sched_action_def weak_valid_sched_action_def in_cur_domain_def ct_in_cur_domain_def switch_in_cur_domain_def valid_blocked_def simple_sched_action_def split: scheduler_action.splits) done lemma set_sched_action_cnt_not_cur_thread[wp]: "\\\ set_scheduler_action choose_new_thread \\rv. not_cur_thread t\" apply (simp add: set_scheduler_action_def | wp)+ apply (simp add: not_cur_thread_def) done lemma reschedule_required_not_cur_thread[wp]: "\\\ reschedule_required \\rv. not_cur_thread t\" by (simp add: reschedule_required_def, wp) lemma possible_switch_to_valid_sched_except: "\valid_sched_except_blocked and valid_blocked_except target and st_tcb_at runnable target and not_cur_thread target and (\s. target \ idle_thread s)\ possible_switch_to target \\rv. valid_sched\" unfolding possible_switch_to_def apply (wpsimp wp: reschedule_required_valid_blocked_except set_scheduler_action_swt_weak_valid_sched_except)+ apply (clarsimp simp: etcb_at'_def not_cur_thread_2_def valid_sched_def valid_sched_action_def in_cur_domain_def ct_in_cur_domain_2_def valid_blocked_def valid_blocked_except_def split: option.splits) done lemma set_priority_valid_sched[wp]: "\valid_sched and (\s. tptr \ idle_thread s)\ set_priority tptr prio \\_. valid_sched\" apply (rule hoare_pre) apply (simp add: set_priority_def thread_set_priority_def) apply (wp gts_wp hoare_vcg_if_lift ethread_set_not_queued_valid_queues hoare_vcg_all_lift tcb_dequeue_not_queued hoare_vcg_imp_lift hoare_vcg_disj_lift tcb_sched_action_enqueue_valid_blocked ethread_set_not_domain_valid_sched_action ethread_set_not_domain_ct_in_cur_domain ethread_set_not_domain_valid_blocked_except ethread_set_not_queued_valid_sched tcb_sched_action_dequeue_valid_blocked_except tcb_sched_action_dequeue_valid_sched_not_runnable reschedule_required_valid_sched_cur_thread ethread_set_valid_idle_etcb possible_switch_to_valid_sched_except | simp add: ct_in_state_def2[symmetric])+ apply (auto simp: valid_sched_def valid_sched_action_def st_tcb_at_def obj_at_def not_cur_thread_def split: thread_state.splits) done context DetSchedSchedule_AI begin crunch simple_sched_action[wp]: set_mcpriority, cap_insert, cap_delete, option_update_thread simple_sched_action crunch idle_thread[wp]: option_update_thread, set_mcpriority, finalise_cap, cap_swap_for_delete "\(s:: det_state). P (idle_thread s)" crunch idle_thread[wp]: preemption_point "\(s:: det_state). P (idle_thread s)" (ignore: OR_choiceE simp: OR_choiceE_def wrap_ext_bool_det_ext_ext_def crunch_simps wp: crunch_wps ignore_del: preemption_point) lemma rec_del_idle_thread[wp]: "\\(s:: det_ext state). P (idle_thread s)\ rec_del call \\rv s. P (idle_thread s)\" apply (rule rec_del_preservation) apply wp+ done crunch idle_thread[wp]: cap_delete "\(s:: det_state). P (idle_thread s)" crunch valid_sched[wp]: cap_delete valid_sched lemma tc_valid_sched[wp]: "\valid_sched and simple_sched_action and (\s. a \ idle_thread s)\ invoke_tcb (ThreadControl a sl b mcp pr e f g) \\rv. valid_sched\" apply clarsimp apply wp apply wpsimp apply wpc apply (clarsimp simp: if_apply_def2 eq_commute[where a=a] option_update_thread_def | wpsimp wp: hoare_drop_imps reschedule_preserves_valid_sched hoare_lift_Pf [where f= "cur_thread" and P="\x s. x \ idle_thread s"] check_cap_inv cap_insert_valid_sched thread_set_not_state_valid_sched cong: conj_cong imp_cong)+ done end lemma possible_switch_to_valid_sched: "\valid_sched and st_tcb_at runnable target and not_cur_thread target and (\s. target \ idle_thread s)\ possible_switch_to target \\rv. valid_sched\" unfolding possible_switch_to_def apply (wpsimp wp: reschedule_required_valid_blocked set_scheduler_action_swt_weak_valid_sched | strengthen valid_blocked_valid_blocked_except)+ by (fastforce simp: etcb_at'_def not_cur_thread_2_def valid_sched_def valid_sched_action_def in_cur_domain_def ct_in_cur_domain_2_def valid_blocked_def valid_blocked_except_def split: option.splits) lemma set_thread_state_not_cur_thread[wp]: "\not_cur_thread thread\ set_thread_state ref ts \\rv. not_cur_thread thread\" apply (simp add: set_thread_state_def) apply (simp add: set_thread_state_def set_thread_state_ext_def set_object_def get_object_def | wp gts_wp)+ done crunch valid_sched[wp]: setup_reply_master valid_sched crunch valid_etcbs[wp]: setup_reply_master valid_etcbs crunch valid_queues[wp]: setup_reply_master valid_queues crunch ct_not_in_q[wp]: setup_reply_master ct_not_in_q crunch valid_sched_action[wp]: setup_reply_master valid_sched_action crunch ct_in_cur_domain[wp]: setup_reply_master ct_in_cur_domain crunch valid_blocked[wp]: setup_reply_master valid_blocked context DetSchedSchedule_AI begin crunch not_cur_thread[wp]: empty_slot "not_cur_thread thread" (wp: crunch_wps) crunch not_cur_thread[wp]: setup_reply_master, cancel_ipc "not_cur_thread thread" (wp: hoare_drop_imps mapM_x_wp simp: unless_def if_fun_split) crunch etcb_at[wp]: setup_reply_master "etcb_at P t" lemma restart_valid_sched[wp]: "\valid_sched and (\s. thread \ idle_thread s)\ restart thread \\rv. valid_sched\" apply (simp add: restart_def | wp set_thread_state_runnable_valid_queues set_thread_state_runnable_valid_sched_action set_thread_state_valid_blocked_except sts_st_tcb_at' cancel_ipc_simple2 possible_switch_to_valid_sched)+ apply (rule_tac Q="\_. valid_sched and not_cur_thread thread and (\s. thread \ idle_thread s)" in hoare_strengthen_post) apply wp apply (simp add: valid_sched_def) apply (simp add: if_fun_split) apply (rule hoare_vcg_conj_lift) apply (simp add: get_thread_state_def thread_get_def) apply wp apply (simp add: get_thread_state_def | wp hoare_drop_imps)+ apply (clarsimp simp: not_cur_thread_def valid_sched_def valid_sched_action_def is_activatable_def) apply (drule_tac test="\ts. \ activatable ts" in st_tcb_at_get_lift) apply simp apply (simp only: st_tcb_at_not) apply simp done end lemma as_user_valid_sched[wp]: "\valid_sched\ as_user tptr f \\rv. valid_sched\" apply (simp add: as_user_def set_object_def get_object_def) apply (wp | wpc)+ apply clarsimp by (fastforce simp: valid_sched_def valid_etcbs_def valid_queues_def valid_sched_action_def is_activatable_def weak_valid_sched_action_def st_tcb_at_kh_if_split st_tcb_def2 valid_blocked_def) crunch valid_sched[wp]: bind_notification "valid_sched" crunch it[wp]: suspend "\ s. P (idle_thread s)" (ignore: tcb_sched_action wp: dxo_wp_weak) context DetSchedSchedule_AI begin lemma invoke_tcb_valid_sched[wp]: "\invs and valid_sched and simple_sched_action and tcb_inv_wf ti\ invoke_tcb ti \\rv. valid_sched\" apply (cases ti, simp_all only:) apply (wp mapM_x_wp | simp | rule subset_refl | rule reschedule_preserves_valid_sched | clarsimp simp:invs_valid_objs invs_valid_global_refs idle_no_ex_cap | intro impI conjI)+ apply (rename_tac option) apply (case_tac option) apply (wp mapM_x_wp | simp | rule subset_refl | clarsimp simp:invs_valid_objs invs_valid_global_refs idle_no_ex_cap | rule reschedule_preserves_valid_sched | intro impI conjI)+ done end lemma runnable_eq_active: "runnable = active" apply (rule ext) apply (case_tac st, simp_all) done lemma handle_yield_valid_sched[wp]: "\valid_sched and ct_active\ handle_yield \\rv. valid_sched\" apply (simp add: handle_yield_def | wp tcb_sched_action_append_valid_blocked tcb_sched_action_dequeue_valid_blocked_except reschedule_required_valid_sched)+ apply (clarsimp simp: valid_sched_def ct_in_state_def valid_sched_action_def runnable_eq_active) done crunch valid_sched[wp]: store_word_offs valid_sched crunch exst[wp]: set_mrs, as_user "\s. P (exst s)" (simp: crunch_simps wp: crunch_wps) crunch valid_sched[wp]: set_mrs valid_sched (wp: valid_sched_lift) lemmas gts_drop_imp = hoare_drop_imp[where f="get_thread_state p" for p] lemma reschedule_required_switch_valid_blocked: "\\s. case scheduler_action s of switch_thread t \ valid_blocked_except t s | _ \ False\ reschedule_required \\_. valid_blocked\" apply (simp add: reschedule_required_def | wp set_scheduler_action_cnt_valid_blocked tcb_sched_action_enqueue_valid_blocked hoare_vcg_all_lift | wpc)+ apply (simp add: tcb_sched_action_def) apply wp+ apply (force simp: etcb_at_def tcb_sched_enqueue_def valid_blocked_def valid_blocked_except_def split: option.splits) done lemma reschedule_required_switch_valid_sched: "\valid_etcbs and valid_queues and weak_valid_sched_action and (\s. case scheduler_action s of switch_thread t \ valid_blocked_except t s | _ \ False) and valid_idle_etcb\ reschedule_required \\_. valid_sched\" by (simp add: valid_sched_def | wp reschedule_required_switch_valid_blocked | force)+ lemma set_scheduler_action_swt_weak_valid_sched': "\valid_sched_except_blocked and valid_blocked_except t and st_tcb_at runnable t and in_cur_domain t and simple_sched_action\ set_scheduler_action (switch_thread t) \\_.valid_sched\" apply (simp add: set_scheduler_action_def) apply wp apply (force simp: valid_sched_def ct_not_in_q_def valid_sched_action_def weak_valid_sched_action_def in_cur_domain_def ct_in_cur_domain_def switch_in_cur_domain_def valid_blocked_def valid_blocked_except_def simple_sched_action_def split: scheduler_action.splits) done lemma enqueue_thread_not_not_queued: "\\s. t = thread \ tcb_sched_action tcb_sched_enqueue thread \\_ s. \ not_queued t s \" apply (simp add: tcb_sched_action_def not_queued_def) apply wp apply (fastforce simp: etcb_at_def tcb_sched_enqueue_def split: option.splits) done crunch scheduler_action[wp]: tcb_sched_action "\s. P (scheduler_action s)" crunch valid_etcbs[wp]: possible_switch_to "valid_etcbs" (* FIXME: Move *) lemma valid_blocked_except_lift: assumes a: "\Q t. \\s. st_tcb_at Q t s\ f \\rv s. st_tcb_at Q t s\" assumes t: "\P T t. \\s. P (typ_at T t s)\ f \\rv s. P (typ_at T t s)\" and c: "\P. \\s. P (scheduler_action s)\ f \\rv s. P (scheduler_action s)\" and e: "\P. \\s. P (cur_thread s)\ f \\rv s. P (cur_thread s)\" and d: "\P. \\s. P (ready_queues s)\ f \\rv s. P (ready_queues s)\" shows "\valid_blocked_except thread\ f \\rv. valid_blocked_except thread\" apply (rule hoare_pre) apply (wps c e d) apply (simp add: valid_blocked_except_def) apply (wp hoare_weak_lift_imp hoare_vcg_ball_lift hoare_vcg_all_lift hoare_vcg_conj_lift a) apply (rule hoare_convert_imp) apply (rule typ_at_st_tcb_at_lift) apply (wp a t)+ apply (simp add: valid_blocked_except_def) done crunch valid_blocked_except[wp]: as_user "valid_blocked_except thread" (wp: valid_blocked_except_lift) (* FIXME: Move *) lemma set_simple_ko_valid_sched_action[wp]: "\valid_sched_action\ set_simple_ko f ptr ntfn \\rv. valid_sched_action\" by (wp hoare_drop_imps valid_sched_action_lift | simp add: set_simple_ko_def)+ crunch not_cur_thread[wp]: cap_insert_ext "not_cur_thread t" crunch not_cur_thread[wp]: cap_insert, set_extra_badge "not_cur_thread t" (wp: hoare_drop_imps) lemma transfer_caps_not_cur_thread[wp]: "\not_cur_thread t\ transfer_caps info caps ep recv recv_buf \\rv. not_cur_thread t\" by (simp add: transfer_caps_def | wp transfer_caps_loop_pres | wpc)+ crunch not_cur_thread[wp]: as_user "not_cur_thread t" (wp: crunch_wps simp: crunch_simps ignore: const_on_failure) crunch (in DetSchedSchedule_AI) not_cur_thread[wp] : do_ipc_transfer "not_cur_thread t" (wp: crunch_wps simp: crunch_simps ignore: const_on_failure) lemmas set_thread_state_active_valid_sched_except_blocked = set_thread_state_runnable_valid_sched_except_blocked[simplified runnable_eq_active] lemma set_thread_state_runnable_valid_blocked: "\valid_blocked and st_tcb_at runnable ref and (\s. runnable ts)\ set_thread_state ref ts \\_. valid_blocked\" apply (simp add: set_thread_state_def) apply (rule hoare_seq_ext[OF _ gets_the_get_tcb_sp]) apply (rule_tac B="\rv. valid_blocked and st_tcb_at runnable ref" in hoare_seq_ext[rotated]) apply (wp set_object_wp) apply (clarsimp simp: valid_blocked_def not_queued_def runnable_eq_active pred_tcb_at_def st_tcb_at_kh_def obj_at_kh_def obj_at_def) apply (simp add: set_thread_state_ext_def) apply (rule hoare_seq_ext[OF _ gts_sp]) apply (rule_tac S="runnable ts" in hoare_gen_asm_spec) apply (clarsimp simp: pred_tcb_at_def obj_at_def) apply clarsimp by wpsimp lemma set_thread_state_runnable_valid_sched: "\valid_sched and st_tcb_at runnable ref and (\s. runnable ts)\ set_thread_state ref ts \\_. valid_sched\" apply (simp add: valid_sched_def | wp set_thread_state_runnable_valid_queues set_thread_state_runnable_valid_sched_action set_thread_state_runnable_valid_blocked)+ done context DetSchedSchedule_AI begin lemma update_waiting_ntfn_valid_sched[wp]: "\ \s. valid_sched s \ hd queue \ idle_thread s \ (scheduler_action s = resume_cur_thread \ hd queue \ cur_thread s)\ update_waiting_ntfn ntfnptr queue badge val \ \_. valid_sched \" apply (simp add: update_waiting_ntfn_def) apply (wp sts_st_tcb_at' possible_switch_to_valid_sched_except set_thread_state_runnable_valid_sched set_thread_state_runnable_valid_queues set_thread_state_runnable_valid_sched_action set_thread_state_valid_blocked_except | clarsimp)+ apply (clarsimp simp add: valid_sched_def not_cur_thread_def ct_not_in_q_def) done end crunch valid_sched[wp]: dec_domain_time valid_sched lemma timer_tick_valid_sched[wp]: "\valid_sched\ timer_tick \\rv. valid_sched\" apply (simp add: timer_tick_def crunch_simps thread_set_time_slice_def trans_state_update'[symmetric] ethread_set_def set_eobject_def | wp gts_wp reschedule_required_valid_sched tcb_sched_action_append_valid_blocked | wpc | rule hoare_strengthen_post, rule dec_domain_time_valid_sched, simp add: valid_sched_def valid_sched_action_def)+ by (fastforce simp: valid_sched_def valid_etcbs_def valid_queues_def pred_tcb_weakenE valid_sched_action_def weak_valid_sched_action_def etcb_at_def is_etcb_at_def get_etcb_def in_cur_domain_def ct_in_cur_domain_2_def switch_in_cur_domain_def valid_idle_etcb_2_def split: option.splits) lemma cancel_badged_sends_filterM_valid_etcbs[wp]: "\valid_etcbs\ filterM (\t. do st \ get_thread_state t; if blocking_ipc_badge st = badge then do y \ set_thread_state t Structures_A.thread_state.Restart; y \ (tcb_sched_action tcb_sched_enqueue t); return False od else return True od) xs \\rv. valid_etcbs\" apply (rule rev_induct[where xs=xs]) apply simp apply (clarsimp simp: filterM_append) apply (wp sts_st_tcb_at' | simp)+ done lemma cancel_badged_sends_filterM_valid_queues[wp]: "\valid_queues\ filterM (\t. do st \ get_thread_state t; if blocking_ipc_badge st = badge then do y \ set_thread_state t Structures_A.thread_state.Restart; y \ (tcb_sched_action tcb_sched_enqueue t); return False od else return True od) xs \\rv. valid_queues\" apply (rule rev_induct[where xs=xs]) apply simp apply (clarsimp simp: filterM_append) apply (wp set_thread_state_runnable_valid_queues sts_st_tcb_at' | simp)+ done lemma cancel_badged_sends_filterM_weak_valid_sched_action[wp]: "\weak_valid_sched_action\ filterM (\t. do st \ get_thread_state t; if blocking_ipc_badge st = badge then do y \ set_thread_state t Structures_A.thread_state.Restart; y \ (tcb_sched_action tcb_sched_enqueue t); return False od else return True od) xs \\rv. weak_valid_sched_action\" apply (rule rev_induct[where xs=xs]) apply simp apply (clarsimp simp: filterM_append) apply (wp set_thread_state_runnable_weak_valid_sched_action sts_st_tcb_at' | simp)+ done lemma cancel_badged_sends_filterM_ct_in_cur_domain[wp]: "\ct_in_cur_domain\ filterM (\t. do st \ get_thread_state t; if blocking_ipc_badge st = badge then do y \ set_thread_state t Structures_A.thread_state.Restart; y \ (tcb_sched_action tcb_sched_enqueue t); return False od else return True od) xs \\rv. ct_in_cur_domain\" apply (rule rev_induct[where xs=xs]) apply simp apply (clarsimp simp: filterM_append) apply (wp | simp)+ done lemma cancel_badged_sends_filterM_valid_blocked[wp]: "\valid_blocked\ filterM (\t. do st \ get_thread_state t; if blocking_ipc_badge st = badge then do y \ set_thread_state t Structures_A.thread_state.Restart; y \ (tcb_sched_action tcb_sched_enqueue t); return False od else return True od) xs \\rv. valid_blocked\" apply (rule rev_induct[where xs=xs]) apply simp apply (clarsimp simp: filterM_append) apply (wp tcb_sched_action_enqueue_valid_blocked set_thread_state_valid_blocked_except | simp)+ done lemma cancel_badged_sends_filterM_valid_idle_etcb[wp]: notes valid_idle_etcb_lift[wp del] shows "\valid_idle_etcb\ filterM (\t. do st \ get_thread_state t; if blocking_ipc_badge st = badge then do y \ set_thread_state t Structures_A.thread_state.Restart; y \ (tcb_sched_action tcb_sched_enqueue t); return False od else return True od) xs \\rv. valid_idle_etcb\" apply (rule rev_induct[where xs=xs]) apply simp apply (clarsimp simp: filterM_append) apply (wp | simp | wp valid_idle_etcb_lift)+ done lemma cancel_badged_sends_valid_sched[wp]: shows "\valid_sched\ cancel_badged_sends epptr badge \\rv. valid_sched\" apply (simp add: cancel_badged_sends_def) apply (wp hoare_drop_imps reschedule_required_valid_sched | wpc | simp)+ apply (simp_all add: valid_sched_def valid_sched_action_def) done context DetSchedSchedule_AI begin lemma cap_revoke_valid_sched[wp]: "\valid_sched and simple_sched_action\ cap_revoke slot \\rv. valid_sched\" apply (rule hoare_strengthen_post) apply (rule validE_valid, rule cap_revoke_preservation) apply (wpsimp wp: preemption_point_inv')+ done lemma cap_revoke_simple_sched_action[wp]: "\simple_sched_action\ cap_revoke slot \\rv. simple_sched_action\" by (wp cap_revoke_preservation preemption_point_inv' | fastforce)+ end lemma thread_set_state_eq_valid_queues: "(\x. tcb_state (f x) = ts) \ \valid_queues and st_tcb_at ((=) ts) tptr\ thread_set f tptr \\rv. valid_queues\" apply (simp add: thread_set_def set_object_def get_object_def) apply wp apply (fastforce simp: valid_queues_def st_tcb_at_kh_if_split st_tcb_def2) done lemma thread_set_state_eq_valid_sched_action: "(\x. tcb_state (f x) = ts) \ \valid_sched_action and st_tcb_at ((=) ts) tptr\ thread_set f tptr \\rv. valid_sched_action\" apply (simp add: thread_set_def set_object_def get_object_def) apply wp apply (fastforce simp: valid_sched_action_def weak_valid_sched_action_def is_activatable_def st_tcb_at_kh_if_split st_tcb_def2) done lemma thread_set_state_eq_ct_in_cur_domain: "(\x. tcb_state (f x) = ts) \ \ct_in_cur_domain and st_tcb_at ((=) ts) tptr\ thread_set f tptr \\rv. ct_in_cur_domain\" apply (simp add: thread_set_def set_object_def get_object_def) apply wp apply (fastforce simp: ct_in_cur_domain_def st_tcb_at_kh_if_split st_tcb_def2) done lemma thread_set_state_eq_valid_blocked: "(\x. tcb_state (f x) = ts) \ \valid_blocked and st_tcb_at ((=) ts) tptr\ thread_set f tptr \\rv. valid_blocked\" apply (simp add: thread_set_def set_object_def get_object_def) apply wp apply (fastforce simp: valid_blocked_def st_tcb_at_kh_if_split st_tcb_def2) done crunch etcb_at[wp]: thread_set "etcb_at P t" context DetSchedSchedule_AI begin lemma thread_set_state_eq_valid_sched: "(\x. tcb_state (f x) = ts) \ \valid_sched and st_tcb_at ((=) ts) tptr\ thread_set f tptr \\rv. valid_sched\" apply (simp add: valid_sched_def) apply (wp thread_set_state_eq_valid_queues thread_set_state_eq_valid_blocked thread_set_state_eq_valid_sched_action thread_set_state_eq_ct_in_cur_domain | simp)+ done end crunch exst[wp]: thread_set "\s. P (exst s)" lemma ethread_set_not_switch_switch_in_cur_domain: "\switch_in_cur_domain and (\s. scheduler_action s \ switch_thread tptr)\ ethread_set f tptr \\_. switch_in_cur_domain\" apply (simp add: ethread_set_def set_eobject_def | wp)+ apply (clarsimp simp: switch_in_cur_domain_def in_cur_domain_def is_etcb_at_def etcb_at_def get_etcb_def) done lemma ethread_set_not_cur_ct_in_cur_domain: "\ct_in_cur_domain and not_cur_thread tptr\ ethread_set f tptr \\_. ct_in_cur_domain\" apply (simp add: ethread_set_def set_eobject_def | wp)+ apply (clarsimp simp: ct_in_cur_domain_def in_cur_domain_def not_cur_thread_def etcb_at_def get_etcb_def) done lemma ethread_set_valid_blocked: "\valid_blocked\ ethread_set f tptr \\_. valid_blocked\" by (wp valid_blocked_lift | simp add: ethread_set_def set_eobject_def)+ lemma ethread_set_inactive_valid_idle_etcb: notes valid_idle_etcb_lift[wp del] shows "\valid_idle_etcb and valid_idle and st_tcb_at inactive tptr\ ethread_set f tptr \\_. valid_idle_etcb\" apply(simp add: ethread_set_def set_eobject_def) apply wp apply(clarsimp simp: get_etcb_def valid_idle_etcb_def etcb_at'_def valid_idle_def pred_tcb_at_def obj_at_def) done lemma ethread_set_inactive_valid_sched: "\valid_sched and valid_idle and st_tcb_at inactive tptr\ ethread_set f tptr \\_. valid_sched\" apply (simp add: valid_sched_def valid_sched_action_def | wp ethread_set_not_queued_valid_queues ethread_set_not_switch_switch_in_cur_domain ethread_set_not_cur_ct_in_cur_domain ethread_set_valid_blocked ethread_set_inactive_valid_idle_etcb)+ apply (force simp: valid_idle_def st_tcb_at_def obj_at_def not_cur_thread_def is_activatable_def weak_valid_sched_action_def valid_queues_def not_queued_def)+ done lemma thread_set_not_idle_valid_idle: "\valid_idle and (\s. tptr \ idle_thread s)\ thread_set f tptr \\_. valid_idle\" apply (simp add: thread_set_def set_object_def get_object_def, wp) apply (clarsimp simp: valid_idle_def pred_tcb_at_def obj_at_def) done lemma thread_set_st_tcb_at: "(\x. P (tcb_state (f x))) \ \\\ thread_set f tptr \\_. st_tcb_at P tptr\" apply (simp add: thread_set_def set_object_def get_object_def, wp) apply (clarsimp simp: valid_idle_def st_tcb_at_def obj_at_def) done crunch valid_sched[wp]: cap_move valid_sched context DetSchedSchedule_AI begin lemma invoke_cnode_valid_sched: "\valid_sched and invs and valid_cnode_inv a and simple_sched_action\ invoke_cnode a \\rv. valid_sched\" apply (simp add: invoke_cnode_def) apply (rule hoare_pre) apply wpc apply (simp add: liftE_def | (wp hoare_vcg_all_lift)+ | wp (once) hoare_drop_imps | wpc)+ apply force done end crunch valid_sched[wp]: set_extra_badge valid_sched (wp: crunch_wps) lemma transfer_caps_valid_sched: "\valid_sched\ transfer_caps info caps ep recv recv_buf \\rv. valid_sched\" apply (simp add: transfer_caps_def | wp transfer_caps_loop_pres | wpc)+ done context DetSchedSchedule_AI begin crunch valid_sched[wp]: do_ipc_transfer, handle_fault_reply valid_sched (wp: crunch_wps) lemma thread_set_ct_active_wp: "\ ct_active \ thread_set (tcb_fault_update u) t \\rv. ct_active \" by (wp ct_in_state_thread_state_lift thread_set_no_change_tcb_state, simp) lemma do_reply_transfer_valid_sched[wp]: "\valid_sched and valid_objs and ct_active and cte_wp_at (is_reply_cap_to t') slot and (\s. receiver \ idle_thread s)\ do_reply_transfer sender receiver slot grant \\rv. valid_sched\" apply (simp add: do_reply_transfer_def) apply (wp set_thread_state_not_runnable_valid_sched sts_st_tcb_at' cap_delete_one_reply_st_tcb_at do_ipc_transfer_non_null_cte_wp_at2 thread_set_not_state_valid_sched thread_set_no_change_tcb_state possible_switch_to_valid_sched_except | wpc | clarsimp split del: if_split)+ apply (wp set_thread_state_runnable_valid_queues set_thread_state_runnable_valid_sched_action set_thread_state_valid_blocked_except sts_st_tcb_at')[1] apply (rule_tac Q="\_. valid_sched and not_cur_thread receiver and (\s. receiver \ idle_thread s)" in hoare_strengthen_post) apply wp apply (simp add: valid_sched_def) apply (wp possible_switch_to_valid_sched_except)+ apply simp apply (rule conjI) apply clarsimp apply (rule_tac P="valid_sched and st_tcb_at (\ts. \ runnable ts) receiver and ct_active and (\s. receiver \ idle_thread s)" in hoare_weaken_pre) apply (wp set_thread_state_runnable_valid_queues set_thread_state_runnable_valid_sched_action set_thread_state_valid_blocked_except sts_st_tcb_at')[1] apply (fastforce simp: valid_sched_def ct_in_state_def st_tcb_at_def not_cur_thread_def obj_at_def) apply clarsimp apply (wp set_thread_state_not_runnable_valid_sched) apply simp apply (wp thread_set_not_state_valid_sched thread_set_no_change_tcb_state cap_delete_one_reply_st_tcb_at thread_set_ct_active_wp | simp add: ct_in_state_def | wps)+ apply (wp hoare_drop_imps hoare_vcg_all_lift)[1] apply (simp add: get_thread_state_def thread_get_def | wp)+ apply (clarsimp simp: ct_in_state_def cte_wp_at_caps_of_state not_cur_thread_def) apply (fastforce simp: st_tcb_def2) done end lemma set_thread_state_not_queued_valid_queues: "\valid_queues and not_queued thread\ set_thread_state thread ts \\rv. valid_queues\" apply (simp add: set_thread_state_def) apply (wp | simp add: set_thread_state_ext_def set_object_def get_object_def)+ apply (fastforce simp: valid_queues_def st_tcb_at_kh_if_split not_queued_def) done lemma set_scheduler_action_cnt_is_activatable'[wp]: "\\\ set_scheduler_action choose_new_thread \\r s. is_activatable (t s) s\" apply (wp set_scheduler_action_wp) apply (simp add: is_activatable_def) done lemma set_scheduler_action_cnt_switch_in_cur_domain[wp]: "\\\ set_scheduler_action choose_new_thread \\_. switch_in_cur_domain\" by (simp add: set_scheduler_action_def, wp, simp) lemma set_thread_state_sched_act_not_valid_sched_action: "\valid_sched_action and scheduler_act_not thread\ set_thread_state thread ts \\rv. valid_sched_action\" apply (simp add: valid_sched_action_def set_thread_state_def) apply (rule hoare_conjI) apply (wp gts_wp | simp add: set_thread_state_ext_def set_object_def get_object_def)+ apply (clarsimp simp: weak_valid_sched_action_def st_tcb_at_kh_if_split scheduler_act_not_def is_activatable_def pred_tcb_at_def obj_at_def) apply (wp gts_wp | simp add: set_thread_state_ext_def set_object_def get_object_def)+ apply (clarsimp simp: weak_valid_sched_action_def st_tcb_at_kh_if_split scheduler_act_not_def is_activatable_def) done lemma set_thread_state_sched_act_not_valid_sched: "\valid_sched and scheduler_act_not thread and not_queued thread and K (\ runnable ts)\ set_thread_state thread ts \\rv. valid_sched\" by (simp add: valid_sched_def | wp set_thread_state_not_queued_valid_queues set_thread_state_not_runnable_valid_blocked set_thread_state_sched_act_not_valid_sched_action)+ lemma setup_caller_cap_sched_act_not_valid_sched: "\valid_sched and scheduler_act_not sender and not_queued sender\ setup_caller_cap sender receiver grant \\rv. valid_sched\" by (auto simp: setup_caller_cap_def | wp set_thread_state_sched_act_not_valid_sched)+ (* strong in case of tcb_domain t = tcb_domain target *) lemma possible_switch_to_sched_act_not[wp]: "\K(t \ target) and scheduler_act_not t\ possible_switch_to target \\_. scheduler_act_not t\" apply (simp add: possible_switch_to_def reschedule_required_def set_scheduler_action_def tcb_sched_action_def split del: if_split | wp | wpc)+ apply (clarsimp simp: etcb_at_def scheduler_act_not_def split: option.splits) done lemma possible_switch_to_not_queued[wp]: "\not_queued t and scheduler_act_not t and K(target \ t)\ possible_switch_to target \\_. not_queued t\" apply (simp add: possible_switch_to_def reschedule_required_def set_scheduler_action_def tcb_sched_action_def split del: if_split | wp | wpc)+ by (fastforce simp: etcb_at_def tcb_sched_enqueue_def simple_sched_action_def not_queued_def scheduler_act_not_def split: option.splits) lemma set_thread_state_ready_queues[wp]: "\\s :: det_state. P (ready_queues s)\ set_thread_state thread ts \\r s. P (ready_queues s)\" apply (simp add: set_thread_state_def) apply (simp add: set_thread_state_ext_def[abs_def] reschedule_required_def set_scheduler_action_def set_object_def get_object_def) apply (wp | wpc | simp add: tcb_sched_action_def)+ done crunch scheduler_act[wp]: set_extra_badge,cap_insert "\s :: det_state. P (scheduler_action s)" (wp: crunch_wps) context DetSchedSchedule_AI begin crunch scheduler_act[wp]: do_ipc_transfer "\s :: det_state. P (scheduler_action s)" (wp: crunch_wps ignore: const_on_failure rule: transfer_caps_loop_pres) crunch ready_queues[wp]: cap_insert,set_extra_badge,do_ipc_transfer, set_simple_ko, thread_set, setup_caller_cap "\s :: det_state. P (ready_queues s)" (wp: crunch_wps ignore: const_on_failure rule: transfer_caps_loop_pres) end crunch sched_act_not[wp]: set_thread_state_ext "scheduler_act_not t" (ignore: set_scheduler_action simp: set_scheduler_action_def if_fun_split scheduler_act_not_def wp: gts_wp) crunch sched_act_not[wp]: set_thread_state, set_simple_ko "scheduler_act_not t" (wp: hoare_drop_imps) context DetSchedSchedule_AI begin lemma send_ipc_valid_sched: "\valid_sched and st_tcb_at active thread and scheduler_act_not thread and not_queued thread and (ct_active or ct_idle) and invs\ send_ipc block call badge can_grant can_grant_reply thread epptr \\rv. valid_sched\" apply (simp add: send_ipc_def) apply (wp set_thread_state_sched_act_not_valid_sched | wpc)+ apply ((wp set_thread_state_sched_act_not_valid_sched setup_caller_cap_sched_act_not_valid_sched possible_switch_to_valid_sched_except hoare_vcg_if_lift2 hoare_drop_imps | simp)+)[5] apply (wp set_thread_state_runnable_valid_queues set_thread_state_runnable_valid_sched_action set_thread_state_valid_blocked_except sts_st_tcb_at') apply (clarsimp simp: conj.commute eq_commute) apply (rename_tac recvr q recv_state) apply (rule_tac Q="\_. valid_sched and scheduler_act_not thread and not_queued thread and (\s. recvr \ cur_thread s) and (\s. recvr \ idle_thread s \ recvr \ thread)" in hoare_strengthen_post) apply ((wp thread_set_ct_active_wp|wpc)+)[1] apply (clarsimp simp: valid_sched_def) apply (clarsimp simp: ct_in_state_def cte_wp_at_caps_of_state not_cur_thread_def) apply (simp | wp gts_wp hoare_vcg_all_lift)+ apply (rename_tac recvr q recv_state) apply (wp hoare_vcg_imp_lift) apply ((simp add: set_simple_ko_def set_object_def | wp hoare_drop_imps | wpc)+)[1] apply (wpsimp wp: hoare_vcg_imp_lift get_object_wp | simp add: get_simple_ko_def | wpc | wp (once) hoare_vcg_all_lift)+ apply (subst st_tcb_at_kh_simp[symmetric]) apply (clarsimp simp: st_tcb_at_kh_if_split pred_tcb_at_def2 obj_at_def a_type_def valid_sched_def valid_sched_action_def weak_valid_sched_action_def scheduler_act_not_def split: kernel_object.splits)+ apply (rename_tac tcb recvr q rtcb ep recv_pl) apply (case_tac "recvr = idle_thread s") subgoal by (fastforce dest: invs_valid_idle simp: valid_idle_def pred_tcb_at_def obj_at_def) apply (case_tac "recvr = cur_thread s") subgoal by (fastforce simp: ct_in_state_def st_tcb_at_def2 obj_at_def) apply (clarsimp simp: is_activatable_def) done end lemma thread_set_tcb_fault_set_invs: "valid_fault f \ \invs\ thread_set (tcb_fault_update (\_. Some f)) t \\rv. invs\" apply (rule thread_set_invs_trivial) apply (clarsimp simp: ran_tcb_cap_cases)+ done context DetSchedSchedule_AI begin lemma send_fault_ipc_valid_sched[wp]: "\valid_sched and st_tcb_at active tptr and scheduler_act_not tptr and not_queued tptr and (ct_active or ct_idle) and invs and (\_. valid_fault fault)\ send_fault_ipc tptr fault \\_. valid_sched\" apply (simp add: send_fault_ipc_def Let_def) apply (wp send_ipc_valid_sched thread_set_not_state_valid_sched thread_set_no_change_tcb_state hoare_gen_asm'[OF thread_set_tcb_fault_set_invs] hoare_drop_imps hoare_vcg_all_lift_R ct_in_state_thread_state_lift thread_set_no_change_tcb_state hoare_vcg_disj_lift | wpc | simp | wps)+ done crunch valid_sched[wp]: delete_caller_cap valid_sched end lemma handle_double_fault_valid_queues: "\valid_queues and not_queued tptr\ handle_double_fault tptr ex1 ex2 \\rv. valid_queues\" apply (simp add: handle_double_fault_def set_thread_state_def) apply (wp | simp add: set_thread_state_ext_def set_object_def get_object_def)+ apply (fastforce simp: valid_queues_def st_tcb_at_kh_if_split not_queued_def) done lemma handle_double_fault_valid_sched_action: "\valid_sched_action and scheduler_act_not tptr\ handle_double_fault tptr ex1 ex2 \\rv. valid_sched_action\" apply (simp add: handle_double_fault_def set_thread_state_def) apply (wp gts_wp | simp add: set_thread_state_ext_def set_object_def get_object_def)+ apply (clarsimp simp: valid_sched_action_def weak_valid_sched_action_def is_activatable_def pred_tcb_at_def obj_at_def st_tcb_at_kh_if_split scheduler_act_not_def split: scheduler_action.splits) done lemma handle_double_fault_valid_sched: "\valid_sched and not_queued tptr and scheduler_act_not tptr\ handle_double_fault tptr ex1 ex2 \\rv. valid_sched\" apply (simp add: valid_sched_def) including no_pre apply (wp handle_double_fault_valid_queues handle_double_fault_valid_sched_action set_thread_state_not_runnable_valid_blocked | rule hoare_conjI | simp add: handle_double_fault_def | fastforce simp: simple_sched_action_def)+ done lemma send_fault_ipc_error_sched_act_not[wp]: "\scheduler_act_not t\ send_fault_ipc tptr fault -, \\rv. scheduler_act_not t\" by (simp add: send_fault_ipc_def Let_def | (wp hoare_drop_imps hoare_vcg_all_lift_R)+ | wpc)+ lemma send_fault_ipc_error_cur_thread[wp]: "\\s. P (cur_thread s)\ send_fault_ipc tptr fault -, \\rv s. P (cur_thread s)\" by (simp add: send_fault_ipc_def Let_def | (wp hoare_drop_imps hoare_vcg_all_lift_R)+ | wpc)+ lemma send_fault_ipc_error_not_queued[wp]: "\not_queued t\ send_fault_ipc tptr fault -, \\rv. not_queued t\" by (simp add: send_fault_ipc_def Let_def | (wp hoare_drop_imps hoare_vcg_all_lift_R)+ | wpc)+ context DetSchedSchedule_AI begin lemma handle_fault_valid_sched: "\valid_sched and st_tcb_at active thread and not_queued thread and (ct_active or ct_idle) and scheduler_act_not thread and invs and (\_. valid_fault ex)\ handle_fault thread ex \\rv. valid_sched\" unfolding handle_fault_def by (simp add: handle_fault_def | wp handle_double_fault_valid_sched send_fault_ipc_valid_sched)+ end lemma idle_not_queued'': "\valid_idle s; sym_refs (state_refs_of s); queue \ {rt} \ state_refs_of s ptr\ \ idle_thread s \ queue" by (frule idle_no_refs, fastforce simp: valid_idle_def sym_refs_def) context DetSchedSchedule_AI begin lemma send_signal_valid_sched[wp]: "\ valid_sched and invs \ send_signal ntfnptr badge \ \_. valid_sched \" apply (simp add: send_signal_def) apply (wp get_simple_ko_wp possible_switch_to_valid_sched_except set_thread_state_runnable_valid_queues set_thread_state_runnable_valid_sched_action set_thread_state_valid_blocked_except sts_st_tcb_at' gts_wp | wpc | clarsimp)+ apply (rename_tac ntfn a st) apply (rule_tac Q="\rv s. valid_sched s \ a \ idle_thread s \ not_cur_thread a s" in hoare_strengthen_post) apply (wp gts_wp get_simple_ko_wp | simp add: valid_sched_def)+ apply (clarsimp) apply (rule conjI, clarsimp, rule conjI) apply (frule invs_valid_idle) apply (clarsimp simp: receive_blocked_def split: thread_state.splits) apply (drule (1) st_tcb_at_idle_thread, clarsimp) apply (clarsimp simp: receive_blocked_def split: thread_state.splits) apply ((auto simp: pred_tcb_at_def obj_at_def valid_sched_action_def is_activatable_def not_cur_thread_def | drule sym)+)[1] apply (clarsimp) apply (frule invs_valid_idle) apply (drule_tac ptr=ntfnptr and rt=NTFNSignal and queue="set x" in idle_not_queued'') apply (clarsimp simp: invs_sym_refs) apply (simp add: state_refs_of_def obj_at_def) apply (frule invs_valid_objs) apply (simp add: valid_objs_def obj_at_def) apply (drule_tac x = ntfnptr in bspec) apply (simp add: dom_def) apply (clarsimp simp: valid_obj_def valid_ntfn_def) apply (drule hd_in_set) apply (rule conjI, clarsimp) apply (clarsimp) apply (cut_tac t="hd x" and P="\st. \activatable st" in ntfn_queued_st_tcb_at) apply ((simp add: obj_at_def ntfn_q_refs_of_def invs_def valid_state_def valid_pspace_def)+)[4] apply simp apply (clarsimp simp add: valid_sched_def valid_sched_action_def is_activatable_def st_tcb_def2) done crunch valid_sched[wp]: complete_signal valid_sched (ignore: resetTimer ackInterrupt wp: gts_wp hoare_drop_imps simp: op_equal pred_tcb_weakenE hoare_if_r_and) lemma receive_ipc_valid_sched: "\valid_sched and st_tcb_at active thread and ct_active and not_queued thread and scheduler_act_not thread and invs\ receive_ipc thread cap is_blocking \\rv. valid_sched\" supply option.case_cong_weak[cong] apply (simp add: receive_ipc_def) including no_pre apply (wp | wpc | simp)+ apply (wp set_thread_state_sched_act_not_valid_sched | wpc)+ apply ((wp set_thread_state_sched_act_not_valid_sched setup_caller_cap_sched_act_not_valid_sched | simp add: do_nbrecv_failed_transfer_def)+)[2] apply ((wp possible_switch_to_valid_sched_except sts_st_tcb_at' hoare_drop_imps set_thread_state_runnable_valid_queues set_thread_state_runnable_valid_sched_action set_thread_state_valid_blocked_except | simp | wpc)+)[3] apply (rule_tac Q="\_. valid_sched and scheduler_act_not (sender) and not_queued (sender) and not_cur_thread (sender) and (\s. sender \ idle_thread s)" in hoare_strengthen_post) apply wp apply (simp add: valid_sched_def) apply ((wp | wpc)+)[1] apply (simp | wp gts_wp hoare_vcg_all_lift)+ apply (wp hoare_vcg_imp_lift) apply ((simp add: set_simple_ko_def set_object_def | wp hoare_drop_imps | wpc)+)[1] apply (wp hoare_vcg_imp_lift get_object_wp set_thread_state_sched_act_not_valid_sched gbn_wp | simp add: get_simple_ko_def do_nbrecv_failed_transfer_def a_type_def split: kernel_object.splits | wpc | wp (once) hoare_vcg_all_lift hoare_vcg_ex_lift)+ apply (subst st_tcb_at_kh_simp[symmetric])+ apply (clarsimp simp: st_tcb_at_kh_if_split default_notification_def default_ntfn_def isActive_def) apply (rename_tac xh xi xj) apply (drule_tac t="hd xh" and P'="\ts. \ active ts" in st_tcb_weakenE) apply clarsimp apply (simp only: st_tcb_at_not) apply (subgoal_tac "hd xh \ idle_thread s") apply (fastforce simp: valid_sched_def valid_sched_action_def weak_valid_sched_action_def valid_queues_def st_tcb_at_not ct_in_state_def not_cur_thread_def runnable_eq_active not_queued_def scheduler_act_not_def split: scheduler_action.splits) (* clag from send_signal_valid_sched *) apply clarsimp apply (frule invs_valid_idle) apply (drule_tac ptr=xc in idle_not_queued) apply (clarsimp simp: invs_sym_refs) apply (simp add: state_refs_of_def obj_at_def) apply (frule invs_valid_objs) apply (simp add: valid_objs_def obj_at_def) apply (drule_tac x = xc in bspec) apply (simp add: dom_def) apply (clarsimp simp: valid_obj_def valid_ntfn_def) apply (drule hd_in_set) apply simp done end crunch valid_sched: receive_signal valid_sched (wp: set_thread_state_sched_act_not_valid_sched) crunch cur_thread[wp]: delete_caller_cap "\s. P (cur_thread s)" context DetSchedSchedule_AI begin crunch sched_act_not[wp]: delete_caller_cap "scheduler_act_not t" (simp: unless_def wp: hoare_drop_imps set_scheduler_action_wp mapM_x_wp ignore: set_scheduler_action) end lemma tcb_sched_action_enqueue_not_queued: "\not_queued t and K (thread \ t)\ tcb_sched_action tcb_sched_enqueue thread \\rv. not_queued t\" apply (simp add: tcb_sched_action_def | wp)+ apply (clarsimp simp: etcb_at_def tcb_sched_enqueue_def not_queued_def split: option.splits) done lemma reschedule_required_not_queued: "\not_queued t and scheduler_act_not t\ reschedule_required \\rv. not_queued t\" apply (simp add: reschedule_required_def tcb_sched_action_def set_scheduler_action_def | wp | wpc)+ apply (clarsimp simp: etcb_at_def tcb_sched_enqueue_def not_queued_def scheduler_act_not_def split: option.splits) done context DetSchedSchedule_AI begin lemma cancel_all_ipc_not_queued: "\st_tcb_at active t and valid_objs and not_queued t and scheduler_act_not t and sym_refs \ state_refs_of\ cancel_all_ipc epptr \\rv. not_queued t\" apply (simp add: cancel_all_ipc_def) apply (wp reschedule_required_not_queued | wpc | simp)+ apply (rule hoare_gen_asm) apply (rule_tac S="set queue - {t}" in mapM_x_wp) apply (wp tcb_sched_action_enqueue_not_queued | clarsimp)+ apply (erule notE, assumption) apply (wp reschedule_required_not_queued | simp add: get_ep_queue_def)+ apply (rule hoare_gen_asm) apply (rule_tac S="set queue - {t}" in mapM_x_wp) apply (wp tcb_sched_action_enqueue_not_queued | clarsimp)+ apply (erule notE, assumption) apply (wp hoare_vcg_imp_lift | simp add: get_ep_queue_def get_simple_ko_def a_type_def get_object_def split: kernel_object.splits | wpc | wp (once) hoare_vcg_all_lift)+ apply safe apply (rename_tac xa) apply (drule_tac P="\ts. \ active ts" and ep="SendEP xa" in ep_queued_st_tcb_at[rotated, rotated]) apply (simp_all only: st_tcb_at_not) apply (simp add: obj_at_def)+ apply (rename_tac xa) apply (drule_tac P="\ts. \ active ts" and ep="RecvEP xa" in ep_queued_st_tcb_at[rotated, rotated]) apply (simp_all only: st_tcb_at_not) apply (fastforce simp: obj_at_def)+ done end crunch not_queued[wp]: set_simple_ko "not_queued t" (wp: hoare_drop_imps) lemma cancel_all_signals_not_queued: "\st_tcb_at active t and valid_objs and not_queued t and scheduler_act_not t and sym_refs \ state_refs_of\ cancel_all_signals epptr \\rv. not_queued t\" apply (simp add: cancel_all_signals_def) apply (wp reschedule_required_not_queued | wpc | simp)+ apply (rename_tac list) apply (rule_tac P="(t \ set list)" in hoare_gen_asm) apply (rule_tac S="set list - {t}" in mapM_x_wp) apply (wp tcb_sched_action_enqueue_not_queued | clarsimp)+ apply (wp hoare_vcg_imp_lift | simp add: get_simple_ko_def get_object_def a_type_def split: kernel_object.splits | wpc | wp (once) hoare_vcg_all_lift)+ apply safe apply (rename_tac ep x y) apply (drule_tac P="\ts. \ active ts" and ep=ep in ntfn_queued_st_tcb_at[rotated, rotated]) apply (simp_all only: st_tcb_at_not) apply (fastforce simp: obj_at_def)+ done lemma unbind_maybe_notification_valid_objs: "\valid_objs\ unbind_maybe_notification ptr \\rv. valid_objs\" unfolding unbind_maybe_notification_def apply (wp thread_set_valid_objs_triv set_simple_ko_valid_objs hoare_drop_imp get_simple_ko_wp | wpc | simp add: tcb_cap_cases_def | strengthen unbind_notification_valid_objs_helper)+ apply (clarsimp) apply (erule (1) obj_at_valid_objsE) apply (clarsimp simp:valid_obj_def valid_tcb_def)+ done lemma unbind_maybe_notification_sym_refs[wp]: "\\s. sym_refs (state_refs_of s) \ valid_objs s\ unbind_maybe_notification a \\rv s. sym_refs (state_refs_of s)\" apply (simp add: unbind_maybe_notification_def) apply (rule hoare_seq_ext [OF _ get_simple_ko_sp]) apply (rule hoare_pre) apply (wp | wpc | clarsimp)+ apply (rule conjI) apply clarsimp apply (rule delta_sym_refs, assumption) apply (clarsimp split: if_split_asm) apply (frule ko_at_state_refs_ofD, simp) apply (clarsimp split: if_split_asm) apply (frule ko_at_state_refs_ofD, simp) apply (fastforce simp: symreftype_inverse' dest!: refs_in_ntfn_q_refs) apply clarsimp apply (rule delta_sym_refs, assumption) apply (clarsimp split: if_split_asm, frule ko_at_state_refs_ofD, simp)+ apply (frule_tac P="(=) (Some a)" in ntfn_bound_tcb_at, simp_all add: obj_at_def)[1] apply (fastforce simp: ntfn_q_refs_no_NTFNBound obj_at_def is_tcb state_refs_of_def tcb_st_refs_of_no_NTFNBound tcb_bound_refs_def2 tcb_ntfn_is_bound_def tcb_st_refs_no_TCBBound dest!: pred_tcb_at_tcb_at bound_tcb_at_state_refs_ofD) apply (frule ko_at_state_refs_ofD, simp) done crunch not_queued[wp]: unbind_maybe_notification, unbind_notification "not_queued t" context DetSchedSchedule_AI begin lemma fast_finalise_not_queued: "\not_queued t and (st_tcb_at active t and valid_objs and scheduler_act_not t and sym_refs \ state_refs_of)\ fast_finalise cap final \\_. not_queued t\" apply (cases cap, simp_all) apply (wp cancel_all_ipc_not_queued cancel_all_signals_not_queued get_simple_ko_wp unbind_maybe_notification_valid_objs | simp)+ done crunch not_queued: delete_caller_cap "not_queued t" (wp: fast_finalise_not_queued hoare_drop_imps simp: if_fun_split unless_def) end lemma set_simple_ko_ct_active: "\ct_active\ set_simple_ko f ptr ep \\rv. ct_active\" apply (simp add: set_simple_ko_def set_object_def | wp get_object_wp)+ apply (clarsimp simp: ct_in_state_def pred_tcb_at_def obj_at_def split: kernel_object.splits) done context DetSchedSchedule_AI begin crunch is_etcb_at[wp]: setup_reply_master, cancel_ipc "is_etcb_at t" (wp: hoare_drop_imps crunch_wps select_inv simp: crunch_simps unless_def) end crunch weak_valid_sched_action[wp]: setup_reply_master "weak_valid_sched_action" crunch is_etcb_at_ext[wp]: set_thread_state_ext, tcb_sched_action, reschedule_required, empty_slot_ext "is_etcb_at t" crunch is_etcb_at[wp]: set_thread_state "is_etcb_at t" (wp: hoare_drop_imps crunch_wps select_inv simp: crunch_simps unless_def) lemma set_eobject_is_etcb_at_ext[wp]: "\is_etcb_at t\ set_eobject ptr etcb \\_. is_etcb_at t\" apply (simp add: set_eobject_def | wp)+ apply (simp add: is_etcb_at_def split: if_split_asm) done crunch is_etcb_at_ext[wp]: ethread_set "is_etcb_at t" crunch valid_etcbs[wp]: set_mrs valid_etcbs (wp: valid_etcbs_lift) lemma cap_insert_check_cap_ext_valid[wp]:" \valid_list\ check_cap_at new_cap src_slot (check_cap_at t slot (cap_insert new_cap src_slot x)) \\rv. valid_list\" apply (simp add: check_cap_at_def) apply (wp get_cap_wp | simp)+ done lemma opt_update_thread_valid_sched[wp]: "(\x a. tcb_state (fn a x) = tcb_state x) \ \valid_sched\ option_update_thread t fn v \\_. valid_sched\" apply (rule hoare_pre) apply (simp add: option_update_thread_def) apply (wp thread_set_not_state_valid_sched | wpc | simp)+ done lemma opt_update_thread_simple_sched_action[wp]: "\simple_sched_action\ option_update_thread t fn v \\_. simple_sched_action\" apply (rule hoare_pre) apply (simp add: option_update_thread_def) apply (wp | wpc | simp)+ done crunch ct_active[wp]: delete_caller_cap ct_active (wp: ct_in_state_thread_state_lift) lemma test: "invs s \ (\y. get_tcb thread s = Some y) \ s \ tcb_ctable (the (get_tcb thread s))" apply (simp add: invs_valid_tcb_ctable_strengthen) done context DetSchedSchedule_AI begin lemma handle_recv_valid_sched: "\valid_sched and invs and ct_active and ct_not_queued and scheduler_act_sane\ handle_recv is_blocking \\rv. valid_sched\" apply (simp add: handle_recv_def Let_def ep_ntfn_cap_case_helper cong: if_cong) apply (wp get_simple_ko_wp handle_fault_valid_sched delete_caller_cap_not_queued receive_ipc_valid_sched receive_signal_valid_sched | simp)+ apply (rule hoare_vcg_E_elim) apply (wpsimp simp: lookup_cap_def lookup_slot_for_thread_def) apply (wp resolve_address_bits_valid_fault2)+ apply (simp add: valid_fault_def) apply (wp hoare_drop_imps hoare_vcg_all_lift_R) apply (wpsimp wp: delete_caller_cap_not_queued | strengthen invs_valid_tcb_ctable_strengthen)+ apply (auto simp: ct_in_state_def tcb_at_invs objs_valid_tcb_ctable invs_valid_objs) done lemma handle_recv_valid_sched': "\invs and valid_sched and ct_active and ct_not_queued and scheduler_act_sane\ handle_recv is_blocking \\_. valid_sched\" apply (wp handle_recv_valid_sched) apply (simp add: invs_def valid_state_def valid_pspace_def) done end crunch valid_sched[wp]: reply_from_kernel valid_sched lemma set_thread_state_active_valid_sched: "active st \ \valid_sched and ct_active and (\s. cur_thread s = ct)\ set_thread_state ct st \\_. valid_sched\" apply (simp add: valid_sched_def valid_queues_def) apply (wp hoare_vcg_all_lift) apply (rule hoare_lift_Pf [where f="\s. ready_queues s", OF _ set_thread_state_ready_queues]) apply (wpsimp wp: hoare_vcg_ball_lift sts_st_tcb_at_cases simp: runnable_eq_active) apply (wp set_thread_state_runnable_valid_sched_action set_thread_state_runnable_valid_blocked) apply (clarsimp simp: ct_in_state_def st_tcb_at_def obj_at_def runnable_eq_active) done lemma set_thread_state_Running_valid_sched: "\valid_sched and ct_active and (\s. cur_thread s = ct)\ set_thread_state ct Running \\_. valid_sched\" by (rule set_thread_state_active_valid_sched) simp lemma set_thread_state_Restart_valid_sched: "\valid_sched and ct_active and (\s. cur_thread s = ct)\ set_thread_state ct Restart \\_. valid_sched\" by (rule set_thread_state_active_valid_sched) simp context DetSchedSchedule_AI begin crunches invoke_irq_control, invoke_irq_handler for valid_sched[wp]: "valid_sched" end lemma simple_sched_action_sched_act_not[simp]: "simple_sched_action s \ scheduler_act_not t s" by (fastforce simp: simple_sched_action_def scheduler_act_not_def) declare valid_idle_etcb_lift[wp del] lemma ethread_set_not_activatable_valid_idle_etcb: "\valid_idle_etcb and valid_idle and st_tcb_at (\ts. \ activatable ts) tptr\ ethread_set f tptr \\_. valid_idle_etcb\" apply(simp add: ethread_set_def set_eobject_def) apply wp apply(clarsimp simp: get_etcb_def valid_idle_etcb_def etcb_at'_def valid_idle_def pred_tcb_at_def obj_at_def) done lemma ethread_set_not_activatable_valid_sched: "\valid_sched and valid_idle and st_tcb_at (\ts. \ activatable ts) tptr\ ethread_set f tptr \\_. valid_sched\" apply (simp add: valid_sched_def valid_sched_action_def | wp ethread_set_not_queued_valid_queues ethread_set_not_switch_switch_in_cur_domain ethread_set_not_cur_ct_in_cur_domain ethread_set_valid_blocked ethread_set_not_activatable_valid_idle_etcb)+ apply (force simp: valid_idle_def st_tcb_at_def obj_at_def not_cur_thread_def is_activatable_def weak_valid_sched_action_def valid_queues_def not_queued_def split: thread_state.splits)+ done lemma ethread_set_not_idle_valid_idle_etcb: "\valid_idle_etcb and valid_idle and (\s. tptr \ idle_thread s)\ ethread_set f tptr \\_. valid_idle_etcb\" apply(simp add: ethread_set_def set_eobject_def) apply wp apply(clarsimp simp: get_etcb_def valid_idle_etcb_def etcb_at'_def valid_idle_def st_tcb_at_def obj_at_def) done lemma ethread_set_not_idle_valid_sched: "\valid_sched and simple_sched_action and not_queued tptr and (\s. tptr \ cur_thread s) and (\s. tptr \ idle_thread s) and valid_idle\ ethread_set f tptr \\_. valid_sched\" apply (simp add: valid_sched_def valid_sched_action_def | wp ethread_set_not_queued_valid_queues ethread_set_not_switch_switch_in_cur_domain ethread_set_not_cur_ct_in_cur_domain ethread_set_valid_blocked ethread_set_not_idle_valid_idle_etcb)+ apply (force simp: simple_sched_action_def st_tcb_at_def obj_at_def not_cur_thread_def is_activatable_def weak_valid_sched_action_def valid_queues_def not_queued_def split: thread_state.splits)+ done lemma ethread_set_ssa_valid_sched_action: "\valid_sched_action and simple_sched_action\ ethread_set f tptr \\_. valid_sched_action\" apply (simp add: valid_sched_action_def | wp ethread_set_not_switch_switch_in_cur_domain)+ apply (force simp: simple_sched_action_def)+ done lemma ethread_set_valid_blocked_except: "\valid_blocked_except t\ ethread_set f tptr \\_. valid_blocked_except t\" by (wp valid_blocked_except_lift | simp add: ethread_set_def set_eobject_def)+ declare tcb_sched_action_valid_idle_etcb[wp] lemma invoke_domain_valid_sched[wp]: "\valid_sched and tcb_at t and (\s. t \ idle_thread s) and simple_sched_action and valid_idle\ invoke_domain t d \\_. valid_sched\" apply (simp add: invoke_domain_def) apply wp apply (simp add: set_domain_def thread_set_domain_def) apply (wp gts_st_tcb_at hoare_vcg_if_lift hoare_vcg_if_lift2 hoare_vcg_imp_lift hoare_vcg_disj_lift ethread_set_not_queued_valid_queues reschedule_required_valid_sched tcb_sched_action_enqueue_valid_blocked ethread_set_valid_blocked_except ethread_set_valid_blocked ethread_set_ssa_valid_sched_action ethread_set_not_cur_ct_in_cur_domain ethread_set_not_idle_valid_sched ethread_set_not_idle_valid_idle_etcb) apply (wp hoare_weak_lift_imp hoare_weak_lift_imp_conj tcb_dequeue_not_queued tcb_sched_action_dequeue_valid_blocked_except) apply simp apply (wp hoare_vcg_disj_lift) apply (rule_tac Q="\_. valid_sched and not_queued t and valid_idle and (\s. t \ idle_thread s)" in hoare_strengthen_post) apply (wp tcb_sched_action_dequeue_valid_sched_not_runnable tcb_dequeue_not_queued) apply (simp add: valid_sched_def valid_sched_action_def) apply simp apply (wp hoare_vcg_disj_lift tcb_dequeue_not_queued tcb_sched_action_dequeue_valid_blocked_except tcb_sched_action_dequeue_valid_sched_not_runnable)+ apply (clarsimp simp: valid_sched_def not_cur_thread_def valid_sched_action_def not_pred_tcb) apply (auto simp: st_tcb_def2 tcb_at_def) done lemma idle_not_reply_cap: "\ valid_idle s; valid_reply_caps s; cte_wp_at (is_reply_cap_to r) p s \ \ r \ idle_thread s" apply (drule_tac p=p in valid_reply_caps_of_stateD',assumption) apply (clarsimp simp: valid_idle_def pred_tcb_at_def obj_at_def) done lemma idle_not_reply_cap': "\ valid_idle s; valid_reply_caps s; cte_wp_at ((=) (ReplyCap r False R)) p s \ \ r \ idle_thread s" apply (simp add: cte_wp_at_caps_of_state) apply (drule_tac p=p in valid_reply_caps_of_stateD,assumption) apply (clarsimp simp: valid_idle_def pred_tcb_at_def obj_at_def) done context DetSchedSchedule_AI begin lemma perform_invocation_valid_sched[wp]: "\invs and valid_invocation i and ct_active and simple_sched_action and valid_sched and (\s. not_queued (cur_thread s) s)\ perform_invocation calling blocking i \\rv. valid_sched\" apply (cases i, simp_all) apply (wp invoke_untyped_valid_sched send_ipc_valid_sched | clarsimp)+ apply (clarsimp simp: ct_in_state_def) apply (wp invoke_cnode_valid_sched send_ipc_valid_sched invoke_domain_valid_sched | simp add: invs_valid_objs idle_not_reply_cap invs_valid_idle invs_valid_reply_caps is_reply_cap_to_def | clarsimp simp: ct_in_state_def)+ done end crunch not_queued[wp]: set_thread_state_ext "not_queued t" crunch ct_not_queued[wp]: set_thread_state ct_not_queued (wp: ct_not_queued_lift) context DetSchedSchedule_AI begin lemma handle_invocation_valid_sched: "\invs and valid_sched and ct_active and (\s. scheduler_action s = resume_cur_thread)\ handle_invocation a b \\rv. valid_sched\" apply (simp add: handle_invocation_def) apply (wp syscall_valid handle_fault_valid_sched | wpc)+ apply (wp set_thread_state_runnable_valid_sched)[1] apply wp+ apply (wp gts_wp hoare_vcg_all_lift) apply (rule_tac Q="\_. valid_sched" and E="\_. valid_sched" in hoare_post_impErr) apply wp apply ((clarsimp simp: st_tcb_at_def obj_at_def)+)[2] apply (wp ct_in_state_set set_thread_state_runnable_valid_sched | simp add: split_def if_apply_def2 split del: if_split)+ apply (simp add: validE_E_def) apply (rule hoare_post_impErr) apply (rule lookup_cap_and_slot_valid_fault) apply (wp | simp)+ apply (auto simp: ct_in_state_def valid_sched_def ct_not_in_q_def valid_queues_def not_queued_def runnable_eq_active elim: st_tcb_ex_cap) done end lemma valid_sched_ct_not_queued: "\valid_sched s; scheduler_action s = resume_cur_thread\ \ not_queued (cur_thread s) s" by (fastforce simp: valid_sched_def ct_not_in_q_def) context DetSchedSchedule_AI begin lemma handle_reply_valid_sched: "\valid_sched and ct_active and invs\ handle_reply \\rv. valid_sched\" apply (simp add: handle_reply_def) apply (wp get_cap_wp | wpc | clarsimp)+ apply (auto simp: invs_valid_objs idle_not_reply_cap' invs_valid_idle invs_valid_reply_caps is_reply_cap_to_def | erule cte_wp_at_lift)+ done end crunch ct_not_queued[wp]: do_machine_op, cap_insert, set_extra_badge "\s. not_queued (cur_thread s) s" (wp: hoare_drop_imps) lemma transfer_caps_ct_not_queued[wp]: "\\s. not_queued (cur_thread s) s\ transfer_caps info caps ep recv recv_buf \\rv s. not_queued (cur_thread s) s\" by (simp add: transfer_caps_def | wp transfer_caps_loop_pres | wpc)+ crunch ct_sched_act_not[wp]: set_thread_state "\s. scheduler_act_not (cur_thread s) s" (wp: set_scheduler_action_wp gts_wp crunch_wps simp: crunch_simps ignore: set_scheduler_action) context DetSchedSchedule_AI begin crunch not_queued[wp]: handle_fault_reply "not_queued t" crunch sched_act_not[wp]: handle_fault_reply "scheduler_act_not t" crunch valid_etcbs[wp]: set_extra_badge, do_ipc_transfer valid_etcbs (wp: crunch_wps const_on_failure_wp simp: crunch_simps rule: transfer_caps_loop_pres) crunch cur[wp]: handle_fault_reply "cur_tcb :: det_ext state \ bool" (wp: crunch_wps simp: crunch_simps) crunch weak_valid_sched_action[wp]: empty_slot_ext, cap_delete_one weak_valid_sched_action (wp: crunch_wps set_thread_state_runnable_weak_valid_sched_action set_bound_notification_weak_valid_sched_action simp: cur_tcb_def unless_def) lemma do_reply_transfer_not_queued[wp]: "\not_queued t and invs and st_tcb_at active t and scheduler_act_not t and K(receiver \ t)\ do_reply_transfer sender receiver slot grant \\_. not_queued t\" apply (simp add: do_reply_transfer_def) apply (wp cap_delete_one_not_queued hoare_vcg_if_lift | wpc | clarsimp split del: if_split | wp (once) hoare_drop_imps)+ apply (simp add: invs_def valid_state_def valid_pspace_def)+ done lemma do_reply_transfer_schedact_not[wp]: "\scheduler_act_not t and K(receiver \ t)\ do_reply_transfer sender receiver slot grant \\_. scheduler_act_not t\" apply (simp add: do_reply_transfer_def) apply (wp hoare_vcg_if_lift | wpc | clarsimp split del: if_split | wp (once) hoare_drop_imps)+ done end lemma assert_false:"\\s. \ P\ assert P \\_ _. False\" apply wp apply simp done lemma do_reply_transfer_add_assert: assumes a: "\st_tcb_at awaiting_reply receiver and P\ do_reply_transfer sender receiver slot grant \\_. Q\" shows "\P\ do_reply_transfer sender receiver slot grant \\_. Q\" apply (rule hoare_name_pre_state) apply (case_tac "st_tcb_at awaiting_reply receiver s") apply (rule hoare_pre) apply (wp a) apply simp apply (simp add: do_reply_transfer_def) apply (rule hoare_seq_ext) apply (rule hoare_seq_ext) prefer 2 apply (rule assert_false) apply simp apply (simp add: get_thread_state_def thread_get_def) apply wp apply (clarsimp simp: get_tcb_def pred_tcb_at_def obj_at_def split: option.splits kernel_object.splits) done context DetSchedSchedule_AI begin lemma do_reply_transfer_ct_not_queued[wp]: "\ct_not_queued and invs and ct_active and scheduler_act_sane\ do_reply_transfer sender receiver slot grant \\_. ct_not_queued\" apply (rule do_reply_transfer_add_assert) apply (rule hoare_pre) apply (wp ct_not_queued_lift) apply (clarsimp simp add: ct_in_state_def pred_tcb_at_def obj_at_def) done lemma do_reply_transfer_scheduler_act_sane[wp]: "\scheduler_act_sane and ct_active\ do_reply_transfer sender receiver slot grant \\_. scheduler_act_sane\" apply (rule do_reply_transfer_add_assert) apply (rule hoare_pre) apply (wp sch_act_sane_lift) apply (clarsimp simp add: ct_in_state_def pred_tcb_at_def obj_at_def) done crunch ct_not_queued[wp]: handle_reply "ct_not_queued" crunch scheduler_act_sane[wp]: handle_reply "scheduler_act_sane" end locale DetSchedSchedule_AI_handle_hypervisor_fault = DetSchedSchedule_AI + assumes handle_hyp_fault_valid_sched[wp]: "\t fault. \valid_sched and invs and st_tcb_at active t and not_queued t and scheduler_act_not t and (ct_active or ct_idle)\ handle_hypervisor_fault t fault \\_. valid_sched\" assumes handle_reserved_irq_valid_sched' [wp]: "\irq. \valid_sched and invs and (\s. irq \ non_kernel_IRQs \ scheduler_act_sane s \ ct_not_queued s)\ handle_reserved_irq irq \\rv. valid_sched\" context DetSchedSchedule_AI_handle_hypervisor_fault begin lemma handle_interrupt_valid_sched[wp]: "\valid_sched and invs and (\s. irq \ non_kernel_IRQs \ scheduler_act_sane s \ ct_not_queued s)\ handle_interrupt irq \\rv. valid_sched\" unfolding handle_interrupt_def by (wpsimp wp: get_cap_wp hoare_drop_imps hoare_vcg_all_lift|rule conjI)+ lemma set_scheduler_action_switch_not_cur_thread [wp]: "\\s. True\ set_scheduler_action (switch_thread target) \\rv. not_cur_thread t\" unfolding set_scheduler_action_def by wp (simp add: not_cur_thread_def) lemma possible_switch_to_not_cur_thread [wp]: "\not_cur_thread t\ possible_switch_to target \\_. not_cur_thread t\" unfolding possible_switch_to_def apply (rule hoare_pre) apply (wp hoare_vcg_imp_lift|wpc)+ apply clarsimp done crunch not_cur_thread[wp]: handle_recv "not_cur_thread target" (wp: crunch_wps simp: crunch_simps) crunch it[wp]: handle_recv "\s::det_ext state. P (idle_thread s)" (wp: crunch_wps simp: crunch_simps) lemma handle_event_valid_sched: "\invs and valid_sched and (\s. e \ Interrupt \ ct_active s) and (\s. scheduler_action s = resume_cur_thread)\ handle_event e \\rv. valid_sched\" apply (cases e, simp_all) apply (rename_tac syscall) apply (case_tac syscall, simp_all add: handle_send_def handle_call_def) apply ((rule hoare_pre, wp handle_invocation_valid_sched handle_recv_valid_sched' handle_reply_valid_sched | fastforce simp: invs_valid_objs invs_sym_refs valid_sched_ct_not_queued)+)[5] apply (wp handle_fault_valid_sched hvmf_active hoare_drop_imps hoare_vcg_disj_lift handle_recv_valid_sched' handle_reply_valid_sched hoare_vcg_all_lift| wpc | clarsimp simp: ct_in_state_def valid_sched_ct_not_queued | fastforce simp: valid_fault_def)+ done crunch valid_list[wp]: activate_thread valid_list crunch valid_list[wp]: guarded_switch_to, switch_to_idle_thread, choose_thread valid_list (wp: crunch_wps) end crunch valid_list[wp]: next_domain valid_list (simp: Let_def) context DetSchedSchedule_AI_handle_hypervisor_fault begin crunch valid_list[wp]: schedule_choose_new_thread valid_list lemma schedule_valid_list[wp]: "\valid_list\ Schedule_A.schedule \\_. valid_list\" apply (simp add: Schedule_A.schedule_def) apply (wp add: tcb_sched_action_valid_list gts_wp hoare_drop_imps del: ethread_get_wp | wpc | simp)+ done lemma call_kernel_valid_list[wp]: "\valid_list\ call_kernel e \\_. valid_list\" apply (simp add: call_kernel_def) by (wp | simp)+ lemma call_kernel_valid_sched: "\invs and valid_sched and (\s. e \ Interrupt \ ct_running s) and (ct_active or ct_idle) and (\s. scheduler_action s = resume_cur_thread)\ call_kernel e \\_. valid_sched\" apply (simp add: call_kernel_def) apply (wp schedule_valid_sched activate_thread_valid_sched | simp)+ apply (rule_tac Q="\rv. invs" in hoare_strengthen_post) apply wp apply (erule invs_valid_idle) apply (rule hoare_strengthen_post [where Q="\irq s. irq \ Some ` non_kernel_IRQs \ valid_sched s \ invs s"]) apply (wpsimp wp: getActiveIRQ_neq_non_kernel) apply auto[1] apply (rule_tac Q="\rv. valid_sched and invs" and E="\rv. valid_sched and invs" in hoare_post_impErr) apply (rule valid_validE) apply (wp handle_event_valid_sched) apply (force intro: active_from_running)+ done end end