crefine: remove lemmas moved into ArchMove_C/Move_C and fix proofs

Signed-off-by: Victor Phan <Victor.Phan@data61.csiro.au>
This commit is contained in:
Victor Phan 2020-03-17 19:50:19 +11:00
parent 70fe3fa943
commit 046a1358f6
71 changed files with 568 additions and 4799 deletions

View File

@ -506,51 +506,6 @@ lemma array_relation_map_conv2:
\<forall>i>n. a i = None \<Longrightarrow> array_map_conv f n c = a"
by (rule ext) (simp add: array_relation_def array_map_conv_def2)
(* FIXME: move to somewhere sensible >>> *)
text \<open>Map inversion (implicitly assuming injectivity).\<close>
definition
"the_inv_map m = (\<lambda>s. if s\<in>ran m then Some (THE x. m x = Some s) else None)"
text \<open>Map inversion can be expressed by function inversion.\<close>
lemma the_inv_map_def2:
"the_inv_map m = (Some \<circ> the_inv_into (dom m) (the \<circ> m)) |` (ran m)"
apply (rule ext)
apply (clarsimp simp: the_inv_map_def the_inv_into_def dom_def)
apply (rule_tac f=The in arg_cong)
apply (rule ext)
apply auto
done
text \<open>The domain of a function composition with Some is the universal set.\<close>
lemma dom_comp_Some[simp]: "dom (comp Some f) = UNIV" by (simp add: dom_def)
text \<open>Assuming injectivity, map inversion produces an inversive map.\<close>
lemma is_inv_the_inv_map:
"inj_on m (dom m) \<Longrightarrow> is_inv m (the_inv_map m)"
apply (simp add: is_inv_def)
apply (intro conjI allI impI)
apply (simp add: the_inv_map_def2)
apply (auto simp add: the_inv_map_def inj_on_def dom_def)
done
lemma the_the_inv_mapI:
"inj_on m (dom m) \<Longrightarrow> m x = Some y \<Longrightarrow> the (the_inv_map m y) = x"
by (auto simp: the_inv_map_def ran_def inj_on_def dom_def)
lemma eq_restrict_map_None[simp]:
"restrict_map m A x = None \<longleftrightarrow> x ~: (A \<inter> dom m)"
by (auto simp: restrict_map_def split: if_split_asm)
lemma eq_the_inv_map_None[simp]: "the_inv_map m x = None \<longleftrightarrow> x\<notin>ran m"
by (simp add: the_inv_map_def2)
lemma is_inv_unique:
"is_inv f g \<Longrightarrow> is_inv f h \<Longrightarrow> g=h"
apply (rule ext)
apply (clarsimp simp: is_inv_def dom_def Collect_eq ran_def)
apply (drule_tac x=x in spec)+
apply (case_tac "g x", clarsimp+)
done
(*<<<*)
definition
casid_map_to_H
:: "(word32[256]) \<Rightarrow> (pde_C ptr \<rightharpoonup> pde_C) \<Rightarrow> asid \<rightharpoonup> hw_asid \<times> obj_ref"
@ -563,11 +518,6 @@ definition
(the_inv_map (array_map_conv (\<lambda>x. if x=0 then None else Some x)
0xFF hw_asid_table) asid))"
lemma eq_option_to_0_rev:
"Some 0 ~: A \<Longrightarrow> \<forall>x. \<forall>y\<in>A.
((=) \<circ> option_to_0) y x \<longrightarrow> (if x = 0 then None else Some x) = y"
by (clarsimp simp: option_to_0_def split: option.splits)
lemma inj_hwasidsI:
"asid_map_pd_to_hwasids m = set_option \<circ> c \<Longrightarrow>
inj_on (map_option snd \<circ> m) (dom m) \<Longrightarrow>
@ -736,12 +686,6 @@ lemma cready_queues_to_H_correct:
(* showing that cpspace_relation is actually unique >>>*)
(* FIXME: move *)
lemma inj_image_inv:
assumes inj_f: "inj f"
shows "f ` A = B \<Longrightarrow> inv f ` B = A"
by (drule sym) (simp add: image_inv_f_f[OF inj_f])
lemma cmap_relation_unique:
assumes inj_f: "inj f"
assumes r: "\<And>x y z. r x z \<Longrightarrow> r y z \<Longrightarrow> x=y"
@ -779,47 +723,6 @@ lemma carch_tcb_relation_imp_eq:
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 =
{(Structures_H.tcbCTable, tcbCTable_update),
(Structures_H.tcbVTable, tcbVTable_update),
(Structures_H.tcbReply, tcbReply_update),
(Structures_H.tcbCaller, tcbCaller_update),
(tcbIPCBufferFrame, tcbIPCBufferFrame_update)}"
by (auto simp add: tcb_cte_cases_def split: if_split_asm)
(* FIXME: move *)
lemma ps_clear_is_aligned_ksPSpace_None:
"\<lbrakk>ps_clear p n s; is_aligned p n; 0<d; d \<le> mask n\<rbrakk>
\<Longrightarrow> ksPSpace s (p + d) = None"
apply (simp add: ps_clear_def add_diff_eq[symmetric] mask_2pm1[symmetric])
apply (drule equals0D[where a="p + d"])
apply (simp add: dom_def word_gt_0 del: word_neq_0_conv)
apply (drule mp)
apply (rule word_plus_mono_right)
apply simp
apply (simp add: mask_2pm1)
apply (erule is_aligned_no_overflow')
apply (drule mp)
apply (case_tac "(0::word32)<2^n")
apply (frule le_m1_iff_lt[of "(2::word32)^n" d, THEN iffD1])
apply (simp add: mask_2pm1[symmetric])
apply (erule (1) is_aligned_no_wrap')
apply (simp add: is_aligned_mask mask_2pm1 not_less word_bits_def
power_overflow)
by assumption
(* FIXME: move *)
lemma ps_clear_is_aligned_ctes_None:
assumes "ps_clear p tcbBlockSizeBits s"
and "is_aligned p tcbBlockSizeBits"
shows "ksPSpace s (p + 2*2^cteSizeBits) = None"
and "ksPSpace s (p + 3*2^cteSizeBits) = None"
and "ksPSpace s (p + 4*2^cteSizeBits) = None"
by (auto intro: assms ps_clear_is_aligned_ksPSpace_None
simp: objBits_defs mask_def)+
lemma map_to_ctes_tcb_ctes:
notes if_cong[cong]
shows
@ -993,10 +896,6 @@ lemma cpspace_ntfn_relation_unique:
kernel.tcb_at_not_NULL tcb_ptr_to_ctcb_ptr_inj
split: ntfn.splits option.splits) (* long *)
(* FIXME: move *)
lemma of_bool_inject[iff]: "of_bool a = of_bool b \<longleftrightarrow> a=b"
by (cases a) (cases b, simp_all)+
(* FIXME: move *)
lemma ap_from_vm_rights_inject[iff]:
"(ap_from_vm_rights a::word32) = ap_from_vm_rights b \<longleftrightarrow> a=b"
@ -1014,23 +913,6 @@ lemma cpspace_pte_relation_unique:
apply (rule cmap_relation_unique[OF inj_Ptr _ assms])
by (simp add: cpte_relation_def s_from_cacheable_def Let_def split: pte.splits bool.splits)
(* FIXME: move *)
lemma Collect_mono2: "Collect P \<subseteq> Collect Q \<longleftrightarrow> (\<forall>x. P x \<longrightarrow> Q x)" by auto
(* FIXME: move to Wellformed, turn valid_asid_pool' into an abbreviation >>>*)
primrec
wf_asid_pool' :: "asidpool \<Rightarrow> bool"
where
"wf_asid_pool' (ASIDPool pool) =
(dom pool \<subseteq> {0 .. 2^asid_low_bits - 1} \<and>
0 \<notin> ran pool \<and> (\<forall>x \<in> ran pool. is_aligned x pdBits))"
lemma valid_eq_wf_asid_pool'[simp]:
"valid_asid_pool' pool = (\<lambda>s. wf_asid_pool' pool)"
by (case_tac pool) simp
declare valid_asid_pool'.simps[simp del]
(*<<<*)
lemma cpspace_asidpool_relation_unique:
assumes invs: "\<forall>x\<in>ran (map_to_asidpools ah). wf_asid_pool' x"
"\<forall>x\<in>ran (map_to_asidpools ah'). wf_asid_pool' x"

View File

@ -8,15 +8,6 @@ theory Arch_C
imports Recycle_C
begin
(* FIXME: move *)
lemma of_bool_from_bool: "of_bool = from_bool"
by (rule ext, simp add: from_bool_def split: bool.split)
lemma ksPSpace_update_eq_ExD:
"s = t\<lparr> ksPSpace := ksPSpace s\<rparr>
\<Longrightarrow> \<exists>ps. s = t \<lparr> ksPSpace := ps \<rparr>"
by (erule exI)
context begin interpretation Arch . (*FIXME: arch_split*)
crunch ctes_of[wp]: unmapPageTable "\<lambda>s. P (ctes_of s)"
(wp: crunch_wps simp: crunch_simps)
@ -1690,18 +1681,6 @@ lemma createMappingEntries_valid_pte_slots'2:
apply (wp | simp add:valid_pte_slots'2_def)+
done
(* FIXME: move *)
(* this one is specialised to a PDE for a supersection *)
lemma vaddr_segment_nonsense6:
"is_aligned (p :: word32) pdBits \<Longrightarrow>
(p + (vaddr >> 20 << 2) && ~~ mask pdBits) = p"
apply (rule is_aligned_add_helper[THEN conjunct2])
apply (erule is_aligned_weaken, simp)
apply (simp add: pdBits_def pdeBits_def)
apply (rule shiftl_less_t2n[where m=14 and n=2 and 'a=machine_word_len, simplified])
apply (rule shiftr_less_t2n'[where m=12 and n=20 and 'a=machine_word_len, simplified])
done
(* replay of proof in Arch_R with stronger validity result *)
lemma createMappingEntries_valid_pde_slots'2:
"\<lbrace>page_directory_at' pd and K (vmsz_aligned' vptr sz \<and> vptr < kernelBase)\<rbrace>
@ -2038,11 +2017,6 @@ lemma ivc_label_flush_case:
= C"
by (auto split: invocation_label.split arch_invocation_label.split)
lemma list_length_less:
"(args = [] \<or> length args \<le> Suc 0) = (length args < 2)"
by (case_tac args,simp_all)
lemma injection_handler_whenE:
"injection_handler Inl (whenE a b)
= (if a then (injection_handler Inl b)
@ -2073,17 +2047,6 @@ lemma flushtype_relation_triv:
ARM_H.isPDFlushLabel_def
split: flush_type.splits invocation_label.splits arch_invocation_label.splits)
lemma at_least_2_args:
"\<not> length args < 2 \<Longrightarrow> \<exists>a b c. args = a#b#c"
apply (case_tac args)
apply simp
apply (case_tac list)
apply simp
apply (case_tac lista)
apply simp
apply simp
done
lemma pte_get_tag_alt:
"pte_lift v = Some pteC
\<Longrightarrow> pte_get_tag v = (case pteC of
@ -2859,13 +2822,6 @@ lemma framesize_from_H_mask2:
Kernel_C.ARMSuperSection_def)+
done
lemma rel_option_alt_def:
"rel_option f a b = (
(a = None \<and> b = None)
\<or> (\<exists>x y. a = Some x \<and> b = Some y \<and> f x y))"
apply (case_tac a, case_tac b, simp, simp, case_tac b, auto)
done
lemma injection_handler_stateAssert_relocate:
"injection_handler Inl (stateAssert ass xs >>= f) >>=E g
= do v \<leftarrow> stateAssert ass xs; injection_handler Inl (f ()) >>=E g od"
@ -3193,12 +3149,6 @@ lemma decodeARMPageDirectoryInvocation_ccorres:
| rule flushtype_relation_triv,simp add:isPageFlush_def isPDFlushLabel_def
| rule word_of_nat_less,simp add: pbfs_less)+
lemma cond_throw_whenE:
"(if P then f else throwError e) = (whenE (\<not> P) (throwError e) >>=E (\<lambda>_. f))"
by (auto split: if_splits
simp: throwError_def bindE_def
whenE_def bind_def returnOk_def return_def)
lemma Arch_decodeInvocation_ccorres:
notes if_cong[cong] tl_drop_1[simp]
shows

View File

@ -24,20 +24,9 @@ lemmas typ_heap_simps' = typ_heap_simps c_guard_clift
lemmas asUser_return = submonad.return [OF submonad_asUser]
lemma setMRs_Nil:
"setMRs thread buffer [] = stateAssert (tcb_at' thread) [] >>= (\<lambda>_. return 0)"
unfolding setMRs_def
by (simp add: zipWithM_x_def sequence_x_def zipWith_def
asUser_return)
lemmas asUser_bind_distrib =
submonad_bind [OF submonad_asUser submonad_asUser submonad_asUser]
lemma ps_clear_upd_None:
"ksPSpace s y = None \<Longrightarrow>
ps_clear x n (ksPSpace_update (\<lambda>a. (ksPSpace s)(y := None)) s') = ps_clear x n s"
by (rule iffI | clarsimp elim!: ps_clear_domE | fastforce)+
lemma ntfnQueue_head_mask_4 :
"ntfnQueue_head_CL (notification_lift ko') && ~~ mask 4 = ntfnQueue_head_CL (notification_lift ko')"
unfolding notification_lift_def
@ -69,13 +58,6 @@ lemma no_overlap_new_cap_addrs_disjoint:
apply auto
done
lemma empty_fail_asUser[iff]:
"empty_fail m \<Longrightarrow> empty_fail (asUser t m)"
apply (simp add: asUser_def split_def)
apply (intro empty_fail_bind, simp_all)
apply (simp add: select_f_def empty_fail_def)
done
declare empty_fail_doMachineOp [simp]
lemma empty_fail_loadWordUser[intro!, simp]:
@ -86,20 +68,6 @@ lemma empty_fail_getMRs[iff]:
"empty_fail (getMRs t buf mi)"
by (auto simp add: getMRs_def split: option.split)
lemma asUser_mapM_x:
"(\<And>x. empty_fail (f x)) \<Longrightarrow>
asUser t (mapM_x f xs) = do stateAssert (tcb_at' t) []; mapM_x (\<lambda>x. asUser t (f x)) xs od"
apply (simp add: mapM_x_mapM asUser_bind_distrib)
apply (subst submonad_mapM [OF submonad_asUser submonad_asUser])
apply simp
apply (simp add: asUser_return bind_assoc o_def)
apply (rule ext)
apply (rule bind_apply_cong [OF refl])+
apply (clarsimp simp: in_monad dest!: fst_stateAssertD)
apply (drule use_valid, rule mapM_wp', rule asUser_typ_ats, assumption)
apply (simp add: stateAssert_def get_def NonDetMonad.bind_def)
done
lemma asUser_get_registers:
"\<lbrace>tcb_at' target\<rbrace>
asUser target (mapM getRegister xs)
@ -130,54 +98,6 @@ lemma projectKO_user_data_device:
by (cases ko)
(auto simp: projectKO_opts_defs split: arch_kernel_object.splits)
lemma device_data_at_ko:
"typ_at' UserDataDeviceT p s \<Longrightarrow> ko_at' UserDataDevice p s"
apply (clarsimp simp: typ_at'_def obj_at'_def ko_wp_at'_def
projectKO_user_data_device projectKO_eq projectKO_eq2)
apply (case_tac ko, auto)
done
(* FIXME: move *)
lemma user_data_at_ko:
"typ_at' UserDataT p s \<Longrightarrow> ko_at' UserData p s"
apply (clarsimp simp: typ_at'_def obj_at'_def ko_wp_at'_def projectKOs)
apply (case_tac ko, auto)
done
(* FIXME: move *)
lemma map_to_ko_atI:
"\<lbrakk>(projectKO_opt \<circ>\<^sub>m ksPSpace s) x = Some v;
pspace_aligned' s; pspace_distinct' s\<rbrakk>
\<Longrightarrow> ko_at' v x s"
apply (clarsimp simp: map_comp_Some_iff)
apply (erule (2) aligned_distinct_obj_atI')
apply (simp add: project_inject)
done
lemma empty_fail_rethrowFailure:
"empty_fail f \<Longrightarrow> empty_fail (rethrowFailure fn f)"
apply (simp add: rethrowFailure_def handleE'_def)
apply (erule empty_fail_bind)
apply (simp split: sum.split)
done
lemma empty_fail_resolveAddressBits:
"empty_fail (resolveAddressBits cap cptr bits)"
proof -
note empty_fail_assertE[iff]
show ?thesis
apply (rule empty_fail_use_cutMon)
apply (induct rule: resolveAddressBits.induct)
apply (subst resolveAddressBits.simps)
apply (unfold Let_def cnode_cap_case_if fun_app_def
K_bind_def haskell_assertE_def split_def)
apply (intro empty_fail_cutMon_intros)
apply (clarsimp simp: empty_fail_drop_cutMon empty_fail_whenEs
locateSlot_conv returnOk_liftE[symmetric]
isCap_simps)+
done
qed
lemma empty_fail_getReceiveSlots:
"empty_fail (getReceiveSlots r rbuf)"
proof -

View File

@ -117,16 +117,6 @@ lemma lookupCapAndSlot_ccorres :
done
(* FIXME: move *)
lemma injection_handler_liftM:
"injection_handler f
= liftM (\<lambda>v. case v of Inl ex \<Rightarrow> Inl (f ex) | Inr rv \<Rightarrow> Inr rv)"
apply (intro ext, simp add: injection_handler_def liftM_def
handleE'_def)
apply (rule bind_apply_cong, rule refl)
apply (simp add: throwError_def split: sum.split)
done
lemma lookupErrorOnFailure_ccorres:
"ccorres (f \<currency> r) xf P P' hs a c
\<Longrightarrow> ccorres ((\<lambda>x y z. \<exists>w. x = FailedLookup isSource w \<and> f w y z) \<currency> r)

View File

@ -35,8 +35,6 @@ lemma maskCapRights_cap_cases:
done
lemma imp_ignore: "B \<Longrightarrow> A \<longrightarrow> B" by blast
lemma wordFromVMRights_spec:
"\<forall>s. \<Gamma> \<turnstile> {s} Call wordFromVMRights_'proc \<lbrace>\<acute>ret__unsigned_long = \<^bsup>s\<^esup>vm_rights\<rbrace>"
by vcg simp?

View File

@ -553,14 +553,6 @@ lemma cap_rights_to_H_from_word_canon [simp]:
apply (simp add: cap_rights_to_H_def)
done
(* MOVE *)
lemma tcb_aligned':
"tcb_at' t s \<Longrightarrow> is_aligned t tcbBlockSizeBits"
apply (drule obj_at_aligned')
apply (simp add: objBits_simps)
apply (simp add: objBits_simps)
done
abbreviation
"lookupSlot_raw_xf \<equiv>
liftxf errstate lookupSlot_raw_ret_C.status_C

View File

@ -126,7 +126,6 @@ lemma det_wp_asUser [wp]:
apply simp
done
(* FIXME move into Refine somewhere *)
lemma wordSize_def':
"wordSize = 4"
unfolding wordSize_def wordBits_def

View File

@ -1295,48 +1295,6 @@ lemma ksPSpace_update_ext:
"(\<lambda>s. s\<lparr>ksPSpace := f (ksPSpace s)\<rparr>) = (ksPSpace_update f)"
by (rule ext) simp
(* FIXME: move *)
lemma valid_untyped':
notes usableUntypedRange.simps[simp del]
assumes pspace_distinct': "pspace_distinct' s" and
pspace_aligned': "pspace_aligned' s" and
al: "is_aligned ptr bits"
shows "valid_untyped' d ptr bits idx s =
(\<forall>p ko. ksPSpace s p = Some ko \<longrightarrow>
obj_range' p ko \<inter> {ptr..ptr + 2 ^ bits - 1} \<noteq> {} \<longrightarrow>
obj_range' p ko \<subseteq> {ptr..ptr + 2 ^ bits - 1} \<and>
obj_range' p ko \<inter>
usableUntypedRange (UntypedCap d ptr bits idx) = {})"
apply (simp add: valid_untyped'_def)
apply (simp add: ko_wp_at'_def)
apply (rule arg_cong[where f=All])
apply (rule ext)
apply (rule arg_cong[where f=All])
apply (rule ext)
apply (case_tac "ksPSpace s ptr' = Some ko", simp_all)
apply (frule pspace_alignedD'[OF _ pspace_aligned'])
apply (frule pspace_distinctD'[OF _ pspace_distinct'])
apply simp
apply (frule aligned_ranges_subset_or_disjoint[OF al])
apply (fold obj_range'_def)
apply (rule iffI)
apply auto[1]
apply (rule conjI)
apply (rule ccontr, simp)
apply (simp add: Set.psubset_eq)
apply (erule conjE)
apply (case_tac "obj_range' ptr' ko \<inter> {ptr..ptr + 2 ^ bits - 1} \<noteq> {}", simp)
apply (cut_tac is_aligned_no_overflow[OF al])
apply (auto simp add: obj_range'_def)[1]
apply (clarsimp simp add: usableUntypedRange.simps Int_commute)
apply (case_tac "obj_range' ptr' ko \<inter> {ptr..ptr + 2 ^ bits - 1} \<noteq> {}", simp+)
apply (cut_tac is_aligned_no_overflow[OF al])
apply (clarsimp simp add: obj_range'_def)
apply (frule is_aligned_no_overflow)
by (metis al intvl_range_conv' le_m1_iff_lt less_is_non_zero_p1
nat_le_linear power_overflow sub_wrap add_0
add_0_right word_add_increasing word_less_1 word_less_sub_1)
lemma hrs_ghost_update_comm:
"(t_hrs_'_update f \<circ> ghost'state_'_update g) =
(ghost'state_'_update g \<circ> t_hrs_'_update f)"

View File

@ -2692,6 +2692,7 @@ lemma fastpath_reply_recv_ccorres:
automatic indentation is improved *)
show ?thesis
using [[goals_limit = 1]]
supply option.case_cong_weak[cong del]
apply (cinit lift: cptr_' msgInfo_')
apply (simp add: catch_liftE_bindE unlessE_throw_catch_If
unifyFailure_catch_If catch_liftE
@ -4195,14 +4196,11 @@ lemma modify_setEndpoint_pivot[unfolded K_bind_def]:
lemma setEndpoint_clearUntypedFreeIndex_pivot[unfolded K_bind_def]:
"do setEndpoint p val; v <- clearUntypedFreeIndex slot; f od
= do v <- clearUntypedFreeIndex slot; setEndpoint p val; f od"
by (simp add: clearUntypedFreeIndex_def bind_assoc
getSlotCap_def
setEndpoint_getCTE_pivot
updateTrackedFreeIndex_def
modify_setEndpoint_pivot
supply option.case_cong_weak[cong del]
by (simp add: clearUntypedFreeIndex_def bind_assoc getSlotCap_def setEndpoint_getCTE_pivot
updateTrackedFreeIndex_def modify_setEndpoint_pivot
split: capability.split
| rule bind_cong[OF refl] allI impI
bind_apply_cong[OF refl])+
| rule bind_cong[OF refl] allI impI bind_apply_cong[OF refl])+
lemma emptySlot_setEndpoint_pivot[unfolded K_bind_def]:
"(do emptySlot slot NullCap; setEndpoint p val; f od) =
@ -4453,8 +4451,9 @@ lemma fastpath_callKernel_SysReplyRecv_corres:
and cnode_caps_gsCNodes')
(callKernel (SyscallEvent SysReplyRecv)) (fastpaths SysReplyRecv)"
including no_pre
supply option.case_cong_weak[cong del]
apply (rule monadic_rewrite_introduce_alternative)
apply ( simp add: callKernel_def)
apply (simp add: callKernel_def)
apply (rule monadic_rewrite_imp)
apply (simp add: handleEvent_def handleReply_def
handleRecv_def liftE_bindE_handle liftE_handle

View File

@ -544,19 +544,6 @@ lemma hasCancelSendRights_spec:
split: capability.splits bool.splits)[1]
done
lemma updateCapData_Untyped:
"isUntypedCap a
\<Longrightarrow> updateCapData b c a = a"
by (clarsimp simp:isCap_simps updateCapData_def)
lemma ctes_of_valid_strengthen:
"(invs' s \<and> ctes_of s p = Some cte) \<longrightarrow> valid_cap' (cteCap cte) s"
apply (case_tac cte)
apply clarsimp
apply (erule ctes_of_valid_cap')
apply fastforce
done
lemma decodeCNodeInvocation_ccorres:
notes gen_invocation_type_eq[simp]
shows
@ -1606,51 +1593,6 @@ lemma t_hrs_update_use_t_hrs:
= (t_hrs_'_update (\<lambda>_. f (t_hrs_' s)) $ s)"
by simp
lemma name_seq_bound_helper:
"(\<not> CP n \<and> (\<forall>n' < n. CP n'))
\<Longrightarrow> (if \<exists>n. \<not> CP n
then simpl_sequence c' (map f [0 ..< (LEAST n. \<not> CP n)])
else c) = (simpl_sequence c' (map f [0 ..< n]))"
apply (simp add: exI[where x=n])
apply (subst Least_equality[where x=n], simp_all)
apply (rule ccontr, simp add: linorder_not_le)
done
lemma reset_name_seq_bound_helper:
fixes sz
fixes v :: "('a :: len) word"
defines "CP \<equiv> (\<lambda>n. ~ (v && ~~ mask sz) + of_nat n * (-1 << sz) =
((-1 :: 'a word) << sz))"
and "n \<equiv> Suc (unat (shiftR v sz))"
assumes vsz: "v + 1 < 2 ^ (len_of TYPE('a) - 1)" "2 ^ sz \<noteq> (0 :: 'a word)"
and vless: "v < v'"
shows "(\<not> CP n \<and> (\<forall>n' < n. CP n'))"
apply (clarsimp simp: shiftl_t2n field_simps less_Suc_eq_le CP_def n_def)
apply (simp add: shiftr_shiftl1[where b=sz and c=sz, simplified, symmetric]
shiftl_t2n)
apply (clarsimp simp: word_sle_msb_le shiftl_t2n[symmetric])
apply (case_tac n', simp_all)
apply (cut_tac vsz(1) order_less_le_trans[OF vless max_word_max])
apply (clarsimp simp: shiftr_shiftl1 dest!: word_add_no_overflow)
apply (drule_tac f="\<lambda>x. x - 2 ^ sz" in arg_cong, simp)
apply (metis less_irrefl order_le_less_trans order_less_trans
word_and_le2[where a=v and y="~~ mask sz"]
word_two_power_neg_ineq[OF vsz(2)])
apply (clarsimp simp add: field_simps)
apply (drule_tac f="\<lambda>x. shiftr x sz" in arg_cong)
apply (simp add: linorder_not_less[symmetric] word_shift_by_n, erule notE)
apply (metis less_Suc_eq_le unat_le_helper and_mask2 word_and_le2)
done
schematic_goal sz8_helper:
"((-1) << 8 :: addr) = ?v"
by (simp add: shiftl_t2n)
lemmas reset_name_seq_bound_helper2
= reset_name_seq_bound_helper[where sz=8 and v="v :: addr" for v,
simplified sz8_helper word_bits_def[symmetric],
THEN name_seq_bound_helper]
lemma reset_untyped_inner_offs_helper:
"\<lbrakk> cteCap cte = UntypedCap dev ptr sz idx;
i \<le> unat ((of_nat idx - 1 :: addr) div 2 ^ sz2);
@ -2471,11 +2413,6 @@ lemma invokeUntyped_Retype_ccorres:
done
qed
lemma injection_handler_whenE:
"injection_handler injf (whenE P f)
= whenE P (injection_handler injf f)"
by (simp add: whenE_def injection_handler_returnOk split: if_split)
lemma fromEnum_object_type_to_H:
"fromEnum x = unat (object_type_from_H x)"
apply (cut_tac eqset_imp_iff[where x=x, OF enum_surj])

View File

@ -562,15 +562,6 @@ lemma ccorres_pre_getQueue:
apply (simp add: maxDom_to_H maxPrio_to_H)+
done
(* FIXME: move *)
lemma threadGet_wp:
"\<lbrace>\<lambda>s. \<forall>tcb. ko_at' tcb thread s \<longrightarrow> P (f tcb) s\<rbrace> threadGet f thread \<lbrace>P\<rbrace>"
apply (rule hoare_post_imp [OF _ tg_sp'])
apply clarsimp
apply (frule obj_at_ko_at')
apply (clarsimp elim: obj_atE')
done
lemma state_relation_queue_update_helper':
"\<lbrakk> (s, s') \<in> rf_sr;
(\<forall>d p. (\<forall>t\<in>set (ksReadyQueues s (d, p)). obj_at' (inQ d p) t s)
@ -733,14 +724,6 @@ lemma invert_l1index_spec:
by vcg
(simp add: word_sle_def sdiv_int_def sdiv_word_def smod_word_def smod_int_def)
lemma unat_ucast_prio_shiftr_simp[simp]:
"unat (ucast (p::priority) >> n :: machine_word) = unat (p >> n)"
by (simp add: shiftr_div_2n')+
lemma unat_ucast_prio_mask_simp[simp]:
"unat (ucast (p::priority) && mask m :: machine_word) = unat (p && mask m)"
by (metis ucast_and_mask unat_ucast_8_32)
lemma unat_ucast_prio_L1_cmask_simp:
"unat (ucast (p::priority) && 0x1F :: machine_word) = unat (p && 0x1F)"
using unat_ucast_prio_mask_simp[where m=5]
@ -778,11 +761,6 @@ lemma prio_ucast_shiftr_wordRadix_helper2:
"(ucast (p::priority) >> wordRadix :: machine_word) < 0x20"
by (rule order_less_trans[OF prio_ucast_shiftr_wordRadix_helper]; simp)
lemma dom_less_0x10_helper:
"d \<le> maxDomain \<Longrightarrow> (ucast d :: machine_word) < 0x10"
unfolding maxDomain_def numDomains_def
by (clarsimp simp add: word_less_nat_alt unat_ucast_upcast is_up word_le_nat_alt)
lemma cready_queues_index_to_C_ucast_helper:
fixes p :: priority
fixes d :: domain
@ -1168,10 +1146,6 @@ lemma rf_sr_drop_bitmaps_dequeue_helper:
carch_state_relation_def cmachine_state_relation_def
by (clarsimp simp: rf_sr_cbitmap_L1_relation rf_sr_cbitmap_L2_relation)
lemma filter_empty_unfiltered_contr:
"\<lbrakk> [x\<leftarrow>xs . x \<noteq> y] = [] ; x' \<in> set xs ; x' \<noteq> y \<rbrakk> \<Longrightarrow> False"
by (induct xs, auto split: if_split_asm)
(* FIXME same proofs as bit_set, maybe can generalise? *)
lemma cbitmap_L1_relation_bit_clear:
fixes p :: priority
@ -1213,11 +1187,6 @@ lemma cbitmap_L2_relationD:
unfolding cbitmap_L2_relation_def l2BitmapSize_def'
by clarsimp
(* FIXME move *)
lemma filter_noteq_op:
"[x \<leftarrow> xs . x \<noteq> y] = filter ((\<noteq>) y) xs"
by (induct xs) auto
lemma cbitmap_L2_relation_bit_clear:
fixes p :: priority
fixes d :: domain

View File

@ -172,72 +172,6 @@ lemma asUser_comm:
apply (rule efa efb)+
done
(* FIXME move the thread_submonad stuff to SubMonad_R and use it for asUser *)
definition
"thread_fetch \<equiv> \<lambda>ext t s. case (ksPSpace s t) of
Some (KOTCB tcb) \<Rightarrow> ext tcb
| None \<Rightarrow> undefined"
definition
"thread_fetch_option \<equiv> \<lambda>ext t s. case (ksPSpace s t) of
Some (KOTCB tcb) \<Rightarrow> ext tcb
| None \<Rightarrow> None"
definition
"thread_replace \<equiv> \<lambda>upd t nv s.
let obj = case (ksPSpace s t) of
Some (KOTCB tcb) \<Rightarrow> Some (KOTCB (upd (\<lambda>_. nv) tcb))
| obj \<Rightarrow> obj
in s \<lparr> ksPSpace := (ksPSpace s) (t := obj) \<rparr>"
lemma thread_submonad_args:
"\<lbrakk> \<And>f v. ext (upd f v) = f (ext v);
\<And>f1 f2 v. upd f1 (upd f2 v) = upd (f1 \<circ> f2) v;
\<And>f v. upd (\<lambda>_. ext v) v = v \<rbrakk> \<Longrightarrow>
submonad_args (thread_fetch ext t) (thread_replace upd t) (tcb_at' t)"
apply unfold_locales
apply (clarsimp simp: thread_fetch_def thread_replace_def
Let_def obj_at'_def projectKOs
split: kernel_object.split option.split)
apply (clarsimp simp: thread_replace_def Let_def
split: kernel_object.split option.split)
apply (clarsimp simp: thread_fetch_def thread_replace_def Let_def
fun_upd_idem
split: kernel_object.splits option.splits)
apply (clarsimp simp: obj_at'_def thread_replace_def Let_def projectKOs
split: kernel_object.splits option.splits)
apply (rename_tac tcb)
apply (case_tac tcb, simp add: objBitsKO_def ps_clear_def)
done
lemma tcbFault_submonad_args:
"submonad_args (thread_fetch tcbFault t) (thread_replace tcbFault_update t)
(tcb_at' t)"
apply (rule thread_submonad_args)
apply (case_tac v, simp)+
done
lemma threadGet_stateAssert_gets:
"threadGet ext t = do stateAssert (tcb_at' t) []; gets (thread_fetch ext t) od"
apply (rule is_stateAssert_gets [OF _ _ empty_fail_threadGet no_fail_threadGet])
apply (clarsimp intro!: obj_at_ko_at'[where P="\<lambda>tcb :: tcb. True", simplified]
| wp threadGet_wp)+
apply (clarsimp simp: obj_at'_def thread_fetch_def projectKOs)
done
lemma threadGet_tcbFault_submonad_fn:
"threadGet tcbFault t =
submonad_fn (thread_fetch tcbFault t) (thread_replace tcbFault_update t)
(tcb_at' t) get"
apply (rule ext)
apply (clarsimp simp: submonad_fn_def bind_assoc split_def)
apply (subst threadGet_stateAssert_gets, simp)
apply (rule bind_apply_cong [OF refl])
apply (clarsimp simp: in_monad bind_def gets_def get_def return_def
submonad_args.args(3) [OF tcbFault_submonad_args]
select_f_def modify_def put_def)
done
crunch inv[wp]: getSanitiseRegisterInfo P
lemma empty_fail_getSanitiseRegisterInfo[wp, simp]:
@ -268,19 +202,6 @@ lemma asUser_mapMloadWordUser_getSanitiseRegisterInfo_comm:
od"
by (rule bind_inv_inv_comm, auto; wp mapM_wp')
lemma getObject_tcb_det:
"(r::tcb,s') \<in> fst (getObject p s) \<Longrightarrow> fst (getObject p s) = {(r,s)} \<and> s' = s"
apply (clarsimp simp add: getObject_def bind_def get_def gets_def
return_def loadObject_default_def split_def)
apply (clarsimp split: kernel_object.split_asm if_split_asm option.split_asm
simp: in_monad typeError_def alignError_def magnitudeCheck_def
objBits_def objBitsKO_def projectKOs
lookupAround2_def Let_def o_def)
apply (simp_all add: bind_def return_def assert_opt_def split_def projectKOs
alignCheck_def is_aligned_mask[symmetric]
unless_def when_def magnitudeCheck_def)
done
lemma asUser_getRegister_discarded:
"(asUser t (getRegister r)) >>= (\<lambda>_. n) =
stateAssert (tcb_at' t) [] >>= (\<lambda>_. n)"
@ -330,17 +251,6 @@ end
context kernel_m begin
(* FIXME move *)
lemma from_bool_to_bool_and_1 [simp]:
assumes r_size: "1 < size r"
shows "from_bool (to_bool (r && 1)) = r && 1"
proof -
from r_size have "r && 1 < 2"
by (simp add: and_mask_less_size [where n=1, unfolded mask_def, simplified])
thus ?thesis
by (fastforce simp add: from_bool_def to_bool_def dest: word_less_cases)
qed
(* FIXME move *)
lemma ccap_relation_ep_helpers:
"\<lbrakk> ccap_relation cap cap'; cap_get_tag cap' = scast cap_endpoint_cap \<rbrakk>
@ -1205,21 +1115,6 @@ lemma copyMRs_register_loop_helper:
rename_tac i))+
done
lemma mab_gt_2 [simp]:
"2 \<le> msg_align_bits" by (simp add: msg_align_bits)
(* FIXME move *)
lemma mapM_only_length:
"do ys \<leftarrow> mapM f xs;
g (length ys)
od =
do _ \<leftarrow> mapM_x f xs;
g (length xs)
od"
by (subst bind_return_subst [OF mapM_length])
(rule mapM_discarded)
(* FIXME move *)
lemma copyMRs_ccorres [corres]:
notes
@ -2163,14 +2058,6 @@ lemma loadCapTransfer_ctReceiveDepth:
apply wpsimp+
done
(* FIXME: move *)
lemma cte_at_0' [dest!]:
"\<lbrakk> cte_at' 0 s; no_0_obj' s \<rbrakk> \<Longrightarrow> False"
apply (clarsimp simp: cte_wp_at_obj_cases')
by (auto simp: tcb_cte_cases_def is_aligned_def objBits_defs
dest!: tcb_aligned'
split: if_split_asm)
lemma getReceiveSlots_ccorres:
"ccorres (\<lambda>a c. (a = [] \<or> (\<exists>slot. a = [slot])) \<and>
((a \<noteq> []) = (c \<noteq> NULL)) \<and> (a\<noteq>[] \<longrightarrow> c = cte_Ptr (hd a) \<and> c \<noteq> NULL))
@ -2949,32 +2836,6 @@ lemma transferCaps_ccorres [corres]:
apply clarsimp
done
(* FIXME: move *)
lemma getMessageInfo_le3:
"\<lbrace>\<top>\<rbrace> getMessageInfo sender \<lbrace>\<lambda>rv s. unat (msgExtraCaps rv) \<le> 3\<rbrace>"
including no_pre
apply (simp add: getMessageInfo_def)
apply wp
apply (rule_tac Q="\<lambda>_. \<top>" in hoare_strengthen_post)
apply wp
apply (simp add: messageInfoFromWord_def Let_def msgExtraCapBits_def)
apply (cut_tac y="r >> Types_H.msgLengthBits" in word_and_le1 [where a=3])
apply (simp add: word_le_nat_alt)
done
lemma getMessageInfo_msgLength:
"\<lbrace>\<top>\<rbrace> getMessageInfo sender \<lbrace>\<lambda>rv. K (unat (msgLength rv) \<le> msgMaxLength)\<rbrace>"
including no_pre
apply (simp add: getMessageInfo_def)
apply wp
apply (rule_tac Q="\<lambda>_. \<top>" in hoare_strengthen_post)
apply wp
apply (simp add: messageInfoFromWord_def Let_def not_less msgMaxLength_def msgLengthBits_def
split: if_split)
apply (cut_tac y="r" in word_and_le1 [where a="0x7F"])
apply (simp add: word_le_nat_alt)
done
definition
mi_from_H :: "Types_H.message_info \<Rightarrow> seL4_MessageInfo_CL"
where
@ -3982,60 +3843,11 @@ lemma handleFaultReply_ccorres [corres]:
apply (fastforce simp: seL4_Faults seL4_Arch_Faults)
done
(* FIXME: move *)
lemma cancelAllIPC_sch_act_wf:
"\<lbrace>\<lambda>s. sch_act_wf (ksSchedulerAction s) s\<rbrace>
cancelAllIPC ep
\<lbrace>\<lambda>rv s. sch_act_wf (ksSchedulerAction s) s\<rbrace>"
apply (simp add: cancelAllIPC_def)
apply (rule hoare_TrueI|wp getEndpoint_wp|wpc|simp)+
apply fastforce?
done
(* FIXME: move *)
lemma cancelAllSignals_sch_act_wf:
"\<lbrace>\<lambda>s. sch_act_wf (ksSchedulerAction s) s\<rbrace>
cancelAllSignals ep
\<lbrace>\<lambda>rv s. sch_act_wf (ksSchedulerAction s) s\<rbrace>"
apply (simp add: cancelAllSignals_def)
apply (rule hoare_TrueI|wp getNotification_wp|wpc|simp)+
apply fastforce?
done
(* FIXME: move *)
lemma cteDeleteOne_sch_act_wf:
"\<lbrace>\<lambda>s. sch_act_wf (ksSchedulerAction s) s\<rbrace>
cteDeleteOne slot
\<lbrace>\<lambda>rv s. sch_act_wf (ksSchedulerAction s) s\<rbrace>"
apply (simp add: cteDeleteOne_def unless_when split_def)
apply (simp add: finaliseCapTrue_standin_def Let_def)
apply (rule hoare_pre)
apply (wp isFinalCapability_inv cancelAllSignals_sch_act_wf
cancelAllIPC_sch_act_wf getCTE_wp' static_imp_wp
| wpc
| simp add: Let_def split: if_split)+
done
(* FIXME: move *)
lemma vp_invs_strg': "invs' s \<longrightarrow> valid_pspace' s" by auto
(* FIXME: move *)
lemma setCTE_tcbFault:
"\<lbrace>obj_at' (\<lambda>tcb. P (tcbFault tcb)) t\<rbrace>
setCTE slot cte
\<lbrace>\<lambda>rv. obj_at' (\<lambda>tcb. P (tcbFault tcb)) t\<rbrace>"
apply (simp add: setCTE_def)
apply (rule setObject_cte_obj_at_tcb', simp_all)
done
crunch tcbFault: emptySlot, tcbSchedEnqueue, rescheduleRequired
"obj_at' (\<lambda>tcb. P (tcbFault tcb)) t"
(wp: threadSet_obj_at'_strongish crunch_wps
simp: crunch_simps unless_def)
(* FIXME: move *)
lemmas threadSet_obj_at' = threadSet_obj_at'_strongish
crunch tcbFault: setThreadState, cancelAllIPC, cancelAllSignals
"obj_at' (\<lambda>tcb. P (tcbFault tcb)) t"
(wp: threadSet_obj_at'_strongish crunch_wps)
@ -4743,34 +4555,6 @@ lemma tcbEPAppend_spec:
split: if_split)
done
(* FIXME: move up to SR_lemmas_C and remove from Retype_C *)
lemma map_to_ko_atI2:
"\<lbrakk>(projectKO_opt \<circ>\<^sub>m (ksPSpace s)) x = Some v; pspace_aligned' s; pspace_distinct' s\<rbrakk> \<Longrightarrow> ko_at' v x s"
apply (clarsimp simp: map_comp_Some_iff)
apply (erule (2) aligned_distinct_obj_atI')
apply (simp add: project_inject)
done
lemma map_to_ko_at_updI':
"\<And>x x' y y' y''.
\<lbrakk> (projectKO_opt \<circ>\<^sub>m (ksPSpace s)) x = Some y;
valid_pspace' s; ko_at' y' x' s;
objBitsKO (injectKO y') = objBitsKO y''; x \<noteq> x' \<rbrakk> \<Longrightarrow>
ko_at' y x (s\<lparr>ksPSpace := ksPSpace s(x' \<mapsto> y'')\<rparr>)"
by (fastforce simp: obj_at'_def projectKOs objBitsKO_def ps_clear_upd
dest: map_to_ko_atI2)
(* FIXME: move *)
lemma state_refs_of'_upd:
"\<lbrakk> valid_pspace' s; ko_wp_at' (\<lambda>ko. objBitsKO ko = objBitsKO ko') ptr s \<rbrakk> \<Longrightarrow>
state_refs_of' (s\<lparr>ksPSpace := ksPSpace s(ptr \<mapsto> ko')\<rparr>) =
(state_refs_of' s)(ptr := refs_of' ko')"
apply (rule ext)
apply (clarsimp simp: ps_clear_upd valid_pspace'_def pspace_aligned'_def
obj_at'_def ko_wp_at'_def state_refs_of'_def
split: option.split if_split)
done
lemma sendIPC_enqueue_ccorres_helper:
"ccorres dc xfdc (valid_pspace'
and (\<lambda>s. sym_refs ((state_refs_of' s)(epptr := set queue \<times> {EPSend})))
@ -4909,11 +4693,6 @@ lemma ctcb_relation_blockingIPCCanGrantD:
thread_state_lift_def from_bool_to_bool_iff mask_eq1_nochoice)+
done
(* FIXME move *)
lemma ex_st_tcb_at'_simp[simp]:
"(\<exists>ts. st_tcb_at' ((=) ts) dest s) = tcb_at' dest s"
by (auto simp add: pred_tcb_at'_def obj_at'_def)
lemma sendIPC_ccorres [corres]:
"ccorres dc xfdc (invs' and st_tcb_at' simple' thread
and sch_act_not thread and ep_at' epptr and
@ -5035,11 +4814,7 @@ lemma sendIPC_ccorres [corres]:
apply (rule_tac ep=IdleEP in sendIPC_enqueue_ccorres_helper)
apply (simp add: valid_ep'_def)
apply (wp sts_st_tcb')
apply (rule_tac Q="\<lambda>rv. ko_wp_at' (\<lambda>x. projectKO_opt x = Some IdleEP
\<and> projectKO_opt x = (None::tcb option)) epptr"
in hoare_post_imp)
apply (clarsimp simp: obj_at'_def ko_wp_at'_def projectKOs)
apply wp
apply (clarsimp simp: obj_at'_def ko_wp_at'_def projectKOs)
apply (clarsimp simp: guard_is_UNIV_def)
apply (rule ccorres_return_Skip)
\<comment> \<open>SendEP case\<close>
@ -5065,12 +4840,7 @@ lemma sendIPC_ccorres [corres]:
apply (rule_tac ep="SendEP list" in sendIPC_enqueue_ccorres_helper)
apply (simp add: valid_ep'_def)
apply (wp sts_st_tcb')
apply (rule_tac Q="\<lambda>rv. ko_wp_at' (\<lambda>x. projectKO_opt x = Some (SendEP list)
\<and> projectKO_opt x = (None::tcb option)) epptr
and K (thread \<notin> set list)"
in hoare_post_imp)
apply (clarsimp simp: obj_at'_def ko_wp_at'_def projectKOs)
apply wp
apply (clarsimp simp: obj_at'_def ko_wp_at'_def projectKOs)
apply (clarsimp simp: guard_is_UNIV_def)
apply (rule ccorres_return_Skip)
apply (clarsimp simp: EPState_Recv_def EPState_Send_def EPState_Idle_def
@ -5111,18 +4881,18 @@ lemma sendIPC_ccorres [corres]:
apply (rule conjI[rotated])
apply (clarsimp simp: isBlockedOnSend_def ko_wp_at'_def obj_at'_def
projectKOs projectKO_opt_tcb objBits_simps')
apply (fastforce split: if_split_asm
elim: delta_sym_refs
simp: pred_tcb_at'_def obj_at'_def projectKOs
tcb_bound_refs'_def eq_sym_conv symreftype_def)
apply clarsimp
apply (fastforce split: if_split_asm
elim: delta_sym_refs
simp: pred_tcb_at'_def obj_at'_def projectKOs
tcb_bound_refs'_def eq_sym_conv symreftype_def)
apply clarsimp
apply (frule(1) sym_refs_obj_atD'[OF _ invs_sym'])
apply (frule simple_st_tcb_at_state_refs_ofD')
apply (case_tac ep, auto simp: st_tcb_at_refs_of_rev' st_tcb_at'_def
obj_at'_def projectKOs)[1]
apply (clarsimp simp: guard_is_UNIV_def)
apply (clarsimp simp: guard_is_UNIV_def)
done
apply (clarsimp simp: guard_is_UNIV_def)
done
lemma ctcb_relation_blockingIPCCanGrantReplyD:
@ -5575,95 +5345,64 @@ lemma receiveIPC_ccorres [corres]:
apply (simp add: cap_endpoint_cap_lift ccap_relation_def cap_to_H_def)
apply ceqv
apply csymbr
apply (rule ccorres_move_c_guard_tcb)
apply (rule_tac xf'=ntfnPtr_'
and r'="\<lambda>rv rv'. rv' = option_to_ptr rv \<and> rv \<noteq> Some 0"
in ccorres_split_nothrow_novcg)
apply (simp add: getBoundNotification_def)
apply (rule_tac P="no_0_obj' and valid_objs'" in threadGet_vcg_corres_P)
apply (rule allI, rule conseqPre, vcg)
apply clarsimp
apply (drule obj_at_ko_at', clarsimp)
apply (drule spec, drule(1) mp, clarsimp)
apply (clarsimp simp: typ_heap_simps ctcb_relation_def)
apply (drule(1) ko_at_valid_objs', simp add: projectKOs)
apply (clarsimp simp: option_to_ptr_def option_to_0_def projectKOs
valid_obj'_def valid_tcb'_def)
apply ceqv
apply (rename_tac ntfnptr ntfnptr')
apply (simp del: Collect_const split del: if_split cong: call_ignore_cong)
apply (rule ccorres_rhs_assoc2)
apply (rule_tac xf'=ret__int_'
and r'="\<lambda>rv rv'. (rv' = 0) = (ntfnptr = None \<or> \<not> isActive rv)"
in ccorres_split_nothrow_novcg)
apply wpc
apply (rule ccorres_from_vcg[where P=\<top> and P'=UNIV])
apply (rule allI, rule conseqPre, vcg)
apply (clarsimp simp: option_to_ptr_def option_to_0_def in_monad Bex_def)
apply (rule ccorres_pre_getNotification[where f=return, simplified])
apply (rule_tac P="\<lambda>s. ko_at' rv (the ntfnptr) s"
in ccorres_from_vcg[where P'=UNIV])
apply (rule ccorres_move_c_guard_tcb)
apply (rule_tac xf'=ntfnPtr_'
and r'="\<lambda>rv rv'. rv' = option_to_ptr rv \<and> rv \<noteq> Some 0"
in ccorres_split_nothrow_novcg)
apply (simp add: getBoundNotification_def)
apply (rule_tac P="no_0_obj' and valid_objs'" in threadGet_vcg_corres_P)
apply (rule allI, rule conseqPre, vcg)
apply clarsimp
apply (drule obj_at_ko_at', clarsimp)
apply (drule spec, drule(1) mp, clarsimp)
apply (clarsimp simp: typ_heap_simps ctcb_relation_def)
apply (drule(1) ko_at_valid_objs', simp add: projectKOs)
apply (clarsimp simp: option_to_ptr_def option_to_0_def projectKOs
valid_obj'_def valid_tcb'_def)
apply ceqv
apply (rename_tac ntfnptr ntfnptr')
apply (simp del: Collect_const split del: if_split cong: call_ignore_cong)
apply (rule ccorres_rhs_assoc2)
apply (rule_tac xf'=ret__int_'
and r'="\<lambda>rv rv'. (rv' = 0) = (ntfnptr = None \<or> \<not> isActive rv)"
in ccorres_split_nothrow_novcg)
apply wpc
apply (rule ccorres_from_vcg[where P=\<top> and P'=UNIV])
apply (rule allI, rule conseqPre, vcg)
apply (clarsimp simp: option_to_ptr_def option_to_0_def in_monad Bex_def)
apply (erule cmap_relationE1[OF cmap_relation_ntfn])
apply (rule ccorres_pre_getNotification[where f=return, simplified])
apply (rule_tac P="\<lambda>s. ko_at' rv (the ntfnptr) s"
in ccorres_from_vcg[where P'=UNIV])
apply (rule allI, rule conseqPre, vcg)
apply (clarsimp simp: option_to_ptr_def option_to_0_def in_monad Bex_def)
apply (erule cmap_relationE1[OF cmap_relation_ntfn])
apply (erule ko_at_projectKO_opt)
apply (clarsimp simp: typ_heap_simps cnotification_relation_def Let_def
notification_state_defs isActive_def
split: Structures_H.ntfn.split_asm Structures_H.notification.splits)
apply ceqv
apply (rule ccorres_cond[where R=\<top>])
apply (simp add: Collect_const_mem)
apply (ctac add: completeSignal_ccorres[unfolded dc_def])
apply (rule_tac xf'=ret__unsigned_'
and val="case ep of IdleEP \<Rightarrow> scast EPState_Idle
| RecvEP _ \<Rightarrow> scast EPState_Recv
| SendEP _ \<Rightarrow> scast EPState_Send"
and R="ko_at' ep (capEPPtr cap)"
in ccorres_symb_exec_r_known_rv_UNIV[where R'=UNIV])
apply (vcg, clarsimp)
apply (erule cmap_relationE1 [OF cmap_relation_ep])
apply (erule ko_at_projectKO_opt)
apply (clarsimp simp: typ_heap_simps cnotification_relation_def Let_def
notification_state_defs isActive_def
split: Structures_H.ntfn.split_asm Structures_H.notification.splits)
apply (clarsimp simp: typ_heap_simps cendpoint_relation_def Let_def
split: endpoint.split_asm)
apply ceqv
apply (rule ccorres_cond[where R=\<top>])
apply (simp add: Collect_const_mem)
apply (ctac add: completeSignal_ccorres[unfolded dc_def])
apply (rule_tac xf'=ret__unsigned_'
and val="case ep of IdleEP \<Rightarrow> scast EPState_Idle
| RecvEP _ \<Rightarrow> scast EPState_Recv
| SendEP _ \<Rightarrow> scast EPState_Send"
and R="ko_at' ep (capEPPtr cap)"
in ccorres_symb_exec_r_known_rv_UNIV[where R'=UNIV])
apply (vcg, clarsimp)
apply (erule cmap_relationE1 [OF cmap_relation_ep])
apply (erule ko_at_projectKO_opt)
apply (clarsimp simp: typ_heap_simps cendpoint_relation_def Let_def
split: endpoint.split_asm)
apply ceqv
apply (rule_tac A="invs' and st_tcb_at' simple' thread
and sch_act_not thread
and (\<lambda>s. \<forall>d p. thread \<notin> set (ksReadyQueues s (d, p)))
and ko_at' ep (capEPPtr cap)"
in ccorres_guard_imp2 [where A'=UNIV])
apply (rule_tac A="invs' and st_tcb_at' simple' thread
and sch_act_not thread
and (\<lambda>s. \<forall>d p. thread \<notin> set (ksReadyQueues s (d, p)))
and ko_at' ep (capEPPtr cap)"
in ccorres_guard_imp2 [where A'=UNIV])
apply wpc
\<comment> \<open>RecvEP case\<close>
apply (rule ccorres_cond_true)
apply csymbr
apply (simp only: case_bool_If from_bool_neq_0)
apply (rule ccorres_Cond_rhs, simp cong: Collect_cong split del: if_split)
apply (intro ccorres_rhs_assoc)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_split_nothrow_novcg)
apply (simp split del: if_split)
apply (rule receiveIPC_block_ccorres_helper[unfolded ptr_val_def, simplified])
apply ceqv
apply simp
apply (rename_tac list NOo)
apply (rule_tac ep="RecvEP list"
in receiveIPC_enqueue_ccorres_helper[simplified, unfolded dc_def])
apply (simp add: valid_ep'_def)
apply (wp sts_st_tcb')
apply (rename_tac list)
apply (rule_tac Q="\<lambda>rv. ko_wp_at' (\<lambda>x. projectKO_opt x = Some (RecvEP list)
\<and> projectKO_opt x = (None::tcb option))
(capEPPtr cap)
and K (thread \<notin> set list)"
in hoare_post_imp)
apply (clarsimp simp: obj_at'_def ko_wp_at'_def projectKOs)
apply wp
apply (clarsimp simp: guard_is_UNIV_def)
apply simp
apply (ctac add: doNBRecvFailedTransfer_ccorres[unfolded dc_def])
\<comment> \<open>IdleEP case\<close>
\<comment> \<open>RecvEP case\<close>
apply (rule ccorres_cond_true)
apply csymbr
apply (simp only: case_bool_If from_bool_neq_0)
@ -5678,47 +5417,81 @@ lemma receiveIPC_ccorres [corres]:
apply (rule receiveIPC_block_ccorres_helper[unfolded ptr_val_def, simplified])
apply ceqv
apply simp
apply (rule_tac ep=IdleEP
in receiveIPC_enqueue_ccorres_helper[simplified, unfolded dc_def])
apply (rename_tac list NOo)
apply (rule_tac ep="RecvEP list"
in receiveIPC_enqueue_ccorres_helper[simplified, unfolded dc_def])
apply (simp add: valid_ep'_def)
apply (wp sts_st_tcb')
apply (rule_tac Q="\<lambda>rv. ko_wp_at' (\<lambda>x. projectKO_opt x = Some IdleEP
\<and> projectKO_opt x = (None::tcb option))
(capEPPtr cap)"
in hoare_post_imp)
apply (clarsimp simp: obj_at'_def ko_wp_at'_def projectKOs)
apply wp
apply (rename_tac list)
apply (clarsimp simp: obj_at'_def ko_wp_at'_def projectKOs)
apply (clarsimp simp: guard_is_UNIV_def)
apply simp
apply (ctac add: doNBRecvFailedTransfer_ccorres[unfolded dc_def])
\<comment> \<open>SendEP case\<close>
apply (thin_tac "isBlockinga = from_bool P" for P)
apply (rule ccorres_cond_false)
apply (ctac add: doNBRecvFailedTransfer_ccorres[unfolded dc_def])
\<comment> \<open>IdleEP case\<close>
apply (rule ccorres_cond_true)
apply (intro ccorres_rhs_assoc)
apply (csymbr, csymbr, csymbr, csymbr, csymbr)
apply wpc
apply (simp only: haskell_fail_def)
apply (rule ccorres_fail)
apply (rename_tac sender rest)
apply csymbr
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_split_nothrow_novcg)
apply simp
apply (rule_tac sender=sender in receiveIPC_dequeue_ccorres_helper[simplified])
apply (simp only: case_bool_If from_bool_neq_0)
apply (rule ccorres_Cond_rhs, simp cong: Collect_cong split del: if_split)
apply (intro ccorres_rhs_assoc)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_split_nothrow_novcg)
apply (simp split del: if_split)
apply (rule receiveIPC_block_ccorres_helper[unfolded ptr_val_def, simplified])
apply ceqv
apply simp
apply (rule_tac ep=IdleEP
in receiveIPC_enqueue_ccorres_helper[simplified, unfolded dc_def])
apply (simp add: valid_ep'_def)
apply (wp sts_st_tcb')
apply (clarsimp simp: obj_at'_def ko_wp_at'_def projectKOs)
apply (clarsimp simp: guard_is_UNIV_def)
apply simp
apply (ctac add: doNBRecvFailedTransfer_ccorres[unfolded dc_def])
\<comment> \<open>SendEP case\<close>
apply (thin_tac "isBlockinga = from_bool P" for P)
apply (rule ccorres_cond_false)
apply (rule ccorres_cond_true)
apply (intro ccorres_rhs_assoc)
apply (csymbr, csymbr, csymbr, csymbr, csymbr)
apply wpc
apply (simp only: haskell_fail_def)
apply (rule ccorres_fail)
apply (rename_tac sender rest)
apply csymbr
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_split_nothrow_novcg)
apply simp
apply (rule_tac sender=sender in receiveIPC_dequeue_ccorres_helper[simplified])
apply ceqv
apply (rename_tac sender')
apply (simp only: K_bind_def haskell_assert_def return_bind)
apply (rule ccorres_move_c_guard_tcb)
apply (rule getThreadState_ccorres_foo)
apply (rename_tac sendState)
apply (rule ccorres_assert)
apply (rule ccorres_rhs_assoc2)
apply (rule_tac val="blockingIPCBadge sendState"
and xf'=badge_'
and R="\<lambda>s. \<exists>t. ko_at' t sender s \<and> tcbState t = sendState"
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 (clarsimp simp: ctcb_relation_def typ_heap_simps
cthread_state_relation_def word_size
isSend_def thread_state_lift_def
split: Structures_H.thread_state.splits)
apply ceqv
apply (rename_tac sender')
apply (simp only: K_bind_def haskell_assert_def return_bind)
apply (simp split del: if_split)
apply (rule ccorres_move_c_guard_tcb)
apply (rule getThreadState_ccorres_foo)
apply (rename_tac sendState)
apply (rule ccorres_assert)
apply (rule ccorres_rhs_assoc2)
apply (rule_tac val="blockingIPCBadge sendState"
and xf'=badge_'
apply (rule_tac val="from_bool (blockingIPCCanGrant sendState)"
and xf'=canGrant_'
and R="\<lambda>s. \<exists>t. ko_at' t sender s \<and> tcbState t = sendState"
in ccorres_symb_exec_r_known_rv_UNIV [where R'=UNIV])
apply (vcg, clarsimp)
@ -5728,20 +5501,6 @@ lemma receiveIPC_ccorres [corres]:
isSend_def thread_state_lift_def
split: Structures_H.thread_state.splits)
apply ceqv
apply (simp split del: if_split)
apply (rule ccorres_move_c_guard_tcb)
apply (rule ccorres_rhs_assoc2)
apply (rule_tac val="from_bool (blockingIPCCanGrant sendState)"
and xf'=canGrant_'
and R="\<lambda>s. \<exists>t. ko_at' t sender s \<and> tcbState t = sendState"
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 (clarsimp simp: ctcb_relation_def typ_heap_simps
cthread_state_relation_def word_size
isSend_def thread_state_lift_def
split: Structures_H.thread_state.splits)
apply ceqv
apply (rule ccorres_rhs_assoc2)
apply (rule_tac xf'=canGrantReply_'
@ -5800,14 +5559,14 @@ lemma receiveIPC_ccorres [corres]:
apply (clarsimp simp: valid_pspace_valid_objs' pred_tcb_at'_def sch_act_wf_weak
obj_at'_def)
apply (wpsimp simp: guard_is_UNIV_def option_to_ptr_def option_to_0_def conj_ac)+
apply (rule_tac Q="\<lambda>rv. valid_queues and valid_pspace'
and cur_tcb' and tcb_at' sender and tcb_at' thread
and sch_act_not sender and K (thread \<noteq> sender)
and ep_at' (capEPPtr cap)
and (\<lambda>s. ksCurDomain s \<le> maxDomain)
and (\<lambda>s. sch_act_wf (ksSchedulerAction s) s \<and>
(\<forall>d p. sender \<notin> set (ksReadyQueues s (d, p))))"
in hoare_post_imp)
apply (rule_tac Q="\<lambda>rv. valid_queues and valid_pspace'
and cur_tcb' and tcb_at' sender and tcb_at' thread
and sch_act_not sender and K (thread \<noteq> sender)
and ep_at' (capEPPtr cap)
and (\<lambda>s. ksCurDomain s \<le> maxDomain)
and (\<lambda>s. sch_act_wf (ksSchedulerAction s) s \<and>
(\<forall>d p. sender \<notin> set (ksReadyQueues s (d, p))))"
in hoare_post_imp)
subgoal by (auto, auto simp: st_tcb_at'_def obj_at'_def)
apply (wp hoare_vcg_all_lift set_ep_valid_objs')
apply (clarsimp simp: guard_is_UNIV_def)

View File

@ -117,59 +117,8 @@ lemma isolate_thread_actions_bind:
lemmas setEndpoint_obj_at_tcb' = setEndpoint_obj_at'_tcb
lemma tcbSchedEnqueue_obj_at_unchangedT:
assumes y: "\<And>f. \<forall>tcb. P (tcbQueued_update f tcb) = P tcb"
shows "\<lbrace>obj_at' P t\<rbrace> tcbSchedEnqueue t' \<lbrace>\<lambda>rv. obj_at' P t\<rbrace>"
apply (simp add: tcbSchedEnqueue_def unless_def)
apply (wp | simp add: y)+
done
(* FIXME: Move to Schedule_R.thy. Make Arch_switchToThread_obj_at a specialisation of this *)
lemma Arch_switchToThread_obj_at_pre:
"\<lbrace>obj_at' (Not \<circ> tcbQueued) t\<rbrace>
Arch.switchToThread t
\<lbrace>\<lambda>rv. obj_at' (Not \<circ> tcbQueued) t\<rbrace>"
apply (simp add: ARM_H.switchToThread_def)
apply (wp doMachineOp_obj_at setVMRoot_obj_at hoare_drop_imps|wpc)+
done
lemma rescheduleRequired_obj_at_unchangedT:
assumes y: "\<And>f. \<forall>tcb. P (tcbQueued_update f tcb) = P tcb"
shows "\<lbrace>obj_at' P t\<rbrace> rescheduleRequired \<lbrace>\<lambda>rv. obj_at' P t\<rbrace>"
apply (simp add: rescheduleRequired_def)
apply (wp tcbSchedEnqueue_obj_at_unchangedT[OF y] | wpc)+
apply simp
done
lemma setThreadState_obj_at_unchangedT:
assumes x: "\<And>f. \<forall>tcb. P (tcbState_update f tcb) = P tcb"
assumes y: "\<And>f. \<forall>tcb. P (tcbQueued_update f tcb) = P tcb"
shows "\<lbrace>obj_at' P t\<rbrace> setThreadState t' ts \<lbrace>\<lambda>rv. obj_at' P t\<rbrace>"
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: "\<And>f. \<forall>tcb. P (tcbBoundNotification_update f tcb) = P tcb"
shows "\<lbrace>obj_at' P t\<rbrace> setBoundNotification t' ts \<lbrace>\<lambda>rv. obj_at' P t\<rbrace>"
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]
lemmas setNotification_tcb = set_ntfn_tcb_obj_at'
(* FIXME: move *)
lemmas threadSet_obj_at' = threadSet_obj_at'_strongish
context kernel_m begin
lemma setObject_modify:
fixes v :: "'a :: pspace_storable" shows
@ -390,13 +339,6 @@ lemma objBits_2n:
pteBits_def pdeBits_def
split: kernel_object.split arch_kernel_object.split)
lemma magnitudeCheck_assert2:
"\<lbrakk> is_aligned x n; (1 :: word32) < 2 ^ n; ksPSpace s x = Some v \<rbrakk> \<Longrightarrow>
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: "\<And>a b c d. (loadObject a b c d :: ('a :: pspace_storable) kernel)
= loadObject_default a b c d"
@ -775,12 +717,6 @@ lemma copyMRs_simple:
apply (simp add: upto_enum_def mapM_Nil)
done
(* FIXME: MOVE *)
lemma returnOK_catch[simp]:
"(returnOk rv <catch> m) = return rv"
unfolding catch_def returnOk_def
by clarsimp
lemma doIPCTransfer_simple_rewrite:
"monadic_rewrite True True
((\<lambda>_. msgExtraCaps (messageInfoFromWord msgInfo) = 0

View File

@ -495,17 +495,6 @@ lemma cpspace_relation_ep_update_ep2:
end
context begin interpretation Arch . (*FIXME: arch_split*)
lemma setObject_tcb_ep_obj_at'[wp]:
"\<lbrace>obj_at' (P :: endpoint \<Rightarrow> bool) ptr\<rbrace> setObject ptr' (tcb :: tcb) \<lbrace>\<lambda>rv. obj_at' P ptr\<rbrace>"
apply (rule obj_at_setObject2, simp_all)
apply (clarsimp simp: updateObject_default_def in_monad)
done
end
crunch ep_obj_at'[wp]: setThreadState "obj_at' (P :: endpoint \<Rightarrow> bool) ptr"
(simp: unless_def)
context kernel_m begin
lemma ccorres_subst_basic_helper:

View File

@ -627,12 +627,6 @@ lemma ccorres_get_registers:
"StrictC'_register_defs")
done
(* FIXME: move *)
lemma st_tcb_at'_opeq_simp:
"st_tcb_at' ((=) Structures_H.thread_state.Running) (ksCurThread s) s
= st_tcb_at' (\<lambda>st. st = Structures_H.thread_state.Running) (ksCurThread s) s"
by (fastforce simp add: st_tcb_at'_def obj_at'_def)
lemma callKernel_withFastpath_corres_C:
"corres_underlying rf_sr False True dc
@ -790,26 +784,6 @@ lemma full_invs_both:
done
end
(* FIXME: move to somewhere sensible *)
lemma dom_eq:
"dom um = dom um' \<longleftrightarrow> (\<forall>a. um a = None \<longleftrightarrow> um' a = None)"
apply (simp add: dom_def del: not_None_eq)
apply (rule iffI)
apply (rule allI)
apply (simp add: set_eq_iff)
apply (drule_tac x=a in spec)
apply auto
done
lemma dom_user_mem':
"dom (user_mem' s) = {p. typ_at' UserDataT (p && ~~ mask pageBits) s}"
by (clarsimp simp:user_mem'_def dom_def pointerInUserData_def split:if_splits)
(* FIXME:move *)
lemma dom_device_mem':
"dom (device_mem' s) = {p. typ_at' UserDataDeviceT (p && ~~ mask pageBits) s}"
by (clarsimp simp: device_mem'_def dom_def pointerInDeviceData_def split: if_splits)
context kernel_m
begin
@ -825,28 +799,29 @@ lemma user_memory_update_corres_C_helper:
[p\<leftarrow>e. p \<in> dom um],
snd (t_hrs_' (globals b)))\<rparr>\<rparr>)
\<in> rf_sr"
apply (induct e)
apply simp
apply (subgoal_tac
"ksMachineState_update (underlying_memory_update (\<lambda>m. m)) a = a")
apply (simp (no_asm_simp))
apply simp
apply (rename_tac x xs)
apply (simp add: foldl_fun_upd_eq_foldr)
apply (case_tac "x \<in> dom um", simp_all)
apply (frule_tac ptr=x and b="the (um x)" in storeByteUser_rf_sr_upd)
apply (induct e)
apply simp
apply simp
apply (thin_tac "(x,y) : rf_sr" for x y)+
apply (fastforce simp add: pointerInUserData_def dom_user_mem')
apply (simp add: o_def hrs_mem_update_def)
done
apply (subgoal_tac
"ksMachineState_update (underlying_memory_update (\<lambda>m. m)) a = a")
apply (simp (no_asm_simp))
apply simp
apply (rename_tac x xs)
apply (simp add: foldl_fun_upd_eq_foldr)
apply (case_tac "x \<in> dom um", simp_all)
apply (frule_tac ptr=x and b="the (um x)" in storeByteUser_rf_sr_upd)
apply simp
apply simp
apply (thin_tac "(x,y) : rf_sr" for x y)+
apply (fastforce simp add: pointerInUserData_def dom_user_mem')
apply (simp add: o_def hrs_mem_update_def)
done
lemma user_memory_update_corres_C:
"corres_underlying rf_sr False nf (%_ _. True)
(\<lambda>s. pspace_aligned' s \<and> pspace_distinct' s \<and> dom um \<subseteq> dom (user_mem' s))
\<top>
(doMachineOp (user_memory_update um)) (setUserMem_C um)"
supply option.case_cong_weak[cong]
apply (clarsimp simp: corres_underlying_def)
apply (rule conjI)
prefer 2
@ -856,7 +831,7 @@ lemma user_memory_update_corres_C:
modify (ksMachineState_update (underlying_memory_update
(\<lambda>m. foldl (\<lambda>f p. f(p := the (um p))) m [p\<leftarrow>enum. p \<in> dom um])))
a")
prefer 2
prefer 2
apply (clarsimp simp add: doMachineOp_def user_memory_update_def
simpler_modify_def simpler_gets_def select_f_def
NonDetMonad.bind_def return_def)

View File

@ -5341,6 +5341,7 @@ lemma pspace_no_overlap_induce_cte:
is_aligned ptr bits; bits < word_bits;
pspace_no_overlap' ptr bits s\<rbrakk>
\<Longrightarrow> {ptr_val xa..+size_of TYPE(cte_C)} \<inter> {ptr..+2 ^ bits} = {}"
supply cteSizeBits_le_cte_level_bits[simp del]
apply (clarsimp simp: cpspace_relation_def)
apply (clarsimp simp: cmap_relation_def size_of_def)
apply (subgoal_tac "xa\<in>cte_Ptr ` dom (ctes_of s)")

View File

@ -1454,15 +1454,6 @@ lemma ctcb_relation_null_queue_ptrs:
apply (simp add: ctcb_relation_def tcb_null_queue_ptrs_def)
done
lemma map_to_ko_atI':
"\<lbrakk>(projectKO_opt \<circ>\<^sub>m (ksPSpace s)) x = Some v; invs' s\<rbrakk> \<Longrightarrow> ko_at' v x s"
apply (clarsimp simp: map_comp_Some_iff)
apply (erule aligned_distinct_obj_atI')
apply clarsimp
apply clarsimp
apply (simp add: project_inject)
done
(* Levity: added (20090419 09:44:27) *)
declare ntfnQueue_head_mask_4 [simp]
@ -1796,10 +1787,6 @@ lemma rf_sr_heap_device_data_relation:
lemma ko_at_projectKO_opt:
"ko_at' ko p s \<Longrightarrow> (projectKO_opt \<circ>\<^sub>m ksPSpace s) p = Some ko"
by (clarsimp elim!: obj_atE' simp: projectKOs)
lemma user_word_at_cross_over:
"\<lbrakk> user_word_at x p s; (s, s') \<in> rf_sr; p' = Ptr p \<rbrakk>
\<Longrightarrow> c_guard p' \<and> hrs_htd (t_hrs_' (globals s')) \<Turnstile>\<^sub>t p'

View File

@ -8,26 +8,6 @@ theory Schedule_C
imports Tcb_C
begin
context begin interpretation Arch . (*FIXME: arch_split*)
(* FIXME move *)
lemma setVMRoot_valid_queues':
"\<lbrace> valid_queues' \<rbrace> setVMRoot a \<lbrace> \<lambda>_. valid_queues' \<rbrace>"
by (rule valid_queues_lift'; wp)
(* FIXME move to REFINE *)
crunches Arch.switchToThread
for valid_queues'[wp]: valid_queues'
(ignore: clearExMonitor)
crunches switchToIdleThread
for ksCurDomain[wp]: "\<lambda>s. P (ksCurDomain s)"
crunches switchToIdleThread, switchToThread
for valid_pspace'[wp]: valid_pspace'
crunches switchToThread
for valid_arch_state'[wp]: valid_arch_state'
end
(*FIXME: arch_split: move up?*)
context Arch begin
context begin global_naming global
@ -39,16 +19,6 @@ end
context kernel_m begin
(* FIXME: move to Refine *)
lemma valid_idle'_tcb_at'_ksIdleThread:
"valid_idle' s \<Longrightarrow> tcb_at' (ksIdleThread s) s"
by (clarsimp simp: valid_idle'_def pred_tcb_at'_def obj_at'_def)
(* FIXME: move to Refine *)
lemma invs_no_cicd'_valid_idle':
"invs_no_cicd' s \<Longrightarrow> valid_idle' s"
by (simp add: invs_no_cicd'_def)
lemma Arch_switchToIdleThread_ccorres:
"ccorres dc xfdc (invs_no_cicd') UNIV []
Arch.switchToIdleThread (Call Arch_switchToIdleThread_'proc)"
@ -58,11 +28,6 @@ lemma Arch_switchToIdleThread_ccorres:
apply (clarsimp simp: valid_idle'_tcb_at'_ksIdleThread[OF invs_no_cicd'_valid_idle'])
done
(* FIXME: move *)
lemma empty_fail_getIdleThread [simp,intro!]:
"empty_fail getIdleThread"
by (simp add: getIdleThread_def)
lemma switchToIdleThread_ccorres:
"ccorres dc xfdc invs_no_cicd' UNIV []
switchToIdleThread (Call switchToIdleThread_'proc)"
@ -345,7 +310,7 @@ lemma scheduleChooseNewThread_ccorres:
chooseThread
od)
(Call scheduleChooseNewThread_'proc)"
apply (cinit)
apply (cinit')
apply (rule ccorres_pre_getDomainTime)
apply (rule ccorres_split_nothrow)
apply (rule_tac R="\<lambda>s. ksDomainTime s = domainTime" in ccorres_when)
@ -372,6 +337,7 @@ lemma isHighestPrio_ccorres:
(* FIXME: these should likely be in simpset for CRefine, or even in general *)
supply from_bool_eq_if[simp] from_bool_eq_if'[simp] from_bool_0[simp] if_1_0_0[simp]
ccorres_IF_True[simp]
supply if_weak_cong[cong]
apply (cinit lift: dom_' prio_')
apply clarsimp
apply (rule ccorres_move_const_guard)

View File

@ -492,34 +492,6 @@ lemma invocationCatch_use_injection_handler:
split: sum.split)
done
lemma injection_handler_returnOk:
"injection_handler injector (returnOk v) = returnOk v"
by (simp add: returnOk_liftE injection_liftE)
lemma injection_handler_If:
"injection_handler injector (If P a b)
= If P (injection_handler injector a)
(injection_handler injector b)"
by (simp split: if_split)
(* FIXME: duplicated in CSpace_All *)
lemma injection_handler_liftM:
"injection_handler f
= liftM (\<lambda>v. case v of Inl ex \<Rightarrow> Inl (f ex) | Inr rv \<Rightarrow> Inr rv)"
apply (intro ext, simp add: injection_handler_def liftM_def
handleE'_def)
apply (rule bind_apply_cong, rule refl)
apply (simp add: throwError_def split: sum.split)
done
lemma injection_handler_throwError:
"injection_handler f (throwError v) = throwError (f v)"
by (simp add: injection_handler_def handleE'_def
throwError_bind)
lemmas injection_handler_bindE = injection_bindE [OF refl refl]
lemmas injection_handler_wp = injection_wp [OF refl]
lemma ccorres_injection_handler_csum1:
"ccorres (f \<currency> r) xf P P' hs a c
\<Longrightarrow> ccorres
@ -755,9 +727,6 @@ lemma capFVMRights_range:
by (simp add: cap_frame_cap_lift_def cap_small_frame_cap_lift_def
cap_lift_def cap_tag_defs word_and_le1 mask_def)+
lemma dumb_bool_for_all: "(\<forall>x. x) = False"
by auto
lemma ccap_relation_page_is_device:
"ccap_relation (capability.ArchObjectCap (arch_capability.PageCap d v0a v1 v2 v3)) c
\<Longrightarrow> (generic_frame_cap_get_capFIsDevice_CL (cap_lift c) \<noteq> 0) = d"

View File

@ -713,10 +713,6 @@ lemma handleFault_ccorres:
apply (clarsimp simp: pred_tcb_at')
done
lemma invs_queues_imp:
"invs' s \<longrightarrow> valid_queues s"
by clarsimp
(* FIXME: move *)
lemma length_CL_from_H [simp]:
"length_CL (mi_from_H mi) = msgLength mi"
@ -1149,16 +1145,6 @@ lemma cap_case_EndpointCap_NotificationCap:
split: capability.split)
(* FIXME: MOVE *)
lemma capFaultOnFailure_if_case_sum:
" (capFaultOnFailure epCPtr b (if c then f else g) >>=
sum.case_sum (handleFault thread) return) =
(if c then ((capFaultOnFailure epCPtr b f)
>>= sum.case_sum (handleFault thread) return)
else ((capFaultOnFailure epCPtr b g)
>>= sum.case_sum (handleFault thread) return))"
by (case_tac c, clarsimp, clarsimp)
lemma invs_valid_objs_strengthen:
"invs' s \<longrightarrow> valid_objs' s" by fastforce

View File

@ -203,14 +203,6 @@ proof -
qed
(* MOVE *)
lemma tcb_aligned':
"tcb_at' t s \<Longrightarrow> is_aligned t tcbBlockSizeBits"
apply (drule obj_at_aligned')
apply (simp add: objBits_simps)
apply (simp add: objBits_simps)
done
lemma tcb_at_not_NULL:
assumes tat: "tcb_at' t s"
shows "tcb_ptr_to_ctcb_ptr t \<noteq> NULL"

View File

@ -505,6 +505,7 @@ lemma invokeTCB_ThreadControl_ccorres:
(invokeTCB (ThreadControl target slot faultep mcp priority cRoot vRoot buf))
(Call invokeTCB_ThreadControl_'proc)"
(is "ccorres ?rvr ?xf (?P and (\<lambda>_. ?P')) ?Q [] ?af ?cf")
supply option.case_cong_weak[cong]
apply (rule ccorres_gen_asm)
apply (cinit lift: target_' slot_' faultep_' mcp_' priority_' cRoot_newCap_' cRoot_srcSlot_'
vRoot_newCap_' vRoot_srcSlot_' bufferAddr_' bufferSrcSlot_' bufferCap_'
@ -532,7 +533,7 @@ lemma invokeTCB_ThreadControl_ccorres:
apply (fastforce intro: typ_heap_simps)
apply simp
apply (erule(1) rf_sr_tcb_update_no_queue2,
(simp add: typ_heap_simps)+)
(simp add: typ_heap_simps)+)
apply (rule ball_tcb_cte_casesI, simp+)
apply (clarsimp simp: ctcb_relation_def option_to_0_def)
apply (rule ccorres_return_Skip)
@ -546,8 +547,7 @@ lemma invokeTCB_ThreadControl_ccorres:
apply (rule ccorres_subgoal_tailE)
apply (rule ccorres_subgoal_tailE)
apply (rule_tac A="invs' and sch_act_simple and tcb_at' target
and (\<lambda>(s::kernel_state). (case priority of None \<Rightarrow> True
| Some x \<Rightarrow> ((\<lambda>y. fst y \<le> maxPriority)) x))
and (\<lambda>(s::kernel_state). (case priority of None \<Rightarrow> True | Some x \<Rightarrow> ((\<lambda>y. fst y \<le> maxPriority)) x))
and case_option \<top> (case_option \<top> (valid_cap' \<circ> fst) \<circ> snd) buf
and case_option \<top> (case_option \<top> (cte_at' \<circ> snd) \<circ> snd) buf
and K (case_option True (swp is_aligned msg_align_bits \<circ> fst) buf)
@ -610,9 +610,9 @@ lemma invokeTCB_ThreadControl_ccorres:
apply (ctac (no_vcg) add: rescheduleRequired_ccorres[unfolded dc_def])
apply (rule ccorres_return_Skip')
apply (rule ccorres_split_nothrow_novcg_dc)
apply(rule ccorres_cond2[where R=\<top>], simp add: Collect_const_mem)
apply(ctac add: setPriority_ccorres)
apply(rule ccorres_return_Skip)
apply (rule ccorres_cond2[where R=\<top>], simp add: Collect_const_mem)
apply (ctac add: setPriority_ccorres)
apply (rule ccorres_return_Skip)
apply (rule ccorres_return_CE, simp+)[1]
apply (wp (once))
apply (clarsimp simp: guard_is_UNIV_def)
@ -620,7 +620,7 @@ lemma invokeTCB_ThreadControl_ccorres:
apply (strengthen sch_act_wf_weak, wp)
apply clarsimp
apply wp
apply (clarsimp simp : guard_is_UNIV_def Collect_const_mem)
apply (clarsimp simp: guard_is_UNIV_def Collect_const_mem)
apply (rule hoare_strengthen_post[
where Q= "\<lambda>rv s.
Invariants_H.valid_queues s \<and>
@ -639,7 +639,6 @@ lemma invokeTCB_ThreadControl_ccorres:
apply (rule ccorres_pre_getCurThread)
apply (rename_tac curThread)
apply (rule ccorres_split_nothrow_novcg_dc)
(* \<not>P *)
apply (simp add: when_def to_bool_def)
apply (rule_tac C'="{s. target = curThread}"
and Q="\<lambda>s. ksCurThread s = curThread"
@ -657,7 +656,7 @@ lemma invokeTCB_ThreadControl_ccorres:
apply (clarsimp simp: guard_is_UNIV_def)
apply (simp add: when_def)
apply (wp hoare_vcg_if_lift2(1) static_imp_wp, strengthen sch_act_wf_weak; wp)
apply (clarsimp simp : guard_is_UNIV_def Collect_const_mem)
apply (clarsimp simp: guard_is_UNIV_def Collect_const_mem)
apply (clarsimp simp: guard_is_UNIV_def Collect_const_mem
tcbBuffer_def size_of_def cte_level_bits_def
tcbIPCBufferSlot_def)
@ -677,14 +676,14 @@ lemma invokeTCB_ThreadControl_ccorres:
apply (ctac(no_vcg) add: rescheduleRequired_ccorres[unfolded dc_def])
apply (rule ccorres_return_Skip')
apply (rule ccorres_split_nothrow_novcg_dc)
apply (rule ccorres_cond2[where R=\<top>], simp add: Collect_const_mem)
apply (ctac add: setPriority_ccorres)
apply (rule ccorres_return_Skip)
apply(rule ccorres_cond2[where R=\<top>], simp add: Collect_const_mem)
apply(ctac add: setPriority_ccorres)
apply(rule ccorres_return_Skip)
apply (rule ccorres_return_CE, simp+)
apply wp
apply (clarsimp simp: guard_is_UNIV_def)
apply (wp hoare_vcg_if_lift2(1) static_imp_wp, strengthen sch_act_wf_weak; wp)
apply (clarsimp simp : guard_is_UNIV_def Collect_const_mem)
apply (clarsimp simp: guard_is_UNIV_def Collect_const_mem)
apply (simp add: guard_is_UNIV_def false_def Collect_const_mem)
apply (clarsimp simp: ccap_relation_def cap_thread_cap_lift cap_to_H_def)
apply simp
@ -712,7 +711,7 @@ lemma invokeTCB_ThreadControl_ccorres:
apply wpsimp
apply (wp static_imp_wp, strengthen sch_act_wf_weak, wp )
apply wp
apply (clarsimp simp : guard_is_UNIV_def Collect_const_mem)
apply (clarsimp simp: guard_is_UNIV_def Collect_const_mem)
apply (simp cong: conj_cong)
apply (rule hoare_strengthen_post[
where Q="\<lambda>a b. (Invariants_H.valid_queues b \<and>
@ -733,25 +732,26 @@ lemma invokeTCB_ThreadControl_ccorres:
(fst (the (snd (the buf))))
(cteCap a))
(snd (the (snd (the buf)))) b \<longrightarrow>
cte_wp_at' (\<lambda>scte. capMasterCap (cteCap scte)
= capMasterCap (fst (the (snd (the buf))))
\<or> is_simple_cap' (fst (the (snd (the buf)))))
(snd (the (snd (the buf)))) b \<and>
valid_mdb' b \<and>
pspace_aligned' b \<and>
cte_wp_at' (\<lambda>c. True) (snd (the (snd (the buf)))) b))"])
cte_wp_at' (\<lambda>scte. capMasterCap (cteCap scte)
= capMasterCap (fst (the (snd (the buf))))
\<or> is_simple_cap' (fst (the (snd (the buf)))))
(snd (the (snd (the buf)))) b \<and>
valid_mdb' b \<and>
pspace_aligned' b \<and>
cte_wp_at' (\<lambda>c. True) (snd (the (snd (the buf)))) b))"])
prefer 2
apply fastforce
apply (strengthen cte_is_derived_capMasterCap_strg
invs_queues invs_weak_sch_act_wf invs_sch_act_wf'
invs_valid_objs' invs_mdb' invs_pspace_aligned',
simp add: o_def)
simp add: o_def)
apply (rule_tac P="is_aligned (fst (the buf)) msg_align_bits"
in hoare_gen_asm)
apply (wp threadSet_ipcbuffer_trivial static_imp_wp
| simp
| strengthen invs_sch_act_wf' invs_valid_objs' invs_weak_sch_act_wf invs_queues
| strengthen invs_sch_act_wf' invs_valid_objs' invs_weak_sch_act_wf invs_queues
invs_valid_queues' | wp hoare_drop_imps)+
(* \<not> P *)
apply (clarsimp simp: guard_is_UNIV_def Collect_const_mem
option_to_0_def
split: option.split_asm)
@ -776,7 +776,7 @@ lemma invokeTCB_ThreadControl_ccorres:
cte_level_bits_def from_bool_def true_def size_of_def case_option_If2 )
apply (rule conjI)
apply (clarsimp simp: case_option_If2 if_n_0_0 objBits_simps' valid_cap'_def
capAligned_def word_bits_conv obj_at'_def projectKOs)
capAligned_def word_bits_conv obj_at'_def projectKOs)
apply (clarsimp simp: invs_valid_objs' invs_valid_queues'
Invariants_H.invs_queues invs_ksCurDomain_maxDomain')
apply (rule ccorres_Cond_rhs_Seq)
@ -784,7 +784,7 @@ lemma invokeTCB_ThreadControl_ccorres:
apply csymbr
apply (rule ccorres_move_array_assertion_tcb_ctes ccorres_Guard_Seq)+
apply (simp add: split_def getThreadVSpaceRoot_def locateSlot_conv
bindE_assoc
bindE_assoc liftE_bindE
del: Collect_const)
apply csymbr
apply (ctac(no_vcg) add: cteDelete_ccorres)
@ -823,9 +823,8 @@ lemma invokeTCB_ThreadControl_ccorres:
apply (simp add: false_def Collect_False ccorres_cond_iffs
del: Collect_const)
apply (rule ccorres_return_Skip[unfolded dc_def])
apply (clarsimp simp: guard_is_UNIV_def Collect_const_mem
tcbVTable_def tcbVTableSlot_def
cte_level_bits_def size_of_def option_to_0_def)
apply (fastforce simp: guard_is_UNIV_def Kernel_C.tcbVTable_def tcbVTableSlot_def
cte_level_bits_def size_of_def)
apply csymbr
apply (simp add: false_def Collect_False
del: Collect_const)
@ -851,7 +850,7 @@ lemma invokeTCB_ThreadControl_ccorres:
apply csymbr
apply (rule ccorres_move_array_assertion_tcb_ctes ccorres_Guard_Seq)+
apply (simp add: split_def getThreadCSpaceRoot_def locateSlot_conv
bindE_assoc
bindE_assoc liftE_bindE
del: Collect_const)
apply csymbr
apply (ctac(no_vcg) add: cteDelete_ccorres)
@ -868,11 +867,10 @@ lemma invokeTCB_ThreadControl_ccorres:
checkCap_inv [where P="valid_cap' c" for c]
checkCap_inv [where P="sch_act_simple"]
| simp)+
apply (clarsimp simp: guard_is_UNIV_def from_bool_def true_def
word_sle_def Collect_const_mem
option_to_0_def Kernel_C.tcbVTable_def tcbVTableSlot_def
cte_level_bits_def size_of_def cintr_def
tcb_cnode_index_defs)
apply (fastforce simp: guard_is_UNIV_def
Kernel_C.tcbVTable_def tcbVTableSlot_def
cte_level_bits_def size_of_def
tcb_cnode_index_defs)
apply (thin_tac "ccorres a1 a2 a3 a4 a5 a6 a7" for a1 a2 a3 a4 a5 a6 a7)
apply (rule ccorres_rhs_assoc)+
apply (rule checkCapAt_ccorres2)
@ -986,8 +984,8 @@ lemma setupReplyMaster_ccorres:
apply clarsimp
apply (rule conjI)
apply (erule(2) cmap_relation_updI)
apply (clarsimp simp: ccte_relation_def cap_reply_cap_lift
cte_lift_def)
apply (clarsimp simp: ccte_relation_def cap_reply_cap_lift cte_lift_def
cong: option.case_cong_weak)
apply (simp add: cte_to_H_def cap_to_H_def mdb_node_to_H_def
nullMDBNode_def c_valid_cte_def)
apply (simp add: cap_reply_cap_lift)
@ -1206,9 +1204,6 @@ lemma invokeTCB_CopyRegisters_ccorres:
apply (clarsimp dest!: global'_no_ex_cap simp: invs'_def valid_state'_def | rule conjI)+
done
(* FIXME: move *)
lemmas mapM_x_append = mapM_x_append2
lemma invokeTCB_WriteRegisters_ccorres_helper:
"\<lbrakk> unat (f (of_nat n)) = incn
\<and> g (of_nat n) = register_from_H reg \<and> n'=incn
@ -1723,6 +1718,7 @@ shows
(doE reply \<leftarrow> invokeTCB (ReadRegisters target susp n archCp);
liftE (replyOnRestart thread reply isCall) odE)
(Call invokeTCB_ReadRegisters_'proc)"
supply option.case_cong_weak[cong]
apply (rule ccorres_gen_asm)
apply (cinit' lift: tcb_src_' suspendSource_' n_' call_'
simp: invokeTCB_def liftE_bindE bind_assoc)
@ -2645,14 +2641,6 @@ lemma slotCapLongRunningDelete_ccorres:
dest!: ccte_relation_ccap_relation)
done
(* FIXME: move *)
lemma empty_fail_slotCapLongRunningDelete:
"empty_fail (slotCapLongRunningDelete slot)"
by (auto simp: slotCapLongRunningDelete_def Let_def
case_Null_If isFinalCapability_def
split: if_split
intro!: empty_fail_bind)
definition
isValidVTableRoot_C :: "cap_C \<Rightarrow> bool"
where
@ -2694,20 +2682,10 @@ lemma updateCapData_spec:
\<lbrace>ccap_relation (RetypeDecls_H.updateCapData preserve newData cap) \<acute>ret__struct_cap_C\<rbrace>"
by (simp add: updateCapData_spec)
lemma if_n_updateCapData_valid_strg:
"s \<turnstile>' cap \<longrightarrow> s \<turnstile>' (if P then cap else updateCapData prs v cap)"
by (simp add: valid_updateCapDataI split: if_split)
lemma length_excaps_map:
"length (excaps_map xcs) = length xcs"
by (simp add: excaps_map_def)
(* FIXME: move *)
lemma from_bool_all_helper:
"(\<forall>bool. from_bool bool = val \<longrightarrow> P bool)
= ((\<exists>bool. from_bool bool = val) \<longrightarrow> P (val \<noteq> 0))"
by (auto simp: from_bool_0)
lemma getSyscallArg_ccorres_foo':
"ccorres (\<lambda>a rv. rv = ucast (args ! n)) (\<lambda>x. ucast (ret__unsigned_long_' x))
(sysargs_rel args buffer and sysargs_rel_n args buffer n)
@ -3929,7 +3907,7 @@ lemma decodeBindNotification_ccorres:
(decodeBindNotification cp extraCaps >>= invocationCatch thread isBlocking isCall InvokeTCB)
(Call decodeBindNotification_'proc)"
apply (simp, rule ccorres_gen_asm)
apply (cinit lift: cap_' excaps_' simp: decodeBindNotification_def)
apply (cinit' lift: cap_' excaps_' simp: decodeBindNotification_def)
apply (simp add: bind_assoc whenE_def bind_bindE_assoc interpret_excaps_test_null
del: Collect_const cong: call_ignore_cong)
apply (rule ccorres_Cond_rhs_Seq)

View File

@ -8,30 +8,6 @@ theory VSpace_C
imports TcbAcc_C CSpace_C PSpace_C TcbQueue_C
begin
context begin interpretation Arch . (*FIXME: arch_split*)
(* FIXME: move *)
lemma empty_fail_findPDForASID[iff]:
"empty_fail (findPDForASID asid)"
apply (simp add: findPDForASID_def liftME_def)
apply (intro empty_fail_bindE, simp_all split: option.split)
apply (simp add: assertE_def split: if_split)
apply (simp add: assertE_def split: if_split)
apply (simp add: empty_fail_getObject)
apply (simp add: assertE_def liftE_bindE checkPDAt_def split: if_split)
done
(* FIXME: move *)
lemma empty_fail_findPDForASIDAssert[iff]:
"empty_fail (findPDForASIDAssert asid)"
apply (simp add: findPDForASIDAssert_def catch_def
checkPDAt_def checkPDUniqueToASID_def
checkPDASIDMapMembership_def)
apply (intro empty_fail_bind, simp_all split: sum.split)
done
end
context kernel_m begin
lemma pageBitsForSize_le:
@ -1316,7 +1292,6 @@ lemma findFreeHWASID_ccorres:
Collect_const_mem word_sless_msb_less
ucast_less[THEN order_less_le_trans]
word_0_sle_from_less)
apply (simp add: option_to_0_def)
apply (frule(1) is_inv_SomeD, clarsimp)
apply (drule subsetD, erule domI)
apply simp
@ -1951,15 +1926,6 @@ lemma flushPage_ccorres:
apply (simp add: pageBits_def mask_eq_iff_w2p word_size)
done
lemma ignoreFailure_liftM:
"ignoreFailure = liftM (\<lambda>v. ())"
apply (rule ext)+
apply (simp add: ignoreFailure_def liftM_def
catch_def)
apply (rule bind_apply_cong[OF refl])
apply (simp split: sum.split)
done
end
context begin interpretation Arch . (*FIXME: arch_split*)
@ -2636,14 +2602,6 @@ lemma ccap_relation_mapped_asid_0:
apply simp
done
(* FIXME: move *)
lemma getSlotCap_wp':
"\<lbrace>\<lambda>s. \<forall>cap. cte_wp_at' (\<lambda>c. cteCap c = cap) p s \<longrightarrow> Q cap s\<rbrace> getSlotCap p \<lbrace>Q\<rbrace>"
apply (simp add: getSlotCap_def)
apply (wp getCTE_wp')
apply (clarsimp simp: cte_wp_at_ctes_of)
done
lemma ccap_relation_PageCap_generics:
"ccap_relation (ArchObjectCap (PageCap d ptr rghts sz mapdata)) cap'
\<Longrightarrow> (mapdata \<noteq> None \<longrightarrow>
@ -2943,27 +2901,6 @@ lemma cap_lift_PDCap_Base:
declare mask_Suc_0[simp]
(* FIXME: move *)
lemma setCTE_asidpool':
"\<lbrace> ko_at' (ASIDPool pool) p \<rbrace> setCTE c p' \<lbrace>\<lambda>_. ko_at' (ASIDPool pool) p\<rbrace>"
apply (clarsimp simp: setCTE_def)
apply (simp add: setObject_def split_def)
apply (rule hoare_seq_ext [OF _ hoare_gets_post])
apply (clarsimp simp: valid_def in_monad)
apply (frule updateObject_type)
apply (clarsimp simp: obj_at'_def projectKOs)
apply (rule conjI)
apply (clarsimp simp: lookupAround2_char1)
apply (clarsimp split: if_split)
apply (case_tac obj', auto)[1]
apply (rename_tac arch_kernel_object)
apply (case_tac arch_kernel_object, auto)[1]
apply (simp add: updateObject_cte)
apply (clarsimp simp: updateObject_cte typeError_def magnitudeCheck_def in_monad
split: kernel_object.splits if_splits option.splits)
apply (clarsimp simp: ps_clear_upd' lookupAround2_char1)
done
(* FIXME: move *)
lemma asid_pool_at_rf_sr:
"\<lbrakk>ko_at' (ASIDPool pool) p s; (s, s') \<in> rf_sr\<rbrakk> \<Longrightarrow>
@ -2974,17 +2911,6 @@ lemma asid_pool_at_rf_sr:
apply clarsimp
done
(* FIXME: move *)
lemma asid_pool_at_ko:
"asid_pool_at' p s \<Longrightarrow> \<exists>pool. ko_at' (ASIDPool pool) p s"
apply (clarsimp simp: typ_at'_def obj_at'_def ko_wp_at'_def projectKOs)
apply (case_tac ko, auto)
apply (rename_tac arch_kernel_object)
apply (case_tac arch_kernel_object, auto)[1]
apply (rename_tac asidpool)
apply (case_tac asidpool, auto)[1]
done
(* FIXME: move *)
lemma setObjectASID_Basic_ccorres:
"ccorres dc xfdc \<top> {s. f s = p \<and> casid_pool_relation pool (asid_pool_C.asid_pool_C (pool' s))} hs
@ -3115,7 +3041,7 @@ lemma performASIDPoolInvocation_ccorres:
apply simp
apply (clarsimp simp: cte_wp_at_ctes_of)
apply (rule conjI)
apply (clarsimp dest!: asid_pool_at_ko simp: obj_at'_def)
apply (clarsimp dest!: asid_pool_at'_ko simp: obj_at'_def)
apply (rule cmap_relationE1[OF cmap_relation_cte], assumption+)
apply (clarsimp simp: typ_heap_simps cap_get_tag_isCap_ArchObject2
isPDCap_def isCap_simps

View File

@ -348,46 +348,9 @@ lemma projectKO_opt_UserData [simp]:
"projectKO_opt KOUserData = Some UserData"
by (simp add: projectKO_opts_defs)
(* FIXME: rewrite using ucast_ucast_mask_shift *)
lemma ucast_ucast_mask_pageBits_shift:
"ucast (ucast (p && mask pageBits >> 2) :: 10 word) = p && mask pageBits >> 2"
apply (rule word_eqI)
apply (auto simp: word_size nth_ucast nth_shiftr pageBits_def)
done
definition
"processMemory s \<equiv> (ksMachineState s) \<lparr>underlying_memory := option_to_0 \<circ> (user_mem' s)\<rparr>"
(* FIXME: rewrite using unat_ucast_mask_shift *)
lemma unat_ucast_mask_pageBits_shift:
"unat (ucast (p && mask pageBits >> 2) :: 10 word) = unat ((p::word32) && mask pageBits >> 2)"
apply (simp only: unat_ucast)
apply (rule Divides.mod_less, simp)
apply (rule unat_less_power)
apply (simp add: word_bits_def)
apply (rule shiftr_less_t2n)
apply (rule order_le_less_trans [OF word_and_le1])
apply (simp add: pageBits_def mask_def)
done
(* FIXME: rewrite using mask_shift_sum *)
lemma mask_pageBits_shift_sum:
"unat n = unat (p && mask 2) \<Longrightarrow>
(p && ~~ mask pageBits) + (p && mask pageBits >> 2) * 4 + n = (p::word32)"
apply (clarsimp simp: word_shift_by_2)
apply (subst word_plus_and_or_coroll)
apply (rule word_eqI)
apply (clarsimp simp: word_size pageBits_def nth_shiftl nth_shiftr word_ops_nth_size)
apply arith
apply (subst word_plus_and_or_coroll)
apply (rule word_eqI)
apply (clarsimp simp: word_size pageBits_def nth_shiftl nth_shiftr word_ops_nth_size)
apply (rule word_eqI)
apply (clarsimp simp: word_size pageBits_def nth_shiftl nth_shiftr word_ops_nth_size)
apply (auto simp: linorder_not_less SucSucMinus)
done
lemma user_mem_C_relation:
"\<lbrakk>cpspace_user_data_relation (ksPSpace s')
(underlying_memory (ksMachineState s')) (t_hrs_' s);
@ -548,53 +511,6 @@ lemma ran_array_map_conv:
"ran (array_map_conv f n c) = {y. \<exists>i\<le>n. f (index c (unat i)) = Some y}"
by (auto simp add: array_map_conv_def2 ran_def Collect_eq)
(* FIXME: move to somewhere sensible >>> *)
text \<open>Map inversion (implicitly assuming injectivity).\<close>
definition
"the_inv_map m = (\<lambda>s. if s\<in>ran m then Some (THE x. m x = Some s) else None)"
text \<open>Map inversion can be expressed by function inversion.\<close>
lemma the_inv_map_def2:
"the_inv_map m = (Some \<circ> the_inv_into (dom m) (the \<circ> m)) |` (ran m)"
apply (rule ext)
apply (clarsimp simp: the_inv_map_def the_inv_into_def dom_def)
apply (rule_tac f=The in arg_cong)
apply (rule ext)
apply auto
done
text \<open>The domain of a function composition with Some is the universal set.\<close>
lemma dom_comp_Some[simp]: "dom (comp Some f) = UNIV" by (simp add: dom_def)
text \<open>Assuming injectivity, map inversion produces an inversive map.\<close>
lemma is_inv_the_inv_map:
"inj_on m (dom m) \<Longrightarrow> is_inv m (the_inv_map m)"
apply (simp add: is_inv_def)
apply (intro conjI allI impI)
apply (simp add: the_inv_map_def2)
apply (auto simp add: the_inv_map_def inj_on_def dom_def)
done
lemma the_the_inv_mapI:
"inj_on m (dom m) \<Longrightarrow> m x = Some y \<Longrightarrow> the (the_inv_map m y) = x"
by (auto simp: the_inv_map_def ran_def inj_on_def dom_def)
lemma eq_restrict_map_None[simp]:
"restrict_map m A x = None \<longleftrightarrow> x ~: (A \<inter> dom m)"
by (auto simp: restrict_map_def split: if_split_asm)
lemma eq_the_inv_map_None[simp]: "the_inv_map m x = None \<longleftrightarrow> x\<notin>ran m"
by (simp add: the_inv_map_def2)
lemma is_inv_unique:
"is_inv f g \<Longrightarrow> is_inv f h \<Longrightarrow> g=h"
apply (rule ext)
apply (clarsimp simp: is_inv_def dom_def Collect_eq ran_def)
apply (drule_tac x=x in spec)+
apply (case_tac "g x", clarsimp+)
done
lemma assumes A: "is_inv f g" shows the_inv_map_eq: "the_inv_map f = g"
by (simp add: is_inv_unique[OF A A[THEN is_inv_inj, THEN is_inv_the_inv_map]])
(*<<<*)
definition
casid_map_to_H
:: "(word32[256]) \<Rightarrow> (pde_C ptr \<rightharpoonup> pde_C) \<Rightarrow> asid \<rightharpoonup> hw_asid \<times> obj_ref"
@ -607,21 +523,6 @@ definition
(the_inv_map (array_map_conv (\<lambda>x. if x=0 then None else Some x)
0xFF hw_asid_table) asid))"
(* FIXME: move *)
lemma ran_map_comp_subset: "ran (map_comp f g) <= (ran f)"
by (fastforce simp: map_comp_def ran_def split: option.splits)
(* FIXME: move *)(* NOTE: unused. *)
lemma inj_on_option_map:
"inj_on (map_option f o m) (dom m) \<Longrightarrow> inj_on m (dom m)"
by (auto simp add: inj_on_def map_option_def dom_def)
lemma eq_option_to_0_rev:
"Some 0 ~: A \<Longrightarrow> \<forall>x. \<forall>y\<in>A.
((=) \<circ> option_to_0) y x \<longrightarrow> (if x = 0 then None else Some x) = y"
by (clarsimp simp: option_to_0_def split: option.splits)
lemma inj_hwasidsI:
"asid_map_pd_to_hwasids m = set_option \<circ> c \<Longrightarrow>
inj_on (map_option snd \<circ> m) (dom m) \<Longrightarrow>
@ -786,12 +687,6 @@ lemma cready_queues_to_H_correct:
(* showing that cpspace_relation is actually unique >>>*)
(* FIXME: move *)
lemma inj_image_inv:
assumes inj_f: "inj f"
shows "f ` A = B \<Longrightarrow> inv f ` B = A"
by (drule sym) (simp add: image_inv_f_f[OF inj_f])
lemma cmap_relation_unique_0:
assumes inj_f: "inj f"
assumes r: "\<And>x y z p . \<lbrakk> r x z; r y z; a p = Some x; a' p = Some y; P p x; P' p y \<rbrakk> \<Longrightarrow> x=y"
@ -854,47 +749,6 @@ 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)
(* FIXME: move *)
lemma ran_tcb_cte_cases:
"ran tcb_cte_cases =
{(Structures_H.tcbCTable, tcbCTable_update),
(Structures_H.tcbVTable, tcbVTable_update),
(Structures_H.tcbReply, tcbReply_update),
(Structures_H.tcbCaller, tcbCaller_update),
(tcbIPCBufferFrame, tcbIPCBufferFrame_update)}"
by (auto simp add: tcb_cte_cases_def split: if_split_asm)
(* FIXME: move *)
lemma ps_clear_is_aligned_ksPSpace_None:
"\<lbrakk>ps_clear p n s; is_aligned p n; 0<d; d \<le> mask n\<rbrakk>
\<Longrightarrow> ksPSpace s (p + d) = None"
apply (simp add: ps_clear_def add_diff_eq[symmetric] mask_2pm1[symmetric])
apply (drule equals0D[where a="p + d"])
apply (simp add: dom_def word_gt_0 del: word_neq_0_conv)
apply (drule mp)
apply (rule word_plus_mono_right)
apply simp
apply (simp add: mask_2pm1)
apply (erule is_aligned_no_overflow')
apply (drule mp)
apply (case_tac "(0::word32)<2^n")
apply (frule le_m1_iff_lt[of "(2::word32)^n" d, THEN iffD1])
apply (simp add: mask_2pm1[symmetric])
apply (erule (1) is_aligned_no_wrap')
apply (simp add: is_aligned_mask mask_2pm1 not_less word_bits_def
power_overflow)
by assumption
(* FIXME: move *)
lemma ps_clear_is_aligned_ctes_None:
assumes "ps_clear p tcbBlockSizeBits s"
and "is_aligned p tcbBlockSizeBits"
shows "ksPSpace s (p + 2*2^cteSizeBits) = None"
and "ksPSpace s (p + 3*2^cteSizeBits) = None"
and "ksPSpace s (p + 4*2^cteSizeBits) = None"
by (auto intro: assms ps_clear_is_aligned_ksPSpace_None
simp: objBits_defs mask_def)+
lemma map_to_ctes_tcb_ctes:
notes if_cong[cong]
shows
@ -1097,10 +951,6 @@ lemma cpspace_ntfn_relation_unique:
kernel.tcb_at_not_NULL tcb_ptr_to_ctcb_ptr_inj
split: ntfn.splits option.splits) (* long *)
(* FIXME: move *)
lemma of_bool_inject[simp]: "of_bool a = of_bool b \<longleftrightarrow> a=b"
by (cases a) (cases b, simp_all)+
(* FIXME: move *)
lemma hap_from_vm_rights_inject[simp]:
"\<lbrakk> a \<noteq> VMNoAccess; b \<noteq> VMNoAccess \<rbrakk> \<Longrightarrow> (hap_from_vm_rights a::word32) = hap_from_vm_rights b \<longleftrightarrow> a=b"
@ -1174,23 +1024,6 @@ lemma cpspace_vcpu_relation_unique:
apply clarsimp
done
(* FIXME: move *)
lemma Collect_mono2: "Collect P \<subseteq> Collect Q \<longleftrightarrow> (\<forall>x. P x \<longrightarrow> Q x)" by auto
(* FIXME: move to Wellformed, turn valid_asid_pool' into an abbreviation >>>*)
primrec
wf_asid_pool' :: "asidpool \<Rightarrow> bool"
where
"wf_asid_pool' (ASIDPool pool) =
(dom pool \<subseteq> {0 .. 2^asid_low_bits - 1} \<and>
0 \<notin> ran pool \<and> (\<forall>x \<in> ran pool. is_aligned x pdBits))"
lemma valid_eq_wf_asid_pool'[simp]:
"valid_asid_pool' pool = (\<lambda>s. wf_asid_pool' pool)"
by (case_tac pool) simp
declare valid_asid_pool'.simps[simp del]
(*<<<*)
lemma cpspace_asidpool_relation_unique:
assumes invs: "\<forall>x\<in>ran (map_to_asidpools ah). wf_asid_pool' x"
"\<forall>x\<in>ran (map_to_asidpools ah'). wf_asid_pool' x"

View File

@ -21,16 +21,6 @@ end
context kernel_m begin
(* FIXME: move to Move *)
lemma length_superSectionPDEOffsets_simp [simp]:
"length superSectionPDEOffsets = 16"
by (clarsimp simp: superSectionPDEOffsets_def table_bits_defs upto_enum_step_def)
(* FIXME: move to Move *)
lemma length_largePagePTEOffsets_simp [simp]:
"length largePagePTEOffsets = 16"
by (clarsimp simp: largePagePTEOffsets_def table_bits_defs upto_enum_step_def)
lemma storePTE_def':
"storePTE slot pte = setObject slot pte"
unfolding storePTE_def
@ -62,7 +52,7 @@ lemma performPageTableInvocationUnmap_ccorres:
apply (rule ccorres_gen_asm)
apply (cinit lift: cap_' ctSlot_')
apply csymbr
apply (simp del: Collect_const)
apply (simp add: del: Collect_const)
apply (rule ccorres_split_nothrow_novcg_dc)
apply (subgoal_tac "capPTMappedAddress cap
= (\<lambda>cp. if to_bool (capPTIsMapped_CL cp)
@ -79,7 +69,7 @@ lemma performPageTableInvocationUnmap_ccorres:
apply csymbr
apply (simp add: storePTE_def' swp_def)
apply (ctac add: clearMemory_setObject_PTE_ccorres[simplified objBits_InvalidPTE,
unfolded dc_def])
unfolded dc_def, simplified])
apply wp
apply (simp del: Collect_const)
apply (vcg exspec=unmapPageTable_modifies)
@ -1227,13 +1217,6 @@ lemma lookupPTSlot_le_0x3C:
apply (simp add: word_bits_def)
done
(* FIXME: ARMHYP move, to SR_Lemmas? *)
lemma isPTE_exclusion:
"isInvalidPTE pte \<Longrightarrow> \<not> (isSmallPagePTE pte) \<and> \<not> (isLargePagePTE pte)"
"isLargePagePTE pte \<Longrightarrow> \<not> (isInvalidPTE pte) \<and> \<not> (isSmallPagePTE pte)"
"isSmallPagePTE pte \<Longrightarrow> \<not> (isInvalidPTE pte) \<and> \<not> (isLargePagePTE pte)"
by (cases pte ; clarsimp simp: isInvalidPTE_def isSmallPagePTE_def isLargePagePTE_def)+
lemma createSafeMappingEntries_PTE_ccorres:
"ccorres (syscall_error_rel \<currency> (\<lambda>rv rv'. isLeft rv \<and> cpte_relation (fst (theLeft rv)) (fst rv')
\<and> pte_range_relation (snd (theLeft rv)) (snd rv')))
@ -1583,22 +1566,6 @@ lemma pte_sadness:
apply (cases pte', cases pte, simp)
done
(* FIXME: move to Lib *)
lemma hd_in_zip_set:
"slots \<noteq> [] \<Longrightarrow> (hd slots, 0) \<in> set (zip slots [0.e.of_nat (length slots - Suc 0)::machine_word])"
by (cases slots; simp add: upto_enum_word upto_0_to_n2 del: upt_Suc)
(* FIXME: move to Lib *)
lemma last_in_zip_set:
"\<lbrakk> slots \<noteq> []; length js = length slots \<rbrakk> \<Longrightarrow> (last slots, last js) \<in> set (zip slots js)"
apply (simp add: in_set_zip last_conv_nth)
apply (rule_tac x="length slots - 1" in exI)
apply clarsimp
apply (subst last_conv_nth)
apply (cases js; simp)
apply simp
done
lemma pte_lift_to_small:
"pte_lift pte = Some (Pte_pte_small pte') \<Longrightarrow> pte_pte_small_lift pte = pte'"
by (simp add: pte_pte_small_lift_def)
@ -1766,7 +1733,7 @@ lemma performPageInvocationMapPTE_ccorres:
apply (simp add: storePTE_def' split_def)
apply (rule obj_at_setObject3)
apply simp
apply (simp add: objBits_simps archObjSize_def pteBits_def)
apply (simp add: objBits_simps archObjSize_def pteBits_def pte_bits_def)
apply (simp add: typ_at_to_obj_at_arches[symmetric])
apply ((wp mapM_x_wp_inv hoare_vcg_ex_lift | simp add: split_def valid_pte_slots'2_def)+)[2]
apply clarsimp
@ -1778,9 +1745,9 @@ lemma performPageInvocationMapPTE_ccorres:
apply (rule conj_assoc[where Q="a \<le> b" for a b, THEN iffD1])+
apply (rule conjI)
(* the inequalities first *)
apply (clarsimp simp: valid_pte_slots'2_def
objBits_simps archObjSize_def hd_conv_nth pteBits_def)
apply (clarsimp simp:pte_range_relation_def ptr_range_to_list_def ptr_add_def)
apply (clarsimp simp: valid_pte_slots'2_def objBits_simps archObjSize_def hd_conv_nth
pteBits_def pte_bits_def)
apply (clarsimp simp: pte_range_relation_def ptr_range_to_list_def ptr_add_def)
apply (frule is_aligned_addrFromPPtr_n,simp)
apply (cut_tac n = "sz + 3" in power_not_zero[where 'a="machine_word_len"])
apply simp
@ -1792,7 +1759,7 @@ lemma performPageInvocationMapPTE_ccorres:
apply (auto simp: valid_pte_slots'2_def upt_conv_Cons[where i=0])[1]
apply (clarsimp simp: guard_is_UNIV_def Collect_const_mem hd_conv_nth last_conv_nth ucast_minus)
apply (clarsimp simp: pte_range_relation_def ptr_range_to_list_def objBits_simps
archObjSize_def pteBits_def)
archObjSize_def pteBits_def pte_bits_def)
apply (simp add: CTypesDefs.ptr_add_def ucast_nat_def word_0_sle_from_less)
apply (clarsimp simp: valid_pte_slots'2_def del: disjCI)
apply (erule disjE, simp_all add: unat_arith_simps)[1]
@ -1840,16 +1807,6 @@ lemma pde_align_ptBits:
apply (simp add: table_bits_defs)
done
(* FIXME: consider generalising and moving to Word_Lemmas *)
lemma vaddr_segment_nonsense3_folded:
"is_aligned (p :: word32) pageBits \<Longrightarrow>
(p + ((vaddr >> pageBits) && mask (pt_bits - pte_bits) << pte_bits) && ~~ mask pt_bits) = p"
apply (rule is_aligned_add_helper[THEN conjunct2])
apply (simp add: vspace_bits_defs mask_def)+
apply (rule shiftl_less_t2n[where m=12 and n=3, simplified, OF and_mask_less'[where n=9, unfolded mask_def, simplified]])
apply simp+
done
lemma createMappingEntries_valid_pte_slots'2:
"\<lbrace>K (vmsz_aligned' vptr sz \<and> vmsz_aligned' base sz) and valid_objs'\<rbrace>
createMappingEntries base vptr sz vm_rights attrib pt
@ -1893,19 +1850,6 @@ lemma createMappingEntries_valid_pte_slots'2:
apply (wp | simp add:valid_pte_slots'2_def)+
done
(* FIXME: rewrite using 'unat_shiftr_shiftl_mask_zero *)
(* FIXME: move *)
(* this one is specialised to a PDE for a supersection *)
lemma vaddr_segment_nonsense6:
"is_aligned (p :: word32) 14 \<Longrightarrow>
(p + (vaddr >> 21 << 3) && ~~ mask 14) = p"
apply (rule is_aligned_add_helper[THEN conjunct2])
apply (erule is_aligned_weaken, simp)
apply simp
apply (rule shiftl_less_t2n[where m=14 and n=3 and 'a=machine_word_len, simplified])
apply (rule shiftr_less_t2n'[where m=11 and n=21 and 'a=machine_word_len, simplified])
done
(* replay of proof in Arch_R with stronger validity result *)
lemma createMappingEntries_valid_pde_slots'2:
"\<lbrace>page_directory_at' pd and K (vmsz_aligned' vptr sz \<and> vptr < kernelBase \<and> vmsz_aligned' base sz)\<rbrace>
@ -1915,7 +1859,7 @@ lemma createMappingEntries_valid_pde_slots'2:
apply (cases sz, simp_all)
apply wpsimp+
apply (clarsimp simp: page_directory_at'_def
lookup_pd_slot_eq[simplified pdBits_eq])
lookup_pd_slot_eq[simplified pd_bits_def pde_bits_def])
apply (clarsimp simp: lookup_pd_slot_def Let_def mask_add_aligned table_bits_defs)
apply (rule conjI, erule less_kernelBase_valid_pde_offset'')
apply (rule conjI, subst vaddr_segment_nonsense6, assumption, blast)
@ -2165,7 +2109,7 @@ lemma performPageInvocationMapPDE_ccorres:
apply (simp add: storePDE_def' split_def)
apply (rule obj_at_setObject3)
apply simp
apply (simp add: objBits_simps archObjSize_def pdeBits_def)
apply (simp add: objBits_simps archObjSize_def pdeBits_def pde_bits_def)
apply (simp add: typ_at_to_obj_at_arches[symmetric])
apply ((wp mapM_x_storePDE_pde_mappings' mapM_x_wp' valid_pde_slots_lift2 | simp add: split_def)+)[2]
apply (clarsimp simp: valid_pde_mapping'_def valid_pde_slots'2_def)
@ -2180,7 +2124,7 @@ lemma performPageInvocationMapPDE_ccorres:
apply (rule conjI)
(* the inequalities first *)
apply (clarsimp simp: valid_pde_slots'2_def
objBits_simps archObjSize_def hd_conv_nth pdeBits_def)
objBits_simps archObjSize_def hd_conv_nth pdeBits_def pde_bits_def)
apply (clarsimp simp:pde_range_relation_def ptr_range_to_list_def ptr_add_def)
apply (frule is_aligned_addrFromPPtr_n,simp)
apply (cut_tac n = "sz + 3" in power_not_zero[where 'a="machine_word_len"])
@ -2193,7 +2137,7 @@ lemma performPageInvocationMapPDE_ccorres:
apply (auto simp: valid_pde_slots'2_def upt_conv_Cons[where i=0])[1]
apply (clarsimp simp: guard_is_UNIV_def Collect_const_mem hd_conv_nth last_conv_nth ucast_minus)
apply (clarsimp simp: pde_range_relation_def ptr_range_to_list_def objBits_simps
archObjSize_def pdeBits_def)
archObjSize_def pdeBits_def pde_bits_def)
apply (simp add: CTypesDefs.ptr_add_def ucast_nat_def word_0_sle_from_less)
apply (clarsimp simp: valid_pde_slots'2_def del: disjCI)
apply (erule disjE, simp_all add: unat_arith_simps)[1]
@ -2201,7 +2145,7 @@ lemma performPageInvocationMapPDE_ccorres:
apply (wp valid_pde_slots_lift2 hoare_drop_imps)
apply vcg
apply (wp valid_pde_slots_lift2 hoare_drop_imps hoare_vcg_all_lift)
apply (vcg)
apply vcg
apply simp
apply (strengthen refl[where t=True] UNIV_I, simp)?
apply (rule conjI, fastforce)
@ -2281,30 +2225,6 @@ lemma framesize_from_H_eq_eq:
lemmas framesize_from_H_eq_eqs = framesize_from_H_eq_eq trans [OF eq_commute framesize_from_H_eq_eq]
(* FIXME: generalise, move to Word_Lib, and/or rewrite using 'shift_then_mask_eq_shift_low_bits' *)
lemma shiftr_asid_low_bits_mask_asid_high_bits:
"(asid :: word32) \<le> mask asid_bits
\<Longrightarrow> (asid >> asid_low_bits) && mask asid_high_bits = asid >> asid_low_bits"
apply (rule iffD2 [OF mask_eq_iff_w2p])
apply (simp add: asid_high_bits_def word_size)
apply (rule shiftr_less_t2n)
apply (simp add: asid_low_bits_def asid_high_bits_def mask_def)
apply (simp add: asid_bits_def)
done
(* FIXME: generalise, move to Word_Lib, and/or rewrite using 'leq_low_bits_iff_zero' *)
lemma shiftr_asid_low_bits_mask_eq_0:
"\<lbrakk> (asid :: word32) \<le> mask asid_bits; asid >> asid_low_bits = 0 \<rbrakk>
\<Longrightarrow> (asid && mask asid_low_bits = 0) = (asid = 0)"
apply (rule iffI[rotated])
apply simp
apply (rule asid_low_high_bits)
apply simp
apply (simp add: ucast_asid_high_bits_is_shift)
apply (simp add: mask_def)
apply simp
done
lemma generic_frame_cap_set_capFMappedAddress_ccap_relation:
"\<lbrakk> cap_lift c'' = generic_frame_cap_set_capFMappedAddress_CL (cap_lift c') asid addr;
ccap_relation c c'; isArchPageCap c; asid \<le> mask asid_bits; asid \<noteq> 0;
@ -2350,11 +2270,6 @@ lemma ivc_label_flush_case:
= C"
by (auto split: invocation_label.split arch_invocation_label.split)
(* FIXME: move to Lib *)
lemma list_length_less:
"(args = [] \<or> length args \<le> Suc 0) = (length args < 2)"
by (case_tac args,simp_all)
lemma injection_handler_whenE:
"injection_handler Inl (whenE a b)
@ -2446,18 +2361,6 @@ lemma setVMRootForFlush_ccorres2:
apply (auto simp: cap_get_tag_isCap_ArchObject2)
done
(* FIXME: move to Lib *)
lemma at_least_2_args:
"\<not> length args < 2 \<Longrightarrow> \<exists>a b c. args = a#b#c"
apply (case_tac args)
apply simp
apply (case_tac list)
apply simp
apply (case_tac lista)
apply simp
apply simp
done
definition
to_option :: "('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a option"
@ -2517,10 +2420,6 @@ lemma cpte_relation_get_tag_simps:
"cpte_relation (ARM_HYP_H.LargePagePTE w x y z) cpte \<Longrightarrow> pte_get_tag cpte = scast pte_pte_small"
by (clarsimp simp: cpte_relation_def pte_get_tag_alt)+
(* FIXME: move *)
lemma valid_objs_valid_pte': "\<lbrakk> valid_objs' s ; ko_at' (ko :: pte) p s \<rbrakk> \<Longrightarrow> valid_pte' ko s"
by (fastforce simp add: obj_at'_def ran_def valid_obj'_def projectKOs valid_objs'_def)
(* Note: needs valid_objs' on ARM Hypervisor since the cpte_relation says nothing about alignment
of the page addresses - therefore the only way to ensure that a C large page is aligned
is to appeal to validity of the PTE on the Haskell side *)
@ -2567,7 +2466,7 @@ lemma resolveVAddr_ccorres:
(2 ^ (ptBits - 3)) (hrs_htd (t_hrs_' (globals s)))}"
in ccorres_from_vcg_might_throw)
apply (rule allI, rule conseqPre, vcg)
apply (clarsimp simp: pteBits_def if_1_0_0 typ_heap_simps' gen_framesize_to_H_def
apply (clarsimp simp: pteBits_def pt_bits_def pte_bits_def typ_heap_simps' gen_framesize_to_H_def
c_page_sizes_noteq)
subgoal for word1 pte \<sigma> x cpte _ tag
@ -2575,7 +2474,8 @@ lemma resolveVAddr_ccorres:
prefer 2
apply (simp add: mask_def ARMLargePage_def)
\<comment> \<open>reduce to resolve_ret_rel goals first\<close>
apply (clarsimp simp: fst_return pte_get_tag_alt true_def false_def split: pte.splits)
apply (clarsimp simp: fst_return pte_get_tag_alt true_def false_def pt_bits_def pte_bits_def
split: pte.splits)
apply (safe ; clarsimp simp: cpte_relation_get_tag_simps c_pages_noteq)
(* 4 subgoals *)
apply (fastforce simp: cpte_relation_def pte_pte_small_lift_def pte_lift_def Let_def mask_def
@ -2585,7 +2485,7 @@ lemma resolveVAddr_ccorres:
done
apply (rule guard_is_UNIVI)
apply (clarsimp simp: typ_heap_simps)
apply (auto simp: ptBits_def' pageBits_def
apply (auto simp: ptBits_def' pageBits_def pt_bits_def pte_bits_def
dest!: page_table_at_rf_sr
elim: clift_array_assertion_imp)[1]
apply (rule_tac P'="{s. \<exists>v. cslift s (pde_Ptr (lookup_pd_slot pd vaddr)) = Some v
@ -2637,13 +2537,6 @@ lemma two_nat_power_pageBitsForSize_le:
by (cases vsz, simp_all add: pageBits_def)
(* FIXME: move *)
lemma is_aligned_pageBitsForSize_minimum:
"\<lbrakk> is_aligned p (pageBitsForSize sz) ; n \<le> pageBits \<rbrakk> \<Longrightarrow> is_aligned p n"
apply (cases sz; clarsimp simp: pageBits_def)
apply (erule is_aligned_weaken, simp)+
done
lemma ptrFromPAddr_add_left:
"ptrFromPAddr (x + y) = ptrFromPAddr x + y"
unfolding ptrFromPAddr_def by simp
@ -3224,18 +3117,6 @@ lemma maskCapRights_eq_Untyped [simp]:
done
(* FIXME: generalise, move to Word_Lib, and/or rewrite using
'leq_high_bits_shiftr_low_bits_leq_bits' *)
lemma le_mask_asid_bits_helper:
"x \<le> 2 ^ asid_high_bits - 1 \<Longrightarrow> (x::word32) << asid_low_bits \<le> mask asid_bits"
apply (simp add: mask_def)
apply (drule le2p_bits_unset_32)
apply (simp add: asid_high_bits_def word_bits_def)
apply (subst upper_bits_unset_is_l2p_32 [symmetric])
apply (simp add: asid_bits_def word_bits_def)
apply (clarsimp simp: asid_bits_def asid_low_bits_def asid_high_bits_def nth_shiftl)
done
declare Word_Lemmas.from_bool_mask_simp [simp]
lemma isPDFlush_fold:
@ -3343,24 +3224,11 @@ lemma framesize_from_H_mask2:
Kernel_C.ARMSuperSection_def)+
done
(* FIXME: move to Lib *)
lemma rel_option_alt_def:
"rel_option f a b = (
(a = None \<and> b = None)
\<or> (\<exists>x y. a = Some x \<and> b = Some y \<and> f x y))"
apply (case_tac a, case_tac b, simp, simp, case_tac b, auto)
done
lemma injection_handler_stateAssert_relocate:
"injection_handler Inl (stateAssert ass xs >>= f) >>=E g
= do v \<leftarrow> stateAssert ass xs; injection_handler Inl (f ()) >>=E g od"
by (simp add: injection_handler_def handleE'_def bind_bindE_assoc bind_assoc)
(* FIXME: move to where is_aligned_ptrFromPAddr is *)
lemma is_aligned_ptrFromPAddr_pageBitsForSize:
"is_aligned p (pageBitsForSize sz) \<Longrightarrow> is_aligned (ptrFromPAddr p) (pageBitsForSize sz)"
by (cases sz ; simp add: is_aligned_ptrFromPAddr_n pageBits_def)
lemma decodeARMPageDirectoryInvocation_ccorres:
notes if_cong[cong] tl_drop_1[simp]
shows

View File

@ -14,32 +14,6 @@ begin
context begin interpretation Arch . (*FIXME: arch_split*)
(* Short-hand for unfolding cumbersome machine constants *)
(* FIXME MOVE these should be in refine, and the _eq forms should NOT be declared [simp]! *)
declare ptBits_eq[simp del] (* used everywhere in CRefine, breaks clarsimp-normal form of rules *)
declare pdBits_eq[simp del] (* used everywhere in CRefine, breaks clarsimp-normal form of rules *)
declare pteBits_eq[simp del] (* used everywhere in CRefine, breaks clarsimp-normal form of rules *)
declare pdeBits_eq[simp del] (* used everywhere in CRefine, breaks clarsimp-normal form of rules *)
declare vcpuBits_eq[simp del] (* used everywhere in CRefine, breaks clarsimp-normal form of rules *)
lemmas pt_bits_def' = pt_bits_def[simplified pte_bits_def, simplified]
lemmas pd_bits_def' = pd_bits_def[simplified pde_bits_def, simplified]
lemmas page_bits_def' = page_bits_def[simplified pageBits_def, simplified]
lemmas ptBits_def' = ptBits_def[simplified pteBits_def, simplified]
lemmas pdBits_def' = pdBits_def[simplified pdeBits_def, simplified]
lemmas pt_index_bits_def' = pt_index_bits_def[simplified pt_bits_def pte_bits_def, simplified]
lemmas vcpuBits_def' = vcpuBits_def[simplified pageBits_def, simplified]
lemmas vcpu_bits_def' = vcpu_bits_def[simplified pageBits_def, simplified]
lemmas table_bits_defs = pt_bits_def' pte_bits_def pd_bits_def' pde_bits_def
pageBits_def page_bits_def'
pteBits_def pdeBits_def
pt_index_bits_def'
ptBits_def' pdBits_def'
lemmas machine_bits_defs = table_bits_defs
vcpuBits_def' vcpu_bits_def'
declare word_neq_0_conv [simp del]
(* Rule previously in the simpset, now not. *)
@ -50,20 +24,9 @@ lemmas typ_heap_simps' = typ_heap_simps c_guard_clift
lemmas asUser_return = submonad.return [OF submonad_asUser]
lemma setMRs_Nil:
"setMRs thread buffer [] = stateAssert (tcb_at' thread) [] >>= (\<lambda>_. return 0)"
unfolding setMRs_def
by (simp add: zipWithM_x_def sequence_x_def zipWith_def
asUser_return)
lemmas asUser_bind_distrib =
submonad_bind [OF submonad_asUser submonad_asUser submonad_asUser]
lemma ps_clear_upd_None:
"ksPSpace s y = None \<Longrightarrow>
ps_clear x n (ksPSpace_update (\<lambda>a. (ksPSpace s)(y := None)) s') = ps_clear x n s"
by (rule iffI | clarsimp elim!: ps_clear_domE | fastforce)+
lemma ntfnQueue_head_mask_4 :
"ntfnQueue_head_CL (notification_lift ko') && ~~ mask 4 = ntfnQueue_head_CL (notification_lift ko')"
unfolding notification_lift_def
@ -95,13 +58,6 @@ lemma no_overlap_new_cap_addrs_disjoint:
apply auto
done
lemma empty_fail_asUser[iff]:
"empty_fail m \<Longrightarrow> empty_fail (asUser t m)"
apply (simp add: asUser_def split_def)
apply (intro empty_fail_bind, simp_all)
apply (simp add: select_f_def empty_fail_def)
done
declare empty_fail_doMachineOp [simp]
lemma empty_fail_loadWordUser[intro!, simp]:
@ -133,20 +89,6 @@ lemma empty_fail_unifyFailure [intro!, simp]:
handleE'_def throwError_def
split: sum.splits)
lemma asUser_mapM_x:
"(\<And>x. empty_fail (f x)) \<Longrightarrow>
asUser t (mapM_x f xs) = do stateAssert (tcb_at' t) []; mapM_x (\<lambda>x. asUser t (f x)) xs od"
apply (simp add: mapM_x_mapM asUser_bind_distrib)
apply (subst submonad_mapM [OF submonad_asUser submonad_asUser])
apply simp
apply (simp add: asUser_return bind_assoc o_def)
apply (rule ext)
apply (rule bind_apply_cong [OF refl])+
apply (clarsimp simp: in_monad dest!: fst_stateAssertD)
apply (drule use_valid, rule mapM_wp', rule asUser_typ_ats, assumption)
apply (simp add: stateAssert_def get_def NonDetMonad.bind_def)
done
lemma asUser_get_registers:
"\<lbrace>tcb_at' target\<rbrace>
asUser target (mapM getRegister xs)
@ -178,54 +120,6 @@ lemma projectKO_user_data_device:
by (cases ko)
(auto simp: projectKO_opts_defs split: arch_kernel_object.splits)
lemma device_data_at_ko:
"typ_at' UserDataDeviceT p s \<Longrightarrow> ko_at' UserDataDevice p s"
apply (clarsimp simp: typ_at'_def obj_at'_def ko_wp_at'_def
projectKO_user_data_device projectKO_eq projectKO_eq2)
apply (case_tac ko, auto)
done
(* FIXME: move *)
lemma user_data_at_ko:
"typ_at' UserDataT p s \<Longrightarrow> ko_at' UserData p s"
apply (clarsimp simp: typ_at'_def obj_at'_def ko_wp_at'_def projectKOs)
apply (case_tac ko, auto)
done
(* FIXME: move *)
lemma map_to_ko_atI:
"\<lbrakk>(projectKO_opt \<circ>\<^sub>m ksPSpace s) x = Some v;
pspace_aligned' s; pspace_distinct' s\<rbrakk>
\<Longrightarrow> ko_at' v x s"
apply (clarsimp simp: map_comp_Some_iff)
apply (erule (2) aligned_distinct_obj_atI')
apply (simp add: project_inject)
done
lemma empty_fail_rethrowFailure:
"empty_fail f \<Longrightarrow> empty_fail (rethrowFailure fn f)"
apply (simp add: rethrowFailure_def handleE'_def)
apply (erule empty_fail_bind)
apply (simp split: sum.split)
done
lemma empty_fail_resolveAddressBits:
"empty_fail (resolveAddressBits cap cptr bits)"
proof -
note empty_fail_assertE[iff]
show ?thesis
apply (rule empty_fail_use_cutMon)
apply (induct rule: resolveAddressBits.induct)
apply (subst resolveAddressBits.simps)
apply (unfold Let_def cnode_cap_case_if fun_app_def
K_bind_def haskell_assertE_def split_def)
apply (intro empty_fail_cutMon_intros)
apply (clarsimp simp: empty_fail_drop_cutMon empty_fail_whenEs
locateSlot_conv returnOk_liftE[symmetric]
isCap_simps)+
done
qed
lemma empty_fail_getReceiveSlots:
"empty_fail (getReceiveSlots r rbuf)"
proof -
@ -265,12 +159,6 @@ lemma exec_Basic_Guard_UNIV:
end
(* FIXME MOVE to where option_to_0 is defined *)
lemma option_to_0_simps [simp]:
"option_to_0 None = 0"
"option_to_0 (Some x) = x"
by (auto simp: option_to_0_def split: option.split)
definition
"option_to_ptr \<equiv> Ptr o option_to_0"
@ -289,8 +177,4 @@ lemma option_to_ptr_not_0:
"\<lbrakk> p \<noteq> 0 ; option_to_ptr v = Ptr p \<rbrakk> \<Longrightarrow> v = Some p"
by (clarsimp simp: option_to_ptr_def option_to_0_def split: option.splits)
(* FIXME: move *)
lemma of_bool_from_bool: "of_bool = from_bool"
by (rule ext, simp add: from_bool_def split: bool.split)
end

View File

@ -117,16 +117,6 @@ lemma lookupCapAndSlot_ccorres :
done
(* FIXME: move *)
lemma injection_handler_liftM:
"injection_handler f
= liftM (\<lambda>v. case v of Inl ex \<Rightarrow> Inl (f ex) | Inr rv \<Rightarrow> Inr rv)"
apply (intro ext, simp add: injection_handler_def liftM_def
handleE'_def)
apply (rule bind_apply_cong, rule refl)
apply (simp add: throwError_def split: sum.split)
done
lemma lookupErrorOnFailure_ccorres:
"ccorres (f \<currency> r) xf P P' hs a c
\<Longrightarrow> ccorres ((\<lambda>x y z. \<exists>w. x = FailedLookup isSource w \<and> f w y z) \<currency> r)

View File

@ -35,8 +35,6 @@ lemma maskCapRights_cap_cases:
done
lemma imp_ignore: "B \<Longrightarrow> A \<longrightarrow> B" by blast
lemma wordFromVMRights_spec:
"\<forall>s. \<Gamma> \<turnstile> {s} Call wordFromVMRights_'proc \<lbrace>\<acute>ret__unsigned_long = \<^bsup>s\<^esup>vm_rights\<rbrace>"
by vcg simp?

View File

@ -599,14 +599,6 @@ lemma to_bool_false [simp]:
unfolding to_bool_def false_def
by simp
(* MOVE *)
lemma tcb_aligned':
"tcb_at' t s \<Longrightarrow> is_aligned t tcbBlockSizeBits"
apply (drule obj_at_aligned')
apply (simp add: objBits_simps)
apply (simp add: objBits_simps)
done
lemma tcb_ptr_to_ctcb_ptr_mask [simp]:
assumes tcbat: "tcb_at' thread s"

View File

@ -8,13 +8,6 @@ theory DetWP
imports "Lib.DetWPLib" "../Include_C"
begin
(* FIXME YUCK where did you come from *)
declare ptBits_eq[simp del] (* used everywhere in CRefine, breaks clarsimp-normal form of rules *)
declare pdBits_eq[simp del] (* used everywhere in CRefine, breaks clarsimp-normal form of rules *)
declare pteBits_eq[simp del] (* used everywhere in CRefine, breaks clarsimp-normal form of rules *)
declare pdeBits_eq[simp del] (* used everywhere in CRefine, breaks clarsimp-normal form of rules *)
declare vcpuBits_eq[simp del] (* used everywhere in CRefine, breaks clarsimp-normal form of rules *)
context begin interpretation Arch . (*FIXME: arch_split*)
lemma det_wp_doMachineOp [wp]:
@ -133,7 +126,6 @@ lemma det_wp_asUser [wp]:
apply simp
done
(* FIXME move into Refine somewhere *)
lemma wordSize_def':
"wordSize = 4"
unfolding wordSize_def wordBits_def

View File

@ -1393,48 +1393,6 @@ lemma ksPSpace_update_ext:
"(\<lambda>s. s\<lparr>ksPSpace := f (ksPSpace s)\<rparr>) = (ksPSpace_update f)"
by (rule ext) simp
(* FIXME: move *)
lemma valid_untyped':
notes usableUntypedRange.simps[simp del]
assumes pspace_distinct': "pspace_distinct' s" and
pspace_aligned': "pspace_aligned' s" and
al: "is_aligned ptr bits"
shows "valid_untyped' d ptr bits idx s =
(\<forall>p ko. ksPSpace s p = Some ko \<longrightarrow>
obj_range' p ko \<inter> {ptr..ptr + 2 ^ bits - 1} \<noteq> {} \<longrightarrow>
obj_range' p ko \<subseteq> {ptr..ptr + 2 ^ bits - 1} \<and>
obj_range' p ko \<inter>
usableUntypedRange (UntypedCap d ptr bits idx) = {})"
apply (simp add: valid_untyped'_def)
apply (simp add: ko_wp_at'_def)
apply (rule arg_cong[where f=All])
apply (rule ext)
apply (rule arg_cong[where f=All])
apply (rule ext)
apply (case_tac "ksPSpace s ptr' = Some ko", simp_all)
apply (frule pspace_alignedD'[OF _ pspace_aligned'])
apply (frule pspace_distinctD'[OF _ pspace_distinct'])
apply simp
apply (frule aligned_ranges_subset_or_disjoint[OF al])
apply (fold obj_range'_def)
apply (rule iffI)
apply auto[1]
apply (rule conjI)
apply (rule ccontr, simp)
apply (simp add: Set.psubset_eq)
apply (erule conjE)
apply (case_tac "obj_range' ptr' ko \<inter> {ptr..ptr + 2 ^ bits - 1} \<noteq> {}", simp)
apply (cut_tac is_aligned_no_overflow[OF al])
apply (auto simp add: obj_range'_def)[1]
apply (clarsimp simp add: usableUntypedRange.simps Int_commute)
apply (case_tac "obj_range' ptr' ko \<inter> {ptr..ptr + 2 ^ bits - 1} \<noteq> {}", simp+)
apply (cut_tac is_aligned_no_overflow[OF al])
apply (clarsimp simp add: obj_range'_def)
apply (frule is_aligned_no_overflow)
by (metis al intvl_range_conv' le_m1_iff_lt less_is_non_zero_p1
nat_le_linear power_overflow sub_wrap add_0
add_0_right word_add_increasing word_less_1 word_less_sub_1)
lemma hrs_ghost_update_comm:
"(t_hrs_'_update f \<circ> ghost'state_'_update g) =
(ghost'state_'_update g \<circ> t_hrs_'_update f)"

View File

@ -1850,22 +1850,17 @@ lemma setObject_vcpuTCB_updated_Basic_ccorres:
(hrs_mem_update (heap_update (Ptr &(vcpu_Ptr vcpuptr\<rightarrow>[''vcpuTCB_C'']))
( option_to_ctcb_ptr tptr :: tcb_C ptr)))) s)))"
apply (rule ccorres_guard_imp2)
apply (rule_tac P="ko_at' (vcpuTCBPtr_update t vcpu) vcpuptr" and P'=UNIV in setObject_ccorres_helper)
apply (simp_all add: objBits_simps archObjSize_def pageBits_def obj_tcb_at' vcpuBits_def)
apply (rule_tac P="ko_at' (vcpuTCBPtr_update t vcpu) vcpuptr" and P'=UNIV in setObject_ccorres_helper)
apply (simp_all add: objBits_simps archObjSize_def pageBits_def obj_tcb_at' vcpuBits_def vcpu_bits_def)
apply (rule conseqPre, vcg, clarsimp)
apply (rule cmap_relationE1[OF cmap_relation_vcpu], assumption, erule ko_at_projectKO_opt)
apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def typ_heap_simps'
cpspace_relation_def update_vcpu_map_tos)
apply (safe ; (clarsimp simp: cpspace_relation_def typ_heap_simps
carch_state_relation_def Let_def
update_vcpu_map_to_vcpu
cmachine_state_relation_def
update_vcpu_map_tos)?)
apply (safe ; (clarsimp simp: cpspace_relation_def typ_heap_simps carch_state_relation_def Let_def
update_vcpu_map_to_vcpu cmachine_state_relation_def update_vcpu_map_tos)?)
apply (erule cmap_relation_updI, erule ko_at_projectKO_opt, simp+)
apply (clarsimp simp add: cvcpu_relation_def cvcpu_vppi_masked_relation_def
option_to_ctcb_ptr_def
cvcpu_regs_relation_def
Let_def vcpuSCTLR_def)
apply (clarsimp simp add: cvcpu_relation_def cvcpu_vppi_masked_relation_def option_to_ctcb_ptr_def
cvcpu_regs_relation_def Let_def vcpuSCTLR_def)
apply simp
done

View File

@ -563,24 +563,6 @@ lemma hasCancelSendRights_spec:
split: capability.splits bool.splits)[1]
done
lemma updateCapData_Untyped:
"isUntypedCap a
\<Longrightarrow> updateCapData b c a = a"
by (clarsimp simp:isCap_simps updateCapData_def)
lemma ctes_of_valid_strengthen:
"(invs' s \<and> ctes_of s p = Some cte) \<longrightarrow> valid_cap' (cteCap cte) s"
apply (case_tac cte)
apply clarsimp
apply (erule ctes_of_valid_cap')
apply fastforce
done
(* FIXME move: *)
lemma hoare_vcg_imp_lift_R:
"\<lbrakk> \<lbrace>P'\<rbrace> f \<lbrace>\<lambda>rv s. \<not> P rv s\<rbrace>, -; \<lbrace>Q'\<rbrace> f \<lbrace>Q\<rbrace>, - \<rbrakk> \<Longrightarrow> \<lbrace>\<lambda>s. P' s \<or> Q' s\<rbrace> f \<lbrace>\<lambda>rv s. P rv s \<longrightarrow> Q rv s\<rbrace>, -"
by (auto simp add: valid_def validE_R_def validE_def split_def split: sum.splits)
lemma decodeCNodeInvocation_ccorres:
notes gen_invocation_type_eq[simp]
shows
@ -1772,57 +1754,6 @@ lemma t_hrs_update_use_t_hrs:
= (t_hrs_'_update (\<lambda>_. f (t_hrs_' s)) $ s)"
by simp
lemma name_seq_bound_helper:
"(\<not> CP n \<and> (\<forall>n' < n. CP n'))
\<Longrightarrow> (if \<exists>n. \<not> CP n
then simpl_sequence c' (map f [0 ..< (LEAST n. \<not> CP n)])
else c) = (simpl_sequence c' (map f [0 ..< n]))"
apply (simp add: exI[where x=n])
apply (subst Least_equality[where x=n], simp_all)
apply (rule ccontr, simp add: linorder_not_le)
done
(* FIXME: what's being proven here? it's a pure word lemma - should it go in Word_Lib? *)
lemma reset_name_seq_bound_helper:
fixes sz
fixes v :: "('a :: len) word"
defines "CP \<equiv> (\<lambda>n. ~ (v && ~~ mask sz) + of_nat n * (-1 << sz) =
((-1 :: 'a word) << sz))"
and "n \<equiv> Suc (unat (shiftR v sz))"
assumes vsz: "v + 1 < 2 ^ (len_of TYPE('a) - 1)" "2 ^ sz \<noteq> (0 :: 'a word)"
and vless: "v < v'"
shows "(\<not> CP n \<and> (\<forall>n' < n. CP n'))"
apply (clarsimp simp: shiftl_t2n field_simps less_Suc_eq_le CP_def n_def)
apply (simp add: shiftr_shiftl1[where b=sz and c=sz, simplified, symmetric]
shiftl_t2n)
apply (clarsimp simp: word_sle_msb_le shiftl_t2n[symmetric])
apply (case_tac n', simp_all)
apply (cut_tac vsz(1) order_less_le_trans[OF vless max_word_max])
apply (clarsimp simp: shiftr_shiftl1 dest!: word_add_no_overflow)
apply (drule_tac f="\<lambda>x. x - 2 ^ sz" in arg_cong, simp)
apply (metis less_irrefl order_le_less_trans order_less_trans
word_and_le2[where a=v and y="~~ mask sz"]
word_two_power_neg_ineq[OF vsz(2)])
apply (clarsimp simp add: field_simps)
apply (drule_tac f="\<lambda>x. shiftr x sz" in arg_cong)
apply simp
apply (simp add: shiftr_div_2n')
apply (simp only: linorder_not_less[symmetric], erule notE)
apply (rule order_le_less_trans[OF div_le_mono])
apply (rule_tac b="v * 2 ^ sz" for v in word_unat_less_le,
simp, rule order_refl)
apply simp
done
schematic_goal sz8_helper:
"((-1) << 8 :: addr) = ?v"
by (simp add: shiftl_t2n)
lemmas reset_name_seq_bound_helper2
= reset_name_seq_bound_helper[where sz=8 and v="v :: addr" for v,
simplified sz8_helper word_bits_def[symmetric],
THEN name_seq_bound_helper]
lemma reset_untyped_inner_offs_helper:
"\<lbrakk> cteCap cte = UntypedCap dev ptr sz idx;
i \<le> unat ((of_nat idx - 1 :: addr) div 2 ^ sz2);
@ -2655,11 +2586,6 @@ lemma ccorres_returnOk_Basic:
apply (clarsimp simp: returnOk_def return_def)
done
lemma injection_handler_whenE:
"injection_handler injf (whenE P f)
= whenE P (injection_handler injf f)"
by (simp add: whenE_def injection_handler_returnOk split: if_split)
lemma fromEnum_object_type_to_H:
"fromEnum x = unat (object_type_from_H x)"
apply (cut_tac eqset_imp_iff[where x=x, OF enum_surj])

View File

@ -680,15 +680,6 @@ lemma ccorres_pre_getQueue:
apply (simp add: maxDom_to_H maxPrio_to_H)+
done
(* FIXME: move *)
lemma threadGet_wp:
"\<lbrace>\<lambda>s. \<forall>tcb. ko_at' tcb thread s \<longrightarrow> P (f tcb) s\<rbrace> threadGet f thread \<lbrace>P\<rbrace>"
apply (rule hoare_post_imp [OF _ tg_sp'])
apply clarsimp
apply (frule obj_at_ko_at')
apply (clarsimp elim: obj_atE')
done
lemma state_relation_queue_update_helper':
"\<lbrakk> (s, s') \<in> rf_sr;
(\<forall>d p. (\<forall>t\<in>set (ksReadyQueues s (d, p)). obj_at' (inQ d p) t s)
@ -885,56 +876,6 @@ lemma cbitmap_L2_relation_update:
by (simp add: rf_sr_def cstate_relation_def Let_def carch_state_relation_def
cmachine_state_relation_def)
lemma unat_ucast_prio_shiftr_simp[simp]:
"unat (ucast (p::priority) >> n :: machine_word) = unat (p >> n)"
by (simp add: shiftr_div_2n')+
lemma unat_ucast_prio_mask_simp[simp]:
"unat (ucast (p::priority) && mask m :: machine_word) = unat (p && mask m)"
by (metis ucast_and_mask unat_ucast_8_32)
lemma unat_ucast_prio_L1_cmask_simp:
"unat (ucast (p::priority) && 0x1F :: machine_word) = unat (p && 0x1F)"
using unat_ucast_prio_mask_simp[where m=5]
by (simp add: mask_def)
lemma machine_word_and_1F_less_20:
"(w :: machine_word) && 0x1F < 0x20"
by (rule word_and_less', simp)
lemma prio_ucast_shiftr_wordRadix_helper: (* FIXME generalise *)
"(ucast (p::priority) >> wordRadix :: machine_word) < 8"
unfolding maxPriority_def numPriorities_def wordRadix_def
using unat_lt2p[where x=p]
apply (clarsimp simp add: word_less_nat_alt shiftr_div_2n' unat_ucast_upcast is_up word_le_nat_alt)
apply arith
done
lemma prio_ucast_shiftr_wordRadix_helper': (* FIXME generalise *)
"(ucast (p::priority) >> wordRadix :: machine_word) \<le> 7"
unfolding maxPriority_def numPriorities_def wordRadix_def
using unat_lt2p[where x=p]
apply (clarsimp simp add: word_less_nat_alt shiftr_div_2n' unat_ucast_upcast is_up word_le_nat_alt)
apply arith
done
lemma prio_unat_shiftr_wordRadix_helper': (* FIXME generalise *)
"unat ((p::priority) >> wordRadix) \<le> 7"
unfolding maxPriority_def numPriorities_def wordRadix_def
using unat_lt2p[where x=p]
apply (clarsimp simp add: word_less_nat_alt shiftr_div_2n' unat_ucast_upcast is_up word_le_nat_alt)
apply arith
done
lemma prio_ucast_shiftr_wordRadix_helper2:
"(ucast (p::priority) >> wordRadix :: machine_word) < 0x20"
by (rule order_less_trans[OF prio_ucast_shiftr_wordRadix_helper]; simp)
lemma dom_less_0x10_helper:
"d \<le> maxDomain \<Longrightarrow> (ucast d :: machine_word) < 0x10"
unfolding maxDomain_def numDomains_def
by (clarsimp simp add: word_less_nat_alt unat_ucast_upcast is_up word_le_nat_alt)
lemma cready_queues_index_to_C_ucast_helper:
fixes p :: priority
fixes d :: domain
@ -1009,30 +950,6 @@ lemma carch_state_relation_enqueue_simp:
unfolding carch_state_relation_def
by clarsimp
(* FIXME move to lib/Eisbach_Methods *)
(* FIXME consider printing error on solve goal apply *)
context
begin
private definition "bool_protect (b::bool) \<equiv> b"
lemma bool_protectD:
"bool_protect P \<Longrightarrow> P"
unfolding bool_protect_def by simp
lemma bool_protectI:
"P \<Longrightarrow> bool_protect P"
unfolding bool_protect_def by simp
(*
When you want to apply a rule/tactic to transform a potentially complex goal into another
one manually, but want to indicate that any fresh emerging goals are solved by a more
brutal method.
E.g. apply (solves_emerging \<open>frule x=... in my_rule\<close>\<open>fastforce simp: ... intro!: ... \<close> *)
method solves_emerging methods m1 m2 = (rule bool_protectD, (m1 ; (rule bool_protectI | (m2; fail))))
end
lemma t_hrs_ksReadyQueues_upd_absorb:
"t_hrs_'_update f (g s) \<lparr>ksReadyQueues_' := rqupd \<rparr> =
t_hrs_'_update f (g s \<lparr>ksReadyQueues_' := rqupd\<rparr>)"
@ -1364,10 +1281,6 @@ lemma rf_sr_drop_bitmaps_dequeue_helper:
carch_state_relation_def cmachine_state_relation_def
by (clarsimp simp: rf_sr_cbitmap_L1_relation rf_sr_cbitmap_L2_relation)
lemma filter_empty_unfiltered_contr:
"\<lbrakk> [x\<leftarrow>xs . x \<noteq> y] = [] ; x' \<in> set xs ; x' \<noteq> y \<rbrakk> \<Longrightarrow> False"
by (induct xs, auto split: if_split_asm)
(* FIXME same proofs as bit_set, maybe can generalise? *)
lemma cbitmap_L1_relation_bit_clear:
fixes p :: priority
@ -1409,16 +1322,6 @@ lemma cbitmap_L2_relationD:
unfolding cbitmap_L2_relation_def l2BitmapSize_def'
by clarsimp
(* FIXME move *)
lemma filter_noteq_op:
"[x \<leftarrow> xs . x \<noteq> y] = filter ((\<noteq>) y) xs"
by (induct xs) auto
(* FIXME move *)
lemma all_filter_propI:
"\<forall>x\<in>set lst. P x \<Longrightarrow> \<forall>x\<in>set (filter Q lst). P x"
by (induct lst, auto)
lemma cbitmap_L2_relation_bit_clear:
fixes p :: priority
fixes d :: domain

View File

@ -183,119 +183,11 @@ lemma asUser_comm:
apply (rule efa efb)+
done
(* FIXME move the thread_submonad stuff to SubMonad_R and use it for asUser *)
definition
"thread_fetch \<equiv> \<lambda>ext t s. case (ksPSpace s t) of
Some (KOTCB tcb) \<Rightarrow> ext tcb
| None \<Rightarrow> undefined"
definition
"thread_fetch_option \<equiv> \<lambda>ext t s. case (ksPSpace s t) of
Some (KOTCB tcb) \<Rightarrow> ext tcb
| None \<Rightarrow> None"
definition
"thread_replace \<equiv> \<lambda>upd t nv s.
let obj = case (ksPSpace s t) of
Some (KOTCB tcb) \<Rightarrow> Some (KOTCB (upd (\<lambda>_. nv) tcb))
| obj \<Rightarrow> obj
in s \<lparr> ksPSpace := (ksPSpace s) (t := obj) \<rparr>"
lemma thread_submonad_args:
"\<lbrakk> \<And>f v. ext (upd f v) = f (ext v);
\<And>f1 f2 v. upd f1 (upd f2 v) = upd (f1 \<circ> f2) v;
\<And>f v. upd (\<lambda>_. ext v) v = v \<rbrakk> \<Longrightarrow>
submonad_args (thread_fetch ext t) (thread_replace upd t) (tcb_at' t)"
apply unfold_locales
apply (clarsimp simp: thread_fetch_def thread_replace_def
Let_def obj_at'_def projectKOs
split: kernel_object.split option.split)
apply (clarsimp simp: thread_replace_def Let_def
split: kernel_object.split option.split)
apply (clarsimp simp: thread_fetch_def thread_replace_def Let_def
fun_upd_idem
split: kernel_object.splits option.splits)
apply (clarsimp simp: obj_at'_def thread_replace_def Let_def projectKOs
split: kernel_object.splits option.splits)
apply (rename_tac tcb)
apply (case_tac tcb, simp add: objBitsKO_def ps_clear_def)
done
lemma tcbFault_submonad_args:
"submonad_args (thread_fetch tcbFault t) (thread_replace tcbFault_update t)
(tcb_at' t)"
apply (rule thread_submonad_args)
apply (case_tac v, simp)+
done
lemma threadGet_stateAssert_gets:
"threadGet ext t = do stateAssert (tcb_at' t) []; gets (thread_fetch ext t) od"
apply (rule is_stateAssert_gets [OF _ _ empty_fail_threadGet no_fail_threadGet])
apply (clarsimp intro!: obj_at_ko_at'[where P="\<lambda>tcb :: tcb. True", simplified]
| wp threadGet_wp)+
apply (clarsimp simp: obj_at'_def thread_fetch_def projectKOs)
done
lemma threadGet_tcbFault_submonad_fn:
"threadGet tcbFault t =
submonad_fn (thread_fetch tcbFault t) (thread_replace tcbFault_update t)
(tcb_at' t) get"
apply (rule ext)
apply (clarsimp simp: submonad_fn_def bind_assoc split_def)
apply (subst threadGet_stateAssert_gets, simp)
apply (rule bind_apply_cong [OF refl])
apply (clarsimp simp: in_monad bind_def gets_def get_def return_def
submonad_args.args(3) [OF tcbFault_submonad_args]
select_f_def modify_def put_def)
done
lemma asUser_threadGet_tcbFault_comm:
"empty_fail im \<Longrightarrow>
do y \<leftarrow> asUser t im;
x \<leftarrow> threadGet tcbFault t';
n x y
od =
do x \<leftarrow> threadGet tcbFault t';
asUser t im >>= n x
od"
apply (rule submonad_comm2 [OF tcbFault_submonad_args
threadGet_tcbFault_submonad_fn
submonad_asUser, symmetric])
apply (clarsimp simp: thread_replace_def asUser_replace_def Let_def
split: option.split)
apply (clarsimp simp: fun_upd_idem fun_upd_twist
split: kernel_object.split)
apply (rename_tac tcb)
apply (case_tac tcb, simp)
apply (clarsimp simp: asUser_replace_def Let_def obj_at'_real_def
ko_wp_at'_def ps_clear_upd_None ps_clear_upd
objBitsKO_def projectKOs
split: option.split kernel_object.split)
apply (clarsimp simp: thread_replace_def Let_def obj_at'_real_def
ko_wp_at'_def ps_clear_upd_None
ps_clear_upd objBitsKO_def projectKOs
split: option.split kernel_object.split)
apply (simp add: get_def empty_fail_def)
apply assumption
done
lemma asUser_getRegister_threadGet_comm:
"do
ra \<leftarrow> asUser a (getRegister r);
rb \<leftarrow> threadGet fb b;
c ra rb
od = do
rb \<leftarrow> threadGet fb b;
ra \<leftarrow> asUser a (getRegister r);
c ra rb
od"
by (rule bind_inv_inv_comm, auto; wp)
crunch inv[wp]: getSanitiseRegisterInfo P
lemma empty_fail_getSanitiseRegisterInfo[wp, simp]:
"empty_fail (getSanitiseRegisterInfo t)"
by (wpsimp simp: getSanitiseRegisterInfo_def2 wp: kernel.empty_fail_archThreadGet)
by (wpsimp simp: getSanitiseRegisterInfo_def2 wp: ArchMove_C.empty_fail_archThreadGet)
lemma asUser_getRegister_getSanitiseRegisterInfo_comm:
"do
@ -333,45 +225,6 @@ lemma asUser_mapMloadWordUser_getSanitiseRegisterInfo_comm:
od"
by (rule bind_inv_inv_comm, auto; wp mapM_wp')
lemma threadGet_tcbFault_doMachineOp_comm:
"\<lbrakk> empty_fail m' \<rbrakk> \<Longrightarrow>
do x \<leftarrow> threadGet tcbFault t; y \<leftarrow> doMachineOp m'; n x y od =
do y \<leftarrow> doMachineOp m'; x \<leftarrow> threadGet tcbFault t; n x y od"
apply (rule submonad_comm2 [OF tcbFault_submonad_args
threadGet_tcbFault_submonad_fn
submonad_doMachineOp])
apply (simp add: thread_replace_def Let_def)
apply simp
apply (rule refl)
apply (simp add: get_def empty_fail_def)
apply assumption
done
lemma getObject_tcb_det:
"(r::tcb,s') \<in> fst (getObject p s) \<Longrightarrow> fst (getObject p s) = {(r,s)} \<and> s' = s"
apply (clarsimp simp add: getObject_def bind_def get_def gets_def
return_def loadObject_default_def split_def)
apply (clarsimp split: kernel_object.split_asm if_split_asm option.split_asm
simp: in_monad typeError_def alignError_def magnitudeCheck_def
objBits_def objBitsKO_def projectKOs
lookupAround2_def Let_def o_def)
apply (simp_all add: bind_def return_def assert_opt_def split_def projectKOs
alignCheck_def is_aligned_mask[symmetric]
unless_def when_def magnitudeCheck_def)
done
lemma threadGet_again:
"\<And>rv s s' n. (rv, s') \<in> fst (threadGet ext t s) \<Longrightarrow>
(threadGet ext t >>= n) s' = n rv s'"
apply (clarsimp simp add: threadGet_def liftM_def in_monad)
apply (frule use_valid [OF _ getObject_obj_at'])
apply (simp add: objBits_simps')+
apply (frule getObject_tcb_det)
apply (clarsimp simp: bind_def split_def)
apply (insert no_fail_getObject_tcb)
apply (clarsimp simp: no_fail_def obj_at'_def is_tcb)
done
lemma asUser_getRegister_discarded:
"(asUser t (getRegister r)) >>= (\<lambda>_. n) =
stateAssert (tcb_at' t) [] >>= (\<lambda>_. n)"
@ -469,17 +322,6 @@ end
context kernel_m begin
(* FIXME move *)
lemma from_bool_to_bool_and_1 [simp]:
assumes r_size: "1 < size r"
shows "from_bool (to_bool (r && 1)) = r && 1"
proof -
from r_size have "r && 1 < 2"
by (simp add: and_mask_less_size [where n=1, unfolded mask_def, simplified])
thus ?thesis
by (fastforce simp add: from_bool_def to_bool_def dest: word_less_cases)
qed
(* FIXME move *)
lemma ccap_relation_ep_helpers:
"\<lbrakk> ccap_relation cap cap'; cap_get_tag cap' = scast cap_endpoint_cap \<rbrakk>
@ -1503,23 +1345,6 @@ lemma copyMRs_register_loop_helper:
rename_tac i))+
done
lemma mab_gt_2 [simp]:
"2 \<le> msg_align_bits" by (simp add: msg_align_bits)
lemma wb_gt_2:
"2 < word_bits" by (simp add: word_bits_conv)
(* FIXME move *)
lemma mapM_only_length:
"do ys \<leftarrow> mapM f xs;
g (length ys)
od =
do _ \<leftarrow> mapM_x f xs;
g (length xs)
od"
by (subst bind_return_subst [OF mapM_length])
(rule mapM_discarded)
(* FIXME move *)
lemma copyMRs_ccorres [corres]:
@ -2665,14 +2490,6 @@ lemma loadCapTransfer_ctReceiveDepth:
apply wpsimp+
done
(* FIXME: move *)
lemma cte_at_0' [dest!]:
"\<lbrakk> cte_at' 0 s; no_0_obj' s \<rbrakk> \<Longrightarrow> False"
apply (clarsimp simp: cte_wp_at_obj_cases')
by (auto simp: tcb_cte_cases_def is_aligned_def objBits_defs
dest!: tcb_aligned'
split: if_split_asm)
lemma getReceiveSlots_ccorres:
"ccorres (\<lambda>a c. (a = [] \<or> (\<exists>slot. a = [slot])) \<and>
((a \<noteq> []) = (c \<noteq> NULL)) \<and> (a\<noteq>[] \<longrightarrow> c = cte_Ptr (hd a) \<and> c \<noteq> NULL))
@ -3502,32 +3319,6 @@ lemma transferCaps_ccorres [corres]:
apply clarsimp
done
(* FIXME: move *)
lemma getMessageInfo_le3:
"\<lbrace>\<top>\<rbrace> getMessageInfo sender \<lbrace>\<lambda>rv s. unat (msgExtraCaps rv) \<le> 3\<rbrace>"
including no_pre
apply (simp add: getMessageInfo_def)
apply wp
apply (rule_tac Q="\<lambda>_. \<top>" in hoare_strengthen_post)
apply wp
apply (simp add: messageInfoFromWord_def Let_def msgExtraCapBits_def)
apply (cut_tac y="r >> Types_H.msgLengthBits" in word_and_le1 [where a=3])
apply (simp add: word_le_nat_alt)
done
lemma getMessageInfo_msgLength:
"\<lbrace>\<top>\<rbrace> getMessageInfo sender \<lbrace>\<lambda>rv. K (unat (msgLength rv) \<le> msgMaxLength)\<rbrace>"
including no_pre
apply (simp add: getMessageInfo_def)
apply wp
apply (rule_tac Q="\<lambda>_. \<top>" in hoare_strengthen_post)
apply wp
apply (simp add: messageInfoFromWord_def Let_def not_less msgMaxLength_def msgLengthBits_def
split: if_split)
apply (cut_tac y="r" in word_and_le1 [where a="0x7F"])
apply (simp add: word_le_nat_alt)
done
definition
mi_from_H :: "Types_H.message_info \<Rightarrow> seL4_MessageInfo_CL"
where
@ -4593,77 +4384,15 @@ lemma handleFaultReply_ccorres [corres]:
apply (fastforce simp: seL4_Faults seL4_Arch_Faults)
done
(* FIXME: move *)
lemma cancelAllIPC_sch_act_wf:
"\<lbrace>\<lambda>s. sch_act_wf (ksSchedulerAction s) s\<rbrace>
cancelAllIPC ep
\<lbrace>\<lambda>rv s. sch_act_wf (ksSchedulerAction s) s\<rbrace>"
apply (simp add: cancelAllIPC_def)
apply (rule hoare_TrueI|wp getEndpoint_wp|wpc|simp)+
apply fastforce?
done
(* FIXME: move *)
lemma cancelAllSignals_sch_act_wf:
"\<lbrace>\<lambda>s. sch_act_wf (ksSchedulerAction s) s\<rbrace>
cancelAllSignals ep
\<lbrace>\<lambda>rv s. sch_act_wf (ksSchedulerAction s) s\<rbrace>"
apply (simp add: cancelAllSignals_def)
apply (rule hoare_TrueI|wp getNotification_wp|wpc|simp)+
apply fastforce?
done
(* FIXME: move *)
lemma cteDeleteOne_sch_act_wf:
"\<lbrace>\<lambda>s. sch_act_wf (ksSchedulerAction s) s\<rbrace>
cteDeleteOne slot
\<lbrace>\<lambda>rv s. sch_act_wf (ksSchedulerAction s) s\<rbrace>"
apply (simp add: cteDeleteOne_def unless_when split_def)
apply (simp add: finaliseCapTrue_standin_def Let_def)
apply (rule hoare_pre)
apply (wp isFinalCapability_inv cancelAllSignals_sch_act_wf
cancelAllIPC_sch_act_wf getCTE_wp' static_imp_wp
| wpc
| simp add: Let_def split: if_split)+
done
(* FIXME: move *)
lemma vp_invs_strg': "invs' s \<longrightarrow> valid_pspace' s" by auto
(* FIXME: move *)
lemma setCTE_tcbFault:
"\<lbrace>obj_at' (\<lambda>tcb. P (tcbFault tcb)) t\<rbrace>
setCTE slot cte
\<lbrace>\<lambda>rv. obj_at' (\<lambda>tcb. P (tcbFault tcb)) t\<rbrace>"
apply (simp add: setCTE_def)
apply (rule setObject_cte_obj_at_tcb', simp_all)
done
crunch tcbFault: emptySlot, tcbSchedEnqueue, rescheduleRequired
"obj_at' (\<lambda>tcb. P (tcbFault tcb)) t"
(wp: threadSet_obj_at'_strongish crunch_wps
simp: crunch_simps unless_def)
(* FIXME: move *)
lemmas threadSet_obj_at' = threadSet_obj_at'_strongish
crunch tcbFault: setThreadState, cancelAllIPC, cancelAllSignals
"obj_at' (\<lambda>tcb. P (tcbFault tcb)) t"
(wp: threadSet_obj_at'_strongish crunch_wps)
(* FIXME: move *)
lemmas setEndpoint_tcb = KHeap_R.setEndpoint_obj_at'_tcb
(* FIXME: move *)
lemma setNotification_tcb:
"\<lbrace>obj_at' (\<lambda>tcb::tcb. P tcb) t\<rbrace>
setNotification ntfn e
\<lbrace>\<lambda>_. obj_at' P t\<rbrace>"
apply (simp add: setNotification_def)
apply (rule obj_at_setObject2)
apply (clarsimp simp: updateObject_default_def in_monad)
done
lemma sbn_tcbFault:
"\<lbrace>obj_at' (\<lambda>tcb. P (tcbFault tcb)) t\<rbrace>
setBoundNotification st t'
@ -5365,34 +5094,6 @@ lemma tcbEPAppend_spec:
split: if_split)
done
(* FIXME: move up to SR_lemmas_C and remove from Retype_C *)
lemma map_to_ko_atI2:
"\<lbrakk>(projectKO_opt \<circ>\<^sub>m (ksPSpace s)) x = Some v; pspace_aligned' s; pspace_distinct' s\<rbrakk> \<Longrightarrow> ko_at' v x s"
apply (clarsimp simp: map_comp_Some_iff)
apply (erule (2) aligned_distinct_obj_atI')
apply (simp add: project_inject)
done
lemma map_to_ko_at_updI':
"\<And>x x' y y' y''.
\<lbrakk> (projectKO_opt \<circ>\<^sub>m (ksPSpace s)) x = Some y;
valid_pspace' s; ko_at' y' x' s;
objBitsKO (injectKO y') = objBitsKO y''; x \<noteq> x' \<rbrakk> \<Longrightarrow>
ko_at' y x (s\<lparr>ksPSpace := ksPSpace s(x' \<mapsto> y'')\<rparr>)"
by (fastforce simp: obj_at'_def projectKOs objBitsKO_def ps_clear_upd
dest: map_to_ko_atI2)
(* FIXME: move *)
lemma state_refs_of'_upd:
"\<lbrakk> valid_pspace' s; ko_wp_at' (\<lambda>ko. objBitsKO ko = objBitsKO ko') ptr s \<rbrakk> \<Longrightarrow>
state_refs_of' (s\<lparr>ksPSpace := ksPSpace s(ptr \<mapsto> ko')\<rparr>) =
(state_refs_of' s)(ptr := refs_of' ko')"
apply (rule ext)
apply (clarsimp simp: ps_clear_upd valid_pspace'_def pspace_aligned'_def
obj_at'_def ko_wp_at'_def state_refs_of'_def
split: option.split if_split)
done
lemma sendIPC_enqueue_ccorres_helper:
"ccorres dc xfdc (valid_pspace'
and (\<lambda>s. sym_refs ((state_refs_of' s)(epptr := set queue \<times> {EPSend})))
@ -5531,11 +5232,6 @@ lemma ctcb_relation_blockingIPCCanGrantD:
thread_state_lift_def from_bool_to_bool_iff mask_eq1_nochoice)+
done
(* FIXME move *)
lemma ex_st_tcb_at'_simp[simp]:
"(\<exists>ts. st_tcb_at' ((=) ts) dest s) = tcb_at' dest s"
by (auto simp add: pred_tcb_at'_def obj_at'_def)
lemma sendIPC_ccorres [corres]:
"ccorres dc xfdc (invs' and st_tcb_at' simple' thread
and sch_act_not thread and ep_at' epptr and
@ -5657,11 +5353,7 @@ lemma sendIPC_ccorres [corres]:
apply (rule_tac ep=IdleEP in sendIPC_enqueue_ccorres_helper)
apply (simp add: valid_ep'_def)
apply (wp sts_st_tcb')
apply (rule_tac Q="\<lambda>rv. ko_wp_at' (\<lambda>x. projectKO_opt x = Some IdleEP
\<and> projectKO_opt x = (None::tcb option)) epptr"
in hoare_post_imp)
apply (clarsimp simp: obj_at'_def ko_wp_at'_def projectKOs)
apply wp
apply (clarsimp simp: obj_at'_def ko_wp_at'_def projectKOs)
apply (clarsimp simp: guard_is_UNIV_def)
apply (rule ccorres_return_Skip)
\<comment> \<open>SendEP case\<close>
@ -5687,12 +5379,7 @@ lemma sendIPC_ccorres [corres]:
apply (rule_tac ep="SendEP list" in sendIPC_enqueue_ccorres_helper)
apply (simp add: valid_ep'_def)
apply (wp sts_st_tcb')
apply (rule_tac Q="\<lambda>rv. ko_wp_at' (\<lambda>x. projectKO_opt x = Some (SendEP list)
\<and> projectKO_opt x = (None::tcb option)) epptr
and K (thread \<notin> set list)"
in hoare_post_imp)
apply (clarsimp simp: obj_at'_def ko_wp_at'_def projectKOs)
apply wp
apply (clarsimp simp: obj_at'_def ko_wp_at'_def projectKOs)
apply (clarsimp simp: guard_is_UNIV_def)
apply (rule ccorres_return_Skip)
apply (clarsimp simp: EPState_Recv_def EPState_Send_def EPState_Idle_def
@ -6195,95 +5882,64 @@ lemma receiveIPC_ccorres [corres]:
apply (simp add: cap_endpoint_cap_lift ccap_relation_def cap_to_H_def)
apply ceqv
apply csymbr
apply (rule ccorres_move_c_guard_tcb)
apply (rule_tac xf'=ntfnPtr_'
and r'="\<lambda>rv rv'. rv' = option_to_ptr rv \<and> rv \<noteq> Some 0"
in ccorres_split_nothrow_novcg)
apply (simp add: getBoundNotification_def)
apply (rule_tac P="no_0_obj' and valid_objs'" in threadGet_vcg_corres_P)
apply (rule allI, rule conseqPre, vcg)
apply clarsimp
apply (drule obj_at_ko_at', clarsimp)
apply (drule spec, drule(1) mp, clarsimp)
apply (clarsimp simp: typ_heap_simps ctcb_relation_def)
apply (drule(1) ko_at_valid_objs', simp add: projectKOs)
apply (clarsimp simp: option_to_ptr_def option_to_0_def projectKOs
valid_obj'_def valid_tcb'_def)
apply ceqv
apply (rename_tac ntfnptr ntfnptr')
apply (simp del: Collect_const split del: if_split cong: call_ignore_cong)
apply (rule ccorres_rhs_assoc2)
apply (rule_tac xf'=ret__int_'
and r'="\<lambda>rv rv'. (rv' = 0) = (ntfnptr = None \<or> \<not> isActive rv)"
in ccorres_split_nothrow_novcg)
apply wpc
apply (rule ccorres_from_vcg[where P=\<top> and P'=UNIV])
apply (rule allI, rule conseqPre, vcg)
apply (clarsimp simp: option_to_ptr_def option_to_0_def in_monad Bex_def)
apply (rule ccorres_pre_getNotification[where f=return, simplified])
apply (rule_tac P="\<lambda>s. ko_at' rv (the ntfnptr) s"
in ccorres_from_vcg[where P'=UNIV])
apply (rule ccorres_move_c_guard_tcb)
apply (rule_tac xf'=ntfnPtr_'
and r'="\<lambda>rv rv'. rv' = option_to_ptr rv \<and> rv \<noteq> Some 0"
in ccorres_split_nothrow_novcg)
apply (simp add: getBoundNotification_def)
apply (rule_tac P="no_0_obj' and valid_objs'" in threadGet_vcg_corres_P)
apply (rule allI, rule conseqPre, vcg)
apply clarsimp
apply (drule obj_at_ko_at', clarsimp)
apply (drule spec, drule(1) mp, clarsimp)
apply (clarsimp simp: typ_heap_simps ctcb_relation_def)
apply (drule(1) ko_at_valid_objs', simp add: projectKOs)
apply (clarsimp simp: option_to_ptr_def option_to_0_def projectKOs
valid_obj'_def valid_tcb'_def)
apply ceqv
apply (rename_tac ntfnptr ntfnptr')
apply (simp del: Collect_const split del: if_split cong: call_ignore_cong)
apply (rule ccorres_rhs_assoc2)
apply (rule_tac xf'=ret__int_'
and r'="\<lambda>rv rv'. (rv' = 0) = (ntfnptr = None \<or> \<not> isActive rv)"
in ccorres_split_nothrow_novcg)
apply wpc
apply (rule ccorres_from_vcg[where P=\<top> and P'=UNIV])
apply (rule allI, rule conseqPre, vcg)
apply (clarsimp simp: option_to_ptr_def option_to_0_def in_monad Bex_def)
apply (erule cmap_relationE1[OF cmap_relation_ntfn])
apply (rule ccorres_pre_getNotification[where f=return, simplified])
apply (rule_tac P="\<lambda>s. ko_at' rv (the ntfnptr) s"
in ccorres_from_vcg[where P'=UNIV])
apply (rule allI, rule conseqPre, vcg)
apply (clarsimp simp: option_to_ptr_def option_to_0_def in_monad Bex_def)
apply (erule cmap_relationE1[OF cmap_relation_ntfn])
apply (erule ko_at_projectKO_opt)
apply (clarsimp simp: typ_heap_simps cnotification_relation_def Let_def
notification_state_defs isActive_def
split: Structures_H.ntfn.split_asm Structures_H.notification.splits)
apply ceqv
apply (rule ccorres_cond[where R=\<top>])
apply (simp add: Collect_const_mem)
apply (ctac add: completeSignal_ccorres[unfolded dc_def])
apply (rule_tac xf'=ret__unsigned_'
and val="case ep of IdleEP \<Rightarrow> scast EPState_Idle
| RecvEP _ \<Rightarrow> scast EPState_Recv
| SendEP _ \<Rightarrow> scast EPState_Send"
and R="ko_at' ep (capEPPtr cap)"
in ccorres_symb_exec_r_known_rv_UNIV[where R'=UNIV])
apply (vcg, clarsimp)
apply (erule cmap_relationE1 [OF cmap_relation_ep])
apply (erule ko_at_projectKO_opt)
apply (clarsimp simp: typ_heap_simps cnotification_relation_def Let_def
notification_state_defs isActive_def
split: Structures_H.ntfn.split_asm Structures_H.notification.splits)
apply (clarsimp simp: typ_heap_simps cendpoint_relation_def Let_def
split: endpoint.split_asm)
apply ceqv
apply (rule ccorres_cond[where R=\<top>])
apply (simp add: Collect_const_mem)
apply (ctac add: completeSignal_ccorres[unfolded dc_def])
apply (rule_tac xf'=ret__unsigned_'
and val="case ep of IdleEP \<Rightarrow> scast EPState_Idle
| RecvEP _ \<Rightarrow> scast EPState_Recv
| SendEP _ \<Rightarrow> scast EPState_Send"
and R="ko_at' ep (capEPPtr cap)"
in ccorres_symb_exec_r_known_rv_UNIV[where R'=UNIV])
apply (vcg, clarsimp)
apply (erule cmap_relationE1 [OF cmap_relation_ep])
apply (erule ko_at_projectKO_opt)
apply (clarsimp simp: typ_heap_simps cendpoint_relation_def Let_def
split: endpoint.split_asm)
apply ceqv
apply (rule_tac A="invs' and st_tcb_at' simple' thread
and sch_act_not thread
and (\<lambda>s. \<forall>d p. thread \<notin> set (ksReadyQueues s (d, p)))
and ko_at' ep (capEPPtr cap)"
in ccorres_guard_imp2 [where A'=UNIV])
apply (rule_tac A="invs' and st_tcb_at' simple' thread
and sch_act_not thread
and (\<lambda>s. \<forall>d p. thread \<notin> set (ksReadyQueues s (d, p)))
and ko_at' ep (capEPPtr cap)"
in ccorres_guard_imp2 [where A'=UNIV])
apply wpc
\<comment> \<open>RecvEP case\<close>
apply (rule ccorres_cond_true)
apply csymbr
apply (simp only: case_bool_If from_bool_neq_0)
apply (rule ccorres_Cond_rhs, simp cong: Collect_cong split del: if_split)
apply (intro ccorres_rhs_assoc)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_split_nothrow_novcg)
apply (simp split del: if_split)
apply (rule receiveIPC_block_ccorres_helper[unfolded ptr_val_def, simplified])
apply ceqv
apply simp
apply (rename_tac list NOo)
apply (rule_tac ep="RecvEP list"
in receiveIPC_enqueue_ccorres_helper[simplified, unfolded dc_def])
apply (simp add: valid_ep'_def)
apply (wp sts_st_tcb')
apply (rename_tac list)
apply (rule_tac Q="\<lambda>rv. ko_wp_at' (\<lambda>x. projectKO_opt x = Some (RecvEP list)
\<and> projectKO_opt x = (None::tcb option))
(capEPPtr cap)
and K (thread \<notin> set list)"
in hoare_post_imp)
apply (clarsimp simp: obj_at'_def ko_wp_at'_def projectKOs)
apply wp
apply (clarsimp simp: guard_is_UNIV_def)
apply simp
apply (ctac add: doNBRecvFailedTransfer_ccorres[unfolded dc_def])
\<comment> \<open>IdleEP case\<close>
\<comment> \<open>RecvEP case\<close>
apply (rule ccorres_cond_true)
apply csymbr
apply (simp only: case_bool_If from_bool_neq_0)
@ -6298,47 +5954,81 @@ lemma receiveIPC_ccorres [corres]:
apply (rule receiveIPC_block_ccorres_helper[unfolded ptr_val_def, simplified])
apply ceqv
apply simp
apply (rule_tac ep=IdleEP
in receiveIPC_enqueue_ccorres_helper[simplified, unfolded dc_def])
apply (rename_tac list NOo)
apply (rule_tac ep="RecvEP list"
in receiveIPC_enqueue_ccorres_helper[simplified, unfolded dc_def])
apply (simp add: valid_ep'_def)
apply (wp sts_st_tcb')
apply (rule_tac Q="\<lambda>rv. ko_wp_at' (\<lambda>x. projectKO_opt x = Some IdleEP
\<and> projectKO_opt x = (None::tcb option))
(capEPPtr cap)"
in hoare_post_imp)
apply (clarsimp simp: obj_at'_def ko_wp_at'_def projectKOs)
apply wp
apply (rename_tac list)
apply (clarsimp simp: obj_at'_def ko_wp_at'_def projectKOs)
apply (clarsimp simp: guard_is_UNIV_def)
apply simp
apply (ctac add: doNBRecvFailedTransfer_ccorres[unfolded dc_def])
\<comment> \<open>IdleEP case\<close>
apply (rule ccorres_cond_true)
apply csymbr
apply (simp only: case_bool_If from_bool_neq_0)
apply (rule ccorres_Cond_rhs, simp cong: Collect_cong split del: if_split)
apply (intro ccorres_rhs_assoc)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_split_nothrow_novcg)
apply (simp split del: if_split)
apply (rule receiveIPC_block_ccorres_helper[unfolded ptr_val_def, simplified])
apply ceqv
apply simp
apply (rule_tac ep=IdleEP
in receiveIPC_enqueue_ccorres_helper[simplified, unfolded dc_def])
apply (simp add: valid_ep'_def)
apply (wp sts_st_tcb')
apply (clarsimp simp: obj_at'_def ko_wp_at'_def projectKOs)
apply (clarsimp simp: guard_is_UNIV_def)
apply simp
apply (ctac add: doNBRecvFailedTransfer_ccorres[unfolded dc_def])
\<comment> \<open>SendEP case\<close>
apply (thin_tac "isBlockinga = from_bool P" for P)
apply (rule ccorres_cond_false)
apply (rule ccorres_cond_true)
apply (intro ccorres_rhs_assoc)
apply (csymbr, csymbr, csymbr, csymbr, csymbr)
apply wpc
apply (simp only: haskell_fail_def)
apply (rule ccorres_fail)
apply (rename_tac sender rest)
apply csymbr
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_split_nothrow_novcg)
apply simp
apply (rule_tac sender=sender in receiveIPC_dequeue_ccorres_helper[simplified])
apply (thin_tac "isBlockinga = from_bool P" for P)
apply (rule ccorres_cond_false)
apply (rule ccorres_cond_true)
apply (intro ccorres_rhs_assoc)
apply (csymbr, csymbr, csymbr, csymbr, csymbr)
apply wpc
apply (simp only: haskell_fail_def)
apply (rule ccorres_fail)
apply (rename_tac sender rest)
apply csymbr
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_split_nothrow_novcg)
apply simp
apply (rule_tac sender=sender in receiveIPC_dequeue_ccorres_helper[simplified])
apply ceqv
apply (rename_tac sender')
apply (simp only: K_bind_def haskell_assert_def return_bind)
apply (rule ccorres_move_c_guard_tcb)
apply (rule getThreadState_ccorres_foo)
apply (rename_tac sendState)
apply (rule ccorres_assert)
apply (rule ccorres_rhs_assoc2)
apply (rule_tac val="blockingIPCBadge sendState"
and xf'=badge_'
and R="\<lambda>s. \<exists>t. ko_at' t sender s \<and> tcbState t = sendState"
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 (clarsimp simp: ctcb_relation_def typ_heap_simps
cthread_state_relation_def word_size
isSend_def thread_state_lift_def
split: Structures_H.thread_state.splits)
apply ceqv
apply (rename_tac sender')
apply (simp only: K_bind_def haskell_assert_def return_bind)
apply (simp split del: if_split)
apply (rule ccorres_move_c_guard_tcb)
apply (rule getThreadState_ccorres_foo)
apply (rename_tac sendState)
apply (rule ccorres_assert)
apply (rule ccorres_rhs_assoc2)
apply (rule_tac val="blockingIPCBadge sendState"
and xf'=badge_'
apply (rule_tac val="from_bool (blockingIPCCanGrant sendState)"
and xf'=canGrant_'
and R="\<lambda>s. \<exists>t. ko_at' t sender s \<and> tcbState t = sendState"
in ccorres_symb_exec_r_known_rv_UNIV [where R'=UNIV])
apply (vcg, clarsimp)
@ -6348,20 +6038,6 @@ lemma receiveIPC_ccorres [corres]:
isSend_def thread_state_lift_def
split: Structures_H.thread_state.splits)
apply ceqv
apply (simp split del: if_split)
apply (rule ccorres_move_c_guard_tcb)
apply (rule ccorres_rhs_assoc2)
apply (rule_tac val="from_bool (blockingIPCCanGrant sendState)"
and xf'=canGrant_'
and R="\<lambda>s. \<exists>t. ko_at' t sender s \<and> tcbState t = sendState"
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 (clarsimp simp: ctcb_relation_def typ_heap_simps
cthread_state_relation_def word_size
isSend_def thread_state_lift_def
split: Structures_H.thread_state.splits)
apply ceqv
apply (rule ccorres_rhs_assoc2)
apply (rule_tac xf'=canGrantReply_'

View File

@ -115,71 +115,8 @@ lemma isolate_thread_actions_bind:
lemmas setEndpoint_obj_at_tcb' = setEndpoint_obj_at'_tcb
lemma tcbSchedEnqueue_obj_at_unchangedT:
assumes y: "\<And>f. \<forall>tcb. P (tcbQueued_update f tcb) = P tcb"
shows "\<lbrace>obj_at' P t\<rbrace> tcbSchedEnqueue t' \<lbrace>\<lambda>rv. obj_at' P t\<rbrace>"
apply (simp add: tcbSchedEnqueue_def unless_def)
apply (wp | simp add: y)+
done
lemma asUser_obj_at_notQ:
"\<lbrace>obj_at' (Not \<circ> tcbQueued) t\<rbrace>
asUser t (setRegister r v)
\<lbrace>\<lambda>rv. obj_at' (Not \<circ> tcbQueued) t\<rbrace>"
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:
"\<lbrace>obj_at' (Not \<circ> tcbQueued) t\<rbrace>
Arch.switchToThread t
\<lbrace>\<lambda>rv. obj_at' (Not \<circ> tcbQueued) t\<rbrace>"
apply (simp add: ARM_HYP_H.switchToThread_def)
apply (wp asUser_obj_at_notQ doMachineOp_obj_at setVMRoot_obj_at'_no_vcpu hoare_drop_imps|wpc)+
done
lemma rescheduleRequired_obj_at_unchangedT:
assumes y: "\<And>f. \<forall>tcb. P (tcbQueued_update f tcb) = P tcb"
shows "\<lbrace>obj_at' P t\<rbrace> rescheduleRequired \<lbrace>\<lambda>rv. obj_at' P t\<rbrace>"
apply (simp add: rescheduleRequired_def)
apply (wp tcbSchedEnqueue_obj_at_unchangedT[OF y] | wpc)+
apply simp
done
lemma setThreadState_obj_at_unchangedT:
assumes x: "\<And>f. \<forall>tcb. P (tcbState_update f tcb) = P tcb"
assumes y: "\<And>f. \<forall>tcb. P (tcbQueued_update f tcb) = P tcb"
shows "\<lbrace>obj_at' P t\<rbrace> setThreadState t' ts \<lbrace>\<lambda>rv. obj_at' P t\<rbrace>"
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: "\<And>f. \<forall>tcb. P (tcbBoundNotification_update f tcb) = P tcb"
shows "\<lbrace>obj_at' P t\<rbrace> setBoundNotification t' ts \<lbrace>\<lambda>rv. obj_at' P t\<rbrace>"
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]
lemmas setNotification_tcb = set_ntfn_tcb_obj_at'
(* FIXME: move *)
lemmas threadSet_obj_at' = threadSet_obj_at'_strongish
context kernel_m begin
context begin interpretation Arch . (*FIXME: arch_split*)
@ -692,7 +629,7 @@ lemma setVCPU_isolatable:
apply (clarsimp simp: thread_actions_isolatable_def monadic_rewrite_def isolate_thread_actions_def)
apply (clarsimp simp: exec_gets getSchedulerAction_def)
apply (subst setObject_assert_modify;
simp add: projectKOs objBits_simps archObjSize_def vcpuBits_def pageBits_def)+
simp add: projectKOs objBits_simps archObjSize_def vcpuBits_def vcpu_bits_def pageBits_def)+
apply (clarsimp simp: select_f_asserts assert_def obj_at_partial_overwrite_id2 split: if_splits)
apply (clarsimp simp: select_f_def simpler_modify_def bind_def o_def)
apply (case_tac s)
@ -1003,12 +940,6 @@ lemma copyMRs_simple:
apply (simp add: upto_enum_def mapM_Nil)
done
(* FIXME: MOVE *)
lemma returnOK_catch[simp]:
"(returnOk rv <catch> m) = return rv"
unfolding catch_def returnOk_def
by clarsimp
lemma doIPCTransfer_simple_rewrite:
"monadic_rewrite True True
((\<lambda>_. msgExtraCaps (messageInfoFromWord msgInfo) = 0

View File

@ -540,16 +540,6 @@ lemma invalidateTLBByASID_ccorres:
apply (clarsimp simp: pde_stored_asid_def to_bool_def)
done
(* FIXME: also in VSpace_C *)
lemma ignoreFailure_liftM:
"ignoreFailure = liftM (\<lambda>v. ())"
apply (rule ext)+
apply (simp add: ignoreFailure_def liftM_def
catch_def)
apply (rule bind_apply_cong[OF refl])
apply (simp split: sum.split)
done
end
lemma option_to_0_user_mem':
@ -669,7 +659,7 @@ lemma clearMemory_setObject_PTE_ccorres:
subgoal
apply (subst unat_of_nat32, simp add: table_bits_defs word_bits_def)
apply (subst unat_power_lower32', simp add: table_bits_defs word_bits_def)
apply (erule (1) page_table_at_rf_sr_dom_s)
apply (erule (1) page_table_at_rf_sr_dom_s[simplified])
done
apply (clarsimp simp add: table_bits_defs
cong: StateSpace.state.fold_congs globals.fold_congs)
@ -834,17 +824,6 @@ lemma cpspace_relation_ep_update_ep2:
end
context begin interpretation Arch . (*FIXME: arch_split*)
lemma setObject_tcb_ep_obj_at'[wp]:
"\<lbrace>obj_at' (P :: endpoint \<Rightarrow> bool) ptr\<rbrace> setObject ptr' (tcb :: tcb) \<lbrace>\<lambda>rv. obj_at' P ptr\<rbrace>"
apply (rule obj_at_setObject2, simp_all)
apply (clarsimp simp: updateObject_default_def in_monad)
done
end
crunch ep_obj_at'[wp]: setThreadState "obj_at' (P :: endpoint \<Rightarrow> bool) ptr"
(simp: unless_def)
context kernel_m begin
lemma ccorres_abstract_h_val:

View File

@ -27,28 +27,6 @@ apply (rule hoare_post_imp)
apply (rule schedule_invs')
done
(* FIXME: This is cheating since ucast from 10 to 16 will never give us 0xFFFF.
However type of 10 word is from irq oracle so it is the oracle that matters not this lemma.
(Xin) *)
lemma ucast_not_helper_cheating:
fixes a:: "10 word"
assumes a: "ucast a \<noteq> (0xFFFF :: word16)"
shows "ucast a \<noteq> (0xFFFF::32 signed word)"
by (word_bitwise,simp)
lemma ucast_helper_not_maxword:
"UCAST(10 \<rightarrow> 32) x \<noteq> 0xFFFF"
apply (subgoal_tac "UCAST(10 \<rightarrow> 32) x \<le> UCAST(10 \<rightarrow> 32) max_word")
apply (rule notI)
defer
apply (rule ucast_up_mono_le)
apply simp
apply simp
by (simp add: max_word_def)
lemmas ucast_helper_simps_32 =
ucast_helper_not_maxword arg_cong[where f="UCAST(16 \<rightarrow> 32)", OF minus_one_norm]
lemma Arch_finaliseInterrupt_ccorres:
"ccorres dc xfdc \<top> UNIV [] (return a) (Call Arch_finaliseInterrupt_'proc)"
apply (cinit')
@ -650,12 +628,6 @@ lemma ccorres_get_registers:
"StrictC'_register_defs")
done
(* FIXME: move *)
lemma st_tcb_at'_opeq_simp:
"st_tcb_at' ((=) Structures_H.thread_state.Running) (ksCurThread s) s
= st_tcb_at' (\<lambda>st. st = Structures_H.thread_state.Running) (ksCurThread s) s"
by (fastforce simp add: st_tcb_at'_def obj_at'_def)
lemma callKernel_withFastpath_corres_C:
"corres_underlying rf_sr False True dc
@ -820,26 +792,6 @@ lemma full_invs_both:
done
end
(* FIXME: move to somewhere sensible *)
lemma dom_eq:
"dom um = dom um' \<longleftrightarrow> (\<forall>a. um a = None \<longleftrightarrow> um' a = None)"
apply (simp add: dom_def del: not_None_eq)
apply (rule iffI)
apply (rule allI)
apply (simp add: set_eq_iff)
apply (drule_tac x=a in spec)
apply auto
done
lemma dom_user_mem':
"dom (user_mem' s) = {p. typ_at' UserDataT (p && ~~ mask pageBits) s}"
by (clarsimp simp:user_mem'_def dom_def pointerInUserData_def split:if_splits)
(* FIXME:move *)
lemma dom_device_mem':
"dom (device_mem' s) = {p. typ_at' UserDataDeviceT (p && ~~ mask pageBits) s}"
by (clarsimp simp: device_mem'_def dom_def pointerInDeviceData_def split: if_splits)
context kernel_m
begin

View File

@ -2042,13 +2042,12 @@ proof (intro impI allI)
have szo: "size_of TYPE(pte_C[512]) = 2 ^ ptBits"
by (simp add: size_of_def size_td_array table_bits_defs)
have szo2: "512 * size_of TYPE(pte_C) = 2 ^ ptBits"
by (simp add: szo[symmetric])
by (simp add: szo[symmetric] pt_bits_def pte_bits_def)
have szo': "size_of TYPE(pte_C) = 2 ^ objBitsKO ko"
by (simp add: objBits_simps ko_def archObjSize_def table_bits_defs)
note rl' = cslift_ptr_retyp_other_inst[where n=1,
simplified, OF empty cover[simplified] szo[symmetric] szo]
simplified, OF empty cover[simplified One_nat_def] szo[symmetric] szo]
have sz_weaken: "objBitsKO ko \<le> ptBits"
by (simp add: objBits_simps ko_def archObjSize_def table_bits_defs)
have cover': "range_cover ptr sz (objBitsKO ko) 512"
@ -2095,11 +2094,10 @@ proof (intro impI allI)
apply (rule projectKO_opt_retyp_same, simp add: ko_def projectKOs)
apply (simp add: h_t_valid_clift_Some_iff dom_def split: if_split)
apply (subst clift_ptr_retyps_gen_prev_memset_same[where n=1, simplified, OF guard],
simp_all only: szo refl empty, simp_all add: zero)[1]
simp_all only: szo refl empty, simp_all add: zero[simplified])[1]
apply (simp add: table_bits_defs word_bits_def)
apply (auto split: if_split)[1]
apply (simp_all add: objBits_simps archObjSize_def table_bits_defs
ko_def word_bits_def)
apply (simp_all add: objBits_simps archObjSize_def table_bits_defs ko_def word_bits_def)
done
from rf have "cpspace_relation (ksPSpace \<sigma>) (underlying_memory (ksMachineState \<sigma>)) (t_hrs_' (globals x))"
@ -2115,7 +2113,7 @@ proof (intro impI allI)
apply (subst clift_ptr_retyps_gen_prev_memset_same[OF guard'], simp_all only: szo2 empty)
apply simp
apply (simp(no_asm) add: table_bits_defs word_bits_def)
apply (simp add: zero)
apply (simp add: zero[simplified])
apply (simp add: rl projectKOs del: pte_C_size)
apply (simp add: rl projectKO_opt_retyp_same ko_def projectKOs Let_def
ptr_add_to_new_cap_addrs [OF szo']
@ -2130,7 +2128,7 @@ proof (intro impI allI)
apply (clarsimp simp: valid_global_refs'_def Let_def
valid_refs'_def ran_def rf_sr_def cstate_relation_def)
apply (erule disjoint_subset)
apply (simp add:kernel_data_refs_disj)
apply (simp add: kernel_data_refs_disj[simplified])
done
ultimately
@ -2238,18 +2236,18 @@ proof (intro impI allI)
have szo: "size_of TYPE(pde_C[2048]) = 2 ^ pdBits"
by (simp add: size_of_def size_td_array table_bits_defs)
have szo2: "2048 * size_of TYPE(pde_C) = 2 ^ pdBits"
by (simp add: szo[symmetric])
by (simp add: szo[symmetric] pdBits_def pd_bits_def pde_bits_def pdeBits_def)
have szo': "size_of TYPE(pde_C) = 2 ^ objBitsKO ko"
by (simp add: objBits_simps ko_def archObjSize_def table_bits_defs)
note rl' = cslift_ptr_retyp_other_inst[where n=1,
simplified, OF empty cover[simplified] szo[symmetric] szo]
simplified, OF empty cover[simplified One_nat_def] szo[symmetric] szo]
have sz_weaken: "objBitsKO ko \<le> pdBits"
by (simp add: objBits_simps ko_def archObjSize_def pdBits_def pageBits_def)
by (simp add: objBits_simps ko_def archObjSize_def pdBits_def pageBits_def pd_bits_def)
have cover': "range_cover ptr sz (objBitsKO ko) 2048"
apply (rule range_cover_rel[OF cover sz_weaken])
apply (simp add: pdBits_def objBits_simps ko_def archObjSize_def pageBits_def)
apply (simp add: pdBits_def objBits_simps ko_def archObjSize_def pageBits_def pd_bits_def)
done
from sz sz_weaken have sz': "objBitsKO ko \<le> sz" by simp
note al' = is_aligned_weaken[OF al sz_weaken]
@ -2264,7 +2262,7 @@ proof (intro impI allI)
have guard: "c_guard ?ptr"
apply (rule is_aligned_c_guard[where n=pdBits and m=2])
apply (simp_all add: al ptr0 align_of_def align_td_array)
apply (simp_all add: al[simplified] ptr0 align_of_def align_td_array)
apply (simp_all add: table_bits_defs)
done
@ -2274,7 +2272,7 @@ proof (intro impI allI)
done
note rl' = cslift_ptr_retyp_other_inst[OF _ cover refl szo,
simplified szo, simplified, OF empty]
simplified szo, simplified, OF empty[simplified]]
from rf have pderl: "cmap_relation (map_to_pdes (ksPSpace \<sigma>)) (cslift x) Ptr cpde_relation"
unfolding rf_sr_def cstate_relation_def by (simp add: Let_def cpspace_relation_def)
@ -2289,7 +2287,7 @@ proof (intro impI allI)
apply (rule projectKO_opt_retyp_same, simp add: ko_def projectKOs)
apply (simp add: h_t_valid_clift_Some_iff dom_def split: if_split)
apply (subst clift_ptr_retyps_gen_prev_memset_same[where n=1, simplified, OF guard],
simp_all only: szo empty, simp_all add: zero)[1]
simp_all only: szo empty, simp_all add: zero[simplified])[1]
apply (simp add: table_bits_defs word_bits_def)
apply (auto split: if_split)[1]
apply (simp_all add: objBits_simps archObjSize_def table_bits_defs
@ -2309,7 +2307,7 @@ proof (intro impI allI)
apply (subst clift_ptr_retyps_gen_prev_memset_same[OF guard'], simp_all only: szo2 empty)
apply simp
apply (simp(no_asm) add: table_bits_defs word_bits_def)
apply (simp add: zero)
apply (simp add: zero[simplified])
apply (simp add: rl projectKOs)
apply (simp add: rl projectKO_opt_retyp_same ko_def projectKOs Let_def
ptr_add_to_new_cap_addrs [OF szo']
@ -2324,7 +2322,7 @@ proof (intro impI allI)
apply (clarsimp simp: valid_global_refs'_def Let_def
valid_refs'_def ran_def rf_sr_def cstate_relation_def)
apply (erule disjoint_subset)
apply (simp add:kernel_data_refs_disj)
apply (simp add: kernel_data_refs_disj[simplified])
done
moreover from rf have stored_asids: "(pde_stored_asid \<circ>\<^sub>m clift ?ks')
@ -2337,7 +2335,7 @@ proof (intro impI allI)
apply (subst clift_ptr_retyps_gen_prev_memset_same[OF guard'], simp_all only: szo2 empty)
apply simp
apply (simp add: table_bits_defs word_bits_def)
apply (simp add: zero)
apply (simp add: zero[simplified])
apply (rule ext)
apply (simp add: map_comp_def stored_asid[simplified] split: option.split if_split)
apply (simp only: o_def CTypesDefs.ptr_add_def' Abs_fnat_hom_mult)
@ -3097,13 +3095,6 @@ lemma tcb_queue_update_other':
unfolding tcb_queue_relation'_def
by (simp add: tcb_queue_update_other)
lemma map_to_ko_atI2:
"\<lbrakk>(projectKO_opt \<circ>\<^sub>m (ksPSpace s)) x = Some v; pspace_aligned' s; pspace_distinct' s\<rbrakk> \<Longrightarrow> ko_at' v x s"
apply (clarsimp simp: map_comp_Some_iff)
apply (erule (2) aligned_distinct_obj_atI')
apply (simp add: project_inject)
done
lemma c_guard_tcb:
assumes al: "is_aligned (ctcb_ptr_to_tcb_ptr p) tcbBlockSizeBits"
and ptr0: "ctcb_ptr_to_tcb_ptr p \<noteq> 0"
@ -5970,7 +5961,7 @@ proof -
apply (ccorres_remove_UNIV_guard)
apply (rule ccorres_rhs_assoc)+
apply (clarsimp simp: hrs_htd_update ptBits_def objBits_simps archObjSize_def
ARM_HYP_H.createObject_def pageBits_def)
ARM_HYP_H.createObject_def pageBits_def pt_bits_def)
apply (ctac pre only: add: placeNewObject_pte[simplified])
apply csymbr
apply (rule ccorres_return_C)
@ -6004,7 +5995,7 @@ proof -
apply (ccorres_remove_UNIV_guard)
apply (rule ccorres_rhs_assoc)+
apply (clarsimp simp: hrs_htd_update ptBits_def objBits_simps archObjSize_def
ARM_HYP_H.createObject_def pageBits_def pdBits_def)
ARM_HYP_H.createObject_def pageBits_def pdBits_def pd_bits_def)
apply (ctac pre only: add: placeNewObject_pde[simplified])
apply (ctac add: copyGlobalMappings_ccorres)
apply csymbr
@ -6647,6 +6638,7 @@ lemma pspace_no_overlap_induce_cte:
is_aligned ptr bits; bits < word_bits;
pspace_no_overlap' ptr bits s\<rbrakk>
\<Longrightarrow> {ptr_val xa..+size_of TYPE(cte_C)} \<inter> {ptr..+2 ^ bits} = {}"
supply cteSizeBits_le_cte_level_bits[simp del]
apply (clarsimp simp: cpspace_relation_def)
apply (clarsimp simp: cmap_relation_def size_of_def)
apply (subgoal_tac "xa\<in>cte_Ptr ` dom (ctes_of s)")
@ -7909,29 +7901,6 @@ end
context kernel_m begin
lemma nat_le_induct [consumes 1, case_names base step]:
assumes le: "i \<le> (k::nat)" and
base: "P(k)" and
step: "\<And>i. \<lbrakk>i \<le> k; P i; 0 < i\<rbrakk> \<Longrightarrow> P(i - 1)"
shows "P i"
proof -
obtain j where jk: "j \<le> k" and j_eq: "i = k - j"
using le
apply (drule_tac x="k - i" in meta_spec)
apply simp
done
have "j \<le> k \<Longrightarrow> P (k - j)"
apply (induct j)
apply (simp add: base)
apply simp
apply (drule step[rotated], simp+)
done
thus "P i" using jk j_eq
by simp
qed
lemma ceqv_restore_as_guard:
"ceqv Gamma xf' rv' t t' d (Guard C_Guard {s. xf' s = rv'} d)"
apply (simp add: ceqv_def)

View File

@ -1561,15 +1561,6 @@ lemma ctcb_relation_null_queue_ptrs:
apply (simp add: ctcb_relation_def tcb_null_queue_ptrs_def)
done
lemma map_to_ko_atI':
"\<lbrakk>(projectKO_opt \<circ>\<^sub>m (ksPSpace s)) x = Some v; invs' s\<rbrakk> \<Longrightarrow> ko_at' v x s"
apply (clarsimp simp: map_comp_Some_iff)
apply (erule aligned_distinct_obj_atI')
apply clarsimp
apply clarsimp
apply (simp add: project_inject)
done
(* Levity: added (20090419 09:44:27) *)
declare ntfnQueue_head_mask_4 [simp]
@ -1934,11 +1925,6 @@ lemma rf_sr_heap_device_data_relation:
lemma ko_at_projectKO_opt:
"ko_at' ko p s \<Longrightarrow> (projectKO_opt \<circ>\<^sub>m ksPSpace s) p = Some ko"
by (clarsimp elim!: obj_atE' simp: projectKOs)
lemma user_word_at_cross_over:
"\<lbrakk> user_word_at x p s; (s, s') \<in> rf_sr; p' = Ptr p \<rbrakk>
\<Longrightarrow> c_guard p' \<and> hrs_htd (t_hrs_' (globals s')) \<Turnstile>\<^sub>t p'
@ -2258,7 +2244,7 @@ lemma page_directory_at_carray_map_relation:
apply (drule_tac x="p' && mask pdBits >> 3" in spec)
apply (clarsimp simp: shiftr_shiftl1)
apply (drule mp)
apply (simp add: shiftr_over_and_dist mask_def pdBits_def'
apply (simp add: shiftr_over_and_dist mask_def pdBits_def' pd_bits_def pde_bits_def
order_le_less_trans[OF word_and_le1])
apply (clarsimp simp: typ_at_to_obj_at_arches objBits_simps archObjSize_def
table_bits_defs

View File

@ -35,16 +35,6 @@ end
context kernel_m begin
(* FIXME: move to Refine *)
lemma valid_idle'_tcb_at'_ksIdleThread:
"valid_idle' s \<Longrightarrow> tcb_at' (ksIdleThread s) s"
by (clarsimp simp: valid_idle'_def pred_tcb_at'_def obj_at'_def)
(* FIXME: move to Refine *)
lemma invs_no_cicd'_valid_idle':
"invs_no_cicd' s \<Longrightarrow> valid_idle' s"
by (simp add: invs_no_cicd'_def)
lemma Arch_switchToIdleThread_ccorres:
"ccorres dc xfdc invs_no_cicd' UNIV []
Arch.switchToIdleThread (Call Arch_switchToIdleThread_'proc)"
@ -56,11 +46,6 @@ lemma Arch_switchToIdleThread_ccorres:
apply (clarsimp simp: invs_no_cicd'_def valid_pspace'_def valid_idle'_tcb_at'_ksIdleThread)
done
(* FIXME: move *)
lemma empty_fail_getIdleThread [simp,intro!]:
"empty_fail getIdleThread"
by (simp add: getIdleThread_def)
lemma switchToIdleThread_ccorres:
"ccorres dc xfdc invs_no_cicd' UNIV hs
switchToIdleThread (Call switchToIdleThread_'proc)"

View File

@ -500,34 +500,6 @@ lemma invocationCatch_use_injection_handler:
split: sum.split)
done
lemma injection_handler_returnOk:
"injection_handler injector (returnOk v) = returnOk v"
by (simp add: returnOk_liftE injection_liftE)
lemma injection_handler_If:
"injection_handler injector (If P a b)
= If P (injection_handler injector a)
(injection_handler injector b)"
by (simp split: if_split)
(* FIXME: duplicated in CSpace_All *)
lemma injection_handler_liftM:
"injection_handler f
= liftM (\<lambda>v. case v of Inl ex \<Rightarrow> Inl (f ex) | Inr rv \<Rightarrow> Inr rv)"
apply (intro ext, simp add: injection_handler_def liftM_def
handleE'_def)
apply (rule bind_apply_cong, rule refl)
apply (simp add: throwError_def split: sum.split)
done
lemma injection_handler_throwError:
"injection_handler f (throwError v) = throwError (f v)"
by (simp add: injection_handler_def handleE'_def
throwError_bind)
lemmas injection_handler_bindE = injection_bindE [OF refl refl]
lemmas injection_handler_wp = injection_wp [OF refl]
lemma ccorres_injection_handler_csum1:
"ccorres (f \<currency> r) xf P P' hs a c
\<Longrightarrow> ccorres
@ -779,9 +751,6 @@ lemma capFVMRights_range:
by (simp add: cap_frame_cap_lift_def cap_small_frame_cap_lift_def
cap_lift_def cap_tag_defs word_and_le1 mask_def)+
lemma dumb_bool_for_all: "(\<forall>x. x) = False"
by auto
lemma dumb_bool_split_for_vcg:
"\<lbrace>d \<longrightarrow> \<acute>ret__unsigned_long \<noteq> 0\<rbrace> \<inter> \<lbrace>\<not> d \<longrightarrow> \<acute>ret__unsigned_long = 0\<rbrace>
= \<lbrace>d = to_bool \<acute>ret__unsigned_long \<rbrace>"

View File

@ -787,10 +787,6 @@ lemma getMessageInfo_less_4:
apply (rule word_and_le1)
done
lemma invs_queues_imp:
"invs' s \<longrightarrow> valid_queues s"
by clarsimp
(* FIXME: move *)
lemma length_CL_from_H [simp]:
"length_CL (mi_from_H mi) = msgLength mi"
@ -1208,32 +1204,6 @@ lemma deleteCallerCap_ccorres [corres]:
done
(* FIXME: MOVE *)
lemma cap_case_EndpointCap_NotificationCap:
"(case cap of EndpointCap v0 v1 v2 v3 v4 v5 \<Rightarrow> f v0 v1 v2 v3 v4 v5
| NotificationCap v0 v1 v2 v3 \<Rightarrow> g v0 v1 v2 v3
| _ \<Rightarrow> h)
= (if isEndpointCap cap
then f (capEPPtr cap) (capEPBadge cap) (capEPCanSend cap) (capEPCanReceive cap)
(capEPCanGrant cap) (capEPCanGrantReply cap)
else if isNotificationCap cap
then g (capNtfnPtr cap) (capNtfnBadge cap) (capNtfnCanSend cap) (capNtfnCanReceive cap)
else h)"
by (simp add: isCap_simps
split: capability.split)
(* FIXME: MOVE *)
lemma capFaultOnFailure_if_case_sum:
" (capFaultOnFailure epCPtr b (if c then f else g) >>=
sum.case_sum (handleFault thread) return) =
(if c then ((capFaultOnFailure epCPtr b f)
>>= sum.case_sum (handleFault thread) return)
else ((capFaultOnFailure epCPtr b g)
>>= sum.case_sum (handleFault thread) return))"
by (case_tac c, clarsimp, clarsimp)
(* FIXME: MOVE to Corres_C.thy *)
lemma ccorres_trim_redundant_throw_break:
@ -1697,17 +1667,6 @@ lemma get_gic_vcpu_ctrl_eisr0_invs'[wp]:
crunch obj_at'_s[wp]: armv_contextSwitch "\<lambda>s. obj_at' P (ksCurThread s) s"
lemma dmo_machine_valid_lift:
"(\<And>s f. P s = P (ksMachineState_update f s)) \<Longrightarrow> \<lbrace>P\<rbrace> doMachineOp f' \<lbrace>\<lambda>rv. P\<rbrace>"
apply (wpsimp simp: split_def doMachineOp_def machine_op_lift_def machine_rest_lift_def in_monad)
done
lemma tcb_runnable_imp_simple:
"obj_at' (\<lambda>s. runnable' (tcbState s)) t s \<Longrightarrow> obj_at' (\<lambda>s. simple' (tcbState s)) t s"
apply (clarsimp simp: obj_at'_def)
apply (case_tac "tcbState obj" ; clarsimp)
done
lemma ccorres_return_void_C_Seq:
"ccorres_underlying sr \<Gamma> r rvxf arrel xf P P' hs X (return_void_C) \<Longrightarrow>
ccorres_underlying sr \<Gamma> r rvxf arrel xf P P' hs X (return_void_C ;; Z)"

View File

@ -37,27 +37,6 @@ lemma ccorres_pre_threadGet:
done
(* FIXME MOVE *)
crunch inv'[wp]: archThreadGet P
(* FIXME MOVE near thm tg_sp' *)
lemma atg_sp':
"\<lbrace>P\<rbrace> archThreadGet f p \<lbrace>\<lambda>t. obj_at' (\<lambda>t'. f (tcbArch t') = t) p and P\<rbrace>"
including no_pre
apply (simp add: archThreadGet_def)
apply wp
apply (rule hoare_strengthen_post)
apply (rule getObject_tcb_sp)
apply clarsimp
apply (erule obj_at'_weakenE)
apply simp
done
(* FIXME: MOVE to EmptyFail *)
lemma empty_fail_archThreadGet [intro!, wp, simp]:
"empty_fail (archThreadGet f p)"
by (simp add: archThreadGet_def getObject_def split_def)
lemma ccorres_pre_archThreadGet:
assumes cc: "\<And>rv. ccorres r xf (P rv) (P' rv) hs (g rv) c"
shows "ccorres r xf
@ -411,25 +390,6 @@ crunches vcpuDisable, vcpuRestore, vcpuSave, vcpuEnable
for ksArch[wp]: "\<lambda>s. P (ksArchState s)"
(wp: crunch_wps)
(* FIXME move to Invariants_H *)
lemma invs_cicd_arch_state' [elim!]:
"all_invs_but_ct_idle_or_in_cur_domain' s \<Longrightarrow> valid_arch_state' s"
by (simp add: all_invs_but_ct_idle_or_in_cur_domain'_def valid_state'_def)
(* FIXME move to Invariants_H *)
lemma invs_cicd_no_0_obj'[elim!]:
"all_invs_but_ct_idle_or_in_cur_domain' s \<Longrightarrow> no_0_obj' s"
by (simp add: all_invs_but_ct_idle_or_in_cur_domain'_def valid_state'_def valid_pspace'_def)
(* FIXME: move *)
lemma vcpu_at_ko:
"vcpu_at' p s \<Longrightarrow> \<exists>vcpu. ko_at' (vcpu::vcpu) p s"
apply (clarsimp simp: typ_at'_def obj_at'_def ko_wp_at'_def projectKOs)
apply (case_tac ko; simp)
apply (rename_tac arch_kernel_object)
apply (case_tac arch_kernel_object, auto)[1]
done
(* FIXME: move *)
lemma vcpu_at_c_guard:
"\<lbrakk>vcpu_at' p s; (s, s') \<in> rf_sr\<rbrakk> \<Longrightarrow> c_guard (vcpu_Ptr p)"

View File

@ -203,14 +203,6 @@ proof -
qed
(* MOVE *)
lemma tcb_aligned':
"tcb_at' t s \<Longrightarrow> is_aligned t tcbBlockSizeBits"
apply (drule obj_at_aligned')
apply (simp add: objBits_simps)
apply (simp add: objBits_simps)
done
lemma tcb_at_not_NULL:
assumes tat: "tcb_at' t s"
shows "tcb_ptr_to_ctcb_ptr t \<noteq> NULL"

View File

@ -348,16 +348,6 @@ lemma getMRs_rel_state:
apply (erule doMachineOp_state)
done
(* FIXME: move *)
lemma setTCB_cur:
"\<lbrace>cur_tcb'\<rbrace> setObject t (v::tcb) \<lbrace>\<lambda>_. cur_tcb'\<rbrace>"
including no_pre
apply (wp cur_tcb_lift)
apply (simp add: setObject_def split_def updateObject_default_def)
apply wp
apply simp
done
lemma setThreadState_getMRs_rel:
"\<lbrace>getMRs_rel args buffer and cur_tcb' and case_option \<top> valid_ipc_buffer_ptr' buffer
and (\<lambda>_. runnable' st)\<rbrace>
@ -1275,9 +1265,6 @@ lemma invokeTCB_CopyRegisters_ccorres:
apply (clarsimp dest!: global'_no_ex_cap simp: invs'_def valid_state'_def | rule conjI)+
done
(* FIXME: move *)
lemmas mapM_x_append = mapM_x_append2
lemma invokeTCB_WriteRegisters_ccorres_helper:
"\<lbrakk> unat (f (of_nat n)) = incn
\<and> g (of_nat n) = register_from_H reg \<and> n'=incn
@ -2747,14 +2734,6 @@ lemma slotCapLongRunningDelete_ccorres:
dest!: ccte_relation_ccap_relation)
done
(* FIXME: move *)
lemma empty_fail_slotCapLongRunningDelete:
"empty_fail (slotCapLongRunningDelete slot)"
by (auto simp: slotCapLongRunningDelete_def Let_def
case_Null_If isFinalCapability_def
split: if_split
intro!: empty_fail_bind)
definition
isValidVTableRoot_C :: "cap_C \<Rightarrow> bool"
where
@ -2796,20 +2775,10 @@ lemma updateCapData_spec:
\<lbrace>ccap_relation (RetypeDecls_H.updateCapData preserve newData cap) \<acute>ret__struct_cap_C\<rbrace>"
by (simp add: updateCapData_spec)
lemma if_n_updateCapData_valid_strg:
"s \<turnstile>' cap \<longrightarrow> s \<turnstile>' (if P then cap else updateCapData prs v cap)"
by (simp add: valid_updateCapDataI split: if_split)
lemma length_excaps_map:
"length (excaps_map xcs) = length xcs"
by (simp add: excaps_map_def)
(* FIXME: move *)
lemma from_bool_all_helper:
"(\<forall>bool. from_bool bool = val \<longrightarrow> P bool)
= ((\<exists>bool. from_bool bool = val) \<longrightarrow> P (val \<noteq> 0))"
by (auto simp: from_bool_0)
lemma getSyscallArg_ccorres_foo':
"ccorres (\<lambda>a rv. rv = ucast (args ! n)) (\<lambda>x. ucast (ret__unsigned_long_' x))
(sysargs_rel args buffer and sysargs_rel_n args buffer n)

View File

@ -36,26 +36,6 @@ lemma ccorres_flip_Guard:
apply (fastforce intro: exec.Guard exec.GuardFault exec_handlers.intros)
done
(* FIXME: move *)
lemma empty_fail_findPDForASID[iff]:
"empty_fail (findPDForASID asid)"
apply (simp add: findPDForASID_def liftME_def)
apply (intro empty_fail_bindE, simp_all split: option.split)
apply (simp add: assertE_def split: if_split)
apply (simp add: assertE_def split: if_split)
apply (simp add: empty_fail_getObject)
apply (simp add: assertE_def liftE_bindE checkPDAt_def split: if_split)
done
(* FIXME: move *)
lemma empty_fail_findPDForASIDAssert[iff]:
"empty_fail (findPDForASIDAssert asid)"
apply (simp add: findPDForASIDAssert_def catch_def
checkPDAt_def checkPDUniqueToASID_def
checkPDASIDMapMembership_def)
apply (intro empty_fail_bind, simp_all split: sum.split)
done
end
@ -586,12 +566,6 @@ lemma leq_asid_bits_shift:
\<Longrightarrow> (x :: word32) >> asid_low_bits \<le> mask asid_high_bits"
by (simp add: leq_mask_shift asid_bits_def asid_low_bits_def asid_high_bits_def)
lemma ucast_asid_high_bits_is_shift:
"asid \<le> mask asid_bits
\<Longrightarrow> ucast (asid_high_bits_of asid) = (asid >> asid_low_bits)"
unfolding asid_bits_def asid_low_bits_def asid_high_bits_of_def
by (rule ucast_ucast_eq_mask_shift, simp)
lemma cap_small_frame_cap_get_capFMappedASID_spec:
"\<forall>s. \<Gamma>\<turnstile> \<lbrace>s. cap_get_tag \<acute>cap = scast cap_small_frame_cap\<rbrace>
Call cap_small_frame_cap_get_capFMappedASID_'proc
@ -1896,7 +1870,7 @@ lemma vcpuWriteReg_obj_at'_vcpuVPPIMasked:
\<lbrace>\<lambda>s. obj_at' (\<lambda>vcpu. P (vcpuVPPIMasked vcpu)) vcpuptr s \<rbrace>"
apply (simp add: vcpuWriteReg_def vcpuUpdate_def obj_at'_real_def)
apply (wp setObject_ko_wp_at[where n="objBits (undefined :: vcpu)"], simp)
apply (simp add: objBits_simps archObjSize_def vcpuBits_def')+
apply (simp add: objBits_simps archObjSize_def vcpuBits_def' vcpu_bits_def pageBits_def)+
apply (wpsimp wp: getVCPU_wp)+
apply (clarsimp simp: pred_conj_def is_vcpu'_def ko_wp_at'_def obj_at'_real_def projectKOs)
done
@ -2712,16 +2686,6 @@ lemma framesize_from_H_mask:
Kernel_C.ARMSection_def Kernel_C.ARMSuperSection_def
split: vmpage_size.splits)
(* FIXME: move *)
lemma dmo_invalidateCacheRange_RAM_invs'[wp]:
"valid invs' (doMachineOp (invalidateCacheRange_RAM vs ve ps)) (\<lambda>rv. invs')"
apply (wp dmo_invs' no_irq no_irq_invalidateCacheRange_RAM)
apply (clarsimp simp: disj_commute[of "pointerInUserData p s" for p s])
apply (erule use_valid)
apply (wp, simp)
done
lemma dmo_flushtype_case:
"(doMachineOp (case t of
ARM_HYP_H.flush_type.Clean \<Rightarrow> f
@ -3141,15 +3105,6 @@ lemma flushPage_ccorres:
apply (simp add: pageBits_def mask_eq_iff_w2p word_size)
done
lemma ignoreFailure_liftM:
"ignoreFailure = liftM (\<lambda>v. ())"
apply (rule ext)+
apply (simp add: ignoreFailure_def liftM_def
catch_def)
apply (rule bind_apply_cong[OF refl])
apply (simp split: sum.split)
done
lemma ccorres_pre_getObject_pte:
assumes cc: "\<And>rv. ccorres r xf (P rv) (P' rv) hs (f rv) c"
shows "ccorres r xf
@ -3352,7 +3307,7 @@ lemma array_assertion_abs_pte_16:
apply (drule(1) page_table_at_rf_sr, clarsimp)
apply (cases ptPtr, simp)
apply (erule clift_array_assertion_imp, simp_all)
apply (rule large_ptSlot_array_constraint, simp_all)
apply (rule large_ptSlot_array_constraint[simplified], simp_all)
done
lemmas ccorres_move_array_assertion_pte_16
@ -3842,14 +3797,6 @@ lemma ccap_relation_mapped_asid_0:
apply simp
done
(* FIXME: move *)
lemma getSlotCap_wp':
"\<lbrace>\<lambda>s. \<forall>cap. cte_wp_at' (\<lambda>c. cteCap c = cap) p s \<longrightarrow> Q cap s\<rbrace> getSlotCap p \<lbrace>Q\<rbrace>"
apply (simp add: getSlotCap_def)
apply (wp getCTE_wp')
apply (clarsimp simp: cte_wp_at_ctes_of)
done
lemma vmsz_aligned_aligned_pageBits:
"vmsz_aligned' ptr sz \<Longrightarrow> is_aligned ptr pageBits"
apply (simp add: vmsz_aligned'_def)
@ -4131,34 +4078,6 @@ lemma cap_lift_PDCap_Base:
declare mask_Suc_0[simp]
(* FIXME: move *)
lemma setCTE_asidpool':
"\<lbrace> ko_at' (ASIDPool pool) p \<rbrace> setCTE c p' \<lbrace>\<lambda>_. ko_at' (ASIDPool pool) p\<rbrace>"
apply (clarsimp simp: setCTE_def)
apply (simp add: setObject_def split_def)
apply (rule hoare_seq_ext [OF _ hoare_gets_post])
apply (clarsimp simp: valid_def in_monad)
apply (frule updateObject_type)
apply (clarsimp simp: obj_at'_def projectKOs)
apply (rule conjI)
apply (clarsimp simp: lookupAround2_char1)
apply (clarsimp split: if_split)
apply (case_tac obj', auto)[1]
apply (rename_tac arch_kernel_object)
apply (case_tac arch_kernel_object, auto)[1]
apply (simp add: updateObject_cte)
apply (clarsimp simp: updateObject_cte typeError_def magnitudeCheck_def in_monad
split: kernel_object.splits if_splits option.splits)
apply (clarsimp simp: ps_clear_upd' lookupAround2_char1)
done
(* FIXME: move *)
lemma udpateCap_asidpool':
"\<lbrace> ko_at' (ASIDPool pool) p \<rbrace> updateCap c p' \<lbrace>\<lambda>_. ko_at' (ASIDPool pool) p\<rbrace>"
apply (simp add: updateCap_def)
apply (wp setCTE_asidpool')
done
(* FIXME: move *)
lemma asid_pool_at_rf_sr:
"\<lbrakk>ko_at' (ASIDPool pool) p s; (s, s') \<in> rf_sr\<rbrakk> \<Longrightarrow>
@ -4169,21 +4088,10 @@ lemma asid_pool_at_rf_sr:
apply clarsimp
done
(* FIXME: move *)
lemma asid_pool_at_ko:
"asid_pool_at' p s \<Longrightarrow> \<exists>pool. ko_at' (ASIDPool pool) p s"
apply (clarsimp simp: typ_at'_def obj_at'_def ko_wp_at'_def projectKOs)
apply (case_tac ko, auto)
apply (rename_tac arch_kernel_object)
apply (case_tac arch_kernel_object, auto)[1]
apply (rename_tac asidpool)
apply (case_tac asidpool, auto)[1]
done
(* FIXME: move *)
lemma asid_pool_at_c_guard:
"\<lbrakk>asid_pool_at' p s; (s, s') \<in> rf_sr\<rbrakk> \<Longrightarrow> c_guard (ap_Ptr p)"
by (fastforce intro: typ_heap_simps dest!: asid_pool_at_ko asid_pool_at_rf_sr)
by (fastforce intro: typ_heap_simps dest!: ArchMove_C.asid_pool_at_ko asid_pool_at_rf_sr)
(* FIXME: move *)
lemma setObjectASID_Basic_ccorres:
@ -4315,7 +4223,7 @@ lemma performASIDPoolInvocation_ccorres:
apply simp
apply (clarsimp simp: cte_wp_at_ctes_of)
apply (rule conjI)
apply (clarsimp dest!: asid_pool_at_ko simp: obj_at'_def)
apply (clarsimp dest!: ArchMove_C.asid_pool_at_ko simp: obj_at'_def)
apply (rule cmap_relationE1[OF cmap_relation_cte], assumption+)
apply (clarsimp simp: typ_heap_simps cap_get_tag_isCap_ArchObject2
isPDCap_def isCap_simps

View File

@ -356,7 +356,7 @@ lemma unat_ucast_mask_pageBits_shift:
lemma mask_pageBits_shift_sum:
"unat n = unat (p && mask 3) \<Longrightarrow>
(p && ~~ mask pageBits) + (p && mask pageBits >> 3) * 8 + n = (p::machine_word)"
apply (clarsimp simp: kernel.word_shift_by_3)
apply (clarsimp simp: ArchMove_C.word_shift_by_3)
apply (subst word_plus_and_or_coroll)
apply (rule word_eqI)
apply (clarsimp simp: word_size pageBits_def nth_shiftl nth_shiftr word_ops_nth_size)
@ -537,64 +537,8 @@ lemma ran_array_map_conv:
"ran (array_map_conv f n c) = {y. \<exists>i\<le>n. f (index c (unat i)) = Some y}"
by (auto simp add: array_map_conv_def2 ran_def Collect_eq)
(* FIXME: move to somewhere sensible >>> *)
text \<open>Map inversion (implicitly assuming injectivity).\<close>
definition
"the_inv_map m = (\<lambda>s. if s\<in>ran m then Some (THE x. m x = Some s) else None)"
text \<open>Map inversion can be expressed by function inversion.\<close>
lemma the_inv_map_def2:
"the_inv_map m = (Some \<circ> the_inv_into (dom m) (the \<circ> m)) |` (ran m)"
apply (rule ext)
apply (clarsimp simp: the_inv_map_def the_inv_into_def dom_def)
apply (rule_tac f=The in arg_cong)
apply (rule ext)
apply auto
done
text \<open>The domain of a function composition with Some is the universal set.\<close>
lemma dom_comp_Some[simp]: "dom (comp Some f) = UNIV" by (simp add: dom_def)
text \<open>Assuming injectivity, map inversion produces an inversive map.\<close>
lemma is_inv_the_inv_map:
"inj_on m (dom m) \<Longrightarrow> is_inv m (the_inv_map m)"
apply (simp add: is_inv_def)
apply (intro conjI allI impI)
apply (simp add: the_inv_map_def2)
apply (auto simp add: the_inv_map_def inj_on_def dom_def)
done
lemma the_the_inv_mapI:
"inj_on m (dom m) \<Longrightarrow> m x = Some y \<Longrightarrow> the (the_inv_map m y) = x"
by (auto simp: the_inv_map_def ran_def inj_on_def dom_def)
lemma eq_restrict_map_None[simp]:
"restrict_map m A x = None \<longleftrightarrow> x ~: (A \<inter> dom m)"
by (auto simp: restrict_map_def split: if_split_asm)
lemma eq_the_inv_map_None[simp]: "the_inv_map m x = None \<longleftrightarrow> x\<notin>ran m"
by (simp add: the_inv_map_def2)
lemma is_inv_unique:
"is_inv f g \<Longrightarrow> is_inv f h \<Longrightarrow> g=h"
apply (rule ext)
apply (clarsimp simp: is_inv_def dom_def Collect_eq ran_def)
apply (drule_tac x=x in spec)+
apply (case_tac "g x", clarsimp+)
done
lemma assumes A: "is_inv f g" shows the_inv_map_eq: "the_inv_map f = g"
by (simp add: is_inv_unique[OF A A[THEN is_inv_inj, THEN is_inv_the_inv_map]])
(*<<<*)
lemmas word_le_p2m1 = word_up_bound[of w for w]
(* FIXME: move *)
lemma ran_map_comp_subset: "ran (map_comp f g) <= (ran f)"
by (fastforce simp: map_comp_def ran_def split: option.splits)
lemma eq_option_to_0_rev:
"Some 0 ~: A \<Longrightarrow> \<forall>x. \<forall>y\<in>A.
((=) \<circ> option_to_0) y x \<longrightarrow> (if x = 0 then None else Some x) = y"
by (clarsimp simp: option_to_0_def split: option.splits)
end
context state_rel begin
@ -700,12 +644,6 @@ lemma cready_queues_to_H_correct:
(* showing that cpspace_relation is actually unique >>>*)
(* FIXME: move *)
lemma inj_image_inv:
assumes inj_f: "inj f"
shows "f ` A = B \<Longrightarrow> inv f ` B = A"
by (drule sym) (simp add: image_inv_f_f[OF inj_f])
lemma cmap_relation_unique_0:
assumes inj_f: "inj f"
assumes r: "\<And>x y z p . \<lbrakk> r x z; r y z; a p = Some x; a' p = Some y; P p x; P' p y \<rbrakk> \<Longrightarrow> x=y"
@ -780,47 +718,6 @@ lemma ccontext_relation_imp_eq:
apply (auto dest: fpu_relation_imp_eq cregs_relation_imp_eq)
done
(* FIXME: move *)
lemma ran_tcb_cte_cases:
"ran tcb_cte_cases =
{(Structures_H.tcbCTable, tcbCTable_update),
(Structures_H.tcbVTable, tcbVTable_update),
(Structures_H.tcbReply, tcbReply_update),
(Structures_H.tcbCaller, tcbCaller_update),
(tcbIPCBufferFrame, tcbIPCBufferFrame_update)}"
by (auto simp add: tcb_cte_cases_def split: if_split_asm)
(* FIXME: move *)
lemma ps_clear_is_aligned_ksPSpace_None:
"\<lbrakk>ps_clear p n s; is_aligned p n; 0<d; d \<le> mask n\<rbrakk>
\<Longrightarrow> ksPSpace s (p + d) = None"
apply (simp add: ps_clear_def add_diff_eq[symmetric] mask_2pm1[symmetric])
apply (drule equals0D[where a="p + d"])
apply (simp add: dom_def word_gt_0 del: word_neq_0_conv)
apply (drule mp)
apply (rule word_plus_mono_right)
apply simp
apply (simp add: mask_2pm1)
apply (erule is_aligned_no_overflow')
apply (drule mp)
apply (case_tac "(0::machine_word)<2^n")
apply (frule le_m1_iff_lt[of "(2::machine_word)^n" d, THEN iffD1])
apply (simp add: mask_2pm1[symmetric])
apply (erule (1) is_aligned_no_wrap')
apply (simp add: is_aligned_mask mask_2pm1 not_less word_bits_def
power_overflow)
by assumption
(* FIXME: move *)
lemma ps_clear_is_aligned_ctes_None:
assumes "ps_clear p tcbBlockSizeBits s"
and "is_aligned p tcbBlockSizeBits"
shows "ksPSpace s (p + 2*2^cteSizeBits) = None"
and "ksPSpace s (p + 3*2^cteSizeBits) = None"
and "ksPSpace s (p + 4*2^cteSizeBits) = None"
by (auto intro: assms ps_clear_is_aligned_ksPSpace_None
simp: objBits_defs mask_def)+
lemma map_to_ctes_tcb_ctes:
notes if_cong[cong]
shows
@ -1068,24 +965,6 @@ lemma option_to_ctcb_ptr_inj:
apply (erule is_aligned_no_overflow_0; simp)
done
(* FIXME: move *)
lemma Collect_mono2: "Collect P \<subseteq> Collect Q \<longleftrightarrow> (\<forall>x. P x \<longrightarrow> Q x)" by auto
(* FIXME: move to Wellformed, turn valid_asid_pool' into an abbreviation >>>*)
primrec
wf_asid_pool' :: "asidpool \<Rightarrow> bool"
where
"wf_asid_pool' (ASIDPool pool) =
(dom pool \<subseteq> {0 .. 2^asid_low_bits - 1} \<and>
0 \<notin> ran pool \<and> (\<forall>x \<in> ran pool. is_aligned x pml4_bits))"
lemma valid_eq_wf_asid_pool'[simp]:
"valid_asid_pool' pool = (\<lambda>s. wf_asid_pool' pool)"
by (case_tac pool) simp
declare valid_asid_pool'.simps[simp del]
(*<<<*)
lemma cpspace_asidpool_relation_unique:
assumes invs: "\<forall>x\<in>ran (map_to_asidpools ah). wf_asid_pool' x"
"\<forall>x\<in>ran (map_to_asidpools ah'). wf_asid_pool' x"

View File

@ -992,12 +992,6 @@ lemma addrFromPPtr_mask_middle_pml4ShiftBits:
apply (word_bitwise, clarsimp)
done
(* FIXME: move *)
lemma obj_at_kernel_mappings':
"\<lbrakk>pspace_in_kernel_mappings' s; obj_at' P p s\<rbrakk> \<Longrightarrow>
p \<in> kernel_mappings"
by (clarsimp simp: pspace_in_kernel_mappings'_def obj_at'_def dom_def)
lemma decodeX64PageTableInvocation_ccorres:
"\<lbrakk>interpret_excaps extraCaps' = excaps_map extraCaps; isPageTableCap cp\<rbrakk> \<Longrightarrow>
ccorres
@ -1522,20 +1516,6 @@ lemma pte_sadness:
apply (cases pte', cases pte, simp)
done
lemma hd_in_zip_set:
"slots \<noteq> [] \<Longrightarrow> (hd slots, 0) \<in> set (zip slots [0.e.of_nat (length slots - Suc 0)::machine_word])"
by (cases slots; simp add: upto_enum_word upto_0_to_n2 del: upt_Suc)
lemma last_in_zip_set:
"\<lbrakk> slots \<noteq> []; length js = length slots \<rbrakk> \<Longrightarrow> (last slots, last js) \<in> set (zip slots js)"
apply (simp add: in_set_zip last_conv_nth)
apply (rule_tac x="length slots - 1" in exI)
apply clarsimp
apply (subst last_conv_nth)
apply (cases js; simp)
apply simp
done
definition
"isVMPTE entry \<equiv> \<exists>x y. entry = (VMPTE x, VMPTEPtr y)"
@ -1869,10 +1849,6 @@ lemma slotcap_in_mem_valid:
apply (erule(1) ctes_of_valid')
done
lemma list_length_less:
"(args = [] \<or> length args \<le> Suc 0) = (length args < 2)"
by (case_tac args,simp_all)
lemma unat_less_iff64:
"\<lbrakk>unat (a::machine_word) = b;c < 2^word_bits\<rbrakk>
\<Longrightarrow> (a < of_nat c) = (b < c)"
@ -1897,26 +1873,11 @@ lemma injection_handler_if_returnOk:
lemma pbfs_less: "pageBitsForSize sz < 31"
by (case_tac sz,simp_all add: bit_simps)
lemma at_least_2_args:
"\<not> length args < 2 \<Longrightarrow> \<exists>a b c. args = a#b#c"
apply (case_tac args)
apply simp
apply (case_tac list)
apply simp
apply (case_tac lista)
apply simp
apply simp
done
definition
to_option :: "('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a option"
where
"to_option f x \<equiv> if f x then Some x else None"
(* FIXME: move *)
lemma valid_objs_valid_pte': "\<lbrakk> valid_objs' s ; ko_at' (ko :: pte) p s \<rbrakk> \<Longrightarrow> valid_pte' ko s"
by (fastforce simp add: obj_at'_def ran_def valid_obj'_def projectKOs valid_objs'_def)
lemma cte_wp_at_eq_gsMaxObjectSize:
"cte_wp_at' ((=) cap o cteCap) slot s
\<Longrightarrow> valid_global_refs' s
@ -1930,13 +1891,6 @@ lemma two_nat_power_pageBitsForSize_le:
"(2 :: nat) ^ pageBits \<le> 2 ^ pageBitsForSize vsz"
by (cases vsz, simp_all add: pageBits_def bit_simps)
(* FIXME: move *)
lemma is_aligned_pageBitsForSize_minimum:
"\<lbrakk> is_aligned p (pageBitsForSize sz) ; n \<le> pageBits \<rbrakk> \<Longrightarrow> is_aligned p n"
apply (cases sz; clarsimp simp: pageBits_def)
apply (erule is_aligned_weaken, simp)+
done
lemma ptrFromPAddr_add_left:
"ptrFromPAddr (x + y) = ptrFromPAddr x + y"
unfolding ptrFromPAddr_def by simp
@ -2904,23 +2858,11 @@ lemma framesize_from_H_mask2:
Kernel_C.X64_HugePage_def)+
done
lemma rel_option_alt_def:
"rel_option f a b = (
(a = None \<and> b = None)
\<or> (\<exists>x y. a = Some x \<and> b = Some y \<and> f x y))"
apply (case_tac a, case_tac b, simp, simp, case_tac b, auto)
done
lemma injection_handler_stateAssert_relocate:
"injection_handler Inl (stateAssert ass xs >>= f) >>=E g
= do v \<leftarrow> stateAssert ass xs; injection_handler Inl (f ()) >>=E g od"
by (simp add: injection_handler_def handleE'_def bind_bindE_assoc bind_assoc)
(* FIXME: move to where is_aligned_ptrFromPAddr is *)
lemma is_aligned_ptrFromPAddr_pageBitsForSize:
"is_aligned p (pageBitsForSize sz) \<Longrightarrow> is_aligned (ptrFromPAddr p) (pageBitsForSize sz)"
by (cases sz ; simp add: is_aligned_ptrFromPAddr_n pageBits_def bit_simps)
lemma makeUserPDPTEPageDirectory_spec:
"\<forall>s. \<Gamma> \<turnstile>
\<lbrace>s. vmsz_aligned (\<acute>paddr) X64SmallPage\<rbrace>
@ -4838,103 +4780,10 @@ lemma invokeX86PortControl_ccorres:
global_ioport_bitmap_relation_def Let_def typ_heap_simps)
done
lemma unat_length_4_helper:
"\<lbrakk>unat (l::machine_word) = length args; \<not> l < 4\<rbrakk> \<Longrightarrow> \<exists>x xa xb xc xs. args = x#xa#xb#xc#xs"
apply (case_tac args; clarsimp simp: unat_eq_0)
by (rename_tac list, case_tac list, clarsimp, unat_arith)+
lemma ucast_drop_big_mask:
"UCAST(64 \<rightarrow> 16) (x && 0xFFFF) = UCAST(64 \<rightarrow> 16) x"
by word_bitwise
lemma first_port_last_port_compare:
"UCAST(16 \<rightarrow> 32 signed) (UCAST(64 \<rightarrow> 16) (xa && 0xFFFF))
<s UCAST(16 \<rightarrow> 32 signed) (UCAST(64 \<rightarrow> 16) (x && 0xFFFF))
= (UCAST(64 \<rightarrow> 16) xa < UCAST(64 \<rightarrow> 16) x)"
apply (clarsimp simp: word_sless_alt ucast_drop_big_mask)
apply (subst sint_ucast_eq_uint, clarsimp simp: is_down)+
by (simp add: word_less_alt)
(* FIXME X64: move to state rel? *)
abbreviation
"x86KSAllocatedIOPorts_ptr == ioport_table_Ptr (symbol_table ''x86KSAllocatedIOPorts'')"
(* FIXME X64: use earlier or replace with earlier def? *)
definition
port_mask :: "16 word \<Rightarrow> 16 word \<Rightarrow> machine_word"
where
"port_mask start end =
mask (unat (end && mask wordRadix)) && ~~ mask (unat (start && mask wordRadix))"
lemma shiftr_le_mask:
fixes w :: "'a::len word"
shows "w >> n \<le> mask (LENGTH('a) - n)"
by (metis and_mask_eq_iff_shiftr_0 le_mask_iff shiftr_mask_eq word_size)
lemma word_minus_1_shiftr:
notes word_unat.Rep_inject[simp del]
fixes w :: "'a::len word"
assumes low_bits_zero: "w && mask n = 0"
assumes neq_zero: "w \<noteq> 0"
shows "(w - 1) >> n = (w >> n) - 1"
using neq_zero low_bits_zero
apply (subgoal_tac "n < LENGTH('a)")
prefer 2
apply (metis not_less ucast_id ucast_mask_drop)
apply (rule_tac t="w - 1" and s="(w && ~~ mask n) - 1" in subst,
fastforce simp: low_bits_zero mask_eq_x_eq_0)
apply (clarsimp simp: mask_eq_0_eq_x neg_mask_is_div lt1_neq0[symmetric])
apply (subst shiftr_div_2n_w, fastforce simp: word_size)+
apply (rule word_uint.Rep_eqD)
apply (simp only: uint_word_ariths uint_div uint_power_lower)
apply (subst mod_pos_pos_trivial, fastforce, fastforce)+
apply (subst mod_pos_pos_trivial)
apply (fastforce simp: le_diff_eq uint_2p_alt word_le_def)
apply (subst uint_1[symmetric])
apply (fastforce intro: uint_sub_lt2p)
apply (subst int_div_sub_1, fastforce)
apply (clarsimp simp: and_mask_dvd low_bits_zero)
apply (subst mod_pos_pos_trivial)
apply (metis linorder_not_less mult_zero_left shiftr_div_2n shiftr_div_2n_w uint_eq_0
uint_le_0_iff word_less_1 word_uint.Rep_inject word_size zle_diff1_eq)
apply (metis shiftr_div_2n uint_1 uint_sub_lt2p)
apply fastforce
done
(* FIXME: move to Word *)
lemma ucast_shiftr:
"UCAST('a::len \<rightarrow> 'b::len) w >> n = UCAST('a \<rightarrow> 'b) ((w && mask LENGTH('b)) >> n)"
apply (rule word_eqI[rule_format]; rule iffI; clarsimp simp: nth_ucast nth_shiftr word_size)
done
(* FIXME: move to Word *)
lemma mask_eq_ucast_shiftr:
assumes mask: "w && mask LENGTH('b) = w"
shows "UCAST('a::len \<rightarrow> 'b::len) w >> n = UCAST('a \<rightarrow> 'b) (w >> n)"
by (rule ucast_shiftr[where 'a='a and 'b='b, of w n, simplified mask])
(* FIXME: move to Word *)
lemma mask_eq_ucast_shiftl:
assumes "w && mask (LENGTH('a) - n) = w"
shows "UCAST('a::len \<rightarrow> 'b::len) w << n = UCAST('a \<rightarrow> 'b) (w << n)"
apply (rule subst[where P="\<lambda>c. ucast c << n = ucast (c << n)", OF assms])
apply (rule word_eqI[rule_format]; rule iffI;
clarsimp simp: nth_ucast nth_shiftl word_size;
drule test_bit_size; simp add: word_size)
done
(* FIXME: replace by mask_mono *)
lemma mask_le_mono:
"m \<le> n \<Longrightarrow> mask m \<le> mask n"
apply (subst and_mask_eq_iff_le_mask[symmetric])
by (auto intro: word_eqI simp: word_size)
(* FIXME: move to Word *)
lemma word_and_mask_eq_le_mono:
"w && mask m = w \<Longrightarrow> m \<le> n \<Longrightarrow> w && mask n = w"
apply (simp add: and_mask_eq_iff_le_mask)
by (erule order.trans, erule mask_le_mono)
lemma first_last_highbits_eq_port_set:
fixes f l :: "16 word"
fixes arr :: "machine_word[1024]"
@ -5035,10 +4884,6 @@ lemma port_set_in_first_word:
by (simp add: unat_of_nat)
done
lemma word_not_exists_nth:
"(w::'a::len word) = 0 \<Longrightarrow> \<forall>i<LENGTH('a). \<not> w !! i"
by (clarsimp simp: nth_0)
lemma bitmap_word_zero_no_bits_set1:
fixes f l :: "16 word"
fixes arr :: "machine_word[1024]"
@ -5062,34 +4907,6 @@ lemma bitmap_word_zero_no_bits_set1:
apply (erule disjE; clarsimp simp:word_less_nat_alt)
done
lemma ucast_shiftl_6_absorb:
fixes f l :: "16 word"
assumes "f \<le> l"
assumes "f >> 6 < l >> 6"
shows "UCAST(16\<rightarrow>32 signed) ((f >> 6) + 1) << 6 = UCAST(16 \<rightarrow> 32 signed) (((f >> 6) + 1) << 6)"
using assms
by (word_bitwise, auto)
(* FIXME: move *)
lemma word_highbits_bounded_highbits_eq:
"\<lbrakk>x \<le> (y :: 'a::len word); y < (x >> n) + 1 << n\<rbrakk> \<Longrightarrow> x >> n = y >> n"
apply (cases "n < LENGTH('a)")
prefer 2
apply (simp add: shiftr_eq_0)
apply (drule_tac n=n in le_shiftr)
apply (subst (asm) word_shiftl_add_distrib)
apply (drule_tac word_less_sub_1)
apply (simp only: add_diff_eq[symmetric] mask_def[symmetric] and_not_mask[symmetric])
apply (drule_tac u=y and n=n in le_shiftr)
apply (subgoal_tac "(x && ~~ mask n) + mask n >> n \<le> x >> n")
apply fastforce
apply (subst aligned_shift')
apply (fastforce simp: mask_lt_2pn)
apply (fastforce simp: is_aligned_neg_mask)
apply fastforce
apply (fastforce simp: mask_shift)
done
lemma bitmap_word_zero_no_bits_set2:
fixes f l :: "16 word"
fixes arr :: "machine_word[1024]"
@ -5108,10 +4925,6 @@ lemma bitmap_word_zero_no_bits_set2:
apply (drule_tac x="unat (port && mask 6)" in spec, clarsimp simp: neg_mask_test_bit not_le word_le_nat_alt)
done
lemma word_eq_cast_unsigned:
"(x = y) = (UCAST ('a signed \<rightarrow> ('a :: len)) x = ucast y)"
by (simp add: word_eq_iff nth_ucast)
lemma isIOPortRangeFree_spec:
notes ucast_mask = ucast_and_mask[where n=6, simplified mask_def, simplified]
notes not_max_word_simps = and_not_max_word shiftr_not_max_word and_mask_not_max_word
@ -5480,11 +5293,6 @@ lemma ct_active_st_tcb_at_minor':
by (clarsimp simp: st_tcb_at'_def ct_in_state'_def obj_at'_def projectKOs,
case_tac "tcbState obj"; clarsimp)+
(* FIXME: move to Lib *)
lemma length_Suc_0_conv:
"length x = Suc 0 = (\<exists>y. x = [y])"
by (induct x; clarsimp)
lemma decodeIOPortInvocation_ccorres:
notes if_cong[cong]
assumes "interpret_excaps extraCaps' = excaps_map extraCaps"

View File

@ -15,31 +15,12 @@ begin
context begin interpretation Arch . (*FIXME: arch_split*)
(* Short-hand for unfolding cumbersome machine constants *)
(* FIXME MOVE these should be in refine, and the _eq forms should NOT be declared [simp]! *)
declare word_neq_0_conv [simp del]
(* Rule previously in the simpset, now not. *)
declare ptr_add_def' [simp]
(* works much better *)
lemmas typ_heap_simps' = typ_heap_simps c_guard_clift
lemmas asUser_return = submonad.return [OF submonad_asUser]
lemma setMRs_Nil:
"setMRs thread buffer [] = stateAssert (tcb_at' thread) [] >>= (\<lambda>_. return 0)"
unfolding setMRs_def
by (simp add: zipWithM_x_def sequence_x_def zipWith_def
asUser_return)
lemmas asUser_bind_distrib =
submonad_bind [OF submonad_asUser submonad_asUser submonad_asUser]
lemma ps_clear_upd_None:
"ksPSpace s y = None \<Longrightarrow>
ps_clear x n (ksPSpace_update (\<lambda>a. (ksPSpace s)(y := None)) s') = ps_clear x n s"
by (rule iffI | clarsimp elim!: ps_clear_domE | fastforce)+
(* Rule previously in the simpset, now not. *)
declare ptr_add_def' [simp]
(* Levity: moved from Ipc_C (20090419 09:44:31) *) (* and remove from Syscall_C *)
lemma empty_fail_doMachineOp[intro!]:
@ -68,13 +49,6 @@ lemma no_overlap_new_cap_addrs_disjoint:
apply auto
done
lemma empty_fail_asUser[iff]:
"empty_fail m \<Longrightarrow> empty_fail (asUser t m)"
apply (simp add: asUser_def split_def)
apply (intro empty_fail_bind, simp_all)
apply (simp add: select_f_def empty_fail_def)
done
declare empty_fail_doMachineOp [simp]
lemma empty_fail_loadWordUser[intro!, simp]:
@ -106,97 +80,6 @@ lemma empty_fail_unifyFailure [intro!, simp]:
handleE'_def throwError_def
split: sum.splits)
lemma asUser_mapM_x:
"(\<And>x. empty_fail (f x)) \<Longrightarrow>
asUser t (mapM_x f xs) = do stateAssert (tcb_at' t) []; mapM_x (\<lambda>x. asUser t (f x)) xs od"
apply (simp add: mapM_x_mapM asUser_bind_distrib)
apply (subst submonad_mapM [OF submonad_asUser submonad_asUser])
apply simp
apply (simp add: asUser_return bind_assoc o_def)
apply (rule ext)
apply (rule bind_apply_cong [OF refl])+
apply (clarsimp simp: in_monad dest!: fst_stateAssertD)
apply (drule use_valid, rule mapM_wp', rule asUser_typ_ats, assumption)
apply (simp add: stateAssert_def get_def NonDetMonad.bind_def)
done
lemma asUser_get_registers:
"\<lbrace>tcb_at' target\<rbrace>
asUser target (mapM getRegister xs)
\<lbrace>\<lambda>rv s. obj_at' (\<lambda>tcb. map ((user_regs o atcbContextGet o tcbArch) tcb) xs = rv) target s\<rbrace>"
apply (induct xs)
apply (simp add: mapM_empty asUser_return)
apply wp
apply simp
apply (simp add: mapM_Cons asUser_bind_distrib asUser_return)
apply wp
apply simp
apply (rule hoare_strengthen_post)
apply (erule hoare_vcg_conj_lift)
apply (rule asUser_inv)
apply (simp add: getRegister_def)
apply (wp mapM_wp')
apply clarsimp
apply (erule(1) obj_at_conj')
apply (wp)
apply (simp add: asUser_def split_def threadGet_def)
apply (wp getObject_tcb_wp)
apply (clarsimp simp: getRegister_def simpler_gets_def
obj_at'_def)
done
lemma projectKO_user_data_device:
"(projectKO_opt ko = Some (t :: user_data_device)) = (ko = KOUserDataDevice)"
by (cases ko)
(auto simp: projectKO_opts_defs split: arch_kernel_object.splits)
lemma device_data_at_ko:
"typ_at' UserDataDeviceT p s \<Longrightarrow> ko_at' UserDataDevice p s"
apply (clarsimp simp: typ_at'_def obj_at'_def ko_wp_at'_def
projectKO_user_data_device projectKO_eq projectKO_eq2)
apply (case_tac ko, auto)
done
(* FIXME: move *)
lemma user_data_at_ko:
"typ_at' UserDataT p s \<Longrightarrow> ko_at' UserData p s"
apply (clarsimp simp: typ_at'_def obj_at'_def ko_wp_at'_def projectKOs)
apply (case_tac ko, auto)
done
(* FIXME: move *)
lemma map_to_ko_atI:
"\<lbrakk>(projectKO_opt \<circ>\<^sub>m ksPSpace s) x = Some v;
pspace_aligned' s; pspace_distinct' s\<rbrakk>
\<Longrightarrow> ko_at' v x s"
apply (clarsimp simp: map_comp_Some_iff)
apply (erule (2) aligned_distinct_obj_atI')
apply (simp add: project_inject)
done
lemma empty_fail_rethrowFailure:
"empty_fail f \<Longrightarrow> empty_fail (rethrowFailure fn f)"
apply (simp add: rethrowFailure_def handleE'_def)
apply (erule empty_fail_bind)
apply (simp split: sum.split)
done
lemma empty_fail_resolveAddressBits:
"empty_fail (resolveAddressBits cap cptr bits)"
proof -
note empty_fail_assertE[iff]
show ?thesis
apply (rule empty_fail_use_cutMon)
apply (induct rule: resolveAddressBits.induct)
apply (subst resolveAddressBits.simps)
apply (unfold Let_def cnode_cap_case_if fun_app_def
K_bind_def haskell_assertE_def split_def)
apply (intro empty_fail_cutMon_intros)
apply (clarsimp simp: empty_fail_drop_cutMon empty_fail_whenEs
locateSlot_conv returnOk_liftE[symmetric]
isCap_simps)+
done
qed
lemma empty_fail_getReceiveSlots:
"empty_fail (getReceiveSlots r rbuf)"
@ -237,12 +120,6 @@ lemma exec_Basic_Guard_UNIV:
end
(* FIXME MOVE to where option_to_0 is defined *)
lemma option_to_0_simps [simp]:
"option_to_0 None = 0"
"option_to_0 (Some x) = x"
by (auto simp: option_to_0_def split: option.split)
definition
"option_to_ptr \<equiv> Ptr o option_to_0"
@ -261,8 +138,4 @@ lemma option_to_ptr_not_0:
"\<lbrakk> p \<noteq> 0 ; option_to_ptr v = Ptr p \<rbrakk> \<Longrightarrow> v = Some p"
by (clarsimp simp: option_to_ptr_def option_to_0_def split: option.splits)
(* FIXME: move *)
lemma of_bool_from_bool: "of_bool = from_bool"
by (rule ext, simp add: from_bool_def split: bool.split)
end

View File

@ -117,16 +117,6 @@ lemma lookupCapAndSlot_ccorres :
done
(* FIXME: move *)
lemma injection_handler_liftM:
"injection_handler f
= liftM (\<lambda>v. case v of Inl ex \<Rightarrow> Inl (f ex) | Inr rv \<Rightarrow> Inr rv)"
apply (intro ext, simp add: injection_handler_def liftM_def
handleE'_def)
apply (rule bind_apply_cong, rule refl)
apply (simp add: throwError_def split: sum.split)
done
lemma lookupErrorOnFailure_ccorres:
"ccorres (f \<currency> r) xf P P' hs a c
\<Longrightarrow> ccorres ((\<lambda>x y z. \<exists>w. x = FailedLookup isSource w \<and> f w y z) \<currency> r)

View File

@ -35,8 +35,6 @@ lemma maskCapRights_cap_cases:
done
lemma imp_ignore: "B \<Longrightarrow> A \<longrightarrow> B" by blast
(* FIXME x64: ucast? see how it goes *)
lemma wordFromVMRights_spec:
"\<forall>s. \<Gamma> \<turnstile> {s} Call wordFromVMRights_'proc \<lbrace>\<acute>ret__unsigned_long = \<^bsup>s\<^esup>vm_rights\<rbrace>"
@ -1167,16 +1165,6 @@ lemma updateMDB_mdbPrev_set_mdbNext:
(* *)
(************************************************************************)
(* FIXME: move *)
lemma cteSizeBits_eq:
"cteSizeBits = cte_level_bits"
by (simp add: cte_level_bits_def cteSizeBits_def)
(* FIXME: move *)
lemma cteSizeBits_le_cte_level_bits[simp]:
"cteSizeBits \<le> cte_level_bits"
by (simp add: cte_level_bits_def cteSizeBits_def)
(* FIXME: rename *)
lemma is_aligned_3_prev:
"\<lbrakk> valid_mdb' s; pspace_aligned' s; ctes_of s p = Some cte \<rbrakk>
@ -2144,137 +2132,6 @@ lemma fold_heap_modify_commute_Array:
apply (clarsimp simp: heap_modify_compose_Array[OF s, THEN fun_cong, simplified] o_def)
done
(* FIXME: move *)
lemma msb_le_mono:
fixes v w :: "'a::len word"
shows "v \<le> w \<Longrightarrow> msb v \<Longrightarrow> msb w"
by (simp add: msb_big)
(* FIXME: move *)
lemma neg_msb_le_mono:
fixes v w :: "'a::len word"
shows "v \<le> w \<Longrightarrow> \<not> msb w \<Longrightarrow> \<not> msb v"
by (simp add: msb_big)
(* FIXME: move *)
lemmas msb_less_mono = msb_le_mono[OF less_imp_le]
lemmas neg_msb_less_mono = neg_msb_le_mono[OF less_imp_le]
(* FIXME: move *)
lemma word_sless_iff_less:
"\<lbrakk> \<not> msb v; \<not> msb w \<rbrakk> \<Longrightarrow> v <s w \<longleftrightarrow> v < w"
by (simp add: word_sless_alt sint_eq_uint word_less_alt)
(* FIXME: move *)
lemmas word_sless_imp_less = word_sless_iff_less[THEN iffD1, rotated 2]
lemmas word_less_imp_sless = word_sless_iff_less[THEN iffD2, rotated 2]
(* FIXME: move *)
lemma word_sle_iff_le:
"\<lbrakk> \<not> msb v; \<not> msb w \<rbrakk> \<Longrightarrow> v <=s w \<longleftrightarrow> v \<le> w"
by (simp add: word_sle_def sint_eq_uint word_le_def)
(* FIXME: move *)
lemmas word_sle_imp_le = word_sle_iff_le[THEN iffD1, rotated 2]
lemmas word_le_imp_sle = word_sle_iff_le[THEN iffD2, rotated 2]
(* FIXME: move to Word_Lib *)
lemma word_upcast_shiftr:
assumes "LENGTH('a::len) \<le> LENGTH('b::len)"
shows "UCAST('a \<rightarrow> 'b) (w >> n) = UCAST('a \<rightarrow> 'b) w >> n"
apply (intro word_eqI impI iffI; clarsimp simp: word_size nth_shiftr nth_ucast)
apply (drule test_bit_size)
using assms by (simp add: word_size)
lemma word_upcast_neg_msb:
"LENGTH('a::len) < LENGTH('b::len) \<Longrightarrow> \<not> msb (UCAST('a \<rightarrow> 'b) w)"
apply (clarsimp simp: ucast_def msb_word_of_int)
apply (drule bin_nth_uint_imp)
by simp
(* FIXME: move to Word_Lib *)
lemma word_upcast_0_sle:
"LENGTH('a::len) < LENGTH('b::len) \<Longrightarrow> 0 <=s UCAST('a \<rightarrow> 'b) w"
by (simp add: word_sle_iff_le[OF word_msb_0 word_upcast_neg_msb])
(* FIXME: move to Word_Lib *)
lemma scast_ucast_up_eq_ucast:
assumes "LENGTH('a::len) < LENGTH('b::len)"
shows "SCAST('b \<rightarrow> 'c) (UCAST('a \<rightarrow> 'b) w) = UCAST('a \<rightarrow> 'c::len) w"
using assms
apply (subst scast_eq_ucast; simp)
apply (clarsimp simp: ucast_def msb_word_of_int)
apply (drule bin_nth_uint_imp)
apply simp
done
lemma not_max_word_iff_less:
"w \<noteq> max_word \<longleftrightarrow> w < max_word"
by (simp add: order_less_le)
lemma ucast_increment:
assumes "w \<noteq> max_word"
shows "UCAST('a::len \<rightarrow> 'b::len) w + 1 = UCAST('a \<rightarrow> 'b) (w + 1)"
apply (cases "LENGTH('b) \<le> LENGTH('a)")
apply (simp add: ucast_down_add is_down)
apply (subgoal_tac "uint w + 1 < 2 ^ LENGTH('a)")
apply (subgoal_tac "uint w + 1 < 2 ^ LENGTH('b)")
apply (subst word_uint_eq_iff)
apply (simp add: uint_arith_simps uint_up_ucast is_up)
apply (erule less_trans, rule power_strict_increasing, simp, simp)
apply (subst less_diff_eq[symmetric])
using assms
apply (simp add: not_max_word_iff_less word_less_alt)
apply (erule less_le_trans)
apply (simp add: max_word_def)
done
lemma max_word_gt_0:
"0 < max_word"
by (simp add: le_neq_trans[OF max_word_max] max_word_not_0)
lemma and_not_max_word:
"m \<noteq> max_word \<Longrightarrow> w && m \<noteq> max_word"
by (simp add: not_max_word_iff_less word_and_less')
lemma mask_not_max_word:
"m < LENGTH('a::len) \<Longrightarrow> mask m \<noteq> (max_word :: 'a word)"
by (metis shiftl_1_not_0 shiftl_mask_is_0 word_bool_alg.conj_one_right)
lemmas and_mask_not_max_word =
and_not_max_word[OF mask_not_max_word]
lemma shiftr_not_max_word:
"0 < n \<Longrightarrow> w >> n \<noteq> max_word"
apply (simp add: not_max_word_iff_less)
apply (cases "n < size w")
apply (cases "w = 0")
apply (simp add: max_word_gt_0)
apply (subst shiftr_div_2n_w, assumption)
apply (rule less_le_trans[OF div_less_dividend_word max_word_max])
apply simp
apply (metis word_size_gt_0 less_numeral_extra(3) mask_def nth_mask power_0 shiftl_t2n)
apply (simp add: not_less word_size)
apply (subgoal_tac "w >> n = 0"; simp add: max_word_gt_0 shiftr_eq_0)
done
lemma word_sandwich1:
fixes a b c :: "'a::len word"
assumes "a < b"
assumes "b <= c"
shows "0 < b - a \<and> b - a <= c"
using assms diff_add_cancel order_less_irrefl add_0 word_le_imp_diff_le
word_le_less_eq word_neq_0_conv
by metis
lemma word_sandwich2:
fixes a b :: "'a::len word"
assumes "0 < a"
assumes "a <= b"
shows "b - a < b"
using assms less_le_trans word_diff_less
by blast
lemma make_pattern_spec:
defines "bits \<equiv> 0x40 :: 32 sword"
notes word_less_imp_less_0 = revcut_rl[OF word_less_imp_sless[OF _ word_msb_0]]
@ -2403,43 +2260,9 @@ lemma word_set_or_clear_test_bit:
shows "i < LENGTH('a) \<Longrightarrow> word_set_or_clear b p w !! i = (if p !! i then b else w !! i)"
by (auto simp: word_set_or_clear_def word_ops_nth_size word_size split: if_splits)
lemma unat_and_mask_less_2p:
fixes w :: "'a::len word"
shows "m < LENGTH('a) \<Longrightarrow> unat (w && mask m) < 2 ^ m"
by (simp add: unat_less_helper and_mask_less')
lemma unat_shiftr_less_2p:
fixes w :: "'a::len word"
shows "n + m = LENGTH('a) \<Longrightarrow> unat (w >> n) < 2 ^ m"
by (cases "n = 0"; simp add: unat_less_helper shiftr_less_t2n3)
lemma nat_div_less_mono:
fixes m n :: nat
shows "m div d < n div d \<Longrightarrow> m < n"
by (meson div_le_mono not_less)
lemma word_shiftr_less_mono:
fixes w :: "'a::len word"
shows "w >> n < v >> n \<Longrightarrow> w < v"
by (auto simp: word_less_nat_alt shiftr_div_2n' elim: nat_div_less_mono)
lemma word_shiftr_less_mask:
fixes w :: "'a::len word"
shows "(w >> n < v >> n) \<longleftrightarrow> (w && ~~mask n < v && ~~mask n)"
by (metis (mono_tags) le_shiftr mask_shift shiftr_eq_neg_mask_eq word_le_less_eq word_le_not_less)
lemma word_shiftr_le_mask:
fixes w :: "'a::len word"
shows "(w >> n \<le> v >> n) \<longleftrightarrow> (w && ~~mask n \<le> v && ~~mask n)"
by (metis (mono_tags) le_shiftr mask_shift shiftr_eq_neg_mask_eq word_le_less_eq word_le_not_less)
lemma word_shiftr_eq_mask:
fixes w :: "'a::len word"
shows "(w >> n = v >> n) \<longleftrightarrow> (w && ~~mask n = v && ~~mask n)"
by (metis (mono_tags) mask_shift shiftr_eq_neg_mask_eq)
lemmas word_shiftr_cmp_mask =
word_shiftr_less_mask word_shiftr_le_mask word_shiftr_eq_mask
lemma heap_modify_fold:
"heap_update p (f (h_val h p)) h = heap_modify p f h"
by (simp add: heap_modify_def)
lemma fold_array_update_index:
fixes arr :: "'a::c_type['b::finite]"
@ -2447,30 +2270,6 @@ lemma fold_array_update_index:
shows "fold (\<lambda>i arr. Arrays.update arr i (f i)) is arr.[i] = (if i \<in> set is then f i else arr.[i])"
using assms by (induct "is" arbitrary: arr) (auto split: if_splits)
lemma if_if_if_same_output:
"(if c1 then if c2 then t else f else if c3 then t else f) = (if c1 \<and> c2 \<or> \<not>c1 \<and> c3 then t else f)"
by (simp split: if_splits)
lemma word_le_split_mask:
"(w \<le> v) \<longleftrightarrow> (w >> n < v >> n \<or> w >> n = v >> n \<and> w && mask n \<le> v && mask n)"
apply (simp add: word_shiftr_eq_mask word_shiftr_less_mask)
apply (rule subst[where P="\<lambda>c. c \<le> d = e" for d e, OF AND_NOT_mask_plus_AND_mask_eq[where n=n]])
apply (rule subst[where P="\<lambda>c. d \<le> c = e" for d e, OF AND_NOT_mask_plus_AND_mask_eq[where n=n]])
apply (rule iffI)
apply safe
apply (fold_subgoals (prefix))[2]
apply (subst atomize_conj)
apply (rule context_conjI)
apply (metis AND_NOT_mask_plus_AND_mask_eq neg_mask_mono_le word_le_less_eq)
apply (metis add.commute word_and_le1 word_bw_comms(1) word_plus_and_or_coroll2 word_plus_mcs_4)
apply (metis Groups.add_ac(2) neg_mask_mono_le word_le_less_eq word_not_le word_plus_and_or_coroll2)
apply (metis add.commute word_and_le1 word_bw_comms(1) word_plus_and_or_coroll2 word_plus_mcs_3)
done
lemma heap_modify_fold:
"heap_update p (f (h_val h p)) h = heap_modify p f h"
by (simp add: heap_modify_def)
lemma t_hrs_'_update_heap_modify_fold:
"gs\<lparr> t_hrs_' := hrs_mem_update (heap_update p (f (h_val (hrs_mem (t_hrs_' gs)) p))) (t_hrs_' gs) \<rparr>
= t_hrs_'_update (hrs_mem_update (heap_modify p f)) gs"

View File

@ -603,14 +603,6 @@ lemma to_bool_false [simp]:
unfolding to_bool_def false_def
by simp
(* MOVE *)
lemma tcb_aligned':
"tcb_at' t s \<Longrightarrow> is_aligned t tcbBlockSizeBits"
apply (drule obj_at_aligned')
apply (simp add: objBits_simps)
apply (simp add: objBits_simps)
done
lemma tcb_ptr_to_ctcb_ptr_mask [simp]:
assumes tcbat: "tcb_at' thread s"
shows "ptr_val (tcb_ptr_to_ctcb_ptr thread) && ~~ mask tcbBlockSizeBits = thread"

View File

@ -126,7 +126,6 @@ lemma det_wp_asUser [wp]:
apply simp
done
(* FIXME move into Refine somewhere *)
lemma wordSize_def':
"wordSize = 8"
unfolding wordSize_def wordBits_def

View File

@ -1387,48 +1387,6 @@ lemma ksPSpace_update_ext:
"(\<lambda>s. s\<lparr>ksPSpace := f (ksPSpace s)\<rparr>) = (ksPSpace_update f)"
by (rule ext) simp
(* FIXME: move *)
lemma valid_untyped':
notes usableUntypedRange.simps[simp del]
assumes pspace_distinct': "pspace_distinct' s" and
pspace_aligned': "pspace_aligned' s" and
al: "is_aligned ptr bits"
shows "valid_untyped' d ptr bits idx s =
(\<forall>p ko. ksPSpace s p = Some ko \<longrightarrow>
obj_range' p ko \<inter> {ptr..ptr + 2 ^ bits - 1} \<noteq> {} \<longrightarrow>
obj_range' p ko \<subseteq> {ptr..ptr + 2 ^ bits - 1} \<and>
obj_range' p ko \<inter>
usableUntypedRange (UntypedCap d ptr bits idx) = {})"
apply (simp add: valid_untyped'_def)
apply (simp add: ko_wp_at'_def)
apply (rule arg_cong[where f=All])
apply (rule ext)
apply (rule arg_cong[where f=All])
apply (rule ext)
apply (case_tac "ksPSpace s ptr' = Some ko", simp_all)
apply (frule pspace_alignedD'[OF _ pspace_aligned'])
apply (frule pspace_distinctD'[OF _ pspace_distinct'])
apply simp
apply (frule aligned_ranges_subset_or_disjoint[OF al])
apply (fold obj_range'_def)
apply (rule iffI)
apply auto[1]
apply (rule conjI)
apply (rule ccontr, simp)
apply (simp add: Set.psubset_eq)
apply (erule conjE)
apply (case_tac "obj_range' ptr' ko \<inter> {ptr..ptr + 2 ^ bits - 1} \<noteq> {}", simp)
apply (cut_tac is_aligned_no_overflow[OF al])
apply (auto simp add: obj_range'_def)[1]
apply (clarsimp simp add: usableUntypedRange.simps Int_commute)
apply (case_tac "obj_range' ptr' ko \<inter> {ptr..ptr + 2 ^ bits - 1} \<noteq> {}", simp+)
apply (cut_tac is_aligned_no_overflow[OF al])
apply (clarsimp simp add: obj_range'_def)
apply (frule is_aligned_no_overflow)
by (metis al intvl_range_conv' le_m1_iff_lt less_is_non_zero_p1
nat_le_linear power_overflow sub_wrap add_0
add_0_right word_add_increasing word_less_1 word_less_sub_1)
lemma hrs_ghost_update_comm:
"(t_hrs_'_update f \<circ> ghost'state_'_update g) =
(ghost'state_'_update g \<circ> t_hrs_'_update f)"

View File

@ -553,24 +553,6 @@ lemma hasCancelSendRights_spec:
split: capability.splits bool.splits)[1]
done
lemma updateCapData_Untyped:
"isUntypedCap a
\<Longrightarrow> updateCapData b c a = a"
by (clarsimp simp:isCap_simps updateCapData_def)
lemma ctes_of_valid_strengthen:
"(invs' s \<and> ctes_of s p = Some cte) \<longrightarrow> valid_cap' (cteCap cte) s"
apply (case_tac cte)
apply clarsimp
apply (erule ctes_of_valid_cap')
apply fastforce
done
(* FIXME move: *)
lemma hoare_vcg_imp_lift_R:
"\<lbrakk> \<lbrace>P'\<rbrace> f \<lbrace>\<lambda>rv s. \<not> P rv s\<rbrace>, -; \<lbrace>Q'\<rbrace> f \<lbrace>Q\<rbrace>, - \<rbrakk> \<Longrightarrow> \<lbrace>\<lambda>s. P' s \<or> Q' s\<rbrace> f \<lbrace>\<lambda>rv s. P rv s \<longrightarrow> Q rv s\<rbrace>, -"
by (auto simp add: valid_def validE_R_def validE_def split_def split: sum.splits)
lemma decodeCNodeInvocation_ccorres:
notes gen_invocation_type_eq[simp]
shows
@ -1683,15 +1665,6 @@ lemma pspace_no_overlap_underlying_zero_update:
apply blast
done
(* FIXME x64: check *)
lemma addrFromPPtr_mask:
"n \<le> 39
\<Longrightarrow> addrFromPPtr ptr && mask n = ptr && mask n"
apply (simp add: addrFromPPtr_def X64.pptrBase_def)
apply word_bitwise
apply simp
done
lemma clearMemory_untyped_ccorres:
"ccorres dc xfdc ((\<lambda>s. invs' s
\<and> (\<exists>cap. cte_wp_at' (\<lambda>cte. cteCap cte = cap) ut_slot s
@ -1748,56 +1721,6 @@ lemma t_hrs_update_use_t_hrs:
= (t_hrs_'_update (\<lambda>_. f (t_hrs_' s)) $ s)"
by simp
lemma name_seq_bound_helper:
"(\<not> CP n \<and> (\<forall>n' < n. CP n'))
\<Longrightarrow> (if \<exists>n. \<not> CP n
then simpl_sequence c' (map f [0 ..< (LEAST n. \<not> CP n)])
else c) = (simpl_sequence c' (map f [0 ..< n]))"
apply (simp add: exI[where x=n])
apply (subst Least_equality[where x=n], simp_all)
apply (rule ccontr, simp add: linorder_not_le)
done
lemma reset_name_seq_bound_helper:
fixes sz
fixes v :: "('a :: len) word"
defines "CP \<equiv> (\<lambda>n. ~ (v && ~~ mask sz) + of_nat n * (-1 << sz) =
((-1 :: 'a word) << sz))"
and "n \<equiv> Suc (unat (shiftR v sz))"
assumes vsz: "v + 1 < 2 ^ (len_of TYPE('a) - 1)" "2 ^ sz \<noteq> (0 :: 'a word)"
and vless: "v < v'"
shows "(\<not> CP n \<and> (\<forall>n' < n. CP n'))"
apply (clarsimp simp: shiftl_t2n field_simps less_Suc_eq_le CP_def n_def)
apply (simp add: shiftr_shiftl1[where b=sz and c=sz, simplified, symmetric]
shiftl_t2n)
apply (clarsimp simp: word_sle_msb_le shiftl_t2n[symmetric])
apply (case_tac n', simp_all)
apply (cut_tac vsz(1) order_less_le_trans[OF vless max_word_max])
apply (clarsimp simp: shiftr_shiftl1 dest!: word_add_no_overflow)
apply (drule_tac f="\<lambda>x. x - 2 ^ sz" in arg_cong, simp)
apply (metis less_irrefl order_le_less_trans order_less_trans
word_and_le2[where a=v and y="~~ mask sz"]
word_two_power_neg_ineq[OF vsz(2)])
apply (clarsimp simp add: field_simps)
apply (drule_tac f="\<lambda>x. shiftr x sz" in arg_cong)
apply simp
apply (simp add: shiftr_div_2n')
apply (simp only: linorder_not_less[symmetric], erule notE)
apply (rule order_le_less_trans[OF div_le_mono])
apply (rule_tac b="v * 2 ^ sz" for v in unat_le_helper,
simp, rule order_refl)
apply simp
done
schematic_goal sz8_helper:
"((-1) << 8 :: addr) = ?v"
by (simp add: shiftl_t2n)
lemmas reset_name_seq_bound_helper2
= reset_name_seq_bound_helper[where sz=8 and v="v :: addr" for v,
simplified sz8_helper word_bits_def[symmetric],
THEN name_seq_bound_helper]
lemma reset_untyped_inner_offs_helper:
"\<lbrakk> cteCap cte = capability.UntypedCap dev ptr sz idx;
i \<le> unat ((of_nat idx - 1 :: addr) div 2 ^ sz2);

View File

@ -689,15 +689,6 @@ lemma ccorres_pre_getQueue:
apply (simp add: maxDom_to_H maxPrio_to_H)+
done
(* FIXME: move *)
lemma threadGet_wp:
"\<lbrace>\<lambda>s. \<forall>tcb. ko_at' tcb thread s \<longrightarrow> P (f tcb) s\<rbrace> threadGet f thread \<lbrace>P\<rbrace>"
apply (rule hoare_post_imp [OF _ tg_sp'])
apply clarsimp
apply (frule obj_at_ko_at')
apply (clarsimp elim: obj_atE')
done
lemma state_relation_queue_update_helper':
"\<lbrakk> (s, s') \<in> rf_sr;
(\<forall>d p. (\<forall>t\<in>set (ksReadyQueues s (d, p)). obj_at' (inQ d p) t s)
@ -896,60 +887,6 @@ lemma cbitmap_L2_relation_update:
by (simp add: rf_sr_def cstate_relation_def Let_def carch_state_relation_def
cmachine_state_relation_def)
lemma unat_ucast_prio_shiftr_simp[simp]:
"unat (ucast (p::priority) >> n :: machine_word) = unat (p >> n)"
by (simp add: shiftr_div_2n')+
lemma unat_ucast_prio_mask_simp[simp]:
"unat (ucast (p::priority) && mask m :: machine_word) = unat (p && mask m)"
by (simp add: ucast_and_mask)
lemma unat_ucast_prio_L1_cmask_simp:
"unat (ucast (p::priority) && 0x3F :: machine_word) = unat (p && 0x3F)"
using unat_ucast_prio_mask_simp[where m=6]
by (simp add: mask_def)
lemma machine_word_and_3F_less_40:
"(w :: machine_word) && 0x3F < 0x40"
by (rule word_and_less', simp)
lemma prio_ucast_shiftr_wordRadix_helper: (* FIXME generalise *)
"(ucast (p::priority) >> wordRadix :: machine_word) < 4"
unfolding maxPriority_def numPriorities_def wordRadix_def
using unat_lt2p[where x=p]
apply (clarsimp simp add: word_less_nat_alt shiftr_div_2n' unat_ucast_upcast is_up word_le_nat_alt)
apply arith
done
lemma prio_ucast_shiftr_wordRadix_helper': (* FIXME generalise *)
"(ucast (p::priority) >> wordRadix :: machine_word) \<le> 3"
unfolding maxPriority_def numPriorities_def wordRadix_def
using unat_lt2p[where x=p]
apply (clarsimp simp add: word_less_nat_alt shiftr_div_2n' unat_ucast_upcast is_up word_le_nat_alt)
apply arith
done
lemma prio_unat_shiftr_wordRadix_helper': (* FIXME generalise *)
"unat ((p::priority) >> wordRadix) \<le> 3"
unfolding maxPriority_def numPriorities_def wordRadix_def
using unat_lt2p[where x=p]
apply (clarsimp simp add: word_less_nat_alt shiftr_div_2n' unat_ucast_upcast is_up word_le_nat_alt)
apply arith
done
lemma prio_ucast_shiftr_wordRadix_helper2: (* FIXME possibly unused *)
"(ucast (p::priority) >> wordRadix :: machine_word) < 0x20"
by (rule order_less_trans[OF prio_ucast_shiftr_wordRadix_helper]; simp)
lemma prio_ucast_shiftr_wordRadix_helper3:
"(ucast (p::priority) >> wordRadix :: machine_word) < 0x40"
by (rule order_less_trans[OF prio_ucast_shiftr_wordRadix_helper]; simp)
lemma dom_less_0x10_helper:
"d \<le> maxDomain \<Longrightarrow> (ucast d :: machine_word) < 0x10"
unfolding maxDomain_def numDomains_def
by (clarsimp simp add: word_less_nat_alt unat_ucast_upcast is_up word_le_nat_alt)
lemma cready_queues_index_to_C_ucast_helper:
fixes p :: priority
fixes d :: domain
@ -1024,30 +961,6 @@ lemma carch_state_relation_enqueue_simp:
unfolding carch_state_relation_def
by clarsimp
(* FIXME move to lib/Eisbach_Methods *)
(* FIXME consider printing error on solve goal apply *)
context
begin
private definition "bool_protect (b::bool) \<equiv> b"
lemma bool_protectD:
"bool_protect P \<Longrightarrow> P"
unfolding bool_protect_def by simp
lemma bool_protectI:
"P \<Longrightarrow> bool_protect P"
unfolding bool_protect_def by simp
(*
When you want to apply a rule/tactic to transform a potentially complex goal into another
one manually, but want to indicate that any fresh emerging goals are solved by a more
brutal method.
E.g. apply (solves_emerging \<open>frule x=... in my_rule\<close>\<open>fastforce simp: ... intro!: ... \<close> *)
method solves_emerging methods m1 m2 = (rule bool_protectD, (m1 ; (rule bool_protectI | (m2; fail))))
end
lemma t_hrs_ksReadyQueues_upd_absorb:
"t_hrs_'_update f (g s) \<lparr>ksReadyQueues_' := rqupd \<rparr> =
t_hrs_'_update f (g s \<lparr>ksReadyQueues_' := rqupd\<rparr>)"
@ -1375,10 +1288,6 @@ lemma rf_sr_drop_bitmaps_dequeue_helper:
carch_state_relation_def cmachine_state_relation_def
by (clarsimp simp: rf_sr_cbitmap_L1_relation rf_sr_cbitmap_L2_relation)
lemma filter_empty_unfiltered_contr:
"\<lbrakk> [x\<leftarrow>xs . x \<noteq> y] = [] ; x' \<in> set xs ; x' \<noteq> y \<rbrakk> \<Longrightarrow> False"
by (induct xs, auto split: if_split_asm)
(* FIXME same proofs as bit_set, maybe can generalise? *)
lemma cbitmap_L1_relation_bit_clear:
fixes p :: priority
@ -1420,16 +1329,6 @@ lemma cbitmap_L2_relationD:
unfolding cbitmap_L2_relation_def l2BitmapSize_def'
by clarsimp
(* FIXME move *)
lemma filter_noteq_op:
"[x \<leftarrow> xs . x \<noteq> y] = filter ((\<noteq>) y) xs"
by (induct xs) auto
(* FIXME move *)
lemma all_filter_propI:
"\<forall>x\<in>set lst. P x \<Longrightarrow> \<forall>x\<in>set (filter Q lst). P x"
by (induct lst, auto)
lemma cbitmap_L2_relation_bit_clear:
fixes p :: priority
fixes d :: domain

View File

@ -183,119 +183,11 @@ lemma asUser_comm:
apply (rule efa efb)+
done
(* FIXME move the thread_submonad stuff to SubMonad_R and use it for asUser *)
definition
"thread_fetch \<equiv> \<lambda>ext t s. case (ksPSpace s t) of
Some (KOTCB tcb) \<Rightarrow> ext tcb
| None \<Rightarrow> undefined"
definition
"thread_fetch_option \<equiv> \<lambda>ext t s. case (ksPSpace s t) of
Some (KOTCB tcb) \<Rightarrow> ext tcb
| None \<Rightarrow> None"
definition
"thread_replace \<equiv> \<lambda>upd t nv s.
let obj = case (ksPSpace s t) of
Some (KOTCB tcb) \<Rightarrow> Some (KOTCB (upd (\<lambda>_. nv) tcb))
| obj \<Rightarrow> obj
in s \<lparr> ksPSpace := (ksPSpace s) (t := obj) \<rparr>"
lemma thread_submonad_args:
"\<lbrakk> \<And>f v. ext (upd f v) = f (ext v);
\<And>f1 f2 v. upd f1 (upd f2 v) = upd (f1 \<circ> f2) v;
\<And>f v. upd (\<lambda>_. ext v) v = v \<rbrakk> \<Longrightarrow>
submonad_args (thread_fetch ext t) (thread_replace upd t) (tcb_at' t)"
apply unfold_locales
apply (clarsimp simp: thread_fetch_def thread_replace_def
Let_def obj_at'_def projectKOs
split: kernel_object.split option.split)
apply (clarsimp simp: thread_replace_def Let_def
split: kernel_object.split option.split)
apply (clarsimp simp: thread_fetch_def thread_replace_def Let_def
fun_upd_idem
split: kernel_object.splits option.splits)
apply (clarsimp simp: obj_at'_def thread_replace_def Let_def projectKOs
split: kernel_object.splits option.splits)
apply (rename_tac tcb)
apply (case_tac tcb, simp add: objBitsKO_def ps_clear_def)
done
lemma tcbFault_submonad_args:
"submonad_args (thread_fetch tcbFault t) (thread_replace tcbFault_update t)
(tcb_at' t)"
apply (rule thread_submonad_args)
apply (case_tac v, simp)+
done
lemma threadGet_stateAssert_gets:
"threadGet ext t = do stateAssert (tcb_at' t) []; gets (thread_fetch ext t) od"
apply (rule is_stateAssert_gets [OF _ _ empty_fail_threadGet no_fail_threadGet])
apply (clarsimp intro!: obj_at_ko_at'[where P="\<lambda>tcb :: tcb. True", simplified]
| wp threadGet_wp)+
apply (clarsimp simp: obj_at'_def thread_fetch_def projectKOs)
done
lemma threadGet_tcbFault_submonad_fn:
"threadGet tcbFault t =
submonad_fn (thread_fetch tcbFault t) (thread_replace tcbFault_update t)
(tcb_at' t) get"
apply (rule ext)
apply (clarsimp simp: submonad_fn_def bind_assoc split_def)
apply (subst threadGet_stateAssert_gets, simp)
apply (rule bind_apply_cong [OF refl])
apply (clarsimp simp: in_monad bind_def gets_def get_def return_def
submonad_args.args(3) [OF tcbFault_submonad_args]
select_f_def modify_def put_def)
done
lemma asUser_threadGet_tcbFault_comm:
"empty_fail im \<Longrightarrow>
do y \<leftarrow> asUser t im;
x \<leftarrow> threadGet tcbFault t';
n x y
od =
do x \<leftarrow> threadGet tcbFault t';
asUser t im >>= n x
od"
apply (rule submonad_comm2 [OF tcbFault_submonad_args
threadGet_tcbFault_submonad_fn
submonad_asUser, symmetric])
apply (clarsimp simp: thread_replace_def asUser_replace_def Let_def
split: option.split)
apply (clarsimp simp: fun_upd_idem fun_upd_twist
split: kernel_object.split)
apply (rename_tac tcb)
apply (case_tac tcb, simp)
apply (clarsimp simp: asUser_replace_def Let_def obj_at'_real_def
ko_wp_at'_def ps_clear_upd_None ps_clear_upd
objBitsKO_def projectKOs
split: option.split kernel_object.split)
apply (clarsimp simp: thread_replace_def Let_def obj_at'_real_def
ko_wp_at'_def ps_clear_upd_None
ps_clear_upd objBitsKO_def projectKOs
split: option.split kernel_object.split)
apply (simp add: get_def empty_fail_def)
apply assumption
done
lemma asUser_getRegister_threadGet_comm:
"do
ra \<leftarrow> asUser a (getRegister r);
rb \<leftarrow> threadGet fb b;
c ra rb
od = do
rb \<leftarrow> threadGet fb b;
ra \<leftarrow> asUser a (getRegister r);
c ra rb
od"
by (rule bind_inv_inv_comm, auto; wp)
crunch inv[wp]: getSanitiseRegisterInfo P
lemma empty_fail_getSanitiseRegisterInfo[wp, simp]:
"empty_fail (getSanitiseRegisterInfo t)"
by (wpsimp simp: getSanitiseRegisterInfo_def wp: kernel.empty_fail_archThreadGet)
by (wpsimp simp: getSanitiseRegisterInfo_def wp: ArchMove_C.empty_fail_archThreadGet)
lemma asUser_getRegister_getSanitiseRegisterInfo_comm:
"do
@ -333,98 +225,6 @@ lemma asUser_mapMloadWordUser_getSanitiseRegisterInfo_comm:
od"
by (rule bind_inv_inv_comm, auto; wp mapM_wp')
lemma threadGet_tcbFault_doMachineOp_comm:
"\<lbrakk> empty_fail m' \<rbrakk> \<Longrightarrow>
do x \<leftarrow> threadGet tcbFault t; y \<leftarrow> doMachineOp m'; n x y od =
do y \<leftarrow> doMachineOp m'; x \<leftarrow> threadGet tcbFault t; n x y od"
apply (rule submonad_comm2 [OF tcbFault_submonad_args
threadGet_tcbFault_submonad_fn
submonad_doMachineOp])
apply (simp add: thread_replace_def Let_def)
apply simp
apply (rule refl)
apply (simp add: get_def empty_fail_def)
apply assumption
done
lemma getObject_tcb_det:
"(r::tcb,s') \<in> fst (getObject p s) \<Longrightarrow> fst (getObject p s) = {(r,s)} \<and> s' = s"
apply (clarsimp simp add: getObject_def bind_def get_def gets_def
return_def loadObject_default_def split_def)
apply (clarsimp split: kernel_object.split_asm if_split_asm option.split_asm
simp: in_monad typeError_def alignError_def magnitudeCheck_def
objBits_def objBitsKO_def projectKOs
lookupAround2_def Let_def o_def)
apply (simp_all add: bind_def return_def assert_opt_def split_def projectKOs
alignCheck_def is_aligned_mask[symmetric]
unless_def when_def magnitudeCheck_def)
done
lemma threadGet_again:
"\<And>rv s s' n. (rv, s') \<in> fst (threadGet ext t s) \<Longrightarrow>
(threadGet ext t >>= n) s' = n rv s'"
apply (clarsimp simp add: threadGet_def liftM_def in_monad)
apply (frule use_valid [OF _ getObject_obj_at'])
apply (simp add: objBits_simps')+
apply (frule getObject_tcb_det)
apply (clarsimp simp: bind_def split_def)
apply (insert no_fail_getObject_tcb)
apply (clarsimp simp: no_fail_def obj_at'_def is_tcb)
done
lemma asUser_getRegister_discarded:
"(asUser t (getRegister r)) >>= (\<lambda>_. n) =
stateAssert (tcb_at' t) [] >>= (\<lambda>_. n)"
apply (rule ext)
apply (clarsimp simp: submonad_asUser.fn_is_sm submonad_fn_def
submonad_asUser.args assert_def select_f_def
gets_def get_def modify_def put_def
getRegister_def bind_def split_def
return_def fail_def stateAssert_def)
done
lemma loadWordUser_submonad_fn:
"loadWordUser p = submonad_fn ksMachineState (ksMachineState_update \<circ> K)
(pointerInUserData p) (loadWord p)"
by (simp add: loadWordUser_def submonad_doMachineOp.fn_is_sm submonad_fn_def)
lemma storeWordUser_submonad_fn:
"storeWordUser p v = submonad_fn ksMachineState (ksMachineState_update \<circ> K)
(pointerInUserData p) (storeWord p v)"
by (simp add: storeWordUser_def submonad_doMachineOp.fn_is_sm submonad_fn_def)
lemma threadGet_tcbFault_loadWordUser_comm:
"do x \<leftarrow> threadGet tcbFault t; y \<leftarrow> loadWordUser p; n x y od =
do y \<leftarrow> loadWordUser p; x \<leftarrow> threadGet tcbFault t; n x y od"
apply (rule submonad_comm [OF tcbFault_submonad_args _
threadGet_tcbFault_submonad_fn
loadWordUser_submonad_fn])
apply (simp add: submonad_args_def pointerInUserData_def)
apply (simp add: thread_replace_def Let_def)
apply simp
apply (clarsimp simp: thread_replace_def Let_def typ_at'_def ko_wp_at'_def
ps_clear_upd ps_clear_upd_None pointerInUserData_def
split: option.split kernel_object.split)
apply (simp add: get_def empty_fail_def)
apply (simp add: ef_loadWord)
done
lemma threadGet_tcbFault_storeWordUser_comm:
"do x \<leftarrow> threadGet tcbFault t; y \<leftarrow> storeWordUser p v; n x y od =
do y \<leftarrow> storeWordUser p v; x \<leftarrow> threadGet tcbFault t; n x y od"
apply (rule submonad_comm [OF tcbFault_submonad_args _
threadGet_tcbFault_submonad_fn
storeWordUser_submonad_fn])
apply (simp add: submonad_args_def pointerInUserData_def)
apply (simp add: thread_replace_def Let_def)
apply simp
apply (clarsimp simp: thread_replace_def Let_def typ_at'_def ko_wp_at'_def
ps_clear_upd ps_clear_upd_None pointerInUserData_def
split: option.split kernel_object.split)
apply (simp add: get_def empty_fail_def)
apply (simp add: ef_storeWord)
done
lemma asUser_loadWordUser_comm:
"empty_fail m \<Longrightarrow>
do x \<leftarrow> asUser t m; y \<leftarrow> loadWordUser p; n x y od =
@ -469,17 +269,6 @@ end
context kernel_m begin
(* FIXME move *)
lemma from_bool_to_bool_and_1 [simp]:
assumes r_size: "1 < size r"
shows "from_bool (to_bool (r && 1)) = r && 1"
proof -
from r_size have "r && 1 < 2"
by (simp add: and_mask_less_size [where n=1, unfolded mask_def, simplified])
thus ?thesis
by (fastforce simp add: from_bool_def to_bool_def dest: word_less_cases)
qed
(* FIXME move *)
lemma ccap_relation_ep_helpers:
"\<lbrakk> ccap_relation cap cap'; cap_get_tag cap' = scast cap_endpoint_cap \<rbrakk>
@ -1487,23 +1276,6 @@ lemma copyMRs_register_loop_helper:
rename_tac i))+
done
lemma mab_gt_2 [simp]:
"2 \<le> msg_align_bits" by (simp add: msg_align_bits)
lemma wb_gt_2:
"2 < word_bits" by (simp add: word_bits_conv)
(* FIXME move *)
lemma mapM_only_length:
"do ys \<leftarrow> mapM f xs;
g (length ys)
od =
do _ \<leftarrow> mapM_x f xs;
g (length xs)
od"
by (subst bind_return_subst [OF mapM_length])
(rule mapM_discarded)
(* FIXME move *)
lemma copyMRs_ccorres [corres]:
@ -2476,12 +2248,6 @@ lemma loadCapTransfer_ctReceiveDepth:
apply wpsimp+
done
(* FIXME: move *)
lemma cte_at_0' [dest!]:
"\<lbrakk> cte_at' 0 s; no_0_obj' s \<rbrakk> \<Longrightarrow> False"
apply (clarsimp simp: cte_wp_at_obj_cases')
by (auto simp: tcb_cte_cases_def is_aligned_def objBits_simps' dest!:tcb_aligned' split: if_split_asm)
lemma getReceiveSlots_ccorres:
"ccorres (\<lambda>a c. (a = [] \<or> (\<exists>slot. a = [slot])) \<and>
((a \<noteq> []) = (c \<noteq> NULL)) \<and> (a\<noteq>[] \<longrightarrow> c = cte_Ptr (hd a) \<and> c \<noteq> NULL))
@ -3317,32 +3083,6 @@ lemma transferCaps_ccorres [corres]:
apply clarsimp
done
(* FIXME: move *)
lemma getMessageInfo_le3:
"\<lbrace>\<top>\<rbrace> getMessageInfo sender \<lbrace>\<lambda>rv s. unat (msgExtraCaps rv) \<le> 3\<rbrace>"
including no_pre
apply (simp add: getMessageInfo_def)
apply wp
apply (rule_tac Q="\<lambda>_. \<top>" in hoare_strengthen_post)
apply wp
apply (simp add: messageInfoFromWord_def Let_def msgExtraCapBits_def)
apply (cut_tac y="r >> Types_H.msgLengthBits" in word_and_le1 [where a=3])
apply (simp add: word_le_nat_alt)
done
lemma getMessageInfo_msgLength:
"\<lbrace>\<top>\<rbrace> getMessageInfo sender \<lbrace>\<lambda>rv. K (unat (msgLength rv) \<le> msgMaxLength)\<rbrace>"
including no_pre
apply (simp add: getMessageInfo_def)
apply wp
apply (rule_tac Q="\<lambda>_. \<top>" in hoare_strengthen_post)
apply wp
apply (simp add: messageInfoFromWord_def Let_def not_less msgMaxLength_def msgLengthBits_def
split: if_split)
apply (cut_tac y="r" in word_and_le1 [where a="0x7F"])
apply (simp add: word_le_nat_alt)
done
definition
mi_from_H :: "Types_H.message_info \<Rightarrow> seL4_MessageInfo_CL"
where
@ -4364,77 +4104,15 @@ lemma handleFaultReply_ccorres [corres]:
apply (fastforce simp: seL4_Faults seL4_Arch_Faults)
done
(* FIXME: move *)
lemma cancelAllIPC_sch_act_wf:
"\<lbrace>\<lambda>s. sch_act_wf (ksSchedulerAction s) s\<rbrace>
cancelAllIPC ep
\<lbrace>\<lambda>rv s. sch_act_wf (ksSchedulerAction s) s\<rbrace>"
apply (simp add: cancelAllIPC_def)
apply (rule hoare_TrueI|wp getEndpoint_wp|wpc|simp)+
apply fastforce?
done
(* FIXME: move *)
lemma cancelAllSignals_sch_act_wf:
"\<lbrace>\<lambda>s. sch_act_wf (ksSchedulerAction s) s\<rbrace>
cancelAllSignals ep
\<lbrace>\<lambda>rv s. sch_act_wf (ksSchedulerAction s) s\<rbrace>"
apply (simp add: cancelAllSignals_def)
apply (rule hoare_TrueI|wp getNotification_wp|wpc|simp)+
apply fastforce?
done
(* FIXME: move *)
lemma cteDeleteOne_sch_act_wf:
"\<lbrace>\<lambda>s. sch_act_wf (ksSchedulerAction s) s\<rbrace>
cteDeleteOne slot
\<lbrace>\<lambda>rv s. sch_act_wf (ksSchedulerAction s) s\<rbrace>"
apply (simp add: cteDeleteOne_def unless_when split_def)
apply (simp add: finaliseCapTrue_standin_def Let_def)
apply (rule hoare_pre)
apply (wp isFinalCapability_inv cancelAllSignals_sch_act_wf
cancelAllIPC_sch_act_wf getCTE_wp' static_imp_wp
| wpc
| simp add: Let_def split: if_split)+
done
(* FIXME: move *)
lemma vp_invs_strg': "invs' s \<longrightarrow> valid_pspace' s" by auto
(* FIXME: move *)
lemma setCTE_tcbFault:
"\<lbrace>obj_at' (\<lambda>tcb. P (tcbFault tcb)) t\<rbrace>
setCTE slot cte
\<lbrace>\<lambda>rv. obj_at' (\<lambda>tcb. P (tcbFault tcb)) t\<rbrace>"
apply (simp add: setCTE_def)
apply (rule setObject_cte_obj_at_tcb', simp_all)
done
crunch tcbFault: emptySlot, tcbSchedEnqueue, rescheduleRequired
"obj_at' (\<lambda>tcb. P (tcbFault tcb)) t"
(wp: threadSet_obj_at'_strongish crunch_wps
simp: crunch_simps unless_def)
(* FIXME: move *)
lemmas threadSet_obj_at' = threadSet_obj_at'_strongish
crunch tcbFault: setThreadState, cancelAllIPC, cancelAllSignals
"obj_at' (\<lambda>tcb. P (tcbFault tcb)) t"
(wp: threadSet_obj_at'_strongish crunch_wps)
(* FIXME: move *)
lemmas setEndpoint_tcb = KHeap_R.setEndpoint_obj_at'_tcb
(* FIXME: move *)
lemma setNotification_tcb:
"\<lbrace>obj_at' (\<lambda>tcb::tcb. P tcb) t\<rbrace>
setNotification ntfn e
\<lbrace>\<lambda>_. obj_at' P t\<rbrace>"
apply (simp add: setNotification_def)
apply (rule obj_at_setObject2)
apply (clarsimp simp: updateObject_default_def in_monad)
done
lemma sbn_tcbFault:
"\<lbrace>obj_at' (\<lambda>tcb. P (tcbFault tcb)) t\<rbrace>
setBoundNotification st t'
@ -5156,34 +4834,6 @@ lemma tcbEPAppend_spec:
apply (subst fpu_state_preservation; simp add: typ_heap_simps')+
done
(* FIXME: move up to SR_lemmas_C and remove from Retype_C *)
lemma map_to_ko_atI2:
"\<lbrakk>(projectKO_opt \<circ>\<^sub>m (ksPSpace s)) x = Some v; pspace_aligned' s; pspace_distinct' s\<rbrakk> \<Longrightarrow> ko_at' v x s"
apply (clarsimp simp: map_comp_Some_iff)
apply (erule (2) aligned_distinct_obj_atI')
apply (simp add: project_inject)
done
lemma map_to_ko_at_updI':
"\<And>x x' y y' y''.
\<lbrakk> (projectKO_opt \<circ>\<^sub>m (ksPSpace s)) x = Some y;
valid_pspace' s; ko_at' y' x' s;
objBitsKO (injectKO y') = objBitsKO y''; x \<noteq> x' \<rbrakk> \<Longrightarrow>
ko_at' y x (s\<lparr>ksPSpace := ksPSpace s(x' \<mapsto> y'')\<rparr>)"
by (fastforce simp: obj_at'_def projectKOs objBitsKO_def ps_clear_upd
dest: map_to_ko_atI2)
(* FIXME: move *)
lemma state_refs_of'_upd:
"\<lbrakk> valid_pspace' s; ko_wp_at' (\<lambda>ko. objBitsKO ko = objBitsKO ko') ptr s \<rbrakk> \<Longrightarrow>
state_refs_of' (s\<lparr>ksPSpace := ksPSpace s(ptr \<mapsto> ko')\<rparr>) =
(state_refs_of' s)(ptr := refs_of' ko')"
apply (rule ext)
apply (clarsimp simp: ps_clear_upd valid_pspace'_def pspace_aligned'_def
obj_at'_def ko_wp_at'_def state_refs_of'_def
split: option.split if_split)
done
lemma sendIPC_enqueue_ccorres_helper:
"ccorres dc xfdc (valid_pspace'
and (\<lambda>s. sym_refs ((state_refs_of' s)(epptr := set queue \<times> {EPSend})))
@ -5344,11 +4994,6 @@ lemma ctcb_relation_blockingIPCCanGrantD:
thread_state_lift_def from_bool_to_bool_iff mask_eq1_nochoice)+
done
(* FIXME move *)
lemma ex_st_tcb_at'_simp[simp]:
"(\<exists>ts. st_tcb_at' ((=) ts) dest s) = tcb_at' dest s"
by (auto simp add: pred_tcb_at'_def obj_at'_def)
lemma sendIPC_ccorres [corres]:
"ccorres dc xfdc (invs' and st_tcb_at' simple' thread
and sch_act_not thread and ep_at' epptr and
@ -5470,11 +5115,7 @@ lemma sendIPC_ccorres [corres]:
apply (rule_tac ep=IdleEP in sendIPC_enqueue_ccorres_helper)
apply (simp add: valid_ep'_def)
apply (wp sts_st_tcb')
apply (rule_tac Q="\<lambda>rv. ko_wp_at' (\<lambda>x. projectKO_opt x = Some IdleEP
\<and> projectKO_opt x = (None::tcb option)) epptr"
in hoare_post_imp)
apply (clarsimp simp: obj_at'_def ko_wp_at'_def projectKOs)
apply wp
apply (clarsimp simp: obj_at'_def ko_wp_at'_def projectKOs)
apply (clarsimp simp: guard_is_UNIV_def)
apply (rule ccorres_return_Skip)
\<comment> \<open>SendEP case\<close>
@ -5500,12 +5141,7 @@ lemma sendIPC_ccorres [corres]:
apply (rule_tac ep="SendEP list" in sendIPC_enqueue_ccorres_helper)
apply (simp add: valid_ep'_def)
apply (wp sts_st_tcb')
apply (rule_tac Q="\<lambda>rv. ko_wp_at' (\<lambda>x. projectKO_opt x = Some (SendEP list)
\<and> projectKO_opt x = (None::tcb option)) epptr
and K (thread \<notin> set list)"
in hoare_post_imp)
apply (clarsimp simp: obj_at'_def ko_wp_at'_def projectKOs)
apply wp
apply (clarsimp simp: obj_at'_def ko_wp_at'_def projectKOs)
apply (clarsimp simp: guard_is_UNIV_def)
apply (rule ccorres_return_Skip)
apply (clarsimp simp: EPState_Recv_def EPState_Send_def EPState_Idle_def
@ -6043,95 +5679,64 @@ lemma receiveIPC_ccorres [corres]:
apply (simp add: cap_endpoint_cap_lift ccap_relation_def cap_to_H_def)
apply ceqv
apply csymbr
apply (rule ccorres_move_c_guard_tcb)
apply (rule_tac xf'=ntfnPtr_'
and r'="\<lambda>rv rv'. rv' = option_to_ptr rv \<and> rv \<noteq> Some 0"
in ccorres_split_nothrow_novcg)
apply (simp add: getBoundNotification_def)
apply (rule_tac P="no_0_obj' and valid_objs'" in threadGet_vcg_corres_P)
apply (rule allI, rule conseqPre, vcg)
apply clarsimp
apply (drule obj_at_ko_at', clarsimp)
apply (drule spec, drule(1) mp, clarsimp)
apply (clarsimp simp: typ_heap_simps ctcb_relation_def)
apply (drule(1) ko_at_valid_objs', simp add: projectKOs)
apply (clarsimp simp: option_to_ptr_def option_to_0_def projectKOs
valid_obj'_def valid_tcb'_def)
apply ceqv
apply (rename_tac ntfnptr ntfnptr')
apply (simp del: Collect_const split del: if_split cong: call_ignore_cong)
apply (rule ccorres_rhs_assoc2)
apply (rule_tac xf'=ret__int_'
and r'="\<lambda>rv rv'. (rv' = 0) = (ntfnptr = None \<or> \<not> isActive rv)"
in ccorres_split_nothrow_novcg)
apply wpc
apply (rule ccorres_from_vcg[where P=\<top> and P'=UNIV])
apply (rule allI, rule conseqPre, vcg)
apply (clarsimp simp: option_to_ptr_def option_to_0_def in_monad Bex_def)
apply (rule ccorres_pre_getNotification[where f=return, simplified])
apply (rule_tac P="\<lambda>s. ko_at' rv (the ntfnptr) s"
in ccorres_from_vcg[where P'=UNIV])
apply (rule ccorres_move_c_guard_tcb)
apply (rule_tac xf'=ntfnPtr_'
and r'="\<lambda>rv rv'. rv' = option_to_ptr rv \<and> rv \<noteq> Some 0"
in ccorres_split_nothrow_novcg)
apply (simp add: getBoundNotification_def)
apply (rule_tac P="no_0_obj' and valid_objs'" in threadGet_vcg_corres_P)
apply (rule allI, rule conseqPre, vcg)
apply clarsimp
apply (drule obj_at_ko_at', clarsimp)
apply (drule spec, drule(1) mp, clarsimp)
apply (clarsimp simp: typ_heap_simps ctcb_relation_def)
apply (drule(1) ko_at_valid_objs', simp add: projectKOs)
apply (clarsimp simp: option_to_ptr_def option_to_0_def projectKOs
valid_obj'_def valid_tcb'_def)
apply ceqv
apply (rename_tac ntfnptr ntfnptr')
apply (simp del: Collect_const split del: if_split cong: call_ignore_cong)
apply (rule ccorres_rhs_assoc2)
apply (rule_tac xf'=ret__int_'
and r'="\<lambda>rv rv'. (rv' = 0) = (ntfnptr = None \<or> \<not> isActive rv)"
in ccorres_split_nothrow_novcg)
apply wpc
apply (rule ccorres_from_vcg[where P=\<top> and P'=UNIV])
apply (rule allI, rule conseqPre, vcg)
apply (clarsimp simp: option_to_ptr_def option_to_0_def in_monad Bex_def)
apply (erule cmap_relationE1[OF cmap_relation_ntfn])
apply (rule ccorres_pre_getNotification[where f=return, simplified])
apply (rule_tac P="\<lambda>s. ko_at' rv (the ntfnptr) s"
in ccorres_from_vcg[where P'=UNIV])
apply (rule allI, rule conseqPre, vcg)
apply (clarsimp simp: option_to_ptr_def option_to_0_def in_monad Bex_def)
apply (erule cmap_relationE1[OF cmap_relation_ntfn])
apply (erule ko_at_projectKO_opt)
apply (clarsimp simp: typ_heap_simps cnotification_relation_def Let_def
notification_state_defs isActive_def
split: Structures_H.ntfn.split_asm Structures_H.notification.splits)
apply ceqv
apply (rule ccorres_cond[where R=\<top>])
apply (simp add: Collect_const_mem)
apply (ctac add: completeSignal_ccorres[unfolded dc_def])
apply (rule_tac xf'=ret__unsigned_longlong_'
and val="case ep of IdleEP \<Rightarrow> scast EPState_Idle
| RecvEP _ \<Rightarrow> scast EPState_Recv
| SendEP _ \<Rightarrow> scast EPState_Send"
and R="ko_at' ep (capEPPtr cap)"
in ccorres_symb_exec_r_known_rv_UNIV[where R'=UNIV])
apply (vcg, clarsimp)
apply (erule cmap_relationE1 [OF cmap_relation_ep])
apply (erule ko_at_projectKO_opt)
apply (clarsimp simp: typ_heap_simps cnotification_relation_def Let_def
notification_state_defs isActive_def
split: Structures_H.ntfn.split_asm Structures_H.notification.splits)
apply (clarsimp simp: typ_heap_simps cendpoint_relation_def Let_def
split: endpoint.split_asm)
apply ceqv
apply (rule ccorres_cond[where R=\<top>])
apply (simp add: Collect_const_mem)
apply (ctac add: completeSignal_ccorres[unfolded dc_def])
apply (rule_tac xf'=ret__unsigned_longlong_'
and val="case ep of IdleEP \<Rightarrow> scast EPState_Idle
| RecvEP _ \<Rightarrow> scast EPState_Recv
| SendEP _ \<Rightarrow> scast EPState_Send"
and R="ko_at' ep (capEPPtr cap)"
in ccorres_symb_exec_r_known_rv_UNIV[where R'=UNIV])
apply (vcg, clarsimp)
apply (erule cmap_relationE1 [OF cmap_relation_ep])
apply (erule ko_at_projectKO_opt)
apply (clarsimp simp: typ_heap_simps cendpoint_relation_def Let_def
split: endpoint.split_asm)
apply ceqv
apply (rule_tac A="invs' and st_tcb_at' simple' thread
and sch_act_not thread
and (\<lambda>s. \<forall>d p. thread \<notin> set (ksReadyQueues s (d, p)))
and ko_at' ep (capEPPtr cap)"
in ccorres_guard_imp2 [where A'=UNIV])
apply (rule_tac A="invs' and st_tcb_at' simple' thread
and sch_act_not thread
and (\<lambda>s. \<forall>d p. thread \<notin> set (ksReadyQueues s (d, p)))
and ko_at' ep (capEPPtr cap)"
in ccorres_guard_imp2 [where A'=UNIV])
apply wpc
\<comment> \<open>RecvEP case\<close>
apply (rule ccorres_cond_true)
apply csymbr
apply (simp only: case_bool_If from_bool_neq_0)
apply (rule ccorres_Cond_rhs, simp cong: Collect_cong split del: if_split)
apply (intro ccorres_rhs_assoc)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_split_nothrow_novcg)
apply (simp split del: if_split)
apply (rule receiveIPC_block_ccorres_helper[unfolded ptr_val_def, simplified])
apply ceqv
apply simp
apply (rename_tac list NOo)
apply (rule_tac ep="RecvEP list"
in receiveIPC_enqueue_ccorres_helper[simplified, unfolded dc_def])
apply (simp add: valid_ep'_def)
apply (wp sts_st_tcb')
apply (rename_tac list)
apply (rule_tac Q="\<lambda>rv. ko_wp_at' (\<lambda>x. projectKO_opt x = Some (RecvEP list)
\<and> projectKO_opt x = (None::tcb option))
(capEPPtr cap)
and K (thread \<notin> set list)"
in hoare_post_imp)
apply (clarsimp simp: obj_at'_def ko_wp_at'_def projectKOs)
apply wp
apply (clarsimp simp: guard_is_UNIV_def)
apply simp
apply (ctac add: doNBRecvFailedTransfer_ccorres[unfolded dc_def])
\<comment> \<open>IdleEP case\<close>
\<comment> \<open>RecvEP case\<close>
apply (rule ccorres_cond_true)
apply csymbr
apply (simp only: case_bool_If from_bool_neq_0)
@ -6146,47 +5751,81 @@ lemma receiveIPC_ccorres [corres]:
apply (rule receiveIPC_block_ccorres_helper[unfolded ptr_val_def, simplified])
apply ceqv
apply simp
apply (rule_tac ep=IdleEP
in receiveIPC_enqueue_ccorres_helper[simplified, unfolded dc_def])
apply (rename_tac list NOo)
apply (rule_tac ep="RecvEP list"
in receiveIPC_enqueue_ccorres_helper[simplified, unfolded dc_def])
apply (simp add: valid_ep'_def)
apply (wp sts_st_tcb')
apply (rule_tac Q="\<lambda>rv. ko_wp_at' (\<lambda>x. projectKO_opt x = Some IdleEP
\<and> projectKO_opt x = (None::tcb option))
(capEPPtr cap)"
in hoare_post_imp)
apply (clarsimp simp: obj_at'_def ko_wp_at'_def projectKOs)
apply wp
apply (rename_tac list)
apply (clarsimp simp: obj_at'_def ko_wp_at'_def projectKOs)
apply (clarsimp simp: guard_is_UNIV_def)
apply simp
apply (ctac add: doNBRecvFailedTransfer_ccorres[unfolded dc_def])
\<comment> \<open>IdleEP case\<close>
apply (rule ccorres_cond_true)
apply csymbr
apply (simp only: case_bool_If from_bool_neq_0)
apply (rule ccorres_Cond_rhs, simp cong: Collect_cong split del: if_split)
apply (intro ccorres_rhs_assoc)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_split_nothrow_novcg)
apply (simp split del: if_split)
apply (rule receiveIPC_block_ccorres_helper[unfolded ptr_val_def, simplified])
apply ceqv
apply simp
apply (rule_tac ep=IdleEP
in receiveIPC_enqueue_ccorres_helper[simplified, unfolded dc_def])
apply (simp add: valid_ep'_def)
apply (wp sts_st_tcb')
apply (clarsimp simp: obj_at'_def ko_wp_at'_def projectKOs)
apply (clarsimp simp: guard_is_UNIV_def)
apply simp
apply (ctac add: doNBRecvFailedTransfer_ccorres[unfolded dc_def])
\<comment> \<open>SendEP case\<close>
apply (thin_tac "isBlockinga = from_bool P" for P)
apply (rule ccorres_cond_false)
apply (rule ccorres_cond_true)
apply (intro ccorres_rhs_assoc)
apply (csymbr, csymbr, csymbr, csymbr, csymbr)
apply wpc
apply (simp only: haskell_fail_def)
apply (rule ccorres_fail)
apply (rename_tac sender rest)
apply csymbr
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_split_nothrow_novcg)
apply simp
apply (rule_tac sender=sender in receiveIPC_dequeue_ccorres_helper[simplified])
apply (thin_tac "isBlockinga = from_bool P" for P)
apply (rule ccorres_cond_false)
apply (rule ccorres_cond_true)
apply (intro ccorres_rhs_assoc)
apply (csymbr, csymbr, csymbr, csymbr, csymbr)
apply wpc
apply (simp only: haskell_fail_def)
apply (rule ccorres_fail)
apply (rename_tac sender rest)
apply csymbr
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_split_nothrow_novcg)
apply simp
apply (rule_tac sender=sender in receiveIPC_dequeue_ccorres_helper[simplified])
apply ceqv
apply (rename_tac sender')
apply (simp only: K_bind_def haskell_assert_def return_bind)
apply (rule ccorres_move_c_guard_tcb)
apply (rule getThreadState_ccorres_foo)
apply (rename_tac sendState)
apply (rule ccorres_assert)
apply (rule ccorres_rhs_assoc2)
apply (rule_tac val="blockingIPCBadge sendState"
and xf'=badge_'
and R="\<lambda>s. \<exists>t. ko_at' t sender s \<and> tcbState t = sendState"
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 (clarsimp simp: ctcb_relation_def typ_heap_simps
cthread_state_relation_def word_size
isSend_def thread_state_lift_def
split: Structures_H.thread_state.splits)
apply ceqv
apply (rename_tac sender')
apply (simp only: K_bind_def haskell_assert_def return_bind)
apply (simp split del: if_split)
apply (rule ccorres_move_c_guard_tcb)
apply (rule getThreadState_ccorres_foo)
apply (rename_tac sendState)
apply (rule ccorres_assert)
apply (rule ccorres_rhs_assoc2)
apply (rule_tac val="blockingIPCBadge sendState"
and xf'=badge_'
apply (rule_tac val="from_bool (blockingIPCCanGrant sendState)"
and xf'=canGrant_'
and R="\<lambda>s. \<exists>t. ko_at' t sender s \<and> tcbState t = sendState"
in ccorres_symb_exec_r_known_rv_UNIV [where R'=UNIV])
apply (vcg, clarsimp)
@ -6196,26 +5835,12 @@ lemma receiveIPC_ccorres [corres]:
isSend_def thread_state_lift_def
split: Structures_H.thread_state.splits)
apply ceqv
apply (simp split del: if_split)
apply (rule ccorres_move_c_guard_tcb)
apply (rule ccorres_rhs_assoc2)
apply (rule_tac val="from_bool (blockingIPCCanGrant sendState)"
and xf'=canGrant_'
and R="\<lambda>s. \<exists>t. ko_at' t sender s \<and> tcbState t = sendState"
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 (clarsimp simp: ctcb_relation_def typ_heap_simps
cthread_state_relation_def word_size
isSend_def thread_state_lift_def
split: Structures_H.thread_state.splits)
apply ceqv
apply (rule ccorres_rhs_assoc2)
apply (rule_tac xf'=canGrantReply_'
and val="from_bool (blockingIPCCanGrantReply sendState)"
and R="st_tcb_at' ((=) sendState) sender" and R'=UNIV
in ccorres_symb_exec_r_known_rv_UNIV)
in ccorres_symb_exec_r_known_rv_UNIV)
apply (clarsimp, rule conseqPre, vcg)
apply (fastforce simp: pred_tcb_at'_def tcb_at_h_t_valid typ_heap_simps'
dest: obj_at_cslift_tcb ctcb_relation_blockingIPCCanGrantReplyD)
@ -6268,73 +5893,73 @@ lemma receiveIPC_ccorres [corres]:
apply (clarsimp simp: valid_pspace_valid_objs' pred_tcb_at'_def sch_act_wf_weak
obj_at'_def)
apply (wpsimp simp: guard_is_UNIV_def option_to_ptr_def option_to_0_def conj_ac)+
apply (rule_tac Q="\<lambda>rv. valid_queues and valid_pspace'
and cur_tcb' and tcb_at' sender and tcb_at' thread
and sch_act_not sender and K (thread \<noteq> sender)
and ep_at' (capEPPtr cap)
and (\<lambda>s. ksCurDomain s \<le> maxDomain)
and (\<lambda>s. sch_act_wf (ksSchedulerAction s) s \<and>
(\<forall>d p. sender \<notin> set (ksReadyQueues s (d, p))))"
in hoare_post_imp)
subgoal by (auto, auto simp: st_tcb_at'_def obj_at'_def)
apply (wp hoare_vcg_all_lift set_ep_valid_objs')
apply (clarsimp simp: guard_is_UNIV_def)
apply (clarsimp simp: EPState_Recv_def EPState_Send_def EPState_Idle_def)
apply (frule(1) ko_at_valid_objs' [OF _ invs_valid_objs'])
apply (clarsimp simp: projectKO_opt_ep split: kernel_object.split_asm)
apply (subgoal_tac "(capEPPtr cap) \<noteq> thread
\<and> state_refs_of' s thread = {r \<in> state_refs_of' s thread. snd r = TCBBound}")
apply (clarsimp simp: valid_obj'_def valid_ep'_def refs_of'_def
split: endpoint.splits)
apply (rename_tac list)
apply (subgoal_tac "state_refs_of' s (capEPPtr cap) = (set list) \<times> {EPRecv}
\<and> thread \<notin> (set list)")
subgoal by (fastforce simp: obj_at'_def is_aligned_neg_mask objBitsKO_def
projectKOs invs'_def valid_state'_def st_tcb_at'_def
valid_tcb_state'_def ko_wp_at'_def invs_valid_objs'
isBlockedOnReceive_def projectKO_opt_tcb
from_bool_def to_bool_def objBits_simps'
elim!: delta_sym_refs
split: if_split_asm bool.splits) (*very long*)
apply (frule(1) sym_refs_obj_atD' [OF _ invs_sym'])
apply (clarsimp simp: st_tcb_at'_def ko_wp_at'_def obj_at'_def projectKOs
split: if_split_asm)
apply (drule(1) bspec)+
apply (case_tac "tcbState obj", simp_all add: tcb_bound_refs'_def)[1]
apply (subgoal_tac "state_refs_of' s (capEPPtr cap) = {}")
apply (rule_tac Q="\<lambda>rv. valid_queues and valid_pspace'
and cur_tcb' and tcb_at' sender and tcb_at' thread
and sch_act_not sender and K (thread \<noteq> sender)
and ep_at' (capEPPtr cap)
and (\<lambda>s. ksCurDomain s \<le> maxDomain)
and (\<lambda>s. sch_act_wf (ksSchedulerAction s) s \<and>
(\<forall>d p. sender \<notin> set (ksReadyQueues s (d, p))))"
in hoare_post_imp)
subgoal by (auto, auto simp: st_tcb_at'_def obj_at'_def)
apply (wp hoare_vcg_all_lift set_ep_valid_objs')
apply (clarsimp simp: guard_is_UNIV_def)
apply (clarsimp simp: EPState_Recv_def EPState_Send_def EPState_Idle_def)
apply (frule(1) ko_at_valid_objs' [OF _ invs_valid_objs'])
apply (clarsimp simp: projectKO_opt_ep split: kernel_object.split_asm)
apply (subgoal_tac "(capEPPtr cap) \<noteq> thread
\<and> state_refs_of' s thread = {r \<in> state_refs_of' s thread. snd r = TCBBound}")
apply (clarsimp simp: valid_obj'_def valid_ep'_def refs_of'_def
split: endpoint.splits)
apply (rename_tac list)
apply (subgoal_tac "state_refs_of' s (capEPPtr cap) = (set list) \<times> {EPRecv}
\<and> thread \<notin> (set list)")
subgoal by (fastforce simp: obj_at'_def is_aligned_neg_mask objBitsKO_def
projectKOs invs'_def valid_state'_def st_tcb_at'_def
valid_tcb_state'_def ko_wp_at'_def invs_valid_objs'
isBlockedOnReceive_def projectKO_opt_tcb objBits_simps'
from_bool_def to_bool_def
elim: delta_sym_refs
split: if_split_asm bool.splits) (*very long *)
apply (clarsimp simp: obj_at'_def state_refs_of'_def projectKOs)
apply (frule(1) sym_refs_ko_atD' [OF _ invs_sym'])
apply (frule invs_queues)
apply clarsimp
apply (rename_tac list x xa)
apply (rule_tac P="x\<in>set list" in case_split)
apply (clarsimp simp:st_tcb_at_refs_of_rev')
apply (erule_tac x=x and P="\<lambda>x. st_tcb_at' P x s" for P in ballE)
apply (drule_tac t=x in valid_queues_not_runnable'_not_ksQ)
apply (clarsimp simp: st_tcb_at'_def obj_at'_def o_def)
apply (subgoal_tac "sch_act_not x s")
prefer 2
apply (frule invs_sch_act_wf')
apply (clarsimp simp:sch_act_wf_def)
apply (clarsimp simp: st_tcb_at'_def obj_at'_def o_def)
apply (clarsimp simp: obj_at'_def st_tcb_at'_def
projectKOs isBlockedOnSend_def
split: list.split | rule conjI)+
apply (clarsimp simp: state_refs_of'_def )
apply (case_tac "tcbState obj", clarsimp+)[1]
apply (clarsimp simp: guard_is_UNIV_def)+
apply (wp getNotification_wp | wpc)+
apply (clarsimp simp add: guard_is_UNIV_def option_to_ptr_def option_to_0_def
split: if_split)
apply (wp gbn_wp' | simp add: guard_is_UNIV_def)+
apply (auto simp: isCap_simps valid_cap'_def)
projectKOs invs'_def valid_state'_def st_tcb_at'_def
valid_tcb_state'_def ko_wp_at'_def invs_valid_objs'
isBlockedOnReceive_def projectKO_opt_tcb
from_bool_def to_bool_def objBits_simps'
elim!: delta_sym_refs
split: if_split_asm bool.splits) (*very long*)
apply (frule(1) sym_refs_obj_atD' [OF _ invs_sym'])
apply (clarsimp simp: st_tcb_at'_def ko_wp_at'_def obj_at'_def projectKOs
split: if_split_asm)
apply (drule(1) bspec)+
apply (case_tac "tcbState obj", simp_all add: tcb_bound_refs'_def)[1]
apply (subgoal_tac "state_refs_of' s (capEPPtr cap) = {}")
subgoal by (fastforce simp: obj_at'_def is_aligned_neg_mask objBitsKO_def
projectKOs invs'_def valid_state'_def st_tcb_at'_def
valid_tcb_state'_def ko_wp_at'_def invs_valid_objs'
isBlockedOnReceive_def projectKO_opt_tcb objBits_simps'
from_bool_def to_bool_def
elim: delta_sym_refs
split: if_split_asm bool.splits) (*very long *)
apply (clarsimp simp: obj_at'_def state_refs_of'_def projectKOs)
apply (frule(1) sym_refs_ko_atD' [OF _ invs_sym'])
apply (frule invs_queues)
apply clarsimp
apply (rename_tac list x xa)
apply (rule_tac P="x\<in>set list" in case_split)
apply (clarsimp simp:st_tcb_at_refs_of_rev')
apply (erule_tac x=x and P="\<lambda>x. st_tcb_at' P x s" for P in ballE)
apply (drule_tac t=x in valid_queues_not_runnable'_not_ksQ)
apply (clarsimp simp: st_tcb_at'_def obj_at'_def o_def)
apply (subgoal_tac "sch_act_not x s")
prefer 2
apply (frule invs_sch_act_wf')
apply (clarsimp simp:sch_act_wf_def)
apply (clarsimp simp: st_tcb_at'_def obj_at'_def o_def)
apply (clarsimp simp: obj_at'_def st_tcb_at'_def
projectKOs isBlockedOnSend_def
split: list.split | rule conjI)+
apply (clarsimp simp: state_refs_of'_def )
apply (case_tac "tcbState obj", clarsimp+)[1]
apply (clarsimp simp: guard_is_UNIV_def)+
apply (wp getNotification_wp | wpc)+
apply (clarsimp simp add: guard_is_UNIV_def option_to_ptr_def option_to_0_def
split: if_split)
apply (wp gbn_wp' | simp add: guard_is_UNIV_def)+
apply (auto simp: isCap_simps valid_cap'_def)
done
lemma sendSignal_dequeue_ccorres_helper:

View File

@ -123,73 +123,8 @@ lemma isolate_thread_actions_bind:
apply (simp add: exec_gets exec_modify)
done
lemmas setEndpoint_obj_at_tcb' = setEndpoint_obj_at'_tcb
lemma tcbSchedEnqueue_obj_at_unchangedT:
assumes y: "\<And>f. \<forall>tcb. P (tcbQueued_update f tcb) = P tcb"
shows "\<lbrace>obj_at' P t\<rbrace> tcbSchedEnqueue t' \<lbrace>\<lambda>rv. obj_at' P t\<rbrace>"
apply (simp add: tcbSchedEnqueue_def unless_def)
apply (wp | simp add: y)+
done
lemma asUser_obj_at_notQ:
"\<lbrace>obj_at' (Not \<circ> tcbQueued) t\<rbrace>
asUser t (setRegister r v)
\<lbrace>\<lambda>rv. obj_at' (Not \<circ> tcbQueued) t\<rbrace>"
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:
"\<lbrace>obj_at' (Not \<circ> tcbQueued) t\<rbrace>
Arch.switchToThread t
\<lbrace>\<lambda>rv. obj_at' (Not \<circ> tcbQueued) t\<rbrace>"
apply (simp add: X64_H.switchToThread_def)
apply (wp asUser_obj_at_notQ doMachineOp_obj_at hoare_drop_imps|wpc)+
done
lemma rescheduleRequired_obj_at_unchangedT:
assumes y: "\<And>f. \<forall>tcb. P (tcbQueued_update f tcb) = P tcb"
shows "\<lbrace>obj_at' P t\<rbrace> rescheduleRequired \<lbrace>\<lambda>rv. obj_at' P t\<rbrace>"
apply (simp add: rescheduleRequired_def)
apply (wp tcbSchedEnqueue_obj_at_unchangedT[OF y] | wpc)+
apply simp
done
lemma setThreadState_obj_at_unchangedT:
assumes x: "\<And>f. \<forall>tcb. P (tcbState_update f tcb) = P tcb"
assumes y: "\<And>f. \<forall>tcb. P (tcbQueued_update f tcb) = P tcb"
shows "\<lbrace>obj_at' P t\<rbrace> setThreadState t' ts \<lbrace>\<lambda>rv. obj_at' P t\<rbrace>"
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: "\<And>f. \<forall>tcb. P (tcbBoundNotification_update f tcb) = P tcb"
shows "\<lbrace>obj_at' P t\<rbrace> setBoundNotification t' ts \<lbrace>\<lambda>rv. obj_at' P t\<rbrace>"
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]
lemmas setNotification_tcb = set_ntfn_tcb_obj_at'
(* FIXME: move *)
lemmas threadSet_obj_at' = threadSet_obj_at'_strongish
end
context kernel_m begin
@ -444,14 +379,6 @@ lemma objBits_2n:
by (simp add: objBits_def objBitsKO_def archObjSize_def pageBits_def objBits_simps'
split: kernel_object.split arch_kernel_object.split)
lemma magnitudeCheck_assert2:
"\<lbrakk> is_aligned x n; (1 :: machine_word) < 2 ^ n; ksPSpace s x = Some v \<rbrakk> \<Longrightarrow>
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: "\<And>a b c d. (loadObject a b c d :: ('a :: pspace_storable) kernel)
= loadObject_default a b c d"
@ -819,12 +746,6 @@ lemma copyMRs_simple:
apply (simp add: upto_enum_def mapM_Nil)
done
(* FIXME: MOVE *)
lemma returnOK_catch[simp]:
"(returnOk rv <catch> m) = return rv"
unfolding catch_def returnOk_def
by clarsimp
lemma doIPCTransfer_simple_rewrite:
"monadic_rewrite True True
((\<lambda>_. msgExtraCaps (messageInfoFromWord msgInfo) = 0

View File

@ -546,16 +546,6 @@ next
done
qed
(* FIXME: also in VSpace_C *)
lemma ignoreFailure_liftM:
"ignoreFailure = liftM (\<lambda>v. ())"
apply (rule ext)+
apply (simp add: ignoreFailure_def liftM_def
catch_def)
apply (rule bind_apply_cong[OF refl])
apply (simp split: sum.split)
done
end
lemma option_to_0_user_mem':
@ -843,19 +833,6 @@ lemma cpspace_relation_ep_update_ep2:
end
context begin interpretation Arch . (*FIXME: arch_split*)
lemma setObject_tcb_ep_obj_at'[wp]:
"\<lbrace>obj_at' (P :: endpoint \<Rightarrow> bool) ptr\<rbrace> setObject ptr' (tcb :: tcb) \<lbrace>\<lambda>rv. obj_at' P ptr\<rbrace>"
apply (rule obj_at_setObject2, simp_all)
apply (clarsimp simp: updateObject_default_def in_monad)
done
end
crunch ep_obj_at'[wp]: setThreadState "obj_at' (P :: endpoint \<Rightarrow> bool) ptr"
(simp: unless_def)
crunch pspace_canonical'[wp]: setThreadState pspace_canonical'
context kernel_m begin
lemma ccorres_abstract_h_val:

View File

@ -602,12 +602,6 @@ lemma ccorres_get_registers:
"StrictC'_register_defs")
done
(* FIXME: move *)
lemma st_tcb_at'_opeq_simp:
"st_tcb_at' ((=) Structures_H.thread_state.Running) (ksCurThread s) s
= st_tcb_at' (\<lambda>st. st = Structures_H.thread_state.Running) (ksCurThread s) s"
by (fastforce simp add: st_tcb_at'_def obj_at'_def)
(* FIXME: fastpath
lemma callKernel_withFastpath_corres_C:
"corres_underlying rf_sr False True dc
@ -769,26 +763,6 @@ lemma full_invs_both:
done
end
(* FIXME: move to somewhere sensible *)
lemma dom_eq:
"dom um = dom um' \<longleftrightarrow> (\<forall>a. um a = None \<longleftrightarrow> um' a = None)"
apply (simp add: dom_def del: not_None_eq)
apply (rule iffI)
apply (rule allI)
apply (simp add: set_eq_iff)
apply (drule_tac x=a in spec)
apply auto
done
lemma dom_user_mem':
"dom (user_mem' s) = {p. typ_at' UserDataT (p && ~~ mask pageBits) s}"
by (clarsimp simp:user_mem'_def dom_def pointerInUserData_def split:if_splits)
(* FIXME:move *)
lemma dom_device_mem':
"dom (device_mem' s) = {p. typ_at' UserDataDeviceT (p && ~~ mask pageBits) s}"
by (clarsimp simp: device_mem'_def dom_def pointerInDeviceData_def split: if_splits)
context kernel_m
begin

View File

@ -3477,13 +3477,6 @@ lemma tcb_queue_update_other':
unfolding tcb_queue_relation'_def
by (simp add: tcb_queue_update_other)
lemma map_to_ko_atI2:
"\<lbrakk>(projectKO_opt \<circ>\<^sub>m (ksPSpace s)) x = Some v; pspace_aligned' s; pspace_distinct' s\<rbrakk> \<Longrightarrow> ko_at' v x s"
apply (clarsimp simp: map_comp_Some_iff)
apply (erule (2) aligned_distinct_obj_atI')
apply (simp add: project_inject)
done
lemma c_guard_tcb:
assumes al: "is_aligned (ctcb_ptr_to_tcb_ptr p) tcbBlockSizeBits"
and ptr0: "ctcb_ptr_to_tcb_ptr p \<noteq> 0"
@ -8105,30 +8098,6 @@ end
context kernel_m begin
(* FIXME: move *)
lemma nat_le_induct [consumes 1, case_names base step]:
assumes le: "i \<le> (k::nat)" and
base: "P(k)" and
step: "\<And>i. \<lbrakk>i \<le> k; P i; 0 < i\<rbrakk> \<Longrightarrow> P(i - 1)"
shows "P i"
proof -
obtain j where jk: "j \<le> k" and j_eq: "i = k - j"
using le
apply (drule_tac x="k - i" in meta_spec)
apply simp
done
have "j \<le> k \<Longrightarrow> P (k - j)"
apply (induct j)
apply (simp add: base)
apply simp
apply (drule step[rotated], simp+)
done
thus "P i" using jk j_eq
by simp
qed
lemma ceqv_restore_as_guard:
"ceqv Gamma xf' rv' t t' d (Guard C_Guard {s. xf' s = rv'} d)"
apply (simp add: ceqv_def)
@ -8440,42 +8409,6 @@ lemma range_cover_n_le':
apply (rule mult_le_mono1, rule one_le_power, simp)
done
(* FIXME: move to AInvs *)
context
fixes ptr sz us n p
assumes cover: "range_cover ptr sz us n"
assumes p: "p < n"
begin
lemma range_cover_mask_offset_bound:
"(ptr && mask sz) + (of_nat p << us) < 2 ^ sz"
proof -
note sz = range_cover.sz[OF cover]
note al = range_cover.aligned[OF cover]
have 1: "unat (ptr && mask sz >> us) + p < 2 ^ (sz - us)"
using sz(3) p by simp
have 2: "(ptr && mask sz >> us) + of_nat p < 2 ^ (sz - us)"
using of_nat_power[OF 1 less_imp_diff_less, OF sz(1)]
of_nat_add word_unat.Rep_inverse
by simp
have 3: "ptr && mask sz >> us << us = ptr && mask sz"
by (rule is_aligned_shiftr_shiftl[OF is_aligned_after_mask[OF al sz(2)]])
have 4: "((ptr && mask sz >> us) + of_nat p) << us < 2 ^ sz"
by (rule shiftl_less_t2n[OF 2 sz(1)])
show ?thesis
by (rule 4[simplified 3 word_shiftl_add_distrib])
qed
lemma range_cover_neg_mask_offset:
"ptr + (of_nat p << us) && ~~ mask sz = ptr && ~~ mask sz"
apply (subst AND_NOT_mask_plus_AND_mask_eq[of ptr sz, symmetric], subst add.assoc)
apply (rule is_aligned_add_helper[THEN conjunct2])
apply (rule Aligned.is_aligned_neg_mask[OF order_refl])
apply (rule range_cover_mask_offset_bound)
done
end
lemma createNewObjects_ccorres:
notes blah[simp del] = atLeastAtMost_iff atLeastatMost_subset_iff atLeastLessThan_iff
Int_atLeastAtMost atLeastatMost_empty_iff split_paired_Ex

View File

@ -10,18 +10,8 @@ imports
"Refine.Invariants_H"
begin
(* FIXME: move to Word_Lib *)
lemma sign_extend_0[simp]:
"sign_extend a 0 = 0"
by (simp add: sign_extend_def)
context begin interpretation Arch . (*FIXME: arch_split*)
(* FIXME: move to ainvs? *)
lemma sign_extend_canonical_address:
"(x = sign_extend 47 x) = canonical_address x"
by (fastforce simp: sign_extended_iff_sign_extend canonical_address_sign_extended)
section "ctes"
subsection "capabilities"
@ -998,7 +988,7 @@ lemma nullCapPointers_def:
lemma valid_mdb_ctes_of_next:
"\<lbrakk> valid_mdb' s; ctes_of s p = Some cte; mdbNext (cteMDBNode cte) \<noteq> 0 \<rbrakk> \<Longrightarrow> cte_at' (mdbNext (cteMDBNode cte)) s"
unfolding valid_mdb'_def valid_mdb_ctes_def
apply clarsimp
apply (erule conjE)
apply (erule (2) valid_dlistE)
apply (simp add: cte_wp_at_ctes_of)
done
@ -1006,7 +996,7 @@ lemma valid_mdb_ctes_of_next:
lemma valid_mdb_ctes_of_prev:
"\<lbrakk> valid_mdb' s; ctes_of s p = Some cte; mdbPrev (cteMDBNode cte) \<noteq> 0 \<rbrakk> \<Longrightarrow> cte_at' (mdbPrev (cteMDBNode cte)) s"
unfolding valid_mdb'_def valid_mdb_ctes_def
apply clarsimp
apply (erule conjE)
apply (erule (2) valid_dlistE)
apply (simp add: cte_wp_at_ctes_of)
done
@ -1551,15 +1541,6 @@ lemma ctcb_relation_null_queue_ptrs:
apply (simp add: ctcb_relation_def tcb_null_queue_ptrs_def)
done
lemma map_to_ko_atI':
"\<lbrakk>(projectKO_opt \<circ>\<^sub>m (ksPSpace s)) x = Some v; invs' s\<rbrakk> \<Longrightarrow> ko_at' v x s"
apply (clarsimp simp: map_comp_Some_iff)
apply (erule aligned_distinct_obj_atI')
apply clarsimp
apply clarsimp
apply (simp add: project_inject)
done
(* FIXME x64: do we still need these?
(* Levity: added (20090419 09:44:27) *)
declare ntfnQueue_head_mask_4 [simp]
@ -1889,22 +1870,6 @@ lemma cmap_relation_updI2:
using cr cof cc inj
by (clarsimp simp add: cmap_relation_def inj_eq split: if_split)
definition
user_word_at :: "machine_word \<Rightarrow> machine_word \<Rightarrow> kernel_state \<Rightarrow> bool"
where
"user_word_at x p \<equiv> \<lambda>s. is_aligned p 3
\<and> pointerInUserData p s
\<and> x = word_rcat (map (underlying_memory (ksMachineState s))
[p + 7, p + 6, p + 5, p + 4, p + 3, p + 2, p + 1, p])"
definition
device_word_at :: "machine_word \<Rightarrow> machine_word \<Rightarrow> kernel_state \<Rightarrow> bool"
where
"device_word_at x p \<equiv> \<lambda>s. is_aligned p 3
\<and> pointerInDeviceData p s
\<and> x = word_rcat (map (underlying_memory (ksMachineState s))
[p + 7, p + 6, p + 5, p + 4, p + 3, p + 2, p + 1, p])"
lemma rf_sr_heap_user_data_relation:
"(s, s') \<in> rf_sr \<Longrightarrow> cmap_relation
(heap_to_user_data (ksPSpace s) (underlying_memory (ksMachineState s)))
@ -1921,21 +1886,6 @@ lemma rf_sr_heap_device_data_relation:
cstate_relation_def Let_def
cpspace_relation_def)
lemma ko_at_projectKO_opt:
"ko_at' ko p s \<Longrightarrow> (projectKO_opt \<circ>\<^sub>m ksPSpace s) p = Some ko"
by (clarsimp elim!: obj_atE' simp: projectKOs)
lemma word_shift_by_3:
"x * 8 = (x::'a::len word) << 3"
by (simp add: shiftl_t2n)
(* FIXME x64: move *)
lemma mask_shiftr_times_simp:
"\<lbrakk>x > n;is_aligned p n\<rbrakk> \<Longrightarrow> (p && mask x >> n) * (2^n) = (p && mask x)"
apply (subst mult.commute)
apply (simp add: shiftl_t2n[symmetric])
by (simp add: is_aligned_andI1 is_aligned_shiftr_shiftl)
lemma user_word_at_cross_over:
"\<lbrakk> user_word_at x p s; (s, s') \<in> rf_sr; p' = Ptr p \<rbrakk>
\<Longrightarrow> c_guard p' \<and> hrs_htd (t_hrs_' (globals s')) \<Turnstile>\<^sub>t p'
@ -2061,16 +2011,6 @@ lemma device_word_at_cross_over:
done
*)
(* FIXME: move to GenericLib *)
lemmas unat64_eq_of_nat = unat_eq_of_nat[where 'a=64, folded word_bits_def]
lemma unat_mask_3_less_8:
"unat (p && mask 3 :: word64) < 8"
apply (rule unat_less_helper)
apply (rule order_le_less_trans, rule word_and_le1)
apply (simp add: mask_def)
done
lemma memory_cross_over:
"\<lbrakk>(\<sigma>, s) \<in> rf_sr; pspace_aligned' \<sigma>; pspace_distinct' \<sigma>;
pointerInUserData ptr \<sigma>\<rbrakk>

View File

@ -8,35 +8,6 @@ theory Schedule_C
imports Tcb_C
begin
context begin interpretation Arch . (*FIXME: arch_split*)
(* FIXME move to REFINE *)
crunches Arch.switchToThread
for valid_queues'[wp]: valid_queues'
crunches switchToIdleThread
for ksCurDomain[wp]: "\<lambda>s. P (ksCurDomain s)"
crunches switchToIdleThread, switchToThread
for valid_pspace'[wp]: valid_pspace'
(simp: whenE_def)
lemma setCurrentUserCR3_valid_arch_state'[wp]:
"\<lbrace>valid_arch_state' and K (valid_cr3' c)\<rbrace> setCurrentUserCR3 c \<lbrace>\<lambda>_. valid_arch_state'\<rbrace>"
by (wpsimp simp: setCurrentUserCR3_def valid_arch_state'_def)
lemma setVMRoot_valid_arch_state':
"\<lbrace>valid_arch_state'\<rbrace> setVMRoot t \<lbrace>\<lambda>_. valid_arch_state'\<rbrace>"
apply (simp add: setVMRoot_def getThreadVSpaceRoot_def setCurrentUserVSpaceRoot_def)
apply (wp hoare_whenE_wp getCurrentUserCR3_wp findVSpaceForASID_vs_at_wp
| wpcw
| clarsimp simp: if_apply_def2 asid_wf_0
| strengthen valid_cr3'_makeCR3)+
done
crunch valid_arch_state'[wp]: switchToThread valid_arch_state'
(wp: crunch_wps simp: crunch_simps)
end
(*FIXME: arch_split: move up?*)
context Arch begin
context begin global_naming global
@ -48,16 +19,6 @@ end
context kernel_m begin
(* FIXME: move to Refine *)
lemma valid_idle'_tcb_at'_ksIdleThread:
"valid_idle' s \<Longrightarrow> tcb_at' (ksIdleThread s) s"
by (clarsimp simp: valid_idle'_def pred_tcb_at'_def obj_at'_def)
(* FIXME: move to Refine *)
lemma invs_no_cicd'_valid_idle':
"invs_no_cicd' s \<Longrightarrow> valid_idle' s"
by (simp add: invs_no_cicd'_def)
lemma Arch_switchToIdleThread_ccorres:
"ccorres dc xfdc invs_no_cicd' UNIV []
Arch.switchToIdleThread (Call Arch_switchToIdleThread_'proc)"
@ -71,11 +32,6 @@ lemma Arch_switchToIdleThread_ccorres:
apply (clarsimp simp: invs_no_cicd'_def valid_pspace'_def valid_idle'_tcb_at'_ksIdleThread)
done
(* FIXME: move *)
lemma empty_fail_getIdleThread [simp,intro!]:
"empty_fail getIdleThread"
by (simp add: getIdleThread_def)
lemma switchToIdleThread_ccorres:
"ccorres dc xfdc invs_no_cicd' UNIV hs
switchToIdleThread (Call switchToIdleThread_'proc)"

View File

@ -44,52 +44,6 @@ proof (rule ext)
qed
qed
(* FIXME x64: figure out where these are needed and adjust appropriately *)
lemma mask_pageBits_inner_beauty:
"is_aligned p 3 \<Longrightarrow>
(p && ~~ mask pageBits) + (ucast ((ucast (p && mask pageBits >> 3)):: 9 word) * 8) = (p::machine_word)"
apply (simp add: is_aligned_nth word_shift_by_3)
apply (subst word_plus_and_or_coroll)
apply (rule word_eqI)
apply (clarsimp simp: word_size word_ops_nth_size nth_ucast nth_shiftr nth_shiftl)
apply (rule word_eqI)
apply (clarsimp simp: word_size word_ops_nth_size nth_ucast nth_shiftr nth_shiftl
pageBits_def)
apply (rule iffI)
apply (erule disjE)
apply clarsimp
apply clarsimp
apply simp
apply clarsimp
apply (rule context_conjI)
apply (rule leI)
apply clarsimp
apply simp
apply arith
done
lemma more_pageBits_inner_beauty:
fixes x :: "9 word"
fixes p :: machine_word
assumes x: "x \<noteq> ucast (p && mask pageBits >> 3)"
shows "(p && ~~ mask pageBits) + (ucast x * 8) \<noteq> p"
apply clarsimp
apply (simp add: word_shift_by_3)
apply (subst (asm) word_plus_and_or_coroll)
apply (clarsimp simp: word_size word_ops_nth_size nth_ucast
nth_shiftl bang_eq)
apply (drule test_bit_size)
apply (clarsimp simp: word_size pageBits_def)
apply arith
apply (insert x)
apply (erule notE)
apply (rule word_eqI)
apply (clarsimp simp: word_size nth_ucast nth_shiftl nth_shiftr bang_eq)
apply (erule_tac x="n+3" in allE)
apply (clarsimp simp: word_ops_nth_size word_size)
apply (clarsimp simp: pageBits_def)
done
lemma byte_to_word_heap_upd_outside_range:
"p \<notin> {(base + ucast off * 8)..+8} \<Longrightarrow>

View File

@ -500,39 +500,6 @@ lemma invocationCatch_use_injection_handler:
split: sum.split)
done
lemma injection_handler_returnOk:
"injection_handler injector (returnOk v) = returnOk v"
by (simp add: returnOk_liftE injection_liftE)
lemma injection_handler_If:
"injection_handler injector (If P a b)
= If P (injection_handler injector a)
(injection_handler injector b)"
by (simp split: if_split)
(* FIXME: duplicated in CSpace_All *)
lemma injection_handler_liftM:
"injection_handler f
= liftM (\<lambda>v. case v of Inl ex \<Rightarrow> Inl (f ex) | Inr rv \<Rightarrow> Inr rv)"
apply (intro ext, simp add: injection_handler_def liftM_def
handleE'_def)
apply (rule bind_apply_cong, rule refl)
apply (simp add: throwError_def split: sum.split)
done
lemma injection_handler_throwError:
"injection_handler f (throwError v) = throwError (f v)"
by (simp add: injection_handler_def handleE'_def
throwError_bind)
lemma injection_handler_whenE:
"injection_handler injf (whenE P f)
= whenE P (injection_handler injf f)"
by (simp add: whenE_def injection_handler_returnOk split: if_split)
lemmas injection_handler_bindE = injection_bindE [OF refl refl]
lemmas injection_handler_wp = injection_wp [OF refl]
lemma ccorres_injection_handler_csum1:
"ccorres (f \<currency> r) xf P P' hs a c
\<Longrightarrow> ccorres
@ -783,9 +750,6 @@ lemma capFVMRights_range:
by (simp add: cap_frame_cap_lift_def
cap_lift_def cap_tag_defs word_and_le1 mask_def)+
lemma dumb_bool_for_all: "(\<forall>x. x) = False"
by auto
lemma dumb_bool_split_for_vcg:
"\<lbrace>d \<longrightarrow> \<acute>ret__unsigned_long \<noteq> 0\<rbrace> \<inter> \<lbrace>\<not> d \<longrightarrow> \<acute>ret__unsigned_long = 0\<rbrace>
= \<lbrace>d = to_bool \<acute>ret__unsigned_long \<rbrace>"

View File

@ -746,23 +746,6 @@ lemma handleFault_ccorres:
apply (clarsimp simp: pred_tcb_at')
done
lemma getMessageInfo_less_4:
"\<lbrace>\<top>\<rbrace> getMessageInfo t \<lbrace>\<lambda>rv s. msgExtraCaps rv < 4\<rbrace>"
including no_pre
apply (simp add: getMessageInfo_def)
apply wp
apply (rule hoare_strengthen_post, rule hoare_vcg_prop)
apply (simp add: messageInfoFromWord_def Let_def
Types_H.msgExtraCapBits_def)
apply (rule word_leq_minus_one_le, simp)
apply simp
apply (rule word_and_le1)
done
lemma invs_queues_imp:
"invs' s \<longrightarrow> valid_queues s"
by clarsimp
(* FIXME: move *)
lemma length_CL_from_H [simp]:
"length_CL (mi_from_H mi) = msgLength mi"
@ -790,16 +773,6 @@ lemma getMRs_length:
apply (simp add: msgMaxLength_def msgLengthBits_def n_msgRegisters_def)
done
lemma getMessageInfo_msgLength':
"\<lbrace>\<top>\<rbrace> getMessageInfo t \<lbrace>\<lambda>rv s. msgLength rv \<le> 0x78\<rbrace>"
including no_pre
apply (simp add: getMessageInfo_def)
apply wp
apply (rule hoare_strengthen_post, rule hoare_vcg_prop)
apply (simp add: messageInfoFromWord_def Let_def msgMaxLength_def not_less
Types_H.msgExtraCapBits_def split: if_split )
done
lemma handleInvocation_ccorres:
"ccorres (K dc \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
(invs' and
@ -1193,19 +1166,6 @@ lemma cap_case_EndpointCap_NotificationCap:
by (simp add: isCap_simps
split: capability.split)
(* FIXME: MOVE *)
lemma capFaultOnFailure_if_case_sum:
" (capFaultOnFailure epCPtr b (if c then f else g) >>=
sum.case_sum (handleFault thread) return) =
(if c then ((capFaultOnFailure epCPtr b f)
>>= sum.case_sum (handleFault thread) return)
else ((capFaultOnFailure epCPtr b g)
>>= sum.case_sum (handleFault thread) return))"
by (case_tac c, clarsimp, clarsimp)
(* FIXME: MOVE to Corres_C.thy *)
lemma ccorres_trim_redundant_throw_break:
"\<lbrakk>ccorres_underlying rf_sr \<Gamma> arrel axf arrel axf G G' (SKIP # hs) a c;
@ -1641,17 +1601,6 @@ lemma ctzl_spec:
apply (simp add: mex_def meq_def)
done
lemma dmo_machine_valid_lift:
"(\<And>s f m. P s = P (ksMachineState_update f s)) \<Longrightarrow> \<lbrace>P\<rbrace> doMachineOp f' \<lbrace>\<lambda>rv. P\<rbrace>"
apply (wpsimp simp: split_def doMachineOp_def machine_op_lift_def machine_rest_lift_def in_monad)
done
lemma tcb_runnable_imp_simple:
"obj_at' (\<lambda>s. runnable' (tcbState s)) t s \<Longrightarrow> obj_at' (\<lambda>s. simple' (tcbState s)) t s"
apply (clarsimp simp: obj_at'_def)
apply (case_tac "tcbState obj" ; clarsimp)
done
lemma ccorres_return_void_C_Seq:
"ccorres_underlying sr \<Gamma> r rvxf arrel xf P P' hs X (return_void_C) \<Longrightarrow>
ccorres_underlying sr \<Gamma> r rvxf arrel xf P P' hs X (return_void_C ;; Z)"
@ -1666,14 +1615,6 @@ lemma ccorres_return_void_C_Seq:
done
lemma scast_specific_plus64:
"scast (of_nat (word_ctz x) + 0x20 :: 64 signed word) = of_nat (word_ctz x) + (0x20 :: machine_word)"
by (simp add: scast_down_add is_down_def target_size_def source_size_def word_size)
lemma scast_specific_plus64_signed:
"scast (of_nat (word_ctz x) + 0x20 :: machine_word) = of_nat (word_ctz x) + (0x20 :: 64 signed word)"
by (simp add: scast_down_add is_down_def target_size_def source_size_def word_size)
lemma ccorres_handleReservedIRQ:
"ccorres dc xfdc
(invs' and (\<lambda>s. irq \<in> non_kernel_IRQs \<longrightarrow> sch_act_not (ksCurThread s) s \<and>

View File

@ -37,27 +37,6 @@ lemma ccorres_pre_threadGet:
done
(* FIXME MOVE *)
crunch inv'[wp]: archThreadGet P
(* FIXME MOVE near thm tg_sp' *)
lemma atg_sp':
"\<lbrace>P\<rbrace> archThreadGet f p \<lbrace>\<lambda>t. obj_at' (\<lambda>t'. f (tcbArch t') = t) p and P\<rbrace>"
including no_pre
apply (simp add: archThreadGet_def)
apply wp
apply (rule hoare_strengthen_post)
apply (rule getObject_tcb_sp)
apply clarsimp
apply (erule obj_at'_weakenE)
apply simp
done
(* FIXME: MOVE to EmptyFail *)
lemma empty_fail_archThreadGet [intro!, wp, simp]:
"empty_fail (archThreadGet f p)"
by (simp add: archThreadGet_def getObject_def split_def)
lemma ccorres_pre_archThreadGet:
assumes cc: "\<And>rv. ccorres r xf (P rv) (P' rv) hs (g rv) c"
shows "ccorres r xf
@ -299,16 +278,6 @@ lemma ccorres_pre_getObject_tcb:
apply simp
done
(* FIXME move to Invariants_H *)
lemma invs_cicd_arch_state' [elim!]:
"all_invs_but_ct_idle_or_in_cur_domain' s \<Longrightarrow> valid_arch_state' s"
by (simp add: all_invs_but_ct_idle_or_in_cur_domain'_def valid_state'_def)
(* FIXME move to Invariants_H *)
lemma invs_cicd_no_0_obj'[elim!]:
"all_invs_but_ct_idle_or_in_cur_domain' s \<Longrightarrow> no_0_obj' s"
by (simp add: all_invs_but_ct_idle_or_in_cur_domain'_def valid_state'_def valid_pspace'_def)
(* FIXME: MOVE, probably to CSpace_RAB *)
lemma ccorres_gen_asm2_state:
assumes rl: "\<And>s. P s \<Longrightarrow> ccorres r xf G G' hs a c"

View File

@ -8,16 +8,6 @@ theory TcbQueue_C
imports Ctac_lemmas_C
begin
(* FIXME: move *)
lemma (in semigroup) foldl_first:
"f x (foldl f y zs) = foldl f (f x y) zs"
by (induction zs arbitrary: x y) (auto simp: assoc)
(* FIXME: move *)
lemma (in monoid) foldl_first':
"f x (foldl f z zs) = foldl f x zs"
using foldl_first by simp
context kernel
begin
@ -213,14 +203,6 @@ proof -
qed
(* MOVE *)
lemma tcb_aligned':
"tcb_at' t s \<Longrightarrow> is_aligned t tcbBlockSizeBits"
apply (drule obj_at_aligned')
apply (simp add: objBits_simps)
apply (simp add: objBits_simps)
done
lemma tcb_at_not_NULL:
assumes tat: "tcb_at' t s"
shows "tcb_ptr_to_ctcb_ptr t \<noteq> NULL"
@ -1186,11 +1168,6 @@ lemma cvariable_relation_upd_const:
= cvariable_array_map_relation m (\<lambda>x. n)"
by (auto simp: fun_eq_iff cvariable_array_map_relation_def)
(* FIXME: move *)
lemma invs'_pspace_domain_valid:
"invs' s \<Longrightarrow> pspace_domain_valid s"
by (simp add: invs'_def valid_state'_def)
lemma ptr_span_ctcb_subset:
"is_aligned p tcbBlockSizeBits \<Longrightarrow> ptr_span (tcb_ptr_to_ctcb_ptr p) \<subseteq> {p .. p + 2^tcbBlockSizeBits-1}"
apply (simp add: tcb_ptr_to_ctcb_ptr_def ctcb_offset_def)

View File

@ -11,8 +11,7 @@ begin
lemma asUser_obj_at' :
"\<lbrace> K(t\<noteq>t') and obj_at' P t' \<rbrace> asUser t f \<lbrace> \<lambda>_. obj_at' (P::Structures_H.tcb \<Rightarrow> bool) t' \<rbrace>"
including no_pre
apply (wpsimp wp: hoare_vcg_ball_lift threadGet_wp simp: split_def asUser_def)
by (drule obj_at_ko_at', fastforce)
by (wpsimp wp: hoare_vcg_ball_lift threadGet_wp simp: split_def asUser_def)
lemma getObject_sched:
"(x::tcb, s') \<in> fst (getObject t s) \<Longrightarrow>
@ -343,16 +342,6 @@ lemma getMRs_rel_state:
apply (erule doMachineOp_state)
done
(* FIXME: move *)
lemma setTCB_cur:
"\<lbrace>cur_tcb'\<rbrace> setObject t (v::tcb) \<lbrace>\<lambda>_. cur_tcb'\<rbrace>"
including no_pre
apply (wp cur_tcb_lift)
apply (simp add: setObject_def split_def updateObject_default_def)
apply wp
apply simp
done
lemma setThreadState_getMRs_rel:
"\<lbrace>getMRs_rel args buffer and cur_tcb' and case_option \<top> valid_ipc_buffer_ptr' buffer
and (\<lambda>_. runnable' st)\<rbrace>
@ -1286,9 +1275,6 @@ lemma invokeTCB_CopyRegisters_ccorres:
apply (clarsimp dest!: global'_no_ex_cap simp: invs'_def valid_state'_def | rule conjI)+
done
(* FIXME: move *)
lemmas mapM_x_append = mapM_x_append2
lemma invokeTCB_WriteRegisters_ccorres_helper:
"\<lbrakk> unat (f (of_nat n)) = incn
\<and> g (of_nat n) = register_from_H reg \<and> n'=incn
@ -2452,15 +2438,6 @@ lemma decodeWriteRegisters_ccorres:
lemma excaps_map_Nil: "(excaps_map caps = []) = (caps = [])"
by (simp add: excaps_map_def)
(* FIXME: move *)
lemma and_eq_0_is_nth:
fixes x :: "('a :: len) word"
shows "y = 1 << n \<Longrightarrow> ((x && y) = 0) = (\<not> (x !! n))"
by (metis (poly_guards_query) and_eq_0_is_nth)
(* FIXME: move *)
lemmas and_neq_0_is_nth = arg_cong [where f=Not, OF and_eq_0_is_nth, simplified]
lemma decodeCopyRegisters_ccorres:
"interpret_excaps extraCaps' = excaps_map extraCaps \<Longrightarrow>
ccorres (intr_and_se_rel \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
@ -2585,11 +2562,6 @@ lemma decodeCopyRegisters_ccorres:
word_and_1_shiftl [where n=3,simplified] cap_get_tag_isCap[symmetric] split: if_split_asm)
done
(* FIXME: move *)
lemma ucast_le_ucast_8_32:
"(ucast x \<le> (ucast y :: word32)) = (x \<le> (y :: word8))"
by (simp add: word_le_nat_alt)
method wrong_cap_throwError_ccorres = solves \<open>
(rule ccorres_from_vcg_split_throws[where P=\<top> and P'=UNIV])
, vcg
@ -2708,14 +2680,6 @@ lemma slotCapLongRunningDelete_ccorres:
dest!: ccte_relation_ccap_relation)
done
(* FIXME: move *)
lemma empty_fail_slotCapLongRunningDelete:
"empty_fail (slotCapLongRunningDelete slot)"
by (auto simp: slotCapLongRunningDelete_def Let_def
case_Null_If isFinalCapability_def
split: if_split
intro!: empty_fail_bind)
definition
isValidVTableRoot_C :: "cap_C \<Rightarrow> bool"
where
@ -2757,20 +2721,10 @@ lemma updateCapData_spec:
\<lbrace>ccap_relation (RetypeDecls_H.updateCapData preserve newData cap) \<acute>ret__struct_cap_C\<rbrace>"
by (simp add: updateCapData_spec)
lemma if_n_updateCapData_valid_strg:
"s \<turnstile>' cap \<longrightarrow> s \<turnstile>' (if P then cap else updateCapData prs v cap)"
by (simp add: valid_updateCapDataI split: if_split)
lemma length_excaps_map:
"length (excaps_map xcs) = length xcs"
by (simp add: excaps_map_def)
(* FIXME: move *)
lemma from_bool_all_helper:
"(\<forall>bool. from_bool bool = val \<longrightarrow> P bool)
= ((\<exists>bool. from_bool bool = val) \<longrightarrow> P (val \<noteq> 0))"
by (auto simp: from_bool_0)
lemma getSyscallArg_ccorres_foo':
"ccorres (\<lambda>a rv. rv = ucast (args ! n)) (\<lambda>x. ucast (ret__unsigned_long_' x))
(sysargs_rel args buffer and sysargs_rel_n args buffer n)

View File

@ -43,25 +43,6 @@ lemma ccorres_flip_Guard:
apply (fastforce intro: exec.Guard exec.GuardFault exec_handlers.intros)
done
(* FIXME: move *)
lemma empty_fail_findVSpaceForASID[iff]:
"empty_fail (findVSpaceForASID asid)"
apply (simp add: findVSpaceForASID_def liftME_def)
apply (intro empty_fail_bindE, simp_all split: option.split)
apply (simp add: assertE_def split: if_split)
apply (simp add: assertE_def split: if_split)
apply (simp add: empty_fail_getObject)
apply (simp add: assertE_def liftE_bindE checkPML4At_def split: if_split)
done
(* FIXME: move *)
lemma empty_fail_findVSpaceForASIDAssert[iff]:
"empty_fail (findVSpaceForASIDAssert asid)"
apply (simp add: findVSpaceForASIDAssert_def catch_def
checkPML4At_def)
apply (intro empty_fail_bind, simp_all split: sum.split)
done
end
context kernel_m begin
@ -216,17 +197,6 @@ lemma vspace_at_asid_cross_over:
by (simp add: asid_map_relation_def bit_simps asid_bits_defs asid_bits_of_defs ucast_ucast_mask
split: option.splits asid_map_CL.splits)
(* FIXME remove *)
lemmas findVSpaceForASIDAssert_pd_at_wp2 = findVSpaceForASIDAssert_vs_at_wp
lemma asid_shiftr_low_bits_less:
"(asid :: machine_word) \<le> mask asid_bits \<Longrightarrow> asid >> asid_low_bits < 0x8"
apply (rule_tac y="2 ^ 3" in order_less_le_trans)
apply (rule shiftr_less_t2n)
apply (simp add: le_mask_iff_lt_2n[THEN iffD1] asid_bits_def asid_low_bits_def)
apply simp
done
lemma array_relation_update:
"\<lbrakk> array_relation R bnd table (arr :: 'a['b :: finite]);
x' = unat (x :: ('td :: len) word); R v v';
@ -243,9 +213,6 @@ where
vmfault_type.X64DataFault \<Rightarrow> scast Kernel_C.X86DataFault
| vmfault_type.X64InstructionFault \<Rightarrow> scast Kernel_C.X86InstructionFault"
lemmas mask_64_id[simp] = mask_len_id[where 'a=64, folded word_bits_def]
mask_len_id[where 'a=64, simplified]
(* FIXME: automate this *)
lemma seL4_Fault_VMFault_new'_spec:
"\<lbrace> \<lambda>s. s = \<sigma> \<rbrace> seL4_Fault_VMFault_new' addr FSR i
@ -477,7 +444,6 @@ lemma corres_symb_exec_unknown_r:
apply (rule corres_symb_exec_r[OF assms]; wp select_inv non_fail_select)
done
(* FIXME: automate this. *)
lemma lookupPML4Slot'_spec:
"\<lbrace> \<lambda>s. s = \<sigma> \<and> array_assertion pm (2 ^ ptTranslationBits) (hrs_htd (t_hrs_' s)) \<rbrace>
lookupPML4Slot' pm vptr
@ -1112,10 +1078,6 @@ next
by (clarsimp elim!: exI)
qed
(* FIXME move *)
lemma fromIntegral_simp_nat[simp]: "(fromIntegral :: nat \<Rightarrow> nat) = id"
by (simp add: fromIntegral_def fromInteger_nat toInteger_nat)
(* FIXME shadows two other identical versions in other files *)
lemma ccorres_abstract_known:
"\<lbrakk> \<And>rv' t t'. ceqv \<Gamma> xf' rv' t t' g (g' rv'); ccorres rvr xf P P' hs f (g' val) \<rbrakk>
@ -1253,11 +1215,6 @@ lemma setCurrentUserVSpaceRoot_ccorres:
apply (clarsimp simp: ccr3_relation_defs Let_def mask_def bit_simps asid_bits_def)
done
(* FIXME: move *)
lemma invs_cicd_valid_objs' [elim!]:
"all_invs_but_ct_idle_or_in_cur_domain' s \<Longrightarrow> valid_objs' s"
by (simp add: all_invs_but_ct_idle_or_in_cur_domain'_def valid_pspace'_def)
lemma setVMRoot_ccorres:
"ccorres dc xfdc
(all_invs_but_ct_idle_or_in_cur_domain' and tcb_at' thread)
@ -1566,15 +1523,6 @@ lemma performPageGetAddress_ccorres:
kernel_all_global_addresses.msgRegisters_def fupdate_def)
done
lemma ignoreFailure_liftM:
"ignoreFailure = liftM (\<lambda>v. ())"
apply (rule ext)+
apply (simp add: ignoreFailure_def liftM_def
catch_def)
apply (rule bind_apply_cong[OF refl])
apply (simp split: sum.split)
done
lemma ccorres_pre_getObject_pte:
assumes cc: "\<And>rv. ccorres r xf (P rv) (P' rv) hs (f rv) c"
shows "ccorres r xf
@ -2071,22 +2019,6 @@ lemma ccap_relation_mapped_asid_0:
apply simp
done
(* FIXME: move *)
lemma getSlotCap_wp':
"\<lbrace>\<lambda>s. \<forall>cap. cte_wp_at' (\<lambda>c. cteCap c = cap) p s \<longrightarrow> Q cap s\<rbrace> getSlotCap p \<lbrace>Q\<rbrace>"
apply (simp add: getSlotCap_def)
apply (wp getCTE_wp')
apply (clarsimp simp: cte_wp_at_ctes_of)
done
lemma vmsz_aligned_aligned_pageBits:
"vmsz_aligned ptr sz \<Longrightarrow> is_aligned ptr pageBits"
apply (simp add: vmsz_aligned_def)
apply (erule is_aligned_weaken)
apply (simp add: pageBits_def pageBitsForSize_def
split: vmpage_size.split)
done
lemma framesize_from_H_bounded:
"framesize_from_H x < 3"
by (clarsimp simp: framesize_from_H_def X86_SmallPage_def X86_LargePage_def X64_HugePage_def