CRefine: Updating crefine for tcb_arch reserved_irq and arch_fault changes

* The majority of changes involved rephrasing references to tcb_context
  for the equivalente (arch_tcb_context_get o tcb_arch) (the same
  applies for set operations)

* Mayor refactoring of setMRs and handleFaultReply C functions both of
  which introduced new function calls when dealing with UserException
  and SyscallException faults

* handleFaultReply' (definition) was modified to include less redundant
  code

* handleFaultReply' (lemma) was completely refactored to use the
  monadicrewrite framework for "easier" and simpler proofs

* IsolatedThreadAction framework striped from Fastpath_C into new
  standalone file

  tags: [VER-623][SELFOUR-413]
This commit is contained in:
Alejandro Gomez-Londono 2016-11-10 15:44:36 +11:00
parent 879ac30249
commit 9a166de8bc
19 changed files with 3199 additions and 2638 deletions

View File

@ -75,7 +75,7 @@ lemma Basic_sem_eq:
lemma setArchTCB_C_corres:
"\<lbrakk> ccontext_relation tc (tcbContext_C tc'); t' = tcb_ptr_to_ctcb_ptr t \<rbrakk> \<Longrightarrow>
corres_underlying rf_sr nf nf' dc (tcb_at' t) \<top>
(threadSet (\<lambda>tcb. tcb \<lparr> tcbContext := tc \<rparr>) t) (setArchTCB_C tc' t')"
(threadSet (\<lambda>tcb. tcb \<lparr> tcbArch := atcbContextSet tc (tcbArch tcb)\<rparr>) 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 \<equiv>
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 \<circ> 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 \<Longrightarrow> ccontext_relation g x \<Longrightarrow> f=g"
by (rule ext) (simp add: ccontext_relation_def)
lemma carch_tcb_relation_imp_eq:
"carch_tcb_relation f x \<Longrightarrow> carch_tcb_relation g x \<Longrightarrow> 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 \<Longrightarrow> cfault_rel y a b \<Longrightarrow> 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 \<Longrightarrow> cthread_state_relation y z \<Longrightarrow> 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

View File

@ -126,7 +126,7 @@ lemma asUser_mapM_x:
lemma asUser_get_registers:
"\<lbrace>tcb_at' target\<rbrace>
asUser target (mapM getRegister xs)
\<lbrace>\<lambda>rv s. obj_at' (\<lambda>tcb. map (tcbContext tcb) xs = rv) target s\<rbrace>"
\<lbrace>\<lambda>rv s. obj_at' (\<lambda>tcb. map ((atcbContextGet o tcbArch) tcb) xs = rv) target s\<rbrace>"
apply (induct xs)
apply (simp add: mapM_empty asUser_return)
apply wp

View File

@ -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 \<equiv> (cap_rights_lift \<^bsup>s\<^esup>cap_rights)"
shows "\<forall>s. \<Gamma> \<turnstile> {s} \<acute>ret__unsigned_long :== PROC wordFromRights(\<acute>cap_rights)
defines "crl s \<equiv> (seL4_CapRights_lift \<^bsup>s\<^esup>seL4_CapRights)"
shows "\<forall>s. \<Gamma> \<turnstile> {s} \<acute>ret__unsigned_long :== PROC wordFromRights(\<acute>seL4_CapRights)
\<lbrace> \<acute>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

View File

@ -59,7 +59,7 @@ lemma maskVMRights_spec:
\<lbrace> \<acute>vm_rights && mask 2 = \<acute>vm_rights \<rbrace>)
Call maskVMRights_'proc
\<lbrace> vmrights_to_H \<acute>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)) \<and>
maskVMRights (vmrights_to_H \<^bsup>s\<^esup>vm_rights) (cap_rights_to_H (seL4_CapRights_lift \<^bsup>s\<^esup>cap_rights_mask)) \<and>
\<acute>ret__unsigned_long && mask 2 = \<acute>ret__unsigned_long \<rbrace>"
apply vcg
apply clarsimp
@ -72,7 +72,7 @@ lemma maskVMRights_spec:
apply clarsimp
apply (subgoal_tac "vm_rights = 0 \<or> vm_rights = 1 \<or> vm_rights = 2 \<or> 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) \<Longrightarrow>
@ -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

View File

@ -605,17 +605,17 @@ abbreviation
lemma rightsFromWord_spec:
shows "\<forall>s. \<Gamma> \<turnstile> {s} \<acute>ret__struct_cap_rights_C :== PROC rightsFromWord(\<acute>w)
\<lbrace>cap_rights_lift \<acute>ret__struct_cap_rights_C = cap_rights_from_word_canon \<^bsup>s\<^esup>w \<rbrace>"
shows "\<forall>s. \<Gamma> \<turnstile> {s} \<acute>ret__struct_seL4_CapRights_C :== PROC rightsFromWord(\<acute>w)
\<lbrace>seL4_CapRights_lift \<acute>ret__struct_seL4_CapRights_C = cap_rights_from_word_canon \<^bsup>s\<^esup>w \<rbrace>"
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' \<equiv> \<lambda>(cte, rm) (cte', rm'). cte' = Ptr cte \<and> rm = cap_rights_to_H (cap_rights_lift rm')"
"lookupSlot_rel' \<equiv> \<lambda>(cte, rm) (cte', rm'). cte' = Ptr cte \<and> rm = cap_rights_to_H (seL4_CapRights_lift rm')"
(* MOVE *)
lemma cap_rights_to_H_from_word_canon [simp]:

File diff suppressed because it is too large Load Diff

View File

@ -2875,7 +2875,7 @@ lemma cancelIPC_ccorres_reply_helper:
cteDeleteOne callerCap
od)
od)
(\<acute>ret__struct_fault_C :== CALL fault_null_fault_new();;
(\<acute>ret__struct_fault_C :== CALL seL4_Fault_NullFault_new();;
Guard C_Guard {s. s \<Turnstile>\<^sub>c tptr_' s}
(Basic (\<lambda>s. s \<lparr> (\<acute>t_hrs :== (hrs_mem_update (heap_update
(Ptr &(\<acute>tptr\<rightarrow>[''tcbFault_C''])) \<acute>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

File diff suppressed because it is too large Load Diff

File diff suppressed because it is too large Load Diff

View File

@ -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 (\<lambda>x. 0))))
(thread_state_C (FCP (\<lambda>x. 0)))
(NULL)
(fault_C (FCP (\<lambda>x. 0)))
(seL4_Fault_C (FCP (\<lambda>x. 0)))
(lookup_fault_C (FCP (\<lambda>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

View File

@ -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:
"\<lbrakk> \<And>cptr msgInfo. ccorres dc xfdc
((\<lambda>s. P s \<and> Q s \<and>
obj_at' (\<lambda>tcb. tcbContext tcb ARM_H.capRegister = cptr
\<and> tcbContext tcb ARM_H.msgInfoRegister = msgInfo)
obj_at' (\<lambda>tcb. (atcbContextGet o tcbArch) tcb ARM_H.capRegister = cptr
\<and> (atcbContextGet o tcbArch) tcb ARM_H.msgInfoRegister = msgInfo)
(ksCurThread s) s) and R)
(UNIV \<inter> \<lbrace>\<acute>cptr = cptr\<rbrace> \<inter> \<lbrace>\<acute>msgInfo = msgInfo\<rbrace>) [] m c \<rbrakk>
\<Longrightarrow>
@ -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':
"\<lbrace>all_invs' e and (\<lambda>s. t = ksCurThread s)\<rbrace>
threadSet (tcbContext_update f) t \<lbrace>\<lambda>_. all_invs' e\<rbrace>"
threadSet (\<lambda>tcb. tcbArch_update (\<lambda>_. atcbContextSet f (tcbArch tcb)) tcb) t \<lbrace>\<lambda>_. all_invs' e\<rbrace>"
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 \<Longrightarrow>
corres_underlying rf_sr False True (op =) (tcb_at' t) \<top>
(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:

View File

@ -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
(\<lambda>_. foldr (\<lambda>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

View File

@ -356,9 +356,9 @@ lemma updateObject_cte_tcb:
done
definition
tcb_no_ctes_proj :: "tcb \<Rightarrow> Structures_H.thread_state \<times> word32 \<times> word32 \<times> (MachineTypes.register \<Rightarrow> word32) \<times> bool \<times> word8 \<times> word8 \<times> word8 \<times> nat \<times> fault option \<times> word32 option"
tcb_no_ctes_proj :: "tcb \<Rightarrow> Structures_H.thread_state \<times> word32 \<times> word32 \<times> arch_tcb \<times> bool \<times> word8 \<times> word8 \<times> word8 \<times> nat \<times> fault option \<times> word32 option"
where
"tcb_no_ctes_proj t \<equiv> (tcbState t, tcbFaultHandler t, tcbIPCBuffer t, tcbContext t, tcbQueued t,
"tcb_no_ctes_proj t \<equiv> (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 \<Rightarrow> cap_rights_CL"
cap_rights_from_word_canon :: "word32 \<Rightarrow> seL4_CapRights_CL"
where
"cap_rights_from_word_canon wd \<equiv>
\<lparr> capAllowGrant_CL = from_bool (wd !! 2),
@ -1758,7 +1758,7 @@ definition
capAllowWrite_CL = from_bool (wd !! 0)\<rparr>"
definition
cap_rights_from_word :: "word32 \<Rightarrow> cap_rights_CL"
cap_rights_from_word :: "word32 \<Rightarrow> seL4_CapRights_CL"
where
"cap_rights_from_word wd \<equiv> SOME cr.
to_bool (capAllowGrant_CL cr) = wd !! 2 \<and>

View File

@ -243,7 +243,7 @@ where
primrec
cthread_state_relation_lifted :: "Structures_H.thread_state \<Rightarrow>
(thread_state_CL \<times> fault_CL option) \<Rightarrow> bool"
(thread_state_CL \<times> seL4_Fault_CL option) \<Rightarrow> 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 \<Rightarrow>
(thread_state_C \<times> fault_C) \<Rightarrow> bool"
(thread_state_C \<times> seL4_Fault_C) \<Rightarrow> bool"
where
"cthread_state_relation \<equiv> \<lambda>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 \<equiv>
(case cf of (Fault_cap_fault _) \<Rightarrow> True
(case cf of (SeL4_Fault_CapFault _) \<Rightarrow> True
| _ \<Rightarrow> False)"
lemma is_cap_fault_simp: "is_cap_fault cf = (\<exists> 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 = (\<exists> 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 \<Rightarrow> lookup_fault_CL \<Rightarrow> fault option"
fault_to_H :: "seL4_Fault_CL \<Rightarrow> lookup_fault_CL \<Rightarrow> 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 \<Rightarrow> fault_CL option \<Rightarrow> lookup_fault_CL option \<Rightarrow> bool"
cfault_rel :: "Fault_H.fault option \<Rightarrow> seL4_Fault_CL option \<Rightarrow> lookup_fault_CL option \<Rightarrow> bool"
where
"cfault_rel af cf lf \<equiv> \<exists>cf'. cf = Some cf' \<and>
(if (is_cap_fault cf') then (\<exists>lf'. lf = Some lf' \<and> fault_to_H cf' lf' = af)
else (fault_to_H cf' undefined = af))"
definition
carch_tcb_relation :: "Structures_H.arch_tcb \<Rightarrow> arch_tcb_C \<Rightarrow> bool"
where
"carch_tcb_relation aarch_tcb carch_tcb \<equiv>
ccontext_relation (atcbContextGet aarch_tcb) (tcbContext_C carch_tcb)"
definition
ctcb_relation :: "Structures_H.tcb \<Rightarrow> tcb_C \<Rightarrow> bool"
where
@ -332,13 +338,13 @@ where
tcbFaultHandler atcb = tcbFaultHandler_C ctcb
\<and> cthread_state_relation (tcbState atcb) (tcbState_C ctcb, tcbFault_C ctcb)
\<and> tcbIPCBuffer atcb = tcbIPCBuffer_C ctcb
\<and> ccontext_relation (tcbContext atcb) (tcbContext_C (tcbArch_C ctcb))
\<and> carch_tcb_relation (tcbArch atcb) (tcbArch_C ctcb)
\<and> tcbQueued atcb = to_bool (tcbQueued_CL (thread_state_lift (tcbState_C ctcb)))
\<and> ucast (tcbDomain atcb) = tcbDomain_C ctcb
\<and> ucast (tcbPriority atcb) = tcbPriority_C ctcb
\<and> ucast (tcbMCP atcb) = tcbMCP_C ctcb
\<and> tcbTimeSlice atcb = unat (tcbTimeSlice_C ctcb)
\<and> cfault_rel (tcbFault atcb) (fault_lift (tcbFault_C ctcb))
\<and> cfault_rel (tcbFault atcb) (seL4_Fault_lift (tcbFault_C ctcb))
(lookup_fault_lift (tcbLookupFailure_C ctcb))
\<and> 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 \<Rightarrow> cte_C ptr \<Rightarrow> (word32[160]) \<Rightarrow> bool"
@ -746,19 +752,27 @@ where
lemma ccap_relation_c_valid_cap: "ccap_relation c c' \<Longrightarrow> c_valid_cap c'"
by (simp add: ccap_relation_def)
context begin interpretation Arch .
fun
arch_fault_to_fault_tag :: "arch_fault \<Rightarrow> word32"
where
"arch_fault_to_fault_tag (VMFault a b) = scast seL4_Fault_VMFault"
end
fun
fault_to_fault_tag :: "fault \<Rightarrow> 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' \<equiv> cr = cap_rights_to_H (cap_rights_lift cr')"
"ccap_rights_relation cr cr' \<equiv> cr = cap_rights_to_H (seL4_CapRights_lift cr')"
lemma (in kernel) syscall_error_to_H_cases_rev:
"\<And>n. syscall_error_to_H e lf = Some (InvalidArgument n) \<Longrightarrow>

View File

@ -639,8 +639,8 @@ lemma msgRegisters_ccorres:
lemma asUser_cur_obj_at':
assumes f: "\<lbrace>P\<rbrace> f \<lbrace>Q\<rbrace>"
shows "\<lbrace>\<lambda>s. obj_at' (\<lambda>tcb. P (tcbContext tcb)) (ksCurThread s) s \<and> t = ksCurThread s\<rbrace>
asUser t f \<lbrace>\<lambda>rv s. obj_at' (\<lambda>tcb. Q rv (tcbContext tcb)) (ksCurThread s) s\<rbrace>"
shows "\<lbrace>\<lambda>s. obj_at' (\<lambda>tcb. P (atcbContextGet (tcbArch tcb))) (ksCurThread s) s \<and> t = ksCurThread s\<rbrace>
asUser t f \<lbrace>\<lambda>rv s. obj_at' (\<lambda>tcb. Q rv (atcbContextGet (tcbArch tcb))) (ksCurThread s) s\<rbrace>"
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:
"\<lbrace>\<lambda>s. n < unat n_msgRegisters \<and> n < unat (msgLength info) \<and> thread = ksCurThread s \<and> cur_tcb' s\<rbrace>
getMRs thread buffer info
\<lbrace>\<lambda>rv s. obj_at' (\<lambda>tcb. tcbContext tcb (ARM_H.msgRegisters ! n) = rv ! n) (ksCurThread s) s\<rbrace>"
\<lbrace>\<lambda>rv s. obj_at' (\<lambda>tcb. atcbContextGet (tcbArch tcb) (ARM_H.msgRegisters ! n) = rv ! n) (ksCurThread s) s\<rbrace>"
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="\<lambda>s. n < unat (scast n_msgRegisters :: word32) \<and> obj_at' (\<lambda>tcb. tcbContext tcb (ARM_H.msgRegisters!n) = x!n) (ksCurThread s) s"
apply (rule_tac P="\<lambda>s. n < unat (scast n_msgRegisters :: word32) \<and> obj_at' (\<lambda>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)

View File

@ -595,7 +595,7 @@ lemma sendFaultIPC_ccorres:
"ccorres (cfault_rel2 \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
(invs' and st_tcb_at' simple' tptr and sch_act_not tptr and
(\<lambda>s. \<forall>p. tptr \<notin> set (ksReadyQueues s p)))
(UNIV \<inter> {s. (cfault_rel (Some fault) (fault_lift(current_fault_' (globals s)))
(UNIV \<inter> {s. (cfault_rel (Some fault) (seL4_Fault_lift(current_fault_' (globals s)))
(lookup_fault_lift(current_lookup_fault_' (globals s))))}
\<inter> {s. tptr_' s = tcb_ptr_to_ctcb_ptr tptr})
[] (sendFaultIPC tptr fault)
@ -646,7 +646,7 @@ lemma sendFaultIPC_ccorres:
apply (rule_tac P=\<top> 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 (\<lambda>s. \<forall>p. t \<notin> set (ksReadyQueues s p)))
(UNIV \<inter> {s. (cfault_rel (Some flt) (fault_lift(current_fault_' (globals s)))
(UNIV \<inter> {s. (cfault_rel (Some flt) (seL4_Fault_lift(current_fault_' (globals s)))
(lookup_fault_lift(current_lookup_fault_' (globals s))) )}
\<inter> {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: "(\<not> (Kernel_C.maxIRQ) <s (ucast \<circ> (uc
(simp add: is_up_def target_size source_size)?)
apply (uint_arith)
done
lemma ccorres_handleReserveIRQ:
"ccorres dc xfdc \<top> 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=\<top> and P'="{s. ret__int_' s = 0 \<and> cap_get_tag cap \<noteq> 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=\<top> and P'="{s. ret__int_' s = 0 \<and> cap_get_tag cap \<noteq> 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

View File

@ -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="\<lambda>u. obj_at' (\<lambda>t. tcbContext t = u) thread" and
apply (rule ccorres_symb_exec_l [where Q="\<lambda>u. obj_at' (\<lambda>t. (atcbContextGet o tcbArch) t = u) thread" and
Q'="\<lambda>rv. {s. thread_' s = tcb_ptr_to_ctcb_ptr thread} \<inter> {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, \<sigma>)" in bexI [rotated])
apply (rule_tac x = "((atcbContextGet o tcbArch) ko reg, \<sigma>)" in bexI [rotated])
apply (simp add: in_monad' asUser_def select_f_def split_def)
apply (subst arg_cong2 [where f = "op \<in>"])
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

View File

@ -172,9 +172,9 @@ lemma getObject_state:
lemma threadGet_state:
"\<lbrakk> (uc, s') \<in> fst (threadGet tcbContext t' s); ko_at' ko t s \<rbrakk> \<Longrightarrow>
"\<lbrakk> (uc, s') \<in> fst (threadGet (atcbContextGet o tcbArch) t' s); ko_at' ko t s \<rbrakk> \<Longrightarrow>
(uc, s'\<lparr>ksPSpace := ksPSpace s(t \<mapsto> KOTCB (tcbState_update (\<lambda>_. st) ko))\<rparr>) \<in>
fst (threadGet tcbContext t' (s\<lparr>ksPSpace := ksPSpace s(t \<mapsto> KOTCB (tcbState_update (\<lambda>_. st) ko))\<rparr>))"
fst (threadGet (atcbContextGet o tcbArch) t' (s\<lparr>ksPSpace := ksPSpace s(t \<mapsto> KOTCB (tcbState_update (\<lambda>_. st) ko))\<rparr>))"
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:
"\<lbrakk> (uc, s') \<in> fst (threadGet tcbContext (ksCurThread s) s); ko_at' ko t s;
"\<lbrakk> (uc, s') \<in> fst (threadGet (atcbContextGet o tcbArch) (ksCurThread s) s); ko_at' ko t s;
t \<noteq> ksCurThread s \<rbrakk> \<Longrightarrow>
(uc, s'\<lparr>ksPSpace := ksPSpace s(t \<mapsto> KOTCB (tcbContext_update (\<lambda>_. st) ko))\<rparr>) \<in>
fst (threadGet tcbContext (ksCurThread s) (s\<lparr>ksPSpace := ksPSpace s(t \<mapsto> KOTCB (tcbContext_update (\<lambda>_. st) ko))\<rparr>))"
(uc, s'\<lparr>ksPSpace := ksPSpace s(t \<mapsto> KOTCB (tcbArch_update (\<lambda>_. atcbContextSet st (tcbArch ko)) ko))\<rparr>) \<in>
fst (threadGet (atcbContextGet o tcbArch) (ksCurThread s) (s\<lparr>ksPSpace := ksPSpace s(t \<mapsto> KOTCB (tcbArch_update (\<lambda>_. atcbContextSet st (tcbArch ko)) ko))\<rparr>))"
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:
"\<lbrakk>(x,s) \<in> fst (asUser (ksCurThread s) f s); ko_at' ko t s; \<And>s. \<lbrace>op = s\<rbrace> f \<lbrace>\<lambda>_. op = s\<rbrace> ;
t \<noteq> ksCurThread s\<rbrakk> \<Longrightarrow>
(x,s\<lparr>ksPSpace := ksPSpace s(t \<mapsto> KOTCB (tcbContext_update (\<lambda>_. st) ko))\<rparr>) \<in>
fst (asUser (ksCurThread s) f (s\<lparr>ksPSpace := ksPSpace s(t \<mapsto> KOTCB (tcbContext_update (\<lambda>_. st) ko))\<rparr>))"
(x,s\<lparr>ksPSpace := ksPSpace s(t \<mapsto> KOTCB (tcbArch_update (\<lambda>_. atcbContextSet st (tcbArch ko)) ko))\<rparr>) \<in>
fst (asUser (ksCurThread s) f (s\<lparr>ksPSpace := ksPSpace s(t \<mapsto> KOTCB (tcbArch_update (\<lambda>_. atcbContextSet st (tcbArch ko)) ko))\<rparr>))"
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:
"\<lbrakk>getMRs_rel args buffer s;
(cur_tcb' and case_option \<top> valid_ipc_buffer_ptr' buffer) s;
ko_at' ko t s ; t \<noteq> ksCurThread s\<rbrakk> \<Longrightarrow>
getMRs_rel args buffer (s\<lparr>ksPSpace := ksPSpace s(t \<mapsto> KOTCB (tcbContext_update (\<lambda>_. st) ko))\<rparr>)"
getMRs_rel args buffer (s\<lparr>ksPSpace := ksPSpace s(t \<mapsto> KOTCB (tcbArch_update (\<lambda>_. atcbContextSet st (tcbArch ko)) ko))\<rparr>)"
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="\<lambda>m s. obj_at' (\<lambda>tcb. map (tcbContext tcb) (genericTake n
apply (rule_tac F="\<lambda>m s. obj_at' (\<lambda>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="\<lambda>m s. obj_at' (\<lambda>tcb. map (tcbContext tcb) (genericTake n
rule_tac F="\<lambda>m s. obj_at' (\<lambda>tcb. map ((atcbContextGet o tcbArch) tcb) (genericTake n
(ARM_H.frameRegisters @ ARM_H.gpRegisters))
= reply) target s
\<and> 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="\<lambda>m s. obj_at' (\<lambda>tcb. map (tcbContext tcb) (genericTake n
apply (rule_tac F="\<lambda>m s. obj_at' (\<lambda>tcb. map ((atcbContextGet o tcbArch) tcb) (genericTake n
(ARM_H.frameRegisters @ ARM_H.gpRegisters))
= reply) target s
\<and> valid_ipc_buffer_ptr' (the rva) s \<and> 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)

View File

@ -472,8 +472,8 @@ lemma mask_32_id [simp]:
lemma handleVMFault_ccorres:
"ccorres ((\<lambda>a ex v. ex = scast EXCEPTION_FAULT \<and> (\<exists>vf.
a = VMFault (fault_vm_fault_CL.address_CL vf) [instructionFault_CL vf, FSR_CL vf] \<and>
errfault v = Some (Fault_vm_fault vf))) \<currency>
a = ArchFault (VMFault (seL4_Fault_VMFault_CL.address_CL vf) [instructionFault_CL vf, FSR_CL vf]) \<and>
errfault v = Some (SeL4_Fault_VMFault vf))) \<currency>
(\<lambda>_. \<bottom>))
(liftxf errstate id (K ()) ret__unsigned_long_')
\<top>
@ -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="\<lambda>tcb. tcbContext tcb = rv"
apply (rule_tac P="\<lambda>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)