(* * 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) *) (* The main theorem *) theory AInvs imports "./$L4V_ARCH/ArchAInvsPre" begin lemma st_tcb_at_nostate_upd: "\ get_tcb t s = Some y; tcb_state y = tcb_state y' \ \ st_tcb_at P t' (s \kheap := kheap s(t \ TCB y')\) = st_tcb_at P t' s" by (clarsimp simp add: pred_tcb_at_def obj_at_def dest!: get_tcb_SomeD) lemma pred_tcb_at_upd_apply: "pred_tcb_at proj P t (s\kheap := p'\) = pred_tcb_at proj P t (s\kheap := (kheap s)(t := p' t)\)" by (simp add: pred_tcb_at_def obj_at_def) text {* The top-level invariance *} lemma akernel_invs: "\invs and (\s. e \ Interrupt \ ct_running s)\ (call_kernel e) :: (unit,unit) s_monad \\rv. invs and (\s. ct_running s \ ct_idle s)\" apply wp apply (simp add: call_kernel_def) apply (wp activate_invs | simp)+ apply (auto simp: active_from_running) done lemma akernel_invs_det_ext: "\invs and (\s. e \ Interrupt \ ct_running s)\ (call_kernel e) :: (unit,det_ext) s_monad \\rv. invs and (\s. ct_running s \ ct_idle s)\" apply wp apply (simp add: call_kernel_def) apply (wp activate_invs | simp)+ apply (auto simp: active_from_running) done (* FIXME: move *) lemma ct_running_machine_op: "\ct_running\ do_machine_op f \\_. ct_running\" apply (simp add: ct_in_state_def pred_tcb_at_def obj_at_def) apply (rule hoare_lift_Pf [where f=cur_thread]) by wp lemma kernel_entry_invs: "\invs and (\s. e \ Interrupt \ ct_running s)\ (kernel_entry e us) :: (register \ 32 word,unit) s_monad \\rv. invs and (\s. ct_running s \ ct_idle s)\" apply (simp add: kernel_entry_def) by (wp akernel_invs thread_set_invs_trivial thread_set_ct_running select_wp ct_running_machine_op static_imp_wp | clarsimp simp add: tcb_cap_cases_def)+ (* FIXME: move to Lib.thy *) lemma Collect_subseteq: "{x. P x} <= {x. Q x} \ (\x. P x \ Q x)" by auto lemma device_update_invs: "\invs and (\s. (dom ds) \ (device_region s))\ do_machine_op (device_memory_update ds) \\_. invs\" apply (simp add: do_machine_op_def device_memory_update_def simpler_modify_def select_f_def gets_def get_def bind_def valid_def return_def) apply (clarsimp simp: invs_def valid_state_def valid_irq_states_def valid_machine_state_def cur_tcb_def pspace_respects_device_region_def cap_refs_respects_device_region_def cong: user_mem_dom_cong simp del: split_paired_All) apply (clarsimp cong: device_mem_dom_cong simp:cap_range_respects_device_region_def simp del: split_paired_All split_paired_Ex) apply (intro conjI) apply fastforce apply fastforce apply (clarsimp simp del: split_paired_All split_paired_Ex) apply (drule_tac x = "(a,b)" in spec) apply (erule notE) apply (erule cte_wp_at_weakenE) apply clarsimp apply (fastforce split: if_splits) (* takes 20 secs *) done crunch device_state_inv[wp]: user_memory_update "\ms. P (device_state ms)" lemma dom_restrict_plus_eq: "a \ b \ b = b" by auto lemma user_memory_update[wp]: "\\s. P (device_region s) \ do_machine_op (user_memory_update a) \\rv s. P (device_region s)\" by (simp add: do_machine_op_def user_memory_update_def simpler_modify_def valid_def bind_def gets_def return_def get_def select_f_def) lemma do_user_op_invs: "\invs and ct_running\ do_user_op f tc \\_. invs and ct_running\" apply (simp add: do_user_op_def split_def) apply (wp device_update_invs) apply (wp ct_running_machine_op select_wp dmo_invs | simp add:dom_restrict_plus_eq)+ apply (clarsimp simp: user_memory_update_def simpler_modify_def restrict_map_def invs_def cur_tcb_def split: option.splits split_if_asm) apply (frule ptable_rights_imp_frame) apply fastforce apply simp apply (clarsimp simp: valid_state_def device_frame_in_device_region) done end