(* * Copyright 2014, General Dynamics C4 Systems * * This software may be distributed and modified according to the terms of * the GNU General Public License version 2. Note that NO WARRANTY is provided. * See "LICENSE_GPLv2.txt" for details. * * @TAG(GD_GPL) *) theory Schedule_AI imports VSpace_AI begin abbreviation "activatable \ \st. runnable st \ idle st" locale Schedule_AI = fixes state_ext :: "('a::state_ext) itself" assumes dmo_mapM_storeWord_0_invs[wp]: "\S. valid invs (do_machine_op (mapM (\p. storeWord p 0) S)) (\_. (invs :: 'a state \ bool))" assumes arch_stt_invs [wp]: "\t'. \invs\ arch_switch_to_thread t' \\_. (invs :: 'a state \ bool)\" assumes arch_stt_tcb [wp]: "\t'. \tcb_at t'\ arch_switch_to_thread t' \\_. (tcb_at t' :: 'a state \ bool)\" assumes arch_stt_runnable: "\t. \st_tcb_at runnable t\ arch_switch_to_thread t \\r . (st_tcb_at runnable t :: 'a state \ bool)\" assumes stit_invs [wp]: "\invs\ switch_to_idle_thread \\rv. (invs :: 'a state \ bool)\" assumes stit_activatable: "\invs\ switch_to_idle_thread \\rv . (ct_in_state activatable :: 'a state \ bool)\" context begin interpretation Arch . (* FIXME arch_split: some of these could be moved to generic theories so they don't need to be unqualified. *) requalify_facts no_irq no_irq_storeWord end crunch inv[wp]: schedule_switch_thread_fastfail P lemma findM_inv'': assumes p: "suffix xs xs'" assumes x: "\x xs. suffix (x # xs) xs' \ \P (x # xs)\ m x \\rv s. (rv \ Q s) \ (\ rv \ P xs s)\" assumes y: "\s. P [] s \ Q s" shows "\P xs\ findM m xs \\rv. Q\" using p apply (induct xs) apply simp apply wp apply (erule y) apply (frule suffix_ConsD) apply simp apply wp apply (erule x) apply simp done lemmas findM_inv' = findM_inv''[OF suffix_order.order.refl] lemma findM_inv: assumes x: "\x xs. \P\ m x \\rv. P\" shows "\P\ findM m xs \\rv. P\" by (rule findM_inv', simp_all add: x) lemma allActiveTCBs_gets: "allActiveTCBs = gets (\state. {x. getActiveTCB x state \ None})" by (simp add: allActiveTCBs_def gets_def) lemma postfix_tails: "\ suffix (xs # ys) (tails zs) \ \ suffix xs zs \ (xs # ys) = tails xs" apply (induct zs arbitrary: xs ys) apply (clarsimp elim!: suffixE) apply (case_tac zs, simp_all)[1] apply (clarsimp elim!: suffixE) apply (case_tac zsa, simp_all) apply clarsimp apply clarsimp apply (erule meta_allE, erule meta_allE, drule meta_mp, rule suffix_appendI[OF suffix_order.order.refl]) apply clarsimp apply (erule suffix_ConsI) done lemma valid_irq_states_cur_thread_update[simp]: "valid_irq_states (cur_thread_update f s) = valid_irq_states s" by(simp add: valid_irq_states_def) lemma sct_invs: "\invs and tcb_at t\ modify (cur_thread_update (%_. t)) \\rv. invs\" by wp (clarsimp simp add: invs_def cur_tcb_def valid_state_def valid_idle_def valid_irq_node_def valid_machine_state_def) lemma storeWord_valid_irq_states: "\\m. valid_irq_states (s\machine_state := m\)\ storeWord x y \\a b. valid_irq_states (s\machine_state := b\)\" apply (simp add: valid_irq_states_def | wp no_irq | simp add: no_irq_storeWord)+ done lemma dmo_storeWord_valid_irq_states[wp]: "\valid_irq_states\ do_machine_op (storeWord x y) \\_. valid_irq_states\" apply (simp add: do_machine_op_def | wp | wpc)+ apply clarsimp apply(erule use_valid[OF _ storeWord_valid_irq_states]) by simp lemma dmo_kheap_arch_state[wp]: "\\s. P (kheap s) (arch_state s)\ do_machine_op a \\_ s. P (kheap s) (arch_state s)\" by (clarsimp simp: do_machine_op_def simpler_gets_def select_f_def simpler_modify_def return_def bind_def valid_def) lemmas do_machine_op_tcb[wp] = do_machine_op_obj_at[where P=id and Q=is_tcb, simplified] lemma (in Schedule_AI) stt_tcb [wp]: "\tcb_at t\ switch_to_thread t \\_. (tcb_at t :: 'a state \ bool)\" apply (simp add: switch_to_thread_def) apply (wp | simp)+ done (* FIXME - Move Invariants_AI *) lemma invs_exst [iff]: "invs (trans_state f s) = invs s" by (simp add: invs_def valid_state_def) lemma (in Schedule_AI) stt_invs [wp]: "\invs :: 'a state \ bool\ switch_to_thread t' \\_. invs\" apply (simp add: switch_to_thread_def) apply wp apply (simp add: trans_state_update[symmetric] del: trans_state_update) apply (rule_tac Q="\_. invs and tcb_at t'" in hoare_strengthen_post, wp) apply (clarsimp simp: invs_def valid_state_def valid_idle_def valid_irq_node_def valid_machine_state_def) apply (fastforce simp: cur_tcb_def obj_at_def elim: valid_pspace_eqI ifunsafe_pspaceI) apply wp+ apply clarsimp apply (simp add: is_tcb_def) done lemma (in Schedule_AI) stt_activatable: "\st_tcb_at runnable t\ switch_to_thread t \\rv . (ct_in_state activatable :: 'a state \ bool) \" apply (simp add: switch_to_thread_def) apply (wp | simp add: ct_in_state_def)+ apply (rule hoare_post_imp [OF _ arch_stt_runnable]) apply (clarsimp elim!: pred_tcb_weakenE) apply (rule assert_inv) apply wp apply assumption done lemma invs_upd_cur_valid: "\\P\ f \\rv. invs\; \Q\ f \\rv. tcb_at thread\\ \ \P and Q\ f \\rv s. invs (s\cur_thread := thread\)\" by (fastforce simp: valid_def invs_def valid_state_def cur_tcb_def valid_machine_state_def) (* FIXME move *) lemma pred_tcb_weaken_strongerE: "\ pred_tcb_at proj P t s; \tcb . P (proj tcb) \ P' (proj' tcb) \ \ pred_tcb_at proj' P' t s" by (auto simp: pred_tcb_at_def elim: obj_at_weakenE) lemma OR_choice_weak_wp: "\P\ f \ g \Q\ \ \P\ OR_choice b f g \Q\" apply (fastforce simp add: OR_choice_def alternative_def valid_def bind_def select_f_def gets_def return_def get_def split: option.splits if_split_asm) done locale Schedule_AI_U = Schedule_AI "TYPE(unit)" lemma (in Schedule_AI_U) schedule_invs[wp]: "\invs\ (Schedule_A.schedule :: (unit,unit) s_monad) \\rv. invs\" apply (simp add: Schedule_A.schedule_def allActiveTCBs_def) apply (wp OR_choice_weak_wp alternative_wp dmo_invs thread_get_inv do_machine_op_tcb select_ext_weak_wp select_wp when_def | clarsimp simp: getActiveTCB_def get_tcb_def)+ done (* FIXME - move *) lemma get_tcb_exst_update: "get_tcb p (trans_state f s) = get_tcb p s" by (simp add: get_tcb_def) lemma ct_in_state_trans_update[simp]: "ct_in_state st (trans_state f s) = ct_in_state st s" apply (simp add: ct_in_state_def) done lemma (in Schedule_AI_U) schedule_ct_activateable[wp]: "\invs\ (Schedule_A.schedule :: (unit,unit) s_monad) \\rv. ct_in_state activatable \" proof - have P: "\t s. ct_in_state activatable (cur_thread_update (\_. t) s) = st_tcb_at activatable t s" by (fastforce simp: ct_in_state_def pred_tcb_at_def intro: obj_at_pspaceI) have Q: "\s. invs s \ idle_thread s = cur_thread s \ ct_in_state activatable s" apply (clarsimp simp: ct_in_state_def dest!: invs_valid_idle) apply (clarsimp simp: valid_idle_def pred_tcb_def2) done show ?thesis apply (simp add: Schedule_A.schedule_def allActiveTCBs_def) apply (wp alternative_wp select_ext_weak_wp select_wp stt_activatable stit_activatable | simp add: P Q)+ apply (clarsimp simp: getActiveTCB_def ct_in_state_def) apply (rule conjI) apply clarsimp apply (case_tac "get_tcb (cur_thread s) s", simp_all add: ct_in_state_def) apply (drule get_tcb_SomeD) apply (clarsimp simp: pred_tcb_at_def obj_at_def split: if_split_asm) apply (case_tac "get_tcb x s", simp_all) apply (drule get_tcb_SomeD) apply (clarsimp simp: pred_tcb_at_def obj_at_def split: if_split_asm) done qed end