diff --git a/proof/crefine/ADT_C.thy b/proof/crefine/ADT_C.thy index ac10b4df6..395b7feef 100644 --- a/proof/crefine/ADT_C.thy +++ b/proof/crefine/ADT_C.thy @@ -75,7 +75,7 @@ lemma Basic_sem_eq: lemma setArchTCB_C_corres: "\ ccontext_relation tc (tcbContext_C tc'); t' = tcb_ptr_to_ctcb_ptr t \ \ corres_underlying rf_sr nf nf' dc (tcb_at' t) \ - (threadSet (\tcb. tcb \ tcbContext := tc \) t) (setArchTCB_C tc' t')" + (threadSet (\tcb. tcb \ tcbArch := atcbContextSet tc (tcbArch tcb)\) t) (setArchTCB_C tc' t')" apply (simp add: setArchTCB_C_def exec_C_def Basic_sem_eq corres_underlying_def) apply clarsimp apply (simp add: threadSet_def bind_assoc split_def exec_gets) @@ -111,7 +111,7 @@ lemma setArchTCB_C_corres: apply (rule ext, simp split: split_if) apply (drule ko_at_projectKO_opt) apply (erule (2) cmap_relation_upd_relI) - apply (simp add: ctcb_relation_def) + apply (simp add: ctcb_relation_def carch_tcb_relation_def) apply assumption apply simp done @@ -294,14 +294,16 @@ where "cirqstate_to_H w \ if w = scast Kernel_C.IRQSignal then irqstate.IRQSignal else if w = scast Kernel_C.IRQTimer then irqstate.IRQTimer - else irqstate.IRQInactive" + else if w = scast Kernel_C.IRQInactive then irqstate.IRQInactive + else irqstate.IRQReserved" lemma cirqstate_cancel: "cirqstate_to_H \ irqstate_to_C = id" apply (rule ext) apply (case_tac x) apply (auto simp: cirqstate_to_H_def Kernel_C.IRQInactive_def - Kernel_C.IRQTimer_def Kernel_C.IRQSignal_def) + Kernel_C.IRQTimer_def Kernel_C.IRQSignal_def + Kernel_C.IRQReserved_def) done definition @@ -812,6 +814,13 @@ lemma ccontext_relation_imp_eq: "ccontext_relation f x \ ccontext_relation g x \ f=g" by (rule ext) (simp add: ccontext_relation_def) +lemma carch_tcb_relation_imp_eq: + "carch_tcb_relation f x \ carch_tcb_relation g x \ f = g" + apply (cases f) + apply (cases g) + apply (simp add: carch_tcb_relation_def ccontext_relation_imp_eq atcbContextGet_def) + done + (* FIXME: move *) lemma ran_tcb_cte_cases: "ran tcb_cte_cases = @@ -915,7 +924,7 @@ lemma map_to_ctes_tcb_ctes: lemma cfault_rel_imp_eq: "cfault_rel x a b \ cfault_rel y a b \ x=y" by (clarsimp simp: cfault_rel_def is_cap_fault_def - split: split_if_asm fault_CL.splits) + split: split_if_asm seL4_Fault_CL.splits) lemma cthread_state_rel_imp_eq: "cthread_state_relation x z \ cthread_state_relation y z \ x=y" @@ -962,7 +971,7 @@ lemma cpspace_tcb_relation_unique: apply (case_tac x, case_tac y, case_tac "the (clift ch (tcb_Ptr (p+0x100)))") apply (clarsimp simp: ctcb_relation_def ran_tcb_cte_cases) apply (clarsimp simp: option_to_ptr_def option_to_0_def split: option.splits) - apply (auto simp: cfault_rel_imp_eq cthread_state_rel_imp_eq + apply (auto simp: cfault_rel_imp_eq cthread_state_rel_imp_eq carch_tcb_relation_imp_eq ccontext_relation_imp_eq up_ucast_inj_eq) done diff --git a/proof/crefine/CLevityCatch.thy b/proof/crefine/CLevityCatch.thy index bf440642f..156b61d7c 100644 --- a/proof/crefine/CLevityCatch.thy +++ b/proof/crefine/CLevityCatch.thy @@ -126,7 +126,7 @@ lemma asUser_mapM_x: lemma asUser_get_registers: "\tcb_at' target\ asUser target (mapM getRegister xs) - \\rv s. obj_at' (\tcb. map (tcbContext tcb) xs = rv) target s\" + \\rv s. obj_at' (\tcb. map ((atcbContextGet o tcbArch) tcb) xs = rv) target s\" apply (induct xs) apply (simp add: mapM_empty asUser_return) apply wp diff --git a/proof/crefine/CSpaceAcc_C.thy b/proof/crefine/CSpaceAcc_C.thy index 31db06b75..e77368826 100644 --- a/proof/crefine/CSpaceAcc_C.thy +++ b/proof/crefine/CSpaceAcc_C.thy @@ -191,8 +191,8 @@ lemmas ccorres_getSlotCap_cte_at = ccorres_guard_from_wp [OF getSlotCap_pre_cte_ ccorres_guard_from_wp_bind [OF getSlotCap_pre_cte_at empty_fail_getSlotCap] lemma wordFromRights_spec: - defines "crl s \ (cap_rights_lift \<^bsup>s\<^esup>cap_rights)" - shows "\s. \ \ {s} \ret__unsigned_long :== PROC wordFromRights(\cap_rights) + defines "crl s \ (seL4_CapRights_lift \<^bsup>s\<^esup>seL4_CapRights)" + shows "\s. \ \ {s} \ret__unsigned_long :== PROC wordFromRights(\seL4_CapRights) \ \ret__unsigned_long = ((capAllowGrant_CL (crl s) << 2) || (capAllowRead_CL (crl s) << 1) @@ -200,7 +200,7 @@ lemma wordFromRights_spec: unfolding crl_def apply vcg apply (simp add: word_sle_def word_sless_def) - apply (simp add: cap_rights_lift_def shiftr_over_and_dist) + apply (simp add: seL4_CapRights_lift_def shiftr_over_and_dist) apply (simp add: mask_shift_simps) apply (simp add: word_ao_dist2[symmetric]) done @@ -211,7 +211,7 @@ lemma errlookup_fault_errstate [simp]: by simp lemma errfault_errstate [simp]: - "errfault (errstate s) = fault_lift (current_fault_' (globals s))" + "errfault (errstate s) = seL4_Fault_lift (current_fault_' (globals s))" unfolding errstate_def by simp diff --git a/proof/crefine/CSpace_C.thy b/proof/crefine/CSpace_C.thy index 077a14f71..33cebc1f7 100644 --- a/proof/crefine/CSpace_C.thy +++ b/proof/crefine/CSpace_C.thy @@ -59,7 +59,7 @@ lemma maskVMRights_spec: \ \vm_rights && mask 2 = \vm_rights \) Call maskVMRights_'proc \ vmrights_to_H \ret__unsigned_long = - maskVMRights (vmrights_to_H \<^bsup>s\<^esup>vm_rights) (cap_rights_to_H (cap_rights_lift \<^bsup>s\<^esup>cap_rights_mask)) \ + maskVMRights (vmrights_to_H \<^bsup>s\<^esup>vm_rights) (cap_rights_to_H (seL4_CapRights_lift \<^bsup>s\<^esup>cap_rights_mask)) \ \ret__unsigned_long && mask 2 = \ret__unsigned_long \" apply vcg apply clarsimp @@ -72,7 +72,7 @@ lemma maskVMRights_spec: apply clarsimp apply (subgoal_tac "vm_rights = 0 \ vm_rights = 1 \ vm_rights = 2 \ vm_rights = 3") apply (auto simp: vmrights_to_H_def maskVMRights_def vmrights_defs - cap_rights_to_H_def cap_rights_lift_def + cap_rights_to_H_def seL4_CapRights_lift_def to_bool_def mask_def split: bool.splits)[1] apply (subst(asm) mask_eq_iff_w2p) @@ -172,14 +172,14 @@ lemma to_bool_mask_to_bool_bf: done lemma to_bool_cap_rights_bf: - "to_bool (capAllowRead_CL (cap_rights_lift R)) = - to_bool_bf (capAllowRead_CL (cap_rights_lift R))" - "to_bool (capAllowWrite_CL (cap_rights_lift R)) = - to_bool_bf (capAllowWrite_CL (cap_rights_lift R))" - "to_bool (capAllowGrant_CL (cap_rights_lift R)) = - to_bool_bf (capAllowGrant_CL (cap_rights_lift R))" + "to_bool (capAllowRead_CL (seL4_CapRights_lift R)) = + to_bool_bf (capAllowRead_CL (seL4_CapRights_lift R))" + "to_bool (capAllowWrite_CL (seL4_CapRights_lift R)) = + to_bool_bf (capAllowWrite_CL (seL4_CapRights_lift R))" + "to_bool (capAllowGrant_CL (seL4_CapRights_lift R)) = + to_bool_bf (capAllowGrant_CL (seL4_CapRights_lift R))" by (subst to_bool_bf_to_bool_mask, - simp add: cap_rights_lift_def mask_def word_bw_assocs, simp)+ + simp add: seL4_CapRights_lift_def mask_def word_bw_assocs, simp)+ lemma to_bool_ntfn_cap_bf: "cap_lift c = Some (Cap_notification_cap cap) \ @@ -2261,7 +2261,8 @@ show ?thesis apply (simp add: from_bool_def) apply (cases irqState, simp_all) apply (simp add: Kernel_C.IRQSignal_def Kernel_C.IRQInactive_def) - apply (simp add: Kernel_C.IRQTimer_def Kernel_C.IRQInactive_def) + apply (simp add: Kernel_C.IRQTimer_def Kernel_C.IRQInactive_def) + apply (simp add: Kernel_C.IRQInactive_def Kernel_C.IRQReserved_def) done qed diff --git a/proof/crefine/CSpace_RAB_C.thy b/proof/crefine/CSpace_RAB_C.thy index a511dbbef..fc7869fd3 100644 --- a/proof/crefine/CSpace_RAB_C.thy +++ b/proof/crefine/CSpace_RAB_C.thy @@ -605,17 +605,17 @@ abbreviation lemma rightsFromWord_spec: - shows "\s. \ \ {s} \ret__struct_cap_rights_C :== PROC rightsFromWord(\w) - \cap_rights_lift \ret__struct_cap_rights_C = cap_rights_from_word_canon \<^bsup>s\<^esup>w \" + shows "\s. \ \ {s} \ret__struct_seL4_CapRights_C :== PROC rightsFromWord(\w) + \seL4_CapRights_lift \ret__struct_seL4_CapRights_C = cap_rights_from_word_canon \<^bsup>s\<^esup>w \" apply vcg - apply (simp add: cap_rights_lift_def nth_shiftr mask_shift_simps nth_shiftr + apply (simp add: seL4_CapRights_lift_def nth_shiftr mask_shift_simps nth_shiftr cap_rights_from_word_canon_def from_bool_def word_and_1 eval_nat_numeral word_sless_def word_sle_def) done abbreviation - "lookupSlot_rel' \ \(cte, rm) (cte', rm'). cte' = Ptr cte \ rm = cap_rights_to_H (cap_rights_lift rm')" + "lookupSlot_rel' \ \(cte, rm) (cte', rm'). cte' = Ptr cte \ rm = cap_rights_to_H (seL4_CapRights_lift rm')" (* MOVE *) lemma cap_rights_to_H_from_word_canon [simp]: diff --git a/proof/crefine/Fastpath_C.thy b/proof/crefine/Fastpath_C.thy index f8af45b24..a7c684041 100644 --- a/proof/crefine/Fastpath_C.thy +++ b/proof/crefine/Fastpath_C.thy @@ -182,98 +182,45 @@ lemma setEndpoint_obj_at_tcb': apply (clarsimp simp: updateObject_default_def in_monad) done -lemma tcbSchedEnqueue_obj_at_unchangedT: - assumes y: "\f. \tcb. P (tcbQueued_update f tcb) = P tcb" - shows "\obj_at' P t\ tcbSchedEnqueue t' \\rv. obj_at' P t\" - apply (simp add: tcbSchedEnqueue_def unless_def) - apply (wp | simp add: y)+ - done - -lemma asUser_obj_at_notQ: - "\obj_at' (Not \ tcbQueued) t\ - asUser t (setRegister r v) - \\rv. obj_at' (Not \ tcbQueued) t\" - apply (simp add: asUser_def) - apply (rule hoare_seq_ext)+ - apply (simp add: split_def) - apply (rule threadSet_obj_at'_really_strongest) - apply (wp threadGet_wp |rule gets_inv|wpc|clarsimp)+ - apply (simp add: setRegister_def simpler_modify_def) - apply (clarsimp simp: obj_at'_def) - done - -(* FIXME: Move to Schedule_R.thy. Make Arch_switchToThread_obj_at a specialisation of this *) -lemma Arch_switchToThread_obj_at_pre: - "\obj_at' (Not \ tcbQueued) t\ - Arch.switchToThread t - \\rv. obj_at' (Not \ tcbQueued) t\" - apply (simp add: ARM_H.switchToThread_def) - apply (wp asUser_obj_at_notQ doMachineOp_obj_at setVMRoot_obj_at hoare_drop_imps|wpc)+ - done - -lemma rescheduleRequired_obj_at_unchangedT: - assumes y: "\f. \tcb. P (tcbQueued_update f tcb) = P tcb" - shows "\obj_at' P t\ rescheduleRequired \\rv. obj_at' P t\" - apply (simp add: rescheduleRequired_def) - apply (wp tcbSchedEnqueue_obj_at_unchangedT[OF y] | wpc)+ - apply simp - done - -lemma setThreadState_obj_at_unchangedT: - assumes x: "\f. \tcb. P (tcbState_update f tcb) = P tcb" - assumes y: "\f. \tcb. P (tcbQueued_update f tcb) = P tcb" - shows "\obj_at' P t\ setThreadState t' ts \\rv. obj_at' P t\" - apply (simp add: setThreadState_def) - apply (wp rescheduleRequired_obj_at_unchangedT[OF y], simp) - apply (wp threadSet_obj_at'_strongish) - apply (clarsimp simp: obj_at'_def projectKOs x cong: if_cong) - done - -lemma setBoundNotification_obj_at_unchangedT: - assumes x: "\f. \tcb. P (tcbBoundNotification_update f tcb) = P tcb" - shows "\obj_at' P t\ setBoundNotification t' ts \\rv. obj_at' P t\" - apply (simp add: setBoundNotification_def) - apply (wp threadSet_obj_at'_strongish) - apply (clarsimp simp: obj_at'_def projectKOs x cong: if_cong) - done - -lemmas setThreadState_obj_at_unchanged - = setThreadState_obj_at_unchangedT[OF all_tcbI all_tcbI] - -lemmas setBoundNotification_obj_at_unchanged - = setBoundNotification_obj_at_unchangedT[OF all_tcbI] - lemma tcbSchedEnqueue_tcbContext[wp]: - "\obj_at' (\tcb. P (tcbContext tcb)) t\ + "\obj_at' (\tcb. P ((atcbContextGet o tcbArch) tcb)) t\ tcbSchedEnqueue t' - \\rv. obj_at' (\tcb. P (tcbContext tcb)) t\" + \\rv. obj_at' (\tcb. P ((atcbContextGet o tcbArch) tcb)) t\" apply (rule tcbSchedEnqueue_obj_at_unchangedT[OF all_tcbI]) apply simp done -lemma setNotification_tcb: - "\obj_at' (\tcb::tcb. P tcb) t\ - setNotification ntfn e - \\_. obj_at' P t\" - apply (simp add: setNotification_def) - apply (rule obj_at_setObject2) - apply (clarsimp simp: updateObject_default_def in_monad) - done - lemma setCTE_tcbContext: - "\obj_at' (\tcb. P (tcbContext tcb)) t\ + "\obj_at' (\tcb. P ((atcbContextGet o tcbArch) tcb)) t\ setCTE slot cte - \\rv. obj_at' (\tcb. P (tcbContext tcb)) t\" + \\rv. obj_at' (\tcb. P ((atcbContextGet o tcbArch) tcb)) t\" apply (simp add: setCTE_def) apply (rule setObject_cte_obj_at_tcb', simp_all) done +lemma seThreadState_tcbContext: + "\obj_at' (\tcb. P ((atcbContextGet o tcbArch) tcb)) t\ + setThreadState a b + \\_. obj_at' (\tcb. P ((atcbContextGet o tcbArch) tcb)) t\" + apply (rule setThreadState_obj_at_unchanged) + apply (clarsimp simp: atcbContext_def)+ + done +lemma setBoundNotification_tcbContext: + "\obj_at' (\tcb. P ((atcbContextGet o tcbArch) tcb)) t\ + setBoundNotification a b + \\_. obj_at' (\tcb. P ((atcbContextGet o tcbArch) tcb)) t\" + apply (rule setBoundNotification_obj_at_unchanged) + apply (clarsimp simp: atcbContext_def)+ + done + +declare comp_apply [simp del] +crunch tcbContext[wp]: deleteCallerCap "obj_at' (\tcb. P ((atcbContextGet o tcbArch) tcb)) t" + (wp: setEndpoint_obj_at_tcb' setBoundNotification_tcbContext + setNotification_tcb crunch_wps seThreadState_tcbContext + ignore: getObject setObject simp: crunch_simps unless_def) +declare comp_apply [simp] -crunch tcbContext[wp]: deleteCallerCap "obj_at' (\tcb. P (tcbContext tcb)) t" - (wp: setEndpoint_obj_at_tcb' setThreadState_obj_at_unchanged - setNotification_tcb crunch_wps setBoundNotification_obj_at_unchanged - simp: crunch_simps unless_def) crunch ksArch[wp]: asUser "\s. P (ksArchState s)" (wp: crunch_wps) @@ -741,10 +688,6 @@ lemma ccap_relation_case_sum_Null_endpoint: by (clarsimp simp: cap_get_tag_isCap isRight_def isCap_simps split: sum.split_asm) -lemma empty_fail_isRunnable: - "empty_fail (isRunnable t)" - by (simp add: isRunnable_def isBlocked_def) - lemma findPDForASID_pd_at_asid_noex: "\pd_at_asid' pd asid\ findPDForASID asid \\rv s. rv = pd\,\\\\" apply (simp add: findPDForASID_def @@ -1044,7 +987,7 @@ lemma armv_contextSwitch_HWASID_fp_rewrite: apply (wp | simp)+ apply (simp add: empty_fail_findPDForASID empty_fail_catch) apply (rule monadic_rewrite_assert monadic_rewrite_gets_l)+ - apply (rule_tac P="x asid \ None \ fst (the (x asid)) = the (pde_stored_asid v)" + apply (rule_tac P="asidMap asid \ None \ fst (the (asidMap asid)) = the (pde_stored_asid v)" in monadic_rewrite_gen_asm) apply (simp only: case_option_If2 simp_thms if_True if_False split_def, simp) @@ -1342,11 +1285,6 @@ lemma endpoint_ptr_set_epQueue_head_np_spec: apply (simp add: limited_and_simps) done -lemmas empty_fail_user_getreg = empty_fail_asUser[OF empty_fail_getRegister] - -lemma empty_fail_getCurThread[iff]: - "empty_fail getCurThread" by (simp add: getCurThread_def) - lemma ccorres_call_hSkip': assumes cul: "ccorres_underlying sr \ r xf' r xf' P (i ` P') [SKIP] a (Call f)" and gsr: "\a b x s t. (x, t) \ sr \ (x, g a b (clean s t)) \ sr" @@ -1578,20 +1516,6 @@ lemma ccorres_cond_both_seq: apply auto done -lemma copyMRs_simple: - "msglen \ of_nat (length ARM_H.msgRegisters) \ - copyMRs sender sbuf receiver rbuf msglen - = forM_x (take (unat msglen) ARM_H.msgRegisters) - (\r. do v \ asUser sender (getRegister r); - asUser receiver (setRegister r v) od) - >>= (\rv. return msglen)" - apply (clarsimp simp: copyMRs_def mapM_discarded) - apply (rule bind_cong[OF refl]) - apply (simp add: length_msgRegisters n_msgRegisters_def min_def - word_le_nat_alt - split: option.split) - apply (simp add: upto_enum_def mapM_Nil) - done lemma unifyFailure_catch_If: "catch (unifyFailure f >>=E g) h @@ -2111,7 +2035,7 @@ lemma ctes_of_Some_cte_wp_at: by (clarsimp simp: cte_wp_at_ctes_of) lemma user_getreg_wp: - "\\s. tcb_at' t s \ (\rv. obj_at' (\tcb. tcbContext tcb r = rv) t s \ Q rv s)\ + "\\s. tcb_at' t s \ (\rv. obj_at' (\tcb. (atcbContextGet o tcbArch) tcb r = rv) t s \ Q rv s)\ asUser t (getRegister r) \Q\" apply (rule_tac Q="\rv s. \rv'. rv' = rv \ Q rv' s" in hoare_post_imp) apply simp @@ -2176,23 +2100,6 @@ lemma setUntypedCapAsFull_replyCap[simp]: end -context begin interpretation Arch . (*FIXME: arch_split*) - -lemma getObject_return: - fixes v :: "'a :: pspace_storable" shows - "\ \a b c d. (loadObject a b c d :: 'a kernel) = loadObject_default a b c d; - ko_at' v p s; (1 :: word32) < 2 ^ objBits v \ \ getObject p s = return v s" - apply (clarsimp simp: getObject_def split_def exec_gets - obj_at'_def projectKOs lookupAround2_known1 - assert_opt_def loadObject_default_def) - apply (simp add: projectKO_def alignCheck_assert) - apply (simp add: project_inject objBits_def) - apply (frule(2) in_magnitude_check[where s'=s]) - apply (simp add: magnitudeCheck_assert in_monad) - done - -end - context kernel_m begin lemma obj_at_bound_tcb_grandD: @@ -2252,8 +2159,8 @@ lemma fastpath_call_ccorres: notes hoare_TrueI[simp] shows "ccorres dc xfdc (\s. invs' s \ ct_in_state' (op = Running) s - \ obj_at' (\tcb. tcbContext tcb ARM_H.capRegister = cptr - \ tcbContext tcb ARM_H.msgInfoRegister = msginfo) + \ obj_at' (\tcb. (atcbContextGet o tcbArch) tcb ARM_H.capRegister = cptr + \ (atcbContextGet o tcbArch) tcb ARM_H.msgInfoRegister = msginfo) (ksCurThread s) s) (UNIV \ {s. cptr_' s = cptr} \ {s. msgInfo_' s = msginfo}) [] (fastpaths SysCall) (Call fastpath_call_'proc)" @@ -2306,7 +2213,7 @@ lemma fastpath_call_ccorres: apply (simp del: Collect_const cong: call_ignore_cong) apply (simp only: ) apply (csymbr, csymbr) - apply (rule_tac r'="\ft ft'. (ft' = scast fault_null_fault) = (ft = None)" + apply (rule_tac r'="\ft ft'. (ft' = scast seL4_Fault_NullFault) = (ft = None)" and xf'="fault_type_'" in ccorres_split_nothrow) apply (rule_tac P="cur_tcb' and (\s. curThread = ksCurThread s)" in ccorres_from_vcg[where P'=UNIV]) @@ -2315,7 +2222,7 @@ lemma fastpath_call_ccorres: apply (drule(1) obj_at_cslift_tcb, clarsimp) apply (clarsimp simp: typ_heap_simps' ctcb_relation_def cfault_rel_def) apply (rule rev_bexI, erule threadGet_eq) - apply (clarsimp simp: fault_lift_def Let_def split: split_if_asm) + apply (clarsimp simp: seL4_Fault_lift_def Let_def split: split_if_asm) apply ceqv apply csymbr apply (simp del: Collect_const cong: call_ignore_cong) @@ -2938,8 +2845,8 @@ lemma fastpath_reply_recv_ccorres: notes hoare_TrueI[simp] shows "ccorres dc xfdc (\s. invs' s \ ct_in_state' (op = Running) s - \ obj_at' (\tcb. tcbContext tcb capRegister = cptr - \ tcbContext tcb msgInfoRegister = msginfo) + \ obj_at' (\tcb. (atcbContextGet o tcbArch) tcb capRegister = cptr + \ (atcbContextGet o tcbArch) tcb msgInfoRegister = msginfo) (ksCurThread s) s) (UNIV \ {s. cptr_' s = cptr} \ {s. msgInfo_' s = msginfo}) [] (fastpaths SysReplyRecv) (Call fastpath_reply_recv_'proc)" @@ -2993,7 +2900,7 @@ lemma fastpath_reply_recv_ccorres: apply (simp del: Collect_const cong: call_ignore_cong) apply (simp only:) apply (csymbr, csymbr) - apply (rule_tac r'="\ft ft'. (ft' = scast fault_null_fault) = (ft = None)" + apply (rule_tac r'="\ft ft'. (ft' = scast seL4_Fault_NullFault) = (ft = None)" and xf'="fault_type_'" in ccorres_split_nothrow) apply (rule_tac P="cur_tcb' and (\s. curThread = ksCurThread s)" in ccorres_from_vcg[where P'=UNIV]) @@ -3002,7 +2909,7 @@ lemma fastpath_reply_recv_ccorres: apply (drule(1) obj_at_cslift_tcb, clarsimp) apply (clarsimp simp: typ_heap_simps' ctcb_relation_def cfault_rel_def) apply (rule rev_bexI, erule threadGet_eq) - apply (clarsimp simp: fault_lift_def Let_def split: split_if_asm) + apply (clarsimp simp: seL4_Fault_lift_def Let_def split: split_if_asm) apply ceqv apply csymbr apply (simp only:) @@ -3154,14 +3061,14 @@ lemma fastpath_reply_recv_ccorres: apply (vcg exspec=slowpath_noreturn_spec) apply (simp del: Collect_const cong: call_ignore_cong) apply (csymbr, csymbr) - apply (rule_tac r'="\ft ft'. (ft' = scast fault_null_fault) = (ft = None)" + apply (rule_tac r'="\ft ft'. (ft' = scast seL4_Fault_NullFault) = (ft = None)" and xf'="fault_type_'" in ccorres_split_nothrow) apply (rule threadGet_vcg_corres) apply (rule allI, rule conseqPre, vcg) apply (clarsimp simp: obj_at_tcbs_of) apply (clarsimp simp: typ_heap_simps' ctcb_relation_def cfault_rel_def ccap_relation_reply_helper) - apply (clarsimp simp: fault_lift_def Let_def split: split_if_asm) + apply (clarsimp simp: seL4_Fault_lift_def Let_def split: split_if_asm) apply ceqv apply (simp del: Collect_const not_None_eq cong: call_ignore_cong) apply (rule ccorres_Cond_rhs_Seq) @@ -3466,7 +3373,7 @@ lemma fastpath_reply_recv_ccorres: apply (simp add:ccap_relation_reply_helper cong:if_cong) apply (rule threadGet_wp) apply (simp add: syscall_from_H_def ccap_relation_reply_helper) - apply (vcg exspec=fault_get_faultType_modifies) + apply (vcg exspec=seL4_Fault_get_seL4_FaultType_modifies) apply simp apply wp apply simp @@ -3566,1589 +3473,11 @@ qed end -datatype tcb_state_regs = TCBStateRegs "thread_state" "MachineTypes.register \ machine_word" - -definition - "tsrContext tsr \ case tsr of TCBStateRegs ts regs \ regs" - -definition - "tsrState tsr \ case tsr of TCBStateRegs ts regs \ ts" - -lemma accessors_TCBStateRegs[simp]: - "TCBStateRegs (tsrState v) (tsrContext v) = v" - by (cases v, simp add: tsrState_def tsrContext_def) - -lemma tsrContext_simp[simp]: - "tsrContext (TCBStateRegs st con) = con" - by (simp add: tsrContext_def) - -lemma tsrState_simp[simp]: - "tsrState (TCBStateRegs st con) = st" - by (simp add: tsrState_def) - -definition - get_tcb_state_regs :: "kernel_object option \ tcb_state_regs" -where - "get_tcb_state_regs oko \ case oko of - Some (KOTCB tcb) \ TCBStateRegs (tcbState tcb) (tcbContext tcb)" - -definition - put_tcb_state_regs_tcb :: "tcb_state_regs \ tcb \ tcb" -where - "put_tcb_state_regs_tcb tsr tcb \ case tsr of - TCBStateRegs st regs \ tcb \ tcbState := st, tcbContext := regs \" - -definition - put_tcb_state_regs :: "tcb_state_regs \ kernel_object option \ kernel_object option" -where - "put_tcb_state_regs tsr oko = Some (KOTCB (put_tcb_state_regs_tcb tsr - (case oko of - Some (KOTCB tcb) \ tcb | _ \ makeObject)))" - -definition - "partial_overwrite idx tcbs ps \ - \x. if x \ range idx - then put_tcb_state_regs (tcbs (inv idx x)) (ps x) - else ps x" - -definition - isolate_thread_actions :: "('x \ word32) \ 'a kernel - \ (('x \ tcb_state_regs) \ ('x \ tcb_state_regs)) - \ (scheduler_action \ scheduler_action) - \ 'a kernel" -where - "isolate_thread_actions idx m t f \ do - s \ gets (ksSchedulerAction_update (\_. ResumeCurrentThread) - o ksPSpace_update (partial_overwrite idx (K undefined))); - tcbs \ gets (\s. get_tcb_state_regs o ksPSpace s o idx); - sa \ getSchedulerAction; - (rv, s') \ select_f (m s); - modify (\s. ksPSpace_update (partial_overwrite idx (t tcbs)) - (s' \ ksSchedulerAction := f sa \)); - return rv - od" - -lemma put_tcb_state_regs_twice[simp]: - "put_tcb_state_regs tsr (put_tcb_state_regs tsr' tcb) - = put_tcb_state_regs tsr tcb" - apply (simp add: put_tcb_state_regs_def put_tcb_state_regs_tcb_def - makeObject_tcb - split: tcb_state_regs.split option.split - Structures_H.kernel_object.split) - apply (intro all_tcbI impI allI) - apply simp - done - -lemma partial_overwrite_twice[simp]: - "partial_overwrite idx f (partial_overwrite idx g ps) - = partial_overwrite idx f ps" - by (rule ext, simp add: partial_overwrite_def) - -lemma get_tcb_state_regs_partial_overwrite[simp]: - "inj idx \ - get_tcb_state_regs (partial_overwrite idx tcbs f (idx x)) - = tcbs x" - apply (simp add: partial_overwrite_def) - apply (simp add: put_tcb_state_regs_def - get_tcb_state_regs_def - put_tcb_state_regs_tcb_def - split: tcb_state_regs.split) - done - -lemma isolate_thread_actions_bind: - "inj idx \ - isolate_thread_actions idx a b c >>= - (\x. isolate_thread_actions idx (d x) e f) - = isolate_thread_actions idx a id id - >>= (\x. isolate_thread_actions idx (d x) (e o b) (f o c))" - apply (rule ext) - apply (clarsimp simp: isolate_thread_actions_def bind_assoc split_def - bind_select_f_bind[symmetric]) - apply (clarsimp simp: exec_gets getSchedulerAction_def) - apply (rule select_bind_eq) - apply (simp add: exec_gets exec_modify o_def) - apply (rule select_bind_eq) - apply (simp add: exec_gets exec_modify) - done - -context kernel_m begin - -lemma setObject_modify: - fixes v :: "'a :: pspace_storable" shows - "\ obj_at' (P :: 'a \ bool) p s; updateObject v = updateObject_default v; - (1 :: word32) < 2 ^ objBits v \ - \ setObject p v s - = modify (ksPSpace_update (\ps. ps (p \ injectKO v))) s" - apply (clarsimp simp: setObject_def split_def exec_gets - obj_at'_def projectKOs lookupAround2_known1 - assert_opt_def updateObject_default_def - bind_assoc) - apply (simp add: projectKO_def alignCheck_assert) - apply (simp add: project_inject objBits_def) - apply (clarsimp simp only: objBitsT_koTypeOf[symmetric] koTypeOf_injectKO) - apply (frule(2) in_magnitude_check[where s'=s]) - apply (simp add: magnitudeCheck_assert in_monad) - apply (simp add: simpler_modify_def) - done - -lemmas getObject_return_tcb - = getObject_return[OF meta_eq_to_obj_eq, OF loadObject_tcb, - unfolded objBits_simps, simplified] - -lemmas setObject_modify_tcb - = setObject_modify[OF _ meta_eq_to_obj_eq, OF _ updateObject_tcb, - unfolded objBits_simps, simplified] - -lemma partial_overwrite_fun_upd: - "inj idx \ - partial_overwrite idx (tsrs (x := y)) - = (\ps. (partial_overwrite idx tsrs ps) (idx x := put_tcb_state_regs y (ps (idx x))))" - apply (intro ext, simp add: partial_overwrite_def) - apply (clarsimp split: split_if) - done - -lemma get_tcb_state_regs_ko_at': - "ko_at' ko p s \ get_tcb_state_regs (ksPSpace s p) - = TCBStateRegs (tcbState ko) (tcbContext ko)" - by (clarsimp simp: obj_at'_def projectKOs get_tcb_state_regs_def) - -lemma put_tcb_state_regs_ko_at': - "ko_at' ko p s \ put_tcb_state_regs tsr (ksPSpace s p) - = Some (KOTCB (ko \ tcbState := tsrState tsr, tcbContext := tsrContext tsr \))" - by (clarsimp simp: obj_at'_def projectKOs put_tcb_state_regs_def - put_tcb_state_regs_tcb_def - split: tcb_state_regs.split) - -lemma partial_overwrite_get_tcb_state_regs: - "\ \x. tcb_at' (idx x) s; inj idx \ \ - partial_overwrite idx (\x. get_tcb_state_regs (ksPSpace s (idx x))) - (ksPSpace s) = ksPSpace s" - apply (rule ext, simp add: partial_overwrite_def - split: split_if) - apply clarsimp - apply (drule_tac x=xa in spec) - apply (clarsimp simp: obj_at'_def projectKOs put_tcb_state_regs_def - get_tcb_state_regs_def put_tcb_state_regs_tcb_def) - apply (case_tac obj, simp) - done - -lemma ksPSpace_update_partial_id: - "\ \ps x. f ps x = ps (idx x) \ f ps x = ksPSpace s (idx x); - \x. tcb_at' (idx x) s; inj idx \ \ - ksPSpace_update (\ps. partial_overwrite idx (\x. get_tcb_state_regs (f ps x)) ps) s - = s" - apply (rule trans, rule kernel_state.fold_congs[OF refl refl]) - apply (erule_tac x="ksPSpace s" in meta_allE) - apply (clarsimp simp: partial_overwrite_get_tcb_state_regs) - apply (rule refl) - apply simp - done - -lemma isolate_thread_actions_asUser: - "\ idx t' = t; inj idx; f = (\s. ({(v, g s)}, False)) \ \ - monadic_rewrite False True (\s. \x. tcb_at' (idx x) s) - (asUser t f) - (isolate_thread_actions idx (return v) - (\tsrs. (tsrs (t' := TCBStateRegs (tsrState (tsrs t')) - (g (tsrContext (tsrs t')))))) - id)" - apply (simp add: asUser_def liftM_def isolate_thread_actions_def split_def - select_f_returns bind_assoc select_f_singleton_return - threadGet_def threadSet_def) - apply (clarsimp simp: monadic_rewrite_def) - apply (frule_tac x=t' in spec) - apply (drule obj_at_ko_at', clarsimp) - apply (simp add: exec_gets getSchedulerAction_def exec_modify - getObject_return_tcb setObject_modify_tcb o_def - cong: bind_apply_cong)+ - apply (simp add: partial_overwrite_fun_upd return_def get_tcb_state_regs_ko_at') - apply (rule kernel_state.fold_congs[OF refl refl]) - apply (clarsimp simp: partial_overwrite_get_tcb_state_regs - put_tcb_state_regs_ko_at') - apply (case_tac ko, simp) - done - -lemma getRegister_simple: - "getRegister r = (\con. ({(con r, con)}, False))" - by (simp add: getRegister_def simpler_gets_def) - -lemma mapM_getRegister_simple: - "mapM getRegister rs = (\con. ({(map con rs, con)}, False))" - apply (induct rs) - apply (simp add: mapM_Nil return_def) - apply (simp add: mapM_Cons getRegister_def simpler_gets_def - bind_def return_def) - done - -lemma setRegister_simple: - "setRegister r v = (\con. ({((), con (r := v))}, False))" - by (simp add: setRegister_def simpler_modify_def) - -lemma zipWithM_setRegister_simple: - "zipWithM_x setRegister rs vs - = (\con. ({((), foldl (\con (r, v). con (r := v)) con (zip rs vs))}, False))" - apply (simp add: zipWithM_x_mapM_x) - apply (induct ("zip rs vs")) - apply (simp add: mapM_x_Nil return_def) - apply (clarsimp simp add: mapM_x_Cons bind_def setRegister_def - simpler_modify_def fun_upd_def[symmetric]) - done - -lemma dom_partial_overwrite: - "\x. tcb_at' (idx x) s \ dom (partial_overwrite idx tsrs (ksPSpace s)) - = dom (ksPSpace s)" - apply (rule set_eqI) - apply (clarsimp simp: dom_def partial_overwrite_def put_tcb_state_regs_def - split: split_if) - apply (fastforce elim!: obj_atE') - done - -lemma map_to_ctes_partial_overwrite: - "\x. tcb_at' (idx x) s \ - map_to_ctes (partial_overwrite idx tsrs (ksPSpace s)) - = ctes_of s" - apply (rule ext) - apply (frule dom_partial_overwrite[where tsrs=tsrs]) - apply (simp add: map_to_ctes_def partial_overwrite_def - Let_def) - apply (case_tac "x \ range idx") - apply (clarsimp simp: put_tcb_state_regs_def) - apply (drule_tac x=xa in spec) - apply (clarsimp simp: obj_at'_def projectKOs objBits_simps - cong: if_cong) - apply (simp add: put_tcb_state_regs_def put_tcb_state_regs_tcb_def - objBits_simps - cong: if_cong option.case_cong) - apply (case_tac obj, simp split: tcb_state_regs.split split_if) - apply simp - apply (rule if_cong[OF refl]) - apply simp - apply (case_tac "x && ~~ mask (objBitsKO (KOTCB undefined)) \ range idx") - apply (clarsimp simp: put_tcb_state_regs_def) - apply (drule_tac x=xa in spec) - apply (clarsimp simp: obj_at'_def projectKOs objBits_simps - cong: if_cong) - apply (simp add: put_tcb_state_regs_def put_tcb_state_regs_tcb_def - objBits_simps - cong: if_cong option.case_cong) - apply (case_tac obj, simp split: tcb_state_regs.split split_if) - apply (intro impI allI) - apply (subgoal_tac "x - idx xa = x && mask 9") - apply (clarsimp simp: tcb_cte_cases_def split: split_if) - apply (drule_tac t = "idx xa" in sym) - apply simp - apply (simp cong: if_cong) - done - -definition - "thread_actions_isolatable idx f = - (inj idx \ monadic_rewrite False True (\s. \x. tcb_at' (idx x) s) - f (isolate_thread_actions idx f id id))" - -lemma getCTE_assert_opt: - "getCTE p = gets (\s. ctes_of s p) >>= assert_opt" - apply (intro ext) - apply (simp add: exec_gets assert_opt_def prod_eq_iff - fail_def return_def - split: option.split) - apply (rule conjI) - apply clarsimp - apply (rule context_conjI) - apply (rule ccontr, clarsimp elim!: nonemptyE) - apply (frule use_valid[OF _ getCTE_sp], rule TrueI) - apply (frule in_inv_by_hoareD[OF getCTE_inv]) - apply (clarsimp simp: cte_wp_at_ctes_of) - apply (simp add: empty_failD[OF empty_fail_getCTE]) - apply clarsimp - apply (simp add: no_failD[OF no_fail_getCTE, OF ctes_of_cte_at]) - apply (subgoal_tac "cte_wp_at' (op = x2) p x") - apply (clarsimp simp: cte_wp_at'_def getCTE_def) - apply (simp add: cte_wp_at_ctes_of) - done - -lemma getCTE_isolatable: - "thread_actions_isolatable idx (getCTE p)" - apply (clarsimp simp: thread_actions_isolatable_def) - apply (simp add: isolate_thread_actions_def bind_assoc split_def) - apply (simp add: getCTE_assert_opt bind_select_f_bind[symmetric] - bind_assoc select_f_returns) - apply (clarsimp simp: monadic_rewrite_def exec_gets getSchedulerAction_def - map_to_ctes_partial_overwrite) - apply (simp add: assert_opt_def select_f_returns select_f_asserts - split: option.split) - apply (clarsimp simp: exec_modify o_def return_def) - apply (simp add: ksPSpace_update_partial_id) - done - -lemma objBits_2n: - "(1 :: word32) < 2 ^ objBits obj" - by (simp add: objBits_def objBitsKO_def archObjSize_def pageBits_def - split: kernel_object.split arch_kernel_object.split) - -lemma magnitudeCheck_assert2: - "\ is_aligned x n; (1 :: word32) < 2 ^ n; ksPSpace s x = Some v \ \ - magnitudeCheck x (snd (lookupAround2 x (ksPSpace (s :: kernel_state)))) n - = assert (ps_clear x n s)" - using in_magnitude_check[where x=x and n=n and s=s and s'=s and v="()"] - by (simp add: magnitudeCheck_assert in_monad) - -lemma getObject_get_assert: - assumes deflt: "\a b c d. (loadObject a b c d :: ('a :: pspace_storable) kernel) - = loadObject_default a b c d" - shows - "(getObject p :: ('a :: pspace_storable) kernel) - = do v \ gets (obj_at' (\x :: 'a. True) p); - assert v; - gets (the o projectKO_opt o the o swp fun_app p o ksPSpace) - od" - apply (rule ext) - apply (simp add: exec_get getObject_def split_def exec_gets - deflt loadObject_default_def projectKO_def2 - alignCheck_assert) - apply (case_tac "ksPSpace x p") - apply (simp add: obj_at'_def assert_opt_def assert_def - split: option.split split_if) - apply (simp add: lookupAround2_known1 assert_opt_def - obj_at'_def projectKO_def2 - split: option.split) - apply (clarsimp simp: fail_def fst_return conj_comms project_inject - objBits_def) - apply (simp only: assert2[symmetric], - rule bind_apply_cong[OF refl]) - apply (clarsimp simp: in_monad) - apply (fold objBits_def) - apply (simp add: magnitudeCheck_assert2[OF _ objBits_2n]) - apply (rule bind_apply_cong[OF refl]) - apply (clarsimp simp: in_monad return_def simpler_gets_def) - apply (simp add: iffD2[OF project_inject refl]) - done - -lemma obj_at_partial_overwrite_If: - "\ \x. tcb_at' (idx x) s \ - \ obj_at' P p (ksPSpace_update (partial_overwrite idx f) s) - = (if p \ range idx - then obj_at' (\tcb. P (put_tcb_state_regs_tcb (f (inv idx p)) tcb)) p s - else obj_at' P p s)" - apply (frule dom_partial_overwrite[where tsrs=f]) - apply (simp add: obj_at'_def ps_clear_def partial_overwrite_def - projectKOs split: split_if) - apply clarsimp - apply (drule_tac x=x in spec) - apply (clarsimp simp: put_tcb_state_regs_def objBits_simps) - done - -lemma obj_at_partial_overwrite_id1: - "\ p \ range idx; \x. tcb_at' (idx x) s \ - \ obj_at' P p (ksPSpace_update (partial_overwrite idx f) s) - = obj_at' P p s" - apply (drule dom_partial_overwrite[where tsrs=f]) - apply (simp add: obj_at'_def ps_clear_def partial_overwrite_def - projectKOs) - done - -lemma obj_at_partial_overwrite_id2: - "\ \x. tcb_at' (idx x) s; \v tcb. P v \ True \ injectKO v \ KOTCB tcb \ - \ obj_at' P p (ksPSpace_update (partial_overwrite idx f) s) - = obj_at' P p s" - apply (frule dom_partial_overwrite[where tsrs=f]) - apply (simp add: obj_at'_def ps_clear_def partial_overwrite_def - projectKOs split: split_if) - apply clarsimp - apply (drule_tac x=x in spec) - apply (clarsimp simp: put_tcb_state_regs_def objBits_simps - project_inject) - done - -lemma getObject_isolatable: - "\ \a b c d. (loadObject a b c d :: 'a kernel) = loadObject_default a b c d; - \tcb. projectKO_opt (KOTCB tcb) = (None :: 'a option) \ \ - thread_actions_isolatable idx (getObject p :: ('a :: pspace_storable) kernel)" - apply (clarsimp simp: thread_actions_isolatable_def) - apply (simp add: getObject_get_assert split_def - isolate_thread_actions_def bind_select_f_bind[symmetric] - bind_assoc select_f_asserts select_f_returns) - apply (clarsimp simp: monadic_rewrite_def exec_gets getSchedulerAction_def) - apply (case_tac "p \ range idx") - apply clarsimp - apply (drule_tac x=x in spec) - apply (clarsimp simp: obj_at'_def projectKOs partial_overwrite_def - put_tcb_state_regs_def) - apply (simp add: obj_at_partial_overwrite_id1) - apply (simp add: partial_overwrite_def) - apply (rule bind_apply_cong[OF refl]) - apply (simp add: exec_modify return_def o_def simpler_gets_def - ksPSpace_update_partial_id in_monad) - done - -lemma gets_isolatable: - "\\g s. \x. tcb_at' (idx x) s \ - f (ksSchedulerAction_update g - (ksPSpace_update (partial_overwrite idx (\_. undefined)) s)) = f s \ \ - thread_actions_isolatable idx (gets f)" - apply (clarsimp simp: thread_actions_isolatable_def) - apply (simp add: isolate_thread_actions_def select_f_returns - liftM_def bind_assoc) - apply (clarsimp simp: monadic_rewrite_def exec_gets - getSchedulerAction_def exec_modify) - apply (simp add: simpler_gets_def return_def - ksPSpace_update_partial_id o_def) - done - -lemma modify_isolatable: - assumes swap:"\tsrs act s. \x. tcb_at' (idx x) s \ - (ksPSpace_update (partial_overwrite idx tsrs) ((f s)\ ksSchedulerAction := act \)) - = f (ksPSpace_update (partial_overwrite idx tsrs) - (s \ ksSchedulerAction := act\))" - shows - "thread_actions_isolatable idx (modify f)" - apply (clarsimp simp: thread_actions_isolatable_def) - apply (simp add: isolate_thread_actions_def select_f_returns - liftM_def bind_assoc) - apply (clarsimp simp: monadic_rewrite_def exec_gets - getSchedulerAction_def) - apply (simp add: simpler_modify_def o_def) - apply (subst swap) - apply (simp add: obj_at_partial_overwrite_If) - apply (simp add: ksPSpace_update_partial_id o_def) - done - -lemma isolate_thread_actions_wrap_bind: - "inj idx \ - do x \ isolate_thread_actions idx a b c; - isolate_thread_actions idx (d x) e f - od = - isolate_thread_actions idx - (do x \ isolate_thread_actions idx a id id; - isolate_thread_actions idx (d x) id id - od) (e o b) (f o c) - " - apply (rule ext) - apply (clarsimp simp: isolate_thread_actions_def bind_assoc split_def - bind_select_f_bind[symmetric] liftM_def - select_f_returns select_f_selects - getSchedulerAction_def) - apply (clarsimp simp: exec_gets getSchedulerAction_def o_def) - apply (rule select_bind_eq) - apply (simp add: exec_gets exec_modify o_def) - apply (rule select_bind_eq) - apply (simp add: exec_modify) - done - -lemma monadic_rewrite_in_isolate_thread_actions: - "\ inj idx; monadic_rewrite F True P a d \ \ - monadic_rewrite F True (\s. P (ksSchedulerAction_update (\_. ResumeCurrentThread) - (ksPSpace_update (partial_overwrite idx (\_. undefined)) s))) - (isolate_thread_actions idx a b c) (isolate_thread_actions idx d b c)" - apply (clarsimp simp: isolate_thread_actions_def split_def) - apply (rule monadic_rewrite_bind_tail)+ - apply (rule_tac P="\_. P s" in monadic_rewrite_bind_head) - apply (simp add: monadic_rewrite_def select_f_def) - apply wp - apply simp - done - -lemma thread_actions_isolatable_bind: - "\ thread_actions_isolatable idx f; \x. thread_actions_isolatable idx (g x); - \t. \tcb_at' t\ f \\rv. tcb_at' t\ \ - \ thread_actions_isolatable idx (f >>= g)" - apply (clarsimp simp: thread_actions_isolatable_def) - apply (rule monadic_rewrite_imp) - apply (rule monadic_rewrite_trans) - apply (erule monadic_rewrite_bind2, assumption) - apply (rule hoare_vcg_all_lift, assumption) - apply (subst isolate_thread_actions_wrap_bind, simp) - apply simp - apply (rule monadic_rewrite_in_isolate_thread_actions, assumption) - apply (rule monadic_rewrite_transverse) - apply (erule monadic_rewrite_bind2, assumption) - apply (rule hoare_vcg_all_lift, assumption) - apply (simp add: bind_assoc id_def) - apply (rule monadic_rewrite_refl) - apply (simp add: obj_at_partial_overwrite_If) - done - -lemma thread_actions_isolatable_return: - "thread_actions_isolatable idx (return v)" - apply (clarsimp simp: thread_actions_isolatable_def - monadic_rewrite_def liftM_def - isolate_thread_actions_def - split_def bind_assoc select_f_returns - exec_gets getSchedulerAction_def) - apply (simp add: exec_modify return_def o_def - ksPSpace_update_partial_id) - done - -lemma thread_actions_isolatable_fail: - "thread_actions_isolatable idx fail" - by (simp add: thread_actions_isolatable_def - isolate_thread_actions_def select_f_asserts - liftM_def bind_assoc getSchedulerAction_def - monadic_rewrite_def exec_gets) - -lemma thread_actions_isolatable_returns: - "thread_actions_isolatable idx (return v)" - "thread_actions_isolatable idx (returnOk v)" - "thread_actions_isolatable idx (throwError v)" - by (simp add: returnOk_def throwError_def - thread_actions_isolatable_return)+ - -lemma thread_actions_isolatable_bindE: - "\ thread_actions_isolatable idx f; \x. thread_actions_isolatable idx (g x); - \t. \tcb_at' t\ f \\rv. tcb_at' t\ \ - \ thread_actions_isolatable idx (f >>=E g)" - apply (simp add: bindE_def) - apply (erule thread_actions_isolatable_bind) - apply (simp add: lift_def thread_actions_isolatable_returns - split: sum.split) - apply assumption - done - -lemma thread_actions_isolatable_catch: - "\ thread_actions_isolatable idx f; \x. thread_actions_isolatable idx (g x); - \t. \tcb_at' t\ f \\rv. tcb_at' t\ \ - \ thread_actions_isolatable idx (f g)" - apply (simp add: catch_def) - apply (erule thread_actions_isolatable_bind) - apply (simp add: thread_actions_isolatable_returns - split: sum.split) - apply assumption - done - -lemma thread_actions_isolatable_if: - "\ P \ thread_actions_isolatable idx a; - \ P \ thread_actions_isolatable idx b \ - \ thread_actions_isolatable idx (if P then a else b)" - by (cases P, simp_all) - -lemma select_f_isolatable: - "thread_actions_isolatable idx (select_f v)" - apply (clarsimp simp: thread_actions_isolatable_def - isolate_thread_actions_def - split_def select_f_selects liftM_def bind_assoc) - apply (rule monadic_rewrite_imp, rule monadic_rewrite_transverse) - apply (rule monadic_rewrite_drop_modify monadic_rewrite_bind_tail)+ - apply wp - apply (simp add: gets_bind_ign getSchedulerAction_def) - apply (rule monadic_rewrite_refl) - apply (simp add: ksPSpace_update_partial_id o_def) - done - -lemma doMachineOp_isolatable: - "thread_actions_isolatable idx (doMachineOp m)" - apply (simp add: doMachineOp_def split_def) - apply (intro thread_actions_isolatable_bind[OF _ _ hoare_pre(1)] - gets_isolatable thread_actions_isolatable_returns - modify_isolatable select_f_isolatable) - apply (simp | wp)+ - done - -lemma page_directory_at_partial_overwrite: - "\x. tcb_at' (idx x) s \ - page_directory_at' p (ksPSpace_update (partial_overwrite idx f) s) - = page_directory_at' p s" - by (simp add: page_directory_at'_def typ_at_to_obj_at_arches - obj_at_partial_overwrite_id2) - -lemma findPDForASID_isolatable: - "thread_actions_isolatable idx (findPDForASID asid)" - apply (simp add: findPDForASID_def liftE_bindE liftME_def bindE_assoc - case_option_If2 assertE_def liftE_def checkPDAt_def - stateAssert_def2 - cong: if_cong) - apply (intro thread_actions_isolatable_bind[OF _ _ hoare_pre(1)] - thread_actions_isolatable_bindE[OF _ _ hoare_pre(1)] - thread_actions_isolatable_if thread_actions_isolatable_returns - thread_actions_isolatable_fail - gets_isolatable getObject_isolatable) - apply (simp add: projectKO_opt_asidpool page_directory_at_partial_overwrite - | wp getASID_wp)+ - done - -lemma getHWASID_isolatable: - "thread_actions_isolatable idx (getHWASID asid)" - apply (simp add: getHWASID_def loadHWASID_def - findFreeHWASID_def - case_option_If2 findPDForASIDAssert_def - checkPDAt_def checkPDUniqueToASID_def - checkPDASIDMapMembership_def - stateAssert_def2 const_def assert_def - findFreeHWASID_def - invalidateASID_def - invalidateHWASIDEntry_def - storeHWASID_def - cong: if_cong) - apply (intro thread_actions_isolatable_bind[OF _ _ hoare_pre(1)] - thread_actions_isolatable_bindE[OF _ _ hoare_pre(1)] - thread_actions_isolatable_catch[OF _ _ hoare_pre(1)] - thread_actions_isolatable_if thread_actions_isolatable_returns - thread_actions_isolatable_fail - gets_isolatable modify_isolatable - findPDForASID_isolatable doMachineOp_isolatable) - apply (wp hoare_drop_imps - | simp add: page_directory_at_partial_overwrite)+ - done - -lemma setVMRoot_isolatable: - "thread_actions_isolatable idx (setVMRoot t)" - apply (simp add: setVMRoot_def getThreadVSpaceRoot_def - locateSlot_conv getSlotCap_def - cap_case_isPageDirectoryCap if_bool_simps - whenE_def liftE_def - checkPDNotInASIDMap_def stateAssert_def2 - checkPDASIDMapMembership_def armv_contextSwitch_def - cong: if_cong) - apply (intro thread_actions_isolatable_bind[OF _ _ hoare_pre(1)] - thread_actions_isolatable_bindE[OF _ _ hoare_pre(1)] - thread_actions_isolatable_catch[OF _ _ hoare_pre(1)] - thread_actions_isolatable_if thread_actions_isolatable_returns - thread_actions_isolatable_fail - gets_isolatable getCTE_isolatable getHWASID_isolatable - findPDForASID_isolatable doMachineOp_isolatable) - apply (simp add: projectKO_opt_asidpool - | wp getASID_wp typ_at_lifts [OF getHWASID_typ_at'])+ - done - -lemma transferCaps_simple: - "transferCaps mi [] ep receiver rcvrBuf = - do - getReceiveSlots receiver rcvrBuf; - return (mi\msgExtraCaps := 0, msgCapsUnwrapped := 0\) - od" - apply (cases mi) - apply (clarsimp simp: transferCaps_def getThreadCSpaceRoot_def locateSlot_conv) - apply (rule ext bind_apply_cong[OF refl])+ - apply (simp add: upto_enum_def - split: option.split) - done - -lemma transferCaps_simple_rewrite: - "monadic_rewrite True True ((\_. caps = []) and \) - (transferCaps mi caps ep r rBuf) - (return (mi \ msgExtraCaps := 0, msgCapsUnwrapped := 0 \))" - apply (rule monadic_rewrite_gen_asm) - apply (rule monadic_rewrite_imp) - apply (rule monadic_rewrite_trans) - apply (simp add: transferCaps_simple, rule monadic_rewrite_refl) - apply (rule monadic_rewrite_symb_exec2, wp empty_fail_getReceiveSlots) - apply (rule monadic_rewrite_refl) - apply simp - done - -lemma lookupExtraCaps_simple_rewrite: - "msgExtraCaps mi = 0 \ - (lookupExtraCaps thread rcvBuf mi = returnOk [])" - by (cases mi, simp add: lookupExtraCaps_def getExtraCPtrs_def - liftE_bindE upto_enum_step_def mapM_Nil - split: option.split) - -lemma lookupIPC_inv: "\P\ lookupIPCBuffer f t \\rv. P\" - by wp - -lemma doIPCTransfer_simple_rewrite: - "monadic_rewrite True True - ((\_. msgExtraCaps (messageInfoFromWord msgInfo) = 0 - \ msgLength (messageInfoFromWord msgInfo) - \ of_nat (length ARM_H.msgRegisters)) - and obj_at' (\tcb. tcbFault tcb = None - \ tcbContext tcb msgInfoRegister = msgInfo) sender) - (doIPCTransfer sender ep badge True rcvr) - (do rv \ mapM_x (\r. do v \ asUser sender (getRegister r); - asUser rcvr (setRegister r v) - od) - (take (unat (msgLength (messageInfoFromWord msgInfo))) ARM_H.msgRegisters); - y \ setMessageInfo rcvr ((messageInfoFromWord msgInfo) \msgCapsUnwrapped := 0\); - asUser rcvr (setRegister ARM_H.badgeRegister badge) - od)" - apply (rule monadic_rewrite_gen_asm) - apply (simp add: doIPCTransfer_def bind_assoc doNormalTransfer_def - getMessageInfo_def - cong: option.case_cong) - apply (rule monadic_rewrite_imp) - apply (rule monadic_rewrite_trans) - apply (rule monadic_rewrite_bind_tail)+ - apply (rule_tac P="fault = None" in monadic_rewrite_gen_asm, simp) - apply (rule monadic_rewrite_bind_tail) - apply (rule_tac x=msgInfo in monadic_rewrite_symb_exec, - wp empty_fail_user_getreg user_getreg_rv) - apply (simp add: lookupExtraCaps_simple_rewrite returnOk_catch_bind) - apply (rule monadic_rewrite_bind) - apply (rule monadic_rewrite_from_simple, rule copyMRs_simple) - apply (rule monadic_rewrite_bind_head) - apply (rule transferCaps_simple_rewrite) - apply (wp threadGet_const) - apply (simp add: bind_assoc) - apply (rule monadic_rewrite_symb_exec2[OF lookupIPC_inv empty_fail_lookupIPCBuffer] - monadic_rewrite_symb_exec2[OF threadGet_inv empty_fail_threadGet] - monadic_rewrite_symb_exec2[OF user_getreg_inv' empty_fail_user_getreg] - monadic_rewrite_bind_head monadic_rewrite_bind_tail - | wp)+ - apply (case_tac "messageInfoFromWord msgInfo") - apply simp - apply (rule monadic_rewrite_refl) - apply wp - apply clarsimp - apply (auto elim!: obj_at'_weakenE) - done - -lemma monadic_rewrite_setSchedulerAction_noop: - "monadic_rewrite F E (\s. ksSchedulerAction s = act) (setSchedulerAction act) (return ())" - unfolding setSchedulerAction_def - apply (rule monadic_rewrite_imp, rule monadic_rewrite_modify_noop) - apply simp - done - -lemma rescheduleRequired_simple_rewrite: - "monadic_rewrite F E - (sch_act_simple) - rescheduleRequired - (setSchedulerAction ChooseNewThread)" - apply (simp add: rescheduleRequired_def getSchedulerAction_def) - apply (simp add: monadic_rewrite_def exec_gets sch_act_simple_def) - apply auto - done - -lemma setThreadState_blocked_rewrite: - "\ runnable' st \ - monadic_rewrite True True - (\s. ksCurThread s = t \ ksSchedulerAction s \ ResumeCurrentThread \ tcb_at' t s) - (setThreadState st t) - (threadSet (tcbState_update (\_. st)) t)" - apply (simp add: setThreadState_def) - apply (rule monadic_rewrite_imp) - apply (rule monadic_rewrite_trans) - apply (rule monadic_rewrite_bind_tail) - apply (rule monadic_rewrite_trans) - apply (rule monadic_rewrite_bind_tail)+ - apply (rule_tac P="\ runnable \ curThread = t - \ (action \ ResumeCurrentThread)" - in monadic_rewrite_gen_asm) - apply (simp add: when_def) - apply (rule monadic_rewrite_refl) - apply wp - apply (rule monadic_rewrite_symb_exec2, - (wp empty_fail_isRunnable - | (simp only: getCurThread_def getSchedulerAction_def - , rule empty_fail_gets))+)+ - apply (rule monadic_rewrite_refl) - apply (simp add: conj_comms, wp) - apply (rule_tac Q="\rv s. obj_at' (Not o runnable' o tcbState) t s" - in hoare_post_imp) - apply (clarsimp simp: obj_at'_def sch_act_simple_def st_tcb_at'_def) - apply (wp) - apply simp - apply (rule monadic_rewrite_refl) - apply clarsimp - done - -lemma setupCallerCap_rewrite: - "monadic_rewrite True True (\s. reply_masters_rvk_fb (ctes_of s)) - (setupCallerCap send rcv) - (do setThreadState BlockedOnReply send; - replySlot \ getThreadReplySlot send; - callerSlot \ getThreadCallerSlot rcv; - replySlotCTE \ getCTE replySlot; - assert (mdbNext (cteMDBNode replySlotCTE) = 0 - \ isReplyCap (cteCap replySlotCTE) - \ capReplyMaster (cteCap replySlotCTE) - \ mdbFirstBadged (cteMDBNode replySlotCTE) - \ mdbRevocable (cteMDBNode replySlotCTE)); - cteInsert (ReplyCap send False) replySlot callerSlot - od)" - apply (simp add: setupCallerCap_def getThreadCallerSlot_def - getThreadReplySlot_def locateSlot_conv - getSlotCap_def) - apply (rule monadic_rewrite_imp) - apply (rule monadic_rewrite_bind_tail)+ - apply (rule monadic_rewrite_assert)+ - apply (rule_tac P="mdbFirstBadged (cteMDBNode masterCTE) - \ mdbRevocable (cteMDBNode masterCTE)" - in monadic_rewrite_gen_asm) - apply simp - apply (rule monadic_rewrite_trans) - apply (rule monadic_rewrite_bind_tail) - apply (rule monadic_rewrite_symb_exec2, (wp | simp)+)+ - apply (rule monadic_rewrite_refl) - apply wp - apply (rule monadic_rewrite_symb_exec2, wp empty_fail_getCTE)+ - apply (rule monadic_rewrite_refl) - apply (wp getCTE_wp' | simp add: cte_wp_at_ctes_of)+ - apply (clarsimp simp: reply_masters_rvk_fb_def) - apply fastforce - done - -lemma attemptSwitchTo_rewrite: - "monadic_rewrite True True - (\s. obj_at' (\tcb. tcbPriority tcb = curPrio) thread s - \ obj_at' (\tcb. tcbPriority tcb = destPrio \ tcbDomain tcb = destDom) t s - \ destPrio \ curPrio - \ ksSchedulerAction s = ResumeCurrentThread - \ ksCurThread s = thread - \ ksCurDomain s = curDom - \ destDom = curDom) - (attemptSwitchTo t) (setSchedulerAction (SwitchToThread t))" - apply (simp add: attemptSwitchTo_def possibleSwitchTo_def) - apply (rule monadic_rewrite_imp) - apply (rule monadic_rewrite_trans) - apply (rule monadic_rewrite_bind_tail) - apply (rule monadic_rewrite_bind_tail) - apply (rule monadic_rewrite_bind_tail) - apply (rule monadic_rewrite_bind_tail) - apply (rule monadic_rewrite_bind_tail) - apply (rule monadic_rewrite_bind_tail) - apply (rule_tac P="curPrio \ targetPrio \ action = ResumeCurrentThread - \ targetDom = curDom" - in monadic_rewrite_gen_asm) - apply (simp add: eq_commute le_less[symmetric]) - apply (rule monadic_rewrite_refl) - apply (wp threadGet_wp) - apply (rule monadic_rewrite_trans) - apply (rule monadic_rewrite_bind_tail) - apply (rule monadic_rewrite_symb_exec2, - wp empty_fail_threadGet | simp add: getSchedulerAction_def curDomain_def)+ - apply (rule monadic_rewrite_refl) - apply wp - apply (rule monadic_rewrite_symb_exec2, simp_all add: getCurThread_def) - apply (rule monadic_rewrite_refl) - apply (auto simp: obj_at'_def) - done - -lemma oblivious_getObject_ksPSpace_default: - "\ \s. ksPSpace (f s) = ksPSpace s; - \a b c ko. (loadObject a b c ko :: 'a kernel) \ loadObject_default a b c ko \ \ - oblivious f (getObject p :: ('a :: pspace_storable) kernel)" - apply (simp add: getObject_def split_def loadObject_default_def - projectKO_def2 alignCheck_assert magnitudeCheck_assert) - apply (intro oblivious_bind, simp_all) - done - -lemmas oblivious_getObject_ksPSpace_tcb[simp] - = oblivious_getObject_ksPSpace_default[OF _ loadObject_tcb] - -lemma oblivious_setObject_ksPSpace_tcb[simp]: - "\ \s. ksPSpace (f s) = ksPSpace s; - \s g. ksPSpace_update g (f s) = f (ksPSpace_update g s) \ \ - oblivious f (setObject p (v :: tcb))" - apply (simp add: setObject_def split_def updateObject_default_def - projectKO_def2 alignCheck_assert magnitudeCheck_assert) - apply (intro oblivious_bind, simp_all) - done - -lemma oblivious_getObject_ksPSpace_cte[simp]: - "\ \s. ksPSpace (f s) = ksPSpace s \ \ - oblivious f (getObject p :: cte kernel)" - apply (simp add: getObject_def split_def loadObject_cte - projectKO_def2 alignCheck_assert magnitudeCheck_assert - typeError_def unless_when - cong: Structures_H.kernel_object.case_cong) - apply (intro oblivious_bind, - simp_all split: Structures_H.kernel_object.split split_if) - by (safe intro!: oblivious_bind, simp_all) - -lemma oblivious_doMachineOp[simp]: - "\ \s. ksMachineState (f s) = ksMachineState s; - \g s. ksMachineState_update g (f s) = f (ksMachineState_update g s) \ - \ oblivious f (doMachineOp oper)" - apply (simp add: doMachineOp_def split_def) - apply (intro oblivious_bind, simp_all) - done - -lemmas oblivious_getObject_ksPSpace_asidpool[simp] - = oblivious_getObject_ksPSpace_default[OF _ loadObject_asidpool] - -lemma oblivious_setVMRoot_schact: - "oblivious (ksSchedulerAction_update f) (setVMRoot t)" - apply (simp add: setVMRoot_def getThreadVSpaceRoot_def locateSlot_conv - getSlotCap_def getCTE_def armv_contextSwitch_def) - by (safe intro!: oblivious_bind oblivious_bindE oblivious_catch - | simp_all add: liftE_def getHWASID_def - findPDForASID_def liftME_def loadHWASID_def - findPDForASIDAssert_def checkPDAt_def - checkPDUniqueToASID_def - checkPDASIDMapMembership_def - findFreeHWASID_def invalidateASID_def - invalidateHWASIDEntry_def storeHWASID_def - checkPDNotInASIDMap_def armv_contextSwitch_def - split: capability.split arch_capability.split option.split)+ - -lemma oblivious_switchToThread_schact: - "oblivious (ksSchedulerAction_update f) (ThreadDecls_H.switchToThread t)" - apply (simp add: switchToThread_def ARM_H.switchToThread_def bind_assoc - getCurThread_def setCurThread_def threadGet_def liftM_def - threadSet_def tcbSchedEnqueue_def unless_when asUser_def - getQueue_def setQueue_def storeWordUser_def setRegister_def - pointerInUserData_def isRunnable_def isBlocked_def - getThreadState_def tcbSchedDequeue_def bitmap_fun_defs) - by (safe intro!: oblivious_bind - | simp_all add: oblivious_setVMRoot_schact)+ - -lemma schedule_rewrite: - notes hoare_TrueI[simp] - shows "monadic_rewrite True True - (\s. ksSchedulerAction s = SwitchToThread t \ ct_in_state' (op = Running) s) - (schedule) - (do curThread \ getCurThread; tcbSchedEnqueue curThread; setSchedulerAction ResumeCurrentThread; switchToThread t od)" - apply (simp add: schedule_def) - apply (rule monadic_rewrite_imp) - apply (rule monadic_rewrite_trans) - apply (rule monadic_rewrite_bind_tail) - apply (rule monadic_rewrite_bind_tail) - apply (rule_tac P="action = SwitchToThread t" in monadic_rewrite_gen_asm, simp) - apply (rule monadic_rewrite_bind_tail) - apply (rule_tac P="curRunnable \ action = SwitchToThread t" in monadic_rewrite_gen_asm, simp) - apply (rule monadic_rewrite_refl) - apply (wp,simp,wp) - apply (rule monadic_rewrite_trans) - apply (rule monadic_rewrite_bind_tail) - apply (rule monadic_rewrite_symb_exec2, wp | simp add: isRunnable_def getSchedulerAction_def)+ - apply (rule monadic_rewrite_refl) - apply (wp) - apply (simp add: setSchedulerAction_def) - apply (subst oblivious_modify_swap[symmetric], rule oblivious_switchToThread_schact) - apply (rule monadic_rewrite_refl) - apply (clarsimp simp: st_tcb_at'_def pred_neg_def o_def obj_at'_def ct_in_state'_def) - done - -lemma schedule_rewrite_ct_not_runnable': - "monadic_rewrite True True - (\s. ksSchedulerAction s = SwitchToThread t \ ct_in_state' (Not \ runnable') s) - (schedule) - (do setSchedulerAction ResumeCurrentThread; switchToThread t od)" - apply (simp add: schedule_def) - apply (rule monadic_rewrite_imp) - apply (rule monadic_rewrite_trans) - apply (rule monadic_rewrite_bind_tail) - apply (rule monadic_rewrite_bind_tail) - apply (rule_tac P="action = SwitchToThread t" in monadic_rewrite_gen_asm, simp) - apply (rule monadic_rewrite_bind_tail) - apply (rule_tac P="\ curRunnable \ action = SwitchToThread t" in monadic_rewrite_gen_asm,simp) - apply (rule monadic_rewrite_refl) - apply (wp,simp,wp) - apply (rule monadic_rewrite_trans) - apply (rule monadic_rewrite_bind_tail) - apply (rule monadic_rewrite_symb_exec2, wp | - simp add: isRunnable_def getSchedulerAction_def | - rule hoare_TrueI)+ - apply (rule monadic_rewrite_refl) - apply (wp) - apply (simp add: setSchedulerAction_def) - apply (subst oblivious_modify_swap[symmetric], rule oblivious_switchToThread_schact) - apply (rule monadic_rewrite_symb_exec2) - apply (wp, simp, rule hoare_TrueI) - apply (rule monadic_rewrite_refl) - apply (clarsimp simp: st_tcb_at'_def pred_neg_def o_def obj_at'_def ct_in_state'_def) - done - -lemma activateThread_simple_rewrite: - "monadic_rewrite True True (ct_in_state' (op = Running)) - (activateThread) (return ())" - apply (simp add: activateThread_def) - apply (rule monadic_rewrite_imp) - apply (rule monadic_rewrite_trans, rule monadic_rewrite_bind_tail)+ - apply (rule_tac P="state = Running" in monadic_rewrite_gen_asm) - apply simp - apply (rule monadic_rewrite_refl) - apply wp - apply (rule monadic_rewrite_symb_exec2, wp empty_fail_getThreadState) - apply (rule monadic_rewrite_refl) - apply wp - apply (rule monadic_rewrite_symb_exec2, - simp_all add: getCurThread_def) - apply (rule monadic_rewrite_refl) - apply (clarsimp simp: ct_in_state'_def elim!: pred_tcb'_weakenE) - done - -end - -lemma setCTE_obj_at_prio[wp]: - "\obj_at' (\tcb. P (tcbPriority tcb)) t\ setCTE p v \\rv. obj_at' (\tcb. P (tcbPriority tcb)) t\" - unfolding setCTE_def - by (rule setObject_cte_obj_at_tcb', simp+) - -crunch obj_at_prio[wp]: cteInsert "obj_at' (\tcb. P (tcbPriority tcb)) t" - (wp: crunch_wps) - -crunch ctes_of[wp]: asUser "\s. P (ctes_of s)" - (wp: crunch_wps) - -lemma tcbSchedEnqueue_tcbPriority[wp]: - "\obj_at' (\tcb. P (tcbPriority tcb)) t\ - tcbSchedEnqueue t' - \\rv. obj_at' (\tcb. P (tcbPriority tcb)) t\" - apply (simp add: tcbSchedEnqueue_def unless_def) - apply (wp | simp cong: if_cong)+ - done - -crunch obj_at_prio[wp]: cteDeleteOne "obj_at' (\tcb. P (tcbPriority tcb)) t" - (wp: crunch_wps setEndpoint_obj_at_tcb' - setThreadState_obj_at_unchanged setNotification_tcb setBoundNotification_obj_at_unchanged - simp: crunch_simps unless_def) - -context kernel_m begin - -lemma setThreadState_no_sch_change: - "\\s. P (ksSchedulerAction s) \ (runnable' st \ t \ ksCurThread s)\ - setThreadState st t - \\rv s. P (ksSchedulerAction s)\" - (is "NonDetMonad.valid ?P ?f ?Q") - apply (simp add: setThreadState_def setSchedulerAction_def) - apply (wp hoare_pre_cont[where a=rescheduleRequired]) - apply (rule_tac Q="\_. ?P and st_tcb_at' (op = st) t" in hoare_post_imp) - apply (clarsimp split: split_if) - apply (clarsimp simp: obj_at'_def st_tcb_at'_def projectKOs) - apply (rule hoare_pre, wp threadSet_pred_tcb_at_state) - apply simp - done - -lemma asUser_obj_at_unchangedT: - assumes x: "\tcb con con'. con' \ fst (m con) - \ P (tcbContext_update (\_. snd con') tcb) = P tcb" shows - "\obj_at' P t\ asUser t' m \\rv. obj_at' P t\" - apply (simp add: asUser_def split_def) - apply (wp threadSet_obj_at' threadGet_wp) - apply (clarsimp simp: obj_at'_def projectKOs x cong: if_cong) - done - -lemmas asUser_obj_at_unchanged - = asUser_obj_at_unchangedT[OF all_tcbI, rule_format] - -lemma bind_assoc: - "do y \ do x \ m; f x od; g y od - = do x \ m; y \ f x; g y od" - by (rule bind_assoc) - -lemma setObject_modify_assert: - "\ updateObject v = updateObject_default v \ - \ setObject p v = do f \ gets (obj_at' (\v'. v = v' \ True) p); - assert f; modify (ksPSpace_update (\ps. ps(p \ injectKO v))) od" - using objBits_2n[where obj=v] - apply (simp add: setObject_def split_def updateObject_default_def - bind_assoc projectKO_def2 alignCheck_assert) - apply (rule ext, simp add: exec_gets) - apply (case_tac "obj_at' (\v'. v = v' \ True) p x") - apply (clarsimp simp: obj_at'_def projectKOs lookupAround2_known1 - assert_opt_def) - apply (clarsimp simp: project_inject) - apply (simp only: objBits_def objBitsT_koTypeOf[symmetric] koTypeOf_injectKO) - apply (simp add: magnitudeCheck_assert2 simpler_modify_def) - apply (clarsimp simp: assert_opt_def assert_def magnitudeCheck_assert2 - split: option.split split_if) - apply (clarsimp simp: obj_at'_def projectKOs) - apply (clarsimp simp: project_inject) - apply (simp only: objBits_def objBitsT_koTypeOf[symmetric] - koTypeOf_injectKO simp_thms) - done - -lemma setEndpoint_isolatable: - "thread_actions_isolatable idx (setEndpoint p e)" - apply (simp add: setEndpoint_def setObject_modify_assert - assert_def) - apply (case_tac "p \ range idx") - apply (clarsimp simp: thread_actions_isolatable_def - monadic_rewrite_def fun_eq_iff - liftM_def isolate_thread_actions_def - bind_assoc exec_gets getSchedulerAction_def - bind_select_f_bind[symmetric]) - apply (simp add: obj_at_partial_overwrite_id2) - apply (drule_tac x=x in spec) - apply (clarsimp simp: obj_at'_def projectKOs select_f_asserts) - apply (intro thread_actions_isolatable_bind[OF _ _ hoare_pre(1)] - thread_actions_isolatable_if - thread_actions_isolatable_return - thread_actions_isolatable_fail) - apply (rule gets_isolatable) - apply (simp add: obj_at_partial_overwrite_id2) - apply (rule modify_isolatable) - apply (clarsimp simp: o_def partial_overwrite_def) - apply (rule kernel_state.fold_congs[OF refl refl]) - apply (clarsimp simp: fun_eq_iff - split: split_if) - apply (wp | simp)+ - done - -lemma setCTE_assert_modify: - "setCTE p v = do c \ gets (real_cte_at' p); - t \ gets (tcb_at' (p && ~~ mask 9) - and K ((p && mask 9) \ dom tcb_cte_cases)); - if c then modify (ksPSpace_update (\ps. ps(p \ KOCTE v))) - else if t then - modify (ksPSpace_update - (\ps. ps (p && ~~ mask 9 \ - KOTCB (snd (the (tcb_cte_cases (p && mask 9))) (K v) - (the (projectKO_opt (the (ps (p && ~~ mask 9))))))))) - else fail od" - apply (clarsimp simp: setCTE_def setObject_def split_def - fun_eq_iff exec_gets) - apply (case_tac "real_cte_at' p x") - apply (clarsimp simp: obj_at'_def projectKOs lookupAround2_known1 - assert_opt_def alignCheck_assert objBits_simps - magnitudeCheck_assert2 updateObject_cte) - apply (simp add: simpler_modify_def) - apply (simp split: split_if, intro conjI impI) - apply (clarsimp simp: obj_at'_def projectKOs) - apply (subgoal_tac "p \ (p && ~~ mask 9) + 2 ^ 9 - 1") - apply (subgoal_tac "fst (lookupAround2 p (ksPSpace x)) - = Some (p && ~~ mask 9, KOTCB obj)") - apply (simp add: assert_opt_def) - apply (subst updateObject_cte_tcb) - apply (fastforce simp add: subtract_mask) - apply (simp add: assert_opt_def alignCheck_assert bind_assoc - magnitudeCheck_assert - is_aligned_neg_mask2 objBits_def) - apply (rule ps_clear_lookupAround2, assumption+) - apply (rule word_and_le2) - apply (simp add: objBits_simps mask_def field_simps) - apply (simp add: simpler_modify_def cong: option.case_cong if_cong) - apply (rule kernel_state.fold_congs[OF refl refl]) - apply (clarsimp simp: projectKO_opt_tcb cong: if_cong) - apply (clarsimp simp: lookupAround2_char1 word_and_le2) - apply (rule ccontr, clarsimp) - apply (erule(2) ps_clearD) - apply (simp add: objBits_simps mask_def field_simps) - apply (rule tcb_cte_cases_in_range2) - apply (simp add: subtract_mask) - apply simp - apply (clarsimp simp: assert_opt_def split: option.split) - apply (rule trans [OF bind_apply_cong[OF _ refl] fun_cong[OF fail_bind]]) - apply (simp add: fail_def prod_eq_iff) - apply (rule context_conjI) - apply (rule ccontr, clarsimp elim!: nonemptyE) - apply (frule(1) updateObject_cte_is_tcb_or_cte[OF _ refl]) - apply (erule disjE) - apply clarsimp - apply (frule(1) tcb_cte_cases_aligned_helpers) - apply (clarsimp simp: domI[where m = cte_cte_cases] field_simps) - apply (clarsimp simp: lookupAround2_char1 obj_at'_def projectKOs - objBits_simps) - apply (clarsimp simp: obj_at'_def lookupAround2_char1 - objBits_simps projectKOs cte_level_bits_def) - apply (erule empty_failD[OF empty_fail_updateObject_cte]) - done - -lemma partial_overwrite_fun_upd2: - "partial_overwrite idx tsrs (f (x := y)) - = (partial_overwrite idx tsrs f) - (x := if x \ range idx then put_tcb_state_regs (tsrs (inv idx x)) y - else y)" - by (simp add: fun_eq_iff partial_overwrite_def split: split_if) - -lemma setCTE_isolatable: - "thread_actions_isolatable idx (setCTE p v)" - apply (simp add: setCTE_assert_modify) - apply (clarsimp simp: thread_actions_isolatable_def - monadic_rewrite_def fun_eq_iff - liftM_def exec_gets - isolate_thread_actions_def - bind_assoc exec_gets getSchedulerAction_def - bind_select_f_bind[symmetric] - obj_at_partial_overwrite_If - obj_at_partial_overwrite_id2 - cong: if_cong) - apply (case_tac "p && ~~ mask 9 \ range idx \ p && mask 9 \ dom tcb_cte_cases") - apply clarsimp - apply (frule_tac x=x in spec, erule obj_atE') - apply (subgoal_tac "\ real_cte_at' p s") - apply (clarsimp simp: select_f_returns select_f_asserts split: split_if) - apply (clarsimp simp: o_def simpler_modify_def partial_overwrite_fun_upd2) - apply (rule kernel_state.fold_congs[OF refl refl]) - apply (rule ext) - apply (clarsimp simp: partial_overwrite_get_tcb_state_regs - split: split_if) - apply (clarsimp simp: projectKOs get_tcb_state_regs_def - put_tcb_state_regs_def put_tcb_state_regs_tcb_def - partial_overwrite_def - split: tcb_state_regs.split) - apply (case_tac obj, simp add: projectKO_opt_tcb) - apply (simp add: tcb_cte_cases_def split: split_if_asm) - apply (drule_tac x=x in spec) - apply (clarsimp simp: obj_at'_def projectKOs objBits_simps subtract_mask(2) [symmetric]) - apply (erule notE[rotated], erule (3) tcb_ctes_clear[rotated]) - apply (simp add: select_f_returns select_f_asserts split: split_if) - apply (intro conjI impI) - apply (clarsimp simp: simpler_modify_def fun_eq_iff - partial_overwrite_fun_upd2 o_def - intro!: kernel_state.fold_congs[OF refl refl]) - apply (clarsimp simp: obj_at'_def projectKOs objBits_simps) - apply (erule notE[rotated], rule tcb_ctes_clear[rotated 2], assumption+) - apply (fastforce simp add: subtract_mask) - apply simp - apply (clarsimp simp: simpler_modify_def - partial_overwrite_fun_upd2 o_def - partial_overwrite_get_tcb_state_regs - intro!: kernel_state.fold_congs[OF refl refl] - split: split_if) - apply (simp add: partial_overwrite_def) - apply (subgoal_tac "p \ range idx") - apply (clarsimp simp: simpler_modify_def - partial_overwrite_fun_upd2 o_def - partial_overwrite_get_tcb_state_regs - intro!: kernel_state.fold_congs[OF refl refl]) - apply clarsimp - apply (drule_tac x=x in spec) - apply (clarsimp simp: obj_at'_def projectKOs) - done - -lemma assert_isolatable: - "thread_actions_isolatable idx (assert P)" - by (simp add: assert_def thread_actions_isolatable_if - thread_actions_isolatable_returns - thread_actions_isolatable_fail) - -lemma cteInsert_isolatable: - "thread_actions_isolatable idx (cteInsert cap src dest)" - apply (simp add: cteInsert_def updateCap_def updateMDB_def - Let_def setUntypedCapAsFull_def) - apply (intro thread_actions_isolatable_bind[OF _ _ hoare_pre(1)] - thread_actions_isolatable_if - thread_actions_isolatable_returns assert_isolatable - getCTE_isolatable setCTE_isolatable) - apply (wp | simp)+ - done - -lemma isolate_thread_actions_threadSet_tcbState: - "\ inj idx; idx t' = t \ \ - monadic_rewrite False True (\s. \x. tcb_at' (idx x) s) - (threadSet (tcbState_update (\_. st)) t) - (isolate_thread_actions idx (return ()) - (\tsrs. (tsrs (t' := TCBStateRegs st (tsrContext (tsrs t'))))) - id)" - apply (simp add: isolate_thread_actions_def bind_assoc split_def - select_f_returns getSchedulerAction_def) - apply (clarsimp simp: monadic_rewrite_def exec_gets threadSet_def - getObject_get_assert bind_assoc liftM_def - setObject_modify_assert) - apply (frule_tac x=t' in spec, drule obj_at_ko_at') - apply (clarsimp simp: exec_gets simpler_modify_def o_def - intro!: kernel_state.fold_congs[OF refl refl]) - apply (simp add: partial_overwrite_fun_upd - partial_overwrite_get_tcb_state_regs) - apply (clarsimp simp: put_tcb_state_regs_def put_tcb_state_regs_tcb_def - projectKOs get_tcb_state_regs_def - elim!: obj_atE') - apply (case_tac ko) - apply (simp add: projectKO_opt_tcb) - done - -lemma thread_actions_isolatableD: - "\ thread_actions_isolatable idx f; inj idx \ - \ monadic_rewrite False True (\s. (\x. tcb_at' (idx x) s)) - f (isolate_thread_actions idx f id id)" - by (clarsimp simp: thread_actions_isolatable_def) - -lemma tcbSchedDequeue_rewrite: - "monadic_rewrite True True (obj_at' (Not \ tcbQueued) t) (tcbSchedDequeue t) (return ())" - apply (simp add: tcbSchedDequeue_def) - apply (rule monadic_rewrite_imp) - apply (rule monadic_rewrite_trans) - apply (rule monadic_rewrite_bind_tail) - apply (rule_tac P="\ queued" in monadic_rewrite_gen_asm) - apply (simp add: when_def) - apply (rule monadic_rewrite_refl) - apply (wp threadGet_const) - apply (rule monadic_rewrite_symb_exec2) - apply wp - apply (rule monadic_rewrite_refl) - apply (clarsimp) - done - -lemma switchToThread_rewrite: - "monadic_rewrite True True - (ct_in_state' (Not \ runnable') and cur_tcb' and obj_at' (Not \ tcbQueued) t) - (switchToThread t) - (do Arch.switchToThread t; setCurThread t od)" - apply (simp add: switchToThread_def) - apply (rule monadic_rewrite_imp) - apply (rule monadic_rewrite_trans) - apply (rule monadic_rewrite_bind_tail) - apply (rule monadic_rewrite_bind) - apply (rule tcbSchedDequeue_rewrite) - apply (rule monadic_rewrite_refl) - apply (wp Arch_switchToThread_obj_at_pre) - apply (rule monadic_rewrite_bind_tail) - apply (rule monadic_rewrite_symb_exec) - apply (wp, simp) - apply (rule monadic_rewrite_refl) - apply (wp) - apply (clarsimp) - done - -lemma threadGet_isolatable: - assumes v: "\tsr. \tcb. f (put_tcb_state_regs_tcb tsr tcb) = f tcb" - shows "thread_actions_isolatable idx (threadGet f t)" - apply (clarsimp simp: threadGet_def thread_actions_isolatable_def - isolate_thread_actions_def split_def - getObject_get_assert liftM_def - bind_select_f_bind[symmetric] - select_f_returns select_f_asserts bind_assoc) - apply (clarsimp simp: monadic_rewrite_def exec_gets - getSchedulerAction_def) - apply (simp add: obj_at_partial_overwrite_If) - apply (rule bind_apply_cong[OF refl]) - apply (clarsimp simp: exec_gets exec_modify o_def - ksPSpace_update_partial_id in_monad) - apply (erule obj_atE') - apply (clarsimp simp: projectKOs - partial_overwrite_def put_tcb_state_regs_def - cong: if_cong) - apply (simp add: projectKO_opt_tcb v split: split_if) - done - - -lemma monadic_rewrite_trans_dup: - "\ monadic_rewrite F E P f g; monadic_rewrite F E P g h \ - \ monadic_rewrite F E P f h" - by (auto simp add: monadic_rewrite_def) - - -lemma setCurThread_isolatable: - "thread_actions_isolatable idx (setCurThread t)" - by (simp add: setCurThread_def modify_isolatable) - -end - crunch tcb2[wp]: "Arch.switchToThread" "tcb_at' t" (ignore: ARM.clearExMonitor) context kernel_m begin -lemma isolate_thread_actions_tcbs_at: - assumes f: "\x. \tcb_at' (idx x)\ f \\rv. tcb_at' (idx x)\" shows - "\\s. \x. tcb_at' (idx x) s\ - isolate_thread_actions idx f f' f'' \\p s. \x. tcb_at' (idx x) s\" - apply (simp add: isolate_thread_actions_def split_def) - apply wp - apply clarsimp - apply (simp add: obj_at_partial_overwrite_If use_valid[OF _ f]) - done - -lemma isolate_thread_actions_rewrite_bind: - "\ inj idx; monadic_rewrite False True (\s. \x. tcb_at' (idx x) s) - f (isolate_thread_actions idx f' f'' f'''); - \x. monadic_rewrite False True (\s. \x. tcb_at' (idx x) s) - (g x) - (isolate_thread_actions idx (g' x) g'' g'''); - thread_actions_isolatable idx f'; \x. thread_actions_isolatable idx (g' x); - \x. \tcb_at' (idx x)\ f' \\rv. tcb_at' (idx x)\ \ - \ monadic_rewrite False True (\s. \x. tcb_at' (idx x) s) - (f >>= g) (isolate_thread_actions idx - (f' >>= g') (g'' o f'') (g''' o f'''))" - apply (rule monadic_rewrite_imp) - apply (rule monadic_rewrite_trans) - apply (rule monadic_rewrite_bind, assumption+) - apply (wp isolate_thread_actions_tcbs_at) - apply simp - apply (subst isolate_thread_actions_wrap_bind, assumption) - apply (rule monadic_rewrite_in_isolate_thread_actions, assumption) - apply (rule monadic_rewrite_transverse) - apply (rule monadic_rewrite_bind2) - apply (erule(1) thread_actions_isolatableD) - apply (rule thread_actions_isolatableD, assumption+) - apply (rule hoare_vcg_all_lift, assumption) - apply (simp add: liftM_def id_def) - apply (rule monadic_rewrite_refl) - apply (simp add: obj_at_partial_overwrite_If) - done - -definition - "copy_register_tsrs src dest r tsrs - = tsrs (dest := TCBStateRegs (tsrState (tsrs dest)) - ((tsrContext (tsrs dest)) (r := tsrContext (tsrs src) r)))" - -lemma tcb_at_KOTCB_upd: - "tcb_at' (idx x) s \ - tcb_at' p (ksPSpace_update (\ps. ps(idx x \ KOTCB tcb)) s) - = tcb_at' p s" - apply (clarsimp simp: obj_at'_def projectKOs objBits_simps - split: split_if) - apply (simp add: ps_clear_def) - done - -definition - "set_register_tsrs dest r v tsrs - = tsrs (dest := TCBStateRegs (tsrState (tsrs dest)) - ((tsrContext (tsrs dest)) (r := v)))" - - -lemma set_register_isolate: - "\ inj idx; idx y = dest \ \ - monadic_rewrite False True - (\s. \x. tcb_at' (idx x) s) - (asUser dest (setRegister r v)) - (isolate_thread_actions idx (return ()) - (set_register_tsrs y r v) id)" - apply (simp add: asUser_def split_def bind_assoc - getRegister_def setRegister_def - select_f_returns isolate_thread_actions_def - getSchedulerAction_def) - apply (simp add: threadGet_def liftM_def getObject_get_assert - bind_assoc threadSet_def - setObject_modify_assert) - apply (clarsimp simp: monadic_rewrite_def exec_gets - exec_modify tcb_at_KOTCB_upd) - apply (clarsimp simp: simpler_modify_def - intro!: kernel_state.fold_congs[OF refl refl]) - apply (clarsimp simp: set_register_tsrs_def o_def - partial_overwrite_fun_upd - partial_overwrite_get_tcb_state_regs) - apply (drule_tac x=y in spec) - apply (clarsimp simp: obj_at'_def projectKOs objBits_simps - cong: if_cong) - apply (case_tac obj) - apply (simp add: projectKO_opt_tcb put_tcb_state_regs_def - put_tcb_state_regs_tcb_def get_tcb_state_regs_def - cong: if_cong) - done - -lemma copy_register_isolate: - "\ inj idx; idx x = src; idx y = dest \ \ - monadic_rewrite False True - (\s. \x. tcb_at' (idx x) s) - (do v \ asUser src (getRegister r); - asUser dest (setRegister r v) od) - (isolate_thread_actions idx (return ()) - (copy_register_tsrs x y r) id)" - apply (simp add: asUser_def split_def bind_assoc - getRegister_def setRegister_def - select_f_returns isolate_thread_actions_def - getSchedulerAction_def) - apply (simp add: threadGet_def liftM_def getObject_get_assert - bind_assoc threadSet_def - setObject_modify_assert) - apply (clarsimp simp: monadic_rewrite_def exec_gets - exec_modify tcb_at_KOTCB_upd) - apply (clarsimp simp: simpler_modify_def - intro!: kernel_state.fold_congs[OF refl refl]) - apply (clarsimp simp: copy_register_tsrs_def o_def - partial_overwrite_fun_upd - partial_overwrite_get_tcb_state_regs) - apply (frule_tac x=x in spec, drule_tac x=y in spec) - apply (clarsimp simp: obj_at'_def projectKOs objBits_simps - cong: if_cong) - apply (case_tac obj, case_tac obja) - apply (simp add: projectKO_opt_tcb put_tcb_state_regs_def - put_tcb_state_regs_tcb_def get_tcb_state_regs_def - cong: if_cong) - apply (auto simp: fun_eq_iff split: split_if) - done - -lemmas monadic_rewrite_bind_alt - = monadic_rewrite_trans[OF monadic_rewrite_bind_tail monadic_rewrite_bind_head, rotated -1] - - -lemma monadic_rewrite_isolate_final2: - assumes mr: "monadic_rewrite F E Q f g" - and eqs: "\s tsrs. \ P s; tsrs = get_tcb_state_regs o ksPSpace s o idx \ - \ f' tsrs = g' tsrs" - "\s. P s \ f'' (ksSchedulerAction s) = g'' (ksSchedulerAction s)" - "\s tsrs sa. R s \ - Q ((ksPSpace_update (partial_overwrite idx tsrs) - s) (| ksSchedulerAction := sa |))" - shows - "monadic_rewrite F E (P and R) - (isolate_thread_actions idx f f' f'') - (isolate_thread_actions idx g g' g'')" - apply (simp add: isolate_thread_actions_def split_def) - apply (rule monadic_rewrite_imp) - apply (rule monadic_rewrite_bind_tail)+ - apply (rule_tac P="\ s'. Q s" in monadic_rewrite_bind) - apply (insert mr)[1] - apply (simp add: monadic_rewrite_def select_f_def) - apply auto[1] - apply (rule_tac P="P and (\s. tcbs = get_tcb_state_regs o ksPSpace s o idx - \ sa = ksSchedulerAction s)" - in monadic_rewrite_refl3) - apply (clarsimp simp: exec_modify eqs return_def) - apply wp - apply (clarsimp simp: o_def eqs) - done - -lemmas monadic_rewrite_isolate_final - = monadic_rewrite_isolate_final2[where R=\, OF monadic_rewrite_refl2, simplified] - -lemma copy_registers_isolate: - "\ inj idx; idx x = t; idx y = t' \ \ - monadic_rewrite False True - (\s. \x. tcb_at' (idx x) s) - (mapM_x (\r. do v \ asUser t (getRegister r); - asUser t' (setRegister r v) - od) - regs) - (isolate_thread_actions idx - (return ()) (foldr (copy_register_tsrs x y) (rev regs)) id)" - apply (induct regs) - apply (simp add: mapM_x_Nil) - apply (clarsimp simp: monadic_rewrite_def liftM_def bind_assoc - isolate_thread_actions_def - split_def exec_gets getSchedulerAction_def - select_f_returns o_def ksPSpace_update_partial_id) - apply (simp add: return_def simpler_modify_def) - apply (simp add: mapM_x_Cons) - apply (rule monadic_rewrite_imp) - apply (rule monadic_rewrite_trans) - apply (rule isolate_thread_actions_rewrite_bind, assumption) - apply (rule copy_register_isolate, assumption+) - apply (rule thread_actions_isolatable_returns)+ - apply wp - apply (rule monadic_rewrite_isolate_final[where P=\], simp+) - done - -lemma setSchedulerAction_isolate: - "inj idx \ - monadic_rewrite False True (\s. \x. tcb_at' (idx x) s) - (setSchedulerAction sa) - (isolate_thread_actions idx (return ()) id (\_. sa))" - apply (clarsimp simp: monadic_rewrite_def liftM_def bind_assoc - isolate_thread_actions_def select_f_returns - exec_gets getSchedulerAction_def o_def - ksPSpace_update_partial_id setSchedulerAction_def) - apply (simp add: simpler_modify_def) - done - -lemma updateMDB_isolatable: - "thread_actions_isolatable idx (updateMDB slot f)" - apply (simp add: updateMDB_def thread_actions_isolatable_return - split: split_if) - apply (intro impI thread_actions_isolatable_bind[OF _ _ hoare_pre(1)] - getCTE_isolatable setCTE_isolatable, - (wp | simp)+) - done - -lemma clearUntypedFreeIndex_isolatable: - "thread_actions_isolatable idx (clearUntypedFreeIndex slot)" - apply (simp add: clearUntypedFreeIndex_def getSlotCap_def) - apply (rule thread_actions_isolatable_bind) - apply (rule getCTE_isolatable) - apply (simp split: capability.split, safe intro!: thread_actions_isolatable_return) - apply (simp add: updateTrackedFreeIndex_def getSlotCap_def) - apply (intro thread_actions_isolatable_bind getCTE_isolatable - modify_isolatable) - apply (wp | simp)+ - done - -lemma emptySlot_isolatable: - "thread_actions_isolatable idx (emptySlot slot None)" - apply (simp add: emptySlot_def updateCap_def case_Null_If - cong: if_cong) - apply (intro thread_actions_isolatable_bind[OF _ _ hoare_pre(1)] - clearUntypedFreeIndex_isolatable - thread_actions_isolatable_if - getCTE_isolatable setCTE_isolatable - thread_actions_isolatable_return - updateMDB_isolatable, - (wp | simp)+) - done - -lemmas fastpath_isolatables - = setEndpoint_isolatable getCTE_isolatable - assert_isolatable cteInsert_isolatable - setCurThread_isolatable - emptySlot_isolatable updateMDB_isolatable - thread_actions_isolatable_returns - -lemmas fastpath_isolate_rewrites - = isolate_thread_actions_threadSet_tcbState isolate_thread_actions_asUser - copy_registers_isolate setSchedulerAction_isolate - fastpath_isolatables[THEN thread_actions_isolatableD] - lemma resolveAddressBits_points_somewhere: "\\s. \slot. Q slot s\ resolveAddressBits cp cptr bits \Q\,-" apply (rule_tac Q'="\rv s. \rv. Q rv s" in hoare_post_imp_R) @@ -5157,7 +3486,7 @@ lemma resolveAddressBits_points_somewhere: done lemma user_getregs_wp: - "\\s. tcb_at' t s \ (\tcb. ko_at' tcb t s \ Q (map (tcbContext tcb) regs) s)\ + "\\s. tcb_at' t s \ (\tcb. ko_at' tcb t s \ Q (map ((atcbContextGet o tcbArch) tcb) regs) s)\ asUser t (mapM getRegister regs) \Q\" apply (rule hoare_strengthen_post) apply (rule hoare_vcg_conj_lift) @@ -5169,7 +3498,7 @@ lemma user_getregs_wp: done lemma foldr_copy_register_tsrs: - "foldr (copy_register_tsrs x y) rs s + "foldr (\r . copy_register_tsrs x y r r (\x. x)) rs s = (s (y := TCBStateRegs (tsrState (s y)) (\r. if r \ set rs then tsrContext (s x) r else tsrContext (s y) r)))" @@ -5721,8 +4050,9 @@ crunch ctes_of[wp]: attemptSwitchTo "\s. P (ctes_of s)" crunch cte_wp_at'[wp]: attemptSwitchTo "cte_wp_at' P p" -crunch tcbContext[wp]: attemptSwitchTo "obj_at' (\tcb. P (tcbContext tcb)) t" - (wp: crunch_wps) + +crunch tcbContext[wp]: attemptSwitchTo "obj_at' (\tcb. P ( (atcbContextGet o tcbArch) tcb)) t" + (wp: crunch_wps simp_del: comp_apply) crunch only_cnode_caps[wp]: doFaultTransfer "\s. P (only_cnode_caps (ctes_of s))" (wp: crunch_wps simp: crunch_simps) @@ -6407,7 +4737,7 @@ lemma fastpath_callKernel_SysReplyRecv_corres: static_imp_wp hoare_vcg_all_lift hoare_vcg_imp_lift static_imp_wp cnode_caps_gsCNodes_lift hoare_vcg_ex_lift - | simp + | simp del: comp_apply | clarsimp simp: obj_at'_weakenE[OF _ TrueI])+) apply (wp getCTE_wp' gts_imp') diff --git a/proof/crefine/IpcCancel_C.thy b/proof/crefine/IpcCancel_C.thy index a9e4a894f..b43bd4204 100644 --- a/proof/crefine/IpcCancel_C.thy +++ b/proof/crefine/IpcCancel_C.thy @@ -2875,7 +2875,7 @@ lemma cancelIPC_ccorres_reply_helper: cteDeleteOne callerCap od) od) - (\ret__struct_fault_C :== CALL fault_null_fault_new();; + (\ret__struct_fault_C :== CALL seL4_Fault_NullFault_new();; Guard C_Guard {s. s \\<^sub>c tptr_' s} (Basic (\s. s \ (\t_hrs :== (hrs_mem_update (heap_update (Ptr &(\tptr\[''tcbFault_C''])) \ret__struct_fault_C)));; @@ -2905,7 +2905,7 @@ lemma cancelIPC_ccorres_reply_helper: apply (clarsimp simp: typ_heap_simps) apply (erule(2) rf_sr_tcb_update_no_queue, simp_all add: typ_heap_simps)[1] apply (rule ball_tcb_cte_casesI, simp_all)[1] - apply (clarsimp simp: ctcb_relation_def fault_lift_null_fault + apply (clarsimp simp: ctcb_relation_def seL4_Fault_lift_NullFault cfault_rel_def cthread_state_relation_def) apply (case_tac "tcbState tcb", simp_all add: is_cap_fault_def)[1] apply ctac @@ -3026,7 +3026,7 @@ lemma cancelIPC_ccorres1: apply (erule(1) rf_sr_tcb_update_no_queue2, (simp add: typ_heap_simps')+)[1] apply (rule ball_tcb_cte_casesI, simp_all)[1] - apply (clarsimp simp: ctcb_relation_def fault_lift_null_fault + apply (clarsimp simp: ctcb_relation_def seL4_Fault_lift_NullFault cfault_rel_def cthread_state_relation_def) apply (case_tac "tcbState tcb", simp_all add: is_cap_fault_def)[1] apply ceqv diff --git a/proof/crefine/Ipc_C.thy b/proof/crefine/Ipc_C.thy index b518864cb..583bf479f 100644 --- a/proof/crefine/Ipc_C.thy +++ b/proof/crefine/Ipc_C.thy @@ -9,7 +9,11 @@ *) theory Ipc_C -imports Finalise_C CSpace_All SyscallArgs_C +imports + Finalise_C + CSpace_All + SyscallArgs_C + IsolatedThreadAction begin context begin interpretation Arch . (*FIXME: arch_split*) @@ -190,7 +194,12 @@ definition | None \ undefined" definition - "thread_replace \ \upd t nv s. + "thread_fetch_option \ \ext t s. case (ksPSpace s t) of + Some (KOTCB tcb) \ ext tcb + | None \ None" + +definition + "thread_replace \ \upd t nv s. let obj = case (ksPSpace s t) of Some (KOTCB tcb) \ Some (KOTCB (upd (\_. nv) tcb)) | obj \ obj @@ -414,10 +423,11 @@ end context kernel_m begin (*FIXME: arch_split: C kernel names hidden by Haskell names *) -abbreviation "syscallMessageC \ kernel_all_substitute.syscallMessage" -lemmas syscallMessageC_def = kernel_all_substitute.syscallMessage_def -abbreviation "exceptionMessageC \ kernel_all_substitute.exceptionMessage" -lemmas exceptionMessageC_def = kernel_all_substitute.exceptionMessage_def +(*FIXME: fupdate simplification issues for 2D arrays *) +abbreviation "syscallMessageC \ kernel_all_global_addresses.fault_messages.[unat MessageID_Syscall]" +lemmas syscallMessageC_def = kernel_all_substitute.fault_messages_def +abbreviation "exceptionMessageC \ kernel_all_substitute.fault_messages.[unat MessageID_Exception]" +lemmas exceptionMessageC_def = kernel_all_substitute.fault_messages_def lemma syscallMessage_ccorres: "n < unat n_syscallMessage @@ -425,10 +435,11 @@ lemma syscallMessage_ccorres: = index syscallMessageC n" apply (simp add: ARM_H.syscallMessage_def syscallMessageC_def ARM.syscallMessage_def + MessageID_Exception_def MessageID_Syscall_def n_syscallMessage_def msgRegisters_unfold) apply (simp add: upto_enum_def fromEnum_def enum_register) apply (simp add: toEnum_def enum_register) - apply (clarsimp simp: fupdate_def + apply (clarsimp simp: fupdate_def Kernel_C.CPSR_def | drule nat_less_cases' | erule disjE)+ done @@ -436,78 +447,56 @@ end context begin interpretation Arch . (*FIXME: arch_split*) +definition + "handleArchFaultReply' f sender receiver tag \ do + label \ return $ msgLabel tag; + mlen \ return $ msgLength tag; + case f of + VMFault _ _ \ do + sendBuf \ lookupIPCBuffer False sender; + stateAssert (tcb_at' sender) []; + case sendBuf of + None \ return () + | Some bufferPtr \ do + mapM loadWordUser + (map (\i. bufferPtr + PPtr (i * 4)) + [(scast n_msgRegisters :: word32) + 1.e. msgMaxLength]); + return () + od; + return True + od + od" + definition "handleFaultReply' f sender receiver \ do tag \ getMessageInfo sender; label \ return $ msgLabel tag; mlen \ return $ msgLength tag; case f of - CapFault _ _ _ \ do - sendBuf \ lookupIPCBuffer False sender; - stateAssert (tcb_at' sender) []; - case sendBuf of - None \ return () - | Some bufferPtr \ do - mapM loadWordUser - (map (\i. bufferPtr + PPtr (i * 4)) - [(scast n_msgRegisters :: word32) + 1.e. msgMaxLength]); - return () - od; - return True - od - | VMFault _ _ \ do - sendBuf \ lookupIPCBuffer False sender; - stateAssert (tcb_at' sender) []; - case sendBuf of - None \ return () - | Some bufferPtr \ do - mapM loadWordUser - (map (\i. bufferPtr + PPtr (i * 4)) - [(scast n_msgRegisters :: word32) + 1.e. msgMaxLength]); - return () - od; - return True - od + CapFault _ _ _ \ return True + | ArchFault af \ handleArchFaultReply' af sender receiver tag | UnknownSyscallException _ \ do - sendBuf \ lookupIPCBuffer False sender; regs \ return $ take (unat mlen) syscallMessage; - stateAssert (tcb_at' sender) []; zipWithM_x (\rs rd. do v \ asUser sender $ getRegister rs; asUser receiver $ setRegister rd $ sanitiseRegister rd v od) msgRegisters regs; - stateAssert (tcb_at' receiver) []; + sendBuf \ lookupIPCBuffer False sender; case sendBuf of None \ return () - | Some bufferPtr \ do + | Some bufferPtr \ zipWithM_x (\i rd. do v \ loadWordUser (bufferPtr + PPtr (i * 4)); asUser receiver $ setRegister rd $ sanitiseRegister rd v od) [(scast n_msgRegisters :: word32) + 1.e. scast n_syscallMessage] (drop (unat (scast n_msgRegisters :: word32)) regs); - mapM loadWordUser - (map (\i. bufferPtr + PPtr (i * 4)) - [(scast n_msgRegisters :: word32) + 1 + of_nat (length (drop (unat (scast n_msgRegisters :: word32)) regs)).e. msgMaxLength]); - return () - od; return (label = 0) od | UserException _ _ \ do - sendBuf \ lookupIPCBuffer False sender; regs \ return $ take (unat mlen) exceptionMessage; - stateAssert (tcb_at' sender) []; zipWithM_x (\rs rd. do v \ asUser sender $ getRegister rs; asUser receiver $ setRegister rd $ sanitiseRegister rd v od) msgRegisters regs; - stateAssert (tcb_at' receiver) []; - case sendBuf of - None \ return () - | Some bufferPtr \ do - mapM loadWordUser - (map (\i. bufferPtr + PPtr (i * 4)) - [(scast n_msgRegisters :: word32) + 1.e. msgMaxLength]); - return () - od; return (label = 0) od od" @@ -538,153 +527,231 @@ lemmas syscallMessage_unfold unfolded fromEnum_def enum_register, simplified, unfolded toEnum_def enum_register, simplified] -lemma handleFaultReply': +lemma handleArchFaultReply': notes option.case_cong_weak [cong] wordSize_def'[simp] - assumes neq: "r \ s" - shows "do - tag \ getMessageInfo s; + shows "(do sb \ lookupIPCBuffer False s; msg \ getMRs s sb tag; - handleFaultReply f r (msgLabel tag) msg - od = handleFaultReply' f s r" - apply (rule ext) - apply (unfold handleFaultReply'_def getMRs_def msgMaxLength_def + handleArchFaultReply f r (msgLabel tag) msg + od) x' = handleArchFaultReply' f s r tag x'" + apply (unfold handleArchFaultReply'_def getMRs_def msgMaxLength_def bit_def msgLengthBits_def msgRegisters_unfold fromIntegral_simp1 fromIntegral_simp2 shiftL_word) apply (simp add: bind_assoc) - apply (rule bind_apply_cong [OF refl], rename_tac tag s') apply (clarsimp simp: mapM_def sequence_def bind_assoc asUser_bind_distrib asUser_return submonad_asUser.fn_stateAssert) apply (case_tac f) - apply (clarsimp simp: handleFaultReply_def zipWithM_x_mapM_x - zip_Cons ARM_H.exceptionMessage_def - ARM.exceptionMessage_def - mapM_x_Cons mapM_x_Nil) - apply (rule bind_apply_cong [OF refl], rename_tac sb s'') - apply (case_tac sb) - apply (case_tac "msgLength tag < scast n_msgRegisters") - apply (fastforce simp: bind_assoc asUser_bind_distrib - zip_Cons mapM_x_Cons mapM_x_Nil - asUser_comm [OF neq] asUser_getRegister_discarded - submonad_asUser.fn_stateAssert asUser_return - bind_subst_lift [OF submonad_asUser.stateAssert_fn] - ARM_H.sanitiseRegister_def - ARM.sanitiseRegister_def - n_msgRegisters_def - dest: word_less_cases) - apply (clarsimp simp: bind_assoc asUser_bind_distrib - zip_Cons mapM_x_Cons mapM_x_Nil - asUser_comm [OF neq] asUser_getRegister_discarded - submonad_asUser.fn_stateAssert asUser_return - bind_subst_lift [OF submonad_asUser.stateAssert_fn] - word_less_nat_alt ARM_H.sanitiseRegister_def - ARM.sanitiseRegister_def - n_msgRegisters_def) - apply (case_tac "msgLength tag < scast n_msgRegisters") - apply (fastforce simp: asUser_bind_distrib - zip_Cons word_le_make_less mapM_x_Cons mapM_x_Nil - asUser_comm [OF neq] asUser_getRegister_discarded - submonad_asUser.fn_stateAssert asUser_return - bind_subst_lift [OF submonad_asUser.stateAssert_fn] - bind_assoc ARM_H.sanitiseRegister_def - ARM.sanitiseRegister_def - bind_comm_mapM_comm [OF asUser_loadWordUser_comm] - word_size stateAssert_mapM_loadWordUser_comm - n_msgRegisters_def - dest: word_less_cases) - apply (clarsimp simp: asUser_bind_distrib - zip_Cons word_le_make_less mapM_x_Cons mapM_x_Nil - asUser_comm [OF neq] asUser_getRegister_discarded - submonad_asUser.fn_stateAssert asUser_return - bind_subst_lift [OF submonad_asUser.stateAssert_fn] - bind_assoc ARM_H.sanitiseRegister_def - ARM.sanitiseRegister_def - bind_comm_mapM_comm [OF asUser_loadWordUser_comm] - word_size stateAssert_mapM_loadWordUser_comm - word_less_nat_alt n_msgRegisters_def) - apply (clarsimp simp: handleFaultReply_def asUser_getRegister_discarded + apply (clarsimp simp: handleArchFaultReply_def asUser_getRegister_discarded bind_subst_lift [OF stateAssert_stateAssert] pred_conj_def) - apply (rule bind_apply_cong [OF refl], rename_tac sb s'') - apply (case_tac sb, simp_all add: word_size n_msgRegisters_def)[1] - apply (clarsimp simp: handleFaultReply_def asUser_getRegister_discarded - bind_subst_lift [OF stateAssert_stateAssert] - pred_conj_def) - apply (rule bind_apply_cong [OF refl], rename_tac sb s'') - apply (case_tac sb, simp_all add: word_size n_msgRegisters_def)[1] - apply (clarsimp simp: handleFaultReply_def zipWithM_x_mapM_x - zip_Cons syscallMessage_unfold - msgRegisters_unfold upto_enum_word - mapM_x_Cons mapM_x_Nil) apply (rule bind_apply_cong [OF refl], rename_tac sb s'') - apply (case_tac sb) - apply (case_tac "msgLength tag < scast n_msgRegisters") - apply (fastforce simp: bind_assoc asUser_bind_distrib - zip_Cons mapM_x_Cons mapM_x_Nil - asUser_comm [OF neq] asUser_getRegister_discarded - submonad_asUser.fn_stateAssert asUser_return - bind_subst_lift [OF submonad_asUser.stateAssert_fn] - ARM_H.sanitiseRegister_def - ARM.sanitiseRegister_def - n_msgRegisters_def - dest: word_less_cases) - apply (clarsimp simp: bind_assoc asUser_bind_distrib - zip_Cons mapM_x_Cons mapM_x_Nil - asUser_comm [OF neq] asUser_getRegister_discarded - submonad_asUser.fn_stateAssert asUser_return - bind_subst_lift [OF submonad_asUser.stateAssert_fn] - word_less_nat_alt ARM_H.sanitiseRegister_def - ARM.sanitiseRegister_def take_Cons - split: nat.split) - apply (case_tac "msgLength tag < scast n_msgRegisters") - apply (fastforce simp: asUser_bind_distrib zipWithM_x_mapM_x - zip_Cons word_le_make_less mapM_x_Cons mapM_x_Nil - asUser_comm [OF neq] asUser_getRegister_discarded - submonad_asUser.fn_stateAssert asUser_return - bind_subst_lift [OF submonad_asUser.stateAssert_fn] - bind_assoc ARM_H.sanitiseRegister_def - ARM.sanitiseRegister_def - bind_comm_mapM_comm [OF asUser_loadWordUser_comm] - word_size stateAssert_mapM_loadWordUser_comm - n_msgRegisters_def - dest: word_less_cases) - apply (simp add: n_msgRegisters_def word_le_nat_alt - linorder_not_less) - apply (clarsimp | frule neq0_conv[THEN iffD2, THEN not0_implies_Suc, - OF order_less_le_trans, rotated])+ - apply (subgoal_tac "\n :: word32. n \ scast n_syscallMessage \ [n .e. msgMaxLength] - = [n .e. scast n_syscallMessage] - @ [scast n_syscallMessage + 1 .e. msgMaxLength]") - apply (simp only: upto_enum_word[where y="scast n_syscallMessage :: word32"] - upto_enum_word[where y="scast n_syscallMessage + 1 :: word32"]) - apply (clarsimp simp: bind_assoc asUser_bind_distrib - mapM_x_Cons mapM_x_Nil - asUser_comm [OF neq] asUser_getRegister_discarded - submonad_asUser.fn_stateAssert - bind_subst_lift [OF submonad_asUser.stateAssert_fn] - word_less_nat_alt ARM_H.sanitiseRegister_def - ARM.sanitiseRegister_def - split_def n_msgRegisters_def msgMaxLength_def - word_size msgLengthBits_def n_syscallMessage_def - split del: split_if cong: if_weak_cong) - apply (simp add: zipWithM_x_mapM_x mapM_x_Nil - stateAssert_mapM_loadWordUser_comm - asUser_loadWordUser_comm word_size - bind_comm_mapM_comm [OF asUser_loadWordUser_comm] - asUser_return submonad_asUser.fn_stateAssert - mapM_x_Cons bind_assoc - mapM_Cons mapM_Nil asUser_bind_distrib - take_Cons - split: nat.split cong: if_weak_cong) - apply (clarsimp simp: upto_enum_word word_le_nat_alt simp del: upt.simps) - apply (cut_tac i="unat n" and j="Suc (unat (scast n_syscallMessage :: word32))" - and k="Suc msgMaxLength" in upt_add_eq_append') - apply simp - apply (simp add: n_syscallMessage_def msgMaxLength_unfold) - apply (simp add: n_syscallMessage_def msgMaxLength_def - msgLengthBits_def shiftL_nat - del: upt.simps upt_rec_numeral) + apply (rule bind_apply_cong [OF refl], rename_tac rv r'') + apply (case_tac sb, simp_all add: word_size n_msgRegisters_def)[1] + done + +lemmas lookup_uset_getreg_swap = bind_inv_inv_comm[OF lookupIPCBuffer_inv' + user_getreg_inv' + empty_fail_lookupIPCBuffer + empty_fail_asUser[OF getRegister_empty_fail]] + +end + +lemma mapM_x_zip_take_Cons_append: + "n = 0 \ zs = [] + \ mapM_x f (zip (x # xs) (take n (y # ys) @ zs)) + = do + when (n > 0) (f (x, y)); + mapM_x f (zip xs (take (n - 1) ys @ zs)) + od" + by (cases n, simp_all add: mapM_x_Cons) + +lemma monadic_rewrite_do_flip: + "monadic_rewrite E F P (do a \ f; b \ g; return (a, b) od) + (do b \ g; a \ f; return (a, b) od) + \ monadic_rewrite E F P (do a \ f; b \ g; h a b od) + (do b \ g; a \ f; h a b od)" + apply (drule_tac h="\(a, b). h a b" in monadic_rewrite_bind_head) + apply (simp add: bind_assoc) + done + +lemma monadic_rewrite_inst: "monadic_rewrite F E P f g \ monadic_rewrite F E P f g" + by simp + +context kernel_m begin interpretation Arch . + +lemma lookupIPCBuffer_isolatable: + "thread_actions_isolatable idx (lookupIPCBuffer w t)" + apply (simp add: lookupIPCBuffer_def) + apply (rule thread_actions_isolatable_bind) + apply (clarsimp simp: put_tcb_state_regs_tcb_def threadGet_isolatable + getThreadBufferSlot_def locateSlot_conv getSlotCap_def + split: tcb_state_regs.split)+ + apply (rule thread_actions_isolatable_bind) + apply (clarsimp simp: thread_actions_isolatable_return + getCTE_isolatable + assert_isolatable + split: capability.split arch_capability.split bool.split)+ + apply (rule thread_actions_isolatable_if) + apply (rule thread_actions_isolatable_bind) + apply (simp add: assert_isolatable thread_actions_isolatable_return | wp)+ + done + +lemma handleFaultReply': + notes option.case_cong_weak [cong] wordSize_def'[simp] take_append[simp del] + assumes neq: "s \ r" + shows "monadic_rewrite True False (tcb_at' s and tcb_at' r) (do + tag \ getMessageInfo s; + sb \ lookupIPCBuffer False s; + msg \ getMRs s sb tag; + handleFaultReply f r (msgLabel tag) msg + od) (handleFaultReply' f s r)" + apply (unfold handleFaultReply'_def getMRs_def msgMaxLength_def + bit_def msgLengthBits_def msgRegisters_unfold + fromIntegral_simp1 fromIntegral_simp2 + shiftL_word) + apply (simp add: bind_assoc del: take_append) + apply (rule monadic_rewrite_bind_tail) + apply (clarsimp simp: mapM_def sequence_def bind_assoc asUser_bind_distrib + asUser_return submonad_asUser.fn_stateAssert + del: take_append) + apply (case_tac f, simp_all add: handleFaultReply_def zipWithM_x_mapM_x zip_take + del: take_append) + (* UserException *) + apply (clarsimp simp: handleFaultReply_def zipWithM_x_mapM_x + zip_Cons ARM_H.exceptionMessage_def + ARM.exceptionMessage_def + mapM_x_Cons mapM_x_Nil) + apply (rule monadic_rewrite_symb_exec_l, wp) + apply (rule_tac P="tcb_at' s and tcb_at' r" in monadic_rewrite_inst) + apply (case_tac rv; (case_tac "msgLength tag < scast n_msgRegisters", + (erule disjE[OF word_less_cases], + ( clarsimp simp: n_msgRegisters_def asUser_bind_distrib + mapM_x_Cons mapM_x_Nil bind_assoc + asUser_getRegister_discarded + asUser_comm[OF neq] + bind_comm_mapM_comm [OF asUser_loadWordUser_comm, symmetric] + word_le_nat_alt[of 4, simplified linorder_not_less[symmetric, of 4]] + asUser_return submonad_asUser.fn_stateAssert + | rule monadic_rewrite_bind_tail monadic_rewrite_refl + monadic_rewrite_symb_exec_l[OF stateAssert_inv] + monadic_rewrite_symb_exec_l[OF mapM_x_mapM_valid[OF mapM_x_wp']] + | wp asUser_tcb_at' lookupIPCBuffer_inv')+)+)) + apply wp + (* capFault *) + apply (rule monadic_rewrite_symb_exec_l, wp empty_fail_asUser)+ + apply(case_tac rv) + apply (clarsimp + | rule monadic_rewrite_bind_tail monadic_rewrite_refl + monadic_rewrite_symb_exec_l[OF mapM_x_mapM_valid[OF mapM_x_wp']] + | wp mapM_x_mapM_valid[OF mapM_x_wp'[OF loadWordUser_inv]] + empty_fail_loadWordUser)+ + (* UnknownSyscallExceptio *) + apply (simp add: zip_append2 mapM_x_append asUser_bind_distrib split_def + bind_assoc) + apply (rule monadic_rewrite_imp) + apply (rule monadic_rewrite_trans[rotated]) + apply (rule monadic_rewrite_do_flip) + apply (rule_tac P="inj (case_bool s r)" in monadic_rewrite_gen_asm) + apply (rule monadic_rewrite_trans[OF _ monadic_rewrite_transverse]) + apply (rule monadic_rewrite_weaken[where F=False and E=True], simp) + apply (rule isolate_thread_actions_rewrite_bind + bool.simps setRegister_simple + zipWithM_setRegister_simple + thread_actions_isolatable_bind lookupIPCBuffer_isolatable + lookupIPCBuffer_isolatable[THEN thread_actions_isolatableD] + copy_registers_isolate_general thread_actions_isolatable_return + thread_actions_isolatable_return[THEN thread_actions_isolatableD] + | assumption + | wp assert_inv)+ + apply (rule monadic_rewrite_isolate_final[where P="\"]) + apply simp+ + (* swap ends *) + apply (clarsimp simp: handleFaultReply_def zipWithM_x_mapM_x + zip_Cons syscallMessage_unfold + n_syscallMessage_def + upto_enum_word mapM_x_Cons mapM_x_Nil) + apply (rule monadic_rewrite_bind_tail) + apply (case_tac sb) + apply (case_tac "msgLength tag < scast n_msgRegisters") + apply (erule disjE[OF word_less_cases], + ( clarsimp simp: n_msgRegisters_def asUser_bind_distrib + mapM_x_Cons mapM_x_Nil bind_assoc + asUser_getRegister_discarded + asUser_comm[OF neq] take_zip + word_le_nat_alt[of 4, simplified linorder_not_less[symmetric, of 4]] + asUser_return submonad_asUser.fn_stateAssert + | rule monadic_rewrite_bind_tail monadic_rewrite_refl + monadic_rewrite_symb_exec_l[OF stateAssert_inv] + | wp asUser_tcb_at')+)+ + apply (case_tac "msgLength tag < scast n_msgRegisters") + apply (erule disjE[OF word_less_cases], + ( clarsimp simp: n_msgRegisters_def asUser_bind_distrib + mapM_x_Cons mapM_x_Nil bind_assoc + zipWithM_x_Nil + asUser_getRegister_discarded + asUser_comm[OF neq] take_zip + bind_comm_mapM_comm [OF asUser_loadWordUser_comm, symmetric] + asUser_return submonad_asUser.fn_stateAssert + | rule monadic_rewrite_bind_tail monadic_rewrite_refl + monadic_rewrite_symb_exec_l[OF mapM_x_mapM_valid[OF mapM_x_wp']] + monadic_rewrite_symb_exec_l[OF stateAssert_inv] + | wp asUser_tcb_at')+)+ + apply (simp add: n_msgRegisters_def word_le_nat_alt n_syscallMessage_def + linorder_not_less syscallMessage_unfold) + apply (clarsimp | frule neq0_conv[THEN iffD2, THEN not0_implies_Suc, + OF order_less_le_trans, rotated])+ + apply (subgoal_tac "\n :: word32. n \ scast n_syscallMessage \ [n .e. msgMaxLength] + = [n .e. scast n_syscallMessage] + @ [scast n_syscallMessage + 1 .e. msgMaxLength]") + apply (simp only: upto_enum_word[where y="scast n_syscallMessage :: word32"] + upto_enum_word[where y="scast n_syscallMessage + 1 :: word32"]) + apply (clarsimp simp: bind_assoc asUser_bind_distrib + mapM_x_Cons mapM_x_Nil + asUser_comm [OF neq] asUser_getRegister_discarded + submonad_asUser.fn_stateAssert take_zip + bind_subst_lift [OF submonad_asUser.stateAssert_fn] + word_less_nat_alt ARM_H.sanitiseRegister_def + split_def n_msgRegisters_def msgMaxLength_def + bind_comm_mapM_comm [OF asUser_loadWordUser_comm, symmetric] + word_size msgLengthBits_def n_syscallMessage_def + split del: split_if cong: if_weak_cong) + apply (rule monadic_rewrite_bind_tail)+ + apply (subst (2) upto_enum_word) + apply (case_tac "ma < unat n_syscallMessage - 4") + apply (erule disjE[OF nat_less_cases'], + ( clarsimp simp: n_syscallMessage_def bind_assoc asUser_bind_distrib + mapM_x_Cons mapM_x_Nil zipWithM_x_mapM_x mapM_Cons + bind_comm_mapM_comm [OF asUser_loadWordUser_comm, symmetric] + asUser_loadWordUser_comm loadWordUser_discarded asUser_return + zip_take_triv2 msgMaxLength_def + cong: if_weak_cong + | rule monadic_rewrite_bind_tail + monadic_rewrite_refl + monadic_rewrite_symb_exec_l[OF stateAssert_inv] + monadic_rewrite_symb_exec_l[OF mapM_x_mapM_valid[OF mapM_x_wp']] + | wp empty_fail_loadWordUser)+)+ + apply (clarsimp simp: upto_enum_word word_le_nat_alt simp del: upt.simps cong: if_weak_cong) + apply (cut_tac i="unat n" and j="Suc (unat (scast n_syscallMessage :: word32))" + and k="Suc msgMaxLength" in upt_add_eq_append') + apply (simp add: n_syscallMessage_def) + apply (simp add: n_syscallMessage_def msgMaxLength_unfold) + apply (simp add: n_syscallMessage_def msgMaxLength_def + msgLengthBits_def shiftL_nat + del: upt.simps upt_rec_numeral) + apply (simp add: upto_enum_word cong: if_weak_cong) + apply wp + (* ArchFault *) + apply (simp add: neq inj_case_bool split: bool.split) + apply (rule monadic_rewrite_imp) + apply (rule monadic_rewrite_is_refl) + apply (rule ext) + apply (unfold handleArchFaultReply'[symmetric] getMRs_def msgMaxLength_def + bit_def msgLengthBits_def msgRegisters_unfold + fromIntegral_simp1 fromIntegral_simp2 shiftL_word) + apply (clarsimp simp: mapM_def sequence_def bind_assoc asUser_bind_distrib + asUser_return submonad_asUser.fn_stateAssert bit_def) + apply wp done end @@ -994,6 +1061,7 @@ end context kernel_m begin lemma setMRs_lookup_failure_ccorres: + shows "ccorres (\r r'. r \ [] \ last r = unat (r' && mask msgLengthBits)) ret__unsigned_' (valid_pspace' @@ -1021,6 +1089,7 @@ lemma setMRs_lookup_failure_ccorres: apply (rule ccorres_guard_imp2) apply csymbr apply (ctac add: setMR_ccorres) + apply (ccorres_rewrite) apply (simp add: ccorres_cond_iffs) apply (rule ccorres_return_C, simp+)[1] apply wp @@ -1032,6 +1101,7 @@ lemma setMRs_lookup_failure_ccorres: apply csymbr apply (ctac add: setMR_ccorres) apply (simp add: ccorres_cond_iffs) + apply (ccorres_rewrite) apply (rule ccorres_rhs_assoc)+ apply csymbr apply (ctac add: setMR_ccorres) @@ -1050,6 +1120,7 @@ lemma setMRs_lookup_failure_ccorres: apply csymbr apply (ctac add: setMR_ccorres) apply (simp add: ccorres_cond_iffs) + apply (ccorres_rewrite) apply (rule ccorres_rhs_assoc)+ apply csymbr apply (ctac add: setMR_ccorres_dc) @@ -1073,6 +1144,7 @@ lemma setMRs_lookup_failure_ccorres: apply csymbr apply (ctac add: setMR_ccorres) apply (simp add: ccorres_cond_iffs) + apply (ccorres_rewrite) apply (rule ccorres_rhs_assoc)+ apply csymbr apply (ctac add: setMR_ccorres_dc) @@ -1358,6 +1430,16 @@ lemma asUser_tcbFault_obj_at: apply simp done +lemma asUser_atcbContext_obj_at: + "t \ t' \ + \obj_at' (\tcb. P ((atcbContextGet o tcbArch) tcb)) t\ + asUser t' m + \\rv. obj_at' (\tcb. P ((atcbContextGet o tcbArch) tcb)) t\" + apply (simp add: asUser_def split_def atcbContextGet_def atcbContextSet_def) + apply (wp threadGet_wp) + apply simp + done + lemma asUser_tcbFault_inv: "\\s. \t. ko_at' t p' s \ tcbFault t = f\ asUser p m \\rv s. \t. ko_at' t p' s \ tcbFault t = f\" @@ -1367,6 +1449,16 @@ lemma asUser_tcbFault_inv: apply (clarsimp simp: obj_at'_def)+ done +lemma setMR_atcbContext_obj_at: + "t \ t' \ + \obj_at' (\tcb. P ((atcbContextGet o tcbArch) tcb)) t\ + setMR t' b r v + \\rv. obj_at' (\tcb. P ((atcbContextGet o tcbArch) tcb)) t\" + apply (simp add: setMR_def split del: split_if) + apply (rule hoare_pre) + apply (wp asUser_atcbContext_obj_at[simplified] | simp | wpc)+ + done + lemma setMR_tcbFault_obj_at: "\obj_at' (\tcb. P (tcbFault tcb)) t\ setMR t' b r v \\rv. obj_at' (\tcb. P (tcbFault tcb)) t\" @@ -1413,10 +1505,10 @@ lemma ccorres_add_getRegister: done lemma user_getreg_rv: - "\obj_at' (\tcb. P (tcbContext tcb r)) t\ asUser t (getRegister r) \\rv s. P rv\" + "\obj_at' (\tcb. P ((atcbContextGet o tcbArch) tcb r)) t\ asUser t (getRegister r) \\rv s. P rv\" apply (simp add: asUser_def split_def) apply (wp threadGet_wp) - apply (clarsimp simp: obj_at'_def projectKOs getRegister_def in_monad) + apply (clarsimp simp: obj_at'_def projectKOs getRegister_def in_monad atcbContextGet_def) done lemma exceptionMessage_ccorres: @@ -1424,7 +1516,7 @@ lemma exceptionMessage_ccorres: \ register_from_H (ARM_H.exceptionMessage ! n) = index exceptionMessageC n" apply (simp add: exceptionMessageC_def ARM_H.exceptionMessage_def - ARM.exceptionMessage_def) + ARM.exceptionMessage_def MessageID_Exception_def) apply (simp add: Arrays.update_def n_exceptionMessage_def fcp_beta nth_Cons' fupdate_def split: split_if) @@ -1438,6 +1530,313 @@ lemma asUser_obj_at_elsewhere: apply clarsimp done +lemma exceptionMessage_length_aux : + "\n. n < length ARM_H.exceptionMessage \ n < unat n_exceptionMessage" + by (simp add: ARM.exceptionMessage_def ARM_H.exceptionMessage_def n_exceptionMessage_def) + +lemma copyMRsFault_ccorres_exception: + "ccorres dc xfdc + (valid_pspace' + and obj_at' (\tcb. map (atcbContext (tcbArch tcb)) ARM_H.exceptionMessage = msg) sender + and K (length msg = 3) + and K (recvBuffer \ Some 0) + and K (sender \ receiver)) + (UNIV \ \\sender = tcb_ptr_to_ctcb_ptr sender\ + \ \\receiver = tcb_ptr_to_ctcb_ptr receiver\ + \ \\receiveIPCBuffer = Ptr (option_to_0 recvBuffer)\ + \ \ \length___unsigned_long = 3 \ + \ \ \id___anonymous_enum = MessageID_Exception \) + hs + (mapM_x (\(x, y). setMR receiver recvBuffer x y) (zip [0..<120] msg)) + (Call copyMRsFault_'proc)" + apply (unfold K_def) + apply (intro ccorres_gen_asm) + apply (cinit' lift: sender_' receiver_' receiveIPCBuffer_' + length___unsigned_long_' id___anonymous_enum_' + simp: whileAnno_def) + apply (simp only: mapM_x_append[where xs="take (unat n_msgRegisters) (zip as bs)" + and ys="drop (unat n_msgRegisters) (zip as bs)" + for as bs, simplified] bind_assoc) + apply (rule ccorres_rhs_assoc2, rule ccorres_split_nothrow_novcg) + + apply (rule_tac F="K $ obj_at' (\tcb. map ((atcbContext o tcbArch) tcb) ARM_H.exceptionMessage = msg) sender" + in ccorres_mapM_x_while) + apply (clarsimp simp: n_msgRegisters_def) + apply (rule ccorres_guard_imp2) + apply (rule_tac t=sender and r="ARM_H.exceptionMessage ! n" + in ccorres_add_getRegister) + apply (ctac(no_vcg)) + apply (rule_tac P="\s. rv = msg ! n" in ccorres_cross_over_guard) + apply (simp add: setMR_def length_msgRegisters n_msgRegisters_def + liftM_def[symmetric]) + apply ctac + apply (wp user_getreg_rv) + apply (clarsimp simp: msgRegisters_ccorres n_msgRegisters_def + syscallMessage_ccorres n_syscallMessage_def + obj_at'_def projectKOs + atcbContextGet_def unat_of_nat32[unfolded word_bits_def]) + apply (clarsimp simp: exceptionMessage_ccorres[simplified,symmetric,OF exceptionMessage_length_aux]) + apply (clarsimp simp: word_of_nat_less MessageID_Exception_def) + apply (clarsimp simp: n_msgRegisters_def foo) + apply (rule allI, rule conseqPre, vcg exspec=setRegister_modifies exspec=getRegister_modifies) + apply simp + apply (simp add: setMR_def split del: split_if) + apply (rule hoare_pre) + apply (wp asUser_obj_at_elsewhere | wpc)+ + apply simp + apply (simp add: word_bits_def) + apply ceqv + apply (rule ccorres_Cond_rhs) + apply (simp del: Collect_const) + apply (simp add: n_msgRegisters_def mapM_x_Nil) + apply (subst ccorres_expand_while_iff[symmetric]) + apply (rule ccorres_cond_false) + apply (rule ccorres_return_Skip') + apply (simp add: n_msgRegisters_def mapM_x_Nil) + apply (rule ccorres_return_Skip') + apply wp + apply (simp add: guard_is_UNIV_def) + apply (clarsimp) + done + +lemma mapM_cong: "\ \x. elem x xs \ f x = g x \ \ mapM_x f xs = mapM_x g xs" + by (induction xs, (simp add: mapM_x_Nil mapM_x_Cons)+) + +lemma copyMRsFault_ccorres_syscall: + "ccorres dc xfdc + (valid_pspace' + and obj_at' (\tcb. map (atcbContext (tcbArch tcb)) ARM_H.syscallMessage = msg) sender + and (case recvBuffer of Some x \ valid_ipc_buffer_ptr' x | None \ \) + and K (length msg = 12) + and K (recvBuffer \ Some 0) + and K (sender \ receiver)) + (UNIV \ \\sender = tcb_ptr_to_ctcb_ptr sender\ + \ \\receiver = tcb_ptr_to_ctcb_ptr receiver\ + \ \\receiveIPCBuffer = Ptr (option_to_0 recvBuffer)\ + \ \\length___unsigned_long = 12 \ + \ \ \id___anonymous_enum = MessageID_Syscall \) + hs + (mapM_x (\(x, y). setMR receiver recvBuffer x y) (zip [0..<120] msg)) + (Call copyMRsFault_'proc)" +proof - + (* auxiliary lemmas *) + have option_to_0_imp : "\ option_to_0 recvBuffer= 0 ; recvBuffer \ Some 0 \ \ recvBuffer = None" + by (simp add: option_to_0_def; cases recvBuffer; simp+) + have drop_n: "\n m. drop n [0..x y n m. elem (x, y) (zip [n..<120] m) \ \ x < n " + by (simp | erule in_set_zipE)+ + have msg_aux: "\p. elem p (zip [4..<120] (drop 4 msg)) + \ (\(x1,y1). setMR receiver None x1 y1) p = (\_ . return (length msgRegisters)) p" + by (fastforce simp add: numeral_eqs setMR_def less_than_4 n_msgRegisters_def length_msgRegisters) + have mapM_x_return_gen: "\v w xs. mapM_x (\_. return v) xs = return w" (* FIXME mapM_x_return *) + by (induct_tac xs; simp add: mapM_x_Nil mapM_x_Cons) + show ?thesis + apply (unfold K_def) + apply (intro ccorres_gen_asm) + apply (cinit' lift: sender_' receiver_' receiveIPCBuffer_' + length___unsigned_long_' id___anonymous_enum_' + simp: whileAnno_def) + apply (simp only: mapM_x_append[where xs="take (unat n_msgRegisters) (zip as bs)" + and ys="drop (unat n_msgRegisters) (zip as bs)" + for as bs, simplified] bind_assoc) + apply (rule ccorres_rhs_assoc2, rule ccorres_split_nothrow_novcg) + apply (rule_tac F="K $ obj_at' (\tcb. map ((atcbContext o tcbArch) tcb) ARM_H.syscallMessage = msg) sender" + in ccorres_mapM_x_while) + apply (clarsimp simp: n_msgRegisters_def) + apply (rule ccorres_guard_imp2) + apply (rule_tac t=sender and r="ARM_H.syscallMessage ! n" + in ccorres_add_getRegister) + apply (ctac(no_vcg)) + apply (rule_tac P="\s. rv = msg ! n" in ccorres_cross_over_guard) + apply (simp add: setMR_def length_msgRegisters n_msgRegisters_def + liftM_def[symmetric]) + apply ctac + apply (wp user_getreg_rv) + apply (clarsimp simp: msgRegisters_ccorres n_msgRegisters_def + syscallMessage_ccorres n_syscallMessage_def + obj_at'_def projectKOs + atcbContextGet_def unat_of_nat32[unfolded word_bits_def]) + apply (clarsimp simp: word_of_nat_less MessageID_Syscall_def) + apply (simp add: n_msgRegisters_def) + apply (rule allI, rule conseqPre, vcg exspec=setRegister_modifies exspec=getRegister_modifies) + apply simp + apply (simp add: setMR_def split del: split_if) + apply (rule hoare_pre) + apply (wp asUser_obj_at_elsewhere | wpc)+ + apply simp + apply (simp add: word_bits_def) + apply ceqv + apply (rule ccorres_Cond_rhs) + apply (simp del: Collect_const) + apply (rule ccorres_rel_imp[where r = "\rv rv'. True", simplified]) + apply (rule_tac F="\_. obj_at' (\tcb. map ((atcbContext o tcbArch) tcb) ARM_H.syscallMessage = msg) + sender and valid_pspace' + and (case recvBuffer of Some x \ valid_ipc_buffer_ptr' x | None \ \)" + in ccorres_mapM_x_while'[where i="unat n_msgRegisters"]) + apply (clarsimp simp: setMR_def n_msgRegisters_def length_msgRegisters + option_to_0_def liftM_def[symmetric] + split: option.split_asm) + apply (rule ccorres_guard_imp2) + apply (rule_tac t=sender and r="ARM_H.syscallMessage ! (n + unat n_msgRegisters)" + in ccorres_add_getRegister) + apply (ctac(no_vcg)) + apply (rule_tac P="\s. rv = msg ! (n + unat n_msgRegisters)" + in ccorres_cross_over_guard) + apply (rule ccorres_move_array_assertion_ipc_buffer + | (rule ccorres_flip_Guard, rule ccorres_move_array_assertion_ipc_buffer))+ + apply (simp add: storeWordUser_def) + apply (rule ccorres_pre_stateAssert) + apply (ctac add: storeWord_ccorres[unfolded fun_app_def]) + apply (simp add: pred_conj_def) + apply (wp user_getreg_rv) + apply (clarsimp simp: n_syscallMessage_def n_msgRegisters_def + syscallMessage_ccorres msgRegisters_ccorres + unat_add_lem[THEN iffD1] unat_of_nat32 + word_bits_def word_size_def) + apply (simp only:field_simps imp_ex imp_conjL) + apply (clarsimp simp: pointerInUserData_c_guard obj_at'_def + pointerInUserData_h_t_valid + atcbContextGet_def + projectKOs objBits_simps word_less_nat_alt + unat_add_lem[THEN iffD1] unat_of_nat) + apply (clarsimp simp: pointerInUserData_h_t_valid rf_sr_def + MessageID_Syscall_def + msg_align_bits valid_ipc_buffer_ptr'_def) + apply (erule aligned_add_aligned) + apply (rule aligned_add_aligned[where n=2]) + apply (simp add: is_aligned_def) + apply (rule is_aligned_mult_triv2 [where n=2, simplified]) + apply (simp add: wb_gt_2)+ + apply (simp add: n_msgRegisters_def) + apply (vcg exspec=getRegister_modifies) + apply simp + apply (simp add: setMR_def n_msgRegisters_def length_msgRegisters) + apply (rule hoare_pre) + apply (wp hoare_case_option_wp | wpc)+ + apply clarsimp + apply (simp add: n_msgRegisters_def word_bits_def) + apply (simp add: n_msgRegisters_def) + apply (frule (1) option_to_0_imp) + apply (subst drop_zip) + apply (subst drop_n) + apply (clarsimp simp: n_msgRegisters_def numeral_eqs + mapM_cong[OF msg_aux, simplified numeral_eqs] + mapM_x_return_gen) + apply (subst mapM_x_return_gen[where w2="()"]) + apply (rule ccorres_return_Skip[simplified dc_def]) + apply (clarsimp) + apply (rule hoare_impI) + apply (wp mapM_x_wp_inv setMR_atcbContext_obj_at[simplified atcbContextGet_def, simplified] + | clarsimp + | wpc)+ + apply (cases recvBuffer; simp add: option_to_0_def ) + apply wp + apply (clarsimp simp: guard_is_UNIV_def n_msgRegisters_def msgLengthBits_def + mask_def)+ + done + qed + +lemma Arch_setMRs_fault_ccorres: + "ccorres (\r r'. r = r' && mask msgLengthBits) ret__unsigned_long_' + (valid_pspace' and obj_at' (\tcb. tcbFault tcb = Some ft) sender + and K (ft = ArchFault aft) + and (case buffer of Some x \ valid_ipc_buffer_ptr' x | None \ \) + and K (buffer \ Some 0) + and K (sender \ receiver)) + (UNIV \ \\sender = tcb_ptr_to_ctcb_ptr sender\ + \ \\receiver = tcb_ptr_to_ctcb_ptr receiver\ + \ \\faultType = fault_to_fault_tag ft \ + \ \\receiveIPCBuffer = Ptr (option_to_0 buffer)\) hs + (makeArchFaultMessage aft sender >>= (\ms. setMRs receiver buffer (snd ms))) + (Call Arch_setMRs_fault_'proc)" +proof - + let ?obj_at_ft = "obj_at' (\tcb. tcbFault tcb = Some ft) sender" + note symb_exec_r_fault = ccorres_symb_exec_r_known_rv_UNIV + [where xf'=ret__unsigned_' and R="?obj_at_ft" and R'=UNIV] + show ?thesis + apply (unfold K_def) + using [[goals_limit = 1]] + apply (intro ccorres_gen_asm) + apply (cinit' lift: sender_' receiver_' faultType_' receiveIPCBuffer_') + apply (simp only: makeArchFaultMessage_def setMRs_to_setMR + del: Collect_const split del: split_if) + apply (rule_tac P="ft = ArchFault aft" in ccorres_gen_asm) + apply wpc + apply (rename_tac list) + apply (rule_tac P="zip [Suc (Suc 0) ..< msgMaxLength] list = [(2, hd list), (3, hd (tl list))]" + in ccorres_gen_asm) + apply (simp add: Collect_True Collect_False ccorres_cond_iffs + zip_upt_Cons msgMaxLength_unfold + zipWithM_mapM mapM_Cons bind_assoc + seL4_Fault_tag_defs + del: Collect_const) + apply (rule ccorres_rhs_assoc)+ + apply simp + apply (rule ccorres_rhs_assoc)+ + apply (ctac (no_vcg)) + apply (rule ccorres_stateAssert) + apply (ctac(no_vcg) add: setMR_ccorres_dc) + apply (rule ccorres_move_c_guard_tcb) + apply (rule_tac val="vmFaultAddress aft" in symb_exec_r_fault) + apply (rule conseqPre, vcg, clarsimp) + apply (drule(1) obj_at_cslift_tcb) + apply (clarsimp simp: ctcb_relation_def typ_heap_simps + cfault_rel_def seL4_Fault_lift_def + seL4_Fault_VMFault_lift_def Let_def + split: split_if_asm) + apply ceqv + apply (ctac(no_vcg) add: setMR_ccorres_dc) + apply (rule_tac val="hd (vmFaultArchData aft)" in symb_exec_r_fault) + apply (rule conseqPre, vcg, clarsimp) + apply (drule(1) obj_at_cslift_tcb) + apply (clarsimp simp: typ_heap_simps') + apply (clarsimp simp: ctcb_relation_def + cfault_rel_def seL4_Fault_lift_def + seL4_Fault_VMFault_lift_def Let_def + split: split_if_asm) + apply ceqv + apply (ctac(no_vcg) add: setMR_ccorres_dc) + apply (rule_tac val="vmFaultArchData aft ! 1" in symb_exec_r_fault) + apply (rule conseqPre, vcg, clarsimp) + apply (drule(1) obj_at_cslift_tcb) + apply (clarsimp simp: typ_heap_simps') + apply (clarsimp simp: ctcb_relation_def + cfault_rel_def seL4_Fault_lift_def + seL4_Fault_VMFault_lift_def Let_def + split: split_if_asm) + apply ceqv + apply (ctac(no_vcg) add: setMR_ccorres) + apply (simp add: mapM_Nil) + apply (rule ccorres_return_C, simp+)[1] + apply wp + apply (clarsimp simp: option_to_ptr_def) + apply (subgoal_tac "case list of [a, b] \ True | _ \ True") + apply (simp add: zip_upt_Cons guard_is_UNIVI seL4_VMFault_FSR_def split: list.split_asm) + apply (simp split: list.split) + apply (wp setMR_tcbFault_obj_at asUser_inv[OF getRestartPC_inv] + hoare_case_option_wp + | simp add: option_to_ptr_def guard_is_UNIVI + seL4_VMFault_PrefetchFault_def + seL4_VMFault_Addr_def + seL4_VMFault_IP_def + msgMaxLength_def + del: Collect_const split del: split_if)+ + apply clarsimp + apply (drule(1) obj_at_cslift_tcb[where thread=sender]) + apply (clarsimp simp: obj_at'_def projectKOs objBits_simps) + apply (clarsimp simp: ctcb_relation_def cfault_rel_def seL4_Fault_lift_def + Let_def zip_upt_Cons + split: split_if_asm) + done +qed + +lemmas seL4_Faults = seL4_Fault_UserException_def + seL4_Fault_UnknownSyscall_def + seL4_Fault_CapFault_def + +lemmas seL4_Arch_Faults = seL4_Fault_VMFault_def + lemma setMRs_fault_ccorres [corres]: "ccorres (\r r'. r = r' && mask msgLengthBits) ret__unsigned_long_' (valid_pspace' and obj_at' (\tcb. tcbFault tcb = Some ft) sender @@ -1454,7 +1853,6 @@ proof - let ?obj_at_ft = "obj_at' (\tcb. tcbFault tcb = Some ft) sender" note symb_exec_r_fault = ccorres_symb_exec_r_known_rv_UNIV [where xf'=ret__unsigned_' and R="?obj_at_ft" and R'=UNIV] - show ?thesis apply (unfold K_def) apply (intro ccorres_gen_asm) @@ -1465,11 +1863,12 @@ proof - apply (vcg, clarsimp) apply (drule(1) obj_at_cslift_tcb) apply (clarsimp simp: ctcb_relation_def typ_heap_simps' - cfault_rel_def fault_lift_def Let_def + cfault_rel_def seL4_Fault_lift_def Let_def split: split_if_asm) apply ceqv + (* UserException *) apply wpc - apply (simp add: bind_assoc fault_tag_defs ccorres_cond_iffs + apply (simp add: bind_assoc seL4_Fault_tag_defs ccorres_cond_iffs Collect_True Collect_False zipWithM_mapM zip_append2 mapM_append del: Collect_const split del: split_if) @@ -1481,45 +1880,16 @@ proof - zip_upt_Cons mapM_Nil mapM_Cons bind_assoc mapM_discarded del: Collect_const upt.simps upt_rec_numeral - split del: split_if) + split del: split_if) apply (rule ccorres_rhs_assoc)+ - apply csymbr - apply (simp del: Collect_const split del: split_if) - apply (rule ccorres_rhs_assoc2, rule ccorres_split_nothrow) - apply (rule_tac F="K $ obj_at' (\tcb. map (tcbContext tcb) ARM_H.exceptionMessage = msg) sender" - in ccorres_mapM_x_while) - apply (clarsimp simp: setMR_def msgRegisters_unfold liftM_def[symmetric] - split del: split_if) - apply (rule ccorres_guard_imp2) - apply (rule_tac t=sender and r="ARM_H.exceptionMessage ! n" - in ccorres_add_getRegister) - apply ctac - apply (rule_tac P="\s. rv = msg ! n" in ccorres_cross_over_guard) - apply ctac - apply (wp user_getreg_rv) - apply simp - apply (vcg exspec=getRegister_modifies) - apply (clarsimp simp: obj_at'_def projectKOs objBits_simps - exceptionMessage_ccorres n_exceptionMessage_def - unat_of_nat32[unfolded word_bits_def]) - apply (simp add: msgRegisters_ccorres[symmetric] n_msgRegisters_def - msgRegisters_unfold word_of_nat_less) - apply simp - apply (rule allI, rule conseqPre) - apply (vcg exspec=setRegister_modifies exspec=getRegister_modifies) - apply clarsimp - apply (simp add: setMR_def split del: split_if) - apply (rule hoare_pre) - apply (wp asUser_obj_at_elsewhere | wpc | simp)+ - apply (simp add: word_bits_def) - apply ceqv + apply (ctac add: copyMRsFault_ccorres_exception) apply (rule ccorres_move_c_guard_tcb) apply (rule_tac val="userExceptionNumber ft" in symb_exec_r_fault) apply (vcg, clarsimp) apply (drule(1) obj_at_cslift_tcb) apply (clarsimp simp: ctcb_relation_def typ_heap_simps - cfault_rel_def fault_lift_def - fault_user_exception_lift_def Let_def + cfault_rel_def seL4_Fault_lift_def + seL4_Fault_UserException_lift_def Let_def split: split_if_asm) apply ceqv apply (ctac add: setMR_ccorres_dc) @@ -1528,8 +1898,8 @@ proof - apply (vcg, clarsimp) apply (drule(1) obj_at_cslift_tcb) apply (clarsimp simp: ctcb_relation_def typ_heap_simps - cfault_rel_def fault_lift_def - fault_user_exception_lift_def Let_def + cfault_rel_def seL4_Fault_lift_def + seL4_Fault_UserException_lift_def Let_def split: split_if_asm) apply ceqv apply (ctac add: setMR_ccorres) @@ -1547,116 +1917,66 @@ proof - apply (wp mapM_x_wp' setMR_tcbFault_obj_at hoare_case_option_wp | simp)+ - apply (fold whileAnno_def[where I=UNIV and V=undefined]) - apply (vcg exspec=getRegister_modifies exspect=setRegister_modifies) - apply simp - apply simp + apply (vcg exspec=copyMRsFault_modifies exspect=setRegister_modifies) apply (wp asUser_inv mapM_wp' getRegister_inv) apply simp - apply (wp asUser_inv mapM_wp' getRegister_inv hoare_drop_imps - asUser_get_registers asUser_const_rv) + apply (wp asUser_inv mapM_wp' getRegister_inv hoare_drop_imps asUser_const_rv + asUser_get_registers[simplified atcbContextGet_def comp_def]) apply simp - apply (rename_tac list) - apply (rule_tac P="zip [Suc (Suc 0) ..< msgMaxLength] list = [(2, hd list), (3, hd (tl list))]" - in ccorres_gen_asm) - apply (simp add: bind_assoc fault_tag_defs ccorres_cond_iffs - Collect_True Collect_False zipWithM_mapM + (* CapFault *) + apply (simp add: Collect_True Collect_False ccorres_cond_iffs zip_upt_Cons msgMaxLength_unfold - mapM_Cons mapM_Nil + zipWithM_mapM mapM_Cons bind_assoc del: Collect_const) apply (rule ccorres_rhs_assoc)+ - apply simp - apply (rule ccorres_rhs_assoc)+ apply (ctac(no_vcg)) apply (rule ccorres_stateAssert) apply (ctac(no_vcg) add: setMR_ccorres_dc) - apply (rule ccorres_move_c_guard_tcb) - apply (rule_tac val="vmFaultAddress ft" in symb_exec_r_fault) + apply (rule_tac val="capFaultAddress ft" in symb_exec_r_fault) apply (rule conseqPre, vcg, clarsimp) apply (drule(1) obj_at_cslift_tcb) - apply (clarsimp simp: ctcb_relation_def typ_heap_simps - cfault_rel_def fault_lift_def - fault_vm_fault_lift_def Let_def + apply (clarsimp simp: typ_heap_simps' ctcb_relation_def + cfault_rel_def seL4_Fault_lift_def Let_def + seL4_Fault_CapFault_lift_def split: split_if_asm) apply ceqv apply (ctac(no_vcg) add: setMR_ccorres_dc) - apply (rule_tac val="hd (vmFaultArchData ft)" in symb_exec_r_fault) + apply (rule_tac val="from_bool (capFaultInReceivePhase ft)" in symb_exec_r_fault) apply (rule conseqPre, vcg, clarsimp) apply (drule(1) obj_at_cslift_tcb) - apply (clarsimp simp: typ_heap_simps') - apply (clarsimp simp: ctcb_relation_def - cfault_rel_def fault_lift_def - fault_vm_fault_lift_def Let_def + apply (clarsimp simp: typ_heap_simps' ctcb_relation_def + cfault_rel_def seL4_Fault_lift_def Let_def + seL4_Fault_CapFault_lift_def + from_bool_to_bool_and_1 word_size split: split_if_asm) apply ceqv apply (ctac(no_vcg) add: setMR_ccorres_dc) - apply (rule_tac val="vmFaultArchData ft ! 1" in symb_exec_r_fault) - apply (rule conseqPre, vcg, clarsimp) - apply (drule(1) obj_at_cslift_tcb) - apply (clarsimp simp: typ_heap_simps') - apply (clarsimp simp: ctcb_relation_def - cfault_rel_def fault_lift_def - fault_vm_fault_lift_def Let_def - split: split_if_asm) - apply ceqv - apply (ctac(no_vcg) add: setMR_ccorres) - apply (rule ccorres_return_C, simp+)[1] - apply wp - apply (clarsimp simp: option_to_ptr_def) - apply (subgoal_tac "case list of [a, b] \ True | _ \ True") - apply (simp add: zip_upt_Cons guard_is_UNIVI split: list.split_asm) - apply (simp split: list.split) - apply (wp setMR_tcbFault_obj_at asUser_inv[OF getRestartPC_inv] - hoare_case_option_wp - | simp add: option_to_ptr_def guard_is_UNIVI - del: Collect_const split del: split_if)+ - apply (simp add: Collect_True Collect_False ccorres_cond_iffs - zip_upt_Cons msgMaxLength_unfold - zipWithM_mapM mapM_Cons bind_assoc - del: Collect_const) - apply (rule ccorres_rhs_assoc)+ - apply (ctac(no_vcg)) - apply (rule ccorres_stateAssert) - apply (ctac(no_vcg) add: setMR_ccorres_dc) - apply (rule_tac val="capFaultAddress ft" in symb_exec_r_fault) - apply (rule conseqPre, vcg, clarsimp) + apply (rule ccorres_move_c_guard_tcb) + apply (rule_tac P="obj_at' (\tcb. tcbFault tcb = Some ft) sender" + in ccorres_cross_over_guard) + apply (ctac(no_vcg) add: setMRs_lookup_failure_ccorres[unfolded msgMaxLength_unfold]) + apply simp + apply (rule ccorres_return_C, simp+)[1] + apply (wp setMR_tcbFault_obj_at hoare_case_option_wp) + apply (clarsimp simp: option_to_ptr_def Collect_const_mem guard_is_UNIV_def) + apply (rule conjI) + apply (simp add: seL4_CapFault_InRecvPhase_def) + apply (rule conjI) + apply (simp add: from_bool_def split: bool.split) + apply clarsimp apply (drule(1) obj_at_cslift_tcb) apply (clarsimp simp: typ_heap_simps' ctcb_relation_def - cfault_rel_def fault_lift_def Let_def - fault_cap_fault_lift_def + cfault_rel_def seL4_Fault_lift_def Let_def + seL4_Fault_CapFault_lift_def is_cap_fault_def + seL4_CapFault_LookupFailureType_def split: split_if_asm) - apply ceqv - apply (ctac(no_vcg) add: setMR_ccorres_dc) - apply (rule_tac val="from_bool (capFaultInReceivePhase ft)" in symb_exec_r_fault) - apply (rule conseqPre, vcg, clarsimp) - apply (drule(1) obj_at_cslift_tcb) - apply (clarsimp simp: typ_heap_simps' ctcb_relation_def - cfault_rel_def fault_lift_def Let_def - fault_cap_fault_lift_def - from_bool_to_bool_and_1 word_size - split: split_if_asm) - apply ceqv - apply (ctac(no_vcg) add: setMR_ccorres_dc) - apply (rule ccorres_move_c_guard_tcb) - apply (rule_tac P="obj_at' (\tcb. tcbFault tcb = Some ft) sender" - in ccorres_cross_over_guard) - apply (ctac(no_vcg) add: setMRs_lookup_failure_ccorres[unfolded msgMaxLength_unfold]) - apply simp - apply (rule ccorres_return_C, simp+)[1] - apply (wp setMR_tcbFault_obj_at hoare_case_option_wp) - apply (clarsimp simp: option_to_ptr_def Collect_const_mem guard_is_UNIV_def) - apply (rule conjI) - apply (simp add: from_bool_def split: bool.split) - apply clarsimp - apply (drule(1) obj_at_cslift_tcb) - apply (clarsimp simp: typ_heap_simps' ctcb_relation_def - cfault_rel_def fault_lift_def Let_def - fault_cap_fault_lift_def is_cap_fault_def - split: split_if_asm) - apply (wp setMR_tcbFault_obj_at hoare_case_option_wp - asUser_inv[OF getRestartPC_inv] - | (rule guard_is_UNIVI, clarsimp simp: option_to_ptr_def))+ - apply (simp add: fault_tag_defs Collect_True Collect_False + apply (wp setMR_tcbFault_obj_at hoare_case_option_wp + asUser_inv[OF getRestartPC_inv] + | (rule guard_is_UNIVI, clarsimp simp: option_to_ptr_def + seL4_CapFault_Addr_def))+ + (* UnknowSyscall *) + apply (rename_tac syscall_number) + apply (simp add: seL4_Fault_tag_defs Collect_True Collect_False ccorres_cond_iffs zipWithM_mapM mapM_append zip_append2 bind_assoc del: Collect_const) @@ -1666,172 +1986,90 @@ proof - apply (simp add: msgMaxLength_unfold n_syscallMessage_def zip_upt_Cons mapM_Cons mapM_Nil mapM_discarded del: Collect_const upt_rec_numeral) - apply (simp only: mapM_x_append[where xs="take (unat n_msgRegisters) (zip as bs)" - and ys="drop (unat n_msgRegisters) (zip as bs)" - for as bs, simplified] bind_assoc) apply (rule ccorres_rhs_assoc)+ - apply csymbr - apply (simp del: Collect_const) - apply (rule ccorres_rhs_assoc2, rule ccorres_split_nothrow_novcg) - apply (simp only: whileAnno_def) - apply (rule_tac F="K $ obj_at' (\tcb. map (tcbContext tcb) ARM_H.syscallMessage = msg) sender" - in ccorres_mapM_x_while) - apply (clarsimp simp: n_msgRegisters_def) - apply (rule ccorres_guard_imp2) - apply (rule_tac t=sender and r="ARM_H.syscallMessage ! n" - in ccorres_add_getRegister) - apply (ctac(no_vcg)) - apply (rule_tac P="\s. rv = msg ! n" in ccorres_cross_over_guard) - apply (simp add: setMR_def length_msgRegisters n_msgRegisters_def - liftM_def[symmetric]) - apply ctac - apply (wp user_getreg_rv) - apply (clarsimp simp: msgRegisters_ccorres n_msgRegisters_def - syscallMessage_ccorres n_syscallMessage_def - unat_of_nat32[unfolded word_bits_def]) - apply (clarsimp simp: obj_at'_def projectKOs word_of_nat_less) - apply (simp add: n_msgRegisters_def) - apply (rule allI, rule conseqPre, vcg exspec=setRegister_modifies exspec=getRegister_modifies) - apply simp - apply (simp add: setMR_def split del: split_if) - apply (rule hoare_pre) - apply (wp asUser_obj_at_elsewhere | wpc)+ - apply simp - apply (simp add: word_bits_def) - apply ceqv - apply (rule ccorres_Cond_rhs) - apply (rule ccorres_rhs_assoc)+ + apply (ctac add: copyMRsFault_ccorres_syscall) + apply (rule ccorres_move_c_guard_tcb) + apply (rule_tac val="syscall_number" in symb_exec_r_fault) + apply (vcg, clarsimp) + apply (drule(1) obj_at_cslift_tcb) + apply (clarsimp simp: ctcb_relation_def typ_heap_simps' + cfault_rel_def seL4_Fault_lift_def Let_def + seL4_Fault_UnknownSyscall_def + seL4_Fault_UnknownSyscall_lift_def + split: split_if_asm) + apply ceqv + apply (ctac add: setMR_ccorres) + apply (rule ccorres_return_C, simp+)[1] + apply wp apply (simp del: Collect_const) - apply (rule ccorres_split_nothrow_novcg) - apply (simp only: whileAnno_def) - apply (rule_tac F="\_. obj_at' (\tcb. map (tcbContext tcb) ARM_H.syscallMessage = msg) - sender and valid_pspace' - and (case buffer of Some x \ valid_ipc_buffer_ptr' x | None \ \)" - in ccorres_mapM_x_while'[where i="unat n_msgRegisters"]) - apply (clarsimp simp: setMR_def n_msgRegisters_def length_msgRegisters - option_to_0_def liftM_def[symmetric] - split: option.split_asm) - apply (rule ccorres_guard_imp2) - apply (rule_tac t=sender and r="ARM_H.syscallMessage ! (n + unat n_msgRegisters)" - in ccorres_add_getRegister) - apply (ctac(no_vcg)) - apply (rule_tac P="\s. rv = msg ! (n + unat n_msgRegisters)" - in ccorres_cross_over_guard) - apply (rule ccorres_move_array_assertion_ipc_buffer - | (rule ccorres_flip_Guard, rule ccorres_move_array_assertion_ipc_buffer))+ - apply (simp add: storeWordUser_def) - apply (rule ccorres_pre_stateAssert) - apply (ctac add: storeWord_ccorres[unfolded fun_app_def]) - apply (simp add: pred_conj_def) - apply (wp user_getreg_rv) - apply (clarsimp simp: n_syscallMessage_def n_msgRegisters_def - syscallMessage_ccorres msgRegisters_ccorres - unat_add_lem[THEN iffD1] unat_of_nat32 - word_bits_def word_size_def) - apply (simp only:field_simps imp_ex imp_conjL) - apply (clarsimp simp: pointerInUserData_c_guard obj_at'_def - pointerInUserData_h_t_valid - projectKOs objBits_simps word_less_nat_alt - unat_add_lem[THEN iffD1] unat_of_nat) - apply (clarsimp simp: pointerInUserData_h_t_valid rf_sr_def - msg_align_bits valid_ipc_buffer_ptr'_def) - apply (erule aligned_add_aligned) - apply (rule aligned_add_aligned[where n=2]) - apply (simp add: is_aligned_def) - apply (rule is_aligned_mult_triv2 [where n=2, simplified]) - apply (simp add: wb_gt_2)+ - apply (simp add: n_msgRegisters_def) - apply (vcg exspec=getRegister_modifies) - apply simp - apply (simp add: setMR_def n_msgRegisters_def length_msgRegisters) - apply (rule hoare_pre) - apply (wp hoare_case_option_wp | wpc)+ - apply clarsimp - apply (simp add: n_msgRegisters_def word_bits_def) - apply ceqv - apply (rule_tac val="unknownSyscallNumber ft" in symb_exec_r_fault) - apply (rule conseqPre, vcg, clarsimp) - apply (drule(1) obj_at_cslift_tcb) - apply (clarsimp simp: typ_heap_simps' ctcb_relation_def - cfault_rel_def fault_lift_def Let_def - fault_unknown_syscall_lift_def - split: split_if_asm) - apply ceqv - apply (simp add: setMR_def msgRegisters_unfold del: Collect_const) - apply wpc - apply (rule ccorres_inst[where P=\ and P'=UNIV], simp add: option_to_0_def) - apply (simp add: bind_assoc storeWordUser_def del: Collect_const) - apply (rule ccorres_pre_stateAssert) - apply (rule ccorres_split_nothrow_novcg) - apply (rule ccorres_move_array_assertion_ipc_buffer - | (rule ccorres_flip_Guard, rule ccorres_move_array_assertion_ipc_buffer))+ - apply (ctac(no_vcg) add: storeWord_ccorres[unfolded fun_app_def]) - apply ceqv - apply ((rule ccorres_return_C ccorres_Guard | simp)+)[1] - apply wp - apply (simp add: msgLengthBits_def mask_def guard_is_UNIV_def) - apply (clarsimp simp: guard_is_UNIV_def option_to_0_def n_msgRegisters_def - word_size_def msg_align_bits) - apply (auto simp add: pointerInUserData_c_guard - pointerInUserData_h_t_valid rf_sr_def)[1] - apply simp - apply (rule_tac Q="\rv. valid_pspace' - and (case buffer of Some x \ valid_ipc_buffer_ptr' x | None \ \) - and ?obj_at_ft" in hoare_post_imp) - apply (clarsimp simp: word_size_def option_to_0_def - pointerInUserData_def valid_ipc_buffer_ptr'_def) - apply (erule aligned_add_aligned) - apply (simp add: is_aligned_def) - apply simp+ - apply (simp add: split_def) - apply (wp mapM_x_wp' setMR_tcbFault_obj_at hoare_case_option_wp | simp)+ - apply (simp add: guard_is_UNIV_def) - apply (subgoal_tac "buffer = None") - apply (simp add: mapM_discarded[symmetric]) - apply (rule_tac P="\a. ccorres r xf Pre Pre' hs (a >>= b) c" for r xf Pre Pre' hs b c in subst) - apply (rule mapM_length_cong[OF refl]) - apply (clarsimp simp: setMR_def length_msgRegisters n_msgRegisters_def - set_zip) - apply (rule refl) - apply (simp add: mapM_return setMR_def msgRegisters_unfold) - apply (rule ccorres_return_C, simp+)[1] - apply (simp add: option_to_0_def split: option.split_asm) - apply simp - apply (wp hoare_drop_imps) - apply (rule mapM_x_wp') - apply (clarsimp simp: setMR_def split del: split_if) - apply (rule hoare_pre) - apply (wp asUser_obj_at_elsewhere hoare_case_option_wp | wpc)+ - apply clarsimp - apply (clarsimp simp: guard_is_UNIV_def n_msgRegisters_def msgLengthBits_def - mask_def) - apply (wp asUser_inv mapM_wp' getRegister_inv hoare_case_option_wp) + apply (vcg exspec=setMR_modifies) + apply (clarsimp simp: option_to_ptr_def guard_is_UNIVI) + apply (wp setMR_tcbFault_obj_at mapM_x_wp_inv hoare_case_option_wp + | clarsimp + | wpc)+ + apply (vcg exspec=copyMRsFault_modifies) + apply (wp asUser_inv mapM_wp' getRegister_inv) apply simp - apply (wp hoare_drop_imps asUser_get_registers - asUser_inv mapM_wp' getRegister_inv - asUser_const_rv) - apply simp - apply (rule guard_is_UNIVI, clarsimp simp: option_to_ptr_def) - apply (clarsimp simp: msgMaxLength_unfold length_syscallMessage + apply (wp asUser_inv mapM_wp' getRegister_inv hoare_drop_imps asUser_const_rv + asUser_get_registers[simplified atcbContextGet_def comp_def]) + apply simp + (* ArchFault *) + apply (rule ccorres_cond_false) + apply (rule ccorres_cond_false) + apply (rule ccorres_cond_false) + apply (simp add: Collect_True Collect_False ccorres_cond_iffs + zip_upt_Cons + mapM_Cons bind_assoc + seL4_Fault_tag_defs + del: Collect_const) + apply (rule ccorres_rhs_assoc)+ + apply (rule ccorres_move_c_guard_tcb) + apply (rule_tac val="fault_to_fault_tag ft" in symb_exec_r_fault) + apply (vcg, clarsimp) + apply (drule(1) obj_at_cslift_tcb) + apply (clarsimp simp: ctcb_relation_def typ_heap_simps' + cfault_rel_def seL4_Fault_lift_def Let_def + seL4_Fault_UnknownSyscall_def + seL4_Fault_UnknownSyscall_lift_def + split: split_if_asm) + apply ceqv + apply (rule ccorres_add_return2) + apply (ctac add: Arch_setMRs_fault_ccorres[simplified setMRs_to_setMR last.simps K_bind_def]) + apply (ctac add: ccorres_return_C) + apply wp + apply (vcg exspec=Arch_setMRs_fault_modifies) + apply (clarsimp simp: guard_is_UNIV_def Collect_const_mem) + apply (rule fault_to_fault_tag.simps(2)[symmetric]) + (* done *) + apply (rule guard_is_UNIVI, clarsimp simp: option_to_ptr_def seL4_CapFault_IP_def) + apply (case_tac x) + apply (simp add: seL4_Faults seL4_Arch_Faults) + apply (clarsimp simp: msgMaxLength_unfold length_syscallMessage ARM_H.exceptionMessage_def ARM.exceptionMessage_def - n_exceptionMessage_def) + n_exceptionMessage_def + n_syscallMessage_def) apply (drule(1) obj_at_cslift_tcb[where thread=sender]) apply (clarsimp simp: obj_at'_def projectKOs objBits_simps) apply (clarsimp simp: msgFromLookupFailure_def split: lookup_failure.split) - apply (clarsimp simp: ctcb_relation_def cfault_rel_def fault_lift_def - Let_def zip_upt_Cons - split: split_if_asm) - done + done qed + +definition + makeArchFaultMessage2 :: "arch_fault \ 32 word" +where + "makeArchFaultMessage2 \ \aft. case aft of VMFault _ _ \ 5" + lemma makeFaultMessage2: "makeFaultMessage ft thread = (do x \ makeFaultMessage ft thread; return (case ft of CapFault _ _ _ \ 1 - | VMFault _ _ \ 2 | UnknownSyscallException _ \ 3 - | UserException _ _ \ 4, snd x) od)" - by (simp add: makeFaultMessage_def split: fault.split) + | UnknownSyscallException _ \ 2 + | UserException _ _ \ 3 + | ArchFault aft \ makeArchFaultMessage2 aft, snd x) od)" + by (simp add: makeFaultMessage_def makeArchFaultMessage2_def makeArchFaultMessage_def + split: fault.split arch_fault.split) lemma doFaultTransfer_ccorres [corres]: "ccorres dc xfdc (valid_pspace' and tcb_at' sender and tcb_at' receiver @@ -1847,6 +2085,7 @@ lemma doFaultTransfer_ccorres [corres]: apply (intro ccorres_gen_asm) apply (simp add: doFaultTransfer_def) apply (subst makeFaultMessage2) + apply (simp only: makeArchFaultMessage2_def) apply (cinit' lift: sender_' receiver_' receiverIPCBuffer_' badge_') apply (rule ccorres_pre_threadGet) apply (rename_tac ft) @@ -1857,15 +2096,15 @@ lemma doFaultTransfer_ccorres [corres]: apply (ctac(no_vcg) add: setMRs_fault_ccorres) apply (rule_tac R="obj_at' (\tcb. tcbFault tcb = ft) sender" and val="case (the ft) of CapFault _ _ _ \ 1 - | VMFault _ _ \ 2 | UnknownSyscallException _ \ 3 - | UserException _ _ \ 4" + | ArchFault (VMFault _ _) \ 5 | UnknownSyscallException _ \ 2 + | UserException _ _ \ 3" in ccorres_symb_exec_r_known_rv_UNIV [where xf'=ret__unsigned_' and R'=UNIV]) apply (rule conseqPre, vcg, clarsimp) apply (drule(1) obj_at_cslift_tcb) apply (clarsimp simp: typ_heap_simps' ctcb_relation_def - cfault_rel_def fault_lift_def Let_def - fault_tag_defs + cfault_rel_def seL4_Fault_lift_def Let_def + seL4_Fault_tag_defs split: split_if_asm) apply ceqv apply csymbr @@ -1876,7 +2115,7 @@ lemma doFaultTransfer_ccorres [corres]: Kernel_C.badgeRegister_def "StrictC'_register_defs") apply (clarsimp simp: message_info_to_H_def guard_is_UNIVI mask_def msgLengthBits_def - split: fault.split) + split: fault.split arch_fault.split) apply wp apply (simp add: setMRs_to_setMR zipWithM_mapM split_def) apply (wp mapM_wp' setMR_tcbFault_obj_at | simp)+ @@ -3107,7 +3346,7 @@ proof - apply (frule lookup_failure_rel_fault_lift, assumption) apply (clarsimp simp: cfault_rel2_def) apply (clarsimp simp: cfault_rel_def) - apply (simp add: fault_cap_fault_lift) + apply (simp add: seL4_Fault_CapFault_lift) apply (clarsimp simp: is_cap_fault_def to_bool_def false_def) apply wp apply (rule hoare_post_imp_R, rule lsft_real_cte) @@ -3379,7 +3618,10 @@ lemma replyFromKernel_error_ccorres [corres]: lemma fault_to_fault_tag_nonzero: "fault_to_fault_tag f \ 0" - by (case_tac f, simp_all add: fault_tag_defs) + apply (case_tac f; simp add: seL4_Fault_tag_defs) + apply (rename_tac af) + apply (case_tac af; simp add: seL4_Fault_tag_defs) + done lemma doIPCTransfer_ccorres [corres]: "ccorres dc xfdc @@ -3401,18 +3643,18 @@ lemma doIPCTransfer_ccorres [corres]: apply (rule ccorres_pre_threadGet) apply (rename_tac fault) apply (rule ccorres_move_c_guard_tcb) - apply (rule_tac val="case_option (scast fault_null_fault) fault_to_fault_tag fault" + apply (rule_tac val="case_option (scast seL4_Fault_NullFault) fault_to_fault_tag fault" and xf'=ret__unsigned_' and R="\s. \t. ko_at' t sender s \ tcbFault t = fault" in ccorres_symb_exec_r_known_rv_UNIV [where R'=UNIV]) apply (vcg, clarsimp) apply (erule(1) cmap_relation_ko_atE [OF cmap_relation_tcb]) apply (fastforce simp: ctcb_relation_def typ_heap_simps - cfault_rel_def fault_lift_def Let_def + cfault_rel_def seL4_Fault_lift_def Let_def split: split_if_asm option.split) apply ceqv apply wpc - apply (clarsimp simp: fault_null_fault_def ccorres_cond_univ_iff) + apply (clarsimp simp: seL4_Fault_NullFault_def ccorres_cond_univ_iff) apply (rule_tac xf'="ptr_coerce \ sendBuffer___ptr_to_void_'" in ccorres_split_nothrow_call_novcg) apply (rule lookupIPCBuffer_ccorres) @@ -3421,7 +3663,7 @@ lemma doIPCTransfer_ccorres [corres]: apply (fold dc_def)[1] apply ctac apply (wp lookupIPCBuffer_not_Some_0 lookupIPCBuffer_aligned) - apply (clarsimp simp: fault_null_fault_def ccorres_cond_iffs + apply (clarsimp simp: seL4_Fault_NullFault_def ccorres_cond_iffs fault_to_fault_tag_nonzero) apply (fold dc_def)[1] apply ctac @@ -3439,21 +3681,362 @@ lemma doIPCTransfer_ccorres [corres]: lemma fault_case_absorb_bind: "(do x \ f; case_fault (p x) (q x) (r x) (s x) ft od) - = case_fault (\a b. f >>= (\x. p x a b)) (\a b. f >>= (\x. q x a b)) - (\a b c. f >>= (\x. r x a b c)) (\a. f >>= (\x. s x a)) ft" + = case_fault (\a b. f >>= (\x. p x a b)) (\a b c. f >>= (\x. q x a b c)) + (\a. f >>= (\x. r x a)) (\a. f >>= (\x. s x a)) ft" by (simp split: fault.split) lemma length_exceptionMessage: "length ARM_H.exceptionMessage = unat n_exceptionMessage" by (simp add: ARM_H.exceptionMessage_def ARM.exceptionMessage_def n_exceptionMessage_def) -declare[[goals_limit = 1]] +lemma copyMRsFaultReply_ccorres_exception: + "ccorres dc xfdc + (valid_pspace' and tcb_at' s + and obj_at' (\t. tcbFault t = Some f) r + and K (r \ s) + and K (len \ unat n_exceptionMessage)) + (UNIV \ \\sender = tcb_ptr_to_ctcb_ptr s\ + \ \\receiver = tcb_ptr_to_ctcb_ptr r\ + \ \\id___anonymous_enum = MessageID_Exception \ + \ \\length___unsigned_long = of_nat len \) hs + (zipWithM_x (\rs rd. do v \ asUser s (getRegister rs); + asUser r (setRegister rd (ARM_H.sanitiseRegister rd v)) + od) + ARM_H.msgRegisters (take len ARM_H.exceptionMessage)) + (Call copyMRsFaultReply_'proc)" +proof - + show ?thesis + apply (unfold K_def, rule ccorres_gen_asm) + apply (cinit' lift: sender_' receiver_' + id___anonymous_enum_' + length___unsigned_long_' + simp: whileAnno_def) + apply (rule ccorres_rhs_assoc2) + apply (simp add: MessageID_Exception_def) + apply ccorres_rewrite + apply (subst bind_return_unit) + apply (rule ccorres_split_nothrow_novcg) + apply (rule ccorres_zipWithM_x_while) + apply clarsimp + apply (intro ccorres_rhs_assoc) + apply (rule ccorres_symb_exec_r) + apply ctac + apply (rule ccorres_symb_exec_r) + apply ctac + apply vcg + apply (rule conseqPre, vcg) + apply fastforce + apply wp + apply vcg + apply vcg + apply (rule conjI, simp add: ARM_H.exceptionMessage_def + ARM.exceptionMessage_def word_of_nat_less) + apply (simp add: msgRegisters_ccorres n_msgRegisters_def length_msgRegisters + unat_of_nat exceptionMessage_ccorres[symmetric,simplified MessageID_Exception_def,simplified] + n_exceptionMessage_def length_exceptionMessage) + apply (simp add: word_less_nat_alt unat_of_nat) + apply (rule conseqPre, vcg) + apply (clarsimp simp: word_of_nat_less ARM_H.exceptionMessage_def + ARM.exceptionMessage_def) + apply (simp add: min_def length_msgRegisters) + apply (clarsimp simp: min_def n_exceptionMessage_def + ARM_H.exceptionMessage_def + ARM.exceptionMessage_def + length_msgRegisters n_msgRegisters_def + message_info_to_H_def word_of_nat_less + split: split_if) + apply (fastforce dest!: le_antisym) + apply clarsimp + apply (vcg spec=TrueI) + apply clarsimp + apply wp + apply simp + apply (clarsimp simp: ARM_H.exceptionMessage_def + ARM.exceptionMessage_def + word_bits_def) + apply unat_arith + apply ceqv + apply (simp add: length_exceptionMessage + length_msgRegisters + n_exceptionMessage_def + msgMaxLength_def + n_msgRegisters_def + of_nat_less_iff) + apply ccorres_rewrite + apply (rule ccorres_return_Skip[simplified dc_def]) + apply (wp mapM_wp') + apply clarsimp+ + apply (clarsimp simp: guard_is_UNIV_def message_info_to_H_def + Collect_const_mem + split: split_if) + apply simp + done +qed + +lemma valid_drop_case: "\ \P\ f \\rv s. P' rv s\ \ + \ \P\ f \\rv s. case rv of None \ True | Some x \ P' rv s\" + apply (simp only: valid_def Ball_def split: prod.split) + apply (rule allI impI)+ + apply (case_tac x1) + apply simp+ + done + +lemma copyMRsFaultReply_ccorres_syscall: + fixes word_size :: "'a::len" + shows "ccorres dc xfdc + (valid_pspace' and tcb_at' s + and obj_at' (\t. tcbFault t = Some f) r + and K (r \ s) + and K (len \ unat n_syscallMessage)) + (UNIV \ \\sender = tcb_ptr_to_ctcb_ptr s\ + \ \\receiver = tcb_ptr_to_ctcb_ptr r\ + \ \\id___anonymous_enum = MessageID_Syscall \ + \ \\length___unsigned_long = of_nat len \) hs + (do a \ zipWithM_x (\rs rd. do v \ asUser s (getRegister rs); + asUser r (setRegister rd (ARM_H.sanitiseRegister rd v)) + od) + ARM_H.msgRegisters (take len ARM_H.syscallMessage); + sendBuf \ lookupIPCBuffer False s; + case sendBuf of None \ return () + | Some bufferPtr \ + zipWithM_x (\i rd. do v \ loadWordUser (bufferPtr + i * 4); + asUser r (setRegister rd (ARM_H.sanitiseRegister rd v)) + od) + [scast n_msgRegisters + 1.e.scast n_syscallMessage] + (drop (unat (scast n_msgRegisters :: word32)) + (take len ARM_H.syscallMessage)) + od) + (Call copyMRsFaultReply_'proc)" + proof - + let ?obj_at_ft = "obj_at' (\tcb. tcbFault tcb = Some f) s" + note symb_exec_r_fault = ccorres_symb_exec_r_known_rv_UNIV + [where xf'=ret__unsigned_' and R="?obj_at_ft" and R'=UNIV] + show ?thesis + apply (unfold K_def, rule ccorres_gen_asm) + apply (cinit' lift: sender_' receiver_' + id___anonymous_enum_' + length___unsigned_long_' + simp: whileAnno_def) + apply (rule ccorres_rhs_assoc2) + apply (simp add: MessageID_Syscall_def) + apply ccorres_rewrite + apply (rule ccorres_split_nothrow_novcg) + apply (rule ccorres_zipWithM_x_while) + apply clarsimp + apply (intro ccorres_rhs_assoc) + apply (rule ccorres_symb_exec_r) + apply ctac + apply (rule ccorres_symb_exec_r) + apply ctac + apply vcg + apply (rule conseqPre, vcg) + apply fastforce + apply wp + apply vcg + apply vcg + apply (rule conjI, simp add: ARM_H.syscallMessage_def + ARM.syscallMessage_def word_of_nat_less + unat_of_nat msgRegisters_ccorres n_msgRegisters_def + length_msgRegisters) + apply (simp add: msgRegisters_ccorres n_msgRegisters_def length_msgRegisters unat_of_nat + syscallMessage_ccorres[symmetric,simplified MessageID_Syscall_def,simplified] + n_syscallMessage_def length_syscallMessage) + apply (simp add: word_less_nat_alt unat_of_nat) + apply (rule conseqPre, vcg) + apply (clarsimp simp: word_of_nat_less syscallMessage_unfold length_msgRegisters + n_syscallMessage_def n_msgRegisters_def) + apply (simp add: min_def length_msgRegisters) + apply (clarsimp simp: min_def n_syscallMessage_def + length_msgRegisters n_msgRegisters_def + length_syscallMessage + message_info_to_H_def word_of_nat_less + split: split_if) + apply (simp add: word_less_nat_alt unat_of_nat not_le) + apply clarsimp + apply (vcg spec=TrueI) + apply clarsimp + apply wp + apply simp + apply (clarsimp simp: length_syscallMessage + length_msgRegisters + n_msgRegisters_def n_syscallMessage_def + word_bits_def min_def + split: split_if) + apply ceqv + apply (rule_tac P'="if 4 < len then _ else _" in ccorres_inst) + apply (cases "4 < len") + apply (clarsimp simp: unat_ucast_less_no_overflow n_syscallMessage_def + length_syscallMessage msgRegisters_unfold + word_of_nat_less unat_of_nat unat_less_helper) + apply ccorres_rewrite + apply (ctac(no_vcg)) + apply (rename_tac sb sb') + apply wpc + apply (simp add: option_to_0_def ccorres_cond_iffs option_to_ptr_def) + apply (rule ccorres_return_Skip') + apply (rule_tac P="sb \ Some 0" in ccorres_gen_asm) + apply (rule_tac P="case_option True (\x. is_aligned x msg_align_bits) sb" + in ccorres_gen_asm) + apply (simp add: option_to_0_def option_to_ptr_def) + apply (subgoal_tac "sb'\ NULL") prefer 2 + apply clarsimp + apply (simp add: ccorres_cond_iffs) + apply (subst ccorres_seq_skip' [symmetric]) + apply (rule_tac r'="\rv rv'. rv' = of_nat (unat n_msgRegisters) + _" in ccorres_rel_imp) + apply (drule_tac s="sb" in sym) + apply (simp only: zipWithM_x_mapM_x) + apply ccorres_rewrite + apply (rule_tac F="\_. valid_pspace' and (case sb of None \ \ + | Some x \ valid_ipc_buffer_ptr' x)" + in ccorres_mapM_x_while') + apply clarsimp + apply (rule ccorres_guard_imp2) + apply (rule ccorres_pre_loadWordUser) + apply (intro ccorres_rhs_assoc) + apply (rule ccorres_symb_exec_r) + apply (rule ccorres_move_array_assertion_ipc_buffer + ccorres_Guard_Seq[where S="{s. h_t_valid (htd s) c_guard (ptr s)}" for ptr htd])+ + apply (rule ccorres_symb_exec_r) + apply (rule ccorres_symb_exec_r) + apply ctac + apply vcg + apply (rule conseqPre, vcg) + apply fastforce + apply vcg + apply (rule conseqPre, vcg) + apply clarsimp + apply vcg + apply (rule conseqPre, vcg) + apply clarsimp + apply clarsimp + apply (subst aligned_add_aligned, assumption) + apply (rule is_aligned_mult_triv2[where n=2, simplified]) + apply (simp add: msg_align_bits) + apply (simp add: of_nat_unat[simplified comp_def]) + apply (simp only: n_msgRegisters_def) + apply (clarsimp simp: n_syscallMessage_def n_msgRegisters_def + word_unat.Rep_inverse[of "scast _ :: 'a word"] + msgRegisters_ccorres[symmetric] + length_msgRegisters[symmetric] + syscallMessage_ccorres[symmetric] + length_msgRegisters length_syscallMessage + syscallMessage_ccorres[symmetric, simplified MessageID_Syscall_def, simplified] + unat_of_nat32 word_bits_def + MessageID_Syscall_def + min_def message_info_to_H_def + upto_enum_def typ_heap_simps' + unat_add_lem[THEN iffD1] + msg_align_bits + simp del: upt_rec_numeral, + simp_all add: word_less_nat_alt unat_add_lem[THEN iffD1] unat_of_nat)[1] + apply (clarsimp simp: n_syscallMessage_def n_msgRegisters_def + msgRegisters_ccorres + syscallMessage_ccorres + length_syscallMessage length_msgRegisters + message_info_to_H_def min_def + split: split_if) + apply (fastforce dest!: le_antisym) + apply (vcg spec=TrueI) + apply clarsimp + apply (simp add: split_def) + apply (wp hoare_case_option_wp) + apply (fastforce elim: aligned_add_aligned + intro: is_aligned_mult_triv2 [where n=2, + simplified] + simp: word_bits_def msg_align_bits) + apply (clarsimp simp: msgRegisters_unfold + n_msgRegisters_def + word_bits_def not_less) + apply (simp add: n_syscallMessage_def) + apply simp + apply (subst option.split[symmetric,where P=id, simplified]) + apply (rule valid_drop_case) + apply (wp lookupIPCBuffer_aligned[simplified K_def] lookupIPCBuffer_not_Some_0[simplified K_def]) + apply (simp add: length_syscallMessage + length_msgRegisters + n_syscallMessage_def + msgMaxLength_def + n_msgRegisters_def + of_nat_less_iff) + apply (rule_tac P'="{s. i_' s = of_nat len}" in ccorres_inst) + apply ccorres_rewrite + apply (rule ccorres_guard_imp) + apply (rule ccorres_symb_exec_l) + apply (case_tac rv ; clarsimp) + apply (rule ccorres_return_Skip[simplified dc_def])+ + apply (wp mapM_x_wp_inv user_getreg_inv' + | clarsimp simp: zipWithM_x_mapM_x split: prod.split)+ + apply (cases "4 < len") + apply (clarsimp simp: guard_is_UNIV_def + msgRegisters_unfold + syscallMessage_unfold + n_syscallMessage_def + n_msgRegisters_def)+ + done +qed + +lemma handleArchFaultReply_corres: + "ccorres (\rv rv'. rv = to_bool rv') ret__unsigned_long_' + (valid_pspace' and tcb_at' sender + and tcb_at' receiver + and K (f = ArchFault af) + and K (sender \ receiver)) + (UNIV \ \ \faultType = fault_to_fault_tag f \ + \ \\sender = tcb_ptr_to_ctcb_ptr sender\ + \ \\receiver = tcb_ptr_to_ctcb_ptr receiver\) + hs + (handleArchFaultReply' af sender receiver msg) + (Call Arch_handleFaultReply_'proc)" + apply (unfold K_def) + apply (rule ccorres_gen_asm)+ + apply (cinit lift: faultType_' sender_' receiver_') + apply (clarsimp simp: bind_assoc seL4_Fault_tag_defs ccorres_cond_iffs + split del: split_if split: arch_fault.split) + apply (rule ccorres_symb_exec_l) + apply (rule ccorres_stateAssert) + apply wpc + apply (clarsimp simp: ccorres_cond_iffs) + apply (rule ccorres_return_C) + apply simp+ + apply (rule ccorres_symb_exec_l) + apply (ctac add: ccorres_return_C) + apply (wp mapM_wp' empty_fail_loadWordUser | clarsimp simp: to_bool_def true_def)+ + done + +(* MOVE *) +lemma monadic_rewrite_ccorres_assemble_nodrop: + assumes cc: "ccorres_underlying sr G r xf ar axf (P and Q) P' hs f c" + assumes mr: "monadic_rewrite True False Q g f" + shows "ccorres_underlying sr G r xf ar axf (P and Q) P' hs g c" +proof - + have snd: "\s. \ Q s; \ snd (g s) \ \ \ snd (f s)" + using mr + by (simp add: monadic_rewrite_def) + + have fst: "\s v. \ Q s; \ snd (g s); v \ fst (f s) \ \ v \ fst (g s)" + using mr + by (auto simp add: monadic_rewrite_def) + + show ?thesis + apply (rule ccorresI') + apply (erule ccorresE[OF cc], (simp add: snd)+) + apply clarsimp + apply (rule rev_bexI[OF fst], assumption+) + apply simp + done +qed + +schematic_goal arch_fault_to_fault_tag_range: + "arch_fault_to_fault_tag x \ ?X" + by (case_tac x, fastforce simp: seL4_Fault_VMFault_def) + lemma handleFaultReply_ccorres [corres]: "ccorres (\rv rv'. rv = to_bool rv') ret__unsigned_long_' - (valid_pspace' and tcb_at' s and - obj_at' (\t. tcbFault t = Some f) r and K (r \ s)) + (valid_pspace' and obj_at' (\t. tcbFault t = Some f) r + and (tcb_at' s and tcb_at' r) + and K (r \ s)) (UNIV \ \cfault_rel (Some f) - (fault_lift (h_val (hrs_mem \t_hrs) + (seL4_Fault_lift (h_val (hrs_mem \t_hrs) (Ptr &(tcb_ptr_to_ctcb_ptr r\[''tcbFault_C''])))) (lookup_fault_lift (h_val (hrs_mem \t_hrs) (Ptr &(tcb_ptr_to_ctcb_ptr r\[''tcbLookupFailure_C'']))))\ @@ -3466,7 +4049,7 @@ lemma handleFaultReply_ccorres [corres]: handleFaultReply f r (msgLabel tag) msg od) (Call handleFaultReply_'proc)" apply (unfold K_def, rule ccorres_gen_asm) - apply (subst handleFaultReply', assumption) + apply (rule monadic_rewrite_ccorres_assemble_nodrop[OF _ handleFaultReply',rotated], simp) apply (cinit lift: sender_' receiver_' simp: whileAnno_def) apply (clarsimp simp del: dc_simp) apply (ctac(c_lines 2) add: getMessageInfo_ccorres') @@ -3478,256 +4061,98 @@ lemma handleFaultReply_ccorres [corres]: apply (rule_tac val="fault_to_fault_tag f" and xf'=ret__unsigned_' and R="\s. \t. ko_at' t r s \ tcbFault t = Some f" - and R'="\cfault_rel (Some f) (fault_lift \fault) + and R'="\cfault_rel (Some f) (seL4_Fault_lift \fault) (lookup_fault_lift (h_val (hrs_mem \t_hrs) (Ptr &(tcb_ptr_to_ctcb_ptr r\[''tcbLookupFailure_C'']))))\" - in ccorres_symb_exec_r_known_rv_UNIV) + in ccorres_symb_exec_r_known_rv) apply (rule conseqPre, vcg, clarsimp) apply (erule(1) cmap_relation_ko_atE [OF cmap_relation_tcb]) apply (clarsimp simp: ctcb_relation_def typ_heap_simps - cfault_rel_def fault_lift_def Let_def + cfault_rel_def seL4_Fault_lift_def Let_def split: split_if_asm) apply ceqv apply (simp add: handleFaultReply_def fault_case_absorb_bind del: Collect_const split del: split_if) apply wpc + (* UserException *) apply (rename_tac number code) - apply (clarsimp simp: bind_assoc fault_tag_defs ccorres_cond_iffs - split del: split_if) - apply (rule ccorres_rhs_assoc)+ - apply csymbr - apply csymbr - apply csymbr - apply (rule ccorres_symb_exec_l) - apply (rule ccorres_stateAssert) - apply (rule ccorres_rhs_assoc2) - apply (rule ccorres_split_nothrow_novcg) - apply (rule ccorres_zipWithM_x_while) - apply clarsimp - apply (intro ccorres_rhs_assoc) - apply (rule ccorres_symb_exec_r) - apply ctac - apply (rule ccorres_symb_exec_r) - apply ctac - apply vcg - apply (rule conseqPre, vcg) - apply fastforce - apply wp - apply vcg - apply vcg - apply (rule conjI, simp add: ARM_H.exceptionMessage_def - ARM.exceptionMessage_def word_of_nat_less) - apply (thin_tac "n < unat n'" for n') - apply (simp add: msgRegisters_ccorres n_msgRegisters_def length_msgRegisters - unat_of_nat exceptionMessage_ccorres[symmetric] - n_exceptionMessage_def length_exceptionMessage) - apply (simp add: word_less_nat_alt unat_of_nat) - apply (rule conseqPre, vcg) - apply (clarsimp simp: word_of_nat_less ARM_H.exceptionMessage_def - ARM.exceptionMessage_def) - apply (clarsimp simp: min_def n_exceptionMessage_def - ARM_H.exceptionMessage_def - ARM.exceptionMessage_def - length_msgRegisters n_msgRegisters_def - message_info_to_H_def - split: split_if) - apply unat_arith - apply clarsimp - apply (vcg spec=TrueI) - apply clarsimp - apply wp - apply simp - apply (clarsimp simp: ARM_H.exceptionMessage_def - ARM.exceptionMessage_def - word_bits_def) - apply unat_arith - apply ceqv - apply (rule ccorres_stateAssert) - apply wpc - apply (simp only: return_bind) - apply (ctac add: ccorres_return_C) - apply (rule ccorres_symb_exec_l) - apply (ctac add: ccorres_return_C) - apply (wp mapM_wp') - apply clarsimp+ - apply (wp hoare_drop_imps) - apply (clarsimp simp: guard_is_UNIV_def message_info_to_H_def - Collect_const_mem - split: split_if) - apply (wp hoare_drop_imps hoare_vcg_all_lift) - apply (clarsimp simp: bind_assoc fault_tag_defs ccorres_cond_iffs - split del: split_if) - apply (rule ccorres_symb_exec_l) - apply (rule ccorres_stateAssert) - apply wpc - apply clarsimp - apply (ctac add: ccorres_return_C) - apply (rule ccorres_symb_exec_l) - apply (ctac add: ccorres_return_C) - apply (wp mapM_wp') - apply clarsimp+ - apply (wp hoare_drop_imps hoare_vcg_all_lift) - apply (clarsimp simp: bind_assoc fault_tag_defs ccorres_cond_iffs - split del: split_if) - apply (rule ccorres_symb_exec_l) - apply (rule ccorres_stateAssert) - apply wpc - apply clarsimp - apply (ctac add: ccorres_return_C) - apply (rule ccorres_symb_exec_l) - apply (ctac add: ccorres_return_C) - apply (wp mapM_wp') - apply clarsimp+ - apply (wp hoare_drop_imps hoare_vcg_all_lift) - apply (rename_tac number) - apply (clarsimp simp: bind_assoc fault_tag_defs ccorres_cond_iffs - split del: split_if) - apply (rule ccorres_rhs_assoc)+ - apply csymbr - apply csymbr - apply csymbr - apply csymbr - apply (ctac(no_vcg)) - apply (rename_tac sb sb') - apply (rule ccorres_stateAssert) - apply (rule ccorres_rhs_assoc2) - apply (rule ccorres_split_nothrow_novcg) - apply (rule ccorres_zipWithM_x_while) - apply clarsimp - apply (intro ccorres_rhs_assoc) - apply (rule ccorres_symb_exec_r) - apply ctac - apply (rule ccorres_symb_exec_r) - apply ctac - apply vcg - apply (rule conseqPre, vcg) - apply fastforce - apply wp - apply vcg - apply vcg - apply (simp add: length_syscallMessage length_msgRegisters) - apply (simp add: msgRegisters_ccorres syscallMessage_ccorres [symmetric] - n_msgRegisters_def unat_of_nat32 word_bits_def - word_of_nat_less n_syscallMessage_def) - apply (rule conseqPre, vcg) - apply (clarsimp simp: word_of_nat_less length_msgRegisters - n_msgRegisters_def) - apply (simp add: length_msgRegisters length_syscallMessage - n_msgRegisters_def n_syscallMessage_def - min_def message_info_to_H_def word_of_nat_less - split: split_if) - apply unat_arith - apply clarsimp - apply (vcg spec=TrueI) - apply clarsimp - apply wp - apply simp - apply (clarsimp simp: msgRegisters_unfold word_bits_def) - apply unat_arith - apply ceqv - apply (rule ccorres_stateAssert) - apply (rule ccorres_split_nothrow_novcg) - apply wpc - apply (simp add: option_to_0_def ccorres_cond_iffs option_to_ptr_def) - apply (rule ccorres_return_Skip) - apply (rule_tac P="sb \ Some 0" in ccorres_gen_asm) - apply (rule_tac P="case_option True (\x. is_aligned x msg_align_bits) sb" - in ccorres_gen_asm) - apply (simp add: option_to_0_def option_to_ptr_def) - apply (subgoal_tac "sb'\ NULL") prefer 2 - apply clarsimp - apply (simp add: ccorres_cond_iffs) - apply (subst ccorres_seq_skip' [symmetric]) - apply (rule ccorres_split_nothrow_novcg) - apply (drule_tac s="sb" in sym) - apply (simp only: zipWithM_x_mapM_x) - apply (rule_tac F="\_. valid_pspace' and (case sb of None \ \ - | Some x \ valid_ipc_buffer_ptr' x)" - in ccorres_mapM_x_while') - apply clarsimp - apply (rule ccorres_guard_imp2) - apply (rule ccorres_pre_loadWordUser) - apply (intro ccorres_rhs_assoc) - apply (rule ccorres_symb_exec_r) - apply (rule ccorres_move_array_assertion_ipc_buffer - ccorres_Guard_Seq[where S="{s. h_t_valid (htd s) c_guard (ptr s)}" for ptr htd] - )+ - apply (rule ccorres_symb_exec_r) - apply (rule ccorres_symb_exec_r) - apply ctac - apply vcg - apply (rule conseqPre, vcg) - apply fastforce - apply vcg - apply (rule conseqPre, vcg) - apply clarsimp - apply vcg - apply (rule conseqPre, vcg) - apply clarsimp - apply clarsimp - apply (subst aligned_add_aligned, assumption) - apply (rule is_aligned_mult_triv2[where n=2, simplified]) - apply (simp add: msg_align_bits) - apply (clarsimp simp: n_syscallMessage_def n_msgRegisters_def - msgRegisters_ccorres[symmetric] - syscallMessage_ccorres[symmetric] - length_msgRegisters length_syscallMessage - unat_of_nat32 word_bits_def - min_def message_info_to_H_def - upto_enum_def typ_heap_simps' - unat_add_lem[THEN iffD1] - msg_align_bits - split: split_if_asm - simp del: upt_rec_numeral, - simp_all add: word_less_nat_alt unat_add_lem[THEN iffD1] unat_of_nat)[1] - apply (clarsimp simp: n_syscallMessage_def n_msgRegisters_def - msgRegisters_ccorres - syscallMessage_ccorres - length_syscallMessage length_msgRegisters - message_info_to_H_def min_def - split: split_if) - apply (elim disjE,unat_arith+)[1] - apply (vcg spec=TrueI) - apply clarsimp - apply (simp add: split_def) - apply (wp hoare_case_option_wp) - apply (fastforce elim: aligned_add_aligned - intro: is_aligned_mult_triv2 [where n=2, - simplified] - simp: word_bits_def msg_align_bits) - apply (clarsimp simp: msgRegisters_unfold word_bits_def) - apply (simp add: n_syscallMessage_def) - apply ceqv - apply (rule ccorres_symb_exec_l) - apply (fold dc_def)[1] - apply (rule ccorres_return_Skip) - apply (wp mapM_wp') - apply clarsimp - apply wp - apply (clarsimp simp: guard_is_UNIV_def) - apply ceqv + apply (clarsimp simp: bind_assoc seL4_Fault_tag_defs ccorres_cond_iffs + split del: split_if) + apply (subst take_min_len[symmetric,where n="unat (msgLength _)"]) + apply (ctac add: copyMRsFaultReply_ccorres_exception) apply (ctac add: ccorres_return_C) apply wp - apply (clarsimp simp: guard_is_UNIV_def message_info_to_H_def - split: split_if) - apply clarsimp - apply (rule_tac Q="\rv. valid_pspace' and - (case_option \ valid_ipc_buffer_ptr' sb)" - in hoare_post_imp) - apply (clarsimp simp: upto_enum_word valid_ipc_buffer_ptr'_def simp del: upt.simps) - apply (simp add: zipWithM_x_mapM_x) - apply (wp mapM_x_wp hoare_case_option_wp | simp add: split_def)+ - apply (rule subset_refl) - apply (clarsimp simp: guard_is_UNIV_def) - apply (wp lookupIPCBuffer_not_Some_0 lookupIPCBuffer_aligned - | wp_once hoare_drop_imps)+ - apply (clarsimp simp: guard_is_UNIV_def to_bool_def true_def false_def) + apply (vcg exspec=copyMRsFaultReply_modifies) + (* CapFault *) + apply (clarsimp simp: bind_assoc seL4_Fault_tag_defs ccorres_cond_iffs + split del: split_if) + apply (ctac add: ccorres_return_C) + (* UnknowSyscall *) + apply (rename_tac number) + apply (clarsimp simp: seL4_Fault_tag_defs ccorres_cond_iffs + split del: split_if) + apply (subst take_min_len[symmetric,where n="unat (msgLength _)"]) + apply (subst take_min_len[symmetric,where n="unat (msgLength _)"]) + apply (fold bind_assoc) + apply (ctac add: copyMRsFaultReply_ccorres_syscall[simplified bind_assoc[symmetric]]) + apply (ctac add: ccorres_return_C) + apply wp + apply (vcg exspec=copyMRsFaultReply_modifies) + (* ArchFault *) + apply (rule ccorres_cond_false) + apply (rule ccorres_cond_false) + apply (rule ccorres_cond_false) + apply (clarsimp simp: ccorres_cond_iffs + split del: split_if) + apply (rule ccorres_rhs_assoc) + (* apply (rule_tac P="\s. \t. ko_at' t r s" in ccorres_cross_over_guard) *) + apply (rule_tac val="fault_to_fault_tag f" + and xf'=ret__unsigned_' + and R="\s. \t. ko_at' t r s \ tcbFault t = Some f" + and R'="\cfault_rel (Some f) (seL4_Fault_lift \fault) + (lookup_fault_lift (h_val (hrs_mem \t_hrs) + (Ptr &(tcb_ptr_to_ctcb_ptr r\[''tcbLookupFailure_C'']))))\" + in ccorres_symb_exec_r_known_rv_UNIV) + apply (rule conseqPre, vcg, clarsimp) + apply (erule(1) cmap_relation_ko_atE [OF cmap_relation_tcb]) + apply (clarsimp simp: ctcb_relation_def typ_heap_simps + cfault_rel_def seL4_Fault_lift_def Let_def + split: split_if_asm) + apply ceqv + apply (rule ccorres_add_return2) + apply (ctac add: handleArchFaultReply_corres) + apply (rule ccorres_return_C) + apply simp+ + apply wp + apply (vcg exspec=Arch_handleFaultReply_modifies) + apply (clarsimp simp: guard_is_UNIV_def) + apply (subst fault_to_fault_tag.simps(2)) + apply (clarsimp split: split_if) + apply simp+ + (* Done *) + apply clarsimp + apply vcg apply vcg - apply (rule conseqPre, vcg) apply clarsimp + apply vcg_step + apply (clarsimp simp: n_exceptionMessage_def n_syscallMessage_def + message_info_to_H_def to_bool_def scast_def + length_exceptionMessage length_syscallMessage + min_def word_less_nat_alt true_def + guard_is_UNIV_def seL4_Faults seL4_Arch_Faults + split: split_if) + apply (simp add: length_exceptionMessage length_syscallMessage) apply wp - apply vcg - apply (clarsimp simp: obj_at'_def n_msgRegisters_def) + apply clarsimp + apply (vcg exspec=getRegister_modifies) + apply (clarsimp simp: n_exceptionMessage_def n_syscallMessage_def + message_info_to_H_def to_bool_def scast_def + length_exceptionMessage length_syscallMessage + min_def word_less_nat_alt true_def + obj_at'_def + split: split_if) + using arch_fault_to_fault_tag_range + apply (fastforce simp: seL4_Faults seL4_Arch_Faults) done (* FIXME: move *) @@ -3931,14 +4356,14 @@ proof - apply (rule ccorres_pre_threadGet) apply (rename_tac fault) apply (rule ccorres_move_c_guard_tcb) - apply (rule_tac val="case_option (scast fault_null_fault) fault_to_fault_tag fault" + apply (rule_tac val="case_option (scast seL4_Fault_NullFault) fault_to_fault_tag fault" and xf'=ret__unsigned_' and R="\s. \t. ko_at' t receiver s \ tcbFault t = fault" in ccorres_symb_exec_r_known_rv_UNIV [where R'=UNIV]) apply (vcg, clarsimp) apply (erule(1) cmap_relation_ko_atE [OF cmap_relation_tcb]) apply (fastforce simp: ctcb_relation_def typ_heap_simps - cfault_rel_def fault_lift_def Let_def + cfault_rel_def seL4_Fault_lift_def Let_def split: split_if_asm option.split) apply ceqv apply wpc @@ -3964,7 +4389,7 @@ proof - invs_valid_queues_strg) apply (simp add: cap_reply_cap_def) apply (wp doIPCTransfer_reply_or_replyslot) - apply (clarsimp simp: fault_null_fault_def ccorres_cond_iffs + apply (clarsimp simp: seL4_Fault_NullFault_def ccorres_cond_iffs fault_to_fault_tag_nonzero split del: split_if) apply (rule ccorres_rhs_assoc)+ @@ -3989,7 +4414,7 @@ proof - (simp add: typ_heap_simps')+, simp_all?)[1] apply (rule ball_tcb_cte_casesI, simp+) apply (clarsimp simp: ctcb_relation_def - fault_lift_null_fault + seL4_Fault_lift_NullFault cfault_rel_def is_cap_fault_def cthread_state_relation_def) apply (case_tac "tcbState tcb", simp_all add: is_cap_fault_def)[1] @@ -4746,8 +5171,8 @@ lemma sendIPC_ccorres [corres]: apply (rule ccorres_symb_exec_r) apply (rule ccorres_cond_false_seq) apply (intro ccorres_rhs_assoc) - apply (rule ccorres_move_c_guard_tcb [simplified]) - apply (rule_tac val="case_option (scast fault_null_fault) + apply (rule ccorres_move_c_guard_tcb [simplified]) + apply (rule_tac val="case_option (scast seL4_Fault_NullFault) fault_to_fault_tag fault" and xf'=ret__unsigned_' and R="\s. \t. ko_at' t thread s \ tcbFault t = fault" @@ -4755,7 +5180,7 @@ lemma sendIPC_ccorres [corres]: apply (rule conseqPre, vcg, clarsimp) apply (erule(1) cmap_relation_ko_atE[OF cmap_relation_tcb]) apply (fastforce simp: ctcb_relation_def typ_heap_simps - cfault_rel_def fault_lift_def Let_def + cfault_rel_def seL4_Fault_lift_def Let_def split: split_if_asm option.split) apply ceqv apply wpc @@ -4777,7 +5202,7 @@ lemma sendIPC_ccorres [corres]: apply (ctac add: setThreadState_ccorres) apply vcg apply (rule conseqPre, vcg, clarsimp) - apply (clarsimp simp: fault_to_fault_tag_nonzero fault_null_fault_def + apply (clarsimp simp: fault_to_fault_tag_nonzero seL4_Fault_NullFault_def guard_is_UNIV_def ThreadState_Inactive_def) apply vcg apply (rule conseqPre, vcg, clarsimp) @@ -5548,7 +5973,7 @@ lemma receiveIPC_ccorres [corres]: apply (rule ccorres_cond_false_seq) apply (intro ccorres_rhs_assoc) apply (rule ccorres_move_c_guard_tcb [simplified]) - apply (rule_tac val="case_option (scast fault_null_fault) + apply (rule_tac val="case_option (scast seL4_Fault_NullFault) fault_to_fault_tag fault" and xf'=ret__unsigned_' and R="\s. \t. ko_at' t sender s \ tcbFault t = fault" @@ -5556,7 +5981,7 @@ lemma receiveIPC_ccorres [corres]: apply (rule conseqPre, vcg, clarsimp) apply (erule(1) cmap_relation_ko_atE [OF cmap_relation_tcb]) apply (auto simp: ctcb_relation_def typ_heap_simps - cfault_rel_def fault_lift_def Let_def + cfault_rel_def seL4_Fault_lift_def Let_def split: split_if_asm option.split)[1] apply ceqv apply wpc @@ -5580,7 +6005,7 @@ lemma receiveIPC_ccorres [corres]: apply ctac apply vcg apply (rule conseqPre, vcg, clarsimp) - apply (clarsimp simp: fault_to_fault_tag_nonzero fault_null_fault_def + apply (clarsimp simp: fault_to_fault_tag_nonzero seL4_Fault_NullFault_def guard_is_UNIV_def ThreadState_Inactive_def mask_def ThreadState_Running_def) apply vcg diff --git a/proof/crefine/IsolatedThreadAction.thy b/proof/crefine/IsolatedThreadAction.thy new file mode 100644 index 000000000..1f8b71b18 --- /dev/null +++ b/proof/crefine/IsolatedThreadAction.thy @@ -0,0 +1,1758 @@ +(* + * 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 IsolatedThreadAction +imports "../../lib/clib/MonadicRewrite_C" Finalise_C CSpace_All SyscallArgs_C +begin + +datatype tcb_state_regs = TCBStateRegs "thread_state" "MachineTypes.register \ machine_word" + +definition + "tsrContext tsr \ case tsr of TCBStateRegs ts regs \ regs" + +definition + "tsrState tsr \ case tsr of TCBStateRegs ts regs \ ts" + +lemma accessors_TCBStateRegs[simp]: + "TCBStateRegs (tsrState v) (tsrContext v) = v" + by (cases v, simp add: tsrState_def tsrContext_def) + +lemma tsrContext_simp[simp]: + "tsrContext (TCBStateRegs st con) = con" + by (simp add: tsrContext_def) + +lemma tsrState_simp[simp]: + "tsrState (TCBStateRegs st con) = st" + by (simp add: tsrState_def) + +definition + get_tcb_state_regs :: "kernel_object option \ tcb_state_regs" +where + "get_tcb_state_regs oko \ case oko of + Some (KOTCB tcb) \ TCBStateRegs (tcbState tcb) ((atcbContextGet o tcbArch) tcb)" + +definition + put_tcb_state_regs_tcb :: "tcb_state_regs \ tcb \ tcb" +where + "put_tcb_state_regs_tcb tsr tcb \ case tsr of + TCBStateRegs st regs \ tcb \ tcbState := st, tcbArch := atcbContextSet regs (tcbArch tcb) \" + +definition + put_tcb_state_regs :: "tcb_state_regs \ kernel_object option \ kernel_object option" +where + "put_tcb_state_regs tsr oko = Some (KOTCB (put_tcb_state_regs_tcb tsr + (case oko of + Some (KOTCB tcb) \ tcb | _ \ makeObject)))" + +definition + "partial_overwrite idx tcbs ps \ + \x. if x \ range idx + then put_tcb_state_regs (tcbs (inv idx x)) (ps x) + else ps x" + +definition + isolate_thread_actions :: "('x \ word32) \ 'a kernel + \ (('x \ tcb_state_regs) \ ('x \ tcb_state_regs)) + \ (scheduler_action \ scheduler_action) + \ 'a kernel" + where + "isolate_thread_actions idx m t f \ do + s \ gets (ksSchedulerAction_update (\_. ResumeCurrentThread) + o ksPSpace_update (partial_overwrite idx (K undefined))); + tcbs \ gets (\s. get_tcb_state_regs o ksPSpace s o idx); + sa \ getSchedulerAction; + (rv, s') \ select_f (m s); + modify (\s. ksPSpace_update (partial_overwrite idx (t tcbs)) + (s' \ ksSchedulerAction := f sa \)); + return rv + od" + +lemma put_tcb_state_regs_twice[simp]: + "put_tcb_state_regs tsr (put_tcb_state_regs tsr' tcb) + = put_tcb_state_regs tsr tcb" + apply (simp add: put_tcb_state_regs_def put_tcb_state_regs_tcb_def + makeObject_tcb + split: tcb_state_regs.split option.split + Structures_H.kernel_object.split) + apply (intro all_tcbI impI allI) + apply simp + done + +lemma partial_overwrite_twice[simp]: + "partial_overwrite idx f (partial_overwrite idx g ps) + = partial_overwrite idx f ps" + by (rule ext, simp add: partial_overwrite_def) + +lemma get_tcb_state_regs_partial_overwrite[simp]: + "inj idx \ + get_tcb_state_regs (partial_overwrite idx tcbs f (idx x)) + = tcbs x" + apply (simp add: partial_overwrite_def) + apply (simp add: put_tcb_state_regs_def + get_tcb_state_regs_def + put_tcb_state_regs_tcb_def + split: tcb_state_regs.split) + done + +lemma isolate_thread_actions_bind: + "inj idx \ + isolate_thread_actions idx a b c >>= + (\x. isolate_thread_actions idx (d x) e f) + = isolate_thread_actions idx a id id + >>= (\x. isolate_thread_actions idx (d x) (e o b) (f o c))" + apply (rule ext) + apply (clarsimp simp: isolate_thread_actions_def bind_assoc split_def + bind_select_f_bind[symmetric]) + apply (clarsimp simp: exec_gets getSchedulerAction_def) + apply (rule select_bind_eq) + apply (simp add: exec_gets exec_modify o_def) + apply (rule select_bind_eq) + apply (simp add: exec_gets exec_modify) + done + +lemma setEndpoint_obj_at_tcb': + "\obj_at' (P :: tcb \ bool) p\ setEndpoint p' val \\rv. obj_at' P p\" + apply (simp add: setEndpoint_def) + apply (rule obj_at_setObject2) + apply (clarsimp simp: updateObject_default_def in_monad) + done + +lemma tcbSchedEnqueue_obj_at_unchangedT: + assumes y: "\f. \tcb. P (tcbQueued_update f tcb) = P tcb" + shows "\obj_at' P t\ tcbSchedEnqueue t' \\rv. obj_at' P t\" + apply (simp add: tcbSchedEnqueue_def unless_def) + apply (wp | simp add: y)+ + done + +lemma asUser_obj_at_notQ: + "\obj_at' (Not \ tcbQueued) t\ + asUser t (setRegister r v) + \\rv. obj_at' (Not \ tcbQueued) t\" + apply (simp add: asUser_def) + apply (rule hoare_seq_ext)+ + apply (simp add: split_def) + apply (rule threadSet_obj_at'_really_strongest) + apply (wp threadGet_wp |rule gets_inv|wpc|clarsimp)+ + apply (clarsimp simp: obj_at'_def) + done + +(* FIXME: Move to Schedule_R.thy. Make Arch_switchToThread_obj_at a specialisation of this *) +lemma Arch_switchToThread_obj_at_pre: + "\obj_at' (Not \ tcbQueued) t\ + Arch.switchToThread t + \\rv. obj_at' (Not \ tcbQueued) t\" + apply (simp add: ARM_H.switchToThread_def) + apply (wp asUser_obj_at_notQ doMachineOp_obj_at setVMRoot_obj_at hoare_drop_imps|wpc)+ + done + +lemma rescheduleRequired_obj_at_unchangedT: + assumes y: "\f. \tcb. P (tcbQueued_update f tcb) = P tcb" + shows "\obj_at' P t\ rescheduleRequired \\rv. obj_at' P t\" + apply (simp add: rescheduleRequired_def) + apply (wp tcbSchedEnqueue_obj_at_unchangedT[OF y] | wpc)+ + apply simp + done + +lemma setThreadState_obj_at_unchangedT: + assumes x: "\f. \tcb. P (tcbState_update f tcb) = P tcb" + assumes y: "\f. \tcb. P (tcbQueued_update f tcb) = P tcb" + shows "\obj_at' P t\ setThreadState t' ts \\rv. obj_at' P t\" + apply (simp add: setThreadState_def) + apply (wp rescheduleRequired_obj_at_unchangedT[OF y], simp) + apply (wp threadSet_obj_at'_strongish) + apply (clarsimp simp: obj_at'_def projectKOs x cong: if_cong) + done + +lemma setBoundNotification_obj_at_unchangedT: + assumes x: "\f. \tcb. P (tcbBoundNotification_update f tcb) = P tcb" + shows "\obj_at' P t\ setBoundNotification t' ts \\rv. obj_at' P t\" + apply (simp add: setBoundNotification_def) + apply (wp threadSet_obj_at'_strongish) + apply (clarsimp simp: obj_at'_def projectKOs x cong: if_cong) + done + +lemmas setThreadState_obj_at_unchanged + = setThreadState_obj_at_unchangedT[OF all_tcbI all_tcbI] + +lemmas setBoundNotification_obj_at_unchanged + = setBoundNotification_obj_at_unchangedT[OF all_tcbI] + +lemma setNotification_tcb: + "\obj_at' (\tcb::tcb. P tcb) t\ + setNotification ntfn e + \\_. obj_at' P t\" + apply (simp add: setNotification_def) + apply (rule obj_at_setObject2) + apply (clarsimp simp: updateObject_default_def in_monad) + done + +(* FIXME: move *) +lemmas threadSet_obj_at' = threadSet_obj_at'_strongish + +context kernel_m begin +lemma setObject_modify: + fixes v :: "'a :: pspace_storable" shows + "\ obj_at' (P :: 'a \ bool) p s; updateObject v = updateObject_default v; + (1 :: word32) < 2 ^ objBits v \ + \ setObject p v s + = modify (ksPSpace_update (\ps. ps (p \ injectKO v))) s" + apply (clarsimp simp: setObject_def split_def exec_gets + obj_at'_def projectKOs lookupAround2_known1 + assert_opt_def updateObject_default_def + bind_assoc) + apply (simp add: projectKO_def alignCheck_assert) + apply (simp add: project_inject objBits_def) + apply (clarsimp simp only: objBitsT_koTypeOf[symmetric] koTypeOf_injectKO) + apply (frule(2) in_magnitude_check[where s'=s]) + apply (simp add: magnitudeCheck_assert in_monad) + apply (simp add: simpler_modify_def) + done + +context begin interpretation Arch . (*FIXME: arch_split*) + +lemma getObject_return: + fixes v :: "'a :: pspace_storable" shows + "\ \a b c d. (loadObject a b c d :: 'a kernel) = loadObject_default a b c d; + ko_at' v p s; (1 :: word32) < 2 ^ objBits v \ \ getObject p s = return v s" + apply (clarsimp simp: getObject_def split_def exec_gets + obj_at'_def projectKOs lookupAround2_known1 + assert_opt_def loadObject_default_def) + apply (simp add: projectKO_def alignCheck_assert) + apply (simp add: project_inject objBits_def) + apply (frule(2) in_magnitude_check[where s'=s]) + apply (simp add: magnitudeCheck_assert in_monad) + done + +end + +lemmas getObject_return_tcb + = getObject_return[OF meta_eq_to_obj_eq, OF loadObject_tcb, + unfolded objBits_simps, simplified] + +lemmas setObject_modify_tcb + = setObject_modify[OF _ meta_eq_to_obj_eq, OF _ updateObject_tcb, + unfolded objBits_simps, simplified] + +lemma partial_overwrite_fun_upd: + "inj idx \ + partial_overwrite idx (tsrs (x := y)) + = (\ps. (partial_overwrite idx tsrs ps) (idx x := put_tcb_state_regs y (ps (idx x))))" + apply (intro ext, simp add: partial_overwrite_def) + apply (clarsimp split: split_if) + done + +lemma get_tcb_state_regs_ko_at': + "ko_at' ko p s \ get_tcb_state_regs (ksPSpace s p) + = TCBStateRegs (tcbState ko) ((atcbContextGet o tcbArch) ko)" + by (clarsimp simp: obj_at'_def projectKOs get_tcb_state_regs_def) + +lemma put_tcb_state_regs_ko_at': + "ko_at' ko p s \ put_tcb_state_regs tsr (ksPSpace s p) + = Some (KOTCB (ko \ tcbState := tsrState tsr + , tcbArch := atcbContextSet (tsrContext tsr) (tcbArch ko)\))" + by (clarsimp simp: obj_at'_def projectKOs put_tcb_state_regs_def + put_tcb_state_regs_tcb_def + split: tcb_state_regs.split) + +lemma partial_overwrite_get_tcb_state_regs: + "\ \x. tcb_at' (idx x) s; inj idx \ \ + partial_overwrite idx (\x. get_tcb_state_regs (ksPSpace s (idx x))) + (ksPSpace s) = ksPSpace s" + apply (rule ext, simp add: partial_overwrite_def + split: split_if) + apply clarsimp + apply (drule_tac x=xa in spec) + apply (clarsimp simp: obj_at'_def projectKOs put_tcb_state_regs_def + get_tcb_state_regs_def put_tcb_state_regs_tcb_def) + apply (case_tac obj, simp) + done + +lemma ksPSpace_update_partial_id: + "\ \ps x. f ps x = ps (idx x) \ f ps x = ksPSpace s (idx x); + \x. tcb_at' (idx x) s; inj idx \ \ + ksPSpace_update (\ps. partial_overwrite idx (\x. get_tcb_state_regs (f ps x)) ps) s + = s" + apply (rule trans, rule kernel_state.fold_congs[OF refl refl]) + apply (erule_tac x="ksPSpace s" in meta_allE) + apply (clarsimp simp: partial_overwrite_get_tcb_state_regs) + apply (rule refl) + apply simp + done + +lemma isolate_thread_actions_asUser: + "\ idx t' = t; inj idx; f = (\s. ({(v, g s)}, False)) \ \ + monadic_rewrite False True (\s. \x. tcb_at' (idx x) s) + (asUser t f) + (isolate_thread_actions idx (return v) + (\tsrs. (tsrs (t' := TCBStateRegs (tsrState (tsrs t')) + (g (tsrContext (tsrs t')))))) + id)" + apply (simp add: asUser_def liftM_def isolate_thread_actions_def split_def + select_f_returns bind_assoc select_f_singleton_return + threadGet_def threadSet_def) + apply (clarsimp simp: monadic_rewrite_def) + apply (frule_tac x=t' in spec) + apply (drule obj_at_ko_at', clarsimp) + apply (simp add: exec_gets getSchedulerAction_def exec_modify + getObject_return_tcb setObject_modify_tcb o_def + cong: bind_apply_cong)+ + apply (simp add: partial_overwrite_fun_upd return_def get_tcb_state_regs_ko_at') + apply (rule kernel_state.fold_congs[OF refl refl]) + apply (clarsimp simp: partial_overwrite_get_tcb_state_regs + put_tcb_state_regs_ko_at') + apply (case_tac ko, simp) + done + +lemma getRegister_simple: + "getRegister r = (\con. ({(con r, con)}, False))" + by (simp add: getRegister_def simpler_gets_def) + +lemma mapM_getRegister_simple: + "mapM getRegister rs = (\con. ({(map con rs, con)}, False))" + apply (induct rs) + apply (simp add: mapM_Nil return_def) + apply (simp add: mapM_Cons getRegister_def simpler_gets_def + bind_def return_def) + done + +lemma setRegister_simple: + "setRegister r v = (\con. ({((), con (r := v))}, False))" + by (simp add: setRegister_def simpler_modify_def) + +lemma zipWithM_setRegister_simple: + "zipWithM_x setRegister rs vs + = (\con. ({((), foldl (\con (r, v). con (r := v)) con (zip rs vs))}, False))" + apply (simp add: zipWithM_x_mapM_x) + apply (induct ("zip rs vs")) + apply (simp add: mapM_x_Nil return_def) + apply (clarsimp simp add: mapM_x_Cons bind_def setRegister_def + simpler_modify_def fun_upd_def[symmetric]) + done + +lemma dom_partial_overwrite: + "\x. tcb_at' (idx x) s \ dom (partial_overwrite idx tsrs (ksPSpace s)) + = dom (ksPSpace s)" + apply (rule set_eqI) + apply (clarsimp simp: dom_def partial_overwrite_def put_tcb_state_regs_def + split: split_if) + apply (fastforce elim!: obj_atE') + done + +lemma map_to_ctes_partial_overwrite: + "\x. tcb_at' (idx x) s \ + map_to_ctes (partial_overwrite idx tsrs (ksPSpace s)) + = ctes_of s" + apply (rule ext) + apply (frule dom_partial_overwrite[where tsrs=tsrs]) + apply (simp add: map_to_ctes_def partial_overwrite_def + Let_def) + apply (case_tac "x \ range idx") + apply (clarsimp simp: put_tcb_state_regs_def) + apply (drule_tac x=xa in spec) + apply (clarsimp simp: obj_at'_def projectKOs objBits_simps + cong: if_cong) + apply (simp add: put_tcb_state_regs_def put_tcb_state_regs_tcb_def + objBits_simps + cong: if_cong option.case_cong) + apply (case_tac obj, simp split: tcb_state_regs.split split_if) + apply simp + apply (rule if_cong[OF refl]) + apply simp + apply (case_tac "x && ~~ mask (objBitsKO (KOTCB undefined)) \ range idx") + apply (clarsimp simp: put_tcb_state_regs_def) + apply (drule_tac x=xa in spec) + apply (clarsimp simp: obj_at'_def projectKOs objBits_simps + cong: if_cong) + apply (simp add: put_tcb_state_regs_def put_tcb_state_regs_tcb_def + objBits_simps + cong: if_cong option.case_cong) + apply (case_tac obj, simp split: tcb_state_regs.split split_if) + apply (intro impI allI) + apply (subgoal_tac "x - idx xa = x && mask 9") + apply (clarsimp simp: tcb_cte_cases_def split: split_if) + apply (drule_tac t = "idx xa" in sym) + apply simp + apply (simp cong: if_cong) + done + +definition + "thread_actions_isolatable idx f = + (inj idx \ monadic_rewrite False True (\s. \x. tcb_at' (idx x) s) + f (isolate_thread_actions idx f id id))" + +lemma getCTE_assert_opt: + "getCTE p = gets (\s. ctes_of s p) >>= assert_opt" + apply (intro ext) + apply (simp add: exec_gets assert_opt_def prod_eq_iff + fail_def return_def + split: option.split) + apply (rule conjI) + apply clarsimp + apply (rule context_conjI) + apply (rule ccontr, clarsimp elim!: nonemptyE) + apply (frule use_valid[OF _ getCTE_sp], rule TrueI) + apply (frule in_inv_by_hoareD[OF getCTE_inv]) + apply (clarsimp simp: cte_wp_at_ctes_of) + apply (simp add: empty_failD[OF empty_fail_getCTE]) + apply clarsimp + apply (simp add: no_failD[OF no_fail_getCTE, OF ctes_of_cte_at]) + apply (subgoal_tac "cte_wp_at' (op = x2) p x") + apply (clarsimp simp: cte_wp_at'_def getCTE_def) + apply (simp add: cte_wp_at_ctes_of) + done + +lemma getCTE_isolatable: + "thread_actions_isolatable idx (getCTE p)" + apply (clarsimp simp: thread_actions_isolatable_def) + apply (simp add: isolate_thread_actions_def bind_assoc split_def) + apply (simp add: getCTE_assert_opt bind_select_f_bind[symmetric] + bind_assoc select_f_returns) + apply (clarsimp simp: monadic_rewrite_def exec_gets getSchedulerAction_def + map_to_ctes_partial_overwrite) + apply (simp add: assert_opt_def select_f_returns select_f_asserts + split: option.split) + apply (clarsimp simp: exec_modify o_def return_def) + apply (simp add: ksPSpace_update_partial_id) + done + +lemma objBits_2n: + "(1 :: word32) < 2 ^ objBits obj" + by (simp add: objBits_def objBitsKO_def archObjSize_def pageBits_def + split: kernel_object.split arch_kernel_object.split) + +lemma magnitudeCheck_assert2: + "\ is_aligned x n; (1 :: word32) < 2 ^ n; ksPSpace s x = Some v \ \ + magnitudeCheck x (snd (lookupAround2 x (ksPSpace (s :: kernel_state)))) n + = assert (ps_clear x n s)" + using in_magnitude_check[where x=x and n=n and s=s and s'=s and v="()"] + by (simp add: magnitudeCheck_assert in_monad) + +lemma getObject_get_assert: + assumes deflt: "\a b c d. (loadObject a b c d :: ('a :: pspace_storable) kernel) + = loadObject_default a b c d" + shows + "(getObject p :: ('a :: pspace_storable) kernel) + = do v \ gets (obj_at' (\x :: 'a. True) p); + assert v; + gets (the o projectKO_opt o the o swp fun_app p o ksPSpace) + od" + apply (rule ext) + apply (simp add: exec_get getObject_def split_def exec_gets + deflt loadObject_default_def projectKO_def2 + alignCheck_assert) + apply (case_tac "ksPSpace x p") + apply (simp add: obj_at'_def assert_opt_def assert_def + split: option.split split_if) + apply (simp add: lookupAround2_known1 assert_opt_def + obj_at'_def projectKO_def2 + split: option.split) + apply (clarsimp simp: fail_def fst_return conj_comms project_inject + objBits_def) + apply (simp only: assert2[symmetric], + rule bind_apply_cong[OF refl]) + apply (clarsimp simp: in_monad) + apply (fold objBits_def) + apply (simp add: magnitudeCheck_assert2[OF _ objBits_2n]) + apply (rule bind_apply_cong[OF refl]) + apply (clarsimp simp: in_monad return_def simpler_gets_def) + apply (simp add: iffD2[OF project_inject refl]) + done + +lemma obj_at_partial_overwrite_If: + "\ \x. tcb_at' (idx x) s \ + \ obj_at' P p (ksPSpace_update (partial_overwrite idx f) s) + = (if p \ range idx + then obj_at' (\tcb. P (put_tcb_state_regs_tcb (f (inv idx p)) tcb)) p s + else obj_at' P p s)" + apply (frule dom_partial_overwrite[where tsrs=f]) + apply (simp add: obj_at'_def ps_clear_def partial_overwrite_def + projectKOs split: split_if) + apply clarsimp + apply (drule_tac x=x in spec) + apply (clarsimp simp: put_tcb_state_regs_def objBits_simps) + done + +lemma obj_at_partial_overwrite_id1: + "\ p \ range idx; \x. tcb_at' (idx x) s \ + \ obj_at' P p (ksPSpace_update (partial_overwrite idx f) s) + = obj_at' P p s" + apply (drule dom_partial_overwrite[where tsrs=f]) + apply (simp add: obj_at'_def ps_clear_def partial_overwrite_def + projectKOs) + done + +lemma obj_at_partial_overwrite_id2: + "\ \x. tcb_at' (idx x) s; \v tcb. P v \ True \ injectKO v \ KOTCB tcb \ + \ obj_at' P p (ksPSpace_update (partial_overwrite idx f) s) + = obj_at' P p s" + apply (frule dom_partial_overwrite[where tsrs=f]) + apply (simp add: obj_at'_def ps_clear_def partial_overwrite_def + projectKOs split: split_if) + apply clarsimp + apply (drule_tac x=x in spec) + apply (clarsimp simp: put_tcb_state_regs_def objBits_simps + project_inject) + done + +lemma getObject_isolatable: + "\ \a b c d. (loadObject a b c d :: 'a kernel) = loadObject_default a b c d; + \tcb. projectKO_opt (KOTCB tcb) = (None :: 'a option) \ \ + thread_actions_isolatable idx (getObject p :: ('a :: pspace_storable) kernel)" + apply (clarsimp simp: thread_actions_isolatable_def) + apply (simp add: getObject_get_assert split_def + isolate_thread_actions_def bind_select_f_bind[symmetric] + bind_assoc select_f_asserts select_f_returns) + apply (clarsimp simp: monadic_rewrite_def exec_gets getSchedulerAction_def) + apply (case_tac "p \ range idx") + apply clarsimp + apply (drule_tac x=x in spec) + apply (clarsimp simp: obj_at'_def projectKOs partial_overwrite_def + put_tcb_state_regs_def) + apply (simp add: obj_at_partial_overwrite_id1) + apply (simp add: partial_overwrite_def) + apply (rule bind_apply_cong[OF refl]) + apply (simp add: exec_modify return_def o_def simpler_gets_def + ksPSpace_update_partial_id in_monad) + done + +lemma gets_isolatable: + "\\g s. \x. tcb_at' (idx x) s \ + f (ksSchedulerAction_update g + (ksPSpace_update (partial_overwrite idx (\_. undefined)) s)) = f s \ \ + thread_actions_isolatable idx (gets f)" + apply (clarsimp simp: thread_actions_isolatable_def) + apply (simp add: isolate_thread_actions_def select_f_returns + liftM_def bind_assoc) + apply (clarsimp simp: monadic_rewrite_def exec_gets + getSchedulerAction_def exec_modify) + apply (simp add: simpler_gets_def return_def + ksPSpace_update_partial_id o_def) + done + +lemma modify_isolatable: + assumes swap:"\tsrs act s. \x. tcb_at' (idx x) s \ + (ksPSpace_update (partial_overwrite idx tsrs) ((f s)\ ksSchedulerAction := act \)) + = f (ksPSpace_update (partial_overwrite idx tsrs) + (s \ ksSchedulerAction := act\))" + shows + "thread_actions_isolatable idx (modify f)" + apply (clarsimp simp: thread_actions_isolatable_def) + apply (simp add: isolate_thread_actions_def select_f_returns + liftM_def bind_assoc) + apply (clarsimp simp: monadic_rewrite_def exec_gets + getSchedulerAction_def) + apply (simp add: simpler_modify_def o_def) + apply (subst swap) + apply (simp add: obj_at_partial_overwrite_If) + apply (simp add: ksPSpace_update_partial_id o_def) + done + +lemma isolate_thread_actions_wrap_bind: + "inj idx \ + do x \ isolate_thread_actions idx a b c; + isolate_thread_actions idx (d x) e f + od = + isolate_thread_actions idx + (do x \ isolate_thread_actions idx a id id; + isolate_thread_actions idx (d x) id id + od) (e o b) (f o c) + " + apply (rule ext) + apply (clarsimp simp: isolate_thread_actions_def bind_assoc split_def + bind_select_f_bind[symmetric] liftM_def + select_f_returns select_f_selects + getSchedulerAction_def) + apply (clarsimp simp: exec_gets getSchedulerAction_def o_def) + apply (rule select_bind_eq) + apply (simp add: exec_gets exec_modify o_def) + apply (rule select_bind_eq) + apply (simp add: exec_modify) + done + +lemma monadic_rewrite_in_isolate_thread_actions: + "\ inj idx; monadic_rewrite F True P a d \ \ + monadic_rewrite F True (\s. P (ksSchedulerAction_update (\_. ResumeCurrentThread) + (ksPSpace_update (partial_overwrite idx (\_. undefined)) s))) + (isolate_thread_actions idx a b c) (isolate_thread_actions idx d b c)" + apply (clarsimp simp: isolate_thread_actions_def split_def) + apply (rule monadic_rewrite_bind_tail)+ + apply (rule_tac P="\_. P s" in monadic_rewrite_bind_head) + apply (simp add: monadic_rewrite_def select_f_def) + apply wp + apply simp + done + +lemma thread_actions_isolatable_bind: + "\ thread_actions_isolatable idx f; \x. thread_actions_isolatable idx (g x); + \t. \tcb_at' t\ f \\rv. tcb_at' t\ \ + \ thread_actions_isolatable idx (f >>= g)" + apply (clarsimp simp: thread_actions_isolatable_def) + apply (rule monadic_rewrite_imp) + apply (rule monadic_rewrite_trans) + apply (erule monadic_rewrite_bind2, assumption) + apply (rule hoare_vcg_all_lift, assumption) + apply (subst isolate_thread_actions_wrap_bind, simp) + apply simp + apply (rule monadic_rewrite_in_isolate_thread_actions, assumption) + apply (rule monadic_rewrite_transverse) + apply (erule monadic_rewrite_bind2, assumption) + apply (rule hoare_vcg_all_lift, assumption) + apply (simp add: bind_assoc id_def) + apply (rule monadic_rewrite_refl) + apply (simp add: obj_at_partial_overwrite_If) + done + +lemma thread_actions_isolatable_return: + "thread_actions_isolatable idx (return v)" + apply (clarsimp simp: thread_actions_isolatable_def + monadic_rewrite_def liftM_def + isolate_thread_actions_def + split_def bind_assoc select_f_returns + exec_gets getSchedulerAction_def) + apply (simp add: exec_modify return_def o_def + ksPSpace_update_partial_id) + done + +lemma thread_actions_isolatable_fail: + "thread_actions_isolatable idx fail" + by (simp add: thread_actions_isolatable_def + isolate_thread_actions_def select_f_asserts + liftM_def bind_assoc getSchedulerAction_def + monadic_rewrite_def exec_gets) + +lemma thread_actions_isolatable_returns: + "thread_actions_isolatable idx (return v)" + "thread_actions_isolatable idx (returnOk v)" + "thread_actions_isolatable idx (throwError v)" + by (simp add: returnOk_def throwError_def + thread_actions_isolatable_return)+ + +lemma thread_actions_isolatable_bindE: + "\ thread_actions_isolatable idx f; \x. thread_actions_isolatable idx (g x); + \t. \tcb_at' t\ f \\rv. tcb_at' t\ \ + \ thread_actions_isolatable idx (f >>=E g)" + apply (simp add: bindE_def) + apply (erule thread_actions_isolatable_bind) + apply (simp add: lift_def thread_actions_isolatable_returns + split: sum.split) + apply assumption + done + +lemma thread_actions_isolatable_catch: + "\ thread_actions_isolatable idx f; \x. thread_actions_isolatable idx (g x); + \t. \tcb_at' t\ f \\rv. tcb_at' t\ \ + \ thread_actions_isolatable idx (f g)" + apply (simp add: catch_def) + apply (erule thread_actions_isolatable_bind) + apply (simp add: thread_actions_isolatable_returns + split: sum.split) + apply assumption + done + +lemma thread_actions_isolatable_if: + "\ P \ thread_actions_isolatable idx a; + \ P \ thread_actions_isolatable idx b \ + \ thread_actions_isolatable idx (if P then a else b)" + by (cases P, simp_all) + +lemma select_f_isolatable: + "thread_actions_isolatable idx (select_f v)" + apply (clarsimp simp: thread_actions_isolatable_def + isolate_thread_actions_def + split_def select_f_selects liftM_def bind_assoc) + apply (rule monadic_rewrite_imp, rule monadic_rewrite_transverse) + apply (rule monadic_rewrite_drop_modify monadic_rewrite_bind_tail)+ + apply wp + apply (simp add: gets_bind_ign getSchedulerAction_def) + apply (rule monadic_rewrite_refl) + apply (simp add: ksPSpace_update_partial_id o_def) + done + +lemma doMachineOp_isolatable: + "thread_actions_isolatable idx (doMachineOp m)" + apply (simp add: doMachineOp_def split_def) + apply (intro thread_actions_isolatable_bind[OF _ _ hoare_pre(1)] + gets_isolatable thread_actions_isolatable_returns + modify_isolatable select_f_isolatable) + apply (simp | wp)+ + done + +lemma page_directory_at_partial_overwrite: + "\x. tcb_at' (idx x) s \ + page_directory_at' p (ksPSpace_update (partial_overwrite idx f) s) + = page_directory_at' p s" + by (simp add: page_directory_at'_def typ_at_to_obj_at_arches + obj_at_partial_overwrite_id2) + +lemma findPDForASID_isolatable: + "thread_actions_isolatable idx (findPDForASID asid)" + apply (simp add: findPDForASID_def liftE_bindE liftME_def bindE_assoc + case_option_If2 assertE_def liftE_def checkPDAt_def + stateAssert_def2 + cong: if_cong) + apply (intro thread_actions_isolatable_bind[OF _ _ hoare_pre(1)] + thread_actions_isolatable_bindE[OF _ _ hoare_pre(1)] + thread_actions_isolatable_if thread_actions_isolatable_returns + thread_actions_isolatable_fail + gets_isolatable getObject_isolatable) + apply (simp add: projectKO_opt_asidpool page_directory_at_partial_overwrite + | wp getASID_wp)+ + done + +lemma getHWASID_isolatable: + "thread_actions_isolatable idx (getHWASID asid)" + apply (simp add: getHWASID_def loadHWASID_def + findFreeHWASID_def + case_option_If2 findPDForASIDAssert_def + checkPDAt_def checkPDUniqueToASID_def + checkPDASIDMapMembership_def + stateAssert_def2 const_def assert_def + findFreeHWASID_def + invalidateASID_def + invalidateHWASIDEntry_def + storeHWASID_def + cong: if_cong) + apply (intro thread_actions_isolatable_bind[OF _ _ hoare_pre(1)] + thread_actions_isolatable_bindE[OF _ _ hoare_pre(1)] + thread_actions_isolatable_catch[OF _ _ hoare_pre(1)] + thread_actions_isolatable_if thread_actions_isolatable_returns + thread_actions_isolatable_fail + gets_isolatable modify_isolatable + findPDForASID_isolatable doMachineOp_isolatable) + apply (wp hoare_drop_imps + | simp add: page_directory_at_partial_overwrite)+ + done + +lemma setVMRoot_isolatable: + "thread_actions_isolatable idx (setVMRoot t)" + apply (simp add: setVMRoot_def getThreadVSpaceRoot_def + locateSlot_conv getSlotCap_def + cap_case_isPageDirectoryCap if_bool_simps + whenE_def liftE_def + checkPDNotInASIDMap_def stateAssert_def2 + checkPDASIDMapMembership_def armv_contextSwitch_def + cong: if_cong) + apply (intro thread_actions_isolatable_bind[OF _ _ hoare_pre(1)] + thread_actions_isolatable_bindE[OF _ _ hoare_pre(1)] + thread_actions_isolatable_catch[OF _ _ hoare_pre(1)] + thread_actions_isolatable_if thread_actions_isolatable_returns + thread_actions_isolatable_fail + gets_isolatable getCTE_isolatable getHWASID_isolatable + findPDForASID_isolatable doMachineOp_isolatable) + apply (simp add: projectKO_opt_asidpool + | wp getASID_wp typ_at_lifts [OF getHWASID_typ_at'])+ + done + + +lemma transferCaps_simple: + "transferCaps mi [] ep receiver rcvrBuf = + do + getReceiveSlots receiver rcvrBuf; + return (mi\msgExtraCaps := 0, msgCapsUnwrapped := 0\) + od" + apply (cases mi) + apply (clarsimp simp: transferCaps_def getThreadCSpaceRoot_def locateSlot_conv) + apply (rule ext bind_apply_cong[OF refl])+ + apply (simp add: upto_enum_def + split: option.split) + done + +lemma transferCaps_simple_rewrite: + "monadic_rewrite True True ((\_. caps = []) and \) + (transferCaps mi caps ep r rBuf) + (return (mi \ msgExtraCaps := 0, msgCapsUnwrapped := 0 \))" + apply (rule monadic_rewrite_gen_asm) + apply (rule monadic_rewrite_imp) + apply (rule monadic_rewrite_trans) + apply (simp add: transferCaps_simple, rule monadic_rewrite_refl) + apply (rule monadic_rewrite_symb_exec2, wp empty_fail_getReceiveSlots) + apply (rule monadic_rewrite_refl) + apply simp + done + +lemma lookupExtraCaps_simple_rewrite: + "msgExtraCaps mi = 0 \ + (lookupExtraCaps thread rcvBuf mi = returnOk [])" + by (cases mi, simp add: lookupExtraCaps_def getExtraCPtrs_def + liftE_bindE upto_enum_step_def mapM_Nil + split: option.split) + +lemma lookupIPC_inv: "\P\ lookupIPCBuffer f t \\rv. P\" + by wp + +lemmas empty_fail_user_getreg = empty_fail_asUser[OF empty_fail_getRegister] + +lemma user_getreg_rv: + "\obj_at' (\tcb. P ((atcbContextGet o tcbArch) tcb r)) t\ asUser t (getRegister r) \\rv s. P rv\" + apply (simp add: asUser_def split_def) + apply (wp threadGet_wp) + apply (clarsimp simp: obj_at'_def projectKOs getRegister_def in_monad atcbContextGet_def) + done + +lemma copyMRs_simple: + "msglen \ of_nat (length ARM_H.msgRegisters) \ + copyMRs sender sbuf receiver rbuf msglen + = forM_x (take (unat msglen) ARM_H.msgRegisters) + (\r. do v \ asUser sender (getRegister r); + asUser receiver (setRegister r v) od) + >>= (\rv. return msglen)" + apply (clarsimp simp: copyMRs_def mapM_discarded) + apply (rule bind_cong[OF refl]) + apply (simp add: length_msgRegisters n_msgRegisters_def min_def + word_le_nat_alt + split: option.split) + apply (simp add: upto_enum_def mapM_Nil) + done + +lemma doIPCTransfer_simple_rewrite: + "monadic_rewrite True True + ((\_. msgExtraCaps (messageInfoFromWord msgInfo) = 0 + \ msgLength (messageInfoFromWord msgInfo) + \ of_nat (length ARM_H.msgRegisters)) + and obj_at' (\tcb. tcbFault tcb = None + \ (atcbContextGet o tcbArch) tcb msgInfoRegister = msgInfo) sender) + (doIPCTransfer sender ep badge True rcvr) + (do rv \ mapM_x (\r. do v \ asUser sender (getRegister r); + asUser rcvr (setRegister r v) + od) + (take (unat (msgLength (messageInfoFromWord msgInfo))) ARM_H.msgRegisters); + y \ setMessageInfo rcvr ((messageInfoFromWord msgInfo) \msgCapsUnwrapped := 0\); + asUser rcvr (setRegister ARM_H.badgeRegister badge) + od)" + apply (rule monadic_rewrite_gen_asm) + apply (simp add: doIPCTransfer_def bind_assoc doNormalTransfer_def + getMessageInfo_def + cong: option.case_cong) + apply (rule monadic_rewrite_imp) + apply (rule monadic_rewrite_trans) + apply (rule monadic_rewrite_bind_tail)+ + apply (rule_tac P="fault = None" in monadic_rewrite_gen_asm, simp) + apply (rule monadic_rewrite_bind_tail) + apply (rule_tac x=msgInfo in monadic_rewrite_symb_exec, + wp empty_fail_user_getreg user_getreg_rv) + apply (simp add: lookupExtraCaps_simple_rewrite returnOk_catch_bind) + apply (rule monadic_rewrite_bind) + apply (rule monadic_rewrite_from_simple, rule copyMRs_simple) + apply (rule monadic_rewrite_bind_head) + apply (rule transferCaps_simple_rewrite) + apply (wp threadGet_const) + apply (simp add: bind_assoc) + apply (rule monadic_rewrite_symb_exec2[OF lookupIPC_inv empty_fail_lookupIPCBuffer] + monadic_rewrite_symb_exec2[OF threadGet_inv empty_fail_threadGet] + monadic_rewrite_symb_exec2[OF user_getreg_inv' empty_fail_user_getreg] + monadic_rewrite_bind_head monadic_rewrite_bind_tail + | wp)+ + apply (case_tac "messageInfoFromWord msgInfo") + apply simp + apply (rule monadic_rewrite_refl) + apply wp + apply clarsimp + apply (auto elim!: obj_at'_weakenE) + done + +lemma monadic_rewrite_setSchedulerAction_noop: + "monadic_rewrite F E (\s. ksSchedulerAction s = act) (setSchedulerAction act) (return ())" + unfolding setSchedulerAction_def + apply (rule monadic_rewrite_imp, rule monadic_rewrite_modify_noop) + apply simp + done + +lemma rescheduleRequired_simple_rewrite: + "monadic_rewrite F E + (sch_act_simple) + rescheduleRequired + (setSchedulerAction ChooseNewThread)" + apply (simp add: rescheduleRequired_def getSchedulerAction_def) + apply (simp add: monadic_rewrite_def exec_gets sch_act_simple_def) + apply auto + done + +lemma empty_fail_isRunnable: + "empty_fail (isRunnable t)" + by (simp add: isRunnable_def isBlocked_def) + +lemma setThreadState_blocked_rewrite: + "\ runnable' st \ + monadic_rewrite True True + (\s. ksCurThread s = t \ ksSchedulerAction s \ ResumeCurrentThread \ tcb_at' t s) + (setThreadState st t) + (threadSet (tcbState_update (\_. st)) t)" + apply (simp add: setThreadState_def) + apply (rule monadic_rewrite_imp) + apply (rule monadic_rewrite_trans) + apply (rule monadic_rewrite_bind_tail) + apply (rule monadic_rewrite_trans) + apply (rule monadic_rewrite_bind_tail)+ + apply (rule_tac P="\ runnable \ curThread = t + \ (action \ ResumeCurrentThread)" + in monadic_rewrite_gen_asm) + apply (simp add: when_def) + apply (rule monadic_rewrite_refl) + apply wp + apply (rule monadic_rewrite_symb_exec2, + (wp empty_fail_isRunnable + | (simp only: getCurThread_def getSchedulerAction_def + , rule empty_fail_gets))+)+ + apply (rule monadic_rewrite_refl) + apply (simp add: conj_comms, wp) + apply (rule_tac Q="\rv s. obj_at' (Not o runnable' o tcbState) t s" + in hoare_post_imp) + apply (clarsimp simp: obj_at'_def sch_act_simple_def st_tcb_at'_def) + apply (wp) + apply simp + apply (rule monadic_rewrite_refl) + apply clarsimp + done + +lemma setupCallerCap_rewrite: + "monadic_rewrite True True (\s. reply_masters_rvk_fb (ctes_of s)) + (setupCallerCap send rcv) + (do setThreadState BlockedOnReply send; + replySlot \ getThreadReplySlot send; + callerSlot \ getThreadCallerSlot rcv; + replySlotCTE \ getCTE replySlot; + assert (mdbNext (cteMDBNode replySlotCTE) = 0 + \ isReplyCap (cteCap replySlotCTE) + \ capReplyMaster (cteCap replySlotCTE) + \ mdbFirstBadged (cteMDBNode replySlotCTE) + \ mdbRevocable (cteMDBNode replySlotCTE)); + cteInsert (ReplyCap send False) replySlot callerSlot + od)" + apply (simp add: setupCallerCap_def getThreadCallerSlot_def + getThreadReplySlot_def locateSlot_conv + getSlotCap_def) + apply (rule monadic_rewrite_imp) + apply (rule monadic_rewrite_bind_tail)+ + apply (rule monadic_rewrite_assert)+ + apply (rule_tac P="mdbFirstBadged (cteMDBNode masterCTE) + \ mdbRevocable (cteMDBNode masterCTE)" + in monadic_rewrite_gen_asm) + apply simp + apply (rule monadic_rewrite_trans) + apply (rule monadic_rewrite_bind_tail) + apply (rule monadic_rewrite_symb_exec2, (wp | simp)+)+ + apply (rule monadic_rewrite_refl) + apply wp + apply (rule monadic_rewrite_symb_exec2, wp empty_fail_getCTE)+ + apply (rule monadic_rewrite_refl) + apply (wp getCTE_wp' | simp add: cte_wp_at_ctes_of)+ + apply (clarsimp simp: reply_masters_rvk_fb_def) + apply fastforce + done + +lemma attemptSwitchTo_rewrite: + "monadic_rewrite True True + (\s. obj_at' (\tcb. tcbPriority tcb = curPrio) thread s + \ obj_at' (\tcb. tcbPriority tcb = destPrio \ tcbDomain tcb = destDom) t s + \ destPrio \ curPrio + \ ksSchedulerAction s = ResumeCurrentThread + \ ksCurThread s = thread + \ ksCurDomain s = curDom + \ destDom = curDom) + (attemptSwitchTo t) (setSchedulerAction (SwitchToThread t))" + apply (simp add: attemptSwitchTo_def possibleSwitchTo_def) + apply (rule monadic_rewrite_imp) + apply (rule monadic_rewrite_trans) + apply (rule monadic_rewrite_bind_tail) + apply (rule monadic_rewrite_bind_tail) + apply (rule monadic_rewrite_bind_tail) + apply (rule monadic_rewrite_bind_tail) + apply (rule monadic_rewrite_bind_tail) + apply (rule monadic_rewrite_bind_tail) + apply (rule_tac P="curPrio \ targetPrio \ action = ResumeCurrentThread + \ targetDom = curDom" + in monadic_rewrite_gen_asm) + apply (simp add: eq_commute le_less[symmetric]) + apply (rule monadic_rewrite_refl) + apply (wp threadGet_wp) + apply (rule monadic_rewrite_trans) + apply (rule monadic_rewrite_bind_tail) + apply (rule monadic_rewrite_symb_exec2, + wp empty_fail_threadGet | simp add: getSchedulerAction_def curDomain_def)+ + apply (rule monadic_rewrite_refl) + apply wp + apply (rule monadic_rewrite_symb_exec2, simp_all add: getCurThread_def) + apply (rule monadic_rewrite_refl) + apply (auto simp: obj_at'_def) + done + +lemma oblivious_getObject_ksPSpace_default: + "\ \s. ksPSpace (f s) = ksPSpace s; + \a b c ko. (loadObject a b c ko :: 'a kernel) \ loadObject_default a b c ko \ \ + oblivious f (getObject p :: ('a :: pspace_storable) kernel)" + apply (simp add: getObject_def split_def loadObject_default_def + projectKO_def2 alignCheck_assert magnitudeCheck_assert) + apply (intro oblivious_bind, simp_all) + done + +lemmas oblivious_getObject_ksPSpace_tcb[simp] + = oblivious_getObject_ksPSpace_default[OF _ loadObject_tcb] + +lemma oblivious_setObject_ksPSpace_tcb[simp]: + "\ \s. ksPSpace (f s) = ksPSpace s; + \s g. ksPSpace_update g (f s) = f (ksPSpace_update g s) \ \ + oblivious f (setObject p (v :: tcb))" + apply (simp add: setObject_def split_def updateObject_default_def + projectKO_def2 alignCheck_assert magnitudeCheck_assert) + apply (intro oblivious_bind, simp_all) + done + +lemma oblivious_getObject_ksPSpace_cte[simp]: + "\ \s. ksPSpace (f s) = ksPSpace s \ \ + oblivious f (getObject p :: cte kernel)" + apply (simp add: getObject_def split_def loadObject_cte + projectKO_def2 alignCheck_assert magnitudeCheck_assert + typeError_def unless_when + cong: Structures_H.kernel_object.case_cong) + apply (intro oblivious_bind, + simp_all split: Structures_H.kernel_object.split split_if) + by (safe intro!: oblivious_bind, simp_all) + +lemma oblivious_doMachineOp[simp]: + "\ \s. ksMachineState (f s) = ksMachineState s; + \g s. ksMachineState_update g (f s) = f (ksMachineState_update g s) \ + \ oblivious f (doMachineOp oper)" + apply (simp add: doMachineOp_def split_def) + apply (intro oblivious_bind, simp_all) + done + +lemmas oblivious_getObject_ksPSpace_asidpool[simp] + = oblivious_getObject_ksPSpace_default[OF _ loadObject_asidpool] + +lemma oblivious_setVMRoot_schact: + "oblivious (ksSchedulerAction_update f) (setVMRoot t)" + apply (simp add: setVMRoot_def getThreadVSpaceRoot_def locateSlot_conv + getSlotCap_def getCTE_def armv_contextSwitch_def) + by (safe intro!: oblivious_bind oblivious_bindE oblivious_catch + | simp_all add: liftE_def getHWASID_def + findPDForASID_def liftME_def loadHWASID_def + findPDForASIDAssert_def checkPDAt_def + checkPDUniqueToASID_def + checkPDASIDMapMembership_def + findFreeHWASID_def invalidateASID_def + invalidateHWASIDEntry_def storeHWASID_def + checkPDNotInASIDMap_def armv_contextSwitch_def + split: capability.split arch_capability.split option.split)+ + +lemma oblivious_switchToThread_schact: + "oblivious (ksSchedulerAction_update f) (ThreadDecls_H.switchToThread t)" + apply (simp add: Thread_H.switchToThread_def ARM_H.switchToThread_def bind_assoc + getCurThread_def setCurThread_def threadGet_def liftM_def + threadSet_def tcbSchedEnqueue_def unless_when asUser_def + getQueue_def setQueue_def storeWordUser_def setRegister_def + pointerInUserData_def isRunnable_def isBlocked_def + getThreadState_def tcbSchedDequeue_def bitmap_fun_defs) + by (safe intro!: oblivious_bind + | simp_all add: oblivious_setVMRoot_schact)+ + +lemma schedule_rewrite: + notes hoare_TrueI[simp] + shows "monadic_rewrite True True + (\s. ksSchedulerAction s = SwitchToThread t \ ct_in_state' (op = Running) s) + (schedule) + (do curThread \ getCurThread; tcbSchedEnqueue curThread; setSchedulerAction ResumeCurrentThread; switchToThread t od)" + apply (simp add: schedule_def) + apply (rule monadic_rewrite_imp) + apply (rule monadic_rewrite_trans) + apply (rule monadic_rewrite_bind_tail) + apply (rule monadic_rewrite_bind_tail) + apply (rule_tac P="action = SwitchToThread t" in monadic_rewrite_gen_asm, simp) + apply (rule monadic_rewrite_bind_tail) + apply (rule_tac P="curRunnable \ action = SwitchToThread t" in monadic_rewrite_gen_asm, simp) + apply (rule monadic_rewrite_refl) + apply (wp,simp,wp) + apply (rule monadic_rewrite_trans) + apply (rule monadic_rewrite_bind_tail) + apply (rule monadic_rewrite_symb_exec2, wp | simp add: isRunnable_def getSchedulerAction_def)+ + apply (rule monadic_rewrite_refl) + apply (wp) + apply (simp add: setSchedulerAction_def) + apply (subst oblivious_modify_swap[symmetric], rule oblivious_switchToThread_schact) + apply (rule monadic_rewrite_refl) + apply (clarsimp simp: st_tcb_at'_def pred_neg_def o_def obj_at'_def ct_in_state'_def) + done + +lemma empty_fail_getCurThread[iff]: + "empty_fail getCurThread" by (simp add: getCurThread_def) + +lemma schedule_rewrite_ct_not_runnable': + "monadic_rewrite True True + (\s. ksSchedulerAction s = SwitchToThread t \ ct_in_state' (Not \ runnable') s) + (schedule) + (do setSchedulerAction ResumeCurrentThread; switchToThread t od)" + apply (simp add: schedule_def) + apply (rule monadic_rewrite_imp) + apply (rule monadic_rewrite_trans) + apply (rule monadic_rewrite_bind_tail) + apply (rule monadic_rewrite_bind_tail) + apply (rule_tac P="action = SwitchToThread t" in monadic_rewrite_gen_asm, simp) + apply (rule monadic_rewrite_bind_tail) + apply (rule_tac P="\ curRunnable \ action = SwitchToThread t" in monadic_rewrite_gen_asm,simp) + apply (rule monadic_rewrite_refl) + apply (wp,simp,wp) + apply (rule monadic_rewrite_trans) + apply (rule monadic_rewrite_bind_tail) + apply (rule monadic_rewrite_symb_exec2, wp | + simp add: isRunnable_def getSchedulerAction_def | + rule hoare_TrueI)+ + apply (rule monadic_rewrite_refl) + apply (wp) + apply (simp add: setSchedulerAction_def) + apply (subst oblivious_modify_swap[symmetric], rule oblivious_switchToThread_schact) + apply (rule monadic_rewrite_symb_exec2) + apply (wp, simp, rule hoare_TrueI) + apply (rule monadic_rewrite_refl) + apply (clarsimp simp: st_tcb_at'_def pred_neg_def o_def obj_at'_def ct_in_state'_def) + done + +lemma activateThread_simple_rewrite: + "monadic_rewrite True True (ct_in_state' (op = Running)) + (activateThread) (return ())" + apply (simp add: activateThread_def) + apply (rule monadic_rewrite_imp) + apply (rule monadic_rewrite_trans, rule monadic_rewrite_bind_tail)+ + apply (rule_tac P="state = Running" in monadic_rewrite_gen_asm) + apply simp + apply (rule monadic_rewrite_refl) + apply wp + apply (rule monadic_rewrite_symb_exec2, wp empty_fail_getThreadState) + apply (rule monadic_rewrite_refl) + apply wp + apply (rule monadic_rewrite_symb_exec2, + simp_all add: getCurThread_def) + apply (rule monadic_rewrite_refl) + apply (clarsimp simp: ct_in_state'_def elim!: pred_tcb'_weakenE) + done + +end + +lemma setCTE_obj_at_prio[wp]: + "\obj_at' (\tcb. P (tcbPriority tcb)) t\ setCTE p v \\rv. obj_at' (\tcb. P (tcbPriority tcb)) t\" + unfolding setCTE_def + by (rule setObject_cte_obj_at_tcb', simp+) + +crunch obj_at_prio[wp]: cteInsert "obj_at' (\tcb. P (tcbPriority tcb)) t" + (wp: crunch_wps) + +crunch ctes_of[wp]: asUser "\s. P (ctes_of s)" + (wp: crunch_wps) + +lemma tcbSchedEnqueue_tcbPriority[wp]: + "\obj_at' (\tcb. P (tcbPriority tcb)) t\ + tcbSchedEnqueue t' + \\rv. obj_at' (\tcb. P (tcbPriority tcb)) t\" + apply (simp add: tcbSchedEnqueue_def unless_def) + apply (wp | simp cong: if_cong)+ + done + +crunch obj_at_prio[wp]: cteDeleteOne "obj_at' (\tcb. P (tcbPriority tcb)) t" + (wp: crunch_wps setEndpoint_obj_at_tcb' + setThreadState_obj_at_unchanged setNotification_tcb setBoundNotification_obj_at_unchanged + simp: crunch_simps unless_def) + +context kernel_m begin + +lemma setThreadState_no_sch_change: + "\\s. P (ksSchedulerAction s) \ (runnable' st \ t \ ksCurThread s)\ + setThreadState st t + \\rv s. P (ksSchedulerAction s)\" + (is "NonDetMonad.valid ?P ?f ?Q") + apply (simp add: setThreadState_def setSchedulerAction_def) + apply (wp hoare_pre_cont[where a=rescheduleRequired]) + apply (rule_tac Q="\_. ?P and st_tcb_at' (op = st) t" in hoare_post_imp) + apply (clarsimp split: split_if) + apply (clarsimp simp: obj_at'_def st_tcb_at'_def projectKOs) + apply (rule hoare_pre, wp threadSet_pred_tcb_at_state) + apply simp + done + +lemma asUser_obj_at_unchangedT: + assumes x: "\tcb con con'. con' \ fst (m con) + \ P (tcbArch_update (\_. atcbContextSet (snd con') (tcbArch tcb)) tcb) = P tcb" shows + "\obj_at' P t\ asUser t' m \\rv. obj_at' P t\" + apply (simp add: asUser_def split_def) + apply (wp threadSet_obj_at' threadGet_wp) + apply (clarsimp simp: obj_at'_def projectKOs x cong: if_cong) + done + +lemmas asUser_obj_at_unchanged + = asUser_obj_at_unchangedT[OF all_tcbI, rule_format] + +lemma bind_assoc: + "do y \ do x \ m; f x od; g y od + = do x \ m; y \ f x; g y od" + by (rule bind_assoc) + +lemma setObject_modify_assert: + "\ updateObject v = updateObject_default v \ + \ setObject p v = do f \ gets (obj_at' (\v'. v = v' \ True) p); + assert f; modify (ksPSpace_update (\ps. ps(p \ injectKO v))) od" + using objBits_2n[where obj=v] + apply (simp add: setObject_def split_def updateObject_default_def + bind_assoc projectKO_def2 alignCheck_assert) + apply (rule ext, simp add: exec_gets) + apply (case_tac "obj_at' (\v'. v = v' \ True) p x") + apply (clarsimp simp: obj_at'_def projectKOs lookupAround2_known1 + assert_opt_def) + apply (clarsimp simp: project_inject) + apply (simp only: objBits_def objBitsT_koTypeOf[symmetric] koTypeOf_injectKO) + apply (simp add: magnitudeCheck_assert2 simpler_modify_def) + apply (clarsimp simp: assert_opt_def assert_def magnitudeCheck_assert2 + split: option.split split_if) + apply (clarsimp simp: obj_at'_def projectKOs) + apply (clarsimp simp: project_inject) + apply (simp only: objBits_def objBitsT_koTypeOf[symmetric] + koTypeOf_injectKO simp_thms) + done + +lemma setEndpoint_isolatable: + "thread_actions_isolatable idx (setEndpoint p e)" + apply (simp add: setEndpoint_def setObject_modify_assert + assert_def) + apply (case_tac "p \ range idx") + apply (clarsimp simp: thread_actions_isolatable_def + monadic_rewrite_def fun_eq_iff + liftM_def isolate_thread_actions_def + bind_assoc exec_gets getSchedulerAction_def + bind_select_f_bind[symmetric]) + apply (simp add: obj_at_partial_overwrite_id2) + apply (drule_tac x=x in spec) + apply (clarsimp simp: obj_at'_def projectKOs select_f_asserts) + apply (intro thread_actions_isolatable_bind[OF _ _ hoare_pre(1)] + thread_actions_isolatable_if + thread_actions_isolatable_return + thread_actions_isolatable_fail) + apply (rule gets_isolatable) + apply (simp add: obj_at_partial_overwrite_id2) + apply (rule modify_isolatable) + apply (clarsimp simp: o_def partial_overwrite_def) + apply (rule kernel_state.fold_congs[OF refl refl]) + apply (clarsimp simp: fun_eq_iff + split: split_if) + apply (wp | simp)+ + done + +lemma setCTE_assert_modify: + "setCTE p v = do c \ gets (real_cte_at' p); + t \ gets (tcb_at' (p && ~~ mask 9) + and K ((p && mask 9) \ dom tcb_cte_cases)); + if c then modify (ksPSpace_update (\ps. ps(p \ KOCTE v))) + else if t then + modify (ksPSpace_update + (\ps. ps (p && ~~ mask 9 \ + KOTCB (snd (the (tcb_cte_cases (p && mask 9))) (K v) + (the (projectKO_opt (the (ps (p && ~~ mask 9))))))))) + else fail od" + apply (clarsimp simp: setCTE_def setObject_def split_def + fun_eq_iff exec_gets) + apply (case_tac "real_cte_at' p x") + apply (clarsimp simp: obj_at'_def projectKOs lookupAround2_known1 + assert_opt_def alignCheck_assert objBits_simps + magnitudeCheck_assert2 updateObject_cte) + apply (simp add: simpler_modify_def) + apply (simp split: split_if, intro conjI impI) + apply (clarsimp simp: obj_at'_def projectKOs) + apply (subgoal_tac "p \ (p && ~~ mask 9) + 2 ^ 9 - 1") + apply (subgoal_tac "fst (lookupAround2 p (ksPSpace x)) + = Some (p && ~~ mask 9, KOTCB obj)") + apply (simp add: assert_opt_def) + apply (subst updateObject_cte_tcb) + apply (fastforce simp add: subtract_mask) + apply (simp add: assert_opt_def alignCheck_assert bind_assoc + magnitudeCheck_assert + is_aligned_neg_mask2 objBits_def) + apply (rule ps_clear_lookupAround2, assumption+) + apply (rule word_and_le2) + apply (simp add: objBits_simps mask_def field_simps) + apply (simp add: simpler_modify_def cong: option.case_cong if_cong) + apply (rule kernel_state.fold_congs[OF refl refl]) + apply (clarsimp simp: projectKO_opt_tcb cong: if_cong) + apply (clarsimp simp: lookupAround2_char1 word_and_le2) + apply (rule ccontr, clarsimp) + apply (erule(2) ps_clearD) + apply (simp add: objBits_simps mask_def field_simps) + apply (rule tcb_cte_cases_in_range2) + apply (simp add: subtract_mask) + apply simp + apply (clarsimp simp: assert_opt_def split: option.split) + apply (rule trans [OF bind_apply_cong[OF _ refl] fun_cong[OF fail_bind]]) + apply (simp add: fail_def prod_eq_iff) + apply (rule context_conjI) + apply (rule ccontr, clarsimp elim!: nonemptyE) + apply (frule(1) updateObject_cte_is_tcb_or_cte[OF _ refl]) + apply (erule disjE) + apply clarsimp + apply (frule(1) tcb_cte_cases_aligned_helpers) + apply (clarsimp simp: domI[where m = cte_cte_cases] field_simps) + apply (clarsimp simp: lookupAround2_char1 obj_at'_def projectKOs + objBits_simps) + apply (clarsimp simp: obj_at'_def lookupAround2_char1 + objBits_simps projectKOs cte_level_bits_def) + apply (erule empty_failD[OF empty_fail_updateObject_cte]) + done + +lemma partial_overwrite_fun_upd2: + "partial_overwrite idx tsrs (f (x := y)) + = (partial_overwrite idx tsrs f) + (x := if x \ range idx then put_tcb_state_regs (tsrs (inv idx x)) y + else y)" + by (simp add: fun_eq_iff partial_overwrite_def split: split_if) + +lemma setCTE_isolatable: + "thread_actions_isolatable idx (setCTE p v)" + apply (simp add: setCTE_assert_modify) + apply (clarsimp simp: thread_actions_isolatable_def + monadic_rewrite_def fun_eq_iff + liftM_def exec_gets + isolate_thread_actions_def + bind_assoc exec_gets getSchedulerAction_def + bind_select_f_bind[symmetric] + obj_at_partial_overwrite_If + obj_at_partial_overwrite_id2 + cong: if_cong) + apply (case_tac "p && ~~ mask 9 \ range idx \ p && mask 9 \ dom tcb_cte_cases") + apply clarsimp + apply (frule_tac x=x in spec, erule obj_atE') + apply (subgoal_tac "\ real_cte_at' p s") + apply (clarsimp simp: select_f_returns select_f_asserts split: split_if) + apply (clarsimp simp: o_def simpler_modify_def partial_overwrite_fun_upd2) + apply (rule kernel_state.fold_congs[OF refl refl]) + apply (rule ext) + apply (clarsimp simp: partial_overwrite_get_tcb_state_regs + split: split_if) + apply (clarsimp simp: projectKOs get_tcb_state_regs_def + put_tcb_state_regs_def put_tcb_state_regs_tcb_def + partial_overwrite_def + split: tcb_state_regs.split) + apply (case_tac obj, simp add: projectKO_opt_tcb) + apply (simp add: tcb_cte_cases_def split: split_if_asm) + apply (drule_tac x=x in spec) + apply (clarsimp simp: obj_at'_def projectKOs objBits_simps subtract_mask(2) [symmetric]) + apply (erule notE[rotated], erule (3) tcb_ctes_clear[rotated]) + apply (simp add: select_f_returns select_f_asserts split: split_if) + apply (intro conjI impI) + apply (clarsimp simp: simpler_modify_def fun_eq_iff + partial_overwrite_fun_upd2 o_def + intro!: kernel_state.fold_congs[OF refl refl]) + apply (clarsimp simp: obj_at'_def projectKOs objBits_simps) + apply (erule notE[rotated], rule tcb_ctes_clear[rotated 2], assumption+) + apply (fastforce simp add: subtract_mask) + apply simp + apply (clarsimp simp: simpler_modify_def + partial_overwrite_fun_upd2 o_def + partial_overwrite_get_tcb_state_regs + intro!: kernel_state.fold_congs[OF refl refl] + split: split_if) + apply (simp add: partial_overwrite_def) + apply (subgoal_tac "p \ range idx") + apply (clarsimp simp: simpler_modify_def + partial_overwrite_fun_upd2 o_def + partial_overwrite_get_tcb_state_regs + intro!: kernel_state.fold_congs[OF refl refl]) + apply clarsimp + apply (drule_tac x=x in spec) + apply (clarsimp simp: obj_at'_def projectKOs) + done + +lemma assert_isolatable: + "thread_actions_isolatable idx (assert P)" + by (simp add: assert_def thread_actions_isolatable_if + thread_actions_isolatable_returns + thread_actions_isolatable_fail) + +lemma cteInsert_isolatable: + "thread_actions_isolatable idx (cteInsert cap src dest)" + apply (simp add: cteInsert_def updateCap_def updateMDB_def + Let_def setUntypedCapAsFull_def) + apply (intro thread_actions_isolatable_bind[OF _ _ hoare_pre(1)] + thread_actions_isolatable_if + thread_actions_isolatable_returns assert_isolatable + getCTE_isolatable setCTE_isolatable) + apply (wp | simp)+ + done + +lemma isolate_thread_actions_threadSet_tcbState: + "\ inj idx; idx t' = t \ \ + monadic_rewrite False True (\s. \x. tcb_at' (idx x) s) + (threadSet (tcbState_update (\_. st)) t) + (isolate_thread_actions idx (return ()) + (\tsrs. (tsrs (t' := TCBStateRegs st (tsrContext (tsrs t'))))) + id)" + apply (simp add: isolate_thread_actions_def bind_assoc split_def + select_f_returns getSchedulerAction_def) + apply (clarsimp simp: monadic_rewrite_def exec_gets threadSet_def + getObject_get_assert bind_assoc liftM_def + setObject_modify_assert) + apply (frule_tac x=t' in spec, drule obj_at_ko_at') + apply (clarsimp simp: exec_gets simpler_modify_def o_def + intro!: kernel_state.fold_congs[OF refl refl]) + apply (simp add: partial_overwrite_fun_upd + partial_overwrite_get_tcb_state_regs) + apply (clarsimp simp: put_tcb_state_regs_def put_tcb_state_regs_tcb_def + projectKOs get_tcb_state_regs_def + elim!: obj_atE') + apply (case_tac ko) + apply (simp add: projectKO_opt_tcb) + done + +lemma thread_actions_isolatableD: + "\ thread_actions_isolatable idx f; inj idx \ + \ monadic_rewrite False True (\s. (\x. tcb_at' (idx x) s)) + f (isolate_thread_actions idx f id id)" + by (clarsimp simp: thread_actions_isolatable_def) + +lemma tcbSchedDequeue_rewrite: + "monadic_rewrite True True (obj_at' (Not \ tcbQueued) t) (tcbSchedDequeue t) (return ())" + apply (simp add: tcbSchedDequeue_def) + apply (rule monadic_rewrite_imp) + apply (rule monadic_rewrite_trans) + apply (rule monadic_rewrite_bind_tail) + apply (rule_tac P="\ queued" in monadic_rewrite_gen_asm) + apply (simp add: when_def) + apply (rule monadic_rewrite_refl) + apply (wp threadGet_const) + apply (rule monadic_rewrite_symb_exec2) + apply wp + apply (rule monadic_rewrite_refl) + apply (clarsimp) + done + +lemma switchToThread_rewrite: + "monadic_rewrite True True + (ct_in_state' (Not \ runnable') and cur_tcb' and obj_at' (Not \ tcbQueued) t) + (switchToThread t) + (do Arch.switchToThread t; setCurThread t od)" + apply (simp add: switchToThread_def Thread_H.switchToThread_def) + apply (rule monadic_rewrite_imp) + apply (rule monadic_rewrite_trans) + apply (rule monadic_rewrite_bind_tail) + apply (rule monadic_rewrite_bind) + apply (rule tcbSchedDequeue_rewrite) + apply (rule monadic_rewrite_refl) + apply (wp Arch_switchToThread_obj_at_pre) + apply (rule monadic_rewrite_bind_tail) + apply (rule monadic_rewrite_symb_exec) + apply (wp, simp) + apply (rule monadic_rewrite_refl) + apply (wp) + apply (clarsimp simp: comp_def) + done + +lemma threadGet_isolatable: + assumes v: "\tsr. \tcb. f (put_tcb_state_regs_tcb tsr tcb) = f tcb" + shows "thread_actions_isolatable idx (threadGet f t)" + apply (clarsimp simp: threadGet_def thread_actions_isolatable_def + isolate_thread_actions_def split_def + getObject_get_assert liftM_def + bind_select_f_bind[symmetric] + select_f_returns select_f_asserts bind_assoc) + apply (clarsimp simp: monadic_rewrite_def exec_gets + getSchedulerAction_def) + apply (simp add: obj_at_partial_overwrite_If) + apply (rule bind_apply_cong[OF refl]) + apply (clarsimp simp: exec_gets exec_modify o_def + ksPSpace_update_partial_id in_monad) + apply (erule obj_atE') + apply (clarsimp simp: projectKOs + partial_overwrite_def put_tcb_state_regs_def + cong: if_cong) + apply (simp add: projectKO_opt_tcb v split: split_if) + done + + lemma switchToThread_isolatable: + "thread_actions_isolatable idx (Arch.switchToThread t)" + apply (simp add: ARM_H.switchToThread_def + storeWordUser_def stateAssert_def2) + apply (intro thread_actions_isolatable_bind[OF _ _ hoare_pre(1)] + gets_isolatable setVMRoot_isolatable + thread_actions_isolatable_if + doMachineOp_isolatable + threadGet_isolatable [OF all_tcbI] + thread_actions_isolatable_returns + thread_actions_isolatable_fail) + apply (wp | + simp add: pointerInUserData_def + typ_at_to_obj_at_arches + obj_at_partial_overwrite_id2 + put_tcb_state_regs_tcb_def + split: tcb_state_regs.split)+ + done + +lemma monadic_rewrite_trans_dup: + "\ monadic_rewrite F E P f g; monadic_rewrite F E P g h \ + \ monadic_rewrite F E P f h" + by (auto simp add: monadic_rewrite_def) + + +lemma setCurThread_isolatable: + "thread_actions_isolatable idx (setCurThread t)" + by (simp add: setCurThread_def modify_isolatable) + +lemma isolate_thread_actions_tcbs_at: + assumes f: "\x. \tcb_at' (idx x)\ f \\rv. tcb_at' (idx x)\" shows + "\\s. \x. tcb_at' (idx x) s\ + isolate_thread_actions idx f f' f'' \\p s. \x. tcb_at' (idx x) s\" + apply (simp add: isolate_thread_actions_def split_def) + apply wp + apply clarsimp + apply (simp add: obj_at_partial_overwrite_If use_valid[OF _ f]) + done + +lemma isolate_thread_actions_rewrite_bind: + "\ inj idx; monadic_rewrite False True (\s. \x. tcb_at' (idx x) s) + f (isolate_thread_actions idx f' f'' f'''); + \x. monadic_rewrite False True (\s. \x. tcb_at' (idx x) s) + (g x) + (isolate_thread_actions idx (g' x) g'' g'''); + thread_actions_isolatable idx f'; \x. thread_actions_isolatable idx (g' x); + \x. \tcb_at' (idx x)\ f' \\rv. tcb_at' (idx x)\ \ + \ monadic_rewrite False True (\s. \x. tcb_at' (idx x) s) + (f >>= g) (isolate_thread_actions idx + (f' >>= g') (g'' o f'') (g''' o f'''))" + apply (rule monadic_rewrite_imp) + apply (rule monadic_rewrite_trans) + apply (rule monadic_rewrite_bind, assumption+) + apply (wp isolate_thread_actions_tcbs_at) + apply simp + apply (subst isolate_thread_actions_wrap_bind, assumption) + apply (rule monadic_rewrite_in_isolate_thread_actions, assumption) + apply (rule monadic_rewrite_transverse) + apply (rule monadic_rewrite_bind2) + apply (erule(1) thread_actions_isolatableD) + apply (rule thread_actions_isolatableD, assumption+) + apply (rule hoare_vcg_all_lift, assumption) + apply (simp add: liftM_def id_def) + apply (rule monadic_rewrite_refl) + apply (simp add: obj_at_partial_overwrite_If) + done + +definition + "copy_register_tsrs src dest r r' rf tsrs + = tsrs (dest := TCBStateRegs (tsrState (tsrs dest)) + ((tsrContext (tsrs dest)) (r' := rf (tsrContext (tsrs src) r))))" + +lemma tcb_at_KOTCB_upd: + "tcb_at' (idx x) s \ + tcb_at' p (ksPSpace_update (\ps. ps(idx x \ KOTCB tcb)) s) + = tcb_at' p s" + apply (clarsimp simp: obj_at'_def projectKOs objBits_simps + split: split_if) + apply (simp add: ps_clear_def) + done + +definition + "set_register_tsrs dest r v tsrs + = tsrs (dest := TCBStateRegs (tsrState (tsrs dest)) + ((tsrContext (tsrs dest)) (r := v)))" + + +lemma set_register_isolate: + "\ inj idx; idx y = dest \ \ + monadic_rewrite False True + (\s. \x. tcb_at' (idx x) s) + (asUser dest (setRegister r v)) + (isolate_thread_actions idx (return ()) + (set_register_tsrs y r v) id)" + apply (simp add: asUser_def split_def bind_assoc + getRegister_def setRegister_def + select_f_returns isolate_thread_actions_def + getSchedulerAction_def) + apply (simp add: threadGet_def liftM_def getObject_get_assert + bind_assoc threadSet_def + setObject_modify_assert) + apply (clarsimp simp: monadic_rewrite_def exec_gets + exec_modify tcb_at_KOTCB_upd) + apply (clarsimp simp: simpler_modify_def + intro!: kernel_state.fold_congs[OF refl refl]) + apply (clarsimp simp: set_register_tsrs_def o_def + partial_overwrite_fun_upd + partial_overwrite_get_tcb_state_regs) + apply (drule_tac x=y in spec) + apply (clarsimp simp: obj_at'_def projectKOs objBits_simps + cong: if_cong) + apply (case_tac obj) + apply (simp add: projectKO_opt_tcb put_tcb_state_regs_def + put_tcb_state_regs_tcb_def get_tcb_state_regs_def + cong: if_cong) + done + +lemma copy_register_isolate: + "\ inj idx; idx x = src; idx y = dest \ \ + monadic_rewrite False True + (\s. \x. tcb_at' (idx x) s) + (do v \ asUser src (getRegister r); + asUser dest (setRegister r' (rf v)) od) + (isolate_thread_actions idx (return ()) + (copy_register_tsrs x y r r' rf) id)" + apply (simp add: asUser_def split_def bind_assoc + getRegister_def setRegister_def + select_f_returns isolate_thread_actions_def + getSchedulerAction_def) + apply (simp add: threadGet_def liftM_def getObject_get_assert + bind_assoc threadSet_def + setObject_modify_assert) + apply (clarsimp simp: monadic_rewrite_def exec_gets + exec_modify tcb_at_KOTCB_upd) + apply (clarsimp simp: simpler_modify_def + intro!: kernel_state.fold_congs[OF refl refl]) + apply (clarsimp simp: copy_register_tsrs_def o_def + partial_overwrite_fun_upd + partial_overwrite_get_tcb_state_regs) + apply (frule_tac x=x in spec, drule_tac x=y in spec) + apply (clarsimp simp: obj_at'_def projectKOs objBits_simps + cong: if_cong) + apply (case_tac obj, case_tac obja) + apply (simp add: projectKO_opt_tcb put_tcb_state_regs_def + put_tcb_state_regs_tcb_def get_tcb_state_regs_def + cong: if_cong) + apply (auto simp: fun_eq_iff split: split_if) + done + +lemmas monadic_rewrite_bind_alt + = monadic_rewrite_trans[OF monadic_rewrite_bind_tail monadic_rewrite_bind_head, rotated -1] + + +lemma monadic_rewrite_isolate_final2: + assumes mr: "monadic_rewrite F E Q f g" + and eqs: "\s tsrs. \ P s; tsrs = get_tcb_state_regs o ksPSpace s o idx \ + \ f' tsrs = g' tsrs" + "\s. P s \ f'' (ksSchedulerAction s) = g'' (ksSchedulerAction s)" + "\s tsrs sa. R s \ + Q ((ksPSpace_update (partial_overwrite idx tsrs) + s) (| ksSchedulerAction := sa |))" + shows + "monadic_rewrite F E (P and R) + (isolate_thread_actions idx f f' f'') + (isolate_thread_actions idx g g' g'')" + apply (simp add: isolate_thread_actions_def split_def) + apply (rule monadic_rewrite_imp) + apply (rule monadic_rewrite_bind_tail)+ + apply (rule_tac P="\ s'. Q s" in monadic_rewrite_bind) + apply (insert mr)[1] + apply (simp add: monadic_rewrite_def select_f_def) + apply auto[1] + apply (rule_tac P="P and (\s. tcbs = get_tcb_state_regs o ksPSpace s o idx + \ sa = ksSchedulerAction s)" + in monadic_rewrite_refl3) + apply (clarsimp simp: exec_modify eqs return_def) + apply wp + apply (clarsimp simp: o_def eqs) + done + +lemmas monadic_rewrite_isolate_final + = monadic_rewrite_isolate_final2[where R=\, OF monadic_rewrite_refl2, simplified] + +lemma copy_registers_isolate_general: + "\ inj idx; idx x = t; idx y = t' \ \ + monadic_rewrite False True + (\s. \x. tcb_at' (idx x) s) + (mapM_x (\r. do v \ asUser t (getRegister (f r)); + asUser t' (setRegister (f' r) (rf r v)) + od) + regs) + (isolate_thread_actions idx + (return ()) (foldr (\r. copy_register_tsrs x y (f r) (f' r) (rf r)) (rev regs)) id)" + apply (induct regs) + apply (simp add: mapM_x_Nil) + apply (clarsimp simp: monadic_rewrite_def liftM_def bind_assoc + isolate_thread_actions_def + split_def exec_gets getSchedulerAction_def + select_f_returns o_def ksPSpace_update_partial_id) + apply (simp add: return_def simpler_modify_def) + apply (simp add: mapM_x_Cons) + apply (rule monadic_rewrite_imp) + apply (rule monadic_rewrite_trans) + apply (rule isolate_thread_actions_rewrite_bind, assumption) + apply (rule copy_register_isolate, assumption+) + apply (rule thread_actions_isolatable_returns)+ + apply wp + apply (rule monadic_rewrite_isolate_final[where P=\], simp+) + done + +lemmas copy_registers_isolate = copy_registers_isolate_general[where f="\x. x" and f'="\x. x" and rf="\_ x. x"] + +lemma setSchedulerAction_isolate: + "inj idx \ + monadic_rewrite False True (\s. \x. tcb_at' (idx x) s) + (setSchedulerAction sa) + (isolate_thread_actions idx (return ()) id (\_. sa))" + apply (clarsimp simp: monadic_rewrite_def liftM_def bind_assoc + isolate_thread_actions_def select_f_returns + exec_gets getSchedulerAction_def o_def + ksPSpace_update_partial_id setSchedulerAction_def) + apply (simp add: simpler_modify_def) + done + +lemma updateMDB_isolatable: + "thread_actions_isolatable idx (updateMDB slot f)" + apply (simp add: updateMDB_def thread_actions_isolatable_return + split: split_if) + apply (intro impI thread_actions_isolatable_bind[OF _ _ hoare_pre(1)] + getCTE_isolatable setCTE_isolatable, + (wp | simp)+) + done + +lemma clearUntypedFreeIndex_isolatable: + "thread_actions_isolatable idx (clearUntypedFreeIndex slot)" + apply (simp add: clearUntypedFreeIndex_def getSlotCap_def) + apply (rule thread_actions_isolatable_bind) + apply (rule getCTE_isolatable) + apply (simp split: capability.split, safe intro!: thread_actions_isolatable_return) + apply (simp add: updateTrackedFreeIndex_def getSlotCap_def) + apply (intro thread_actions_isolatable_bind getCTE_isolatable + modify_isolatable) + apply (wp | simp)+ + done + +lemma emptySlot_isolatable: + "thread_actions_isolatable idx (emptySlot slot None)" + apply (simp add: emptySlot_def updateCap_def case_Null_If + cong: if_cong) + apply (intro thread_actions_isolatable_bind[OF _ _ hoare_pre(1)] + clearUntypedFreeIndex_isolatable + thread_actions_isolatable_if + getCTE_isolatable setCTE_isolatable + thread_actions_isolatable_return + updateMDB_isolatable, + (wp | simp)+) + done + +lemmas fastpath_isolatables + = setEndpoint_isolatable getCTE_isolatable + assert_isolatable cteInsert_isolatable + switchToThread_isolatable setCurThread_isolatable + emptySlot_isolatable updateMDB_isolatable + thread_actions_isolatable_returns + +lemmas fastpath_isolate_rewrites + = isolate_thread_actions_threadSet_tcbState isolate_thread_actions_asUser + copy_registers_isolate setSchedulerAction_isolate + fastpath_isolatables[THEN thread_actions_isolatableD] + +lemma lookupIPCBuffer_isolatable: + "thread_actions_isolatable idx (lookupIPCBuffer w t)" + apply (simp add: lookupIPCBuffer_def) + apply (rule thread_actions_isolatable_bind) + apply (clarsimp simp: put_tcb_state_regs_tcb_def threadGet_isolatable + getThreadBufferSlot_def locateSlot_conv getSlotCap_def + split: tcb_state_regs.split)+ + apply (rule thread_actions_isolatable_bind) + apply (clarsimp simp: thread_actions_isolatable_return + getCTE_isolatable + assert_isolatable + split: capability.split arch_capability.split bool.split)+ + apply (rule thread_actions_isolatable_if) + apply (rule thread_actions_isolatable_bind) + apply (simp add: assert_isolatable thread_actions_isolatable_return | wp)+ + done + +end + +end diff --git a/proof/crefine/Recycle_C.thy b/proof/crefine/Recycle_C.thy index 5c9eeee62..ef2995b25 100644 --- a/proof/crefine/Recycle_C.thy +++ b/proof/crefine/Recycle_C.thy @@ -43,7 +43,7 @@ lemma coerce_memset_to_heap_update_user_data: apply (simp add: ti_typ_pad_combine_empty_ti ti_typ_pad_combine_td align_of_def padup_def final_pad_def size_td_lt_ti_typ_pad_combine Let_def size_of_def) apply (simp add: typ_info_simps - user_context_C_tag_def thread_state_C_tag_def fault_C_tag_def + user_context_C_tag_def thread_state_C_tag_def seL4_Fault_C_tag_def lookup_fault_C_tag_def update_ti_t_ptr_0s ti_typ_pad_combine_empty_ti ti_typ_pad_combine_td ti_typ_combine_empty_ti ti_typ_combine_td @@ -394,7 +394,7 @@ lemma coerce_memset_to_heap_update_asidpool: apply (simp add: ti_typ_pad_combine_empty_ti ti_typ_pad_combine_td align_of_def padup_def final_pad_def size_td_lt_ti_typ_pad_combine Let_def size_of_def) apply (simp add: typ_info_simps - user_context_C_tag_def thread_state_C_tag_def fault_C_tag_def + user_context_C_tag_def thread_state_C_tag_def seL4_Fault_C_tag_def lookup_fault_C_tag_def update_ti_t_ptr_0s ti_typ_pad_combine_empty_ti ti_typ_pad_combine_td ti_typ_combine_empty_ti ti_typ_combine_td @@ -1236,7 +1236,7 @@ lemma coerce_memset_to_heap_update: (tcb_C (arch_tcb_C (user_context_C (FCP (\x. 0)))) (thread_state_C (FCP (\x. 0))) (NULL) - (fault_C (FCP (\x. 0))) + (seL4_Fault_C (FCP (\x. 0))) (lookup_fault_C (FCP (\x. 0))) 0 0 0 0 0 0 NULL NULL NULL NULL)" apply (intro ext, simp add: heap_update_def) @@ -1245,7 +1245,7 @@ lemma coerce_memset_to_heap_update: apply (simp add: ti_typ_pad_combine_empty_ti ti_typ_pad_combine_td align_of_def padup_def final_pad_def size_td_lt_ti_typ_pad_combine Let_def size_of_def) apply (simp add: typ_info_simps - user_context_C_tag_def thread_state_C_tag_def fault_C_tag_def + user_context_C_tag_def thread_state_C_tag_def seL4_Fault_C_tag_def lookup_fault_C_tag_def update_ti_t_ptr_0s arch_tcb_C_tag_def ti_typ_pad_combine_empty_ti ti_typ_pad_combine_td ti_typ_combine_empty_ti ti_typ_combine_td diff --git a/proof/crefine/Refine_C.thy b/proof/crefine/Refine_C.thy index 8f8eb583e..5d1192302 100644 --- a/proof/crefine/Refine_C.thy +++ b/proof/crefine/Refine_C.thy @@ -123,7 +123,7 @@ lemma handleUnknownSyscall_ccorres: apply (clarsimp simp: ct_in_state'_def) apply (frule st_tcb_idle'[rotated]) apply (erule invs_valid_idle') - apply (clarsimp simp: cfault_rel_def fault_unknown_syscall_lift is_cap_fault_def) + apply (clarsimp simp: cfault_rel_def seL4_Fault_UnknownSyscall_lift is_cap_fault_def) done lemma handleVMFaultEvent_ccorres: @@ -210,7 +210,7 @@ lemma handleUserLevelFault_ccorres: apply (frule st_tcb_idle'[rotated]) apply (erule invs_valid_idle') apply simp - apply (clarsimp simp: cfault_rel_def fault_user_exception_lift) + apply (clarsimp simp: cfault_rel_def seL4_Fault_UserException_lift) apply (simp add: is_cap_fault_def) done @@ -552,8 +552,8 @@ lemma ccorres_add_gets: lemma ccorres_get_registers: "\ \cptr msgInfo. ccorres dc xfdc ((\s. P s \ Q s \ - obj_at' (\tcb. tcbContext tcb ARM_H.capRegister = cptr - \ tcbContext tcb ARM_H.msgInfoRegister = msgInfo) + obj_at' (\tcb. (atcbContextGet o tcbArch) tcb ARM_H.capRegister = cptr + \ (atcbContextGet o tcbArch) tcb ARM_H.msgInfoRegister = msgInfo) (ksCurThread s) s) and R) (UNIV \ \\cptr = cptr\ \ \\msgInfo = msgInfo\) [] m c \ \ @@ -566,14 +566,15 @@ lemma ccorres_get_registers: apply (rule ccorres_assume_pre) apply (clarsimp simp: ct_in_state'_def st_tcb_at'_def) apply (drule obj_at_ko_at', clarsimp) - apply (erule_tac x="tcbContext ko ARM_H.capRegister" in meta_allE) - apply (erule_tac x="tcbContext ko ARM_H.msgInfoRegister" in meta_allE) + apply (erule_tac x="(atcbContextGet o tcbArch) ko ARM_H.capRegister" in meta_allE) + apply (erule_tac x="(atcbContextGet o tcbArch) ko ARM_H.msgInfoRegister" in meta_allE) apply (erule ccorres_guard_imp2) apply (clarsimp simp: rf_sr_ksCurThread) apply (drule(1) obj_at_cslift_tcb, clarsimp simp: obj_at'_def projectKOs) apply (clarsimp simp: ctcb_relation_def ccontext_relation_def ARM_H.msgInfoRegister_def ARM_H.capRegister_def ARM.msgInfoRegister_def ARM.capRegister_def + carch_tcb_relation_def "StrictC'_register_defs") done @@ -618,12 +619,13 @@ lemma callKernel_withFastpath_corres_C: lemma threadSet_all_invs_triv': "\all_invs' e and (\s. t = ksCurThread s)\ - threadSet (tcbContext_update f) t \\_. all_invs' e\" + threadSet (\tcb. tcbArch_update (\_. atcbContextSet f (tcbArch tcb)) tcb) t \\_. all_invs' e\" unfolding all_invs'_def apply (rule hoare_pre) apply (rule wp_from_corres_unit) - apply (rule threadset_corresT [where f="tcb_context_update f"]) - apply (simp add: tcb_relation_def) + apply (rule threadset_corresT [where f="tcb_arch_update (arch_tcb_context_set f)"]) + apply (simp add: tcb_relation_def arch_tcb_context_set_def + atcbContextSet_def arch_tcb_relation_def) apply (simp add: tcb_cap_cases_def) apply (simp add: tcb_cte_cases_def) apply (simp add: exst_same_def) @@ -642,7 +644,7 @@ lemma threadSet_all_invs_triv': lemma getContext_corres: "t' = tcb_ptr_to_ctcb_ptr t \ corres_underlying rf_sr False True (op =) (tcb_at' t) \ - (threadGet tcbContext t) (gets (getContext_C t'))" + (threadGet (atcbContextGet o tcbArch) t) (gets (getContext_C t'))" apply (clarsimp simp: corres_underlying_def simpler_gets_def) apply (drule obj_at_ko_at') apply clarsimp @@ -654,7 +656,7 @@ lemma getContext_corres: apply (clarsimp simp: getContext_C_def) apply (drule cmap_relation_ko_atD [rotated]) apply fastforce - apply (clarsimp simp: typ_heap_simps ctcb_relation_def from_user_context_C) + apply (clarsimp simp: typ_heap_simps ctcb_relation_def carch_tcb_relation_def from_user_context_C) done lemma callKernel_cur: diff --git a/proof/crefine/Retype_C.thy b/proof/crefine/Retype_C.thy index d705459d9..812a9ebfe 100644 --- a/proof/crefine/Retype_C.thy +++ b/proof/crefine/Retype_C.thy @@ -3548,9 +3548,9 @@ proof - (thread_state_C.words_C (tcbState_C undefined))) (tcbState_C undefined), tcbFault_C := - fault_C.words_C_update + seL4_Fault_C.words_C_update (\_. foldr (\n arr. Arrays.update arr n 0) [0..<2] - (fault_C.words_C (tcbFault_C undefined))) + (seL4_Fault_C.words_C (tcbFault_C undefined))) (tcbFault_C undefined), tcbLookupFailure_C := lookup_fault_C.words_C_update @@ -3569,7 +3569,7 @@ proof - final_pad_def size_td_lt_ti_typ_pad_combine Let_def size_of_def)(* takes ages *) apply (simp add: update_ti_adjust_ti update_ti_t_word32_0s typ_info_simps - user_context_C_tag_def thread_state_C_tag_def fault_C_tag_def + user_context_C_tag_def thread_state_C_tag_def seL4_Fault_C_tag_def lookup_fault_C_tag_def update_ti_t_ptr_0s arch_tcb_C_tag_def ti_typ_pad_combine_empty_ti ti_typ_pad_combine_td ti_typ_combine_empty_ti ti_typ_combine_td @@ -3586,14 +3586,14 @@ proof - apply (intro conjI) apply (simp add: cthread_state_relation_def thread_state_lift_def eval_nat_numeral ThreadState_Inactive_def) - apply (simp add: ccontext_relation_def newContext_def2) + apply (simp add: ccontext_relation_def newContext_def2 carch_tcb_relation_def) apply rule - apply (case_tac r, simp_all add: "StrictC'_register_defs" eval_nat_numeral)[1] -- "takes ages" - apply (simp add: thread_state_lift_def eval_nat_numeral) + apply (case_tac r, simp_all add: "StrictC'_register_defs" eval_nat_numeral atcbContext_def newArchTCB_def newContext_def initContext_def)[1] -- "takes ages" + apply (simp add: thread_state_lift_def eval_nat_numeral atcbContextGet_def)+ apply (simp add: timeSlice_def) - apply (simp add: cfault_rel_def fault_lift_def fault_get_tag_def Let_def + apply (simp add: cfault_rel_def seL4_Fault_lift_def seL4_Fault_get_tag_def Let_def lookup_fault_lift_def lookup_fault_get_tag_def lookup_fault_invalid_root_def - eval_nat_numeral fault_null_fault_def option_to_ptr_def option_to_0_def + eval_nat_numeral seL4_Fault_NullFault_def option_to_ptr_def option_to_0_def split: split_if)+ done diff --git a/proof/crefine/SR_lemmas_C.thy b/proof/crefine/SR_lemmas_C.thy index f98cf8962..8d84ecb40 100644 --- a/proof/crefine/SR_lemmas_C.thy +++ b/proof/crefine/SR_lemmas_C.thy @@ -356,9 +356,9 @@ lemma updateObject_cte_tcb: done definition - tcb_no_ctes_proj :: "tcb \ Structures_H.thread_state \ word32 \ word32 \ (MachineTypes.register \ word32) \ bool \ word8 \ word8 \ word8 \ nat \ fault option \ word32 option" + tcb_no_ctes_proj :: "tcb \ Structures_H.thread_state \ word32 \ word32 \ arch_tcb \ bool \ word8 \ word8 \ word8 \ nat \ fault option \ word32 option" where - "tcb_no_ctes_proj t \ (tcbState t, tcbFaultHandler t, tcbIPCBuffer t, tcbContext t, tcbQueued t, + "tcb_no_ctes_proj t \ (tcbState t, tcbFaultHandler t, tcbIPCBuffer t, tcbArch t, tcbQueued t, tcbMCP t, tcbPriority t, tcbDomain t, tcbTimeSlice t, tcbFault t, tcbBoundNotification t)" lemma tcb_cte_cases_proj_eq [simp]: @@ -1750,7 +1750,7 @@ lemma ctcb_ptr_to_ctcb_ptr [simp]: declare ucast_id [simp] definition - cap_rights_from_word_canon :: "word32 \ cap_rights_CL" + cap_rights_from_word_canon :: "word32 \ seL4_CapRights_CL" where "cap_rights_from_word_canon wd \ \ capAllowGrant_CL = from_bool (wd !! 2), @@ -1758,7 +1758,7 @@ definition capAllowWrite_CL = from_bool (wd !! 0)\" definition - cap_rights_from_word :: "word32 \ cap_rights_CL" + cap_rights_from_word :: "word32 \ seL4_CapRights_CL" where "cap_rights_from_word wd \ SOME cr. to_bool (capAllowGrant_CL cr) = wd !! 2 \ diff --git a/proof/crefine/StateRelation_C.thy b/proof/crefine/StateRelation_C.thy index 03111d1d1..03bdd4246 100644 --- a/proof/crefine/StateRelation_C.thy +++ b/proof/crefine/StateRelation_C.thy @@ -243,7 +243,7 @@ where primrec cthread_state_relation_lifted :: "Structures_H.thread_state \ - (thread_state_CL \ fault_CL option) \ bool" + (thread_state_CL \ seL4_Fault_CL option) \ bool" where "cthread_state_relation_lifted (Structures_H.Running) ts' = (tsType_CL (fst ts') = scast ThreadState_Running)" @@ -271,17 +271,17 @@ where definition cthread_state_relation :: "Structures_H.thread_state \ - (thread_state_C \ fault_C) \ bool" + (thread_state_C \ seL4_Fault_C) \ bool" where "cthread_state_relation \ \a (cs, cf). - cthread_state_relation_lifted a (thread_state_lift cs, fault_lift cf)" + cthread_state_relation_lifted a (thread_state_lift cs, seL4_Fault_lift cf)" definition "is_cap_fault cf \ - (case cf of (Fault_cap_fault _) \ True + (case cf of (SeL4_Fault_CapFault _) \ True | _ \ False)" -lemma is_cap_fault_simp: "is_cap_fault cf = (\ x. cf=Fault_cap_fault x)" - by (simp add: is_cap_fault_def split:fault_CL.splits) +lemma is_cap_fault_simp: "is_cap_fault cf = (\ x. cf=SeL4_Fault_CapFault x)" + by (simp add: is_cap_fault_def split:seL4_Fault_CL.splits) definition @@ -306,25 +306,31 @@ fun (MissingCapability (unat (lookup_fault_missing_capability_CL.bitsLeft_CL lf)))" fun - fault_to_H :: "fault_CL \ lookup_fault_CL \ fault option" + fault_to_H :: "seL4_Fault_CL \ lookup_fault_CL \ fault option" where - "fault_to_H Fault_null_fault lf = None" - | "fault_to_H (Fault_cap_fault cf) lf - = Some (CapFault (fault_cap_fault_CL.address_CL cf) (to_bool (inReceivePhase_CL cf)) (lookup_fault_to_H lf))" - | "fault_to_H (Fault_vm_fault vf) lf - = Some (VMFault (fault_vm_fault_CL.address_CL vf) [instructionFault_CL vf, FSR_CL vf])" - | "fault_to_H (Fault_unknown_syscall us) lf + "fault_to_H SeL4_Fault_NullFault lf = None" + | "fault_to_H (SeL4_Fault_CapFault cf) lf + = Some (CapFault (seL4_Fault_CapFault_CL.address_CL cf) (to_bool (inReceivePhase_CL cf)) (lookup_fault_to_H lf))" + | "fault_to_H (SeL4_Fault_VMFault vf) lf + = Some (ArchFault (VMFault (seL4_Fault_VMFault_CL.address_CL vf) [instructionFault_CL vf, FSR_CL vf]))" + | "fault_to_H (SeL4_Fault_UnknownSyscall us) lf = Some (UnknownSyscallException (syscallNumber_CL us))" - | "fault_to_H (Fault_user_exception ue) lf + | "fault_to_H (SeL4_Fault_UserException ue) lf = Some (UserException (number_CL ue) (code_CL ue))" definition - cfault_rel :: "Fault_H.fault option \ fault_CL option \ lookup_fault_CL option \ bool" + cfault_rel :: "Fault_H.fault option \ seL4_Fault_CL option \ lookup_fault_CL option \ bool" where "cfault_rel af cf lf \ \cf'. cf = Some cf' \ (if (is_cap_fault cf') then (\lf'. lf = Some lf' \ fault_to_H cf' lf' = af) else (fault_to_H cf' undefined = af))" +definition + carch_tcb_relation :: "Structures_H.arch_tcb \ arch_tcb_C \ bool" +where + "carch_tcb_relation aarch_tcb carch_tcb \ + ccontext_relation (atcbContextGet aarch_tcb) (tcbContext_C carch_tcb)" + definition ctcb_relation :: "Structures_H.tcb \ tcb_C \ bool" where @@ -332,13 +338,13 @@ where tcbFaultHandler atcb = tcbFaultHandler_C ctcb \ cthread_state_relation (tcbState atcb) (tcbState_C ctcb, tcbFault_C ctcb) \ tcbIPCBuffer atcb = tcbIPCBuffer_C ctcb - \ ccontext_relation (tcbContext atcb) (tcbContext_C (tcbArch_C ctcb)) + \ carch_tcb_relation (tcbArch atcb) (tcbArch_C ctcb) \ tcbQueued atcb = to_bool (tcbQueued_CL (thread_state_lift (tcbState_C ctcb))) \ ucast (tcbDomain atcb) = tcbDomain_C ctcb \ ucast (tcbPriority atcb) = tcbPriority_C ctcb \ ucast (tcbMCP atcb) = tcbMCP_C ctcb \ tcbTimeSlice atcb = unat (tcbTimeSlice_C ctcb) - \ cfault_rel (tcbFault atcb) (fault_lift (tcbFault_C ctcb)) + \ cfault_rel (tcbFault atcb) (seL4_Fault_lift (tcbFault_C ctcb)) (lookup_fault_lift (tcbLookupFailure_C ctcb)) \ option_to_ptr (tcbBoundNotification atcb) = tcbBoundNotification_C ctcb" @@ -614,7 +620,7 @@ fun "irqstate_to_C IRQInactive = scast Kernel_C.IRQInactive" | "irqstate_to_C IRQSignal = scast Kernel_C.IRQSignal" | "irqstate_to_C IRQTimer = scast Kernel_C.IRQTimer" - + | "irqstate_to_C irqstate.IRQReserved = scast Kernel_C.IRQReserved" definition cinterrupt_relation :: "interrupt_state \ cte_C ptr \ (word32[160]) \ bool" @@ -746,19 +752,27 @@ where lemma ccap_relation_c_valid_cap: "ccap_relation c c' \ c_valid_cap c'" by (simp add: ccap_relation_def) +context begin interpretation Arch . +fun + arch_fault_to_fault_tag :: "arch_fault \ word32" + where + "arch_fault_to_fault_tag (VMFault a b) = scast seL4_Fault_VMFault" +end + + fun fault_to_fault_tag :: "fault \ word32" - where - "fault_to_fault_tag (CapFault a b c) = scast fault_cap_fault" - | "fault_to_fault_tag (VMFault a b) = scast fault_vm_fault" - | "fault_to_fault_tag (UnknownSyscallException a) = scast fault_unknown_syscall" - | "fault_to_fault_tag (UserException a b) = scast fault_user_exception" +where + " fault_to_fault_tag (CapFault a b c) = scast seL4_Fault_CapFault" + | "fault_to_fault_tag (ArchFault f) = arch_fault_to_fault_tag f" + | "fault_to_fault_tag (UnknownSyscallException a) = scast seL4_Fault_UnknownSyscall" + | "fault_to_fault_tag (UserException a b) = scast seL4_Fault_UserException" (* Return relations *) record errtype = - errfault :: "fault_CL option" + errfault :: "seL4_Fault_CL option" errlookup_fault :: "lookup_fault_CL option" errsyscall :: syscall_error_C @@ -843,7 +857,7 @@ definition (to_bool (capAllowGrant_CL rs))" definition - "ccap_rights_relation cr cr' \ cr = cap_rights_to_H (cap_rights_lift cr')" + "ccap_rights_relation cr cr' \ cr = cap_rights_to_H (seL4_CapRights_lift cr')" lemma (in kernel) syscall_error_to_H_cases_rev: "\n. syscall_error_to_H e lf = Some (InvalidArgument n) \ diff --git a/proof/crefine/SyscallArgs_C.thy b/proof/crefine/SyscallArgs_C.thy index 32c4333a5..7cc15d15a 100644 --- a/proof/crefine/SyscallArgs_C.thy +++ b/proof/crefine/SyscallArgs_C.thy @@ -639,8 +639,8 @@ lemma msgRegisters_ccorres: lemma asUser_cur_obj_at': assumes f: "\P\ f \Q\" - shows "\\s. obj_at' (\tcb. P (tcbContext tcb)) (ksCurThread s) s \ t = ksCurThread s\ - asUser t f \\rv s. obj_at' (\tcb. Q rv (tcbContext tcb)) (ksCurThread s) s\" + shows "\\s. obj_at' (\tcb. P (atcbContextGet (tcbArch tcb))) (ksCurThread s) s \ t = ksCurThread s\ + asUser t f \\rv s. obj_at' (\tcb. Q rv (atcbContextGet (tcbArch tcb))) (ksCurThread s) s\" apply (simp add: asUser_def split_def) apply (wp) apply (rule hoare_lift_Pf2 [where f=ksCurThread]) @@ -678,7 +678,7 @@ lemma asUser_const_rv: lemma getMRs_tcbContext: "\\s. n < unat n_msgRegisters \ n < unat (msgLength info) \ thread = ksCurThread s \ cur_tcb' s\ getMRs thread buffer info - \\rv s. obj_at' (\tcb. tcbContext tcb (ARM_H.msgRegisters ! n) = rv ! n) (ksCurThread s) s\" + \\rv s. obj_at' (\tcb. atcbContextGet (tcbArch tcb) (ARM_H.msgRegisters ! n) = rv ! n) (ksCurThread s) s\" apply (rule hoare_assume_pre) apply (elim conjE) apply (thin_tac "thread = t" for t) @@ -1276,7 +1276,7 @@ lemma getSyscallArg_ccorres_foo: apply (simp add: word_less_nat_alt split: split_if) apply (rule ccorres_add_return2) apply (rule ccorres_symb_exec_l) - apply (rule_tac P="\s. n < unat (scast n_msgRegisters :: word32) \ obj_at' (\tcb. tcbContext tcb (ARM_H.msgRegisters!n) = x!n) (ksCurThread s) s" + apply (rule_tac P="\s. n < unat (scast n_msgRegisters :: word32) \ obj_at' (\tcb. atcbContextGet (tcbArch tcb) (ARM_H.msgRegisters!n) = x!n) (ksCurThread s) s" and P' = UNIV in ccorres_from_vcg_split_throws) apply vcg @@ -1285,7 +1285,9 @@ lemma getSyscallArg_ccorres_foo: apply (clarsimp simp: rf_sr_ksCurThread) apply (drule (1) obj_at_cslift_tcb) apply (clarsimp simp: typ_heap_simps' msgRegisters_scast) - apply (clarsimp simp: ctcb_relation_def ccontext_relation_def msgRegisters_ccorres) + apply (clarsimp simp: ctcb_relation_def ccontext_relation_def + msgRegisters_ccorres atcbContextGet_def + carch_tcb_relation_def) apply (subst (asm) msgRegisters_ccorres) apply (clarsimp simp: n_msgRegisters_def) apply (simp add: n_msgRegisters_def word_less_nat_alt) diff --git a/proof/crefine/Syscall_C.thy b/proof/crefine/Syscall_C.thy index 42e81d4a3..7c8df9099 100644 --- a/proof/crefine/Syscall_C.thy +++ b/proof/crefine/Syscall_C.thy @@ -595,7 +595,7 @@ lemma sendFaultIPC_ccorres: "ccorres (cfault_rel2 \ dc) (liftxf errstate id (K ()) ret__unsigned_long_') (invs' and st_tcb_at' simple' tptr and sch_act_not tptr and (\s. \p. tptr \ set (ksReadyQueues s p))) - (UNIV \ {s. (cfault_rel (Some fault) (fault_lift(current_fault_' (globals s))) + (UNIV \ {s. (cfault_rel (Some fault) (seL4_Fault_lift(current_fault_' (globals s))) (lookup_fault_lift(current_lookup_fault_' (globals s))))} \ {s. tptr_' s = tcb_ptr_to_ctcb_ptr tptr}) [] (sendFaultIPC tptr fault) @@ -646,7 +646,7 @@ lemma sendFaultIPC_ccorres: apply (rule_tac P=\ and P'=invs' and R="{s. (cfault_rel (Some fault) - (fault_lift(current_fault_' (globals s))) + (seL4_Fault_lift(current_fault_' (globals s))) (lookup_fault_lift(original_lookup_fault_' s)))}" in threadSet_ccorres_lemma4) apply vcg @@ -669,7 +669,7 @@ lemma sendFaultIPC_ccorres: apply (case_tac "tcbState tcb", simp+) apply (simp add: cfault_rel_def) apply (clarsimp) - apply (clarsimp simp: fault_lift_def Let_def is_cap_fault_def + apply (clarsimp simp: seL4_Fault_lift_def Let_def is_cap_fault_def split: split_if_asm) apply ceqv @@ -695,7 +695,7 @@ lemma sendFaultIPC_ccorres: apply (rule conseqPre, vcg) apply (clarsimp simp: EXCEPTION_FAULT_def EXCEPTION_NONE_def) apply (simp add: cfault_rel2_def cfault_rel_def EXCEPTION_FAULT_def) - apply (simp add: fault_cap_fault_lift) + apply (simp add: seL4_Fault_CapFault_lift) apply (simp add: lookup_fault_missing_capability_lift is_cap_fault_def) apply vcg @@ -713,7 +713,7 @@ lemma sendFaultIPC_ccorres: apply (rule conseqPre, vcg) apply (clarsimp simp: EXCEPTION_FAULT_def EXCEPTION_NONE_def) apply (simp add: cfault_rel2_def cfault_rel_def EXCEPTION_FAULT_def) - apply (simp add: fault_cap_fault_lift is_cap_fault_def) + apply (simp add: seL4_Fault_CapFault_lift is_cap_fault_def) apply (erule lookup_failure_rel_fault_lift [rotated, unfolded EXCEPTION_NONE_def, simplified], assumption) @@ -746,7 +746,7 @@ lemma sendFaultIPC_ccorres: lemma handleFault_ccorres: "ccorres dc xfdc (invs' and st_tcb_at' simple' t and sch_act_not t and (\s. \p. t \ set (ksReadyQueues s p))) - (UNIV \ {s. (cfault_rel (Some flt) (fault_lift(current_fault_' (globals s))) + (UNIV \ {s. (cfault_rel (Some flt) (seL4_Fault_lift(current_fault_' (globals s))) (lookup_fault_lift(current_lookup_fault_' (globals s))) )} \ {s. tptr_' s = tcb_ptr_to_ctcb_ptr t}) [] (handleFault t flt) @@ -1038,7 +1038,7 @@ lemma handleInvocation_ccorres: Kernel_C.capRegister_def ARM_H.capRegister_def ARM.capRegister_def Kernel_C.R0_def) apply (clarsimp simp: cfault_rel_def option_to_ptr_def) - apply (simp add: fault_cap_fault_lift is_cap_fault_def) + apply (simp add: seL4_Fault_CapFault_lift is_cap_fault_def) apply (frule lookup_failure_rel_fault_lift, assumption) apply clarsimp apply (clarsimp simp: ct_in_state'_def pred_tcb_at') @@ -1510,12 +1510,12 @@ lemma handleRecv_ccorres: apply (frule tcb_aligned'[OF tcb_at_invs']) apply clarsimp apply (intro conjI impI allI) - apply (clarsimp simp: cfault_rel_def fault_cap_fault_lift + apply (clarsimp simp: cfault_rel_def seL4_Fault_CapFault_lift lookup_fault_missing_capability_lift is_cap_fault_def)+ apply (clarsimp simp: cap_get_tag_NotificationCap) apply (rule cmap_relationE1[OF cmap_relation_ntfn], assumption, erule ko_at_projectKO_opt) apply (clarsimp simp: cnotification_relation_def Let_def) - apply (clarsimp simp: cfault_rel_def fault_cap_fault_lift + apply (clarsimp simp: cfault_rel_def seL4_Fault_CapFault_lift lookup_fault_missing_capability_lift is_cap_fault_def)+ apply (clarsimp simp: cap_get_tag_NotificationCap) apply (simp add: ccap_relation_def to_bool_def) @@ -1523,7 +1523,7 @@ lemma handleRecv_ccorres: apply (drule obj_at_ko_at', clarsimp) apply (rule cmap_relationE1[OF cmap_relation_ntfn], assumption, erule ko_at_projectKO_opt) apply (clarsimp simp: typ_heap_simps) - apply (clarsimp simp: cfault_rel_def fault_cap_fault_lift + apply (clarsimp simp: cfault_rel_def seL4_Fault_CapFault_lift lookup_fault_missing_capability_lift is_cap_fault_def)+ apply (case_tac w, clarsimp+) done @@ -1780,7 +1780,14 @@ lemma scast_maxIRQ_is_not_less: "(\ (Kernel_C.maxIRQ) (uc (simp add: is_up_def target_size source_size)?) apply (uint_arith) done - + +lemma ccorres_handleReserveIRQ: + "ccorres dc xfdc \ UNIV hs (handleReservedIRQ irq) (Call handleReservedIRQ_'proc)" + apply cinit + apply (rule ccorres_return_Skip) + apply simp + done + lemma handleInterrupt_ccorres: "ccorres dc xfdc (invs') @@ -1810,71 +1817,80 @@ lemma handleInterrupt_ccorres: apply (vcg exspec=maskInterrupt_modifies) apply (simp add: scast_maxIRQ_is_not_less Platform_maxIRQ del: Collect_const) apply (rule ccorres_pre_getIRQState) - apply wpc - apply simp - apply (rule ccorres_fail) - apply (simp add: bind_assoc cong: call_ignore_cong) - apply (rule ccorres_move_const_guards)+ - apply (rule ccorres_cond_true_seq) - apply (rule ccorres_rhs_assoc)+ - apply csymbr - apply (rule getIRQSlot_ccorres3) - apply (rule ccorres_getSlotCap_cte_at) - apply (rule_tac P="cte_at' rv" in ccorres_cross_over_guard) - apply (rule ptr_add_assertion_irq_guard[unfolded dc_def]) - apply (rule ccorres_move_array_assertion_irq ccorres_move_c_guard_cte)+ - apply ctac - apply csymbr - apply csymbr - apply (rule ccorres_ntfn_cases) - apply (clarsimp cong: call_ignore_cong simp del: Collect_const) - apply (rule_tac b=send in ccorres_case_bools) + apply wpc + apply simp + apply (rule ccorres_fail) + apply (simp add: bind_assoc cong: call_ignore_cong) + apply (rule ccorres_move_const_guards)+ + apply (rule ccorres_cond_true_seq) + apply (rule ccorres_rhs_assoc)+ + apply csymbr + apply (rule getIRQSlot_ccorres3) + apply (rule ccorres_getSlotCap_cte_at) + apply (rule_tac P="cte_at' rv" in ccorres_cross_over_guard) + apply (rule ptr_add_assertion_irq_guard[unfolded dc_def]) + apply (rule ccorres_move_array_assertion_irq ccorres_move_c_guard_cte)+ + apply ctac + apply csymbr + apply csymbr + apply (rule ccorres_ntfn_cases) + apply (clarsimp cong: call_ignore_cong simp del: Collect_const) + apply (rule_tac b=send in ccorres_case_bools) + apply simp + apply (rule ccorres_cond_true_seq) + apply (rule ccorres_rhs_assoc)+ + apply csymbr + apply csymbr + apply (rule ccorres_cond_true_seq) + apply (rule ccorres_rhs_assoc)+ + apply csymbr + apply csymbr + apply (ctac (no_vcg) add: sendSignal_ccorres) + apply (ctac (no_vcg) add: maskInterrupt_ccorres) + apply (ctac add: ackInterrupt_ccorres [unfolded dc_def]) + apply wp + apply (simp del: Collect_const) + apply (rule ccorres_cond_true_seq) + apply (rule ccorres_rhs_assoc)+ + apply csymbr+ + apply (rule ccorres_cond_false_seq) apply simp - apply (rule ccorres_cond_true_seq) - apply (rule ccorres_rhs_assoc)+ - apply csymbr - apply csymbr - apply (rule ccorres_cond_true_seq) - apply (rule ccorres_rhs_assoc)+ - apply csymbr - apply csymbr - apply (ctac (no_vcg) add: sendSignal_ccorres) - apply (ctac (no_vcg) add: maskInterrupt_ccorres) - apply (ctac add: ackInterrupt_ccorres [unfolded dc_def]) - apply wp - apply (simp del: Collect_const) - apply (rule ccorres_cond_true_seq) - apply (rule ccorres_rhs_assoc)+ - apply csymbr+ - apply (rule ccorres_cond_false_seq) - apply simp - apply (ctac (no_vcg) add: maskInterrupt_ccorres) - apply (ctac add: ackInterrupt_ccorres [unfolded dc_def]) - apply wp - apply (rule_tac P=\ and P'="{s. ret__int_' s = 0 \ cap_get_tag cap \ scast cap_notification_cap}" in ccorres_inst) - apply (clarsimp simp: isCap_simps simp del: Collect_const) - apply (case_tac rva, simp_all del: Collect_const)[1] - prefer 3 - apply metis - apply ((rule ccorres_guard_imp2, - rule ccorres_cond_false_seq, simp, - rule ccorres_cond_false_seq, simp, - ctac (no_vcg) add: maskInterrupt_ccorres, - ctac (no_vcg) add: ackInterrupt_ccorres [unfolded dc_def], - wp, simp)+) - apply (wp getSlotCap_wp) - apply simp - apply vcg - apply (simp add: bind_assoc) + apply (ctac (no_vcg) add: maskInterrupt_ccorres) + apply (ctac add: ackInterrupt_ccorres [unfolded dc_def]) + apply wp + apply (rule_tac P=\ and P'="{s. ret__int_' s = 0 \ cap_get_tag cap \ scast cap_notification_cap}" in ccorres_inst) + apply (clarsimp simp: isCap_simps simp del: Collect_const) + apply (case_tac rva, simp_all del: Collect_const)[1] + prefer 3 + apply metis + apply ((rule ccorres_guard_imp2, + rule ccorres_cond_false_seq, simp, + rule ccorres_cond_false_seq, simp, + ctac (no_vcg) add: maskInterrupt_ccorres, + ctac (no_vcg) add: ackInterrupt_ccorres [unfolded dc_def], + wp, simp)+) + apply (wp getSlotCap_wp) + apply simp + apply vcg + apply (simp add: bind_assoc) + apply (rule ccorres_move_const_guards)+ + apply (rule ccorres_cond_false_seq) + apply (rule ccorres_cond_true_seq) + apply (fold dc_def)[1] + apply (rule ccorres_rhs_assoc)+ + apply (ctac (no_vcg) add: timerTick_ccorres) + apply (ctac (no_vcg) add: resetTimer_ccorres) + apply (ctac add: ackInterrupt_ccorres ) + apply wp + apply (simp add: Platform_maxIRQ maxIRQ_def del: Collect_const) apply (rule ccorres_move_const_guards)+ apply (rule ccorres_cond_false_seq) + apply (rule ccorres_cond_false_seq) apply (rule ccorres_cond_true_seq) - apply (fold dc_def)[1] - apply (rule ccorres_rhs_assoc)+ - apply (ctac (no_vcg) add: timerTick_ccorres) - apply (ctac (no_vcg) add: resetTimer_ccorres) - apply (ctac add: ackInterrupt_ccorres ) + apply (ctac add: ccorres_handleReserveIRQ) + apply (ctac (no_vcg) add: ackInterrupt_ccorres [unfolded dc_def]) apply wp + apply vcg apply (simp add: sint_ucast_eq_uint is_down uint_up_ucast is_up ) apply (clarsimp simp: word_sless_alt word_less_alt word_le_def Kernel_C.maxIRQ_def uint_up_ucast is_up_def @@ -1886,21 +1902,23 @@ lemma handleInterrupt_ccorres: apply (clarsimp simp: Kernel_C.IRQTimer_def Kernel_C.IRQSignal_def cte_wp_at_ctes_of ucast_ucast_b is_up) apply (intro conjI impI) - apply clarsimp - apply (erule(1) cmap_relationE1[OF cmap_relation_cte]) - apply (clarsimp simp: typ_heap_simps') - apply (simp add: cap_get_tag_isCap) - apply (clarsimp simp: isCap_simps) - apply (frule cap_get_tag_isCap_unfolded_H_cap) - apply (frule cap_get_tag_to_H, assumption) - apply (clarsimp simp: to_bool_def) - apply (cut_tac un_ui_le[where b = 159 and a = irq, - simplified word_size]) - apply (simp add: ucast_eq_0 is_up_def source_size_def - target_size_def word_size unat_gt_0 - | subst array_assertion_abs_irq[rule_format, OF conjI])+ - apply (erule(1) rf_sr_cte_at_valid[OF ctes_of_cte_at]) - apply (clarsimp simp:nat_le_iff) + apply clarsimp + apply (erule(1) cmap_relationE1[OF cmap_relation_cte]) + apply (clarsimp simp: typ_heap_simps') + apply (simp add: cap_get_tag_isCap) + apply (clarsimp simp: isCap_simps) + apply (frule cap_get_tag_isCap_unfolded_H_cap) + apply (frule cap_get_tag_to_H, assumption) + apply (clarsimp simp: to_bool_def) + apply (cut_tac un_ui_le[where b = 159 and a = irq, + simplified word_size]) + apply (simp add: ucast_eq_0 is_up_def source_size_def + target_size_def word_size unat_gt_0 + | subst array_assertion_abs_irq[rule_format, OF conjI])+ + apply (erule exE conjE)+ + apply (erule(1) rf_sr_cte_at_valid[OF ctes_of_cte_at]) + apply (clarsimp simp:nat_le_iff) + apply (clarsimp simp: IRQReserved_def)+ done end end diff --git a/proof/crefine/TcbAcc_C.thy b/proof/crefine/TcbAcc_C.thy index d03b5e7ce..e039792ee 100644 --- a/proof/crefine/TcbAcc_C.thy +++ b/proof/crefine/TcbAcc_C.thy @@ -98,7 +98,7 @@ lemma getRegister_ccorres [corres]: (asUser thread (getRegister reg)) (Call getRegister_'proc)" apply (unfold asUser_def) apply (rule ccorres_guard_imp) - apply (rule ccorres_symb_exec_l [where Q="\u. obj_at' (\t. tcbContext t = u) thread" and + apply (rule ccorres_symb_exec_l [where Q="\u. obj_at' (\t. (atcbContextGet o tcbArch) t = u) thread" and Q'="\rv. {s. thread_' s = tcb_ptr_to_ctcb_ptr thread} \ {s. reg_' s = register_from_H reg}"]) apply (rule ccorres_from_vcg) apply (rule allI, rule conseqPre) @@ -107,13 +107,13 @@ lemma getRegister_ccorres [corres]: apply (drule (1) obj_at_cslift_tcb) apply (clarsimp simp: typ_heap_simps register_from_H_less register_from_H_sless) apply (clarsimp simp: getRegister_def typ_heap_simps) - apply (rule_tac x = "(tcbContext ko reg, \)" in bexI [rotated]) + apply (rule_tac x = "((atcbContextGet o tcbArch) ko reg, \)" in bexI [rotated]) apply (simp add: in_monad' asUser_def select_f_def split_def) apply (subst arg_cong2 [where f = "op \"]) defer apply (rule refl) apply (erule threadSet_eq) - apply (clarsimp simp: ctcb_relation_def ccontext_relation_def) + apply (clarsimp simp: ctcb_relation_def ccontext_relation_def carch_tcb_relation_def) apply (wp threadGet_obj_at2) apply simp apply simp diff --git a/proof/crefine/Tcb_C.thy b/proof/crefine/Tcb_C.thy index 88ef56704..eade9687b 100644 --- a/proof/crefine/Tcb_C.thy +++ b/proof/crefine/Tcb_C.thy @@ -172,9 +172,9 @@ lemma getObject_state: lemma threadGet_state: - "\ (uc, s') \ fst (threadGet tcbContext t' s); ko_at' ko t s \ \ + "\ (uc, s') \ fst (threadGet (atcbContextGet o tcbArch) t' s); ko_at' ko t s \ \ (uc, s'\ksPSpace := ksPSpace s(t \ KOTCB (tcbState_update (\_. st) ko))\) \ - fst (threadGet tcbContext t' (s\ksPSpace := ksPSpace s(t \ KOTCB (tcbState_update (\_. st) ko))\))" + fst (threadGet (atcbContextGet o tcbArch) t' (s\ksPSpace := ksPSpace s(t \ KOTCB (tcbState_update (\_. st) ko))\))" apply (clarsimp simp: threadGet_def liftM_def in_monad) apply (drule (1) getObject_state [where st=st]) apply (rule exI) @@ -1229,10 +1229,10 @@ lemma getObject_context: done lemma threadGet_context: - "\ (uc, s') \ fst (threadGet tcbContext (ksCurThread s) s); ko_at' ko t s; + "\ (uc, s') \ fst (threadGet (atcbContextGet o tcbArch) (ksCurThread s) s); ko_at' ko t s; t \ ksCurThread s \ \ - (uc, s'\ksPSpace := ksPSpace s(t \ KOTCB (tcbContext_update (\_. st) ko))\) \ - fst (threadGet tcbContext (ksCurThread s) (s\ksPSpace := ksPSpace s(t \ KOTCB (tcbContext_update (\_. st) ko))\))" + (uc, s'\ksPSpace := ksPSpace s(t \ KOTCB (tcbArch_update (\_. atcbContextSet st (tcbArch ko)) ko))\) \ + fst (threadGet (atcbContextGet o tcbArch) (ksCurThread s) (s\ksPSpace := ksPSpace s(t \ KOTCB (tcbArch_update (\_. atcbContextSet st (tcbArch ko)) ko))\))" apply (clarsimp simp: threadGet_def liftM_def in_monad) apply (drule (1) getObject_context [where st=st]) apply (rule exI) @@ -1244,8 +1244,8 @@ done lemma asUser_context: "\(x,s) \ fst (asUser (ksCurThread s) f s); ko_at' ko t s; \s. \op = s\ f \\_. op = s\ ; t \ ksCurThread s\ \ - (x,s\ksPSpace := ksPSpace s(t \ KOTCB (tcbContext_update (\_. st) ko))\) \ - fst (asUser (ksCurThread s) f (s\ksPSpace := ksPSpace s(t \ KOTCB (tcbContext_update (\_. st) ko))\))" + (x,s\ksPSpace := ksPSpace s(t \ KOTCB (tcbArch_update (\_. atcbContextSet st (tcbArch ko)) ko))\) \ + fst (asUser (ksCurThread s) f (s\ksPSpace := ksPSpace s(t \ KOTCB (tcbArch_update (\_. atcbContextSet st (tcbArch ko)) ko))\))" apply (clarsimp simp: asUser_def in_monad select_f_def) apply (frule use_valid, rule threadGet_inv [where P="op = s"], rule refl) apply (frule use_valid, assumption, rule refl) @@ -1316,7 +1316,7 @@ lemma getMRs_rel_context: "\getMRs_rel args buffer s; (cur_tcb' and case_option \ valid_ipc_buffer_ptr' buffer) s; ko_at' ko t s ; t \ ksCurThread s\ \ - getMRs_rel args buffer (s\ksPSpace := ksPSpace s(t \ KOTCB (tcbContext_update (\_. st) ko))\)" + getMRs_rel args buffer (s\ksPSpace := ksPSpace s(t \ KOTCB (tcbArch_update (\_. atcbContextSet st (tcbArch ko)) ko))\)" apply (clarsimp simp: getMRs_rel_def) apply (rule exI, erule conjI) apply (subst (asm) det_wp_use, rule det_wp_getMRs) @@ -1677,7 +1677,7 @@ shows = min (unat n) (unat n_frameRegisters + unat n_gpRegisters)" in ccorres_gen_asm) apply (rule ccorres_split_nothrow_novcg) - apply (rule_tac F="\m s. obj_at' (\tcb. map (tcbContext tcb) (genericTake n + apply (rule_tac F="\m s. obj_at' (\tcb. map ((atcbContextGet o tcbArch) tcb) (genericTake n (ARM_H.frameRegisters @ ARM_H.gpRegisters)) = reply) target s" in ccorres_mapM_x_while) @@ -1744,7 +1744,7 @@ shows apply (rule ccorres_Cond_rhs) apply (simp del: Collect_const) apply (rule ccorres_rel_imp, - rule_tac F="\m s. obj_at' (\tcb. map (tcbContext tcb) (genericTake n + rule_tac F="\m s. obj_at' (\tcb. map ((atcbContextGet o tcbArch) tcb) (genericTake n (ARM_H.frameRegisters @ ARM_H.gpRegisters)) = reply) target s \ valid_ipc_buffer_ptr' (the rva) s @@ -1858,7 +1858,7 @@ shows apply (simp add: drop_zip del: Collect_const) apply (rule ccorres_Cond_rhs) apply (simp del: Collect_const) - apply (rule_tac F="\m s. obj_at' (\tcb. map (tcbContext tcb) (genericTake n + apply (rule_tac F="\m s. obj_at' (\tcb. map ((atcbContextGet o tcbArch) tcb) (genericTake n (ARM_H.frameRegisters @ ARM_H.gpRegisters)) = reply) target s \ valid_ipc_buffer_ptr' (the rva) s \ valid_pspace' s" @@ -2007,8 +2007,8 @@ shows apply wp apply (simp cong: rev_conj_cong) apply wp - apply (wp asUser_inv mapM_wp' getRegister_inv - asUser_get_registers static_imp_wp) + apply (wp asUser_inv mapM_wp' getRegister_inv + asUser_get_registers[simplified] static_imp_wp) apply (rule hoare_strengthen_post, rule asUser_get_registers) apply (clarsimp simp: obj_at'_def genericTake_def frame_gp_registers_convs) diff --git a/proof/crefine/VSpace_C.thy b/proof/crefine/VSpace_C.thy index c72e036fe..3d7ed089c 100644 --- a/proof/crefine/VSpace_C.thy +++ b/proof/crefine/VSpace_C.thy @@ -472,8 +472,8 @@ lemma mask_32_id [simp]: lemma handleVMFault_ccorres: "ccorres ((\a ex v. ex = scast EXCEPTION_FAULT \ (\vf. - a = VMFault (fault_vm_fault_CL.address_CL vf) [instructionFault_CL vf, FSR_CL vf] \ - errfault v = Some (Fault_vm_fault vf))) \ + a = ArchFault (VMFault (seL4_Fault_VMFault_CL.address_CL vf) [instructionFault_CL vf, FSR_CL vf]) \ + errfault v = Some (SeL4_Fault_VMFault vf))) \ (\_. \)) (liftxf errstate id (K ()) ret__unsigned_long_') \ @@ -497,7 +497,7 @@ lemma handleVMFault_ccorres: apply vcg apply (clarsimp simp: errstate_def) apply (clarsimp simp: EXCEPTION_FAULT_def EXCEPTION_NONE_def) - apply (simp add: fault_vm_fault_lift false_def) + apply (simp add: seL4_Fault_VMFault_lift false_def) apply wp apply (simp add: vm_fault_type_from_H_def Kernel_C.ARMDataAbort_def Kernel_C.ARMPrefetchAbort_def) apply (simp add: ccorres_cond_univ_iff ccorres_cond_empty_iff) @@ -512,7 +512,7 @@ lemma handleVMFault_ccorres: apply vcg apply (clarsimp simp: errstate_def) apply (clarsimp simp: EXCEPTION_FAULT_def EXCEPTION_NONE_def) - apply (simp add: fault_vm_fault_lift true_def mask_def) + apply (simp add: seL4_Fault_VMFault_lift true_def mask_def) apply wp apply simp done @@ -1845,7 +1845,7 @@ lemma setRegister_ccorres: apply (rule ccorres_pre_threadGet) apply (rule ccorres_Guard) apply (simp add: setRegister_def simpler_modify_def exec_select_f_singleton) - apply (rule_tac P="\tcb. tcbContext tcb = rv" + apply (rule_tac P="\tcb. (atcbContextGet o tcbArch) tcb = rv" in threadSet_ccorres_lemma2 [unfolded dc_def]) apply vcg apply (clarsimp simp: setRegister_def HaskellLib_H.runState_def @@ -1859,7 +1859,9 @@ lemma setRegister_ccorres: (simp add: typ_heap_simps')+) apply (rule ball_tcb_cte_casesI, simp+) apply (clarsimp simp: ctcb_relation_def ccontext_relation_def - split: split_if) + atcbContextSet_def atcbContextGet_def + carch_tcb_relation_def + split: split_if) apply (clarsimp simp: Collect_const_mem register_from_H_sless register_from_H_less) apply (auto intro: typ_heap_simps elim: obj_at'_weakenE)