From 86731939f21dca112aebdb9f96f47971cdc46d7b Mon Sep 17 00:00:00 2001 From: Thomas Sewell Date: Wed, 21 Sep 2016 10:33:03 +1000 Subject: [PATCH] SELFOUR-444: CRefine proof for preemptible retype. --- lib/SIMPL_Lemmas.thy | 3 +- lib/clib/Corres_UL_C.thy | 19 + lib/clib/Ctac.thy | 233 ++++- lib/clib/ctac-method.ML | 17 +- proof/crefine/ADT_C.thy | 68 +- proof/crefine/Arch_C.thy | 108 ++- proof/crefine/CSpace_C.thy | 395 +++++--- proof/crefine/Detype_C.thy | 29 +- proof/crefine/Fastpath_C.thy | 164 +++- proof/crefine/Finalise_C.thy | 29 +- proof/crefine/Invoke_C.thy | 1306 +++++++++++++++++++-------- proof/crefine/IpcCancel_C.thy | 188 +++- proof/crefine/Ipc_C.thy | 247 +++-- proof/crefine/Recycle_C.thy | 118 ++- proof/crefine/Retype_C.thy | 1394 +++++++++++++++-------------- proof/crefine/SR_lemmas_C.thy | 81 ++ proof/crefine/Schedule_C.thy | 6 +- proof/crefine/StateRelation_C.thy | 26 + proof/crefine/StoreWord_C.thy | 52 +- proof/crefine/Syscall_C.thy | 29 +- proof/crefine/TcbQueue_C.thy | 21 +- proof/crefine/Tcb_C.thy | 31 +- proof/crefine/VSpace_C.thy | 7 +- proof/drefine/Untyped_DR.thy | 191 ++-- 24 files changed, 3104 insertions(+), 1658 deletions(-) diff --git a/lib/SIMPL_Lemmas.thy b/lib/SIMPL_Lemmas.thy index 6f74effc3..7b38b00e0 100644 --- a/lib/SIMPL_Lemmas.thy +++ b/lib/SIMPL_Lemmas.thy @@ -227,7 +227,8 @@ lemma Guard_Seq_semantic_equiv: lemma exec_Seq_cong: "\ \s''. \ \ \a, Normal s\ \ s'' = \ \ \c, Normal s\ \ s''; - \s''. \ \ \b, Normal s''\ \ s' = \ \ \d, Normal s''\ \ s' \ + \s''. \ \ \c, Normal s\ \ Normal s'' + \ \ \ \b, Normal s''\ \ s' = \ \ \d, Normal s''\ \ s' \ \ \ \ \a ;; b, Normal s\ \ s' = \ \ \c ;; d, Normal s\ \ s'" apply (rule iffI) apply (erule exec_Normal_elim_cases) diff --git a/lib/clib/Corres_UL_C.thy b/lib/clib/Corres_UL_C.thy index 0c91abcf0..abba9bdd5 100644 --- a/lib/clib/Corres_UL_C.thy +++ b/lib/clib/Corres_UL_C.thy @@ -1598,4 +1598,23 @@ lemma ccorres_cond_seq: apply assumption done +lemma ccorres_assume_pre: + assumes "\s. P s \ ccorres_underlying sr \ r xf r' xf' (P and (\s'. s' = s)) P' hs H C" + shows "ccorres_underlying sr \ r xf r' xf' P P' hs H C" + apply (clarsimp simp: ccorres_underlying_def) + apply (frule assms) + apply (simp add: ccorres_underlying_def) + apply blast + done + +lemma ccorres_name_pre: + "(\s. P s \ ccorres_underlying sr \ r xf r' xf' (\s'. s' = s) P' hs H C) + \ ccorres_underlying sr \ r xf r' xf' P P' hs H C" + apply (rule ccorres_assume_pre) + apply (rule ccorres_guard_imp) + apply fastforce + apply simp + apply simp + done + end diff --git a/lib/clib/Ctac.thy b/lib/clib/Ctac.thy index 24081365b..a4bea8652 100644 --- a/lib/clib/Ctac.thy +++ b/lib/clib/Ctac.thy @@ -11,8 +11,8 @@ (* Automation framework for general C refinement *) theory Ctac -imports - Corres_C +imports + Corres_C "../XPres" begin @@ -360,8 +360,8 @@ lemma semantic_equiv_While_cong: done lemma semantic_equiv_Seq_cong: - assumes sea: "\s s'. semantic_equiv Gamma s s' a a'" - and seb: "\s s'. semantic_equiv Gamma s s' b b'" + assumes sea: "\s'. semantic_equiv Gamma s s' a a'" + and seb: "\s. semantic_equiv Gamma s s' b b'" shows "semantic_equiv Gamma s s' (a ;; b) (a' ;; b')" using sea seb apply (simp add: semantic_equiv_def2) @@ -399,8 +399,8 @@ lemma semantic_equiv_refl: by (rule semantic_equivI, simp) lemma semantic_equiv_trans: - assumes sea: "\s s'. semantic_equiv Gamma s s' a b" - and seb: "\s s'. semantic_equiv Gamma s s' b c" + assumes sea: "semantic_equiv Gamma s s' a b" + and seb: "semantic_equiv Gamma s s' b c" shows "semantic_equiv Gamma s s' a c" using sea seb by (simp add: semantic_equiv_def2) @@ -1049,8 +1049,8 @@ lemmas ceqv_rules = ceqv_refl [where xf' = xfdc] -- "Any ceqv with xfdc should b ceqv_refl [where c = catchbrk_C] definition - ceqv_xpres :: "(int \ ('s, int, strictc_errortype) com) \ ('s \ 'a) \ 'a - \ bool \ ('s, int, strictc_errortype) com \ bool \ ('s, int, strictc_errortype) com \ bool" + ceqv_xpres :: "('p \ ('s, 'p, 'x) com) \ ('s \ 'a) \ 'a + \ bool \ ('s, 'p, 'x) com \ bool \ ('s, 'p, 'x) com \ bool" where "ceqv_xpres \ xf v pres c pres' c' \ \s s' s''. (pres \ xf s = v) @@ -1335,6 +1335,223 @@ lemma ceqv_xpres_eq_If_rules: "ceqv_xpres_eq_If True x y x" by (simp add: ceqv_xpres_eq_If_def)+ +definition + "simpl_sequence f xs + = foldr (Seq) (map f xs) Skip" + +lemma simpl_sequence_Cons: + "simpl_sequence f (x # xs) = Seq (f x) (simpl_sequence f xs)" + by (simp add: simpl_sequence_def) + +fun(sequential) + simpl_final_basic_opts :: "('s, 'p, 'x) com \ ('s \ 's) option" +where + "simpl_final_basic_opts (x ;; y) + = (case (simpl_final_basic_opts y) of None \ simpl_final_basic_opts x + | Some v \ Some v)" + | "simpl_final_basic_opts (Basic f) = Some f" + | "simpl_final_basic_opts (Guard E F c) = simpl_final_basic_opts c" + | "simpl_final_basic_opts Skip = None" + | "simpl_final_basic_opts Throw = None" + | "simpl_final_basic_opts c = Some id" + +definition + "simpl_final_basic c = (case simpl_final_basic_opts c + of Some v \ v | None \ id)" + +lemmas simpl_final_basic_simps[simp] + = simpl_final_basic_def[where c="Seq c c'" for c c'] + simpl_final_basic_def[where c="Basic f" for f] + simpl_final_basic_def[where c="Guard E F c" for E F c] + +lemma simpl_final_basic_opts_exec[OF _ refl refl]: + "\ \ \c, xs\ \ xs' \ xs = Normal t \ xs' = Normal t' + \ (case simpl_final_basic_opts c of None \ t' = t + | Some f \ \s. t' = f s)" + apply (induct arbitrary: t t' rule: exec.induct, simp_all) + apply metis + apply atomize + apply clarsimp + apply (case_tac s') + apply (auto split: option.split_asm)[1] + apply (auto elim!: exec_elim_cases) + done + +lemma simpl_final_basic_exec: + "\ \ \c, Normal t\ \ Normal t' + \ \s. t' = simpl_final_basic c s" + apply (frule simpl_final_basic_opts_exec) + apply (simp add: simpl_final_basic_def split: option.split_asm) + done + +lemma ceqv_xpres_to_simpl_sequence: + fixes v :: "'a :: ring_1" + assumes c: "\v. ceqv_xpres \ xf' v True c pres (c' v)" + and v: "\v s. xf' (simpl_final_basic (c' v) s) - v = offs" + shows "\ CP (v + of_nat n * offs) + \ ceqv_xpres \ xf' v True (While {s. CP (xf' s)} c) False + (simpl_sequence c' (takeWhile CP (map (\x. v + of_nat x * offs) [0 ..< n])))" + (is "_ \ ceqv_xpres _ _ _ _ (While ?S _) _ _") +proof (induct n arbitrary: v) + case 0 + show ?case using c[where v=v] 0 + apply (simp add: simpl_sequence_def) + apply (simp add: ceqv_xpres_eq_imp ceqv_def) + apply (auto elim!: exec_Normal_elim_cases intro: exec.intros)[1] + done +next + case (Suc n) + have foo: "\t t'. (\ \ \c' v, Normal t\ \ Normal t') \ xf' t' = v + offs" + using v + by (clarsimp simp: field_simps dest!: simpl_final_basic_exec) + + show ?case using c[where v=v] Suc.hyps[where v="v + offs"] Suc.prems + apply (subst upt_conv_Cons, simp) + apply (simp only: map_Suc_upt[symmetric] list.simps) + apply (cases "CP v") + apply (simp add: o_def field_simps simpl_sequence_Cons + ceqv_xpres_eq_imp) + apply (clarsimp, rule ceqv_trans[where c'="c ;; While ?S c"]) + apply (simp add: ceqv_def) + apply (auto elim!: exec_Normal_elim_cases intro: exec.Seq exec.WhileTrue)[1] + apply (rule ceqv_trans[where c'="c' v ;; While ?S c"]) + apply (simp add: ceqv_def) + apply (auto elim!: exec_Normal_elim_cases intro: exec.Seq)[1] + apply (simp add: ceqv_def) + apply (intro impI exec_Seq_cong refl) + apply (simp add: foo) + apply (simp add: simpl_sequence_def field_simps) + apply (simp add: ceqv_xpres_eq_imp ceqv_def) + apply (auto intro: exec.WhileFalse exec.Skip elim!: exec_Normal_elim_cases)[1] + done +qed + +lemma ceqv_xpres_While_simpl_sequence: + fixes v :: "'a :: ring_1" + assumes c: "\v. ceqv_xpres \ xf' v True c pres (c' v)" + shows "ceqv_xpres \ xf' v True (While {s. CP (xf' s)} c) False + (if \n offs. (\s v. (xf' (simpl_final_basic (c' v) s) - v = offs)) \ \ CP (v + of_nat n * offs) + then simpl_sequence c' (map (\x. v + of_nat x + * (THE offs. \s v. (xf' (simpl_final_basic (c' v) s) - v = offs))) + [0 ..< (LEAST n. \ CP (v + of_nat n + * (THE offs. \s v. (xf' (simpl_final_basic (c' v) s) - v = offs))))]) + else While {s. CP (xf' s)} c)" + apply (split split_if, simp add: ceqv_xpres_def[where c=c and c'=c for c]) + apply (clarsimp simp: ceqv_xpres_eq_ceqv) + apply (rule ceqv_trans) + apply (rule_tac n="LEAST n. \ CP (v + of_nat n * offs)" + in ceqv_xpres_to_simpl_sequence[simplified ceqv_xpres_eq_ceqv, rule_format]) + apply (rule c) + apply simp + apply simp + apply (rule LeastI_ex) + apply blast + apply (subst takeWhile_eq_all_conv[THEN iffD2]) + apply (clarsimp dest!: not_less_Least) + apply (simp add: ceqv_def) + done + +lemma ccorres_underlying_name_seq_bound: + "(\ CP n \ (\n' < n. CP n')) + \ ccorres_underlying srel \ rrel xf arrel axf G G' hs m + (simpl_sequence c' (map f [0 ..< n])) + \ ccorres_underlying srel \ rrel xf arrel axf + G G' hs m (if \n. \ CP n + then simpl_sequence c' (map f [0 ..< (LEAST n. \ CP n)]) + else c)" + apply (subst if_P, blast) + apply (subst Least_equality[where x=n], simp_all) + apply (rule ccontr, simp add: linorder_not_le) + done + +lemma sequenceE_simpl_sequence_nth_corres': + "\ length xs = length ys; + \zs. length zs < length xs \ + ccorres_underlying sr \ (inr_rrel (\rv rv'. r' (prev_xs @ zs @ [rv]) rv')) xf' + (inl_rrel arrel) axf + (P and F (length prev_xs + length zs)) (Q \ {s. r' (prev_xs @ zs) (xf' s)}) hs + (xs ! length zs) (f (ys ! length zs)); + \s \. s \ Q \ P \ \ + (\, s) \ sr \ \y \ set ys. \\\<^bsub>/UNIV\<^esub> {s} (f y) Q,UNIV; + \n. Suc n < length xs \ \P and F (length prev_xs + n)\ xs ! n \\_. P and F (length prev_xs + Suc n)\, - + \ + \ ccorres_underlying sr \ (inr_rrel (\rv rv'. r' (prev_xs @ rv) rv')) xf' + (inl_rrel arrel) axf + (\s. xs \ [] \ P s \ F (length prev_xs) s) (Q \ {s. r' prev_xs (xf' s)}) hs + (sequenceE xs) + (simpl_sequence f ys)" +proof (induct xs ys arbitrary: prev_xs rule: list_induct2) + case Nil + show ?case + apply (simp add: sequenceE_def simpl_sequence_def) + apply (rule ccorres_guard_imp2, rule ccorres_returnOk_skip) + apply simp + done +next + case (Cons x xs y ys) + show ?case + apply (simp add: simpl_sequence_Cons sequenceE_Cons) + apply (rule ccorres_guard_imp2) + apply (rule ccorres_splitE) + apply (simp add: inl_rrel_inl_rrel) + apply (rule Cons.prems(1)[where zs=Nil, simplified]) + apply (rule ceqv_refl) + apply (simp add: liftME_def[symmetric] liftME_liftM) + apply (rule ccorres_rel_imp2, rule Cons.hyps(2)[where prev_xs="prev_xs @ [rv]" for rv]) + apply (rule ccorres_guard_imp2, rule ccorres_rel_imp2, + rule Cons.prems(1)[where zs="z # zs" for z zs, simplified]) + apply simp+ + apply (blast dest: Cons.prems[simplified]) + apply simp + apply (cut_tac n="Suc n" in Cons.prems(3), simp, simp) + apply (clarsimp elim!: inl_inrE) + apply assumption + apply (clarsimp elim!: inl_inrE) + apply simp + apply (rule hoare_vcg_const_imp_lift_R) + apply (rule hoare_gen_asmE) + apply (erule Cons.prems(3)[where n=0, simplified]) + apply (rule_tac P="Q \ {s. \\. P \ \ (\, s) \ sr}" + in HoarePartial.conseq_exploit_pre) + apply (clarsimp, rule conseqPost, rule Cons.prems(2)[simplified, THEN conjunct1], + simp+) + apply (clarsimp simp: ccHoarePost_def elim!: inl_inrE) + apply simp + apply auto + done +qed + +lemmas sequenceE_simpl_sequence_nth_corres + = sequenceE_simpl_sequence_nth_corres'[where prev_xs=Nil, simplified] + +lemma mapME_x_simpl_sequence_fun_related: + "\ ys = map yf xs; + \n x. x \ set xs \ + ccorres_underlying sr \ (inr_rrel dc) xfdc (inl_rrel arrel) axf + (P and F n (n < length xs) x) Q hs + (f x) (f' (yf x)); + \s \. s \ Q \ P \ \ + (\, s) \ sr \ \x \ set xs. \\\<^bsub>/UNIV\<^esub> {s} (f' (yf x)) Q,UNIV; + \n. Suc n < length xs \ \P and F n True (xs ! n)\ f (xs ! n) \\_. P and F (Suc n) (Suc n < length xs) (xs ! Suc n)\, - + \ + \ ccorres_underlying sr \ (inr_rrel dc) xfdc + (inl_rrel arrel) axf + (P and F 0 (xs \ []) (xs ! 0)) Q hs + (mapME_x f xs) + (simpl_sequence f' ys)" + apply (simp add: mapME_x_sequenceE liftME_def[symmetric] + liftME_liftM) + apply (rule ccorres_rel_imp2, rule ccorres_guard_imp2, + rule sequenceE_simpl_sequence_nth_corres[where r'=dc and xf'=xfdc + and P=P and F="\i. F i (i < length xs) (xs ! i)" and Q=Q and arrel=arrel and axf=axf]; + clarsimp elim!: inl_inrE) + apply (erule_tac x="length zs" in meta_allE + | erule_tac x="xs ! length zs" in meta_allE)+ + apply (simp add: dc_def) + done + +lemmas mapME_x_simpl_sequence_same + = mapME_x_simpl_sequence_fun_related[where yf=id, simplified] lemma call_ignore_cong: "call i f g r = call i f g r" by (rule refl) diff --git a/lib/clib/ctac-method.ML b/lib/clib/ctac-method.ML index 1c690c577..9a0e1826f 100644 --- a/lib/clib/ctac-method.ML +++ b/lib/clib/ctac-method.ML @@ -123,6 +123,10 @@ val ctac_post_del = Thm.declaration_attribute (fn thm => (ctac_post.map (Thm.del_thm thm))); +val ceqv_simpl_sequence_pair = Attrib.config_bool + @{binding ceqv_simpl_sequence} (K false) +fun ceqv_simpl_seq ctxt = Config.get ctxt (fst ceqv_simpl_sequence_pair) + val setup = Attrib.setup @{binding "corres"} (Attrib.add_del ctac_add ctac_del) "correspondence rules" @@ -131,7 +135,8 @@ val setup = "correspondence preprocessing rules" #> Attrib.setup @{binding "corres_post"} (Attrib.add_del ctac_post_add ctac_post_del) - "correspondence postprocessing rules"; + "correspondence postprocessing rules" + #> snd ceqv_simpl_sequence_pair; (* tacticals *) @@ -327,14 +332,20 @@ val ceqv_xpres_eq_If_rules = @{thms ceqv_xpres_eq_If_rules}; val apply_ceqv_xpres_rules_trace = ref @{term True}; -fun apply_ceqv_xpres_rules ctxt _ n = DETERM - (resolve_tac ctxt ceqv_xpres_rules n +fun apply_ceqv_xpres_rules1 ctxt rules _ n = DETERM + (resolve_tac ctxt rules n ORELSE (SUBGOAL (fn (t, n) => (warning ("apply_ceqv_xpres_rules: no rule applied" ^ " - see CtacImpl.apply_ceqv_xpres_rules_trace"); apply_ceqv_xpres_rules_trace := t; resolve_tac ctxt [ceqv_xpres_FalseI] n)) n)); +fun apply_ceqv_xpres_rules ctxt = let + val seq = ceqv_simpl_seq ctxt + val ex_rules = if seq then [@{thm ceqv_xpres_While_simpl_sequence}] + else [] + in apply_ceqv_xpres_rules1 ctxt (ex_rules @ ceqv_xpres_rules) end + fun addcongs thms ss = foldl (uncurry Simplifier.add_cong) ss thms fun delsplits thms ss = foldl (uncurry Splitter.del_split) ss thms diff --git a/proof/crefine/ADT_C.thy b/proof/crefine/ADT_C.thy index b18684f47..cb9991be3 100644 --- a/proof/crefine/ADT_C.thy +++ b/proof/crefine/ADT_C.thy @@ -1440,12 +1440,45 @@ lemma cbitmap_L2_to_H_correct: apply clarsimp done +definition + mk_gsUntypedZeroRanges +where + "mk_gsUntypedZeroRanges s + = ran (untypedZeroRange \\<^sub>m (option_map cteCap o map_to_ctes (cstate_to_pspace_H s)))" + +lemma cpspace_user_data_relation_user_mem'[simp]: + "\pspace_aligned' as;pspace_distinct' as\ \ cpspace_user_data_relation (ksPSpace as) (option_to_0 \ user_mem' as) (t_hrs_' cs) + = cpspace_user_data_relation (ksPSpace as) (underlying_memory (ksMachineState as)) (t_hrs_' cs)" + by (simp add: cmap_relation_def) + +lemma cpspace_device_data_relation_user_mem'[simp]: + "cpspace_device_data_relation (ksPSpace as) (option_to_0 \ user_mem' as) (t_hrs_' cs) + = cpspace_device_data_relation (ksPSpace as) (underlying_memory (ksMachineState as)) (t_hrs_' cs)" + apply (clarsimp simp: cmap_relation_def cuser_user_data_device_relation_def heap_to_device_data_def) + apply (rule_tac arg_cong[where f = "%x. x = y" for y]) + by auto + +lemma (in kernel) mk_gsUntypedZeroRanges_correct: + assumes valid: "valid_state' as" + assumes cstate_rel: "cstate_relation as cs" + shows "mk_gsUntypedZeroRanges cs = gsUntypedZeroRanges as" + using assms + apply (clarsimp simp: valid_state'_def untyped_ranges_zero_inv_def + mk_gsUntypedZeroRanges_def cteCaps_of_def) + apply (subst cstate_to_pspace_H_correct[where c=cs], simp_all) + apply (clarsimp simp: cstate_relation_def Let_def) + apply (subst cstate_to_machine_H_correct, assumption, simp_all) + apply (clarsimp simp: cpspace_relation_def)+ + apply (clarsimp simp: observable_memory_def valid_pspace'_def) + done + definition (in state_rel) cstate_to_H :: "globals \ kernel_state" where "cstate_to_H s \ \ksPSpace = cstate_to_pspace_H s, gsUserPages = fst (ghost'state_' s), gsCNodes = fst (snd (ghost'state_' s)), + gsUntypedZeroRanges = mk_gsUntypedZeroRanges s, gsMaxObjectSize = (let v = unat (gs_get_assn cap_get_capSizeBits_'proc (ghost'state_' s)) in if v = 0 then card (UNIV :: word32 set) else v), ksDomScheduleIdx = unat (ksDomScheduleIdx_' s), @@ -1468,36 +1501,25 @@ where lemma trivial_eq_conj: "B = C \ (A \ B) = (A \ C)" by simp -lemma cpspace_user_data_relation_user_mem'[simp]: - "\pspace_aligned' as;pspace_distinct' as\ \ cpspace_user_data_relation (ksPSpace as) (option_to_0 \ user_mem' as) (t_hrs_' cs) - = cpspace_user_data_relation (ksPSpace as) (underlying_memory (ksMachineState as)) (t_hrs_' cs)" - by (simp add: cmap_relation_def) - -lemma cpspace_device_data_relation_user_mem'[simp]: - "cpspace_device_data_relation (ksPSpace as) (option_to_0 \ user_mem' as) (t_hrs_' cs) - = cpspace_device_data_relation (ksPSpace as) (underlying_memory (ksMachineState as)) (t_hrs_' cs)" - apply (clarsimp simp: cmap_relation_def cuser_user_data_device_relation_def heap_to_device_data_def) - apply (rule_tac arg_cong[where f = "%x. x = y" for y]) - by auto - - lemma (in kernel_m) cstate_to_H_correct: assumes valid: "valid_state' as" assumes cstate_rel: "cstate_relation as cs" shows "cstate_to_H cs = as \ksMachineState:= observable_memory (ksMachineState as) (user_mem' as)\" apply (subgoal_tac "cstate_to_machine_H cs = observable_memory (ksMachineState as) (user_mem' as)") apply (rule kernel_state.equality, simp_all add: cstate_to_H_def) - apply (rule cstate_to_pspace_H_correct) - using valid - apply (simp add: valid_state'_def) - using cstate_rel valid - apply (clarsimp simp: cstate_relation_def cpspace_relation_def Let_def - observable_memory_def valid_state'_def - valid_pspace'_def) + apply (rule cstate_to_pspace_H_correct) + using valid + apply (simp add: valid_state'_def) + using cstate_rel valid + apply (clarsimp simp: cstate_relation_def cpspace_relation_def Let_def + observable_memory_def valid_state'_def + valid_pspace'_def) + using cstate_rel + apply (clarsimp simp: cstate_relation_def cpspace_relation_def Let_def prod_eq_iff) using cstate_rel - apply (clarsimp simp: cstate_relation_def cpspace_relation_def Let_def prod_eq_iff) - using cstate_rel - apply (clarsimp simp: cstate_relation_def cpspace_relation_def Let_def prod_eq_iff) + apply (clarsimp simp: cstate_relation_def cpspace_relation_def Let_def prod_eq_iff) + using valid cstate_rel + apply (rule mk_gsUntypedZeroRanges_correct) using cstate_rel apply (fastforce simp: cstate_relation_def cpspace_relation_def Let_def ghost_size_rel_def unat_eq_0 diff --git a/proof/crefine/Arch_C.thy b/proof/crefine/Arch_C.thy index 716aac8f5..3aac81bf2 100644 --- a/proof/crefine/Arch_C.thy +++ b/proof/crefine/Arch_C.thy @@ -88,9 +88,10 @@ lemma performPageTableInvocationUnmap_ccorres: apply (clarsimp simp: typ_heap_simps cap_get_tag_isCap_ArchObject) apply (rule fst_setCTE [OF ctes_of_cte_at], assumption) apply (erule rev_bexI) - apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def) + apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def + typ_heap_simps') apply (rule conjI) - apply (clarsimp simp: cpspace_relation_def) + apply (clarsimp simp: cpspace_relation_def typ_heap_simps') apply (subst setCTE_tcb_case, assumption+) apply (clarsimp dest!: ksPSpace_update_eq_ExD) apply (erule cmap_relation_updI, assumption) @@ -106,7 +107,7 @@ lemma performPageTableInvocationUnmap_ccorres: cap_page_table_cap_lift_def) apply simp apply (clarsimp simp: carch_state_relation_def cmachine_state_relation_def - h_t_valid_clift_Some_iff + typ_heap_simps' cvariable_array_map_const_add_map_option[where f="tcb_no_ctes_proj"] dest!: ksPSpace_update_eq_ExD) apply (simp add: cte_wp_at_ctes_of) @@ -146,6 +147,7 @@ lemma createObjects_asidpool_ccorres: shows "ccorres dc xfdc ((\s. \p. cte_wp_at' (\cte. cteCap cte = UntypedCap isdev frame pageBits idx ) p s) and pspace_aligned' and pspace_distinct' and valid_objs' + and ret_zero frame (2 ^ pageBits) and valid_global_refs' and pspace_no_overlap' frame pageBits) ({s. region_actually_is_bytes frame (2^pageBits) s}) hs @@ -156,6 +158,7 @@ proof - have helper: "\\ x. (\, x) \ rf_sr \ is_aligned frame pageBits \ frame \ 0 \ pspace_aligned' \ \ pspace_distinct' \ \ pspace_no_overlap' frame pageBits \ + \ ret_zero frame (2 ^ pageBits) \ \ region_actually_is_bytes frame (2 ^ pageBits) x \ {frame ..+ 2 ^ pageBits} \ kernel_data_refs = {} \ @@ -178,9 +181,10 @@ proof - hence rf: "(\, x) \ rf_sr" and al: "is_aligned frame pageBits" and ptr0: "frame \ 0" and pal: "pspace_aligned' \" and pdst: "pspace_distinct' \" and pno: "pspace_no_overlap' frame pageBits \" + and zro: "ret_zero frame (2 ^ pageBits) \" and actually: "region_actually_is_bytes frame (2 ^ pageBits) x" and kdr: "{frame ..+ 2 ^ pageBits} \ kernel_data_refs = {}" - by (simp_all) + by simp_all note empty = region_actually_is_bytes[OF actually] @@ -307,7 +311,7 @@ proof - apply (rule relrl) done - thus ?thesis using rf empty kdr + thus ?thesis using rf empty kdr zro apply (simp add: rf_sr_def cstate_relation_def Let_def rl' tag_disj_via_td_name ko_def[symmetric]) apply (simp add: carch_state_relation_def cmachine_state_relation_def) @@ -316,6 +320,7 @@ proof - kernel_data_refs_domain_eq_rotate cvariable_array_ptr_retyps[OF szo] foldr_upd_app_if [folded data_map_insert_def] + zero_ranges_ptr_retyps rl empty projectKOs) done qed @@ -425,17 +430,31 @@ shows apply (rule ccorres_Guard_Seq[where S=UNIV])? apply (rule deleteObjects_ccorres) apply ceqv - apply (rule ccorres_move_c_guard_cte) - apply csymbr + apply (rule ccorres_rhs_assoc2) apply (rule ccorres_abstract_cleanup) apply (rule ccorres_symb_exec_l) apply (rule_tac P = "rva = (capability.UntypedCap isdev frame pageBits idx)" in ccorres_gen_asm) apply (simp add: hrs_htd_update del:fun_upd_apply) - apply (rule ccorres_move_Guard_Seq_strong[OF c_guard_abs_cte]) - apply (ctac add:update_freeIndex) + apply (rule ccorres_split_nothrow) + + apply (rule_tac cap'="UntypedCap isdev frame pageBits idx" in updateFreeIndex_ccorres) + apply (rule allI, rule conseqPre, vcg) + apply (rule subsetI, clarsimp simp: typ_heap_simps' pageBits_def isCap_simps) + apply (frule ccte_relation_ccap_relation, clarsimp) + apply (frule cap_get_tag_isCap_unfolded_H_cap) + apply (clarsimp simp: isCap_simps cap_lift_untyped_cap + cap_to_H_simps cap_untyped_cap_lift_def + ccap_relation_def modify_map_def + fun_eq_iff + dest!: word_unat.Rep_inverse' split: split_if) + apply (rule exI, strengthen refl) + apply (case_tac cte', simp add: cap_lift_untyped_cap max_free_index_def mask_def) + apply (simp add: mex_def meq_def del: split_paired_Ex) + apply blast + apply ceqv apply ccorres_remove_UNIV_guard apply (rule ccorres_Guard_Seq[where F=ShiftError])+ - apply (ctac (c_lines 2) add:createObjects_asidpool_ccorres + apply (ctac (c_lines 2) add:createObjects_asidpool_ccorres[where idx="max_free_index pageBits"] pre del: ccorres_Guard_Seq) apply csymbr apply (ctac (no_vcg) add: cteInsert_ccorres) @@ -471,15 +490,16 @@ shows apply vcg apply (clarsimp simp:conj_comms objBits_simps archObjSize_def | strengthen valid_pspace_mdb' vp_strgs' invs_valid_pspace' - valid_pspace_valid_objs' invs_valid_global')+ + valid_pspace_valid_objs' invs_valid_global' + invs_urz)+ apply (wp updateFreeIndex_forward_invs' - updateFreeIndex_caps_no_overlap''[where cap = "UntypedCap isdev frame pageBits idx",simplified] - updateFreeIndex_pspace_no_overlap'[where cap = "UntypedCap isdev frame pageBits idx",simplified] - updateFreeIndex_caps_overlap_reserved'[where cap = "UntypedCap isdev frame pageBits idx",simplified]) - apply (rule_tac P1 = "\x. x = UntypedCap isdev frame pageBits (max_free_index pageBits)" - in hoare_strengthen_post[OF updateCap_weak_cte_wp_at[where p = parent]]) - apply (rule exI) - apply assumption + updateFreeIndex_caps_no_overlap''[where sz=pageBits] + updateFreeIndex_pspace_no_overlap'[where sz=pageBits] + updateFreeIndex_caps_overlap_reserved + updateFreeIndex_cte_wp_at + ) + apply (strengthen exI[where x=parent]) + apply (wp updateFreeIndex_cte_wp_at) apply clarsimp apply vcg apply wp @@ -496,19 +516,23 @@ shows pageBits_def max_free_index_def asid_low_bits_def) apply (case_tac cte,clarsimp simp:invs_valid_pspace') apply (frule(1) ctes_of_valid_cap'[OF _ invs_valid_objs']) - apply (clarsimp simp:valid_cap'_def asid_low_bits_def) + apply (clarsimp simp:valid_cap'_def asid_low_bits_def invs_urz) + apply (strengthen descendants_range_in_subseteq'[mk_strg I E] refl) + apply simp apply (intro conjI) - apply (simp add:is_aligned_def) + apply (simp add:is_aligned_def) apply (rule descendants_range_caps_no_overlapI'[where d=isdev and cref = parent]) - apply simp - apply (fastforce simp:cte_wp_at_ctes_of is_aligned_neg_mask_eq) - apply (clarsimp simp:is_aligned_neg_mask_eq) - apply (rule le_m1_iff_lt[THEN iffD1,THEN iffD1]) - apply (simp add:asid_bits_def) - apply (simp add:mask_def) + apply simp + apply (fastforce simp:cte_wp_at_ctes_of is_aligned_neg_mask_eq) + apply (clarsimp simp:is_aligned_neg_mask_eq) + apply (rule le_m1_iff_lt[THEN iffD1,THEN iffD1]) + apply (simp add:asid_bits_def) + apply (simp add:mask_def) + apply (clarsimp dest!: upto_intvl_eq) apply (wp deleteObjects_cte_wp_at'[where d=isdev and idx = idx and p = parent] deleteObjects_descendants[where d=isdev and p = parent and idx = idx] deleteObjects_invs'[where d=isdev and p = parent and idx = idx] + Detype_R.deleteObjects_descendants[where p = parent and idx = idx] deleteObjects_ct_active'[where d=isdev and cref = parent and idx = idx]) apply clarsimp apply vcg @@ -516,7 +540,7 @@ shows apply (frule cte_wp_at_valid_objs_valid_cap', fastforce) apply (clarsimp simp:valid_cap'_def capAligned_def cte_wp_at_ctes_of descendants_range'_def2 empty_descendants_range_in') - apply (intro conjI) + apply (intro conjI; (rule refl)?) apply clarsimp apply (drule(1) cte_cap_in_untyped_range[where ptr = frame]) apply (fastforce simp: cte_wp_at_ctes_of) @@ -532,18 +556,15 @@ shows apply (clarsimp simp: typ_heap_simps cap_get_tag_isCap dest!: ccte_relation_ccap_relation) apply (clarsimp simp: is_aligned_mask max_free_index_def pageBits_def gen_framesize_to_H_def) + apply (rule conjI, rule UNIV_I)? + apply clarsimp? apply (erule_tac s = sa in rf_sr_ctes_of_cliftE) apply assumption apply (frule_tac s = sa in rf_sr_cte_relation) apply simp+ - apply (clarsimp simp:typ_heap_simps' ) - apply (rule context_conjI) - apply (rule cap_get_tag_isCap_unfolded_H_cap) - apply (fastforce dest!:ccte_relation_ccap_relation) - apply (frule cap_get_tag_isCap_unfolded_H_cap) - apply (intro conjI) - apply ((drule cap_get_tag_UntypedCap, - clarsimp simp:unat32_eq_of_nat word_bits_conv)+)[2] + apply (clarsimp simp:typ_heap_simps' region_is_bytes'_def[where sz=0]) + apply (frule ccte_relation_ccap_relation) + apply (clarsimp simp: cap_get_tag_isCap) apply (clarsimp simp: hrs_htd_update_def split_def pageBits_def split: split_if) @@ -553,18 +574,19 @@ shows apply assumption apply (simp add: ghost_assertion_data_get_gs_clear_region[unfolded o_def]) apply (drule valid_global_refsD_with_objSize, clarsimp)+ - apply clarsimp+ + apply (clarsimp simp: isCap_simps dest!: ccte_relation_ccap_relation) apply (cut_tac ptr=frame and bits=12 and s="globals_update (t_hrs_'_update (hrs_htd_update (typ_region_bytes frame 12))) s'" in typ_region_bytes_actually_is_bytes) apply (simp add: hrs_htd_update) + apply clarsimp apply (intro conjI) apply (clarsimp elim!:is_aligned_weaken) apply (simp add:is_aligned_def) apply (erule is_aligned_no_wrap',simp) apply (drule region_actually_is_bytes_dom_s[OF _ order_refl]) apply (simp add: hrs_htd_update_def split_def) - apply (clarsimp simp: region_actually_is_bytes_def hrs_htd_def - hrs_htd_update_def split_def) + apply (clarsimp simp: region_actually_is_bytes_def hrs_htd_update) + apply (simp add: hrs_htd_def hrs_htd_update_def split_def) apply (clarsimp simp: ccap_relation_def) apply (clarsimp simp: cap_asid_pool_cap_lift) apply (clarsimp simp: cap_to_H_def) @@ -860,7 +882,8 @@ lemma decodeARMPageTableInvocation_ccorres: apply (clarsimp simp: rf_sr_ksCurThread "StrictC'_thread_state_defs" mask_eq_iff_w2p word_size ct_in_state'_def st_tcb_at'_def - word_sle_def word_sless_def) + word_sle_def word_sless_def + typ_heap_simps') apply (clarsimp simp: cap_get_tag_isCap_ArchObject isCap_simps word_sle_def word_sless_def word_less_nat_alt) @@ -3703,7 +3726,7 @@ lemma decodeARMPageDirectoryInvocation_ccorres: apply (case_tac st,simp+) apply (frule cap_get_tag_isCap_unfolded_H_cap(15)) apply (clarsimp simp: cap_lift_page_directory_cap hd_conv_nth - cap_lift_page_table_cap + cap_lift_page_table_cap typ_heap_simps' cap_to_H_def cap_page_directory_cap_lift_def to_bool_def cap_page_table_cap_lift_def typ_heap_simps' shiftl_t2n[where n=2] field_simps @@ -3711,7 +3734,7 @@ lemma decodeARMPageDirectoryInvocation_ccorres: apply (intro conjI impI allI) apply (clarsimp simp:ThreadState_Restart_def less_mask_eq rf_sr_ksCurThread resolve_ret_rel_def framesize_from_to_H framesize_from_H_mask2 - to_option_def rel_option_alt_def to_bool_def + to_option_def rel_option_alt_def to_bool_def typ_heap_simps' split:option.splits if_splits | fastforce simp: mask_def | rule flushtype_relation_triv,simp add:isPageFlush_def isPDFlushLabel_def @@ -3727,6 +3750,7 @@ lemma decodeARMPageDirectoryInvocation_ccorres: apply (clarsimp simp:ThreadState_Restart_def less_mask_eq rf_sr_ksCurThread resolve_ret_rel_def framesize_from_to_H framesize_from_H_mask2 to_option_def rel_option_alt_def to_bool_def + typ_heap_simps' split:option.splits if_splits | fastforce simp: mask_def | rule flushtype_relation_triv,simp add:isPageFlush_def isPDFlushLabel_def @@ -3742,6 +3766,7 @@ lemma decodeARMPageDirectoryInvocation_ccorres: apply (clarsimp simp:ThreadState_Restart_def less_mask_eq rf_sr_ksCurThread resolve_ret_rel_def framesize_from_to_H framesize_from_H_mask2 to_option_def rel_option_alt_def to_bool_def + typ_heap_simps' split:option.splits if_splits | fastforce simp: mask_def | rule flushtype_relation_triv,simp add:isPageFlush_def isPDFlushLabel_def @@ -3757,6 +3782,7 @@ lemma decodeARMPageDirectoryInvocation_ccorres: by (clarsimp simp:ThreadState_Restart_def less_mask_eq rf_sr_ksCurThread resolve_ret_rel_def framesize_from_to_H framesize_from_H_mask2 to_option_def rel_option_alt_def to_bool_def + typ_heap_simps' split:option.splits if_splits | fastforce simp: mask_def | rule flushtype_relation_triv,simp add:isPageFlush_def isPDFlushLabel_def diff --git a/proof/crefine/CSpace_C.thy b/proof/crefine/CSpace_C.thy index 84b4ae658..077a14f71 100644 --- a/proof/crefine/CSpace_C.thy +++ b/proof/crefine/CSpace_C.thy @@ -638,7 +638,8 @@ lemma ccorres_updateMDB_set_mdbNext [corres]: apply (erule bexI [rotated]) apply (clarsimp simp add: rf_sr_def cstate_relation_def Let_def cpspace_relation_def cte_wp_at_ctes_of heap_to_user_data_def - cvariable_array_map_const_add_map_option[where f="tcb_no_ctes_proj"]) + cvariable_array_map_const_add_map_option[where f="tcb_no_ctes_proj"] + typ_heap_simps') apply (rule conjI) apply (erule (2) cspace_cte_relation_upd_mdbI) apply (simp add: cmdbnode_relation_def) @@ -685,7 +686,8 @@ lemma ccorres_updateMDB_set_mdbPrev [corres]: apply (erule bexI [rotated]) apply (clarsimp simp add: rf_sr_def cstate_relation_def Let_def cpspace_relation_def cte_wp_at_ctes_of heap_to_user_data_def - cvariable_array_map_const_add_map_option[where f="tcb_no_ctes_proj"]) + cvariable_array_map_const_add_map_option[where f="tcb_no_ctes_proj"] + typ_heap_simps') apply (rule conjI) apply (erule (2) cspace_cte_relation_upd_mdbI) apply (simp add: cmdbnode_relation_def) @@ -693,7 +695,7 @@ lemma ccorres_updateMDB_set_mdbPrev [corres]: subgoal for _ s' by (cases "v_' s' = 0"; simp) apply (erule_tac t = s'a in ssubst) apply (simp add: carch_state_relation_def cmachine_state_relation_def - h_t_valid_clift_Some_iff) + h_t_valid_clift_Some_iff typ_heap_simps') apply (erule (1) setCTE_tcb_case) apply clarsimp @@ -1935,48 +1937,6 @@ done (************************************************************************) -(*-------------------------------------------------------------------------------------*) -(* redefining Haskell definition of emptySlot so that it corresponds to the C code ----*) -(*-------------------------------------------------------------------------------------*) - -lemma emptySlot_def: - "emptySlot slot irq = (do - newCTE \ getCTE slot; - (if ((cteCap newCTE) \ NullCap) - then (do - - mdbNode \ return ( cteMDBNode newCTE); - prev \ return ( mdbPrev mdbNode); - next \ return ( mdbNext mdbNode); - updateMDB prev (\ mdb. mdb \ mdbNext := next \); - updateMDB next (\ mdb. mdb \ - mdbPrev := prev, - mdbFirstBadged := - mdbFirstBadged mdb \ mdbFirstBadged mdbNode \); - updateCap slot NullCap; - updateMDB slot (const nullMDBNode); - (case irq of - Some irq \ deletedIRQHandler irq - | None \ return () - ) - od) - else return () - ) - od)" - apply (simp add: emptySlot_def) - apply (rule bind_eqI) - apply (rule ext) - apply (case_tac "cteCap newCTE") - apply (simp_all add: bind_assoc) - done - - - - - - - - declare split_if [split del] @@ -2070,6 +2030,7 @@ lemma emptySlot_helper: apply (erule_tac t = s' in ssubst) apply (simp add: carch_state_relation_def cmachine_state_relation_def h_t_valid_clift_Some_iff cvariable_array_map_const_add_map_option[where f="tcb_no_ctes_proj"] + typ_heap_simps' cong: lifth_update) apply (erule (1) setCTE_tcb_case) @@ -2100,6 +2061,7 @@ lemma emptySlot_helper: apply (simp add: carch_state_relation_def cmachine_state_relation_def h_t_valid_clift_Some_iff cvariable_array_map_const_add_map_option[where f="tcb_no_ctes_proj"] + typ_heap_simps' cong: lifth_update) apply (erule (1) setCTE_tcb_case) @@ -2124,7 +2086,7 @@ lemma emptySlot_helper: apply (rule ccorres_return_Skip) apply simp apply (clarsimp simp: cmdbnode_relation_def) -done +done @@ -2358,78 +2320,256 @@ ML_command {*show_abbrevs true*} ML_command {*show_abbrevs false*} *) +lemmas ccorres_split_noop_lhs + = ccorres_split_nothrow[where c=Skip, OF _ ceqv_refl _ _ hoarep.Skip, + simplified ccorres_seq_skip] +(* FIXME: to SR_Lemmas *) +lemma region_is_bytes_subset: + "region_is_bytes' ptr sz htd + \ {ptr' ..+ sz'} \ {ptr ..+ sz} + \ region_is_bytes' ptr' sz' htd" + by (auto simp: region_is_bytes'_def) + +lemma intvl_both_le: + "\ a \ x; unat x + y \ unat a + b \ + \ {x ..+ y} \ {a ..+ b}" + apply (rule order_trans[OF _ intvl_sub_offset[where x="x - a"]]) + apply (simp, rule order_refl) + apply unat_arith + done + +lemma untypedZeroRange_idx_forward_helper: + "isUntypedCap cap + \ capFreeIndex cap \ idx + \ idx \ 2 ^ capBlockSize cap + \ valid_cap' cap s + \ (case (untypedZeroRange cap, untypedZeroRange (capFreeIndex_update (\_. idx) cap)) + of (Some (a, b), Some (a', b')) \ {a' ..+ unat (b' + 1 - a')} \ {a ..+ unat (b + 1 - a)} + | _ \ True)" + apply (clarsimp split: option.split) + apply (clarsimp simp: untypedZeroRange_def max_free_index_def Let_def + isCap_simps valid_cap_simps' capAligned_def + split: split_if_asm) + apply (erule subsetD[rotated], rule intvl_both_le) + apply (clarsimp simp: getFreeRef_def) + apply (rule word_plus_mono_right) + apply (rule PackedTypes.of_nat_mono_maybe_le) + apply (erule order_le_less_trans, rule power_strict_increasing, simp_all) + apply (erule is_aligned_no_wrap') + apply (rule word_of_nat_less, simp) + apply (simp add: getFreeRef_def) + apply (simp add: unat_plus_simple[THEN iffD1, OF is_aligned_no_wrap'] + word_of_nat_less) + apply (simp add: word_of_nat_le unat_sub + order_le_less_trans[OF _ power_strict_increasing] + unat_of_nat_eq[where 'a=32, folded word_bits_def]) + done + +lemma intvl_close_Un: + "y = x + of_nat n + \ ({x ..+ n} \ {y ..+ m}) = {x ..+ n + m}" + apply ((simp add: intvl_def, safe, simp_all, + simp_all only: of_nat_add[symmetric]); (rule exI, strengthen refl)) + apply simp_all + apply (rule ccontr) + apply (drule_tac x="k - n" in spec) + apply simp + done + +lemma untypedZeroRange_idx_backward_helper: + "isUntypedCap cap + \ idx \ capFreeIndex cap + \ idx \ 2 ^ capBlockSize cap + \ valid_cap' cap s + \ (case untypedZeroRange (capFreeIndex_update (\_. idx) cap) + of None \ True | Some (a', b') \ + {a' ..+ unat (b' + 1 - a')} \ {capPtr cap + of_nat idx ..+ (capFreeIndex cap - idx)} + \ (case untypedZeroRange cap + of Some (a, b) \ {a ..+ unat (b + 1 - a)} + | None \ {}) + )" + apply (clarsimp split: option.split, intro impI conjI allI) + apply (rule intvl_both_le; clarsimp simp: untypedZeroRange_def + max_free_index_def Let_def + isCap_simps valid_cap_simps' capAligned_def + split: split_if_asm) + apply (clarsimp simp: getFreeRef_def) + apply (clarsimp simp: getFreeRef_def) + apply (simp add: word_of_nat_le unat_sub + order_le_less_trans[OF _ power_strict_increasing] + unat_of_nat_eq[where 'a=32, folded word_bits_def]) + apply (subst intvl_close_Un) + apply (clarsimp simp: untypedZeroRange_def + max_free_index_def Let_def + getFreeRef_def + split: split_if_asm) + apply (clarsimp simp: untypedZeroRange_def + max_free_index_def Let_def + getFreeRef_def isCap_simps valid_cap_simps' + split: split_if_asm) + apply (simp add: word_of_nat_le unat_sub capAligned_def + order_le_less_trans[OF _ power_strict_increasing] + order_le_less_trans[where x=idx] + unat_of_nat_eq[where 'a=32, folded word_bits_def]) + done + +lemma ctes_of_untyped_zero_rf_sr_case: + "\ ctes_of s p = Some cte; (s, s') \ rf_sr; + untyped_ranges_zero' s \ + \ case untypedZeroRange (cteCap cte) + of None \ True + | Some (start, end) \ region_is_zero_bytes start (unat ((end + 1) - start)) s'" + by (simp split: option.split add: ctes_of_untyped_zero_rf_sr) + +lemma gsUntypedZeroRanges_update_helper: + "(\, s) \ rf_sr + \ (zero_ranges_are_zero (gsUntypedZeroRanges \) (t_hrs_' (globals s)) + \ zero_ranges_are_zero (f (gsUntypedZeroRanges \)) (t_hrs_' (globals s))) + \ (gsUntypedZeroRanges_update f \, s) \ rf_sr" + by (clarsimp simp: rf_sr_def cstate_relation_def Let_def) + +lemma heap_list_zero_Ball_intvl: + "heap_list_is_zero hmem ptr n = (\x \ {ptr ..+ n}. hmem x = 0)" + apply safe + apply (erule heap_list_h_eq_better) + apply (simp add: heap_list_rpbs) + apply (rule trans[OF heap_list_h_eq2 heap_list_rpbs]) + apply simp + done + +lemma untypedZeroRange_not_device: + "untypedZeroRange cap = Some r + \ \ capIsDevice cap" + by (clarsimp simp: untypedZeroRange_def) + +lemma updateTrackedFreeIndex_noop_ccorres: + "ccorres dc xfdc (cte_wp_at' ((\cap. isUntypedCap cap + \ idx \ 2 ^ capBlockSize cap + \ (capFreeIndex cap \ idx \ cap' = cap)) o cteCap) slot + and valid_objs' and untyped_ranges_zero') + {s. \ capIsDevice cap' \ region_is_zero_bytes (capPtr cap' + of_nat idx) (capFreeIndex cap' - idx) s} hs + (updateTrackedFreeIndex slot idx) Skip" + (is "ccorres dc xfdc ?P ?P' _ _ _") + apply (simp add: updateTrackedFreeIndex_def getSlotCap_def) + apply (rule ccorres_guard_imp) + apply (rule ccorres_pre_getCTE[where P="\rv. + cte_wp_at' (op = rv) slot and ?P" and P'="K ?P'"]) + apply (rule ccorres_from_vcg) + apply (rule allI, rule conseqPre, vcg) + apply (clarsimp simp: cte_wp_at_ctes_of) + apply (frule(1) ctes_of_valid') + apply (frule(2) ctes_of_untyped_zero_rf_sr_case) + apply (clarsimp simp: simpler_modify_def bind_def cte_wp_at_ctes_of) + apply (erule gsUntypedZeroRanges_update_helper) + apply (clarsimp simp: zero_ranges_are_zero_def + split: split_if) + apply (case_tac "(a, b) \ gsUntypedZeroRanges \") + apply (drule(1) bspec, simp) + apply (erule disjE_L) + apply (frule(3) untypedZeroRange_idx_forward_helper) + apply (clarsimp simp: isCap_simps valid_cap_simps') + apply (case_tac "untypedZeroRange (cteCap cte)") + apply (clarsimp simp: untypedZeroRange_def + valid_cap_simps' + max_free_index_def Let_def + split: split_if_asm) + apply clarsimp + apply (thin_tac "\ capIsDevice cap' \ P" for P) + apply (clarsimp split: option.split_asm) + apply (subst region_is_bytes_subset, simp+) + apply (subst heap_list_is_zero_mono2, simp+) + apply (frule untypedZeroRange_idx_backward_helper[where idx=idx], + simp+) + apply (clarsimp simp: isCap_simps valid_cap_simps') + apply (clarsimp split: option.split_asm) + apply (clarsimp dest!: untypedZeroRange_not_device) + apply (subst region_is_bytes_subset, simp+) + apply (subst heap_list_is_zero_mono2, simp+) + apply (simp add: region_is_bytes'_def heap_list_zero_Ball_intvl) + apply (clarsimp dest!: untypedZeroRange_not_device) + apply blast + apply (clarsimp simp: cte_wp_at_ctes_of) + apply clarsimp + done + +lemma updateTrackedFreeIndex_forward_noop_ccorres: + "ccorres dc xfdc (cte_wp_at' ((\cap. isUntypedCap cap + \ capFreeIndex cap \ idx \ idx \ 2 ^ capBlockSize cap) o cteCap) slot + and valid_objs' and untyped_ranges_zero') UNIV hs + (updateTrackedFreeIndex slot idx) Skip" + (is "ccorres dc xfdc ?P UNIV _ _ _") + apply (rule ccorres_name_pre) + apply (rule ccorres_guard_imp2, + rule_tac cap'="cteCap (the (ctes_of s slot))" in updateTrackedFreeIndex_noop_ccorres) + apply (clarsimp simp: cte_wp_at_ctes_of region_is_bytes'_def) + done + +lemma clearUntypedFreeIndex_noop_ccorres: + "ccorres dc xfdc (valid_objs' and untyped_ranges_zero') UNIV hs + (clearUntypedFreeIndex p) Skip" + apply (simp add: clearUntypedFreeIndex_def getSlotCap_def) + apply (rule ccorres_guard_imp) + apply (rule ccorres_pre_getCTE[where P="\rv. cte_wp_at' (op = rv) p + and valid_objs' and untyped_ranges_zero'" and P'="K UNIV"]) + apply (case_tac "cteCap cte", simp_all add: ccorres_guard_imp[OF ccorres_return_Skip])[1] + apply (rule ccorres_guard_imp, rule updateTrackedFreeIndex_forward_noop_ccorres) + apply (clarsimp simp: cte_wp_at_ctes_of max_free_index_def) + apply (frule(1) Finalise_R.ctes_of_valid') + apply (clarsimp simp: valid_cap_simps') + apply simp + apply (clarsimp simp: cte_wp_at_ctes_of) + apply simp + done lemma emptySlot_ccorres: "ccorres dc xfdc - (valid_mdb' and pspace_aligned') + (valid_mdb' and valid_objs' and pspace_aligned' + and untyped_ranges_zero') (UNIV \ {s. slot_' s = Ptr slot} \ {s. irq_opt_relation irq (irq_' s)} ) [] (emptySlot slot irq) (Call emptySlot_'proc)" - apply (cinit lift: slot_' irq_' simp del: return_bind) + apply (cinit lift: slot_' irq_' simp: case_Null_If) - -- "***Main goal***" - apply (unfold emptySlot_def) -- "unfolding the new definition" + -- "--- handle the clearUntypedFreeIndex" + apply (rule ccorres_split_noop_lhs, rule clearUntypedFreeIndex_noop_ccorres) -- "--- instruction: newCTE \ getCTE slot; ---" - apply (rule ccorres_pre_getCTE) + apply (rule ccorres_pre_getCTE) -- "--- instruction: CALL on C side" -(* apply (rule ccorres_Guard_Seq) *) - apply (rule ccorres_move_c_guard_cte) - apply csymbr + apply (rule ccorres_move_c_guard_cte) + apply csymbr + apply (rule ccorres_abstract_cleanup) + apply (rename_tac cap_tag) + apply (rule_tac P="(cap_tag = scast cap_null_cap) + = (cteCap newCTE = NullCap)" in ccorres_gen_asm2) + apply (simp del: Collect_const) - -- "--- instruction: if-then-else / IF-THEN-ELSE " - apply (rule ccorres_guard_imp2) - - -- "*** Main Goal ***" - apply (rule ccorres_cond2) + -- "--- instruction: if-then-else / IF-THEN-ELSE " + apply (rule ccorres_cond2'[where R=\]) -- "*** link between abstract and concrete conditionals ***" - - apply clarsimp - apply (subgoal_tac "\ cap'. ret__unsigned = cap_get_tag cap' - \ ccap_relation (cteCap rv) cap' ") - apply (insert not_NullCap_eq_not_cap_null_cap)[1] - subgoal by (clarsimp) - apply assumption -(* apply clarsimp - apply (subgoal_tac "\ cap'. ret__unsigned_long = cap_get_tag cap' - \ ccap_relation (cteCap newCTE) cap' ") - prefer 2 apply assumption - apply clarsimp - apply (simp add: not_NullCap_eq_not_cap_null_cap)*) + apply (clarsimp split: split_if) -- "*** proof for the 'else' branch (return () and SKIP) ***" prefer 2 - apply (ctac add: ccorres_return_Skip) + apply (ctac add: ccorres_return_Skip[unfolded dc_def]) -- "*** proof for the 'then' branch ***" - -- "---instructions: multiple undefined on C side" + -- "---instructions: multiple on C side, including mdbNode fetch" apply (rule ccorres_rhs_assoc)+ -- "we have to do it here because the first assoc did not apply inside the then block" - apply csymbr - apply csymbr - apply csymbr + apply (rule ccorres_move_c_guard_cte | csymbr)+ + apply (rule ccorres_symb_exec_r) + apply (rule_tac xf'="mdbNode_'" in ccorres_abstract, ceqv) + apply (rename_tac "cmdbNode") + apply (rule_tac P="cmdbnode_relation (cteMDBNode newCTE) cmdbNode" + in ccorres_gen_asm2) + apply csymbr+ - -- "--- instruction : mdbNode \ return (cteMDBNode rv)" - apply ctac - -- "*** Main goal ***" - -- "--- instruction: CALL on C side" - apply csymbr - - -- "--- instruction: prev \ return (mdbPrev mdbNode); ---" - -- "--- next \ return (mdbNext mdbNode); ---" - apply (simp del: Collect_const) -- "to simplify the returns and replace them in the rest\" - - -- "--- instruction: assignements on C side" - apply csymbr - apply (erule_tac t = ret__unsigned in ssubst) - apply csymbr - apply (erule_tac t = ret__unsigned in ssubst) - apply csymbr -- "--- instruction: updateMDB (mdbPrev rva) (mdbNext_update \) but with Ptr\\ NULL on C side" apply (simp only:Ptr_not_null_pointer_not_zero) --"replaces Ptr p \ NULL with p\0" @@ -2439,7 +2579,7 @@ lemma emptySlot_ccorres: -- "here ctac alone does not apply because the subgoal generated by the rule are not solvable by simp" -- "so we have to use (no_simp) (or apply (rule ccorres_split_nothrow))" - apply (simp add: cmdbnode_relation_def) + apply (simp add: cmdbnode_relation_def) apply assumption -- "*** Main goal ***" -- "--- instruction: updateMDB (mdbNext rva) @@ -2476,61 +2616,32 @@ lemma emptySlot_ccorres: -- "C pre/post for the two nested updates " apply (simp add: Collect_const_mem ccap_relation_NullCap_iff) -- "Haskell pre/post for (updateMDB (mdbPrev rva) (mdbNext_update (\_. mdbNext rva)))" - apply wp + apply (simp, wp) -- "C pre/post for (updateMDB (mdbPrev rva) (mdbNext_update (\_. mdbNext rva)))" apply simp+ - -- "Haskell pre/post for mdbNode \ return (cteMDBNode rv)" - - apply wp - -- "C pre/post for mdbNode \ return (cteMDBNode rv)" - apply vcg - - - -- "*************************************************************" - -- "*** generalised precondition for ccorres_cond application ***" - -- "*************************************************************" - --- "from here, really big goals" - - apply (clarsimp simp: typ_heap_simps Collect_const_mem split del: split_if) - - -- "instantiating the Haskell precondition that has been generalised" - apply (subgoal_tac "ctes_of s slot = Some rv \ (cte_at' slot and valid_mdb' and pspace_aligned') s") - prefer 2 - apply assumption - apply clarsimp - - -- "instantiating the C precondition that has been generalised" - apply (subgoal_tac "s' \ UNIV") - prefer 2 - apply assumption - - -- "generating cslift" - apply (frule (1) rf_sr_ctes_of_clift) - apply (clarsimp simp: typ_heap_simps) - - -- "generating the hypotheses regarding validity, c_guard and cslift for Prev and Next" - apply (intro conjI) - apply (rule_tac x="cte_C.cap_C cte'" in exI) - apply (drule (2) rf_sr_cte_relation) - apply (drule ccte_relation_ccap_relation, simp) - apply (frule (2) is_aligned_3_next) - apply (frule (2) is_aligned_3_prev) + apply vcg + apply (rule conseqPre, vcg) apply clarsimp - prefer 2 - apply (clarsimp, rule refl) apply simp - apply (clarsimp simp: cte_wp_at_ctes_of) -done - - - - - - + apply (wp hoare_vcg_all_lift hoare_vcg_imp_lift) + -- "final precondition proof" + apply (clarsimp simp: typ_heap_simps Collect_const_mem + cte_wp_at_ctes_of + split del: split_if) + apply (rule conjI) + -- "Haskell side" + apply (simp add: is_aligned_3_prev is_aligned_3_next) + -- "C side" + apply (clarsimp simp: map_comp_Some_iff typ_heap_simps) + apply (subst cap_get_tag_isCap) + apply (rule ccte_relation_ccap_relation) + apply (simp add: ccte_relation_def c_valid_cte_def + cl_valid_cte_def c_valid_cap_def) + apply simp + done (************************************************************************) diff --git a/proof/crefine/Detype_C.thy b/proof/crefine/Detype_C.thy index 6a9dc26fa..974ebd20d 100644 --- a/proof/crefine/Detype_C.thy +++ b/proof/crefine/Detype_C.thy @@ -1137,24 +1137,6 @@ lemma tcb_ptr_to_ctcb_ptr_comp: apply (simp add: tcb_ptr_to_ctcb_ptr_def) done -lemma ccorres_assume_pre: - assumes "\s. P s \ ccorres r xf (P and (\s'. s' = s)) P' hs H C" - shows "ccorres r xf P P' hs H C" - apply (clarsimp simp: ccorres_underlying_def) - apply (frule assms) - apply (simp add: ccorres_underlying_def) - apply blast - done - -lemma ccorres_name_pre: - "(\s. P s \ ccorresG rf_sr \ r xf (\s'. s' = s) P' hs H C) \ ccorresG rf_sr \ r xf P P' hs H C" - apply (rule ccorres_assume_pre) - apply (rule ccorres_guard_imp) - apply fastforce - apply simp - apply simp - done - lemma tcb_ptr_to_ctcb_ptr_to_Ptr: "tcb_ptr_to_ctcb_ptr ` {p..+b} = Ptr ` {p + ctcb_offset..+b}" apply (simp add: tcb_ptr_to_ctcb_ptr_comp image_comp [symmetric]) @@ -1631,6 +1613,14 @@ lemma cvariable_array_map_relation_detype: apply (simp add: h_t_array_valid_typ_region_bytes) done +lemma zero_ranges_are_zero_typ_region_bytes: + "zero_ranges_are_zero rs hrs + \ zero_ranges_are_zero rs (hrs_htd_update (typ_region_bytes ptr bits) hrs)" + apply (clarsimp simp: zero_ranges_are_zero_def) + apply (drule(1) bspec) + apply (clarsimp simp: region_is_bytes'_def typ_region_bytes_def hrs_htd_update) + done + lemma deleteObjects_ccorres': notes if_cong[cong] shows @@ -2091,7 +2081,8 @@ proof - by (clarsimp simp: rf_sr_def cstate_relation_def Let_def psu_restrict cpspace_relation_def carch_state_relation_def cmachine_state_relation_def - hrs_htd_update htd_safe_typ_region_bytes) + hrs_htd_update htd_safe_typ_region_bytes + zero_ranges_are_zero_typ_region_bytes) qed abbreviation (input) diff --git a/proof/crefine/Fastpath_C.thy b/proof/crefine/Fastpath_C.thy index 217e4db23..8d0d8be3b 100644 --- a/proof/crefine/Fastpath_C.thy +++ b/proof/crefine/Fastpath_C.thy @@ -1196,9 +1196,10 @@ lemma thread_state_ptr_set_tsType_np_spec: tsType_CL (thread_state_lift thread_state) = tsType_' s \ tcbQueued_CL (thread_state_lift thread_state) = tcbQueued_CL (thread_state_lift (tcbState_C (the (cslift s (ptr s))))) \ - cslift t = cslift s(ptr s \ the (cslift s (ptr s))\tcbState_C := thread_state\)) - \ types_proofs.cslift_all_but_tcb_C t s - \ hrs_htd (t_hrs_' (globals t)) = hrs_htd (t_hrs_' (globals s))}" + t_hrs_' (globals t) = hrs_mem_update (heap_update (ptr s) + (the (cslift s (ptr s))\tcbState_C := thread_state\)) + (t_hrs_' (globals s)) + )}" apply (intro allI, rule conseqPre, vcg) apply (clarsimp simp: ptr_def) apply (clarsimp simp: h_t_valid_clift_Some_iff) @@ -1224,9 +1225,10 @@ lemma thread_state_ptr_mset_blockingObject_tsType_spec: \ blockingObject_CL (thread_state_lift thread_state) = ep_ref_' s \ tcbQueued_CL (thread_state_lift thread_state) = tcbQueued_CL (thread_state_lift (tcbState_C (the (cslift s (ptr s))))) - \ cslift t = cslift s(ptr s \ the (cslift s (ptr s))\tcbState_C := thread_state\)) - \ types_proofs.cslift_all_but_tcb_C t s - \ hrs_htd (t_hrs_' (globals t)) = hrs_htd (t_hrs_' (globals s))}" + \ t_hrs_' (globals t) = hrs_mem_update (heap_update (ptr s) + (the (cslift s (ptr s))\tcbState_C := thread_state\)) + (t_hrs_' (globals s)) + )}" apply (intro allI, rule conseqPre, vcg) apply (clarsimp simp: ptr_def) apply (frule h_t_valid_c_guard_cparent, simp+) @@ -1253,9 +1255,10 @@ lemma mdb_node_ptr_mset_mdbNext_mdbRevocable_mdbFirstBadged_spec: mdb_node_lift mdb_node = mdb_node_lift (cteMDBNode_C (the (cslift s (ptr s)))) \ mdbNext_CL := mdbNext___unsigned_long_' s, mdbRevocable_CL := mdbRevocable___unsigned_long_' s, mdbFirstBadged_CL := mdbFirstBadged___unsigned_long_' s \ - \ cslift t = cslift s(ptr s \ the (cslift s (ptr s)) \ cteMDBNode_C := mdb_node \)) - \ types_proofs.cslift_all_but_cte_C t s - \ hrs_htd (t_hrs_' (globals t)) = hrs_htd (t_hrs_' (globals s))}" + \ t_hrs_' (globals t) = hrs_mem_update (heap_update (ptr s) + (the (cslift s (ptr s)) \ cteMDBNode_C := mdb_node \)) + (t_hrs_' (globals s)) + )}" apply (intro allI, rule conseqPre, vcg) apply (clarsimp simp: ptr_def) apply (clarsimp simp: h_t_valid_clift_Some_iff) @@ -1280,9 +1283,10 @@ lemma mdb_node_ptr_set_mdbPrev_np_spec: {t. (\mdb_node. mdb_node_lift mdb_node = mdb_node_lift (cteMDBNode_C (the (cslift s (ptr s)))) \ mdbPrev_CL := mdbPrev___unsigned_long_' s \ - \ cslift t = cslift s(ptr s \ the (cslift s (ptr s)) \ cteMDBNode_C := mdb_node \)) - \ types_proofs.cslift_all_but_cte_C t s - \ hrs_htd (t_hrs_' (globals t)) = hrs_htd (t_hrs_' (globals s))}" + \ t_hrs_' (globals t) = hrs_mem_update (heap_update (ptr s) + (the (cslift s (ptr s)) \ cteMDBNode_C := mdb_node \)) + (t_hrs_' (globals s)) + )}" apply (intro allI, rule conseqPre, vcg) apply (clarsimp simp: ptr_def) apply (clarsimp simp: h_t_valid_clift_Some_iff) @@ -1305,9 +1309,10 @@ lemma cap_reply_cap_ptr_new_np_spec2: {t. (\cap. cap_lift cap = Some (Cap_reply_cap \ capReplyMaster_CL = capReplyMaster___unsigned_long_' s, capTCBPtr_CL = capTCBPtr___unsigned_long_' s \) - \ cslift t = cslift s(ptr s \ the (cslift s (ptr s)) \ cte_C.cap_C := cap \)) - \ types_proofs.cslift_all_but_cte_C t s - \ hrs_htd (t_hrs_' (globals t)) = hrs_htd (t_hrs_' (globals s))}" + \ t_hrs_' (globals t) = hrs_mem_update (heap_update (ptr s) + (the (cslift s (ptr s)) \ cte_C.cap_C := cap \)) + (t_hrs_' (globals s)) + )}" apply (intro allI, rule conseqPre, vcg) apply (clarsimp simp: ptr_def) apply (clarsimp simp: h_t_valid_clift_Some_iff word_sle_def) @@ -1331,9 +1336,10 @@ lemma endpoint_ptr_mset_epQueue_tail_state_spec: {t. (\endpoint. endpoint_lift endpoint = endpoint_lift (the (cslift s (ep_ptr_' s))) \ endpoint_CL.state_CL := state_' s, epQueue_tail_CL := epQueue_tail_' s \ - \ cslift t = cslift s(ep_ptr_' s \ endpoint)) - \ types_proofs.cslift_all_but_endpoint_C t s - \ hrs_htd (t_hrs_' (globals t)) = hrs_htd (t_hrs_' (globals s))}" + \ t_hrs_' (globals t) = hrs_mem_update (heap_update (ep_ptr_' s) + endpoint) + (t_hrs_' (globals s)) + )}" apply (intro allI, rule conseqPre, vcg) apply (clarsimp simp: h_t_valid_clift_Some_iff typ_heap_simps' word_sle_def word_sless_def) @@ -1350,9 +1356,10 @@ lemma endpoint_ptr_set_epQueue_head_np_spec: {t. (\endpoint. endpoint_lift endpoint = endpoint_lift (the (cslift s (ep_ptr_' s))) \ epQueue_head_CL := epQueue_head_' s \ - \ cslift t = cslift s(ep_ptr_' s \ endpoint)) - \ types_proofs.cslift_all_but_endpoint_C t s - \ hrs_htd (t_hrs_' (globals t)) = hrs_htd (t_hrs_' (globals s))}" + \ t_hrs_' (globals t) = hrs_mem_update (heap_update (ep_ptr_' s) + endpoint) + (t_hrs_' (globals s)) + )}" apply (intro allI, rule conseqPre, vcg) apply (clarsimp simp: h_t_valid_clift_Some_iff typ_heap_simps' word_sless_def word_sle_def) @@ -1675,19 +1682,22 @@ lemma fastpath_dequeue_ccorres: typ_heap_simps' endpoint_state_defs) apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def) apply (rule conjI) - apply (clarsimp simp: cpspace_relation_def update_ep_map_tos) + apply (clarsimp simp: cpspace_relation_def update_ep_map_tos + typ_heap_simps') apply (erule(1) cpspace_relation_ep_update_ep2) apply (simp add: cendpoint_relation_def endpoint_state_defs) apply simp apply (simp add: carch_state_relation_def cmachine_state_relation_def - h_t_valid_clift_Some_iff update_ep_map_tos) + h_t_valid_clift_Some_iff update_ep_map_tos + typ_heap_simps') apply (clarsimp simp: neq_Nil_conv cendpoint_relation_def Let_def isRecvEP_endpoint_case tcb_queue_relation'_def typ_heap_simps' endpoint_state_defs) apply (clarsimp simp: is_aligned_weaken[OF is_aligned_tcb_ptr_to_ctcb_ptr] tcb_at_not_NULL) apply (drule(1) obj_at_cslift_tcb)+ - apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def) + apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def + typ_heap_simps' tcb_at_not_NULL[OF obj_at'_weakenE, OF _ TrueI]) apply (rule conjI) apply (clarsimp simp: cpspace_relation_def update_ep_map_tos update_tcb_map_tos typ_heap_simps') @@ -1978,21 +1988,21 @@ lemma ccorres_updateCap [corres]: apply clarsimp done - lemma setCTE_rf_sr: "\ (\, s) \ rf_sr; ctes_of \ ptr = Some cte''; - cslift s' = ((cslift s)(cte_Ptr ptr \ cte')); + t_hrs_' (globals s') = hrs_mem_update + (heap_update (cte_Ptr ptr) cte') + (t_hrs_' (globals s)); ccte_relation cte cte'; - types_proofs.cslift_all_but_cte_C s' s; - hrs_htd (t_hrs_' (globals s')) = hrs_htd (t_hrs_' (globals s)); (globals s')\ t_hrs_' := undefined \ = (globals s)\ t_hrs_' := undefined \ \ - \ \x\fst (setCTE ptr cte \). (snd x, s') \ rf_sr" apply (rule fst_setCTE[OF ctes_of_cte_at], assumption) apply (erule rev_bexI) + apply clarsimp + apply (frule(1) rf_sr_ctes_of_clift) apply (subgoal_tac "\hrs. globals s' = globals s \ t_hrs_' := hrs \") apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def @@ -2012,10 +2022,10 @@ lemma setCTE_rf_sr: lemma getCTE_setCTE_rf_sr: "\ (\, s) \ rf_sr; ctes_of \ ptr = Some cte; - cslift s' = ((cslift s)(cte_Ptr ptr \ cte')); + t_hrs_' (globals s') = hrs_mem_update + (heap_update (cte_Ptr ptr) cte') + (t_hrs_' (globals s)); ccte_relation (f cte) cte'; - types_proofs.cslift_all_but_cte_C s' s; - hrs_htd (t_hrs_' (globals s')) = hrs_htd (t_hrs_' (globals s)); (globals s')\ t_hrs_' := undefined \ = (globals s)\ t_hrs_' := undefined \ \ @@ -2310,7 +2320,6 @@ lemma fastpath_call_ccorres: done show ?thesis - using [[goals_limit = 3]] apply (cinit lift: cptr_' msgInfo_') apply (simp add: catch_liftE_bindE unlessE_throw_catch_If unifyFailure_catch_If catch_liftE @@ -2568,10 +2577,10 @@ lemma fastpath_call_ccorres: apply (simp add: ctcb_relation_def cthread_state_relation_def) apply simp apply (rule conjI, erule cready_queues_relation_not_queue_ptrs) - apply (rule ext, simp split: split_if) - apply (rule ext, simp split: split_if) + apply (rule ext, simp split: split_if add: typ_heap_simps') + apply (rule ext, simp split: split_if add: typ_heap_simps') apply (simp add: carch_state_relation_def cmachine_state_relation_def - h_t_valid_clift_Some_iff map_comp_update projectKO_opt_tcb + typ_heap_simps' map_comp_update projectKO_opt_tcb cvariable_relation_upd_const ko_at_projectKO_opt) apply ceqv apply (rule ccorres_abstract_ksCurThread, ceqv) @@ -2679,7 +2688,7 @@ lemma fastpath_call_ccorres: apply (rule ext, simp split: split_if) apply (rule ext, simp split: split_if) apply (simp add: carch_state_relation_def cmachine_state_relation_def - h_t_valid_clift_Some_iff map_comp_update projectKO_opt_tcb + typ_heap_simps' map_comp_update projectKO_opt_tcb cvariable_relation_upd_const ko_at_projectKO_opt) apply ceqv apply (simp only: bind_assoc[symmetric]) @@ -3076,7 +3085,7 @@ lemma fastpath_reply_recv_ccorres: del: Collect_const cong: call_ignore_cong) apply (rule ccorres_Cond_rhs_Seq) apply simp - apply (rule ccorres_split_throws) + apply (rule ccorres_split_throws) apply (fold dc_def)[1] apply (rule ccorres_call_hSkip) apply (rule slowpath_ccorres) @@ -3324,7 +3333,7 @@ lemma fastpath_reply_recv_ccorres: apply (rule ext, simp split: split_if) apply (rule ext, simp split: split_if) apply (simp add: carch_state_relation_def cmachine_state_relation_def - h_t_valid_clift_Some_iff map_comp_update projectKO_opt_tcb + typ_heap_simps' map_comp_update projectKO_opt_tcb cvariable_relation_upd_const ko_at_projectKO_opt) apply ceqv apply (rule ccorres_rhs_assoc2, rule ccorres_rhs_assoc2) @@ -3377,7 +3386,8 @@ lemma fastpath_reply_recv_ccorres: apply (rule cmap_relationE1[OF cmap_relation_cte], assumption+) apply (clarsimp simp: typ_heap_simps' split_def tcbCallerSlot_def tcb_cnode_index_defs tcb_ptr_to_ctcb_ptr_mask - cte_level_bits_def size_of_def) + cte_level_bits_def size_of_def + packed_heap_update_collapse_hrs) apply (rule setCTE_rf_sr, simp_all add: typ_heap_simps')[1] apply (clarsimp simp: ccte_relation_eq_ccap_relation makeObject_cte mdb_node_to_H_def nullMDBNode_def @@ -3402,7 +3412,7 @@ lemma fastpath_reply_recv_ccorres: apply (rule ext, simp split: split_if) apply (rule ext, simp split: split_if) apply (simp add: carch_state_relation_def cmachine_state_relation_def - h_t_valid_clift_Some_iff map_comp_update projectKO_opt_tcb + typ_heap_simps' map_comp_update projectKO_opt_tcb cvariable_relation_upd_const ko_at_projectKO_opt) apply ceqv apply (simp only: bind_assoc[symmetric]) @@ -5101,11 +5111,24 @@ lemma updateMDB_isolatable: (wp | simp)+) done +lemma clearUntypedFreeIndex_isolatable: + "thread_actions_isolatable idx (clearUntypedFreeIndex slot)" + apply (simp add: clearUntypedFreeIndex_def getSlotCap_def) + apply (rule thread_actions_isolatable_bind) + apply (rule getCTE_isolatable) + apply (simp split: capability.split, safe intro!: thread_actions_isolatable_return) + apply (simp add: updateTrackedFreeIndex_def getSlotCap_def) + apply (intro thread_actions_isolatable_bind getCTE_isolatable + modify_isolatable) + apply (wp | simp)+ + done + lemma emptySlot_isolatable: "thread_actions_isolatable idx (emptySlot slot None)" - apply (simp add: emptySlot_def updateCap_def + apply (simp add: emptySlot_def updateCap_def case_Null_If cong: if_cong) apply (intro thread_actions_isolatable_bind[OF _ _ hoare_pre(1)] + clearUntypedFreeIndex_isolatable thread_actions_isolatable_if getCTE_isolatable setCTE_isolatable thread_actions_isolatable_return @@ -5825,6 +5848,36 @@ lemma setEndpoint_updateCap_pivot[unfolded K_bind_def]: setEndpoint_getCTE_pivot setEndpoint_setCTE_pivot) +lemma modify_setEndpoint_pivot[unfolded K_bind_def]: + "\ \ksf s. ksPSpace_update ksf (sf s) = sf (ksPSpace_update ksf s) \ + \ (do modify sf; setEndpoint p val; f od) = + (do setEndpoint p val; modify sf; f od)" + apply (subgoal_tac "\s. ep_at' p (sf s) = ep_at' p s") + apply (simp add: setEndpoint_def setObject_modify_assert + bind_assoc fun_eq_iff + exec_gets exec_modify assert_def + split: split_if) + apply atomize + apply clarsimp + apply (drule_tac x="\_. ksPSpace s" in spec) + apply (drule_tac x="s" in spec) + apply (drule_tac f="ksPSpace" in arg_cong) + apply simp + apply (metis obj_at'_pspaceI) + done + +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 + split: capability.split + | rule bind_cong[OF refl] allI impI + bind_apply_cong[OF refl])+ + lemma emptySlot_setEndpoint_pivot[unfolded K_bind_def]: "(do emptySlot slot None; setEndpoint p val; f od) = (do setEndpoint p val; emptySlot slot None; f od)" @@ -5833,6 +5886,8 @@ lemma emptySlot_setEndpoint_pivot[unfolded K_bind_def]: setEndpoint_getCTE_pivot setEndpoint_updateCap_pivot setEndpoint_updateMDB_pivot + case_Null_If + setEndpoint_clearUntypedFreeIndex_pivot split: split_if | rule bind_apply_cong[OF refl])+ done @@ -5907,12 +5962,29 @@ lemma setCTE_updateCapMDB: apply (auto simp: mask_out_sub_mask) done +lemma clearUntypedFreeIndex_simple_rewrite: + "monadic_rewrite True False + (cte_wp_at' (Not o isUntypedCap o cteCap) slot) + (clearUntypedFreeIndex slot) (return ())" + apply (simp add: clearUntypedFreeIndex_def getSlotCap_def) + apply (rule monadic_rewrite_name_pre) + apply (clarsimp simp: cte_wp_at_ctes_of) + apply (rule monadic_rewrite_imp) + apply (rule_tac rv=cte in monadic_rewrite_symb_exec_l_known, wp) + apply (simp split: capability.split, + strengthen monadic_rewrite_refl, simp) + apply clarsimp + apply (wp getCTE_wp') + apply (clarsimp simp: cte_wp_at_ctes_of) + done + lemma emptySlot_replymaster_rewrite[OF refl]: "mdbn = cteMDBNode cte \ monadic_rewrite True False ((\_. mdbNext mdbn = 0 \ mdbPrev mdbn \ 0) and ((\_. cteCap cte \ NullCap) and (cte_wp_at' (op = cte) slot + and cte_wp_at' (\cte. isReplyCap (cteCap cte)) slot and cte_wp_at' (\cte. isReplyCap (cteCap cte) \ capReplyMaster (cteCap cte)) (mdbPrev mdbn) and (\s. reply_masters_rvk_fb (ctes_of s)) @@ -5926,8 +5998,12 @@ lemma emptySlot_replymaster_rewrite[OF refl]: apply (rule monadic_rewrite_imp) apply (rule_tac P="slot \ 0" in monadic_rewrite_gen_asm) apply (clarsimp simp: emptySlot_def setCTE_updateCapMDB) + apply (rule monadic_rewrite_trans) + apply (rule monadic_rewrite_bind_head) + apply (rule clearUntypedFreeIndex_simple_rewrite) + apply simp apply (rule_tac rv=cte in monadic_rewrite_symb_exec_l_known, wp empty_fail_getCTE) - apply (simp add: updateMDB_def Let_def bind_assoc makeObject_cte) + apply (simp add: updateMDB_def Let_def bind_assoc makeObject_cte case_Null_If) apply (rule monadic_rewrite_bind_tail) apply (rule monadic_rewrite_bind) apply (rule_tac P="mdbFirstBadged (cteMDBNode ctea) \ mdbRevocable (cteMDBNode ctea)" @@ -5938,7 +6014,7 @@ lemma emptySlot_replymaster_rewrite[OF refl]: apply (rule monadic_rewrite_refl) apply (wp getCTE_wp') apply (clarsimp simp: cte_wp_at_ctes_of reply_masters_rvk_fb_def) - apply fastforce + apply (fastforce simp: isCap_simps) done (* FIXME: Move *) diff --git a/proof/crefine/Finalise_C.thy b/proof/crefine/Finalise_C.thy index e5d9865fe..c6482f087 100644 --- a/proof/crefine/Finalise_C.thy +++ b/proof/crefine/Finalise_C.thy @@ -340,7 +340,7 @@ lemma cancelAllIPC_ccorres: Let_def carch_state_relation_def carch_globals_def cmachine_state_relation_def) apply (clarsimp simp: cpspace_relation_def - update_ep_map_tos) + update_ep_map_tos typ_heap_simps') apply (erule(2) cpspace_relation_ep_update_ep) subgoal by (simp add: cendpoint_relation_def endpoint_state_defs) subgoal by simp @@ -389,7 +389,7 @@ lemma cancelAllIPC_ccorres: apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def carch_state_relation_def carch_globals_def cmachine_state_relation_def) - apply (clarsimp simp: cpspace_relation_def + apply (clarsimp simp: cpspace_relation_def typ_heap_simps' update_ep_map_tos) apply (erule(2) cpspace_relation_ep_update_ep) subgoal by (simp add: cendpoint_relation_def endpoint_state_defs) @@ -476,7 +476,7 @@ lemma cancelAllSignals_ccorres: apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def carch_state_relation_def carch_globals_def cmachine_state_relation_def) - apply (clarsimp simp: cpspace_relation_def + apply (clarsimp simp: cpspace_relation_def typ_heap_simps' update_ntfn_map_tos) apply (erule(2) cpspace_relation_ntfn_update_ntfn) subgoal by (simp add: cnotification_relation_def notification_state_defs Let_def) @@ -668,7 +668,8 @@ lemma doUnbindNotification_ccorres: apply (clarsimp simp: setNotification_def split_def) apply (rule bexI [OF _ setObject_eq]) apply (simp add: rf_sr_def cstate_relation_def Let_def init_def - cpspace_relation_def update_ntfn_map_tos) + typ_heap_simps' + cpspace_relation_def update_ntfn_map_tos) apply (elim conjE) apply (intro conjI) -- "tcb relation" @@ -676,7 +677,7 @@ lemma doUnbindNotification_ccorres: apply (clarsimp simp: cnotification_relation_def Let_def mask_def [where n=2] NtfnState_Waiting_def) apply (case_tac "ntfnObj rv", ((simp add: option_to_ctcb_ptr_def)+)[4]) - subgoal by (simp add: carch_state_relation_def) + subgoal by (simp add: carch_state_relation_def typ_heap_simps') subgoal by (simp add: cmachine_state_relation_def) subgoal by (simp add: h_t_valid_clift_Some_iff) subgoal by (simp add: objBits_simps) @@ -690,7 +691,7 @@ lemma doUnbindNotification_ccorres: apply vcg apply simp apply (erule(1) rf_sr_tcb_update_no_queue2) - apply (simp add: typ_heap_simps)+ + apply (simp add: typ_heap_simps')+ apply (simp add: tcb_cte_cases_def) apply (simp add: ctcb_relation_def option_to_ptr_def option_to_0_def) apply (simp add: invs'_def valid_state'_def) @@ -715,7 +716,8 @@ lemma doUnbindNotification_ccorres': apply (clarsimp simp: setNotification_def split_def) apply (rule bexI [OF _ setObject_eq]) apply (simp add: rf_sr_def cstate_relation_def Let_def init_def - cpspace_relation_def update_ntfn_map_tos) + typ_heap_simps' + cpspace_relation_def update_ntfn_map_tos) apply (elim conjE) apply (intro conjI) -- "tcb relation" @@ -725,7 +727,7 @@ lemma doUnbindNotification_ccorres': apply (fold_subgoals (prefix))[2] subgoal premises prems using prems by (case_tac "ntfnObj ntfn", (simp add: option_to_ctcb_ptr_def)+) - subgoal by (simp add: carch_state_relation_def) + subgoal by (simp add: carch_state_relation_def typ_heap_simps') subgoal by (simp add: cmachine_state_relation_def) subgoal by (simp add: h_t_valid_clift_Some_iff) subgoal by (simp add: objBits_simps) @@ -739,7 +741,7 @@ lemma doUnbindNotification_ccorres': apply vcg apply simp apply (erule(1) rf_sr_tcb_update_no_queue2) - apply (simp add: typ_heap_simps)+ + apply (simp add: typ_heap_simps')+ apply (simp add: tcb_cte_cases_def) apply (simp add: ctcb_relation_def option_to_ptr_def option_to_0_def) apply (simp add: invs'_def valid_state'_def) @@ -1221,7 +1223,7 @@ lemma deleteASID_ccorres: simp_all add: objBits_simps archObjSize_def pageBits_def)[1] apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def typ_heap_simps) apply (rule conjI) - apply (clarsimp simp: cpspace_relation_def typ_heap_simps + apply (clarsimp simp: cpspace_relation_def typ_heap_simps' update_asidpool_map_tos update_asidpool_map_to_asidpools) apply (rule cmap_relation_updI, simp_all)[1] @@ -1234,7 +1236,7 @@ lemma deleteASID_ccorres: subgoal by (simp add: asid_low_bits_def) subgoal by (simp add: carch_state_relation_def cmachine_state_relation_def carch_globals_def update_asidpool_map_tos - typ_heap_simps) + typ_heap_simps') apply (rule ccorres_pre_getCurThread) apply (ctac add: setVMRoot_ccorres) apply (simp add: cur_tcb'_def[symmetric]) @@ -1746,8 +1748,9 @@ lemma cteDeleteOne_ccorres: apply (simp add: dc_def[symmetric]) apply (ctac add: emptySlot_ccorres) apply (simp add: pred_conj_def finaliseCapTrue_standin_simple_def) - apply (strengthen invs_mdb_strengthen') - apply (wp typ_at_lifts isFinalCapability_inv) + apply (strengthen invs_mdb_strengthen' invs_urz) + apply (wp typ_at_lifts isFinalCapability_inv + | strengthen invs_valid_objs')+ apply (clarsimp simp: from_bool_def true_def irq_opt_relation_def invs_pspace_aligned' cte_wp_at_ctes_of) apply (erule(1) cmap_relationE1 [OF cmap_relation_cte]) diff --git a/proof/crefine/Invoke_C.thy b/proof/crefine/Invoke_C.thy index 8cb573a46..2c30bcce6 100644 --- a/proof/crefine/Invoke_C.thy +++ b/proof/crefine/Invoke_C.thy @@ -58,8 +58,7 @@ lemma setDomain_ccorres: apply vcg apply clarsimp apply (erule(1) rf_sr_tcb_update_no_queue2, - (simp add: typ_heap_simps)+, simp_all)[1] - apply (subst heap_update_field_hrs | fastforce intro: typ_heap_simps)+ + (simp add: typ_heap_simps')+, simp_all?)[1] apply (rule ball_tcb_cte_casesI, simp+) apply (simp add: ctcb_relation_def) apply (ctac(no_vcg) add: isRunnable_ccorres) @@ -1499,7 +1498,7 @@ lemma resetUntypedCap_gsCNodes_at_pt: \\rv s. P (gsCNodes s ptr)\, -" apply (simp add: resetUntypedCap_def unlessE_def) apply (rule hoare_pre) - apply (wp mapME_x_wp') + apply (wp mapME_x_wp' | simp add: unless_def)+ apply (wp hoare_vcg_const_imp_lift deleteObjects_gsCNodes_at_pt getSlotCap_wp) @@ -1688,38 +1687,700 @@ lemma cNodeNoOverlap_retype_have_size: apply clarsimp done +lemma range_cover_compare_bound_word: + "range_cover ptr sz sbit n + \ (of_nat n * 2 ^ sbit) + (ptr && mask sz) \ 2 ^ sz" + apply (simp add: word_le_nat_alt range_cover_unat + add.commute) + apply (frule range_cover.range_cover_compare_bound) + apply (simp add: range_cover.sz range_cover.unat_of_nat_shift) + done + +lemma from_to_bool_last_bit: + "from_bool (to_bool (x && 1)) = x && 1" + apply (simp add: from_bool_def to_bool_and_1 + split: bool.split) + apply (safe intro!: word_eqI, auto) + done + +lemma isUntypedCap_ccap_relation_helper: + "ccap_relation cap ccap + \ isUntypedCap cap + \ cap_get_tag ccap = scast cap_untyped_cap + \ cap_lift ccap = Some (Cap_untyped_cap (cap_untyped_cap_lift ccap)) + \ cap_untyped_cap_lift ccap = + \ capFreeIndex_CL = of_nat (capFreeIndex cap) >> 4, + capIsDevice_CL = from_bool (capIsDevice cap), + capBlockSize_CL = of_nat (capBlockSize cap), + capPtr_CL = capPtr cap\" + apply (simp add: cap_get_tag_isCap[symmetric]) + apply (frule(1) cap_get_tag_UntypedCap[THEN iffD1]) + apply (frule cap_lift_untyped_cap) + apply (simp add: cap_untyped_cap_lift_def) + apply (clarsimp simp: shiftl_shiftr1 word_size from_to_bool_last_bit) + apply (simp add: mask_def word_bw_assocs ) + done + +lemma pspace_no_overlap_underlying_zero_update: + "pspace_no_overlap' ptr sz s + \ invs' s + \ S \ {ptr .. (ptr && ~~ mask sz) + 2 ^ sz - 1} + \ s\ksMachineState := underlying_memory_update + (\m x. if x \ S then 0 else m x) (ksMachineState s)\ + = s" + apply (subgoal_tac "\x \ S. underlying_memory (ksMachineState s) x = 0") + apply (cases "ksMachineState s") + apply (cases s, simp add: fun_eq_iff split: split_if) + apply (clarsimp split: split_if_asm) + apply (erule pspace_no_overlap_underlying_zero) + apply (simp add: invs'_def valid_state'_def) + apply blast + done + +lemma addrFromPPtr_mask: + "n \ 28 + \ addrFromPPtr ptr && mask n = ptr && mask n" + apply (simp add: addrFromPPtr_def physMappingOffset_def kernelBase_addr_def + ARM.physBase_def) + apply word_bitwise + apply simp + done + +lemma clearMemory_untyped_ccorres: + "ccorres dc xfdc ((\s. invs' s + \ (\cap. cte_wp_at' (\cte. cteCap cte = cap) ut_slot s + \ isUntypedCap cap + \ {ptr ..+ 2 ^ sz} \ untypedRange cap + \ pspace_no_overlap' (capPtr cap) (capBlockSize cap) s) + \ 2 ^ sz \ gsMaxObjectSize s) + and (\_. is_aligned ptr sz \ sz \ 2 \ sz \ reset_chunk_bits)) + ({s. region_actually_is_bytes ptr (2 ^ sz) s} + \ {s. bits_' s = of_nat sz} + \ {s. ptr___ptr_to_unsigned_long_' s = Ptr ptr}) + [] + (doMachineOp (clearMemory ptr (2 ^ sz))) (Call clearMemory_'proc)" + (is "ccorres dc xfdc ?P ?P' [] ?m ?c") + apply (rule ccorres_gen_asm) + apply (cinit' lift: bits_' ptr___ptr_to_unsigned_long_') + apply (rule_tac P="ptr \ 0 \ sz < word_bits" + in ccorres_gen_asm) + apply (rule ccorres_Guard_Seq) + apply (simp add: clearMemory_def) + apply (simp add: doMachineOp_bind ef_storeWord) + apply (rule ccorres_split_nothrow_novcg_dc) + apply (rule_tac P="?P" and P'="{s. region_actually_is_bytes ptr (2 ^ sz) s}" + in ccorres_from_vcg) + apply (rule allI, rule conseqPre, vcg) + apply clarsimp + apply (clarsimp simp: isCap_simps valid_cap'_def capAligned_def + is_aligned_no_wrap'[OF _ word32_power_less_1] + unat_of_nat_eq word_bits_def) + apply (simp add: is_aligned_weaken is_aligned_triv[THEN is_aligned_weaken]) + apply (clarsimp simp: ghost_assertion_size_logic[unfolded o_def] + region_actually_is_bytes_dom_s) + apply (clarsimp simp: field_simps word_size_def mapM_x_storeWord_step + word_bits_def cte_wp_at_ctes_of) + apply (frule ctes_of_valid', clarify+) + apply (simp add: doMachineOp_def split_def exec_gets) + apply (simp add: select_f_def simpler_modify_def bind_def + valid_cap_simps' capAligned_def) + apply (subst pspace_no_overlap_underlying_zero_update, simp+) + apply (case_tac sz, simp_all)[1] + apply (case_tac nat, simp_all)[1] + apply (clarsimp dest!: region_actually_is_bytes) + apply (drule(1) rf_sr_rep0) + apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def + carch_state_relation_def cmachine_state_relation_def) + + apply csymbr + apply (ctac add: cleanCacheRange_PoU_ccorres[unfolded dc_def]) + apply wp + apply (simp add: guard_is_UNIV_def unat_of_nat + word_bits_def capAligned_def word_of_nat_less) + apply (clarsimp simp: cte_wp_at_ctes_of) + apply (frule ctes_of_valid', clarify+) + apply (clarsimp simp: isCap_simps valid_cap_simps' capAligned_def + word_of_nat_less reset_chunk_bits_def + word_bits_def unat_2p_sub_1) + apply (strengthen is_aligned_no_wrap'[where sz=sz] is_aligned_addrFromPPtr_n)+ + apply (simp add: addrFromPPtr_mask) + apply (cases "ptr = 0") + apply (drule subsetD, rule intvl_self, simp) + apply (simp split: split_if_asm) + apply simp + done + +lemma t_hrs_update_use_t_hrs: + "t_hrs_'_update f s + = (t_hrs_'_update (\_. f (t_hrs_' s)) $ s)" + by simp + +lemma name_seq_bound_helper: + "(\ CP n \ (\n' < n. CP n')) + \ (if \n. \ CP n + then simpl_sequence c' (map f [0 ..< (LEAST n. \ 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 word_two_power_neg_ineq: + "2 ^ m \ (0 :: 'a word) \ 2 ^ n \ - (2 ^ m :: ('a :: len) word)" + apply (cases "n < len_of TYPE('a)", simp_all add: power_overflow) + apply (cases "m < len_of TYPE('a)", simp_all add: power_overflow) + apply (simp add: word_le_nat_alt Aligned.unat_minus word_size) + apply (cases "len_of TYPE('a)", simp_all) + apply (simp add: less_Suc_eq_le) + apply (drule power_increasing[where a=2 and n=n] + power_increasing[where a=2 and n=m], simp)+ + apply (drule(1) add_le_mono) + apply simp + done + +lemma reset_name_seq_bound_helper: + fixes sz + fixes v :: "('a :: len) word" + defines "CP \ (\n. ~ (v && ~~ mask sz) + of_nat n * (-1 << sz) = + ((-1 :: 'a word) << sz))" + and "n \ Suc (unat (shiftR v sz))" + assumes vsz: "v + 1 < 2 ^ (len_of TYPE('a) - 1)" "2 ^ sz \ (0 :: 'a word)" + and vless: "v < v'" + shows "(\ CP n \ (\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="\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="\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 n="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: + "\ cteCap cte = capability.UntypedCap dev ptr sz idx; + i \ unat ((of_nat idx - 1 :: addr) div 2 ^ sz2); + sz2 \ sz; idx \ 0; + valid_cap' (cteCap cte) s + \ + \ of_nat i * 2 ^ sz2 < (2 ^ sz :: addr)" + apply (clarsimp simp: valid_cap_simps') + apply (rule word_less_power_trans2, simp_all) + apply (rule word_of_nat_less) + apply (erule order_le_less_trans) + apply (simp only: word_less_nat_alt[symmetric]) + apply (simp add: shiftr_div_2n_w[symmetric] word_size) + apply (rule shiftr_less_t2n) + apply (simp add: word_of_nat_le) + apply (rule of_nat_neq_0, simp) + apply (erule order_le_less_trans) + apply (rule power_strict_increasing, simp_all) + done + +lemma typ_region_bytes_dom_s: + "S \ {ptr ..+ 2 ^ bits} + \ S \ {SIndexVal, SIndexTyp 0} \ dom_s (typ_region_bytes ptr bits htd)" + apply (clarsimp simp: typ_region_bytes_def dom_s_def) + apply fastforce + done + +lemma aligned_intvl_offset_subset: + assumes al: "is_aligned (ptr :: 'a :: len word) sz" and al': "is_aligned x sz'" + and szv: "sz' \ sz" and xsz: "x < 2 ^ sz" + shows "{ptr + x ..+ 2 ^ sz'} \ {ptr ..+ 2 ^ sz}" + apply (simp only: upto_intvl_eq al aligned_add_aligned[OF al al' szv]) + apply (rule aligned_range_offset_subset[OF al al' szv xsz]) + done + +lemma aligned_intvl_offset_subset_ran: + assumes al: "is_aligned (ptr :: 'a :: len word) sz" and al': "is_aligned x sz'" + and szv: "sz' \ sz" and xsz: "x < 2 ^ sz" + shows "{ptr + x ..+ 2 ^ sz'} \ {ptr .. ptr + 2 ^ sz - 1}" + apply (simp only: upto_intvl_eq al aligned_add_aligned[OF al al' szv]) + apply (rule aligned_range_offset_subset[OF al al' szv xsz]) + done + +lemma ccorres_req_Ex: + assumes v: "\s s'. \ (s, s') \ sr; P s; s' \ P' \ \ \v. Q v s \ Q' v s' \ V v" + and cc: "\v. V v \ ccorres_underlying sr \ r xf r' xf' (P and Q v) (P' \ {s. Q' v s}) hs H C" + shows "ccorres_underlying sr \ r xf r' xf' P P' hs H C" + apply (rule ccorres_name_pre) + apply (rule ccorres_name_pre_C) + apply (case_tac "(s, sa) \ sr") + apply (drule(2) v, clarsimp) + apply (rule ccorres_guard_imp2, erule cc) + apply auto[1] + apply (rule ccorresI', simp) + done + +lemma is_aligned_mask_out_add_eq: + "is_aligned p n + \ (p + x) && ~~ mask n = p + (x && ~~ mask n)" + by (simp add: mask_out_sub_mask mask_add_aligned) + +lemmas is_aligned_mask_out_add_eq_sub + = is_aligned_mask_out_add_eq[where x="a - b" for a b, simplified field_simps] + +lemma aligned_bump_down: + "is_aligned x n + \ (x - 1) && ~~ mask n = x - 2 ^ n" + apply (frule is_aligned_mask_out_add_eq[where x="-1"]) + apply (simp add: NOT_mask) + done + +lemma nat_diff_diff_le_lhs: + "a + c - b \ d \ a - (b - c) \ (d :: nat)" + by arith + +lemma region_actually_is_bytes_subset_t_hrs: + "region_actually_is_bytes ptr sz s' + \ {ptr' ..+ sz'} \ {ptr ..+ sz} + \ t_hrs_' (globals s') = t_hrs_' (globals s) + \ region_actually_is_bytes ptr' sz' s" + by (auto simp: region_actually_is_bytes_def) + +lemma eq_UntypedCap_helper: + "isUntypedCap cap \ capIsDevice cap = dev + \ capPtr cap = ptr \ capBlockSize cap = sz + \ capFreeIndex cap = idx + \ cap = UntypedCap dev ptr sz idx" + by (clarsimp simp: isCap_simps) + lemma resetUntypedCap_ccorres: + notes upt.simps[simp del] Collect_const[simp del] replicate_numeral[simp del] + shows "ccorres (cintr \ dc) (liftxf errstate id (K ()) ret__unsigned_long_') - (invs' and ct_active' and ex_cte_cap_to' cnodeptr) + (invs' and sch_act_simple and ct_active' and cte_wp_at' (isUntypedCap o cteCap) slot + and (\s. descendants_of' slot (ctes_of s) = {})) (UNIV \ {s. srcSlot_' s = Ptr slot}) [] (resetUntypedCap slot) (Call resetUntypedCap_'proc)" + using [[ceqv_simpl_sequence = true]] apply (cinit lift: srcSlot_') - apply (simp add: liftE_bindE getSlotCap_def del: Collect_const) + apply (simp add: liftE_bindE getSlotCap_def + Collect_True extra_sle_sless_unfolds) apply (rule ccorres_getCTE, rename_tac cte) - apply (rule ccorres_move_c_guard_cte ccorres_rhs_assoc - | (csymbr, rule ccorres_abstract_cleanup))+ -sorry - \to_bool \deviceMemory = isdev\ \ - \ isFrameType newType))\ - ) - [] (do invokeUntyped (Retype cref ptr_base ptr newType us destSlots isdev); - "isdev \ (newType = APIObjectType ArchTypes_H.apiobject_type.Untyped \ isFrameType newType)" - \to_bool \deviceMemory = isdev\ \ + apply (rule ccorres_move_c_guard_cte) + apply (rule ccorres_symb_exec_r) + apply (rule_tac xf'="prev_cap_'" in ccorres_abstract, ceqv) + apply (rename_tac prev_cap) + apply (rule_tac P="ccap_relation (cteCap cte) prev_cap" + in ccorres_gen_asm2) + apply (csymbr | rule ccorres_Guard_Seq[where S=UNIV])+ + apply (rule_tac P="isUntypedCap (cteCap cte) + \ capFreeIndex (cteCap cte) < 2 ^ word_bits + \ capFreeIndex (cteCap cte) < 2 ^ (word_bits - 1) + \ is_aligned (of_nat (capFreeIndex (cteCap cte)) :: addr) 4 + \ capBlockSize (cteCap cte) < 2 ^ word_bits" + in ccorres_gen_asm) + apply clarsimp + apply (frule(1) isUntypedCap_ccap_relation_helper) + apply (clarsimp simp: shiftr_shiftl1) + apply (rule ccorres_Cond_rhs_Seq) + apply (frule of_nat_0, simp add: word_bits_def) + apply (simp add: unlessE_def) + apply (rule ccorres_split_throws) + apply (rule ccorres_return_CE, simp+) + apply vcg + apply clarsimp + apply (clarsimp simp: unat_of_nat32) + apply (frule of_nat_gt_0) + apply (simp add: unlessE_def) + apply (simp add: hrs_htd_update) + apply (rule ccorres_Guard_Seq[where S=UNIV])? + apply (rule ccorres_rhs_assoc2) + apply (rule ccorres_split_nothrow) + apply (rule_tac idx="capFreeIndex (cteCap cte)" + in deleteObjects_ccorres[where p=slot, unfolded o_def]) + apply ceqv + apply clarsimp + apply (simp only: ccorres_seq_cond_raise) + apply (rule ccorres_cond[where R="\"]) + apply (clarsimp simp: reset_chunk_bits_def) + apply (simp add: word_less_nat_alt unat_of_nat32 from_bool_0) + apply blast + apply (simp add: liftE_def bind_assoc shiftL_nat unless_def + when_def) + apply (rule ccorres_rhs_assoc)+ + apply (rule ccorres_split_nothrow[where xf'=xfdc and r'=dc]) + apply (rule ccorres_cond2[where R=\]) + apply (simp add: from_bool_0) + apply (ctac add: clearMemory_untyped_ccorres[where ut_slot=slot]) + apply (rule ccorres_return_Skip) + apply ceqv + apply (rule ccorres_rhs_assoc2)+ + apply (rule ccorres_split_nothrow_novcg) + apply (rule_tac cap'="cteCap cte" in updateFreeIndex_ccorres) + apply (rule allI, rule conseqPre, vcg) + apply (clarsimp simp: typ_heap_simps' cap_get_tag_isCap + dest!: ccte_relation_ccap_relation) + apply (drule(1) isUntypedCap_ccap_relation_helper)+ + apply (rule exI, strengthen refl, simp) + apply (simp only: t_hrs_update_use_t_hrs mex_def meq_def) + apply blast + apply ceqv + apply (rule ccorres_return_CE'[unfolded returnOk_def o_apply], simp+) + apply wp + apply (simp add: guard_is_UNIV_def) + apply wp + apply simp + apply (vcg exspec=cleanCacheRange_PoU_modifies) + apply (rule_tac P="reset_chunk_bits \ capBlockSize (cteCap cte) + \ of_nat (capFreeIndex (cteCap cte)) - 1 + < (2 ^ capBlockSize (cteCap cte) :: addr)" + in ccorres_gen_asm) + apply (elim conjE) + apply (simp add: whileAnno_def) + apply (rule ccorres_Guard_Seq ccorres_rhs_assoc)+ + apply csymbr + apply (simp add: reset_name_seq_bound_helper2 word_sle_def word_sless_def + msb_big linorder_not_le word_bits_def word_of_nat_less + reset_name_seq_bound_helper2[simplified simp_thms] + Collect_True) + apply ((rule ccorres_Guard_Seq[where S=UNIV])+)? + apply (rule ccorres_add_returnOk) + apply (rule ccorres_splitE_novcg) + apply (rule_tac P="capPtr (cteCap cte) \ getFreeRef (capPtr (cteCap cte)) + (capFreeIndex (cteCap cte)) - 1" + in ccorres_gen_asm) + apply (rule_tac P="(\s. valid_cap' (cteCap cte) s) + \ \ capIsDevice (cteCap cte)" in ccorres_gen_asm) + apply (rule_tac yf="\ptr. ptr - (capPtr (cteCap cte))" + and P="\s. 2 ^ reset_chunk_bits \ gsMaxObjectSize s" + and F="\n b idx. cte_wp_at' (\cte'. \idx'. cteCap cte' + = (cteCap cte)\ capFreeIndex := idx' \ + \ idx = (getFreeRef (capPtr (cteCap cte)) idx') - 1 + && ~~ mask reset_chunk_bits) slot + and invs' + and (\s. descendants_of' slot (ctes_of s) = {}) + and pspace_no_overlap' (capPtr (cteCap cte)) (capBlockSize (cteCap cte))" + and Q="{s. \ capIsDevice (cteCap cte) + \ region_actually_is_bytes (capPtr (cteCap cte)) + (2 ^ (capBlockSize (cteCap cte))) s}" + in mapME_x_simpl_sequence_fun_related) + apply (rule nth_equalityI) + apply (simp add: length_upto_enum_step) + apply (simp add: getFreeRef_def shiftr_div_2n_w reset_chunk_bits_def + word_size) + apply (simp add: length_upto_enum_step upto_enum_step_nth + less_Suc_eq_le nth_rev getFreeRef_def + reset_chunk_bits_def shiftr_div_2n_w word_size + and_not_mask shiftl_t2n) + apply clarify + apply (rule_tac Q="\v. cte_wp_at' (\cte. capFreeIndex (cteCap cte) = v) slot" + and Q'="\\" and V="\v. x = (getFreeRef (capPtr (cteCap cte)) v) - 1 + && ~~ mask reset_chunk_bits" + in ccorres_req_Ex) + apply (clarsimp simp: cte_wp_at_ctes_of isCap_simps) + apply (clarsimp simp add: shiftL_nat) + apply (rename_tac prior_idx) + apply (rule ccorres_guard_imp2) + apply (rule ccorres_rhs_assoc)+ + apply (ctac add: clearMemory_untyped_ccorres[where ut_slot=slot]) + apply (rule ccorres_Guard_Seq)+ + apply csymbr + apply (rule ccorres_move_c_guard_cte) + apply (rule ccorres_split_nothrow_novcg) + apply (rule_tac cap'="(cteCap cte)\ capFreeIndex := prior_idx \" + in updateFreeIndex_ccorres) + apply (rule allI, rule conseqPre, vcg) + apply (clarsimp simp: typ_heap_simps' cap_get_tag_isCap + dest!: ccte_relation_ccap_relation) + apply (drule(1) isUntypedCap_ccap_relation_helper)+ + apply (drule isUntypedCap_ccap_relation_helper, clarsimp simp: isCap_simps) + apply (rule exI, strengthen refl, simp) + apply (simp only: t_hrs_update_use_t_hrs mex_def meq_def, + simp only: fun_app_def, strengthen exI[mk_strg I], strengthen refl) + apply (clarsimp simp: isCap_simps) + apply (simp add: getFreeIndex_def) + apply (clarsimp simp: in_set_conv_nth + length_upto_enum_step upto_enum_step_nth + less_Suc_eq_le getFreeRef_def) + apply (frule(2) reset_untyped_inner_offs_helper, simp+) + apply (clarsimp simp: valid_cap_simps' capAligned_def + is_aligned_mask_out_add_eq_sub[OF is_aligned_weaken]) + apply (rule less_mask_eq, rule shiftr_less_t2n, + erule order_less_le_trans, rule two_power_increasing, + simp_all)[1] + + apply ceqv + apply (rule ccorres_add_returnOk) + apply (ctac add: preemptionPoint_ccorres) + apply (rule ccorres_from_vcg[where P=\ and P'=UNIV]) + apply (rule allI, rule conseqPre, vcg) + apply (clarsimp simp: returnOk_def return_def) + apply (rule ccorres_from_vcg_throws[where P=\ and P'=UNIV]) + apply (rule allI, rule conseqPre, vcg) + apply (clarsimp simp: throwError_def return_def cintr_def) + apply wp + apply (simp, vcg exspec=preemptionPoint_modifies) + apply (wp updateFreeIndex_clear_invs') + apply (simp add: guard_is_UNIV_def) + apply (wp hoare_vcg_ex_lift doMachineOp_psp_no_overlap) + apply clarsimp + apply (vcg exspec=cleanCacheRange_PoU_modifies) + apply clarify + apply (rule conjI) + apply (clarsimp simp: invs_valid_objs' cte_wp_at_ctes_of + invs_urz + getFreeIndex_def isCap_simps + invs_pspace_aligned' + invs_pspace_distinct' + simp del: ) + apply (frule valid_global_refsD_with_objSize, clarsimp) + apply (clarsimp simp: conj_comms in_set_conv_nth + length_upto_enum_step upto_enum_step_nth + less_Suc_eq_le getFreeRef_def) + apply (frule(2) reset_untyped_inner_offs_helper, simp+) + apply (clarsimp simp: valid_cap_simps' capAligned_def + aligned_offset_non_zero cteCaps_of_def + is_aligned_mask_out_add_eq_sub[OF is_aligned_weaken] + split_if[where P="\z. a \ z" for a]) + apply (strengthen is_aligned_mult_triv2[THEN is_aligned_weaken] + aligned_sub_aligned[OF _ _ order_refl] + aligned_intvl_offset_subset_ran + unat_le_helper Aligned.is_aligned_neg_mask) + apply (simp add: order_less_imp_le reset_chunk_bits_def) + + apply (clarsimp simp: in_set_conv_nth isCap_simps + length_upto_enum_step upto_enum_step_nth + less_Suc_eq_le getFreeRef_def + cte_wp_at_ctes_of getFreeIndex_def + hrs_mem_update) + apply (frule valid_global_refsD_with_objSize, clarsimp) + apply (frule(2) reset_untyped_inner_offs_helper, simp+) + apply (frule ctes_of_valid', clarify+) + apply (clarsimp simp: valid_cap_simps') + apply (strengthen ghost_assertion_size_logic[unfolded o_def, rotated, mk_strg I E] + is_aligned_weaken[where y=2 and x=reset_chunk_bits] + is_aligned_weaken[where y=8 and x=reset_chunk_bits] + is_aligned_no_overflow'[where n=8, simplified] + power_increasing[where a=2 and n=8, simplified] + region_actually_is_bytes_dom_s[mk_strg I E] + aligned_intvl_offset_subset[where sz'=8, simplified] + is_aligned_mult_triv2[THEN is_aligned_weaken] + region_actually_is_bytes_subset_t_hrs[mk_strg I E] + | simp)+ + apply (clarsimp simp: capAligned_def + aligned_offset_non_zero + is_aligned_add_multI conj_comms + is_aligned_mask_out_add_eq_sub[OF is_aligned_weaken]) + apply (strengthen region_is_bytes_subset[OF region_actually_is_bytes, mk_strg I E] + heap_list_is_zero_mono[OF heap_list_update_eq] + order_trans [OF intvl_start_le + aligned_intvl_offset_subset[where sz'=reset_chunk_bits]] + | simp add: is_aligned_mult_triv2)+ + apply (simp add: unat_sub word_le_nat_alt unat_sub[OF word_and_le2] + mask_out_sub_mask word_and_le2 + unat_of_nat32[OF order_le_less_trans, rotated, + OF power_strict_increasing]) + apply (case_tac idx') + (* must be contradictory *) + apply (simp add: is_aligned_def addr_card_def card_word + reset_chunk_bits_def) + apply (simp add: unat_of_nat32[OF order_le_less_trans, rotated, + OF power_strict_increasing]) + apply (simp add: word_mod_2p_is_mask[symmetric] reset_chunk_bits_def + unat_mod unat_of_nat mod_mod_cancel) + apply (strengthen nat_le_Suc_less_imp[OF mod_less_divisor, THEN order_trans]) + apply (simp add: is_aligned_def addr_card_def card_word) + + apply clarsimp + apply (rule conseqPre, vcg exspec=cleanCacheRange_PoU_modifies + exspec=preemptionPoint_modifies) + apply (clarsimp simp: in_set_conv_nth isCap_simps + length_upto_enum_step upto_enum_step_nth + less_Suc_eq_le getFreeRef_def + cte_wp_at_ctes_of getFreeIndex_def + hrs_mem_update) + apply (frule(2) reset_untyped_inner_offs_helper, simp+) + apply (clarsimp simp: valid_cap_simps') + apply (strengthen is_aligned_weaken[where y=2 and x=reset_chunk_bits] + ghost_assertion_size_logic[unfolded o_def, rotated, mk_strg I E] + is_aligned_weaken[where y=8 and x=reset_chunk_bits] + is_aligned_no_overflow'[where n=8, simplified] + power_increasing[where a=2 and n=8, simplified] + region_actually_is_bytes_dom_s[mk_strg I E] + aligned_intvl_offset_subset[where sz'=8, simplified] + is_aligned_mult_triv2[THEN is_aligned_weaken] + | simp)+ + apply (clarsimp simp: capAligned_def + aligned_offset_non_zero + is_aligned_add_multI conj_comms + region_actually_is_bytes_def) + apply (simp add: reset_chunk_bits_def is_aligned_def) + + apply (rule hoare_pre) + apply (wp updateFreeIndex_cte_wp_at updateFreeIndex_clear_invs' + updateFreeIndex_pspace_no_overlap' + updateFreeIndex_descendants_of2 + doMachineOp_psp_no_overlap + hoare_vcg_ex_lift + | (wp_once preemptionPoint_inv, simp, simp add: pspace_no_overlap'_def) + | simp)+ + apply (simp add: cte_wp_at_ctes_of isCap_simps | clarify)+ + apply (clarsimp simp: length_upto_enum_step upto_enum_step_nth + less_Suc_eq_le getFreeRef_def + getFreeIndex_def nth_rev + conj_comms invs_pspace_aligned' invs_pspace_distinct' + invs_valid_pspace') + apply (frule(1) reset_untyped_inner_offs_helper[OF _ order_refl], simp+) + apply (frule ctes_of_valid', clarify+) + apply (clarsimp simp: valid_cap_simps' capAligned_def + is_aligned_mask_out_add_eq[OF is_aligned_weaken] + aligned_bump_down Aligned.is_aligned_neg_mask + is_aligned_mask_out_add_eq_sub[OF is_aligned_weaken]) + apply (simp add: field_simps) + apply (strengthen Aligned.is_aligned_neg_mask unat_le_helper) + apply (simp add: reset_chunk_bits_def[THEN arg_cong[where f="\x. n \ x" for n]]) + apply (rule order_less_imp_le, erule order_le_less_trans[rotated], + rule olen_add_eqv[THEN iffD2]) + apply (rule order_trans, rule word_mult_le_mono1, rule word_of_nat_le, + erule order_trans[rotated], simp, simp add: reset_chunk_bits_def) + apply (simp only: unat_power_lower32 shiftr_div_2n_w[symmetric] + word_size word_bits_def[symmetric]) + apply (rule nat_less_power_trans2) + apply (rule order_less_le_trans[OF word_shiftr_lt]) + apply (simp add: word_bits_def) + apply (simp add: word_bits_def reset_chunk_bits_def) + apply (simp add: field_simps) + apply ceqv + apply (rule ccorres_return_CE, simp+)[1] + apply wp + apply (simp add: ccHoarePost_def guard_is_UNIV_def) + apply simp + + apply (strengthen invs_valid_objs' invs_urz) + apply ((rule_tac d="capIsDevice (cteCap cte)" and idx="capFreeIndex (cteCap cte)" in + deleteObject_no_overlap + | rule_tac d="capIsDevice (cteCap cte)" and idx="capFreeIndex (cteCap cte)" in + deleteObjects_cte_wp_at' + | wp_once hoare_vcg_const_imp_lift + hoare_vcg_conj_lift + | wp_once deleteObjects_invs'[where p=slot] + deleteObjects_descendants[where p=slot] + | strengthen exI[mk_strg I])+)[1] + apply (simp add: word_sle_def) + apply vcg + apply simp + apply vcg + apply (rule conseqPre, vcg, clarsimp) + apply (rule conjI) + apply clarsimp + apply (frule if_unsafe_then_capD', clarsimp+) + apply (clarsimp simp: cte_wp_at_ctes_of) + apply ((strengthen refl eq_UntypedCap_helper + eq_UntypedCap_helper[symmetric] | simp)+)? + apply (frule ctes_of_valid', clarsimp+) + apply (simp add: exI)? + apply (clarsimp simp: isCap_simps valid_cap_simps' capAligned_def + conj_comms invs_valid_pspace' + descendants_range'_def2 + empty_descendants_range_in' + getFreeRef_def upto_intvl_eq) + apply (frule valid_global_refsD_with_objSize, clarsimp+) + apply (strengthen order_le_less_trans[where z="2 ^ n" for n, mk_strg I E] + order_trans[rotated, where z="gsMaxObjectSize s" for s, mk_strg I E]) + apply (strengthen power_strict_increasing + | simp)+ + apply (clarsimp simp: word_bits_def) + apply (subgoal_tac "capPtr (cteCap cte) \ getFreeRef (capPtr (cteCap cte)) + (capFreeIndex (cteCap cte)) - 1") + apply (case_tac "the (ctes_of s slot)", simp) + apply (frule(3) ex_cte_not_in_untyped_range, clarsimp+) + apply (strengthen is_aligned_no_wrap'[where off="a - b" for a b, + simplified field_simps, mk_strg I E]) + apply (simp add: getFreeRef_def nth_rev length_upto_enum_step + upto_enum_step_nth word_of_nat_le + is_aligned_mask_out_add_eq_sub[OF is_aligned_weaken]) + apply (simp add: neg_mask_is_div' reset_chunk_bits_def word_size) + apply (safe, simp_all)[1] + + apply (simp add: getFreeRef_def) + apply (strengthen is_aligned_no_wrap'[where off="a - b" for a b, + simplified field_simps, mk_strg I E]) + apply (simp add: word_of_nat_le) + + apply (clarsimp simp: cte_wp_at_ctes_of) + apply (frule(1) rf_sr_ctes_of_clift, clarsimp) + apply (frule(2) rf_sr_cte_relation) + apply (clarsimp simp: typ_heap_simps' + dest!: ccte_relation_ccap_relation) + apply (strengthen typ_region_bytes_actually_is_bytes) + apply (simp add: hrs_htd_update hrs_mem_update exI) + apply (frule(1) isUntypedCap_ccap_relation_helper) + apply (frule ctes_of_valid', clarify+) + apply (frule valid_global_refsD_with_objSize, clarsimp) + apply (clarsimp simp: valid_cap_simps' isCap_simps capAligned_def + from_bool_0 cap_to_H_simps) + apply (frule(1) ghost_assertion_size_logic_no_unat) + apply (simp add: ghost_assertion_data_get_def gs_clear_region_def) + apply (strengthen is_aligned_no_overflow' + typ_region_bytes_dom_s + aligned_intvl_offset_subset + region_is_bytes_typ_region_bytes + intvl_start_le is_aligned_power2 + heap_list_is_zero_mono[OF heap_list_update_eq] + | simp add: unat_of_nat)+ + apply (clarsimp simp: order_trans[OF power_increasing[where a=2]] + addr_card_def card_word + is_aligned_weaken from_bool_0) + done lemma ccorres_cross_retype_zero_bytes_over_guard: "range_cover ptr sz (APIType_capBits newType userSize) num_ret \ ccorres_underlying rf_sr Gamm rvr xf arrel axf P' Q hs af cf \ ccorres_underlying rf_sr Gamm rvr xf arrel axf ((\s. invs' s - \ cte_wp_at' (\cte. \idx. cteCap cte = UntypedCap (ptr &&~~ mask sz) sz idx + \ cte_wp_at' (\cte. \idx. cteCap cte = UntypedCap dev (ptr && ~~ mask sz) sz idx \ idx \ unat (ptr && mask sz)) p s) and P') - {s'. region_is_zero_bytes ptr - (num_ret * 2 ^ APIType_capBits newType userSize) s' \ s' \ Q} hs af cf" + {s'. (\ dev \ region_is_zero_bytes ptr + (num_ret * 2 ^ APIType_capBits newType userSize) s') + \ (\cte cte' idx. cslift s' (cte_Ptr p) = Some cte' + \ ccte_relation cte cte' \ cteCap cte = UntypedCap dev (ptr && ~~ mask sz) sz idx) + \ s' \ Q} hs af cf" apply (erule ccorres_guard_imp2) apply (clarsimp simp: cte_wp_at_ctes_of) - apply (frule(2) retype_offs_region_is_zero_bytes, simp+) + apply (frule(1) rf_sr_ctes_of_clift, clarsimp) + apply (frule(2) rf_sr_cte_relation) + apply (case_tac dev) + apply fastforce + apply (frule(1) retype_offs_region_is_zero_bytes, (simp | clarsimp)+) + apply fastforce + done + +lemma zero_bytes_heap_update: + "heap_list_is_zero (hrs_mem hrs) ptr n + \ region_is_bytes' ptr n (hrs_htd hrs) + \ h_t_valid (hrs_htd hrs) c_guard (cptr :: 'a ptr) + \ typ_uinfo_t TYPE ('a :: mem_type) \ typ_uinfo_t TYPE(8 word) + \ heap_list_is_zero (heap_update cptr v (hrs_mem hrs)) ptr n" + apply (frule(2) region_is_bytes_disjoint) + apply (clarsimp simp: heap_update_def) + apply (subst heap_list_update_disjoint_same, simp_all) + apply (simp add: Int_commute) done lemma invokeUntyped_Retype_ccorres: @@ -1727,45 +2388,50 @@ lemma invokeUntyped_Retype_ccorres: (invs' and ct_active' and ex_cte_cap_to' cnodeptr and (\s. case gsCNodes s cnodeptr of None \ False | Some n \ length destSlots + unat start \ 2 ^ n) - and valid_untyped_inv' (Retype cref reset ptr_base ptr newType us destSlots)) + and valid_untyped_inv' (Retype cref reset ptr_base ptr newType us destSlots isdev) + and K (isdev \ (newType = APIObjectType ArchTypes_H.apiobject_type.Untyped + \ isFrameType newType)) + ) (UNIV \ {s. retypeBase_' s = Ptr ptr} \ {s. srcSlot_' s = Ptr cref} \ {s. reset_' s = from_bool reset} \ {s. newType_' s = object_type_from_H newType } \ {s. unat (userSize_' s) = us } - \ {s. to_bool (deviceMemory_' s) = isdev} + \ {s. deviceMemory_' s = from_bool isdev} \ \\destSlots = slot_range_C (cte_Ptr cnodeptr) start (of_nat (length destSlots)) \ (\n - (isdev \ (newType = APIObjectType ArchTypes_H.apiobject_type.Untyped - \ isFrameType newType))\ + destSlots ! n = cnodeptr + (start + of_nat n) * 0x10)\ ) [] - (invokeUntyped (Retype cref reset ptr_base ptr newType us destSlots)) + (invokeUntyped (Retype cref reset ptr_base ptr newType us destSlots isdev)) (Call invokeUntyped_Retype_'proc)" + (is "ccorres _ _ _ ?P' [] _ _") apply (rule ccorres_name_pre) apply (clarsimp simp only: valid_untyped_inv_wcap') - apply clarsimp+ - apply (subgoal_tac - "invokeUntyped_proofs s cref reset ptr_base ptr newType us - destSlots ptr_base sz idx") - prefer 2 - apply (clarsimp simp:invokeUntyped_proofs_def cte_wp_at_ctes_of) - (* This is the main case involving deleteObjects. *) - (* FIXME: proof very similar to invokeUntyped_Retype_ccorres_side_case. *) proof - - fix s sz idx cte ptr' - assume proofs: - "invokeUntyped_proofs s cref reset ptr_base ptr newType us destSlots ptr_base sz idx" - and vui: - "valid_untyped_inv_wcap' (Retype cref reset ptr_base ptr newType us destSlots) - (Some (UntypedCap ptr' sz idx)) s" - and misc1: - "ct_active' s" "invs' s" "ex_cte_cap_to' cnodeptr s" + fix s sz idx cte + assume vui1: "valid_untyped_inv_wcap' + (Invocations_H.untyped_invocation.Retype cref reset ptr_base ptr newType us destSlots isdev) + (Some (case Invocations_H.untyped_invocation.Retype cref reset ptr_base ptr newType us destSlots + isdev of + Invocations_H.untyped_invocation.Retype slot reset ptr_base ptr ty us slots d \ + capability.UntypedCap d (ptr && ~~ mask sz) sz idx)) + s" + and misc1[simplified]: "ct_active' s" "invs' s" "ex_cte_cap_to' cnodeptr s" "case gsCNodes s cnodeptr of None \ False | Some n \ length destSlots + unat start \ 2 ^ n" - "isdev \ (newType = APIObjectType ArchTypes_H.apiobject_type.Untyped \ isFrameType newType)" + "K (isdev \ (newType = APIObjectType ArchTypes_H.apiobject_type.Untyped \ isFrameType newType)) s" + + have vui: "valid_untyped_inv_wcap' (Retype cref reset ptr_base ptr newType us destSlots isdev) + (Some (UntypedCap isdev (ptr && ~~ mask sz) sz idx)) s" + using vui1 + by (clarsimp simp: cte_wp_at_ctes_of) + + have proofs: "invokeUntyped_proofs s cref reset ptr_base ptr newType us destSlots sz idx isdev" + using vui misc1 + by (clarsimp simp: cte_wp_at_ctes_of invokeUntyped_proofs_def) + note no_simps[simp del] = untyped_range.simps usable_untyped_range.simps atLeastAtMost_iff atLeastatMost_subset_iff atLeastLessThan_iff Int_atLeastAtMost atLeastatMost_empty_iff split_paired_Ex @@ -1778,7 +2444,7 @@ lemma invokeUntyped_Retype_ccorres: have us_misc: "newType = APIObjectType apiobject_type.CapTableObject \ 0 < us" - "newType = APIObjectType apiobject_type.Untyped \ 4 \ us \ us \ 30" + "newType = APIObjectType apiobject_type.Untyped \ 4 \ us \ us \ 29" using vui by (auto simp: cte_wp_at_ctes_of) @@ -1802,40 +2468,39 @@ lemma invokeUntyped_Retype_ccorres: have ptr_base_eq: "ptr_base = ptr && ~~ mask sz" using vui - by (clarsimp simp: cte_wp_at_ctes_of) + by (clarsimp simp: cte_wp_at_ctes_of)+ -(* - have cte_size_inter_empty: - "{cref..cref + 0xF} \ {ptr_base..ptr_base + 2 ^ sz - 1} = {}" - using cover misc - apply - - apply (rule disjoint_subset2[rotated]) - apply (erule(1) cte_size_inter_empty) - apply (clarsimp simp:isCap_simps) - apply (cut_tac invokeUntyped_proofs.cref_inv[OF proofs])\<^sup> - apply (simp add: ptr_base_eq) -find_theorems name: subs intro - apply (erule contra_subsetD[rotated]) - apply (simp add: atLeastatMost_subset_iff) -find_theorems range_cover -defer - apply simp + have sz_bound: + "sz \ 29" + using vui misc + apply (clarsimp simp: cte_wp_at_ctes_of) + apply (frule Finalise_R.ctes_of_valid', clarsimp) + apply (clarsimp simp: valid_cap_simps') + done + + have some_range_cover_arithmetic: + "(ptr + (of_nat (length destSlots) << unat (of_nat (APIType_capBits newType us) :: addr)) + - ptr_base >> 4) && mask 26 + = of_nat (getFreeIndex ptr_base + (ptr + of_nat (shiftL (length destSlots) + (APIType_capBits newType us)))) >> 4" + using cover range_cover_sz'[OF cover] + apply (simp add: getFreeIndex_def shiftl_t2n + unat_of_nat_eq shiftL_nat) + apply (rule less_mask_eq) + apply (rule shiftr_less_t2n) + apply (rule order_le_less_trans[where y="2 ^ sz"]) + apply (rule order_trans[OF _ range_cover_compare_bound_word[OF cover]]) + apply (simp add: ptr_base_eq mask_out_sub_mask mult.commute) + apply (simp add: word_less_nat_alt order_le_less_trans[OF sz_bound]) + apply (rule order_less_le_trans, rule power_strict_increasing, + rule order_le_less_trans[OF sz_bound lessI], simp+) done -*) show "ccorres (cintr \ dc) - (liftxf errstate id (K ()) ret__unsigned_long_') (\s'. s' = s) - (UNIV \ \\retypeBase = Ptr ptr\ \ - \\srcSlot = cte_Ptr cref\ \ - \\reset = from_bool reset\ \ - \\newType = object_type_from_H newType\ \ - \unat \userSize = us\ \ - \to_bool \deviceMemory = isdev\ \ - \\destSlots = slot_range_C (cte_Ptr cnodeptr) start - (of_nat (length destSlots)) \ - (\n) - [] (invokeUntyped (Retype cref reset ptr_base ptr newType us destSlots)) + (liftxf errstate id (K ()) ret__unsigned_long_') (\s'. s' = s) ?P' + [] (invokeUntyped (Retype cref reset ptr_base ptr newType us destSlots isdev)) (Call invokeUntyped_Retype_'proc)" apply (cinit lift: retypeBase_' srcSlot_' reset_' newType_' userSize_' deviceMemory_' destSlots_' @@ -1843,6 +2508,8 @@ defer apply (rule ccorres_move_c_guard_cte) apply csymbr apply (rule ccorres_abstract_cleanup) + apply (rename_tac ptr_fetch, + rule_tac P="ptr_fetch = ptr_base" in ccorres_gen_asm2) apply csymbr apply csymbr apply (rule ccorres_move_c_guard_cte) @@ -1866,57 +2533,66 @@ defer apply (simp add: whenE_def) apply (rule ccorres_returnOk_skip) apply ceqv + apply (simp add: liftE_def bind_assoc) + apply csymbr + apply (rule ccorres_Guard_Seq) + apply csymbr + apply csymbr + apply (rule ccorres_move_c_guard_cte) + apply (rule ccorres_Guard_Seq)+ + apply (rule ccorres_stateAssert) + apply (rule ccorres_cross_retype_zero_bytes_over_guard[where + dev=isdev and p=cref, OF cover]) + apply (rule ccorres_rhs_assoc2) + apply (rule ccorres_split_nothrow[where r'=dc and xf'=xfdc]) + apply (rule_tac cap'="UntypedCap isdev ptr_base sz (if reset then 0 else idx)" + in updateFreeIndex_ccorres) + apply (rule allI, rule conseqPre, vcg) + apply (rule subsetI, clarsimp simp: typ_heap_simps' isCap_simps) + apply (frule ccte_relation_ccap_relation) + apply clarsimp + apply (frule cap_get_tag_isCap_unfolded_H_cap) + apply (cut_tac some_range_cover_arithmetic) + apply (case_tac cte', clarsimp simp: modify_map_def fun_eq_iff split: split_if) + apply (simp add: mex_def meq_def ptr_base_eq del: split_paired_Ex) + apply (rule exI, strengthen refl, simp) + apply (strengthen globals.fold_congs, simp add: field_simps) + apply ceqv + apply (ctac add: createNewObjects_ccorres[where sz = sz and + start = start and cnodeptr=cnodeptr and + num = "of_nat (length destSlots)" + and idx = "getFreeIndex ptr_base + (ptr + of_nat (shiftL (length destSlots) (APIType_capBits newType us)))"]) + apply (rule ccorres_from_vcg_throws[where P=\ and P'=UNIV]) + apply (rule allI, rule conseqPre, vcg) + apply (clarsimp simp: return_def) + apply wp + apply (vcg exspec=createNewObjects_modifies) + apply simp + apply (wp updateFreeIndex_forward_invs' sch_act_simple_lift + updateFreeIndex_cte_wp_at hoare_vcg_const_Ball_lift + updateFreeIndex_pspace_no_overlap' + updateFreeIndex_caps_no_overlap'' + updateFreeIndex_caps_overlap_reserved + updateFreeIndex_descendants_range_in' + | simp)+ + apply (clarsimp simp: misc unat_of_nat_eq[OF range_cover.weak, OF cover]) + apply (vcg exspec=cap_untyped_cap_ptr_set_capFreeIndex_modifies) + apply simp + apply (rule validE_validE_R, rule hoare_post_impErr, + rule hoare_vcg_conj_liftE1[rotated, where Q="\_ s. + case gsCNodes s cnodeptr of None \ False + | Some n \ length destSlots + unat start \ 2 ^ n"], + rule whenE_reset_resetUntypedCap_invs_etc + [where ui="Retype cref reset ptr_base ptr newType us destSlots isdev" + and dev=isdev and ptr="ptr && ~~ mask sz" and ptr'=ptr and sz=sz and idx=idx]) + apply (simp add: whenE_def, wp resetUntypedCap_gsCNodes_at_pt)[1] + prefer 2 apply simp - apply (simp add: liftE_def bind_assoc) - apply csymbr - apply (rule ccorres_Guard_Seq) - apply csymbr - apply csymbr - apply (rule ccorres_move_c_guard_cte) - apply (rule ccorres_Guard_Seq)+ - apply (rule ccorres_stateAssert) - apply (rule ccorres_cross_retype_zero_bytes_over_guard[where p=cref, OF cover]) - apply (rule ccorres_split_nothrow[where r'=dc and xf'=xfdc]) - - apply (rule_tac P=\ and P'=UNIV in ccorres_inst) - defer - apply ceqv - - apply (ctac add: createNewObjects_ccorres[where sz = sz and - start = start and cnodeptr=cnodeptr and - num = "of_nat (length destSlots)" - and idx = "getFreeIndex ptr_base - (ptr + of_nat (shiftL (length destSlots) (APIType_capBits newType us)))"]) - apply (rule ccorres_from_vcg_throws[where P=\ and P'=UNIV]) - apply (rule allI, rule conseqPre, vcg) - apply (clarsimp simp: return_def) - apply wp - apply (vcg exspec=createNewObjects_modifies) - apply (wp updateFreeIndex_forward_invs' sch_act_simple_lift - updateFreeIndex_cte_wp_at hoare_vcg_const_Ball_lift - updateFreeIndex_pspace_no_overlap' - updateFreeIndex_caps_no_overlap'' - updateFreeIndex_caps_overlap_reserved - updateFreeIndex_descendants_range_in' - | simp)+ - apply (clarsimp simp: misc - unat_of_nat_eq[OF range_cover.weak, OF cover]) - apply (vcg exspec=cap_untyped_cap_ptr_set_capFreeIndex_modifies) - apply simp - apply (rule validE_validE_R, rule hoare_post_impErr, - rule hoare_vcg_conj_liftE1[rotated, where Q="\_ s. - case gsCNodes s cnodeptr of None \ False - | Some n \ length destSlots + unat start \ 2 ^ n"], - rule whenE_reset_resetUntypedCap_invs_etc - [where ui="Retype cref reset ptr_base ptr newType us destSlots" - and ptr="ptr && ~~ mask sz" and ptr'=ptr and sz=sz and idx=idx]) - apply (simp add: whenE_def, wp resetUntypedCap_gsCNodes_at_pt)[1] - prefer 2 - apply simp - apply (clarsimp simp only: ) - apply (frule(2) invokeUntyped_proofs.intro) - apply (cut_tac bits_low us_misc us_misc') - apply (clarsimp simp: cte_wp_at_ctes_of + apply (clarsimp simp only: ) + apply (frule(2) invokeUntyped_proofs.intro) + apply (cut_tac bits_low us_misc us_misc') + apply (clarsimp simp: cte_wp_at_ctes_of invokeUntyped_proofs.caps_no_overlap' invokeUntyped_proofs.ps_no_overlap' invokeUntyped_proofs.descendants_range @@ -1927,123 +2603,97 @@ defer invokeUntyped_proofs.not_0_ptr atLeastAtMost_iff[where i=0] cong: if_cong) - apply (frule invokeUntyped_proofs.idx_le_new_offs) - apply (frule invokeUntyped_proofs.szw) - apply (frule invokeUntyped_proofs.descendants_range(2), simp) - apply (simp add: cNodeNoOverlap_retype_have_size shiftL_nat mult.commute) - apply (clarsimp simp: getFreeIndex_def conj_comms shiftL_nat + apply (frule invokeUntyped_proofs.idx_le_new_offs) + apply (frule invokeUntyped_proofs.szw) + apply (frule invokeUntyped_proofs.descendants_range(2), simp) + apply (simp add: cNodeNoOverlap_retype_have_size shiftL_nat mult.commute) + apply (clarsimp simp: getFreeIndex_def conj_comms shiftL_nat is_aligned_weaken[OF range_cover.funky_aligned] invs_valid_pspace' isCap_simps arg_cong[OF mask_out_sub_mask, where f="\y. x - y" for x] field_simps unat_of_nat_eq[OF range_cover.weak, OF cover] - if_apply_def2) + if_apply_def2 invs_valid_objs' ptr_base_eq + invs_urz) - apply (intro conjI) - (* pspace_no_overlap *) - apply (cases reset, simp_all)[1] + apply (intro conjI) + (* pspace_no_overlap *) + apply (cases reset, simp_all)[1] + (* is_aligned 4 *) + apply (erule is_aligned_weaken[OF range_cover.aligned]) + apply (clarsimp simp: APIType_capBits_low) + (* new idx le *) + apply (clarsimp split: split_if) + (* cnodeptr not in area *) + apply (rule contra_subsetD[rotated], + rule invokeUntyped_proofs.ex_cte_no_overlap'[OF proofs], rule misc) + apply (simp add: shiftl_t2n mult.commute) + apply (rule order_trans, erule range_cover_subset', simp_all)[1] + (* gsCNodes *) + apply (clarsimp split: option.split_asm) + (* kernel data refs *) + apply (drule(1) valid_global_refsD'[OF _ invs_valid_global']) + apply clarsimp + apply (subst Int_commute, erule disjoint_subset2[rotated]) + apply (rule order_trans, erule invokeUntyped_proofs.subset_stuff) + apply (simp add: atLeastatMost_subset_iff word_and_le2) + (* offset bounds *) + apply (frule range_cover.unat_of_nat_n_shift, rule order_refl) + apply (rule order_trans[rotated], erule range_cover.range_cover_compare_bound) + apply (subst unat_plus_simple[THEN iffD1]) + apply (rule order_trans, erule range_cover.range_cover_base_le, + simp add: shiftl_t2n field_simps) + apply (simp add: shiftl_t2n field_simps) + (* subsets *) + apply (rule order_trans, erule invokeUntyped_proofs.subset_stuff) + apply (simp add: atLeastatMost_subset_iff word_and_le2) + (* destSlots *) + apply (clarsimp split: split_if) + apply (frule invokeUntyped_proofs.slots_invD[OF proofs]) + apply (simp add: conj_comms) + (* usableUntyped *) + apply (drule invokeUntyped_proofs.usableRange_disjoint[where d=isdev]) + apply (clarsimp simp: field_simps mask_out_sub_mask) - (* is_aligned 4 *) - apply (erule is_aligned_weaken[OF range_cover.aligned]) - apply (clarsimp simp: APIType_capBits_low) + (* clean up the C postcondition before applying VCG *) + apply (rule conseqPost[where Q'=UNIV and Q'=UNIV]) + apply (vcg exspec=resetUntypedCap_modifies) + apply (cut_tac range_cover.sz[OF cover] + invokeUntyped_proofs.idx_le_new_offs[OF proofs]) + apply (clarsimp simp: ccHoarePost_def hrs_mem_update + object_type_from_H_bound + typ_heap_simps' word_sle_def + word_of_nat_less zero_bytes_heap_update) + apply (frule ccte_relation_ccap_relation) + apply (cut_tac vui) + apply (clarsimp simp: cap_get_tag_isCap getFreeIndex_def + cte_wp_at_ctes_of shiftL_nat + split: split_if) + apply (simp add: mask_out_sub_mask field_simps region_is_bytes'_def) + apply (rule order_refl) - (* new idx le *) - apply (clarsimp split: split_if) - - (* cnodeptr not in area *) - apply (rule contra_subsetD[rotated], - rule invokeUntyped_proofs.ex_cte_no_overlap'[OF proofs], rule misc) - apply (simp add: shiftl_t2n mult.commute) - apply (rule order_trans, erule range_cover_subset', simp_all)[1] - (* gsCNodes *) - apply (clarsimp split: option.split_asm) - - (* kernel data refs *) - - apply (drule(1) valid_global_refsD'[OF _ invs_valid_global']) - apply clarsimp - apply (subst Int_commute, erule disjoint_subset2[rotated]) - apply (rule order_trans, erule invokeUntyped_proofs.subset_stuff) - apply (simp add: atLeastatMost_subset_iff word_and_le2) - - (* offset bounds *) - apply (frule range_cover.unat_of_nat_n_shift, rule order_refl) - apply (rule order_trans[rotated], erule range_cover.range_cover_compare_bound) - apply (subst unat_plus_simple[THEN iffD1]) - apply (rule order_trans, erule range_cover.range_cover_base_le, - simp add: shiftl_t2n field_simps) - apply (simp add: shiftl_t2n field_simps) - - (* subsets *) - apply (rule order_trans, erule invokeUntyped_proofs.subset_stuff) - apply (simp add: atLeastatMost_subset_iff word_and_le2) - -(* destSlots *) -apply (clarsimp split: split_if) -apply (frule invokeUntyped_proofs.slots_invD[OF proofs]) -apply (simp add: conj_comms) - - -(* usableUntyped *) - apply (drule invokeUntyped_proofs.usableRange_disjoint) - apply (clarsimp simp: field_simps mask_out_sub_mask) - -apply (clarsimp simp: ccHoarePost_def) -apply (vcg exspec=resetUntypedCap_modifies) - - - apply simp apply (cut_tac misc us_misc' proofs us_misc bits_low invokeUntyped_proofs.cref_inv[OF proofs]) apply (clarsimp simp: cte_wp_at_ctes_of invokeUntyped_proofs_def descendants_range'_def2 sch_act_simple_def invs_valid_pspace' range_cover.sz) - apply (simp add: APIType_capBits_high word_of_nat_less - word_sle_def object_type_from_H_bound) apply (frule ctes_of_valid', fastforce) apply (clarsimp simp: valid_cap'_def capAligned_def ct_in_state'_def invs_valid_objs' inr_rrel_def) - apply (intro conjI) - apply fast - apply clarsimp - apply (drule invokeUntyped_proofs.ex_cte_no_overlap'[OF proofs]) - apply simp - -apply (clarsimp split: sum.split) - apply (erule(1) rf_sr_ctes_of_cliftE) apply (frule(2) rf_sr_cte_relation) apply (clarsimp simp: cap_get_tag_isCap typ_heap_simps dest!: ccte_relation_ccap_relation) + apply (frule cap_get_tag_isCap_unfolded_H_cap) -defer - - - apply (frule rf_sr_cpspace_relation) - apply (frule cap_CL_lift[symmetric]) - apply (frule(2) rf_sr_cte_relation) - apply (subst (asm) cap_untyped_cap_lift[THEN iffD1]) - apply (subst cap_get_tag_isCap[OF ccte_relation_ccap_relation]) - apply simp + apply (intro conjI) + apply clarsimp + apply (drule invokeUntyped_proofs.ex_cte_no_overlap'[OF proofs]) apply simp - apply (simp add: cap_to_H_simps) - apply (frule typ_clear_region_eq[rotated -1], simp+) - apply (simp add: Int_ac cte_size_inter_empty) - apply (clarsimp simp: range_cover.unat_of_nat_n - region_is_typeless_def[symmetric]) - apply (intro conjI impI allI) - apply (rule of_nat_power[where x = 5,simplified]) - apply (rule APIType_capBits_high) - apply (clarsimp simp: unat_of_nat32[where x=us] word_bits_def)+ - apply (clarsimp simp:getFreeIndex_def) - apply (simp add: shiftL_nat word_bits_conv shiftl_t2n) - apply (clarsimp dest!: range_cover_sz' - simp: unat_of_nat32 word_bits_def) - apply (rule object_type_from_H_bound) - apply (subst cap_get_tag_isCap) - apply (erule ccte_relation_ccap_relation) - apply (simp add:isCap_simps cap_to_H_def) - done - qed + apply (frule(1) cap_get_tag_to_H) + apply (simp add: cap_lift_untyped_cap) + apply clarsimp + done +qed lemma ccorres_returnOk_Basic: "\ \\ s. (\, s) \ sr \ r (Inr v) (xf (f s)) @@ -2165,59 +2815,44 @@ lemma alignUp_spec: lemma checkFreeIndex_ccorres: "ccap_relation cp cap \ - ccorresG rf_sr \ (intr_and_se_rel \ (\r r'. r = unat (r' << 4))) (liftxf errstate (K (scast EXCEPTION_NONE)) id freeIndex_') + ccorresG rf_sr \ (intr_and_se_rel \ (\r (fi, r'). r' = from_bool r + \ (case r of True \ fi = 0 | False \ capFreeIndex cp = unat (fi << 4)))) + (liftxf errstate (K (scast EXCEPTION_NONE)) id (\s. (freeIndex_' s, reset_' s))) (cte_wp_at' (\cte. (cteCap cte = cp \ isUntypedCap cp)) slot and valid_objs' and valid_mdb') UNIV hs - (liftE $ constOnFailure (capFreeIndex cp) (doE y \ ensureNoChildren slot; - returnOk 0 odE)) + (liftE $ constOnFailure False (doE y \ ensureNoChildren slot; returnOk True odE)) (\status :== CALL ensureNoChildren(cte_Ptr slot);; - (Cond \\status \ scast EXCEPTION_NONE\ (\freeIndex :== CALL cap_untyped_cap_get_capFreeIndex(cap)) - (\freeIndex :== 0)))" + (Cond \\status \ scast EXCEPTION_NONE\ + (\freeIndex :== CALL cap_untyped_cap_get_capFreeIndex(cap) + ;; \reset :== scast false) + (\freeIndex :== 0 + ;; \reset :== scast true)))" apply (simp add: constOnFailure_def catch_def liftE_def bindE_bind_linearise bind_assoc case_sum_distrib) apply (rule ccorres_guard_imp2) apply (rule ccorres_split_nothrow_case_sum) apply (ctac add:ensureNoChildren_ccorres) apply (ceqv) - apply (rule ccorres_from_vcg[where P' = UNIV]) - apply (clarsimp simp:returnOk_def return_def bind_def) - apply (rule conseqPre) - apply vcg - apply clarsimp - apply simp - apply (rule ccorres_from_vcg[where P'= UNIV]) - apply (clarsimp simp:return_def) + apply (rule ccorres_from_vcg[where P' = UNIV]) + apply (clarsimp simp add: returnOk_def, simp add: return_def) apply (rule conseqPre) apply vcg apply clarsimp - apply (rule context_conjI) - apply (clarsimp simp:cap_get_tag_isCap) - apply assumption - apply (clarsimp simp:ccap_relation_def isCap_simps cap_untyped_cap_lift_def - cap_lift_def cap_to_H_def - split:if_splits) - apply (rule ensureNoChildren_wp[where P = dc]) - apply clarsimp - apply vcg - apply (clarsimp simp:cte_wp_at_ctes_of rf_sr_def cstate_relation_def cpspace_relation_def Let_def) - apply (rule cmap_relationE1,assumption+) - apply (rule exI)+ - apply (rule conjI,assumption) - apply (rule conjI,simp add:typ_heap_simps') - apply (clarsimp simp:typ_heap_simps') - apply (subst (asm) mdbNext_not_zero_eq_simpler[symmetric]) - apply (erule ccte_relation_cmdbnode_relation) - apply (frule(2) valid_mdbD1') - apply (drule_tac s'= s' in valid_mdb_cslift_next) - apply (simp add:rf_sr_def Let_def cstate_relation_def cpspace_relation_def) - apply simp - apply simp + apply simp + apply (rule ccorres_from_vcg[where P'= UNIV]) + apply (simp, clarsimp simp:return_def) + apply (rule conseqPre) + apply vcg + apply clarsimp + apply (rule context_conjI) + apply (clarsimp simp:cap_get_tag_isCap) + apply assumption + apply (clarsimp simp:ccap_relation_def isCap_simps cap_untyped_cap_lift_def + cap_lift_def cap_to_H_def + split:if_splits) + apply (rule ensureNoChildren_wp[where P = dc]) apply clarsimp - apply (erule_tac y = c in cmap_relationE1) - apply assumption - apply (clarsimp simp:cmdb_node_relation_mdbNext[OF ccte_relation_cmdbnode_relation]) - apply (intro conjI exI,assumption) - apply (erule(1) valid_capAligned[OF ctes_of_valid']) - apply (erule(1) ctes_of_valid') - done + apply (vcg exspec=ensureNoChildren_modifies) + apply (clarsimp simp: cte_wp_at_ctes_of) + done lemma ccap_relation_untyped_CL_simps: "\ccap_relation cp cap;isUntypedCap cp\ @@ -2274,13 +2909,13 @@ lemma unat_of_nat_APIType_capBits: done lemma valid_untyped_inv'_D: - "valid_untyped_inv' (Retype slot ptr_base ptr ty us slots isdev) s + "valid_untyped_inv' (Retype cref reset ptr_base ptr ty us destSlots isdev) s \ APIType_capBits ty us < 32" - apply (clarsimp simp:valid_untyped_inv'.simps) + apply clarsimp apply (drule range_cover_sz') apply (simp add:word_bits_def) done - + lemma object_type_from_to_H: "unat v \ (fromEnum::object_type \ nat) maxBound \ v = object_type_from_H (object_type_to_H v)" @@ -2355,7 +2990,7 @@ lemma Arch_isFrameType_spec: lemma decodeUntypedInvocation_ccorres_helper: notes TripleSuc[simp] -notes valid_untyped_inv'.simps[simp del] tl_drop_1[simp] +notes valid_untyped_inv_wcap'.simps[simp del] tl_drop_1[simp] shows "interpret_excaps extraCaps' = excaps_map extraCaps \ ccorres (intr_and_se_rel \ dc) (liftxf errstate id (K ()) ret__unsigned_long_') @@ -2459,7 +3094,7 @@ shows apply arith apply (rule syscall_error_throwError_ccorres_n) apply (simp add: syscall_error_to_H_cases) - apply (rule ccorres_Guard_Seq)+ + apply ((rule ccorres_Guard_Seq)+)? apply (rule ccorres_split_when_throwError_cond [where Q=\ and Q'=\, rotated -1]) apply vcg @@ -2705,7 +3340,11 @@ shows apply simp apply (rule ceqv_refl) apply (ctac (c_lines 2) add:checkFreeIndex_ccorres[unfolded fun_app_def]) - apply (rule_tac P = "rvb \ (capFreeIndex cp)" in ccorres_gen_asm) + apply (rename_tac reset reset_fi_tup) + apply (rule_tac xf'=reset_' in ccorres_abstract, ceqv) + apply (rule_tac xf'=freeIndex_' in ccorres_abstract, ceqv) + apply (rename_tac reset' fi', rule_tac P="reset_fi_tup = (fi', reset')" + in ccorres_gen_asm2) apply csymbr apply (rule ccorres_Guard_Seq)+ apply csymbr+ @@ -2715,19 +3354,24 @@ shows apply csymbr apply (rule ccorres_symb_exec_r) apply (rule_tac xf'=ret__int_' in ccorres_abstract, ceqv) - apply (rule_tac P = "rv'b = (if (unat (2 ^ capBlockSize cp - (xfdc << 4) + apply (rule_tac P = "rv'b = (if (unat (2 ^ capBlockSize cp - (fi' << 4) >> (APIType_capBits (toEnum (unat (hd args))) (unat (args ! Suc 0)))) < unat (args ! 5)) then 1 else 0)" in ccorres_gen_asm2) apply (rule ccorres_split_when_throwError_cond[where Q = \ and Q' = \]) - apply (clarsimp simp: ccap_relation_untyped_CL_simps shiftL_nat + apply (case_tac reset; + clarsimp simp: ccap_relation_untyped_CL_simps shiftL_nat valid_untyped_capBlockSize_misc + valid_untyped_capBlockSize_misc[where z=0, simplified] of_nat_shiftR) + apply (clarsimp simp:toEnum_object_type_to_H + unat_of_nat_APIType_capBits word_size hd_conv_nth length_ineq_not_Nil + split:if_splits) apply (clarsimp simp:toEnum_object_type_to_H unat_of_nat_APIType_capBits word_size hd_conv_nth length_ineq_not_Nil split:if_splits) apply (rule syscall_error_throwError_ccorres_n) - apply (clarsimp simp: syscall_error_rel_def + apply (case_tac reset; clarsimp simp: syscall_error_rel_def ccap_relation_untyped_CL_simps shiftL_nat syscall_error_to_H_cases valid_untyped_capBlockSize_misc) apply csymbr @@ -2754,18 +3398,18 @@ shows apply (simp (no_asm) add: ccorres_invocationCatch_Inr split_def performInvocation_def liftE_bindE bind_assoc) apply (ctac add: setThreadState_ccorres) - apply csymbr apply (rule ccorres_trim_returnE, (simp (no_asm))+) apply (simp (no_asm) add: o_def dc_def[symmetric] bindE_assoc id_def[symmetric] bind_bindE_assoc) - apply (simp (no_asm) only:alternative_distrib) + apply (rule ccorres_seq_skip'[THEN iffD1]) + apply (ctac(no_vcg) add: invokeUntyped_Retype_ccorres[where start = "args!4"]) apply (rule ccorres_alternative2) - apply (rule ccorres_call) - apply (rule_tac cnodeptr="capCNodePtr rv" in invokeUntyped_Retype_ccorres[where start = "args!4"]) - apply simp - apply simp - apply simp - apply (wp sts_invs_minor' setThreadStateRestart_ct_active' sts_valid_untyped_inv') + apply (rule ccorres_returnOk_skip) + apply (simp(no_asm) add: throwError_def, rule ccorres_return_Skip') + apply (rule hoare_vcg_conj_lift + | rule_tac p="capCNodePtr rv" in setThreadState_cap_to' + | wp_once sts_invs_minor' setThreadStateRestart_ct_active' + sts_valid_untyped_inv')+ apply (clarsimp simp: ccap_relation_untyped_CL_simps shiftL_nat toEnum_object_type_to_H unat_of_nat_APIType_capBits word_size valid_untyped_capBlockSize_misc getFreeRef_def hd_conv_nth length_ineq_not_Nil) @@ -2776,17 +3420,17 @@ shows apply (frule_tac cap = rv in cap_get_tag_to_H(5)) apply (simp add: cap_get_tag_isCap) apply (simp add: field_simps Suc_unat_diff_1) + apply (rule conjI) + apply (clarsimp split: bool.split_asm) apply (frule iffD2[OF olen_add_eqv]) + apply (frule(1) isUntypedCap_ccap_relation_helper) apply (clarsimp simp: unat_plus_simple[THEN iffD1]) apply (case_tac slots,simp) - apply (rule conjI) - apply clarsimp - apply (subst upto_enum_word) - apply (subst nth_map_upt) - apply (clarsimp simp: field_simps Suc_unat_diff_1 unat_plus_simple[THEN iffD1]) - apply (clarsimp simp: cte_level_bits_def) - apply (clarsimp simp: isFrameType_def - fromAPIType_def ARM_H.fromAPIType_def) + apply clarsimp + apply (subst upto_enum_word) + apply (subst nth_map_upt) + apply (clarsimp simp: field_simps Suc_unat_diff_1 unat_plus_simple[THEN iffD1]) + apply (clarsimp simp: cte_level_bits_def) apply simp apply wp apply simp @@ -2818,27 +3462,20 @@ shows apply (rule_tac Q' ="{sa. ksCurThread_' (globals sa) = tcb_ptr_to_ctcb_ptr (ksCurThread s)}" in conseqPost[where A' = "{}"]) - apply vcg - apply (clarsimp simp: toEnum_object_type_to_H not_le word_sle_def + apply (vcg exspec=ensureNoChildren_modifies + exspec=cap_untyped_cap_get_capFreeIndex_modifies) + apply (rule subsetI, + clarsimp simp:toEnum_object_type_to_H not_le word_sle_def enum_apiobject_type enum_object_type maxBound_is_length unat_of_nat_APIType_capBits word_size hd_conv_nth length_ineq_not_Nil - not_less word_le_nat_alt - split: if_splits) - apply (clarsimp simp: from_bool_0 ccap_relation_isDeviceCap2) - apply (intro conjI allI impI) - apply (clarsimp simp: shiftL_nat isCap_simps valid_cap_simps', simp add: word_bits_def) - apply (rule conjI allI) - apply (rule ccontr) - apply (clarsimp simp: not_less shiftr_overflow) - apply (intro conjI allI impI) - apply ((clarsimp simp: shiftL_nat isCap_simps valid_cap_simps' isFrameType_def - ThreadState_Restart_def not_less mask_def - split:split_if_asm | - simp add: false_def of_nat_power[where x = 5,simplified] not_less)+)[8] - apply (intro conjI allI impI APIType_capBits_high word_of_nat_less) - apply ((clarsimp simp: shiftL_nat isCap_simps valid_cap_simps' isFrameType_def - ThreadState_Restart_def not_less mask_def isFrameType_def split:split_if_asm | - simp add: false_def of_nat_power[where x = 5,simplified] not_less)+)[5] + not_less word_le_nat_alt isCap_simps valid_cap_simps') + apply (strengthen word_of_nat_less) + apply (clarsimp simp: StrictC'_thread_state_defs mask_def true_def false_def + from_bool_0 ccap_relation_isDeviceCap2 + split: split_if) + apply (intro conjI impI; clarsimp simp:not_less shiftr_overflow) + apply simp + apply simp apply (rule_tac Q'="\r. cte_wp_at' (\cte. cteCap cte = cp) slot and invs' and (\s. ksCurThread s = thread) and ex_cte_cap_to' (capCNodePtr rv) @@ -2865,10 +3502,12 @@ shows apply (clarsimp split:if_splits simp: not_less toEnum_object_type_to_H word_size hd_conv_nth length_ineq_not_Nil) apply (subgoal_tac "tcbQueued obja \ runnable' (tcbState obja)") - apply (simp add: trans [OF olen_add_eqv[symmetric] unat_plus_simple]) + apply (simp add: trans [OF olen_add_eqv[symmetric] unat_plus_simple] + fromAPIType_def) apply (clarsimp simp: word_le_nat_alt unat_2tp_if - valid_tcb_state'_def + valid_tcb_state'_def split: option.split_asm split_if_asm) + apply blast apply (case_tac "tcbState obja", (simp add: runnable'_def valid_tcb_state'_def)+)[1] apply simp @@ -2961,7 +3600,7 @@ shows signed_shift_guard_simpler_32 extra_sle_sless_unfolds elim!: inl_inrE - simp del: rf_sr_upd_safe) + simp del: rf_sr_upd_safe imp_disjL) apply (clarsimp simp:cap_get_tag_isCap[symmetric]) apply (rule conjI) apply (clarsimp simp: cap_get_tag_isCap[symmetric] @@ -2970,58 +3609,19 @@ shows apply (drule(1) cap_get_tag_to_H)+ apply (clarsimp simp: isCap_simps capAligned_def[unfolded capUntypedPtr_def, split_simps capability.split] objBits_simps word_bits_def) - apply (rule rf_sr_ctes_of_cliftE,assumption+) - apply (intro exI)+ - apply (rule conjI,assumption) - apply (rule conjI,erule(2) rf_sr_cte_relation) - apply (frule(1) h_t_valid_and_cslift_and_c_guard_field_mdbNext_CL[rotated -1]) - apply (clarsimp simp:cte_wp_at_ctes_of) - apply fastforce - apply (clarsimp simp: typ_heap_simps) - apply (frule_tac p = slot in valid_mdb_ctes_of_next[rotated]) - apply simp - apply fastforce - apply (clarsimp simp: cte_wp_at_ctes_of) - apply (frule_tac src = "(mdbNext_CL (cteMDBNode_CL ctel'))" in rf_sr_cte_relation) - apply simp - apply simp - apply (intro exI conjI,assumption) - apply (erule valid_capAligned[OF ctes_of_valid']) - apply fastforce - apply simp - apply (clarsimp elim!:inl_inrE) - apply (drule(1) cap_get_tag_to_H)+ - apply (rule conjI) - apply (clarsimp simp: cap_get_tag_isCap[symmetric] - capCNodeRadix_CL_less_32s rf_sr_ksCurThread not_le - elim!: inl_inrE) - apply (rule rf_sr_ctes_of_cliftE,assumption+) - apply (intro exI)+ - apply (rule conjI,assumption) - apply (rule conjI,erule(2) rf_sr_cte_relation) - apply (frule(1) h_t_valid_and_cslift_and_c_guard_field_mdbNext_CL[rotated -1]) - apply (clarsimp simp: cte_wp_at_ctes_of simp del: rf_sr_upd_safe) - apply fastforce - apply (clarsimp simp: typ_heap_simps simp del: rf_sr_upd_safe) - apply (frule_tac p = slot in valid_mdb_ctes_of_next[rotated]) - apply (simp del: rf_sr_upd_safe) - apply fastforce - apply (clarsimp simp: cte_wp_at_ctes_of) - apply (frule_tac src = "(mdbNext_CL (cteMDBNode_CL ctel'))" in rf_sr_cte_relation) - apply simp - apply simp - apply (intro exI conjI,assumption) - apply (erule valid_capAligned[OF ctes_of_valid']) - apply fastforce - apply simp - apply (clarsimp simp: cap_get_tag_isCap[symmetric]) + + apply (clarsimp simp: cap_get_tag_isCap[symmetric] + capCNodeRadix_CL_less_32s rf_sr_ksCurThread not_le + elim!: inl_inrE) apply (drule(1) cap_get_tag_to_H)+ - apply (clarsimp simp: capAligned_def objBits_simps word_bits_def) + apply (clarsimp simp: isCap_simps capAligned_def[unfolded capUntypedPtr_def, split_simps capability.split] + objBits_simps word_bits_def) + done lemma decodeUntypedInvocation_ccorres: notes TripleSuc[simp] -notes valid_untyped_inv'.simps[simp del] +notes valid_untyped_inv_wcap'.simps[simp del] shows "interpret_excaps extraCaps' = excaps_map extraCaps \ ccorres (intr_and_se_rel \ dc) (liftxf errstate id (K ()) ret__unsigned_long_') diff --git a/proof/crefine/IpcCancel_C.thy b/proof/crefine/IpcCancel_C.thy index 2e03b3094..a9e4a894f 100644 --- a/proof/crefine/IpcCancel_C.thy +++ b/proof/crefine/IpcCancel_C.thy @@ -115,6 +115,17 @@ lemma ntfn_ptr_get_queue_spec: declare td_names_word8[simp] +abbreviation + "cslift_all_but_tcb_C s t \ (cslift s :: cte_C typ_heap) = cslift t + \ (cslift s :: endpoint_C typ_heap) = cslift t + \ (cslift s :: notification_C typ_heap) = cslift t + \ (cslift s :: asid_pool_C typ_heap) = cslift t + \ (cslift s :: pte_C typ_heap) = cslift t + \ (cslift s :: pde_C typ_heap) = cslift t + \ (cslift s :: user_data_C typ_heap) = cslift t + \ (cslift s :: user_data_device_C typ_heap) = cslift t +" + lemma tcbEPDequeue_spec: "\s queue. \ \ \s. \t. (t, s) \ rf_sr \ (\tcb\set queue. tcb_at' tcb t) \ distinct queue @@ -136,7 +147,10 @@ lemma tcbEPDequeue_spec: (cslift s |` (- tcb_ptr_to_ctcb_ptr ` set queue)) \ option_map tcb_null_ep_ptrs \ (cslift t) = option_map tcb_null_ep_ptrs \ (cslift s)) - \ cslift_all_but_tcb_C t s \ (hrs_htd \<^bsup>t\<^esup>t_hrs) = (hrs_htd \<^bsup>s\<^esup>t_hrs)}" + \ cslift_all_but_tcb_C t s + \ (\rs. zero_ranges_are_zero rs (\<^bsup>t\<^esup>t_hrs) + = zero_ranges_are_zero rs (\<^bsup>s\<^esup>t_hrs)) + \ (hrs_htd \<^bsup>t\<^esup>t_hrs) = (hrs_htd \<^bsup>s\<^esup>t_hrs)}" apply (intro allI) apply (rule conseqPre) apply vcg @@ -177,11 +191,11 @@ lemma ntfn_ptr_set_queue_spec: "\s. \ \ \s. s \\<^sub>c \ntfnPtr\ Call ntfn_ptr_set_queue_'proc {t. (\ntfn'. notification_lift ntfn' = (notification_lift (the (cslift s (\<^bsup>s\<^esup>ntfnPtr))))\ ntfnQueue_head_CL := ptr_val (head_C \<^bsup>s\<^esup>ntfn_queue) && ~~ mask 4, - ntfnQueue_tail_CL := ptr_val (end_C \<^bsup>s\<^esup>ntfn_queue) && ~~ mask 4 \ \ - (cslift t :: notification_C typ_heap) = (cslift s)(\<^bsup>s\<^esup>ntfnPtr \ ntfn')) - \ cslift_all_but_notification_C t s \ (hrs_htd \<^bsup>t\<^esup>t_hrs) = (hrs_htd \<^bsup>s\<^esup>t_hrs)}" - apply vcg - apply (auto simp: split_def h_t_valid_clift_Some_iff) + ntfnQueue_tail_CL := ptr_val (end_C \<^bsup>s\<^esup>ntfn_queue) && ~~ mask 4 \ + \ t_hrs_' (globals t) = hrs_mem_update (heap_update (\<^bsup>s\<^esup>ntfnPtr) ntfn') + (t_hrs_' (globals s)))}" + apply (rule allI, rule conseqPre, vcg) + apply (clarsimp simp: packed_heap_update_collapse_hrs typ_heap_simps') done lemma cancelSignal_ccorres_helper: @@ -230,7 +244,9 @@ lemma cancelSignal_ccorres_helper: apply simp apply (simp add: setNotification_def split_def) apply (rule bexI [OF _ setObject_eq]) - apply (simp add: remove1_empty rf_sr_def cstate_relation_def Let_def cpspace_relation_def update_ntfn_map_tos) + apply (simp add: remove1_empty rf_sr_def cstate_relation_def Let_def + cpspace_relation_def update_ntfn_map_tos + typ_heap_simps') apply (elim conjE) apply (intro conjI) -- "tcb relation" @@ -249,13 +265,15 @@ lemma cancelSignal_ccorres_helper: -- "queue relation" apply (rule cready_queues_relation_null_queue_ptrs, assumption+) apply (clarsimp simp: comp_def) - apply (simp add: carch_state_relation_def carch_globals_def) + apply (simp add: carch_state_relation_def carch_globals_def + typ_heap_simps') apply (simp add: cmachine_state_relation_def) apply (simp add: h_t_valid_clift_Some_iff) apply (simp add: objBits_simps) apply (simp add: objBits_simps) apply assumption + -- "non empty case" apply (frule tcb_queue_head_empty_iff [OF tcb_queue_relation'_queue_rel]) apply (rule ballI, erule bspec) @@ -264,7 +282,9 @@ lemma cancelSignal_ccorres_helper: apply (simp add: setNotification_def split_def) apply (rule bexI [OF _ setObject_eq]) apply (frule (1) st_tcb_at_h_t_valid) - apply (simp add: remove1_empty rf_sr_def cstate_relation_def Let_def cpspace_relation_def update_ntfn_map_tos) + apply (simp add: remove1_empty rf_sr_def cstate_relation_def Let_def + cpspace_relation_def update_ntfn_map_tos + typ_heap_simps') apply (elim conjE) apply (intro conjI) -- "tcb relation" @@ -291,7 +311,8 @@ lemma cancelSignal_ccorres_helper: -- "queue relation" apply (rule cready_queues_relation_null_queue_ptrs, assumption+) apply (clarsimp simp: comp_def) - apply (simp add: carch_state_relation_def carch_globals_def) + apply (simp add: carch_state_relation_def carch_globals_def + typ_heap_simps') apply (simp add: cmachine_state_relation_def) apply (simp add: h_t_valid_clift_Some_iff) apply (simp add: objBits_simps) @@ -299,6 +320,8 @@ lemma cancelSignal_ccorres_helper: apply assumption done +lemmas rf_sr_tcb_update_no_queue_gen + = rf_sr_tcb_update_no_queue[where t="t''\ globals := gs \ t_hrs_' := th \\" for th, simplified] lemma threadSet_tcbState_simple_corres: "ccorres dc xfdc (tcb_at' thread) @@ -308,10 +331,10 @@ lemma threadSet_tcbState_simple_corres: apply (rule threadSet_corres_lemma) apply (rule thread_state_ptr_set_tsType_spec) apply (rule thread_state_ptr_set_tsType_modifies) - apply clarsimp - apply (frule (1) obj_at_cslift_tcb) apply clarsimp - apply (rule rf_sr_tcb_update_no_queue, assumption+, simp_all) + apply (frule (1) obj_at_cslift_tcb) + apply (clarsimp simp: typ_heap_simps') + apply (rule rf_sr_tcb_update_no_queue_gen, assumption+, simp, simp_all) apply (rule ball_tcb_cte_casesI, simp_all) apply (frule cmap_relation_tcb) apply (frule (1) cmap_relation_ko_atD) @@ -321,7 +344,6 @@ lemma threadSet_tcbState_simple_corres: apply (clarsimp simp: typ_heap_simps) done - lemma ko_at_obj_congD': "\ko_at' k p s; ko_at' k' p s\ \ k = k'" apply (erule obj_atE')+ @@ -643,7 +665,7 @@ lemma threadSet_queued_ccorres [corres]: apply clarsimp apply (frule (1) obj_at_cslift_tcb) apply clarsimp - apply (rule rf_sr_tcb_update_no_queue, assumption+, simp_all) + apply (rule rf_sr_tcb_update_no_queue_gen, assumption+, simp, simp_all) apply (rule ball_tcb_cte_casesI, simp_all) apply (simp add: ctcb_relation_def cthread_state_relation_def) apply (case_tac "tcbState ko", simp_all add: Word_Lemmas.from_bool_mask_simp)[1] @@ -699,6 +721,8 @@ lemma state_relation_queue_update_helper': option_map tcb_null_sched_ptrs \ cslift t = option_map tcb_null_sched_ptrs \ cslift s'; cslift_all_but_tcb_C t s'; + zero_ranges_are_zero (gsUntypedZeroRanges s) (f (t_hrs_' (globals s'))) + = zero_ranges_are_zero (gsUntypedZeroRanges s) (t_hrs_' (globals s')); hrs_htd (t_hrs_' (globals t)) = hrs_htd (t_hrs_' (globals s')); prio' = cready_queues_index_to_C qdom prio; \x \ S. obj_at' (inQ qdom prio) x s @@ -770,6 +794,8 @@ lemma state_relation_queue_update_helper: option_map tcb_null_sched_ptrs \ cslift t = option_map tcb_null_sched_ptrs \ cslift s'; cslift_all_but_tcb_C t s'; + zero_ranges_are_zero (gsUntypedZeroRanges s) (f (t_hrs_' (globals s'))) + = zero_ranges_are_zero (gsUntypedZeroRanges s) (t_hrs_' (globals s')); hrs_htd (t_hrs_' (globals t)) = hrs_htd (t_hrs_' (globals s')); prio' = cready_queues_index_to_C qdom prio; \x \ S. obj_at' (inQ qdom prio) x s @@ -781,7 +807,7 @@ lemma state_relation_queue_update_helper: \ (s \ksReadyQueues := (ksReadyQueues s)((qdom, prio) := q)\, t) \ rf_sr" apply (subgoal_tac "\d p. (\t\set (ksReadyQueues s (d, p)). obj_at' (inQ d p) t s) \ distinct(ksReadyQueues s (d, p))") - apply (erule(12) state_relation_queue_update_helper') + apply (erule(5) state_relation_queue_update_helper', simp_all) apply (clarsimp simp: valid_queues_def valid_queues_no_bitmap_def) apply (drule_tac x=d in spec) apply (drule_tac x=p in spec) @@ -1160,6 +1186,7 @@ proof - simp_all add: valid_queues_valid_q)[1] apply (rule tcb_at_not_NULL, erule obj_at'_weakenE, simp) apply (erule(1) state_relation_queue_update_helper[where S="{t}"], + (simp | rule globals.equality)+, simp_all add: cready_queues_index_to_C_def2 numPriorities_def t_hrs_ksReadyQueues_upd_absorb upd_unless_null_def typ_heap_simps)[1] @@ -1202,6 +1229,7 @@ proof - apply (frule(2) obj_at_cslift_tcb[OF valid_queues_obj_at'D]) apply (clarsimp simp: h_val_field_clift' h_t_valid_clift) apply (erule_tac S="{t, v}" for v in state_relation_queue_update_helper, + (simp | rule globals.equality)+, simp_all add: typ_heap_simps if_Some_helper numPriorities_def cready_queues_index_to_C_def2 upd_unless_null_def cong: if_cong split del: split_if @@ -1559,7 +1587,8 @@ proof - apply (clarsimp simp: h_t_valid_c_guard [OF h_t_valid_field, OF h_t_valid_clift] h_t_valid_field[OF h_t_valid_clift] h_t_valid_clift) apply (erule_tac S="set (ksReadyQueues \ (tcbDomain ko, tcbPriority ko))" - in state_relation_queue_update_helper', assumption, + in state_relation_queue_update_helper', + (simp | rule globals.equality)+, simp_all add: clift_field_update if_Some_helper numPriorities_def cready_queues_index_to_C_def2 typ_heap_simps maxDom_to_H maxPrio_to_H @@ -1567,14 +1596,14 @@ proof - apply (fold_subgoals (prefix))[2] subgoal premises prems using prems by (fastforce simp: tcb_null_sched_ptrs_def)+ apply (erule_tac S="set (ksReadyQueues \ (tcbDomain ko, tcbPriority ko))" - in state_relation_queue_update_helper', assumption, + in state_relation_queue_update_helper', + (simp | rule globals.equality)+, simp_all add: clift_field_update if_Some_helper numPriorities_def cready_queues_index_to_C_def2 maxDom_to_H maxPrio_to_H - cong: if_cong split del: split_if)[1] - subgoal by (fastforce simp: typ_heap_simps) - subgoal by (fastforce simp: tcb_null_sched_ptrs_def) - subgoal by (simp add: typ_heap_simps) + cong: if_cong split del: split_if, + simp_all add: typ_heap_simps')[1] + subgoal by (fastforce simp: tcb_null_sched_ptrs_def) subgoal by fastforce apply clarsimp @@ -1587,7 +1616,8 @@ proof - elim: obj_at'_weaken)+ apply (clarsimp simp: h_val_field_clift' h_t_valid_clift) apply (erule_tac S="set (ksReadyQueues \ (tcbDomain ko, tcbPriority ko))" - in state_relation_queue_update_helper', assumption, + in state_relation_queue_update_helper', + (simp | rule globals.equality)+, simp_all add: clift_field_update if_Some_helper numPriorities_def cready_queues_index_to_C_def2 maxDom_to_H maxPrio_to_H @@ -1656,14 +1686,16 @@ proof - apply (clarsimp simp: typ_heap_simps) apply (rule conjI; clarsimp simp: typ_heap_simps) apply (erule_tac S="set (ksReadyQueues \ (tcbDomain ko, tcbPriority ko))" - in state_relation_queue_update_helper', assumption, + in state_relation_queue_update_helper', + (simp | rule globals.equality)+, simp_all add: clift_field_update if_Some_helper numPriorities_def cready_queues_index_to_C_def2 maxDom_to_H maxPrio_to_H cong: if_cong split del: split_if)[1] apply (fastforce simp: typ_heap_simps tcb_null_sched_ptrs_def)+ apply (erule_tac S="set (ksReadyQueues \ (tcbDomain ko, tcbPriority ko))" - in state_relation_queue_update_helper', assumption, + in state_relation_queue_update_helper', + (simp | rule globals.equality)+, simp_all add: clift_field_update if_Some_helper numPriorities_def cready_queues_index_to_C_def2 maxDom_to_H maxPrio_to_H @@ -1683,12 +1715,13 @@ proof - apply (fold_subgoals (prefix))[2] subgoal premises prems using prems by (fastforce elim: obj_at'_weaken)+ apply (erule_tac S="set (ksReadyQueues \ (tcbDomain ko, tcbPriority ko))" - in state_relation_queue_update_helper', assumption, + in state_relation_queue_update_helper', + (simp | rule globals.equality)+, simp_all add: clift_field_update if_Some_helper numPriorities_def cready_queues_index_to_C_def2 typ_heap_simps maxDom_to_H maxPrio_to_H cong: if_cong split del: split_if)[1] - apply (fold_subgoals (prefix))[3] + apply (fold_subgoals (prefix))[2] subgoal premises prems using prems by (fastforce simp: typ_heap_simps tcb_null_sched_ptrs_def)+ apply (simp add: guard_is_UNIV_def) @@ -1857,6 +1890,7 @@ proof - apply (subst rf_sr_drop_bitmaps_enqueue_helper, assumption) apply (fastforce intro: cbitmap_L1_relation_bit_set cbitmap_L2_relation_bit_set)+ apply (erule(1) state_relation_queue_update_helper[where S="{t}"], + (simp | rule globals.equality)+, simp_all add: cready_queues_index_to_C_def2 numPriorities_def t_hrs_ksReadyQueues_upd_absorb upd_unless_null_def typ_heap_simps)[1] @@ -1897,6 +1931,7 @@ proof - apply (frule(2) obj_at_cslift_tcb[OF valid_queues_obj_at'D]) apply (clarsimp simp: h_val_field_clift' h_t_valid_clift) apply (erule_tac S="{t, v}" for v in state_relation_queue_update_helper, + (simp | rule globals.equality)+, simp_all add: typ_heap_simps if_Some_helper numPriorities_def cready_queues_index_to_C_def2 upd_unless_null_def cong: if_cong split del: split_if @@ -2487,11 +2522,12 @@ lemma ep_ptr_set_queue_spec: "\s. \ \ \s. s \\<^sub>c \epptr\ Call ep_ptr_set_queue_'proc {t. (\ep'. endpoint_lift ep' = (endpoint_lift (the (cslift s (\<^bsup>s\<^esup>epptr))))\ epQueue_head_CL := ptr_val (head_C \<^bsup>s\<^esup>queue) && ~~ mask 4, - epQueue_tail_CL := ptr_val (end_C \<^bsup>s\<^esup>queue) && ~~ mask 4 \ \ - (cslift t :: endpoint_C typ_heap) = (cslift s)(\<^bsup>s\<^esup>epptr \ ep')) - \ cslift_all_but_endpoint_C t s \ (hrs_htd \<^bsup>t\<^esup>t_hrs) = (hrs_htd \<^bsup>s\<^esup>t_hrs)}" + epQueue_tail_CL := ptr_val (end_C \<^bsup>s\<^esup>queue) && ~~ mask 4 \ + \ t_hrs_' (globals t) = hrs_mem_update (heap_update (\<^bsup>s\<^esup>epptr) ep') + (t_hrs_' (globals s)))}" apply vcg - apply (auto simp: split_def h_t_valid_clift_Some_iff) + apply (auto simp: split_def h_t_valid_clift_Some_iff + typ_heap_simps packed_heap_update_collapse_hrs) done lemma valid_ep_blockedD: @@ -2724,7 +2760,8 @@ lemma cancelIPC_ccorres_helper: subgoal by simp apply (simp add: setEndpoint_def split_def) apply (rule bexI [OF _ setObject_eq]) - apply (simp add: remove1_empty rf_sr_def cstate_relation_def Let_def cpspace_relation_def update_ep_map_tos) + apply (simp add: remove1_empty rf_sr_def cstate_relation_def Let_def + cpspace_relation_def update_ep_map_tos typ_heap_simps') apply (elim conjE) apply (intro conjI) -- "tcb relation" @@ -2744,7 +2781,8 @@ lemma cancelIPC_ccorres_helper: -- "queue relation" apply (rule cready_queues_relation_null_queue_ptrs, assumption+) subgoal by (clarsimp simp: comp_def) - subgoal by (simp add: carch_state_relation_def carch_globals_def) + subgoal by (simp add: carch_state_relation_def carch_globals_def + typ_heap_simps') subgoal by (simp add: cmachine_state_relation_def) subgoal by (simp add: h_t_valid_clift_Some_iff) subgoal by (simp add: objBits_simps) @@ -2759,7 +2797,8 @@ lemma cancelIPC_ccorres_helper: apply (simp add: setEndpoint_def split_def) apply (rule bexI [OF _ setObject_eq]) apply (frule (1) st_tcb_at_h_t_valid) - apply (simp add: remove1_empty rf_sr_def cstate_relation_def Let_def cpspace_relation_def update_ep_map_tos) + apply (simp add: remove1_empty rf_sr_def cstate_relation_def Let_def + cpspace_relation_def update_ep_map_tos typ_heap_simps') apply (elim conjE) apply (intro conjI) -- "tcb relation" @@ -2786,7 +2825,8 @@ lemma cancelIPC_ccorres_helper: -- "queue relation" apply (rule cready_queues_relation_null_queue_ptrs, assumption+) subgoal by (clarsimp simp: comp_def) - subgoal by (simp add: carch_state_relation_def carch_globals_def) + subgoal by (simp add: carch_state_relation_def carch_globals_def + typ_heap_simps') subgoal by (simp add: cmachine_state_relation_def) subgoal by (simp add: h_t_valid_clift_Some_iff) subgoal by (simp add: objBits_simps) @@ -2813,6 +2853,8 @@ lemma getThreadState_ccorres_foo: apply (clarsimp simp: ctcb_relation_def obj_at'_def) done +(* +thm cancelIPC_body_def lemma cancelIPC_ccorres_reply_helper: assumes cteDeleteOne_ccorres: "\w slot. ccorres dc xfdc @@ -2833,7 +2875,10 @@ lemma cancelIPC_ccorres_reply_helper: cteDeleteOne callerCap od) od) - (CALL fault_null_fault_ptr_new(Ptr &(tcb_ptr_to_ctcb_ptr thread\[''tcbFault_C'']));; + (\ret__struct_fault_C :== CALL fault_null_fault_new();; + Guard C_Guard {s. s \\<^sub>c tptr_' s} + (Basic (\s. s \ (\t_hrs :== (hrs_mem_update (heap_update + (Ptr &(\tptr\[''tcbFault_C''])) \ret__struct_fault_C)));; (Guard MemorySafety \ptr_add_assertion (cte_Ptr (ptr_val (tcb_ptr_to_ctcb_ptr thread) && 0xFFFFFE00)) (sint Kernel_C.tcbReply) False (hrs_htd \t_hrs)\ @@ -2900,6 +2945,7 @@ lemma cancelIPC_ccorres_reply_helper: Kernel_C.tcbReply_def mask_def tcbCNodeEntries_def) apply (fastforce simp: pred_tcb_at' inQ_def tcb_aligned'[OF pred_tcb_at']) done +*) lemma ep_blocked_in_queueD_recv: "\st_tcb_at' (op = (Structures_H.thread_state.BlockedOnReceive x)) thread \; ko_at' ep' x \; invs' \\ \ thread \ set (epQueue ep') \ isRecvEP ep'" @@ -2971,10 +3017,60 @@ lemma cancelIPC_ccorres1: apply csymbr apply csymbr apply (unfold comp_def)[1] - apply (rule ccorres_Guard ccorres_Guard_Seq)+ - apply (clarsimp simp del: dc_simp simp: of_int_sint) - apply ccorres_remove_UNIV_guard - apply (rule cancelIPC_ccorres_reply_helper [OF cteDeleteOne_ccorres]) + apply csymbr + apply (rule ccorres_move_c_guard_tcb)+ + apply (rule ccorres_split_nothrow_novcg) + apply (rule_tac P=\ in threadSet_ccorres_lemma2) + apply vcg + apply (clarsimp simp: typ_heap_simps') + apply (erule(1) rf_sr_tcb_update_no_queue2, + (simp add: typ_heap_simps')+)[1] + apply (rule ball_tcb_cte_casesI, simp_all)[1] + apply (clarsimp simp: ctcb_relation_def fault_lift_null_fault + cfault_rel_def cthread_state_relation_def) + apply (case_tac "tcbState tcb", simp_all add: is_cap_fault_def)[1] + apply ceqv + apply ccorres_remove_UNIV_guard + apply (rule ccorres_move_array_assertion_tcb_ctes) + apply (rule_tac P="tcb_at' thread" in ccorres_cross_over_guard) + apply (simp add: getThreadReplySlot_def) + apply ctac + apply (simp only: liftM_def bind_assoc return_bind del: Collect_const) + apply (rule ccorres_pre_getCTE) + apply (rename_tac slot slot' cte) + apply (rule ccorres_move_c_guard_cte) + apply (rule_tac xf'=ret__unsigned_' and val="mdbNext (cteMDBNode cte)" + and R="cte_wp_at' (op = cte) slot and invs'" + in ccorres_symb_exec_r_known_rv_UNIV[where R'=UNIV]) + apply vcg + apply (clarsimp simp: cte_wp_at_ctes_of) + apply (erule(1) cmap_relationE1[OF cmap_relation_cte]) + apply (clarsimp simp: typ_heap_simps) + apply (clarsimp simp: ccte_relation_def map_option_Some_eq2) + apply ceqv + apply csymbr + apply (rule ccorres_Cond_rhs) + apply (simp add: nullPointer_def when_def) + apply (rule ccorres_symb_exec_l[OF _ _ _ empty_fail_stateAssert]) + apply (simp only: dc_def[symmetric]) + apply (rule ccorres_symb_exec_r) + apply (ctac add: cteDeleteOne_ccorres[where w1="scast cap_reply_cap"]) + apply vcg + apply (rule conseqPre, vcg, clarsimp simp: rf_sr_def + gs_set_assn_Delete_cstate_relation[unfolded o_def]) + apply (wp | simp)+ + apply (simp add: when_def nullPointer_def dc_def[symmetric]) + apply (rule ccorres_return_Skip) + apply (simp add: guard_is_UNIV_def ghost_assertion_data_get_def + ghost_assertion_data_set_def cap_tag_defs) + apply (simp add: locateSlot_conv, wp) + apply vcg + apply (rule_tac Q="\rv. tcb_at' thread and invs'" in hoare_post_imp) + apply (clarsimp simp: cte_wp_at_ctes_of capHasProperty_def + cap_get_tag_isCap ucast_id) + apply (wp threadSet_invs_trivial | simp)+ + apply (clarsimp simp add: guard_is_UNIV_def tcbReplySlot_def + Kernel_C.tcbReply_def tcbCNodeEntries_def) -- "BlockedOnNotification" apply (simp add: word_sle_def "StrictC'_thread_state_defs" ccorres_cond_iffs dc_def [symmetric] cong: call_ignore_cong) apply (rule ccorres_symb_exec_r) @@ -3046,11 +3142,15 @@ lemma cancelIPC_ccorres1: cthread_state_relation_def sch_act_wf_weak valid_ep'_def dest!: valid_queues_not_runnable'_not_ksQ[where t=thread] split: thread_state.splits endpoint.splits) apply (rule conjI) - apply (clarsimp simp: isBlockedOnReply_def pred_tcb_at'_def obj_at'_def) - apply (rule conjI, clarsimp) + apply (clarsimp simp: inQ_def) + apply (rule conjI) + apply clarsimp apply clarsimp apply (rule conjI) - subgoal by (auto simp: pred_tcb_at'_def isBlockedOnReceive_def isBlockedOnSend_def obj_at'_def projectKOs split: thread_state.splits)[1] + subgoal by (auto simp: obj_at'_def projectKOs pred_tcb_at'_def invs'_def valid_state'_def + isTS_defs cte_wp_at_ctes_of + cthread_state_relation_def sch_act_wf_weak valid_ep'_def + dest!: valid_queues_not_runnable'_not_ksQ[where t=thread] split: thread_state.splits) apply clarsimp apply (rule conjI) subgoal by (auto simp: obj_at'_def projectKOs pred_tcb_at'_def invs'_def valid_state'_def @@ -3066,7 +3166,7 @@ lemma cancelIPC_ccorres1: dest!: valid_queues_not_runnable'_not_ksQ[where t=thread] split: thread_state.splits endpoint.splits)[1] apply (auto simp: isTS_defs cthread_state_relation_def typ_heap_simps weak_sch_act_wf_def) apply (case_tac ts, - auto simp: isTS_defs cthread_state_relation_def typ_heap_simps) + auto simp: isTS_defs cthread_state_relation_def typ_heap_simps) done end diff --git a/proof/crefine/Ipc_C.thy b/proof/crefine/Ipc_C.thy index c0758bdd9..b518864cb 100644 --- a/proof/crefine/Ipc_C.thy +++ b/proof/crefine/Ipc_C.thy @@ -3532,6 +3532,7 @@ lemma handleFaultReply_ccorres [corres]: message_info_to_H_def split: split_if) apply unat_arith + apply clarsimp apply (vcg spec=TrueI) apply clarsimp apply wp @@ -3616,6 +3617,7 @@ lemma handleFaultReply_ccorres [corres]: min_def message_info_to_H_def word_of_nat_less split: split_if) apply unat_arith + apply clarsimp apply (vcg spec=TrueI) apply clarsimp apply wp @@ -3728,27 +3730,6 @@ lemma handleFaultReply_ccorres [corres]: apply (clarsimp simp: obj_at'_def n_msgRegisters_def) done -lemma fault_null_fault_ptr_new_ccorres [corres]: - "ccorres dc xfdc (tcb_at' receiver) UNIV hs - (threadSet (tcbFault_update empty) receiver) - (CALL fault_null_fault_ptr_new(Ptr - &(tcb_ptr_to_ctcb_ptr receiver\[''tcbFault_C''])))" - apply (rule ccorres_from_vcg) - apply (rule allI) - apply (rule conseqPre, vcg) - apply clarsimp - apply (frule(1) obj_at_cslift_tcb) - apply (rule conjI) - apply (clarsimp simp: typ_heap_simps) - apply clarsimp - apply (rule rev_bexI [OF threadSet_eq], assumption) - apply clarsimp - apply (erule(2) rf_sr_tcb_update_no_queue, simp_all add: tcb_cte_cases_def) - apply (clarsimp simp: ctcb_relation_def cthread_state_relation_def - fault_lift_def cfault_rel_def) - apply (case_tac "tcbState ko", simp_all add: is_cap_fault_def) - done - (* FIXME: move *) lemma cancelAllIPC_sch_act_wf: "\\s. sch_act_wf (ksSchedulerAction s) s\ @@ -3795,58 +3776,20 @@ lemma setCTE_tcbFault: apply (rule setObject_cte_obj_at_tcb', simp_all) done -(* FIXME: move *) -lemma emptySlot_tcbFault: - "\obj_at' (\tcb. P (tcbFault tcb)) t\ - emptySlot slot irq - \\_. obj_at' (\tcb. P (tcbFault tcb)) t\" - apply (simp add: emptySlot_def) - apply (wp setCTE_tcbFault getCTE_wp' hoare_drop_imps - | wpc - | simp add: updateMDB_def updateCap_def Let_def - | rule conjI impI)+ - done +crunch tcbFault: emptySlot, tcbSchedEnqueue, rescheduleRequired + "obj_at' (\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 -(* FIXME: move *) -lemma tcbSchedEnqueue_tcbFault: - "\obj_at' (\tcb. P (tcbFault tcb)) t\ - tcbSchedEnqueue t' - \\_. obj_at' (\tcb. P (tcbFault tcb)) t\" - apply (simp add: tcbSchedEnqueue_def unless_when) - apply (wp threadSet_obj_at' hoare_drop_imps threadGet_wp - |simp split: split_if)+ - done - -lemma rescheduleRequired_tcbFault: - "\obj_at' (\tcb. P (tcbFault tcb)) t\ - rescheduleRequired - \\_. obj_at' (\tcb. P (tcbFault tcb)) t\" - apply (simp add: rescheduleRequired_def) - apply (wp tcbSchedEnqueue_tcbFault | wpc)+ - apply simp - done +crunch tcbFault: setThreadState, cancelAllIPC, cancelAllSignals + "obj_at' (\tcb. P (tcbFault tcb)) t" + (wp: threadSet_obj_at'_strongish crunch_wps) (* FIXME: move *) -lemma sts_tcbFault: - "\obj_at' (\tcb. P (tcbFault tcb)) t\ - setThreadState st t' - \\_. obj_at' (\tcb. P (tcbFault tcb)) t\" - apply (simp add: setThreadState_def) - apply (wp threadSet_obj_at' rescheduleRequired_tcbFault | simp)+ - done - -(* FIXME: move *) -lemma setEndpoint_tcb: - "\obj_at' (\tcb::tcb. P tcb) t\ - setEndpoint ep e - \\_. obj_at' P t\" - apply (simp add: setEndpoint_def) - apply (rule obj_at_setObject2) - apply (clarsimp simp: updateObject_default_def in_monad) - done +lemmas setEndpoint_tcb = KHeap_R.setEndpoint_obj_at'_tcb (* FIXME: move *) lemma setNotification_tcb: @@ -3858,30 +3801,6 @@ lemma setNotification_tcb: apply (clarsimp simp: updateObject_default_def in_monad) done -(* FIXME: move *) -lemma cancelAllIPC_tcbFault: - "\obj_at' (\tcb. P (tcbFault tcb)) t\ - cancelAllIPC ep - \\_. obj_at' (\tcb. P (tcbFault tcb)) t\" - apply (simp add: cancelAllIPC_def) - apply (rule order_refl | - wp mapM_x_wp tcbSchedEnqueue_tcbFault sts_tcbFault - setEndpoint_tcb getEndpoint_wp rescheduleRequired_tcbFault | - wpc | simp)+ - done - -(* FIXME: move *) -lemma cancelAllSignals_tcbFault: - "\obj_at' (\tcb. P (tcbFault tcb)) t\ - cancelAllSignals ep - \\_. obj_at' (\tcb. P (tcbFault tcb)) t\" - apply (simp add: cancelAllSignals_def) - apply (rule order_refl | - wp mapM_x_wp tcbSchedEnqueue_tcbFault sts_tcbFault - setNotification_tcb getNotification_wp rescheduleRequired_tcbFault | - wpc | simp)+ - done - lemma sbn_tcbFault: "\obj_at' (\tcb. P (tcbFault tcb)) t\ setBoundNotification st t' @@ -4061,9 +3980,19 @@ proof - where xf'=restart_']) apply simp_all[3] apply ceqv - apply (rule ccorres_move_c_guard_tcb) + apply csymbr apply (rule ccorres_split_nothrow_novcg) - apply (rule fault_null_fault_ptr_new_ccorres) + apply (rule threadSet_ccorres_lemma2[where P=\]) + apply vcg + apply (clarsimp simp: typ_heap_simps') + apply (erule(1) rf_sr_tcb_update_no_queue2, + (simp add: typ_heap_simps')+, simp_all?)[1] + apply (rule ball_tcb_cte_casesI, simp+) + apply (clarsimp simp: ctcb_relation_def + fault_lift_null_fault + cfault_rel_def is_cap_fault_def + cthread_state_relation_def) + apply (case_tac "tcbState tcb", simp_all add: is_cap_fault_def)[1] apply ceqv apply (rule_tac R=\ in ccorres_cond2) apply (clarsimp simp: to_bool_def Collect_const_mem) @@ -4256,7 +4185,8 @@ lemma sendIPC_dequeue_ccorres_helper: apply (rule conjI) apply (rule bexI [OF _ setObject_eq]) apply (simp add: rf_sr_def cstate_relation_def Let_def - cpspace_relation_def update_ep_map_tos) + cpspace_relation_def update_ep_map_tos + typ_heap_simps') apply (elim conjE) apply (intro conjI) -- "tcb relation" @@ -4276,7 +4206,7 @@ lemma sendIPC_dequeue_ccorres_helper: -- "queue relation" apply (rule cready_queues_relation_null_queue_ptrs, assumption+) apply (clarsimp simp: comp_def) - apply (simp add: carch_state_relation_def) + apply (simp add: carch_state_relation_def typ_heap_simps') apply (simp add: cmachine_state_relation_def) apply (simp add: h_t_valid_clift_Some_iff) apply (simp add: objBits_simps) @@ -4295,7 +4225,8 @@ lemma sendIPC_dequeue_ccorres_helper: apply (rule bexI [OF _ setObject_eq]) apply (frule(1) st_tcb_at_h_t_valid) apply (simp add: rf_sr_def cstate_relation_def Let_def - cpspace_relation_def update_ep_map_tos) + cpspace_relation_def update_ep_map_tos + typ_heap_simps') apply (elim conjE) apply (intro conjI) -- "tcb relation" @@ -4323,7 +4254,7 @@ lemma sendIPC_dequeue_ccorres_helper: -- "queue relation" apply (rule cready_queues_relation_null_queue_ptrs, assumption+) apply (clarsimp simp: comp_def) - apply (simp add: carch_state_relation_def) + apply (simp add: carch_state_relation_def typ_heap_simps') apply (simp add: cmachine_state_relation_def) apply (simp add: h_t_valid_clift_Some_iff) apply (simp add: objBits_simps) @@ -4332,6 +4263,18 @@ lemma sendIPC_dequeue_ccorres_helper: apply (clarsimp simp: cendpoint_relation_def Let_def tcb_queue_relation'_def) done +lemma rf_sr_tcb_update_twice: + "h_t_valid (hrs_htd hrs) c_guard ptr + \ ((s, s'\globals := globals s''\t_hrs_' := + hrs_mem_update (heap_update (ptr :: tcb_C ptr) v) + (hrs_mem_update (heap_update ptr v') hrs)\\) \ rf_sr) + = ((s, s'\globals := globals s''\t_hrs_' := + hrs_mem_update (heap_update (ptr :: tcb_C ptr) v) hrs\\) \ rf_sr)" + by (simp add: rf_sr_def cstate_relation_def Let_def + cpspace_relation_def typ_heap_simps' + carch_state_relation_def + cmachine_state_relation_def) + lemma sendIPC_block_ccorres_helper: "ccorres dc xfdc (tcb_at' thread and valid_queues and valid_objs' and sch_act_not thread and ep_at' epptr and @@ -4377,9 +4320,10 @@ lemma sendIPC_block_ccorres_helper: apply clarsimp apply (frule(1) tcb_at_h_t_valid) apply (frule h_t_valid_c_guard) - apply (clarsimp simp: typ_heap_simps) - apply (erule(1) rf_sr_tcb_update_no_queue) - apply (simp add: typ_heap_simps')+ + apply (clarsimp simp: typ_heap_simps' + rf_sr_tcb_update_twice) + apply (erule(1) rf_sr_tcb_update_no_queue_gen, + (simp add: typ_heap_simps')+)[1] apply (simp add: tcb_cte_cases_def) apply (simp add: ctcb_relation_def cthread_state_relation_def ThreadState_BlockedOnSend_def mask_def @@ -4515,7 +4459,9 @@ lemma tcbEPAppend_spec: ((ctcb_ptr_to_tcb_ptr \<^bsup>s\<^esup>tcb) # queue))) \ option_map tcb_null_ep_ptrs \ (cslift t) = option_map tcb_null_ep_ptrs \ (cslift s) - \ cslift_all_but_tcb_C t s \ (hrs_htd \<^bsup>t\<^esup>t_hrs) = (hrs_htd \<^bsup>s\<^esup>t_hrs)}" + \ cslift_all_but_tcb_C t s \ (hrs_htd \<^bsup>t\<^esup>t_hrs) = (hrs_htd \<^bsup>s\<^esup>t_hrs) + \ (\rs. zero_ranges_are_zero rs (\<^bsup>s\<^esup>t_hrs) + \ zero_ranges_are_zero rs (\<^bsup>t\<^esup>t_hrs))}" apply (intro allI) apply (rule conseqPre, vcg) apply (clarsimp split del: split_if) @@ -4536,13 +4482,13 @@ lemma tcbEPAppend_spec: apply assumption+ apply (drule tcb_at_not_NULL, simp) apply (unfold upd_unless_null_def) - apply (simp add: typ_heap_simps) - apply (intro impI conjI allI) - apply simp - apply simp - apply (rule ext) - apply (clarsimp simp add: typ_heap_simps tcb_null_ep_ptrs_def - split: split_if) + apply (clarsimp split: split_if_asm) + apply (simp add: typ_heap_simps') + apply (rule ext) + apply (clarsimp simp add: typ_heap_simps tcb_null_ep_ptrs_def + split: split_if) + apply (simp add: typ_heap_simps') + apply (intro conjI) apply (clarsimp simp add: typ_heap_simps h_t_valid_clift_Some_iff) apply (erule iffD1 [OF tcb_queue_relation'_cong, OF refl refl refl, rotated -1]) apply (clarsimp split: split_if) @@ -4634,7 +4580,8 @@ lemma sendIPC_enqueue_ccorres_helper: apply (simp add: setEndpoint_def split_def) apply (rule bexI [OF _ setObject_eq]) apply (simp add: rf_sr_def cstate_relation_def init_def Let_def - cpspace_relation_def update_ep_map_tos) + cpspace_relation_def update_ep_map_tos + typ_heap_simps') apply (elim conjE) apply (intro conjI) -- "tcb relation" @@ -4660,9 +4607,9 @@ lemma sendIPC_enqueue_ccorres_helper: -- "queue relation" apply (rule cready_queues_relation_null_queue_ptrs, assumption+) apply (clarsimp simp: comp_def) - apply (simp add: carch_state_relation_def) + apply (simp add: carch_state_relation_def typ_heap_simps') apply (simp add: cmachine_state_relation_def) - apply (simp add: h_t_valid_clift_Some_iff) + apply (simp add: typ_heap_simps') apply (simp add: objBits_simps) apply (simp add: objBits_simps) apply assumption @@ -4672,7 +4619,8 @@ lemma sendIPC_enqueue_ccorres_helper: apply (simp add: setEndpoint_def split_def) apply (rule bexI [OF _ setObject_eq]) apply (simp add: rf_sr_def cstate_relation_def Let_def init_def - cpspace_relation_def update_ep_map_tos) + cpspace_relation_def update_ep_map_tos + typ_heap_simps') apply (elim conjE) apply (intro conjI) -- "tcb relation" @@ -4700,7 +4648,7 @@ lemma sendIPC_enqueue_ccorres_helper: -- "queue relation" apply (rule cready_queues_relation_null_queue_ptrs, assumption+) apply (clarsimp simp: comp_def) - apply (simp add: carch_state_relation_def) + apply (simp add: carch_state_relation_def typ_heap_simps') apply (simp add: cmachine_state_relation_def) apply (simp add: h_t_valid_clift_Some_iff) apply (simp add: objBits_simps) @@ -4998,9 +4946,8 @@ lemma receiveIPC_block_ccorres_helper: apply clarsimp apply (frule(1) tcb_at_h_t_valid) apply (frule h_t_valid_c_guard) - apply (clarsimp simp: typ_heap_simps) - apply (erule(1) rf_sr_tcb_update_no_queue) - apply (simp add: typ_heap_simps)+ + apply (clarsimp simp: typ_heap_simps' rf_sr_tcb_update_twice) + apply (erule(1) rf_sr_tcb_update_no_queue_gen, (simp add: typ_heap_simps)+) apply (simp add: tcb_cte_cases_def) apply (simp add: ctcb_relation_def cthread_state_relation_def ThreadState_BlockedOnReceive_def mask_def @@ -5069,7 +5016,8 @@ lemma receiveIPC_enqueue_ccorres_helper: apply (simp add: setEndpoint_def split_def) apply (rule bexI [OF _ setObject_eq]) apply (simp add: rf_sr_def cstate_relation_def Let_def init_def - cpspace_relation_def update_ep_map_tos) + cpspace_relation_def update_ep_map_tos + typ_heap_simps') apply (elim conjE) apply (intro conjI) -- "tcb relation" @@ -5097,7 +5045,7 @@ lemma receiveIPC_enqueue_ccorres_helper: -- "queue relation" apply (rule cready_queues_relation_null_queue_ptrs, assumption+) apply (clarsimp simp: comp_def) - apply (simp add: carch_state_relation_def) + apply (simp add: carch_state_relation_def typ_heap_simps') apply (simp add: cmachine_state_relation_def) apply (simp add: h_t_valid_clift_Some_iff) apply (simp add: objBits_simps) @@ -5109,7 +5057,8 @@ lemma receiveIPC_enqueue_ccorres_helper: apply (simp add: setEndpoint_def split_def) apply (rule bexI [OF _ setObject_eq]) apply (simp add: rf_sr_def cstate_relation_def init_def Let_def - cpspace_relation_def update_ep_map_tos) + cpspace_relation_def update_ep_map_tos + typ_heap_simps') apply (elim conjE) apply (intro conjI) -- "tcb relation" @@ -5134,9 +5083,9 @@ lemma receiveIPC_enqueue_ccorres_helper: -- "queue relation" apply (rule cready_queues_relation_null_queue_ptrs, assumption+) apply (clarsimp simp: comp_def) - apply (simp add: carch_state_relation_def) + apply (simp add: carch_state_relation_def typ_heap_simps') apply (simp add: cmachine_state_relation_def) - apply (simp add: h_t_valid_clift_Some_iff) + apply (simp add: typ_heap_simps') apply (simp add: objBits_simps) apply (simp add: objBits_simps) apply assumption @@ -5198,7 +5147,8 @@ lemma receiveIPC_dequeue_ccorres_helper: apply (rule conjI) apply (rule bexI [OF _ setObject_eq]) apply (simp add: rf_sr_def cstate_relation_def Let_def - cpspace_relation_def update_ep_map_tos) + cpspace_relation_def update_ep_map_tos + typ_heap_simps') apply (elim conjE) apply (intro conjI) -- "tcb relation" @@ -5218,9 +5168,9 @@ lemma receiveIPC_dequeue_ccorres_helper: -- "queue relation" apply (rule cready_queues_relation_null_queue_ptrs, assumption+) apply (clarsimp simp: comp_def) - apply (simp add: carch_state_relation_def) + apply (simp add: carch_state_relation_def typ_heap_simps') apply (simp add: cmachine_state_relation_def) - apply (simp add: h_t_valid_clift_Some_iff) + apply (simp add: typ_heap_simps') apply (simp add: objBits_simps) apply (simp add: objBits_simps) apply assumption @@ -5237,7 +5187,8 @@ lemma receiveIPC_dequeue_ccorres_helper: apply (rule bexI [OF _ setObject_eq]) apply (frule(1) st_tcb_at_h_t_valid) apply (simp add: rf_sr_def cstate_relation_def Let_def - cpspace_relation_def update_ep_map_tos) + cpspace_relation_def update_ep_map_tos + typ_heap_simps') apply (elim conjE) apply (intro conjI) -- "tcb relation" @@ -5265,9 +5216,9 @@ lemma receiveIPC_dequeue_ccorres_helper: -- "queue relation" apply (rule cready_queues_relation_null_queue_ptrs, assumption+) apply (clarsimp simp: comp_def) - apply (simp add: carch_state_relation_def) + apply (simp add: carch_state_relation_def typ_heap_simps') apply (simp add: cmachine_state_relation_def) - apply (simp add: h_t_valid_clift_Some_iff) + apply (simp add: typ_heap_simps') apply (simp add: objBits_simps) apply (simp add: objBits_simps) apply assumption @@ -5335,13 +5286,13 @@ lemma completeSignal_ccorres: apply (clarsimp simp: typ_heap_simps setNotification_def) apply (rule bexI [OF _ setObject_eq]) apply (simp add: rf_sr_def cstate_relation_def Let_def update_ntfn_map_tos - cpspace_relation_def) + cpspace_relation_def typ_heap_simps') apply (elim conjE) apply (intro conjI) apply (rule cpspace_relation_ntfn_update_ntfn, assumption+) apply (simp add: cnotification_relation_def Let_def NtfnState_Idle_def mask_def) apply simp - apply (simp add: carch_state_relation_def) + apply (simp add: carch_state_relation_def typ_heap_simps') apply (simp add: cmachine_state_relation_def) apply (simp add: h_t_valid_clift_Some_iff) apply (simp add: objBits_simps) @@ -5772,7 +5723,8 @@ lemma sendSignal_dequeue_ccorres_helper: apply (rule conjI) apply (rule bexI [OF _ setObject_eq]) apply (simp add: rf_sr_def cstate_relation_def Let_def - cpspace_relation_def update_ntfn_map_tos) + cpspace_relation_def update_ntfn_map_tos + typ_heap_simps') apply (elim conjE) apply (intro conjI) -- "tcb relation" @@ -5792,7 +5744,7 @@ lemma sendSignal_dequeue_ccorres_helper: -- "queue relation" apply (rule cready_queues_relation_null_queue_ptrs, assumption+) apply (clarsimp simp: comp_def) - apply (simp add: carch_state_relation_def) + apply (simp add: carch_state_relation_def typ_heap_simps') apply (simp add: cmachine_state_relation_def) apply (simp add: h_t_valid_clift_Some_iff) apply (simp add: objBits_simps) @@ -5813,7 +5765,8 @@ lemma sendSignal_dequeue_ccorres_helper: apply (rule bexI [OF _ setObject_eq]) apply (frule(1) st_tcb_at_h_t_valid) apply (simp add: rf_sr_def cstate_relation_def Let_def - cpspace_relation_def update_ntfn_map_tos) + cpspace_relation_def update_ntfn_map_tos + typ_heap_simps') apply (elim conjE) apply (intro conjI) -- "tcb relation" @@ -5841,7 +5794,7 @@ lemma sendSignal_dequeue_ccorres_helper: -- "queue relation" apply (rule cready_queues_relation_null_queue_ptrs, assumption+) apply (clarsimp simp: comp_def) - apply (simp add: carch_state_relation_def) + apply (simp add: carch_state_relation_def typ_heap_simps') apply (simp add: cmachine_state_relation_def) apply (simp add: h_t_valid_clift_Some_iff) apply (simp add: objBits_simps) @@ -5868,7 +5821,7 @@ lemma ntfn_set_active_ccorres: apply (rule cmap_relation_ko_atE [OF cmap_relation_ntfn], assumption+) apply (clarsimp simp: typ_heap_simps) apply (rule bexI[OF _ setObject_eq], simp_all) - apply (clarsimp simp: typ_heap_simps rf_sr_def cstate_relation_def Let_def + apply (clarsimp simp: typ_heap_simps' rf_sr_def cstate_relation_def Let_def cpspace_relation_def update_ntfn_map_tos carch_state_relation_def cmachine_state_relation_def) apply (rule cpspace_relation_ntfn_update_ntfn, assumption+) @@ -5977,14 +5930,15 @@ lemma sendSignal_ccorres [corres]: apply (clarsimp simp: typ_heap_simps) apply (rule bexI [OF _ setObject_eq]) apply (simp add: rf_sr_def cstate_relation_def Let_def - cpspace_relation_def update_ntfn_map_tos) + cpspace_relation_def update_ntfn_map_tos + typ_heap_simps') apply (elim conjE) apply (intro conjI) apply (rule cpspace_relation_ntfn_update_ntfn, assumption+) apply (simp add: cnotification_relation_def Let_def NtfnState_Active_def mask_def word_bw_comms) apply simp - apply (simp add: carch_state_relation_def) + apply (simp add: carch_state_relation_def typ_heap_simps') apply (simp add: cmachine_state_relation_def) apply (simp add: h_t_valid_clift_Some_iff) apply (simp add: objBits_simps) @@ -6067,9 +6021,9 @@ lemma receiveSignal_block_ccorres_helper: apply clarsimp apply (frule(1) tcb_at_h_t_valid) apply (frule h_t_valid_c_guard) - apply (clarsimp simp: typ_heap_simps) - apply (erule(1) rf_sr_tcb_update_no_queue) - apply (simp add: typ_heap_simps)+ + apply (clarsimp simp: typ_heap_simps' rf_sr_tcb_update_twice) + apply (erule(1) rf_sr_tcb_update_no_queue_gen, + (simp add: typ_heap_simps')+) apply (simp add: tcb_cte_cases_def) apply (simp add: ctcb_relation_def cthread_state_relation_def ThreadState_BlockedOnNotification_def mask_def @@ -6191,7 +6145,8 @@ lemma receiveSignal_enqueue_ccorres_helper: apply (simp add: setNotification_def split_def) apply (rule bexI [OF _ setObject_eq]) apply (simp add: rf_sr_def cstate_relation_def Let_def init_def - cpspace_relation_def update_ntfn_map_tos) + cpspace_relation_def update_ntfn_map_tos + typ_heap_simps') apply (elim conjE) apply (intro conjI) -- "tcb relation" @@ -6219,7 +6174,7 @@ lemma receiveSignal_enqueue_ccorres_helper: -- "queue relation" apply (rule cready_queues_relation_null_queue_ptrs, assumption+) subgoal by (clarsimp simp: comp_def) - apply (simp add: carch_state_relation_def) + apply (simp add: carch_state_relation_def typ_heap_simps') apply (simp add: cmachine_state_relation_def) apply (simp add: h_t_valid_clift_Some_iff) apply (simp add: objBits_simps) @@ -6231,7 +6186,8 @@ lemma receiveSignal_enqueue_ccorres_helper: apply (simp add: setNotification_def split_def) apply (rule bexI [OF _ setObject_eq]) apply (simp add: rf_sr_def cstate_relation_def init_def Let_def - cpspace_relation_def update_ntfn_map_tos) + cpspace_relation_def update_ntfn_map_tos + typ_heap_simps') apply (elim conjE) apply (intro conjI) -- "tcb relation" @@ -6258,7 +6214,7 @@ lemma receiveSignal_enqueue_ccorres_helper: -- "queue relation" apply (rule cready_queues_relation_null_queue_ptrs, assumption+) apply (clarsimp simp: comp_def) - apply (simp add: carch_state_relation_def) + apply (simp add: carch_state_relation_def typ_heap_simps') apply (simp add: cmachine_state_relation_def) apply (simp add: h_t_valid_clift_Some_iff) apply (simp add: objBits_simps) @@ -6352,14 +6308,15 @@ lemma receiveSignal_ccorres [corres]: apply (clarsimp simp: typ_heap_simps setNotification_def) apply (rule bexI [OF _ setObject_eq]) apply (simp add: rf_sr_def cstate_relation_def Let_def - cpspace_relation_def update_ntfn_map_tos) + cpspace_relation_def update_ntfn_map_tos + typ_heap_simps') apply (elim conjE) apply (intro conjI) apply (rule cpspace_relation_ntfn_update_ntfn, assumption+) apply (simp add: cnotification_relation_def Let_def NtfnState_Idle_def mask_def) apply simp - apply (simp add: carch_state_relation_def) + apply (simp add: carch_state_relation_def typ_heap_simps') apply (simp add: cmachine_state_relation_def) apply (simp add: h_t_valid_clift_Some_iff) apply (simp add: objBits_simps) diff --git a/proof/crefine/Recycle_C.thy b/proof/crefine/Recycle_C.thy index 0efd5a7c4..cd611ea68 100644 --- a/proof/crefine/Recycle_C.thy +++ b/proof/crefine/Recycle_C.thy @@ -194,6 +194,31 @@ lemma flex_user_data_at_rf_sr_dom_s: apply simp done +lemma hrs_mem_update_fold_eq: + "hrs_mem_update (fold f xs) + = fold (hrs_mem_update o f) xs" + apply (rule sym, induct xs) + apply (simp add: hrs_mem_update_def) + apply (simp add: hrs_mem_update_def fun_eq_iff) + done + +lemma power_user_page_foldl_zero_ranges: + " \p<2 ^ (pageBitsForSize sz - pageBits). + hrs_htd hrs \\<^sub>t (Ptr (ptr + of_nat p * 0x1000) :: user_data_C ptr) + \ zero_ranges_are_zero rngs hrs + \ zero_ranges_are_zero rngs + (hrs_mem_update (\s. foldl (\s x. heap_update (Ptr x) (user_data_C (arr x)) s) s + (map (\n. ptr + of_nat n * 0x1000) [0..<2 ^ (pageBitsForSize sz - pageBits)])) + hrs)" + apply (simp add: foldl_conv_fold hrs_mem_update_fold_eq) + apply (rule conjunct1) + apply (rule fold_invariant[where P="\hrs'. zero_ranges_are_zero rngs hrs' + \ hrs_htd hrs' = hrs_htd hrs" + and xs=xs and Q="\x. x \ set xs" for xs], simp_all) + apply (subst zero_ranges_are_zero_update, simp_all) + apply clarsimp + done + lemma heap_to_device_data_disj_mdf': "\is_aligned ptr (pageBitsForSize sz); ksPSpace \ a = Some obj; objBitsKO obj = pageBits; pspace_aligned' \; pspace_distinct' \; pspace_no_overlap' ptr (pageBitsForSize sz) \\ @@ -291,6 +316,7 @@ lemma clearMemory_PageCap_ccorres: carch_state_relation_def cmachine_state_relation_def foldl_fun_upd_const[unfolded fun_upd_def] + power_user_page_foldl_zero_ranges dom_heap_to_device_data) apply (rule conjI[rotated]) apply (simp add:pageBitsForSize_mess_multi) @@ -1409,7 +1435,8 @@ lemma cancelBadgedSends_ccorres: cmachine_state_relation_def ) apply (clarsimp simp: cpspace_relation_def - update_ep_map_tos) + update_ep_map_tos + typ_heap_simps') apply (erule(1) cpspace_relation_ep_update_ep2) apply (simp add: cendpoint_relation_def endpoint_state_defs) subgoal by simp @@ -1447,7 +1474,7 @@ lemma cancelBadgedSends_ccorres: apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def carch_state_relation_def cmachine_state_relation_def - cpspace_relation_def + cpspace_relation_def typ_heap_simps' update_ep_map_tos) apply (erule(1) cpspace_relation_ep_update_ep2) subgoal by (simp add: cendpoint_relation_def Let_def) @@ -1459,7 +1486,7 @@ lemma cancelBadgedSends_ccorres: apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def carch_state_relation_def cmachine_state_relation_def - cpspace_relation_def + cpspace_relation_def typ_heap_simps' update_ep_map_tos) apply (erule(1) cpspace_relation_ep_update_ep2) apply (simp add: cendpoint_relation_def Let_def) @@ -1936,6 +1963,91 @@ lemma cteRecycle_ccorres: apply (clarsimp simp: from_bool_def true_def exception_defs cintr_def) done +lemma cte_lift_ccte_relation: + "cte_lift cte' = Some ctel' + \ c_valid_cte cte' + \ ccte_relation (cte_to_H ctel') cte'" + by (simp add: ccte_relation_def) + +lemma updateFreeIndex_ccorres: + "\s. \ \ ({s} \ {s. \cte cte'. cslift s (cte_Ptr srcSlot) = Some cte' + \ cteCap cte = cap' \ ccte_relation cte cte'}) + c + {t. \cap. cap_untyped_cap_lift cap = (cap_untyped_cap_lift + (cte_C.cap_C (the (cslift s (cte_Ptr srcSlot))))) + \ cap_untyped_cap_CL.capFreeIndex_CL := ((of_nat idx') >> 4) \ + \ cap_get_tag cap = scast cap_untyped_cap + \ t_hrs_' (globals t) = hrs_mem_update (heap_update (cte_Ptr srcSlot) + (cte_C.cap_C_update (\_. cap) (the (cslift s (cte_Ptr srcSlot))))) + (t_hrs_' (globals s)) + \ t may_only_modify_globals s in [t_hrs] + } + \ ccorres dc xfdc + (valid_objs' and cte_wp_at' (\cte. isUntypedCap (cteCap cte) + \ cap' = (cteCap cte)) srcSlot + and untyped_ranges_zero' + and (\_. is_aligned (of_nat idx' :: word32) 4 \ idx' \ 2 ^ (capBlockSize cap'))) + {s. \ capIsDevice cap' + \ region_is_zero_bytes (capPtr cap' + of_nat idx') (capFreeIndex cap' - idx') s} hs + (updateFreeIndex srcSlot idx') c" + (is "_ \ ccorres dc xfdc (valid_objs' and ?cte_wp_at' and _ and _) ?P' hs ?a c") + apply (rule ccorres_gen_asm) + apply (simp add: updateFreeIndex_def getSlotCap_def updateCap_def) + apply (rule ccorres_guard_imp2) + apply (rule ccorres_split_noop_lhs, rule_tac cap'=cap' in updateTrackedFreeIndex_noop_ccorres) + apply (rule ccorres_pre_getCTE)+ + apply (rename_tac cte cte2) + apply (rule_tac P = "\s. ?cte_wp_at' s \ cte2 = cte \ cte_wp_at' (op = cte) srcSlot s" + and P'="{s. \cte cte'. cslift s (cte_Ptr srcSlot) = Some cte' + \ cteCap cte = cap' \ ccte_relation cte cte'} \ ?P'" in ccorres_from_vcg) + apply (rule allI, rule HoarePartial.conseq_exploit_pre, clarify) + apply (drule_tac x=s in spec, rule conseqPre, erule conseqPost) + defer + apply clarsimp + apply clarsimp + apply (simp add: cte_wp_at_ctes_of) + apply wp + apply (clarsimp simp: isCap_simps cte_wp_at_ctes_of) + apply (frule(1) rf_sr_ctes_of_clift) + apply clarsimp + apply (frule(1) cte_lift_ccte_relation) + apply (rule exI, intro conjI[rotated], assumption, simp_all)[1] + apply (clarsimp simp: cte_wp_at_ctes_of) + apply (erule(1) rf_sr_ctes_of_cliftE) + apply (frule(1) rf_sr_ctes_of_clift) + apply clarsimp + apply (subgoal_tac "ccap_relation (capFreeIndex_update (\_. idx') + (cteCap (the (ctes_of \ srcSlot)))) cap") + apply (rule fst_setCTE [OF ctes_of_cte_at], assumption) + apply (erule bexI [rotated]) + apply (clarsimp simp add: rf_sr_def cstate_relation_def Let_def + cvariable_array_map_const_add_map_option[where f="tcb_no_ctes_proj"] + isCap_simps) + apply (simp add:cpspace_relation_def) + apply (clarsimp simp:typ_heap_simps' modify_map_def mex_def meq_def) + apply (rule conjI) + apply (rule cpspace_cte_relation_upd_capI, assumption+) + apply (rule conjI) + apply (rule setCTE_tcb_case, assumption+) + apply (case_tac s', clarsimp) + subgoal by (simp add: carch_state_relation_def cmachine_state_relation_def + typ_heap_simps') + + apply (clarsimp simp: isCap_simps) + apply (drule(1) cte_lift_ccte_relation, + drule ccte_relation_ccap_relation) + apply (simp add: cte_to_H_def) + apply (frule cap_get_tag_isCap_unfolded_H_cap) + apply (clarsimp simp: ccap_relation_def cap_lift_untyped_cap + cap_to_H_simps cap_untyped_cap_lift_def + is_aligned_shiftr_shiftl + dest!: ccte_relation_ccap_relation) + apply (rule unat_of_nat_eq unat_of_nat_eq[symmetric], + erule order_le_less_trans, + rule power_strict_increasing, simp_all) + apply (rule unat_less_helper, rule order_le_less_trans[OF word_and_le1], simp) + done + end (* FIXME: Move *) diff --git a/proof/crefine/Retype_C.thy b/proof/crefine/Retype_C.thy index ca1f1195d..b748947e6 100644 --- a/proof/crefine/Retype_C.thy +++ b/proof/crefine/Retype_C.thy @@ -561,21 +561,6 @@ lemma ptr_retyps_gen_out: apply (clarsimp simp: ptr_arr_retyps_def htd_update_list_same2) done -definition - region_is_bytes' :: "word32 \ nat \ heap_typ_desc \ bool" -where - "region_is_bytes' ptr sz htd \ \z\{ptr ..+ sz}. \ td. td \ typ_uinfo_t TYPE (word8) \ - (\n b. snd (htd z) n \ Some (td, b))" - -abbreviation - region_is_bytes :: "word32 \ nat \ globals myvars \ bool" -where - "region_is_bytes ptr sz s \ region_is_bytes' ptr sz (hrs_htd (t_hrs_' (globals s)))" - -lemma map_leD: - "\ map_le m m'; m x = Some y \ \ m' x = Some y" - by (simp add: map_le_def dom_def) - lemma h_t_valid_intvl_htd_contains_uinfo_t: "h_t_valid d g (p :: ('a :: c_type) ptr) \ x \ {ptr_val p ..+ size_of TYPE('a)} \ (\n. snd (d x) n \ None \ fst (the (snd (d x) n)) = typ_uinfo_t TYPE ('a))" @@ -950,19 +935,6 @@ next done qed -lemma region_is_bytes_disjoint: - assumes cleared: "region_is_bytes' p (n * size_of TYPE('a :: c_type)) (hrs_htd hrs)" - and not_byte: "typ_uinfo_t TYPE('b :: mem_type) \ typ_uinfo_t TYPE(word8)" - shows "hrs_htd hrs \\<^sub>t p' \ {p..+n * size_of TYPE('a)} \ ptr_span (p' :: 'b ptr) = {}" - apply (clarsimp simp: h_t_valid_def valid_footprint_def Let_def) - apply (clarsimp simp: set_eq_iff dest!: intvlD[where p="ptr_val p'"]) - apply (drule_tac x="of_nat k" in spec, clarsimp simp: size_of_def) - apply (cut_tac m=k in typ_slice_t_self[where td="typ_uinfo_t TYPE('b)"]) - apply (clarsimp simp: in_set_conv_nth) - apply (drule_tac x=i in map_leD, simp) - apply (simp add: cleared[unfolded region_is_bytes'_def] not_byte size_of_def) - done - lemma clift_ptr_retyps_gen_memset_same: assumes guard: "\n' < n. c_guard (CTypesDefs.ptr_add (Ptr p :: 'a :: mem_type ptr) (of_nat n'))" assumes cleared: "region_is_bytes' p (n * size_of TYPE('a :: mem_type)) (hrs_htd hrs)" @@ -1518,12 +1490,64 @@ lemma clift_eq_h_t_valid_eq: = h_t_valid (hrs_htd hp') c_guard" by (rule ext, simp add: h_t_valid_clift_Some_iff) -abbreviation - "heap_list_is_zero hp ptr n \ heap_list hp n ptr = replicate n 0" +lemma region_is_bytes_typ_region_bytes: + "{ptr ..+ len} \ {ptr' ..+ 2 ^ bits} + \ region_is_bytes' ptr len (typ_region_bytes ptr' bits htd)" + apply (clarsimp simp: region_is_bytes'_def typ_region_bytes_def hrs_htd_update) + apply (simp add: subsetD split: split_if_asm) + done + +lemma region_is_bytes_retyp_disjoint: + "{ptr ..+ sz} \ {ptr_val (p :: 'a ptr)..+n * size_of TYPE('a :: mem_type)} = {} + \ region_is_bytes' ptr sz htd + \ region_is_bytes' ptr sz (ptr_retyps_gen n p arr htd)" + apply (clarsimp simp: region_is_bytes'_def del: impI) + apply (subst ptr_retyps_gen_out) + apply blast + apply simp + done + +lemma intvl_plus_unat_eq: + "p \ p + x - 1 \ x \ 0 + \ {p ..+ unat x} = {p .. p + x - 1}" + apply (subst upto_intvl_eq', simp_all add: unat_eq_0 field_simps) + apply (rule order_less_imp_le, simp) + done + +lemma zero_ranges_ptr_retyps: + "zero_ranges_are_zero (gsUntypedZeroRanges s) hrs + \ caps_overlap_reserved' {ptr_val (p :: 'a ptr) ..+ n * size_of TYPE ('a :: mem_type)} s + \ untyped_ranges_zero' s + \ valid_objs' s + \ zero_ranges_are_zero (gsUntypedZeroRanges s) + (hrs_htd_update (ptr_retyps_gen n p arr) hrs)" + apply (clarsimp simp: zero_ranges_are_zero_def untyped_ranges_zero_inv_def + hrs_htd_update) + apply (drule(1) bspec, clarsimp) + apply (rule region_is_bytes_retyp_disjoint, simp_all) + apply (clarsimp simp: map_comp_Some_iff cteCaps_of_def + elim!: ranE) + apply (frule(1) ctes_of_valid') + apply (simp add: caps_overlap_reserved'_def, + drule bspec, erule ranI) + apply (frule(1) untypedZeroRange_to_usableCapRange) + apply (clarsimp simp: isCap_simps untypedZeroRange_def + getFreeRef_def max_free_index_def + split: split_if_asm) + apply (erule disjoint_subset[rotated]) + apply (subst intvl_plus_unat_eq) + apply clarsimp + apply clarsimp + apply (clarsimp simp: word_unat.Rep_inject[symmetric] + valid_cap_simps' capAligned_def + unat_of_nat + simp del: word_unat.Rep_inject) + apply clarsimp + done abbreviation - "region_is_zero_bytes ptr n x \ region_is_bytes ptr n x - \ heap_list_is_zero (hrs_mem (t_hrs_' (globals x))) ptr n" + "ret_zero ptr sz + \ valid_objs' and untyped_ranges_zero' and caps_overlap_reserved' {ptr ..+ sz}" lemma createObjects_ccorres_ep: defines "ko \ (KOEndpoint (makeObject :: endpoint))" @@ -1531,6 +1555,7 @@ lemma createObjects_ccorres_ep: \ ptr \ 0 \ pspace_aligned' \ \ pspace_distinct' \ \ pspace_no_overlap' ptr sz \ + \ ret_zero ptr (n * (2 ^ objBitsKO ko)) \ \ region_is_zero_bytes ptr (n * (2 ^ objBitsKO ko)) x \ range_cover ptr sz (objBitsKO ko) n \ {ptr ..+ n * (2 ^ objBitsKO ko)} \ kernel_data_refs = {} @@ -1555,7 +1580,8 @@ proof (intro impI allI) and sz: "objBitsKO ko \ sz" and szb: "sz < word_bits" and pal: "pspace_aligned' \" and pdst: "pspace_distinct' \" - and pno: "pspace_no_overlap' ptr sz \" + and pno: "pspace_no_overlap' ptr sz \" + and rzo: "ret_zero ptr (n * (2 ^ objBitsKO ko)) \" and empty: "region_is_bytes ptr (n * (2 ^ objBitsKO ko)) x" and zero: "heap_list_is_zero (hrs_mem (t_hrs_' (globals x))) ptr (n * (2 ^ objBitsKO ko))" and rc: "range_cover ptr sz (objBitsKO ko) n" @@ -1622,7 +1648,7 @@ proof (intro impI allI) apply (rule relrl[simplified szo ko_def]) done - thus ?thesis using rf empty kdr + thus ?thesis using rf empty kdr rzo apply (simp add: rf_sr_def cstate_relation_def Let_def rl' tag_disj_via_td_name) apply (simp add: carch_state_relation_def cmachine_state_relation_def) @@ -1631,6 +1657,7 @@ proof (intro impI allI) kernel_data_refs_domain_eq_rotate ht_rl foldr_upd_app_if [folded data_map_insert_def] rl projectKOs cvariable_array_ptr_retyps[OF szo] + zero_ranges_ptr_retyps simp del: endpoint_C_size) done qed @@ -1640,6 +1667,7 @@ lemma createObjects_ccorres_ntfn: shows "\\ x. (\, x) \ rf_sr \ ptr \ 0 \ pspace_aligned' \ \ pspace_distinct' \ \ pspace_no_overlap' ptr sz \ + \ ret_zero ptr (n * (2 ^ objBitsKO ko)) \ \ region_is_zero_bytes ptr (n * 2 ^ objBitsKO ko) x \ range_cover ptr sz (objBitsKO ko) n \ {ptr ..+ n * (2 ^ objBitsKO ko)} \ kernel_data_refs = {} @@ -1666,6 +1694,7 @@ proof (intro impI allI) and szb: "sz < word_bits" and pal: "pspace_aligned' \" and pdst: "pspace_distinct' \" and pno: "pspace_no_overlap' ptr sz \" + and rzo: "ret_zero ptr (n * (2 ^ objBitsKO ko)) \" and empty: "region_is_bytes ptr (n * (2 ^ objBitsKO ko)) x" and zero: "heap_list_is_zero (hrs_mem (t_hrs_' (globals x))) ptr (n * (2 ^ objBitsKO ko))" and rc: "range_cover ptr sz (objBitsKO ko) n" @@ -1733,7 +1762,7 @@ proof (intro impI allI) apply (rule relrl[simplified szo ko_def]) done - thus ?thesis using rf empty kdr + thus ?thesis using rf empty kdr rzo apply (simp add: rf_sr_def cstate_relation_def Let_def rl' tag_disj_via_td_name) apply (simp add: carch_state_relation_def cmachine_state_relation_def) apply (simp add: rl' cterl tag_disj_via_td_name h_t_valid_clift_Some_iff ) @@ -1741,8 +1770,8 @@ proof (intro impI allI) kernel_data_refs_domain_eq_rotate ht_rl foldr_upd_app_if [folded data_map_insert_def] rl projectKOs cvariable_array_ptr_retyps[OF szo] + zero_ranges_ptr_retyps simp del: notification_C_size) - done qed @@ -1790,6 +1819,7 @@ lemma createObjects_ccorres_cte: shows "\\ x. (\, x) \ rf_sr \ ptr \ 0 \ pspace_aligned' \ \ pspace_distinct' \ \ pspace_no_overlap' ptr sz \ + \ ret_zero ptr (n * 2 ^ objBitsKO ko) \ \ region_is_zero_bytes ptr (n * 2 ^ objBitsKO ko) x \ range_cover ptr sz (objBitsKO ko) n \ {ptr ..+ n * (2 ^ objBitsKO ko)} \ kernel_data_refs = {} @@ -1814,7 +1844,8 @@ proof (intro impI allI) and sz: "objBitsKO ko \ sz" and szb: "sz < word_bits" and pal: "pspace_aligned' \" and pdst: "pspace_distinct' \" - and pno: "pspace_no_overlap' ptr sz \" + and pno: "pspace_no_overlap' ptr sz \" + and rzo: "ret_zero ptr (n * 2 ^ objBitsKO ko) \" and empty: "region_is_bytes ptr (n * (2 ^ objBitsKO ko)) x" and zero: "heap_list_is_zero (hrs_mem (t_hrs_' (globals x))) ptr (n * (2 ^ objBitsKO ko))" and rc: "range_cover ptr sz (objBitsKO ko) n" @@ -1878,13 +1909,14 @@ proof (intro impI allI) apply (rule relrl[simplified szo ko_def]) done - thus ?thesis using rf empty kdr irq + thus ?thesis using rf empty kdr irq rzo apply (simp add: rf_sr_def cstate_relation_def Let_def rl' tag_disj_via_td_name) apply (simp add: carch_state_relation_def cmachine_state_relation_def) apply (simp add: rl' cterl tag_disj_via_td_name h_t_valid_clift_Some_iff) apply (clarsimp simp: hrs_htd_update ptr_retyps_htd_safe_neg szo kernel_data_refs_domain_eq_rotate rl foldr_upd_app_if [folded data_map_insert_def] projectKOs + zero_ranges_ptr_retyps ht_rl cvariable_array_ptr_retyps[OF szo]) done qed @@ -2016,6 +2048,7 @@ lemma createObjects_ccorres_pte: shows "\\ x. (\, x) \ rf_sr \ ptr \ 0 \ pspace_aligned' \ \ pspace_distinct' \ \ pspace_no_overlap' ptr sz \ + \ ret_zero ptr (2 ^ ptBits) \ \ region_is_zero_bytes ptr (2 ^ ptBits) x \ range_cover ptr sz ptBits 1 \ valid_global_refs' s @@ -2042,6 +2075,7 @@ proof (intro impI allI) and pal: "pspace_aligned' \" and pdst: "pspace_distinct' \" and pno: "pspace_no_overlap' ptr sz \" + and rzo: "ret_zero ptr (2 ^ ptBits) \" and empty: "region_is_bytes ptr (2 ^ ptBits) x" and zero: "heap_list_is_zero (hrs_mem (t_hrs_' (globals x))) ptr (2 ^ ptBits)" and kernel_data_refs_disj : "kernel_data_refs \ {ptr..+ 2 ^ ptBits} = {}" @@ -2165,12 +2199,13 @@ proof (intro impI allI) done ultimately - show ?thesis using rf empty kernel_data_refs_disj + show ?thesis using rf empty kernel_data_refs_disj rzo apply (simp add: rf_sr_def cstate_relation_def Let_def rl' tag_disj_via_td_name) apply (simp add: carch_state_relation_def cmachine_state_relation_def) apply (clarsimp simp add: rl' cterl tag_disj_via_td_name hrs_htd_update ht_rl foldr_upd_app_if [folded data_map_insert_def] rl projectKOs - cvariable_array_ptr_retyps[OF szo]) + cvariable_array_ptr_retyps[OF szo] + zero_ranges_ptr_retyps[where p="pt_Ptr ptr", simplified szo]) apply (subst h_t_valid_ptr_retyps_gen_disjoint, assumption) apply (simp add:szo cte_C_size cte_level_bits_def) apply (erule disjoint_subset) @@ -2189,6 +2224,7 @@ lemma createObjects_ccorres_pde: shows "\\ x. (\, x) \ rf_sr \ ptr \ 0 \ pspace_aligned' \ \ pspace_distinct' \ \ pspace_no_overlap' ptr sz \ + \ ret_zero ptr (2 ^ pdBits) \ \ region_is_zero_bytes ptr (2 ^ pdBits) x \ range_cover ptr sz pdBits 1 \ valid_global_refs' s @@ -2213,6 +2249,7 @@ proof (intro impI allI) and szb: "sz < word_bits" and pal: "pspace_aligned' \" and pdst: "pspace_distinct' \" and pno: "pspace_no_overlap' ptr sz \" + and rzo: "ret_zero ptr (2 ^ pdBits) \" and empty: "region_is_bytes ptr (2 ^ pdBits) x" and zero: "heap_list_is_zero (hrs_mem (t_hrs_' (globals x))) ptr (2 ^ pdBits)" and kernel_data_refs_disj : "kernel_data_refs \ {ptr..+ 2 ^ pdBits} = {}" @@ -2376,12 +2413,13 @@ proof (intro impI allI) done ultimately - show ?thesis using rf empty kernel_data_refs_disj + show ?thesis using rf empty kernel_data_refs_disj rzo apply (simp add: rf_sr_def cstate_relation_def Let_def rl' tag_disj_via_td_name) apply (simp add: carch_state_relation_def cmachine_state_relation_def) apply (clarsimp simp add: rl' cte_C_size cterl tag_disj_via_td_name hrs_htd_update ht_rl foldr_upd_app_if [folded data_map_insert_def] - projectKOs rl cvariable_array_ptr_retyps[OF szo]) + projectKOs rl cvariable_array_ptr_retyps[OF szo] + zero_ranges_ptr_retyps[where p="pd_Ptr ptr", simplified szo]) apply (subst h_t_valid_ptr_retyps_gen_disjoint) apply assumption apply (simp add:szo cte_C_size cte_level_bits_def) @@ -2595,55 +2633,202 @@ lemma insertNewCap_ccorres_helper: typ_heap_simps cvariable_array_map_const_add_map_option[where f="tcb_no_ctes_proj"]) -lemma insertNewCap_ccorres [corres]: - "ccorres dc xfdc (pspace_aligned' and valid_mdb') - (UNIV \ {s. ccap_relation cap (cap_' s)} \ {s. parent_' s = Ptr parent} \ {s. slot_' s = Ptr slot}) [] +definition + byte_regions_unmodified :: "heap_raw_state \ heap_raw_state \ bool" +where + "byte_regions_unmodified hrs hrs' \ \x. (\n td b. snd (hrs_htd hrs x) n = Some (td, b) + \ td = typ_uinfo_t TYPE (word8)) + \ hrs_mem hrs' x = hrs_mem hrs x" + +abbreviation + byte_regions_unmodified' :: "globals myvars \ globals myvars \ bool" +where + "byte_regions_unmodified' s t \ byte_regions_unmodified (t_hrs_' (globals s)) + (t_hrs_' (globals t))" + +lemma byte_regions_unmodified_refl[iff]: + "byte_regions_unmodified hrs hrs" + by (simp add: byte_regions_unmodified_def) + +lemma byte_regions_unmodified_trans: + "byte_regions_unmodified hrs hrs' + \ byte_regions_unmodified hrs' hrs'' + \ hrs_htd hrs' = hrs_htd hrs + \ byte_regions_unmodified hrs hrs''" + by (simp add: byte_regions_unmodified_def) + +lemma byte_regions_unmodified_hrs_mem_update1: + "byte_regions_unmodified hrs hrs' + \ hrs_htd hrs \\<^sub>t (p :: ('a :: wf_type) ptr) + \ hrs_htd hrs' = hrs_htd hrs + \ typ_uinfo_t TYPE ('a) \ typ_uinfo_t TYPE (word8) + \ byte_regions_unmodified hrs + (hrs_mem_update (heap_update p v) hrs')" + apply (erule byte_regions_unmodified_trans, simp_all) + apply (clarsimp simp: byte_regions_unmodified_def hrs_mem_update + heap_update_def h_t_valid_def + valid_footprint_def Let_def) + apply (rule heap_update_nmem_same) + apply (clarsimp simp: size_of_def intvl_def) + apply (drule spec, drule(1) mp, clarsimp) + apply (cut_tac s="(typ_uinfo_t TYPE('a))" and n=k in ladder_set_self) + apply (clarsimp dest!: in_set_list_map) + apply (drule(1) map_le_trans) + apply (simp add: map_le_def) + apply metis + done + +lemma byte_regions_unmodified_hrs_mem_update2: + "byte_regions_unmodified hrs hrs' + \ hrs_htd hrs \\<^sub>t (p :: ('a :: wf_type) ptr) + \ typ_uinfo_t TYPE ('a) \ typ_uinfo_t TYPE (word8) + \ byte_regions_unmodified (hrs_mem_update (heap_update p v) hrs) hrs'" + apply (erule byte_regions_unmodified_trans[rotated], simp_all) + apply (clarsimp simp: byte_regions_unmodified_def hrs_mem_update + heap_update_def h_t_valid_def + valid_footprint_def Let_def) + apply (rule sym, rule heap_update_nmem_same) + apply (clarsimp simp: size_of_def intvl_def) + apply (drule spec, drule(1) mp, clarsimp) + apply (cut_tac s="(typ_uinfo_t TYPE('a))" and n=k in ladder_set_self) + apply (clarsimp dest!: in_set_list_map) + apply (drule(1) map_le_trans) + apply (simp add: map_le_def) + apply metis + done + +lemmas byte_regions_unmodified_hrs_mem_update + = byte_regions_unmodified_hrs_mem_update1 + byte_regions_unmodified_hrs_mem_update2 + +lemma byte_regions_unmodified_hrs_htd_update[iff]: + "byte_regions_unmodified + (hrs_htd_update h hrs) hrs" + by (clarsimp simp: byte_regions_unmodified_def) + +lemma byte_regions_unmodified_flip: + "byte_regions_unmodified (hrs_htd_update (\_. hrs_htd hrs) hrs') hrs + \ byte_regions_unmodified hrs hrs'" + by (simp add: byte_regions_unmodified_def hrs_htd_update) + +lemma mdb_node_ptr_set_mdbPrev_preserves_bytes: + "\s. \\\<^bsub>/UNIV\<^esub> {s} Call mdb_node_ptr_set_mdbPrev_'proc + {t. hrs_htd (t_hrs_' (globals t)) = hrs_htd (t_hrs_' (globals s)) + \ byte_regions_unmodified' s t}" + apply (hoare_rule HoarePartial.ProcNoRec1) + apply (rule allI, rule conseqPre, vcg) + apply (clarsimp simp: ) + apply (intro byte_regions_unmodified_hrs_mem_update byte_regions_unmodified_refl, + simp_all add: typ_heap_simps) + done + +lemma mdb_node_ptr_set_mdbNext_preserves_bytes: + "\s. \\\<^bsub>/UNIV\<^esub> {s} Call mdb_node_ptr_set_mdbNext_'proc + {t. hrs_htd (t_hrs_' (globals t)) = hrs_htd (t_hrs_' (globals s)) + \ byte_regions_unmodified' s t}" + apply (hoare_rule HoarePartial.ProcNoRec1) + apply (rule allI, rule conseqPre, vcg) + apply (clarsimp simp: ) + apply (intro byte_regions_unmodified_hrs_mem_update byte_regions_unmodified_refl, + simp_all add: typ_heap_simps) + done + +lemma updateNewFreeIndex_noop_ccorres: + "ccorres dc xfdc (valid_objs' and cte_wp_at' (\cte. cteCap cte = cap) slot) + {s. (case untypedZeroRange cap of None \ True + | Some (a, b) \ region_is_zero_bytes a (unat ((b + 1) - a)) s)} hs + (updateNewFreeIndex slot) Skip" + (is "ccorres _ _ ?P ?P' hs _ _") + apply (simp add: updateNewFreeIndex_def getSlotCap_def) + apply (rule ccorres_guard_imp) + apply (rule ccorres_pre_getCTE[where P="\rv. cte_wp_at' (op = rv) slot and ?P" + and P'="K ?P'"]) + apply (case_tac "cteCap cte", simp_all add: ccorres_guard_imp[OF ccorres_return_Skip])[1] + defer + apply (clarsimp simp: cte_wp_at_ctes_of) + apply simp + apply (simp add: updateTrackedFreeIndex_def getSlotCap_def) + apply (rule ccorres_guard_imp) + apply (rule_tac P="\rv. cte_wp_at' (op = rv) slot and K (rv = cte) and ?P" + in ccorres_pre_getCTE[where P'="K ?P'"]) + defer + apply (clarsimp simp: cte_wp_at_ctes_of) + apply simp + apply (rule ccorres_from_vcg) + apply (rule allI, rule conseqPre, vcg) + apply (clarsimp simp: bind_def simpler_modify_def) + apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def) + apply (clarsimp simp: zero_ranges_are_zero_def + cte_wp_at_ctes_of + split: option.split) + done + +lemma byte_regions_unmodified_region_is_bytes: + "byte_regions_unmodified hrs hrs' + \ region_is_bytes' y n (hrs_htd hrs) + \ x \ {y ..+ n} + \ hrs_mem hrs' x = hrs_mem hrs x" + apply (clarsimp simp: byte_regions_unmodified_def) + apply (drule spec, erule mp) + apply (clarsimp simp: region_is_bytes'_def) + apply metis + done + +lemma insertNewCap_ccorres1: + "ccorres dc xfdc (pspace_aligned' and valid_mdb' and valid_objs' and valid_cap' cap) + ({s. (case untypedZeroRange cap of None \ True + | Some (a, b) \ region_is_zero_bytes a (unat ((b + 1) - a)) s)} + \ {s. ccap_relation cap (cap_' s)} \ {s. parent_' s = Ptr parent} + \ {s. slot_' s = Ptr slot}) [] (insertNewCap parent slot cap) (Call insertNewCap_'proc)" apply (cinit (no_ignore_call) lift: cap_' parent_' slot_') apply (rule ccorres_liftM_getCTE_cte_at) apply (rule ccorres_move_c_guard_cte) apply (simp only: ) - apply (rule ccorres_split_nothrow_novcg [OF mdb_node_get_mdbNext_heap_ccorres]) - apply ceqv - apply (erule_tac s = "next" in subst) + apply (rule ccorres_split_nothrow [OF mdb_node_get_mdbNext_heap_ccorres]) + apply ceqv + apply (erule_tac s = "next" in subst) apply csymbr - apply (ctac (no_vcg, c_lines 3) pre: ccorres_pre_getCTE ccorres_assert add: insertNewCap_ccorres_helper) + apply (ctac (c_lines 3) pre: ccorres_pre_getCTE ccorres_assert add: insertNewCap_ccorres_helper) apply (simp only: Ptr_not_null_pointer_not_zero) - apply (ctac (no_vcg) add: updateMDB_set_mdbPrev) - apply (ctac (no_vcg) add: updateMDB_set_mdbNext) - apply simp - apply wp - apply simp - apply wp - apply simp - apply (wp getCTE_wp) - apply simp - apply (rule guard_is_UNIVI) - apply simp - apply simp - apply (clarsimp simp: cte_wp_at_ctes_of) - apply (erule (2) is_aligned_3_next) + apply (ctac add: updateMDB_set_mdbPrev) + apply (rule ccorres_seq_skip'[THEN iffD1]) + apply (ctac add: updateMDB_set_mdbNext) + apply (rule updateNewFreeIndex_noop_ccorres[where cap=cap]) + apply (wp updateMDB_weak_cte_wp_at) + apply simp + apply (vcg exspec=mdb_node_ptr_set_mdbNext_preserves_bytes) + apply (wp updateMDB_weak_cte_wp_at) + apply clarsimp + apply (vcg exspec=mdb_node_ptr_set_mdbPrev_preserves_bytes) + apply (wp setCTE_weak_cte_wp_at) + apply (clarsimp simp: hrs_mem_update Collect_const_mem + simp del: imp_disjL) + apply vcg + apply simp + apply (wp getCTE_wp') + apply (clarsimp simp: hrs_mem_update) + apply vcg + apply (rule conjI) + apply (clarsimp simp: cte_wp_at_ctes_of is_aligned_3_next) + apply (clarsimp split: option.split) + apply (intro allI conjI impI; simp; clarsimp) + apply (erule trans[OF heap_list_h_eq2, rotated]) + apply (rule byte_regions_unmodified_region_is_bytes) + apply (erule byte_regions_unmodified_trans[rotated] + | simp + | rule byte_regions_unmodified_hrs_mem_update + | simp add: typ_heap_simps')+ + apply (erule trans[OF heap_list_h_eq2, rotated]) + apply (rule byte_regions_unmodified_region_is_bytes) + apply (erule byte_regions_unmodified_trans[rotated] + | simp + | rule byte_regions_unmodified_hrs_mem_update + | simp add: typ_heap_simps')+ done -lemma insertNewCap_ccorres_with_Guard: - "ccorres dc xfdc (pspace_aligned' and valid_mdb' and cte_wp_at' (\_. True) slot) - (UNIV \ {s. ccap_relation cap (cap_' s)} \ {s. parent_' s = Ptr parent} \ {s. slot_' s = Ptr slot}) [] - (insertNewCap parent slot cap) - (Guard C_Guard \hrs_htd \t_hrs \\<^sub>t \slot \ - (Call insertNewCap_'proc))" - apply (rule ccorres_guard_imp - [where Q = "(pspace_aligned' and valid_mdb' and cte_at' slot)" - and Q' = "(UNIV \ {s. ccap_relation cap (cap_' s)} \ {s. parent_' s = Ptr parent} \ {s. slot_' s = Ptr slot} \ {s. slot_' s = Ptr slot})"]) - apply (cinitlift slot_') - apply (rule ccorres_guard_imp) - apply (rule_tac ccorres_move_c_guard_cte) - apply (ctac) - apply clarsimp - apply clarsimp+ - done - -lemma insertNewCap_pre_cte_at: + lemma insertNewCap_pre_cte_at: "\\s. \ (cte_at' p s \ cte_at' p' s) \ insertNewCap p p' cap \ \_ _. False \" unfolding insertNewCap_def apply simp @@ -3027,15 +3212,6 @@ lemma tcb_ptr_orth_cte_ptrs': apply (simp add: unat_arith_simps unat_of_nat) done -lemma intvl_both_le: - "\ a \ x; unat x + y \ unat a + b \ - \ {x ..+ y} \ {a ..+ b}" - apply (clarsimp simp: intvl_def) - apply (rule_tac x="unat (x - a) + k" in exI) - apply (clarsimp simp: field_simps) - apply unat_arith - done - lemma region_is_typeless_weaken: "\ region_is_typeless a b s'; (t_hrs_' (globals s)) = (t_hrs_' (globals s')); a \ x; unat x + y \ unat a + b \ \ region_is_typeless x y s" by (clarsimp simp: region_is_typeless_def subsetD[OF intvl_both_le]) @@ -3057,6 +3233,7 @@ lemma cnc_tcb_helper: and pds: "pspace_distinct' (\\ksPSpace := ks\)" and symref: "sym_refs (state_refs_of' (\\ksPSpace := ks\))" and kssub: "dom (ksPSpace \) \ dom ks" + and rzo: "ret_zero (ctcb_ptr_to_tcb_ptr p) (2 ^ objBitsKO kotcb) \" and empty: "region_is_bytes (ctcb_ptr_to_tcb_ptr p) (2 ^ 9) x" and rep0: "heap_list (fst (t_hrs_' (globals x))) (2 ^ 9) (ctcb_ptr_to_tcb_ptr p) = replicate (2 ^ 9) 0" and kdr: "{ctcb_ptr_to_tcb_ptr p..+2 ^ 9} \ kernel_data_refs = {}" @@ -3240,9 +3417,6 @@ proof - Arrays.update (registers_C (tcbContext_C (tcbArch_C (from_bytes (replicate (size_of TYPE(tcb_C)) 0))))) (unat Kernel_C.CPSR) 0x150\\, tcbTimeSlice_C := 5\)" - have help_me: "\p v hm th. (heap_update p v hm, th) = hrs_mem_update (heap_update p v) (hm, th)" - by (simp add: hrs_mem_update_def) - have tdisj': "\y. hrs_htd (t_hrs_' (globals x)) \\<^sub>t y \ ptr_span p \ ptr_span y \ {} \ y = p" using tdisj by (auto simp: h_t_valid_clift_Some_iff) @@ -3499,6 +3673,31 @@ proof - apply (simp add: size_of_def) done + have zro: + "zero_ranges_are_zero (gsUntypedZeroRanges \) (t_hrs_' (globals x))" + using rfsr + by (clarsimp simp: rf_sr_def cstate_relation_def Let_def) + + have h_t_valid_p: + "h_t_valid (hrs_htd (t_hrs_' ?gs)) c_guard p" + using fun_cong[OF cl_tcb, where x=p] + by (clarsimp dest!: h_t_valid_clift) + + have zro': + "zero_ranges_are_zero (gsUntypedZeroRanges \) (t_hrs_' ?gs)" + using zro h_t_valid_p rzo al + apply clarsimp + apply (simp add: hrs_htd_update typ_heap_simps') + apply (intro zero_ranges_ptr_retyps, simp_all) + apply (erule caps_overlap_reserved'_subseteq) + apply (rule order_trans, rule tcb_ptr_to_ctcb_ptr_in_range') + apply (simp add: objBits_simps kotcb_def) + apply (simp add: objBits_simps kotcb_def) + apply (erule caps_overlap_reserved'_subseteq) + apply (rule intvl_start_le) + apply (simp add: cte_C_size kotcb_def objBits_simps) + done + note ht_rest = clift_eq_h_t_valid_eq[OF cl_rest, simplified] note irq = h_t_valid_eq_array_valid[where 'a=cte_C and p="ptr_coerce x" for x] @@ -3558,7 +3757,7 @@ proof - done ultimately show ?thesis - using rfsr + using rfsr zro' apply (simp add: rf_sr_def cstate_relation_def Let_def h_t_valid_clift_Some_iff tag_disj_via_td_name carch_state_relation_def cmachine_state_relation_def irq) apply (simp add: cl_cte [simplified] cl_tcb [simplified] cl_rest [simplified] tag_disj_via_td_name) @@ -3693,6 +3892,26 @@ lemma cslift_bytes_mem_update: apply simp done +lemma heap_list_eq_replicate_eq_eq: + "(heap_list hp n ptr = replicate n v) + = (\p \ {ptr ..+ n}. hp p = v)" + by (induct n arbitrary: ptr, simp_all add: intvl_Suc_right) + +lemma heap_update_list_replicate_eq: + "(heap_update_list x (replicate n v) hp y) + = (if y \ {x ..+ n} then v else hp y)" + apply (induct n arbitrary: x hp, simp_all add: intvl_Suc_right) + apply (simp split: split_if) + done + +lemma zero_ranges_are_zero_update_zero[simp]: + "zero_ranges_are_zero rs hrs + \ zero_ranges_are_zero rs (hrs_mem_update (heap_update_list ptr (replicate n 0)) hrs)" + apply (clarsimp simp: zero_ranges_are_zero_def hrs_mem_update) + apply (drule(1) bspec) + apply (clarsimp simp: heap_list_eq_replicate_eq_eq heap_update_list_replicate_eq) + done + lemma rf_sr_rep0: assumes sr: "(\, x) \ rf_sr" assumes empty: "region_is_bytes ptr sz x" @@ -3841,6 +4060,14 @@ lemma pspace_no_overlap_underlying_zero: apply (clarsimp simp: valid_machine_state'_def) apply (drule_tac x=x in spec, clarsimp simp: pointerInUserData_def) apply (clarsimp simp: typ_at'_def ko_wp_at'_def koTypeOf_eq_UserDataT) + apply (case_tac "pointerInDeviceData x \") + apply (clarsimp simp: pointerInDeviceData_def + ko_wp_at'_def obj_at'_def projectKOs + dest!: device_data_at_ko) + apply (drule(1) pspace_no_overlapD') + apply (drule_tac x=x in eqset_imp_iff) + apply (simp add: objBits_simps) + apply clarsimp apply (drule(1) pspace_no_overlapD') apply (drule_tac x=x in eqset_imp_iff, simp) apply (simp add: objBits_simps) @@ -3921,12 +4148,19 @@ lemma pageBitsForSize_mess_multi: apply (case_tac sz,simp+) done +lemma word_le_mask_out_plus_2sz: + "x \ (x && ~~ mask sz) + 2 ^ sz - 1" + using mask_in_range[where ptr'=x and bits=sz, + OF is_aligned_neg_mask2[where a=x]] + by simp + lemma createObjects_ccorres_user_data: defines "ko \ KOUserData" shows "\\ x. (\, x) \ rf_sr \ range_cover ptr sz (gbits + pageBits) n \ ptr \ 0 \ pspace_aligned' \ \ pspace_distinct' \ \ valid_machine_state' \ + \ ret_zero ptr (n * 2 ^ (gbits + pageBits)) \ \ pspace_no_overlap' ptr sz \ \ region_is_zero_bytes ptr (n * 2 ^ (gbits + pageBits)) x \ {ptr ..+ n * (2 ^ (gbits + pageBits))} \ kernel_data_refs = {} @@ -3957,6 +4191,7 @@ proof (intro impI allI) and pal: "pspace_aligned' \" and pdst: "pspace_distinct' \" and pno: "pspace_no_overlap' ptr sz \" and vms: "valid_machine_state' \" + and rzo: "ret_zero ptr (n * 2 ^ (gbits + pageBits)) \" and empty: "region_is_bytes ptr (n * 2 ^ (gbits + pageBits)) x" and zero: "heap_list_is_zero (hrs_mem (t_hrs_' (globals x))) ptr (n * 2 ^ (gbits + pageBits))" and rc: "range_cover ptr sz (gbits + pageBits) n" @@ -4055,6 +4290,8 @@ proof (intro impI allI) have zero_app: "\x. x \ {ptr..+ n * 2 ^ (gbits + pageBits) } \ underlying_memory (ksMachineState \) x = 0" + apply (cases "n = 0") + apply simp apply (rule pspace_no_overlap_underlying_zero[OF pno vms]) apply (erule subsetD[rotated]) apply (cases "n = 0") @@ -4074,7 +4311,7 @@ proof (intro impI allI) apply (rule allI) apply (subst i1) apply (simp add: byte_to_word_heap_def Let_def - zero_app nca nca [where x2 = 0, simplified]) + zero_app nca nca [where x3 = 0, simplified]) apply (simp add: word_rcat_bl) done @@ -4089,7 +4326,7 @@ proof (intro impI allI) Some v; xa \ set (new_cap_addrs (n*2^gbits) ptr KOUserData); heap_to_user_data (ksPSpace \) (underlying_memory (ksMachineState \)) xa = Some y \ \ y = v" using range_cover_intvl[OF rc] - by (clarsimp simp add: heap_to_page_data_def Let_def sz2 + by (clarsimp simp add: heap_to_user_data_def Let_def sz2 byte_to_word_heap_def[abs_def] map_comp_Some_iff projectKOs) have relrl: "cmap_relation (heap_to_user_data (ksPSpace \) (underlying_memory (ksMachineState \))) @@ -4169,12 +4406,10 @@ proof (intro impI allI) apply (simp add: objBits_simps ptr_add_to_new_cap_addrs[OF szo] ko_def cong: if_cong) apply (simp add: p2dist[symmetric]) - apply (rule conjI) - apply (erule relrl[simplified]) - apply (simp add:cmap_relation_def cuser_user_data_device_relation_def local.dom_heap_to_device_data) + apply (erule relrl[simplified]) done - thus ?thesis using rf empty kdr + thus ?thesis using rf empty kdr rzo apply (simp add: rf_sr_def cstate_relation_def Let_def rl' tag_disj_via_td_name ) apply (simp add: carch_state_relation_def cmachine_state_relation_def) apply (simp add: tag_disj_via_td_name rl' tcb_C_size h_t_valid_clift_Some_iff) @@ -4182,7 +4417,8 @@ proof (intro impI allI) apply (simp add:szo hrs_htd_def p2dist objBits_simps ko_def ptr_retyps_htd_safe_neg kernel_data_refs_domain_eq_rotate rl foldr_upd_app_if [folded data_map_insert_def] - projectKOs cvariable_array_ptr_retyps) + projectKOs cvariable_array_ptr_retyps + zero_ranges_ptr_retyps) done qed @@ -4511,6 +4747,7 @@ lemma ccorres_placeNewObject_endpoint: \ ccorresG rf_sr \ dc xfdc (pspace_aligned' and pspace_distinct' and pspace_no_overlap' regionBase (objBits ko) + and ret_zero regionBase (2 ^ objBits ko) and (\s. 2 ^ (objBits ko) \ gsMaxObjectSize s) and K (regionBase \ 0 \ range_cover regionBase (objBits ko) (objBits ko) 1 \ {regionBase..+ 2 ^ (objBits ko)} \ kernel_data_refs = {})) @@ -4548,6 +4785,7 @@ lemma ccorres_placeNewObject_notification: "ccorresG rf_sr \ dc xfdc (pspace_aligned' and pspace_distinct' and pspace_no_overlap' regionBase 4 and (\s. 16 \ gsMaxObjectSize s) + and ret_zero regionBase 16 and K (regionBase \ 0 \ {regionBase..+16} \ kernel_data_refs = {} \ range_cover regionBase 4 4 1)) @@ -4606,6 +4844,7 @@ lemma ccorres_placeNewObject_captable: "ccorresG rf_sr \ dc xfdc (pspace_aligned' and pspace_distinct' and pspace_no_overlap' regionBase (unat userSize + 4) and (\s. 2 ^ (unat userSize + 4) \ gsMaxObjectSize s) + and ret_zero regionBase (2 ^ (unat userSize + 4)) and K (regionBase \ 0 \ range_cover regionBase (unat userSize + 4) (unat userSize + 4) 1 \ ({regionBase..+2 ^ (unat userSize + 4)} \ kernel_data_refs = {}))) ({s. region_actually_is_zero_bytes regionBase (2 ^ (unat userSize + 4)) s}) @@ -4658,6 +4897,7 @@ lemma ccorres_placeNewObject_tcb: "ccorresG rf_sr \ dc xfdc (pspace_aligned' and pspace_distinct' and pspace_no_overlap' regionBase 9 and valid_queues and (\s. sym_refs (state_refs_of' s)) and (\s. 2 ^ 9 \ gsMaxObjectSize s) + and ret_zero regionBase (2 ^ 9) and K (regionBase \ 0 \ range_cover regionBase 9 9 1 \ {regionBase..+2^9} \ kernel_data_refs = {})) ({s. region_actually_is_zero_bytes regionBase 0x200 s}) @@ -4705,24 +4945,24 @@ lemma ccorres_placeNewObject_tcb: apply (clarsimp simp: split_def new_cap_addrs_def) apply (cut_tac \=\ and x=x and ks="ksPSpace \" and p="tcb_Ptr (regionBase + 0x100)" in cnc_tcb_helper) - apply clarsimp - apply (clarsimp simp: ctcb_ptr_to_tcb_ptr_def - ctcb_offset_def objBitsKO_def range_cover.aligned) - apply (clarsimp simp: ctcb_ptr_to_tcb_ptr_def ctcb_offset_def objBitsKO_def) - apply (simp add:olen_add_eqv[symmetric]) - apply (erule is_aligned_no_wrap'[OF range_cover.aligned]) + apply clarsimp + apply (clarsimp simp: ctcb_ptr_to_tcb_ptr_def + ctcb_offset_def objBitsKO_def range_cover.aligned) + apply (clarsimp simp: ctcb_ptr_to_tcb_ptr_def ctcb_offset_def objBitsKO_def) + apply (simp add:olen_add_eqv[symmetric]) + apply (erule is_aligned_no_wrap'[OF range_cover.aligned]) + apply simp apply simp - apply simp - apply (clarsimp) - apply (clarsimp simp: ctcb_ptr_to_tcb_ptr_def ctcb_offset_def objBitsKO_def) - apply (clarsimp) - apply simp - apply clarsimp + apply (clarsimp) + apply (clarsimp simp: ctcb_ptr_to_tcb_ptr_def ctcb_offset_def objBitsKO_def) + apply (clarsimp) + apply simp + apply clarsimp + apply (clarsimp simp: objBits_simps ctcb_ptr_to_tcb_ptr_def ctcb_offset_def) apply (frule region_actually_is_bytes) apply (clarsimp simp: region_is_bytes'_def ctcb_ptr_to_tcb_ptr_def ctcb_offset_def split_def hrs_mem_update_def hrs_htd_def) apply (clarsimp simp: ctcb_ptr_to_tcb_ptr_def ctcb_offset_def hrs_mem_update_def split_def) - apply (rule heap_list_update', auto simp: word_bits_conv)[1] apply (simp add: hrs_mem_def) apply (simp add: ctcb_ptr_to_tcb_ptr_def ctcb_offset_def) apply (clarsimp simp: ctcb_ptr_to_tcb_ptr_def ctcb_offset_def hrs_mem_update_def split_def) @@ -4740,6 +4980,7 @@ lemma placeNewObject_pte: "ccorresG rf_sr \ dc xfdc ( valid_global_refs' and pspace_aligned' and pspace_distinct' and pspace_no_overlap' regionBase 10 and (\s. 2 ^ 10 \ gsMaxObjectSize s) + and ret_zero regionBase (2 ^ 10) and K (regionBase \ 0 \ range_cover regionBase 10 10 1 \ ({regionBase..+2 ^ 10} \ kernel_data_refs = {}) )) @@ -4783,6 +5024,7 @@ lemma placeNewObject_pde: "ccorresG rf_sr \ dc xfdc (valid_global_refs' and pspace_aligned' and pspace_distinct' and pspace_no_overlap' regionBase 14 and (\s. 2 ^ 14 \ gsMaxObjectSize s) + and ret_zero regionBase (2 ^ 14) and K (regionBase \ 0 \ range_cover regionBase 14 14 1 \ ({regionBase..+2 ^ 14} \ kernel_data_refs = {}) @@ -4839,6 +5081,7 @@ lemma createObjects_ccorres_user_data_device: \ ptr \ 0 \ pspace_aligned' \ \ pspace_distinct' \ \ pspace_no_overlap' ptr sz \ + \ ret_zero ptr (n * 2 ^ (gbits + pageBits)) \ \ region_is_bytes ptr (n * 2 ^ (gbits + pageBits)) x \ {ptr ..+ n * (2 ^ (gbits + pageBits))} \ kernel_data_refs = {} \ @@ -4865,7 +5108,8 @@ proof (intro impI allI) and sz: "gbits + pageBits \ sz" and szb: "sz < word_bits" and pal: "pspace_aligned' \" and pdst: "pspace_distinct' \" - and pno: "pspace_no_overlap' ptr sz \" + and pno: "pspace_no_overlap' ptr sz \" + and rzo: "ret_zero ptr (n * 2 ^ (gbits + pageBits)) \" and empty: "region_is_bytes ptr (n * 2 ^ (gbits + pageBits)) x" and rc: "range_cover ptr sz (gbits + pageBits) n" and rc': "range_cover ptr sz (objBitsKO ko) (n * 2^ gbits)" @@ -4968,11 +5212,13 @@ proof (intro impI allI) apply (simp add: heap_to_device_data_def cuser_user_data_device_relation_def) done (* dont need to track all the device memory *) - thus ?thesis using rf empty kdr + thus ?thesis using rf empty kdr rzo apply (simp add: rf_sr_def cstate_relation_def Let_def rl' tag_disj_via_td_name ) apply (simp add: carch_state_relation_def cmachine_state_relation_def) apply (simp add: tag_disj_via_td_name rl' tcb_C_size h_t_valid_clift_Some_iff) apply (clarsimp simp: hrs_htd_update szo'[symmetric] cvariable_array_ptr_retyps[OF szo] rb') + apply (subst zero_ranges_ptr_retyps, simp_all only: szo'[symmetric] power_add, + simp) apply (simp add:szo p2dist objBits_simps ko_def ptr_retyps_htd_safe_neg kernel_data_refs_domain_eq_rotate rl foldr_upd_app_if [folded data_map_insert_def] @@ -4980,7 +5226,7 @@ proof (intro impI allI) apply (subst cvariable_array_ptr_retyps[OF szo]) apply (simp add: rb' ptr_retyps_htd_safe_neg)+ apply (erule ptr_retyps_htd_safe_neg) - apply (simp add:pageBits_def) + apply (simp add:pageBits_def field_simps) done qed @@ -4988,6 +5234,7 @@ lemma placeNewObject_user_data: "ccorresG rf_sr \ dc xfdc (pspace_aligned' and pspace_distinct' and pspace_no_overlap' regionBase (pageBits+us) and valid_queues and valid_machine_state' + and ret_zero regionBase (2 ^ (pageBits+us)) and (\s. sym_refs (state_refs_of' s)) and (\s. 2^(pageBits + us) \ gsMaxObjectSize s) and K (regionBase \ 0 \ range_cover regionBase (pageBits + us) (pageBits+us) (Suc 0) @@ -5031,7 +5278,8 @@ definition createObject_hs_preconds :: "word32 \ ArchTypes_H.object_type \ nat \ bool \ kernel_state \ bool" where "createObject_hs_preconds regionBase newType userSize d \ - (invs' and (pspace_no_overlap' regionBase (getObjectSize newType userSize)) + (invs' and pspace_no_overlap' regionBase (getObjectSize newType userSize) + and caps_overlap_reserved' {regionBase ..+ 2 ^ (getObjectSize newType userSize)} and (\s. 2 ^ (getObjectSize newType userSize) \ gsMaxObjectSize s) and K(regionBase \ 0 \ ({regionBase..+2 ^ (getObjectSize newType userSize)} \ kernel_data_refs = {}) @@ -5043,19 +5291,24 @@ where \ (d \ newType = APIObjectType apiobject_type.Untyped \ isFrameType newType) ))" +abbreviation + "region_actually_is_dev_bytes ptr len devMem s + \ region_actually_is_bytes ptr len s + \ (\ devMem \ heap_list_is_zero (hrs_mem (t_hrs_' (globals s))) ptr len)" + (* these preconds actually used throughout the proof *) abbreviation(input) - createObject_c_preconds1 :: "word32 \ ArchTypes_H.object_type \ nat \ (globals myvars) set" + createObject_c_preconds1 :: "word32 \ ArchTypes_H.object_type \ nat \ bool \ (globals myvars) set" where - "createObject_c_preconds1 regionBase newType userSize \ - {s. region_actually_is_zero_bytes regionBase (2 ^ getObjectSize newType userSize) s}" + "createObject_c_preconds1 regionBase newType userSize deviceMemory \ + {s. region_actually_is_dev_bytes regionBase (2 ^ getObjectSize newType userSize) deviceMemory s}" (* these preconds used at start of proof *) definition createObject_c_preconds :: "word32 \ ArchTypes_H.object_type \ nat \ bool \ (globals myvars) set" where "createObject_c_preconds regionBase newType userSize deviceMemory \ - (createObject_c_preconds1 regionBase newType userSize + (createObject_c_preconds1 regionBase newType userSize deviceMemory \ {s. object_type_from_H newType = t_' s} \ {s. Ptr regionBase = regionBase_' s} \ {s. unat (scast (userSize_' s) :: word32) = userSize} @@ -5110,9 +5363,6 @@ lemma range_coverI: apply simp done - od; - apply (simp add: clearMemory_def word_size createObject_def createPageObject_def - placeNewObject_with_memset_def (* FIXME: with the current state of affairs, we could simplify gs_new_frames *) lemma gsUserPages_update_ccorres: "ccorresG rf_sr G dc xf (\_. sz = pageBitsForSize pgsz) UNIV hs @@ -5131,7 +5381,9 @@ lemma gsUserPages_update_ccorres: lemma placeNewObject_user_data_device: "ccorresG rf_sr \ dc xfdc - (pspace_aligned' and pspace_distinct' and pspace_no_overlap' regionBase (pageBits+us) and valid_queues + (pspace_aligned' and pspace_distinct' + and ret_zero regionBase (2 ^ (pageBits + us)) + and pspace_no_overlap' regionBase (pageBits+us) and valid_queues and (\s. sym_refs (state_refs_of' s)) and (\s. 2^(pageBits + us) \ gsMaxObjectSize s) and K (regionBase \ 0 \ range_cover regionBase (pageBits + us) (pageBits+us) (Suc 0) @@ -5139,31 +5391,37 @@ lemma placeNewObject_user_data_device: ({s. region_actually_is_bytes regionBase (2^(pageBits+us)) s}) hs (placeNewObject regionBase UserDataDevice us ) - (Basic (\s. globals_update - (t_hrs_'_update (hrs_htd_update (ptr_retyps (2^us) (Ptr regionBase :: (user_data_device_C) ptr)))) s))" + (global_htd_update (\s. (ptr_retyps (2^us) (Ptr regionBase :: user_data_device_C ptr))))" apply (rule ccorres_from_vcg_nofail) - apply clarsimp + apply (clarsimp simp:) apply (rule conseqPre) apply vcg apply (clarsimp simp: rf_sr_htd_safe) - apply (subst simpler_placeNewObject_def) - apply (simp add: range_cover_def[where 'a=32, folded word_bits_def]) - apply ((simp add: objBits_simps range_cover.aligned)+)[3] - apply (subst(asm) simpler_placeNewObject_def) - apply (simp add: range_cover_def[where 'a=32, folded word_bits_def]) - apply ((simp add: objBits_simps range_cover.aligned)+)[3] - apply (simp add: simpler_modify_def) - apply (cut_tac ptr=regionBase and sz="pageBits + us" and gbits=us and arr=False - in createObjects_ccorres_user_data_device[rule_format]) - apply (fastforce simp: pageBits_def field_simps region_actually_is_bytes) - apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def fun_upd_def - carch_state_relation_def cmachine_state_relation_def - ghost_size_rel_def ghost_assertion_data_get_def - ptr_retyps_gen_def - cong: if_cong) + apply (intro conjI allI impI) + apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def + kernel_data_refs_domain_eq_rotate + elim!: ptr_retyps_htd_safe_neg[where arr=False, + unfolded ptr_retyps_gen_def, simplified]) + apply (simp add: size_of_def pageBits_def power_add mult.commute mult.left_commute) + apply (frule range_cover.unat_of_nat_shift[where gbits = "pageBits + us"]) + apply simp + apply (clarsimp simp:size_of_def power_add pageBits_def + rf_sr_def cstate_relation_def Let_def field_simps) + apply blast + apply (frule range_cover.aligned) + apply (frule range_cover.sz(1), fold word_bits_def) + apply (rule bexI [OF _ placeNewObject_eq], simp_all) + apply (cut_tac ptr=regionBase and sz="pageBits + us" and gbits=us and arr=False + in createObjects_ccorres_user_data_device[rule_format]) + apply (rule conjI, assumption, clarsimp) + apply (fastforce simp: pageBits_def field_simps region_actually_is_bytes) + apply (clarsimp elim!: rsubst[where P="\x. (\, x) \ rf_sr" for \] + simp: field_simps objBitsKO_def ptr_retyps_gen_def) + apply (simp add: objBitsKO_def field_simps) + apply (rule no_fail_pre, rule no_fail_placeNewObject) + apply (clarsimp simp: objBitsKO_def) done - lemma createObjects'_page_directory_at_global: "\ \s. n \ 0 \ range_cover ptr sz (objBitsKO val + gbits) n \ pspace_aligned' s \ pspace_distinct' s \ pspace_no_overlap' ptr sz s @@ -5181,314 +5439,64 @@ lemma gsUserPages_update: "\f. (\s. s\gsUserPages := f(gsUserPages s)\) = gsUserPages_update f" by (rule ext) simp +lemma modify_gsUserPages_update: + "modify (\s. s\gsUserPages := f(gsUserPages s)\) = modify (gsUserPages_update f)" + by (simp only: gsUserPages_update) + method arch_create_data_obj_corres_helper = (match conclusion in "ccorres ?rel ?var ?P ?P' ?hs (ARM_H.createObject object_type.SmallPageObject ?regionBase sz ?deviceMemory) (Call Arch_createObject_'proc) " for sz \ \(simp add: toAPIType_def ARM_H.createObject_def - createPageObject_def bind_assoc + placeNewDataObject_def bind_assoc ARMLargePageBits_def),subst gsUserPages_update,((rule ccorres_gen_asm)+) \) -lemma place_small_page_ccorres: - "ccorres (\a. ccap_relation (capability.ArchObjectCap a)) ret__struct_cap_C_' - (invs' and pspace_no_overlap' regionBase (APIType_capBits object_type.SmallPageObject userSize) - and (\s. 2 ^ APIType_capBits object_type.SmallPageObject userSize \ gsMaxObjectSize s) - and K({regionBase..+2 ^ (getObjectSize object_type.SmallPageObject userSize)} \ kernel_data_refs = {}) - and K(range_cover regionBase (getObjectSize object_type.SmallPageObject userSize) - (getObjectSize object_type.SmallPageObject userSize) (Suc 0)) - and K(regionBase \ 0)) - (Collect (region_actually_is_bytes regionBase (2 ^ APIType_capBits object_type.SmallPageObject userSize)) \ - \object_type_from_H object_type.SmallPageObject = \t\ \ - \\regionBase = Ptr regionBase \ \ - \unat \userSize = userSize\ \ - \to_bool \deviceMemory = deviceMemory\) - hs (ARM_H.createObject object_type.SmallPageObject regionBase userSize deviceMemory) (Call Arch_createObject_'proc)" -proof - - note if_cong[cong] +lemma placeNewDataObject_ccorres: + "ccorresG rf_sr \ dc xfdc + (createObject_hs_preconds regionBase newType us devMem + and K (APIType_capBits newType us = pageBits + us)) + ({s. region_actually_is_bytes regionBase (2 ^ (pageBits + us)) s + \ (\ devMem \ heap_list_is_zero (hrs_mem (t_hrs_' (globals s))) regionBase + (2 ^ (pageBits + us)))}) + hs + (placeNewDataObject regionBase us devMem) + (Cond {s. devMem} + (global_htd_update (\s. (ptr_retyps (2^us) (Ptr regionBase :: user_data_device_C ptr)))) + (global_htd_update (\s. (ptr_retyps (2^us) (Ptr regionBase :: user_data_C ptr)))) + )" + apply (cases devMem) + apply (simp add: placeNewDataObject_def ccorres_cond_univ_iff) + apply (rule ccorres_guard_imp, rule placeNewObject_user_data_device, simp_all) + apply (clarsimp simp: createObject_hs_preconds_def invs'_def + valid_state'_def valid_pspace'_def) + apply (simp add: placeNewDataObject_def ccorres_cond_empty_iff) + apply (rule ccorres_guard_imp, rule placeNewObject_user_data, simp_all) + apply (clarsimp simp: createObject_hs_preconds_def invs'_def + valid_state'_def valid_pspace'_def) + apply (frule range_cover.sz(1), simp add: word_bits_def) + done - show ?thesis - apply (simp add: toAPIType_def toAPIType_def ARM_H.createObject_def - createPageObject_def bind_assoc - ArchRetype_H.createObject_def bind_assoc - ARMLargePageBits_def) - apply (subst gsUserPages_update) - apply (rule ccorres_gen_asm)+ - apply (cinit' lift: t_' regionBase_' userSize_' deviceMemory_') - apply (simp add: object_type_from_H_def Kernel_C_defs from_bool_neq_0 - ccorres_cond_univ_iff ccorres_cond_empty_iff - asidInvalid_def sle_positive APIType_capBits_def - shiftL_nat ARMSmallPageBits_def) - apply (ccorres_remove_UNIV_guard) - apply (rule ccorres_if_lhs ccorres_rhs_assoc ccorres_return_C - hoare_vcg_prop[where P=True] - apply (ctac pre only: add: placeNewObject_user_data[where us = 0,simplified]) - | (rule match_ccorres, ctac add: placeNewObject_user_data_device[where us = 0, - unfolded pageBits_def,simplified] - gsUserPages_update_ccorres gsUserPages_update_ccorres - cleanCacheRange_PoU_ccorres) - | (simp add: placeNewObject_with_memset_no_device_eq[where us = 0,simplified] - dmo'_gsUserPages_upd_comm) - | (rule match_ccorres, - ctac (c_lines 2) add: placeNewObject_user_data[where us = 0, - | (rule match_ccorres[where c="call a b c d ;; e" for a b c d e], csymbr) - | wp - apply (rule ccorres_return_C) - apply simp - apply simp - apply simp - apply wp - apply vcg - apply (frule is_aligned_weaken[where y = 5,OF is_aligned_addrFromPPtr_n - [OF range_cover.aligned]],simp,simp,simp add:mask_zero) - apply (frule local.rf_sr_htd_safe, - simp add: local.rf_sr_domain_eq ptr_retyp_htd_safe_neg) - apply (clarsimp simp: invs'_def valid_state'_def valid_pspace'_def - APIType_capBits_def pageBits_def - word_bits_def ) - APIType_capBits_def cap_to_H_simps cap_small_frame_cap_lift - is_aligned_def vmrights_to_H_def mask_def to_bool_eq_0 - Kernel_C.VMReadWrite_def Kernel_C.VMNoAccess_def - Kernel_C.VMKernelOnly_def Kernel_C.VMReadOnly_def) - - done -qed +lemma cond_second_eq_seq_ccorres: + "ccorres_underlying sr Gamm r xf arrel axf G G' hs m + (Cond P (a ;; c) (b ;; c) ;; d) + = ccorres_underlying sr Gamm r xf arrel axf G G' hs m + (Cond P a b ;; c ;; d)" + apply (rule ccorres_semantic_equiv) + apply (rule semantic_equivI) + apply (auto elim!: exec_Normal_elim_cases intro: exec.Seq exec.CondTrue exec.CondFalse) + done -(* clag, but really hard to avoid *) -lemma place_large_page_ccorres: - "ccorres (\a. ccap_relation (capability.ArchObjectCap a)) ret__struct_cap_C_' - (invs' and pspace_no_overlap' regionBase (APIType_capBits object_type.LargePageObject userSize) - and (\s. 2 ^ APIType_capBits object_type.LargePageObject userSize \ gsMaxObjectSize s) - and K({regionBase..+2 ^ (getObjectSize object_type.LargePageObject userSize)} \ kernel_data_refs = {}) - and K(range_cover regionBase (getObjectSize object_type.LargePageObject userSize) - (getObjectSize object_type.LargePageObject userSize) (Suc 0)) - and K(regionBase \ 0)) - (Collect (region_actually_is_bytes regionBase (2 ^ APIType_capBits object_type.LargePageObject userSize)) \ - \object_type_from_H object_type.LargePageObject = \t\ \ - \\regionBase = Ptr regionBase \ \ - \unat \userSize = userSize\ \ - \to_bool \deviceMemory = deviceMemory\) - hs (ARM_H.createObject object_type.LargePageObject regionBase userSize deviceMemory) (Call Arch_createObject_'proc)" -proof - - note if_cong[cong] - - show ?thesis - apply (simp add: toAPIType_def toAPIType_def ARM_H.createObject_def - createPageObject_def bind_assoc - ARMLargePageBits_def) - apply (subst gsUserPages_update) - apply (rule ccorres_gen_asm)+ - apply (cinit' lift: t_' regionBase_' userSize_' deviceMemory_') - apply (simp add: object_type_from_H_def Kernel_C_defs from_bool_neq_0 - ccorres_cond_univ_iff ccorres_cond_empty_iff - asidInvalid_def sle_positive APIType_capBits_def - shiftL_nat ARMLargePageBits_def) - apply (ccorres_remove_UNIV_guard) - apply (rule ccorres_if_lhs ccorres_rhs_assoc ccorres_return_C - hoare_vcg_prop[where P=True] - | clarsimp simp: hrs_htd_update rf_sr_htd_safe - | (rule match_ccorres, ctac add: placeNewObject_user_data_device[where us = 4, - unfolded pageBits_def,simplified] - gsUserPages_update_ccorres gsUserPages_update_ccorres - cleanCacheRange_PoU_ccorres) - | (simp add: placeNewObject_with_memset_no_device_eq[where us = 4,simplified] - dmo'_gsUserPages_upd_comm) - | (rule match_ccorres, - ctac (c_lines 2) add: placeNewObject_user_data[where us = 4, - unfolded pageBits_def,simplified]) - | (rule match_ccorres[where c="call a b c d ;; e" for a b c d e], csymbr) - | wp - | vcg)+ - apply(clarsimp simp: pageBits_def ccap_relation_def invs_pspace_aligned' - APIType_capBits_def invs_queues pageBits_def - APIType_capBits_def cap_to_H_simps cap_frame_cap_lift - is_aligned_neg_mask_eq[OF range_cover.aligned] vmrights_to_H_def local.rf_sr_domain_eq - ptr_retyp_htd_safe_neg invs_pspace_distinct' is_aligned_no_wrap' - Kernel_C.VMReadWrite_def Kernel_C.VMNoAccess_def - Kernel_C.VMKernelOnly_def Kernel_C.ARMLargePage_def - Kernel_C.VMReadWrite_def Kernel_C.VMNoAccess_def from_bool_def - Kernel_C.VMReadOnly_def field_simps cl_valid_cap_def c_valid_cap_def - is_aligned_no_wrap'[OF range_cover.aligned] is_aligned_weaken[OF range_cover.aligned] - is_aligned_no_wrap'[OF is_aligned_addrFromPPtr_n[OF range_cover.aligned]]) - apply (frule is_aligned_weaken[where y = 5,OF range_cover.aligned],simp,simp add:mask_zero) - apply (frule is_aligned_weaken[where y = 5,OF is_aligned_addrFromPPtr_n - [OF range_cover.aligned]],simp,simp,simp add:mask_zero) - apply (frule is_aligned_weaken[where y = 5,OF range_cover.aligned],simp,simp add:mask_zero) - apply (frule is_aligned_weaken[where y = 5,OF is_aligned_addrFromPPtr_n - [OF range_cover.aligned]],simp,simp,simp add:mask_zero) - apply (frule local.rf_sr_htd_safe, - simp add: local.rf_sr_domain_eq ptr_retyp_htd_safe_neg) - apply (frule region_actually_is_bytes_dom_s,fastforce+)[1] - apply (frule(1) ghost_assertion_size_logic_no_unat) - apply (simp add: o_def APIType_capBits_def) - apply (auto simp: pageBits_def ccap_relation_def - APIType_capBits_def cap_to_H_simps cap_small_frame_cap_lift - is_aligned_def vmrights_to_H_def mask_def framesize_to_H_def - Kernel_C.VMReadWrite_def Kernel_C.VMNoAccess_def to_bool_eq_0 - Kernel_C.ARMLargePage_def Kernel_C.ARMSmallPage_def - Kernel_C.VMKernelOnly_def Kernel_C.VMReadOnly_def) - apply (fastforce elim!: ptr_retyps_htd_safe - [where arr=False,unfolded ptr_retyps_gen_def,simplified])+ - done -qed - -(* clag, but really hard to avoid *) -lemma place_section_ccorres: - "ccorres (\a. ccap_relation (capability.ArchObjectCap a)) ret__struct_cap_C_' - (invs' and pspace_no_overlap' regionBase (APIType_capBits object_type.SectionObject userSize) - and (\s. 2 ^ APIType_capBits object_type.SectionObject userSize \ gsMaxObjectSize s) - and K({regionBase..+2 ^ (getObjectSize object_type.SectionObject userSize)} \ kernel_data_refs = {}) - and K(range_cover regionBase (getObjectSize object_type.SectionObject userSize) - (getObjectSize object_type.SectionObject userSize) (Suc 0)) - and K(regionBase \ 0)) - (Collect (region_actually_is_bytes regionBase (2 ^ APIType_capBits object_type.SectionObject userSize)) \ - \object_type_from_H object_type.SectionObject = \t\ \ - \\regionBase = Ptr regionBase \ \ - \unat \userSize = userSize\ \ - \to_bool \deviceMemory = deviceMemory\) - hs (ARM_H.createObject object_type.SectionObject regionBase userSize deviceMemory) (Call Arch_createObject_'proc)" -proof - - note if_cong[cong] - - show ?thesis - apply (simp add: toAPIType_def ARM_H.createObject_def - createPageObject_def bind_assoc - ARMLargePageBits_def) - apply (subst gsUserPages_update) - apply (rule ccorres_gen_asm)+ - apply (cinit' lift: t_' regionBase_' userSize_' deviceMemory_') - apply (simp add: object_type_from_H_def Kernel_C_defs from_bool_neq_0 - ccorres_cond_univ_iff ccorres_cond_empty_iff - asidInvalid_def sle_positive APIType_capBits_def - shiftL_nat ARMSectionBits_def dmo'_gsUserPages_upd_comm word_sle_def - if_to_top_of_bind word_sless_def to_bool_def) - apply (ccorres_remove_UNIV_guard) - apply (rule ccorres_if_lhs ccorres_rhs_assoc ccorres_return_C - hoare_vcg_prop[where P=True] - | clarsimp simp: hrs_htd_update rf_sr_htd_safe - | (rule match_ccorres, ctac add: placeNewObject_user_data_device[where us = 8, - unfolded pageBits_def,simplified] - gsUserPages_update_ccorres gsUserPages_update_ccorres - cleanCacheRange_PoU_ccorres) - | (simp add: placeNewObject_with_memset_no_device_eq[where us = 8,simplified] - dmo'_gsUserPages_upd_comm) - | (rule match_ccorres, - ctac (c_lines 2) add: placeNewObject_user_data[where us = 8, - unfolded pageBits_def,simplified]) - | (rule match_ccorres[where c="call a b c d ;; e" for a b c d e], csymbr) - | wp - | vcg)+ - apply(clarsimp simp: pageBits_def ccap_relation_def invs_pspace_aligned' - APIType_capBits_def invs_queues pageBits_def - APIType_capBits_def cap_to_H_simps cap_frame_cap_lift - is_aligned_neg_mask_eq[OF range_cover.aligned] vmrights_to_H_def local.rf_sr_domain_eq - ptr_retyp_htd_safe_neg invs_pspace_distinct' is_aligned_no_wrap' - Kernel_C.VMReadWrite_def Kernel_C.VMNoAccess_def - Kernel_C.VMKernelOnly_def Kernel_C.ARMLargePage_def - Kernel_C.VMReadWrite_def Kernel_C.VMNoAccess_def from_bool_def - Kernel_C.VMReadOnly_def field_simps cl_valid_cap_def c_valid_cap_def - is_aligned_no_wrap'[OF range_cover.aligned] is_aligned_weaken[OF range_cover.aligned] - is_aligned_no_wrap'[OF is_aligned_addrFromPPtr_n[OF range_cover.aligned]]) - apply (frule is_aligned_weaken[where y = 5,OF range_cover.aligned],simp,simp add:mask_zero) - apply (frule is_aligned_weaken[where y = 5,OF is_aligned_addrFromPPtr_n - [OF range_cover.aligned]],simp,simp,simp add:mask_zero) - apply (frule is_aligned_weaken[where y = 5,OF range_cover.aligned],simp,simp add:mask_zero) - apply (frule is_aligned_weaken[where y = 5,OF is_aligned_addrFromPPtr_n - [OF range_cover.aligned]],simp,simp,simp add:mask_zero) - apply (frule local.rf_sr_htd_safe, - simp add: local.rf_sr_domain_eq ptr_retyp_htd_safe_neg) - apply (frule region_actually_is_bytes_dom_s,fastforce+)[1] - apply (frule(1) ghost_assertion_size_logic_no_unat) - apply (simp add: o_def APIType_capBits_def) - apply (auto simp: pageBits_def ccap_relation_def - APIType_capBits_def cap_to_H_simps cap_small_frame_cap_lift - is_aligned_def vmrights_to_H_def mask_def framesize_to_H_def - Kernel_C.VMReadWrite_def Kernel_C.VMNoAccess_def Kernel_C.ARMSection_def - Kernel_C.ARMLargePage_def Kernel_C.ARMSmallPage_def to_bool_eq_0 - Kernel_C.VMKernelOnly_def Kernel_C.VMReadOnly_def) - apply (fastforce elim!: ptr_retyps_htd_safe - [where arr=False,unfolded ptr_retyps_gen_def,simplified])+ - done -qed - -(* clag, but really hard to avoid *) -lemma place_super_section_ccorres: - "ccorres (\a. ccap_relation (capability.ArchObjectCap a)) ret__struct_cap_C_' - (invs' and pspace_no_overlap' regionBase (APIType_capBits object_type.SuperSectionObject userSize) - and (\s. 2 ^ APIType_capBits object_type.SuperSectionObject userSize \ gsMaxObjectSize s) - and K({regionBase..+2 ^ (getObjectSize object_type.SuperSectionObject userSize)} \ kernel_data_refs = {}) - and K(range_cover regionBase (getObjectSize object_type.SuperSectionObject userSize) - (getObjectSize object_type.SuperSectionObject userSize) (Suc 0)) - and K(regionBase \ 0)) - (Collect (region_actually_is_bytes regionBase (2 ^ APIType_capBits object_type.SuperSectionObject userSize)) \ - \object_type_from_H object_type.SuperSectionObject = \t\ \ - \\regionBase = Ptr regionBase \ \ - \unat \userSize = userSize\ \ - \to_bool \deviceMemory = deviceMemory\) - hs (ARM_H.createObject object_type.SuperSectionObject regionBase userSize deviceMemory) (Call Arch_createObject_'proc)" -proof - - note if_cong[cong] - - show ?thesis - apply (simp add: toAPIType_def toAPIType_def ARM_H.createObject_def - createPageObject_def bind_assoc - ARMLargePageBits_def) - apply (subst gsUserPages_update) - apply (rule ccorres_gen_asm)+ - apply (cinit' lift: t_' regionBase_' userSize_' deviceMemory_') - apply (simp add: object_type_from_H_def Kernel_C_defs from_bool_neq_0 - ccorres_cond_univ_iff ccorres_cond_empty_iff - asidInvalid_def sle_positive APIType_capBits_def - shiftL_nat ARMSuperSectionBits_def dmo'_gsUserPages_upd_comm word_sle_def - if_to_top_of_bind word_sless_def to_bool_def) - apply (ccorres_remove_UNIV_guard) - apply (rule ccorres_if_lhs ccorres_rhs_assoc ccorres_return_C - hoare_vcg_prop[where P=True] - | clarsimp simp: hrs_htd_update rf_sr_htd_safe - | (rule match_ccorres, ctac add: placeNewObject_user_data_device[where us = 12, - unfolded pageBits_def,simplified] - gsUserPages_update_ccorres gsUserPages_update_ccorres - cleanCacheRange_PoU_ccorres) - | (simp add: placeNewObject_with_memset_no_device_eq[where us = 12,simplified] - dmo'_gsUserPages_upd_comm) - | (rule match_ccorres, - ctac (c_lines 2) add: placeNewObject_user_data[where us = 12, - unfolded pageBits_def,simplified]) - | (rule match_ccorres[where c="call a b c d ;; e" for a b c d e], csymbr) - | wp - | vcg)+ - apply(clarsimp simp: pageBits_def ccap_relation_def invs_pspace_aligned' - APIType_capBits_def invs_queues pageBits_def - APIType_capBits_def cap_to_H_simps cap_frame_cap_lift - is_aligned_neg_mask_eq[OF range_cover.aligned] vmrights_to_H_def local.rf_sr_domain_eq - ptr_retyp_htd_safe_neg invs_pspace_distinct' is_aligned_no_wrap' - Kernel_C.VMReadWrite_def Kernel_C.VMNoAccess_def - Kernel_C.VMKernelOnly_def Kernel_C.ARMLargePage_def - Kernel_C.VMReadWrite_def Kernel_C.VMNoAccess_def from_bool_def - Kernel_C.VMReadOnly_def field_simps cl_valid_cap_def c_valid_cap_def - is_aligned_no_wrap'[OF range_cover.aligned] is_aligned_weaken[OF range_cover.aligned] - is_aligned_no_wrap'[OF is_aligned_addrFromPPtr_n[OF range_cover.aligned]]) - apply (frule is_aligned_weaken[where y = 5,OF range_cover.aligned],simp,simp add:mask_zero) - apply (frule is_aligned_weaken[where y = 5,OF is_aligned_addrFromPPtr_n - [OF range_cover.aligned]],simp,simp,simp add:mask_zero) - apply (frule is_aligned_weaken[where y = 5,OF range_cover.aligned],simp,simp add:mask_zero) - apply (frule is_aligned_weaken[where y = 5,OF is_aligned_addrFromPPtr_n - [OF range_cover.aligned]],simp,simp,simp add:mask_zero) - apply (frule local.rf_sr_htd_safe, - simp add: local.rf_sr_domain_eq ptr_retyp_htd_safe_neg) - apply (frule region_actually_is_bytes_dom_s,fastforce+)[1] - apply (frule(1) ghost_assertion_size_logic_no_unat) - apply (simp add: o_def APIType_capBits_def) - apply (auto simp: pageBits_def ccap_relation_def - APIType_capBits_def cap_to_H_simps cap_small_frame_cap_lift to_bool_eq_0 - is_aligned_def vmrights_to_H_def mask_def framesize_to_H_def - Kernel_C.VMReadWrite_def Kernel_C.VMNoAccess_def Kernel_C.ARMSuperSection_def - Kernel_C.ARMLargePage_def Kernel_C.ARMSmallPage_def Kernel_C.ARMSection_def - Kernel_C.VMKernelOnly_def Kernel_C.VMReadOnly_def) - apply (fastforce elim!: ptr_retyps_htd_safe - [where arr=False,unfolded ptr_retyps_gen_def,simplified])+ - done -qed +lemma fold_eq_0_to_bool: + "(v = 0) = (\ to_bool v)" + by (simp add: to_bool_def) +lemma is_aligned_neg_mask_eq_concrete: + "is_aligned p n + \ msk && ~~ mask n = ~~ mask n + \ p && msk = p" + apply (drule is_aligned_neg_mask_eq) + apply (metis word_bw_assocs word_bw_comms) + done lemma Arch_createObject_ccorres: assumes t: "toAPIType newType = None" @@ -5501,10 +5509,6 @@ lemma Arch_createObject_ccorres: proof - note if_cong[cong] - have gsUserPages_update: - "\f. (\s. s\gsUserPages := f(gsUserPages s)\) = gsUserPages_update f" - by (rule ext) simp - show ?thesis apply (clarsimp simp: createObject_c_preconds_def createObject_hs_preconds_def) @@ -5514,17 +5518,116 @@ proof - apply (cut_tac t) apply (case_tac newType, simp_all add: toAPIType_def - createPageObject_def bind_assoc + bind_assoc ARMLargePageBits_def) - -- "SmallPageObject" - apply (rule ccorres_guard_imp[OF place_small_page_ccorres],simp+)[1] - -- "LargePageObject" - apply (rule ccorres_guard_imp[OF place_large_page_ccorres],simp+)[1] - -- "SectionObject" - apply (rule ccorres_guard_imp[OF place_section_ccorres],simp+)[1] - -- "Super Section" - apply (rule ccorres_guard_imp[OF place_super_section_ccorres],simp+)[1] + apply (cinit' lift: t_' regionBase_' userSize_' deviceMemory_') + apply (simp add: object_type_from_H_def Kernel_C_defs) + apply (simp add: ccorres_cond_univ_iff ccorres_cond_empty_iff + ARMLargePageBits_def ARMSmallPageBits_def + ARMSectionBits_def ARMSuperSectionBits_def asidInvalid_def + sle_positive APIType_capBits_def shiftL_nat objBits_simps + ptBits_def archObjSize_def pageBits_def word_sle_def word_sless_def + fold_eq_0_to_bool) + apply (ccorres_remove_UNIV_guard) + apply (clarsimp simp: hrs_htd_update ptBits_def objBits_simps archObjSize_def + ARM_H.createObject_def pageBits_def + cond_second_eq_seq_ccorres modify_gsUserPages_update + intro!: ccorres_rhs_assoc) + apply ((rule ccorres_return_C | simp | wp | vcg + | (rule match_ccorres, ctac add: + placeNewDataObject_ccorres[where us=0 and newType=newType, simplified] + gsUserPages_update_ccorres[folded modify_gsUserPages_update]) + | (rule match_ccorres, csymbr))+)[1] + apply (intro conjI) + apply (clarsimp simp: createObject_hs_preconds_def + APIType_capBits_def pageBits_def) + apply (clarsimp simp: pageBits_def ccap_relation_def APIType_capBits_def + framesize_to_H_def cap_to_H_simps cap_small_frame_cap_lift + vmrights_to_H_def mask_def vm_rights_defs) + + -- "Page objects: could possibly fix the duplication here" + apply (cinit' lift: t_' regionBase_' userSize_' deviceMemory_') + apply (simp add: object_type_from_H_def Kernel_C_defs) + apply (simp add: ccorres_cond_univ_iff ccorres_cond_empty_iff + ARMLargePageBits_def ARMSmallPageBits_def + ARMSectionBits_def ARMSuperSectionBits_def asidInvalid_def + sle_positive APIType_capBits_def shiftL_nat objBits_simps + ptBits_def archObjSize_def pageBits_def word_sle_def word_sless_def + fold_eq_0_to_bool) + apply (ccorres_remove_UNIV_guard) + apply (clarsimp simp: hrs_htd_update ptBits_def objBits_simps archObjSize_def + ARM_H.createObject_def pageBits_def + cond_second_eq_seq_ccorres modify_gsUserPages_update + intro!: ccorres_rhs_assoc) + apply ((rule ccorres_return_C | simp | wp | vcg + | (rule match_ccorres, ctac add: + placeNewDataObject_ccorres[where us=4 and newType=newType, simplified] + gsUserPages_update_ccorres[folded modify_gsUserPages_update]) + | (rule match_ccorres, csymbr))+)[1] + apply (intro conjI) + apply (clarsimp simp: createObject_hs_preconds_def + APIType_capBits_def pageBits_def) + apply (clarsimp simp: pageBits_def ccap_relation_def APIType_capBits_def + framesize_to_H_def cap_to_H_simps cap_frame_cap_lift + vmrights_to_H_def mask_def vm_rights_defs vm_page_size_defs + cl_valid_cap_def c_valid_cap_def + is_aligned_neg_mask_eq_concrete[THEN sym]) + + apply (cinit' lift: t_' regionBase_' userSize_' deviceMemory_') + apply (simp add: object_type_from_H_def Kernel_C_defs) + apply (simp add: ccorres_cond_univ_iff ccorres_cond_empty_iff + ARMLargePageBits_def ARMSmallPageBits_def + ARMSectionBits_def ARMSuperSectionBits_def asidInvalid_def + sle_positive APIType_capBits_def shiftL_nat objBits_simps + ptBits_def archObjSize_def pageBits_def word_sle_def word_sless_def + fold_eq_0_to_bool) + apply (ccorres_remove_UNIV_guard) + apply (clarsimp simp: hrs_htd_update ptBits_def objBits_simps archObjSize_def + ARM_H.createObject_def pageBits_def + cond_second_eq_seq_ccorres modify_gsUserPages_update + intro!: ccorres_rhs_assoc) + apply ((rule ccorres_return_C | simp | wp | vcg + | (rule match_ccorres, ctac add: + placeNewDataObject_ccorres[where us=8 and newType=newType, simplified] + gsUserPages_update_ccorres[folded modify_gsUserPages_update]) + | (rule match_ccorres, csymbr))+)[1] + apply (intro conjI) + apply (clarsimp simp: createObject_hs_preconds_def + APIType_capBits_def pageBits_def) + apply (clarsimp simp: pageBits_def ccap_relation_def APIType_capBits_def + framesize_to_H_def cap_to_H_simps cap_frame_cap_lift + vmrights_to_H_def mask_def vm_rights_defs vm_page_size_defs + cl_valid_cap_def c_valid_cap_def + is_aligned_neg_mask_eq_concrete[THEN sym]) + + apply (cinit' lift: t_' regionBase_' userSize_' deviceMemory_') + apply (simp add: object_type_from_H_def Kernel_C_defs) + apply (simp add: ccorres_cond_univ_iff ccorres_cond_empty_iff + ARMLargePageBits_def ARMSmallPageBits_def + ARMSectionBits_def ARMSuperSectionBits_def asidInvalid_def + sle_positive APIType_capBits_def shiftL_nat objBits_simps + ptBits_def archObjSize_def pageBits_def word_sle_def word_sless_def + fold_eq_0_to_bool) + apply (ccorres_remove_UNIV_guard) + apply (clarsimp simp: hrs_htd_update ptBits_def objBits_simps archObjSize_def + ARM_H.createObject_def pageBits_def + cond_second_eq_seq_ccorres modify_gsUserPages_update + intro!: ccorres_rhs_assoc) + apply ((rule ccorres_return_C | simp | wp | vcg + | (rule match_ccorres, ctac add: + placeNewDataObject_ccorres[where us=12 and newType=newType, simplified] + gsUserPages_update_ccorres[folded modify_gsUserPages_update]) + | (rule match_ccorres, csymbr))+)[1] + apply (intro conjI) + apply (clarsimp simp: createObject_hs_preconds_def + APIType_capBits_def pageBits_def) + apply (clarsimp simp: pageBits_def ccap_relation_def APIType_capBits_def + framesize_to_H_def cap_to_H_simps cap_frame_cap_lift + vmrights_to_H_def mask_def vm_rights_defs vm_page_size_defs + cl_valid_cap_def c_valid_cap_def + is_aligned_neg_mask_eq_concrete[THEN sym]) + -- "PageTableObject" apply (cinit' lift: t_' regionBase_' userSize_' deviceMemory_') apply (simp add: object_type_from_H_def Kernel_C_defs) @@ -5536,6 +5639,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_H.createObject_def pageBits_def) apply (ctac pre only: add: placeNewObject_pte[simplified]) apply csymbr apply (rule ccorres_return_C) @@ -5547,14 +5651,15 @@ proof - apply clarify apply (intro conjI) apply (clarsimp simp: invs_pspace_aligned' invs_pspace_distinct' invs_valid_global' - APIType_capBits_def invs_queues) + APIType_capBits_def invs_queues invs_valid_objs' + invs_urz) apply clarsimp apply (clarsimp simp: pageBits_def ccap_relation_def APIType_capBits_def framesize_to_H_def cap_to_H_simps cap_page_table_cap_lift is_aligned_neg_mask_eq vmrights_to_H_def Kernel_C.VMReadWrite_def Kernel_C.VMNoAccess_def Kernel_C.VMKernelOnly_def Kernel_C.VMReadOnly_def) - apply (simp add: to_bool_def false_def) + apply (simp add: to_bool_def false_def isFrameType_def) -- "PageDirectoryObject" apply (cinit' lift: t_' regionBase_' userSize_' deviceMemory_') @@ -5569,7 +5674,6 @@ proof - ARM_H.createObject_def pageBits_def pdBits_def) apply (ctac pre only: add: placeNewObject_pde[simplified]) apply (ctac add: copyGlobalMappings_ccorres) -sorry (* apply csymbr apply (ctac add: cleanCacheRange_PoU_ccorres) apply csymbr @@ -5595,29 +5699,19 @@ sorry (* apply (clarsimp simp: invs_pspace_aligned' invs_pspace_distinct' archObjSize_def invs_valid_global' makeObject_pde pdBits_def pageBits_def range_cover.aligned projectKOs APIType_capBits_def - object_type_from_H_def objBits_simps) + object_type_from_H_def objBits_simps + invs_valid_objs' isFrameType_def) apply (frule invs_arch_state') apply (frule range_cover.aligned) apply (frule is_aligned_addrFromPPtr_n, simp) - apply (intro conjI) - apply fastforce - apply simp+ - apply (clarsimp simp: pageBits_def - valid_arch_state'_def page_directory_at'_def pdBits_def) - apply assumption - apply (clarsimp simp: is_aligned_no_overflow'[where n=14, simplified] - field_simps is_aligned_mask[symmetric] mask_AND_less_0)+ - apply (clarsimp elim!: is_aligned_weaken - simp: is_aligned_no_wrap' APIType_capBits_def - dest!: range_cover.aligned)+ - apply (clarsimp simp: is_aligned_def) - apply (erule region_actually_is_bytes_dom_s) - apply (clarsimp simp: APIType_capBits_def rf_sr_def cstate_relation_def - Let_def) - apply (frule(1) ghost_assertion_size_logic_no_unat) - apply (simp add: o_def APIType_capBits_def) + apply (intro conjI, simp_all) + apply fastforce + apply fastforce + apply (clarsimp simp: pageBits_def + valid_arch_state'_def page_directory_at'_def pdBits_def) + apply (clarsimp simp: is_aligned_no_overflow'[where n=14, simplified] + field_simps is_aligned_mask[symmetric] mask_AND_less_0)+ done -*) qed lemma add_ge0_weak: @@ -5732,9 +5826,9 @@ proof - (* Architecture specific objects. *) apply (rule_tac Q="createObject_hs_preconds regionBase newType userSize isdev" and - S="createObject_c_preconds1 regionBase newType userSize" and + S="createObject_c_preconds1 regionBase newType userSize isdev" and R="createObject_hs_preconds regionBase newType userSize isdev" and - T="createObject_c_preconds1 regionBase newType userSize" + T="createObject_c_preconds1 regionBase newType userSize isdev" in ccorres_Cond_rhs) apply (subgoal_tac "toAPIType newType = None") apply clarsimp @@ -5799,7 +5893,7 @@ proof - A ="createObject_hs_preconds regionBase (APIObjectType apiobject_type.TCBObject) (unat userSizea) isdev" and A'="createObject_c_preconds1 regionBase - (APIObjectType apiobject_type.TCBObject) (unat userSizea)" in + (APIObjectType apiobject_type.TCBObject) (unat userSizea) isdev" in ccorres_guard_imp2) apply (rule ccorres_symb_exec_r) apply (ccorres_remove_UNIV_guard) @@ -5841,7 +5935,8 @@ proof - aligned_add_aligned split: option.splits) apply (clarsimp simp: ctcb_ptr_to_tcb_ptr_def ctcb_offset_def - tcb_ptr_to_ctcb_ptr_def) + tcb_ptr_to_ctcb_ptr_def + invs_valid_objs' invs_urz isFrameType_def) apply (subst is_aligned_neg_mask) apply (rule aligned_add_aligned_simple [where n=8]) apply (clarsimp elim!: is_aligned_weaken @@ -5849,7 +5944,7 @@ proof - apply (clarsimp simp: is_aligned_def) apply (clarsimp simp: word_bits_def) apply simp - apply simp + apply clarsimp (* Endpoint *) apply (clarsimp simp: Kernel_C_defs object_type_from_H_def @@ -5862,7 +5957,7 @@ proof - (unat (userSizea :: word32)) isdev" and A'="createObject_c_preconds1 regionBase (APIObjectType apiobject_type.EndpointObject) - (unat userSizea)" in + (unat userSizea) isdev" in ccorres_guard_imp2) apply (ccorres_remove_UNIV_guard) apply (simp add: hrs_htd_update) @@ -5878,7 +5973,7 @@ proof - Collect_const_mem cap_endpoint_cap_lift to_bool_def true_def split: option.splits dest!: range_cover.aligned) - apply (clarsimp simp: createObject_hs_preconds_def) + apply (clarsimp simp: createObject_hs_preconds_def isFrameType_def) apply (frule invs_pspace_aligned') apply (frule invs_pspace_distinct') apply (frule invs_queues) @@ -5903,7 +5998,7 @@ proof - (unat (userSizea :: word32)) isdev" and A'="createObject_c_preconds1 regionBase (APIObjectType apiobject_type.NotificationObject) - (unat userSizea)" in + (unat userSizea) isdev" in ccorres_guard_imp2) apply (ccorres_remove_UNIV_guard) apply (simp add: hrs_htd_update) @@ -5918,7 +6013,7 @@ proof - apiGetObjectSize_def ntfnSizeBits_def objBits_simps Collect_const_mem cap_notification_cap_lift to_bool_def true_def dest!: range_cover.aligned split: option.splits) - apply (clarsimp simp: createObject_hs_preconds_def) + apply (clarsimp simp: createObject_hs_preconds_def isFrameType_def) apply (frule invs_pspace_aligned') apply (frule invs_pspace_distinct') apply (frule invs_queues) @@ -5944,7 +6039,7 @@ proof - (unat (userSizea :: word32)) isdev" and A'="createObject_c_preconds1 regionBase (APIObjectType apiobject_type.CapTableObject) - (unat userSizea)" in + (unat userSizea) isdev" in ccorres_guard_imp2) apply (simp add:field_simps hrs_htd_update) apply (ccorres_remove_UNIV_guard) @@ -5960,7 +6055,7 @@ proof - apply wp apply vcg apply (rule conjI) - apply (clarsimp simp: createObject_hs_preconds_def) + apply (clarsimp simp: createObject_hs_preconds_def isFrameType_def) apply (frule invs_pspace_aligned') apply (frule invs_pspace_distinct') apply (frule invs_queues) @@ -5970,9 +6065,10 @@ proof - ARM_H.getObjectSize_def apiGetObjectSize_def cteSizeBits_def word_bits_conv add.commute createObject_c_preconds_def region_actually_is_bytes_def + invs_valid_objs' invs_urz elim!: is_aligned_no_wrap' dest: word_of_nat_le intro!: range_coverI)[1] - apply (clarsimp simp: createObject_hs_preconds_def hrs_htd_update) + apply (clarsimp simp: createObject_hs_preconds_def hrs_htd_update isFrameType_def) apply (frule range_cover.strong_times_32[folded addr_card_wb], simp+) apply (subst h_t_array_valid_retyp, simp+) apply (simp add: power_add cte_C_size) @@ -5994,8 +6090,7 @@ proof - apply auto[1] apply (clarsimp simp: createObject_c_preconds_def) apply (clarsimp simp:nAPIOBjects_object_type_from_H)? - apply (auto simp: createObject_c_preconds_def objBits_simps field_simps - split: apiobject_type.splits)[1] + apply (intro impI conjI, simp_all)[1] apply (clarsimp simp: nAPIObjects_def object_type_from_H_def Kernel_C_defs split: object_type.splits) apply (clarsimp simp: createObject_c_preconds_def @@ -6727,7 +6822,8 @@ lemma ccorres_typ_region_bytes_dummy: tcb_ctes_typ_region_bytes[OF _ _ invs_pspace_aligned']) apply (simp add: cmap_array_typ_region_bytes_triv invs_pspace_aligned' pdBits_def pageBits_def ptBits_def - objBitsT_simps word_bits_def) + objBitsT_simps word_bits_def + zero_ranges_are_zero_typ_region_bytes) apply (rule htd_safe_typ_region_bytes, simp) apply blast done @@ -7370,15 +7466,6 @@ lemma range_cover_bound3: apply (clarsimp simp: field_simps) done -lemma region_is_bytes_update: - "{ptr ..+ len} \ {ptr' ..+ 2 ^ bits} - \ region_is_bytes ptr len - (globals_update (t_hrs_'_update - (hrs_htd_update (typ_region_bytes ptr' bits))) s)" - apply (clarsimp simp: region_is_bytes'_def typ_region_bytes_def hrs_htd_update) - apply (simp add: subsetD split: split_if_asm) - done - lemma range_cover_gsMaxObjectSize: "cte_wp_at' (\cte. cteCap cte = UntypedCap dev (ptr &&~~ mask sz) sz idx) srcSlot s \ range_cover ptr sz (APIType_capBits newType userSize) (length destSlots) @@ -7412,7 +7499,8 @@ context begin interpretation Arch . (*FIXME: arch_split*) crunch gsCNodes[wp]: insertNewCap, Arch_createNewCaps, threadSet, "Arch.createObject" "\s. P (gsCNodes s)" (wp: crunch_wps setObject_ksPSpace_only - simp: unless_def updateObject_default_def ignore: getObject setObject) + simp: unless_def updateObject_default_def crunch_simps + ignore: getObject setObject) lemma createNewCaps_1_gsCNodes_p: "\\s. P (gsCNodes s p) \ p \ ptr\ createNewCaps newType ptr 1 n dev\\rv s. P (gsCNodes s p)\" @@ -7456,11 +7544,11 @@ lemma range_cover_not_in_neqD: apply simp done -end - crunch gsMaxObjectSize[wp]: createObject "\s. P (gsMaxObjectSize s)" (simp: crunch_simps unless_def wp: crunch_wps) +end + context kernel_m begin lemma nat_le_induct [consumes 1, case_names base step]: @@ -7486,155 +7574,12 @@ proof - by simp qed -lemma gsUntypedZeroRanges_rf_sr: - "\ (start, end) \ gsUntypedZeroRanges s; (s, s') \ rf_sr \ - \ region_is_zero_bytes start (unat ((end + 1) - start)) s'" - sorry - -lemma ctes_of_untyped_zero_rf_sr: - "\ ctes_of s p = Some cte; (s, s') \ rf_sr; invs' s; - untypedZeroRange (cteCap cte) = Some (start, end) \ - \ region_is_zero_bytes start (unat ((end + 1) - start)) s'" - apply (erule gsUntypedZeroRanges_rf_sr[rotated]) - apply (clarsimp simp: invs'_def valid_state'_def untyped_ranges_zero_inv_def) - apply (rule_tac a=p in ranI) - apply (simp add: map_comp_def cteCaps_of_def) - done - 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) apply (auto elim!: exec_Normal_elim_cases intro: exec.Guard) done -lemma heap_list_is_zero_mono: - "heap_list_is_zero hmem p n \ n' \ n - \ heap_list_is_zero hmem p n'" - apply (induct n arbitrary: n' p) - apply simp - apply clarsimp - apply (case_tac n', simp_all) - done - -lemma heap_list_h_eq_better: - "\p. \ x \ {p..+q}; heap_list h q p = heap_list h' q p \ - \ h x = h' x" -proof (induct q) - case 0 thus ?case by simp -next - case (Suc n) thus ?case by (force dest: intvl_neq_start) -qed - -lemma heap_list_is_zero_mono2: - "heap_list_is_zero hmem p n - \ {p' ..+ n'} \ {p ..+ n} - \ heap_list_is_zero hmem p' n'" - using heap_list_h_eq2[where h'="\_. 0"] - heap_list_h_eq_better[where h'="\_. 0"] - apply (simp(no_asm_use) add: heap_list_rpbs) - apply blast - done - -definition - byte_regions_unmodified :: "heap_raw_state \ heap_raw_state \ bool" -where - "byte_regions_unmodified hrs hrs' \ \x. (\n td b. snd (hrs_htd hrs x) n = Some (td, b) - \ td = typ_uinfo_t TYPE (word8)) - \ hrs_mem hrs' x = hrs_mem hrs x" - -abbreviation - byte_regions_unmodified' :: "globals myvars \ globals myvars \ bool" -where - "byte_regions_unmodified' s t \ byte_regions_unmodified (t_hrs_' (globals s)) - (t_hrs_' (globals t))" - -lemma byte_regions_unmodified_refl[iff]: - "byte_regions_unmodified hrs hrs" - by (simp add: byte_regions_unmodified_def) - -lemma byte_regions_unmodified_trans: - "byte_regions_unmodified hrs hrs' - \ byte_regions_unmodified hrs' hrs'' - \ hrs_htd hrs' = hrs_htd hrs - \ byte_regions_unmodified hrs hrs''" - by (simp add: byte_regions_unmodified_def) - -lemma byte_regions_unmodified_hrs_mem_update1: - "byte_regions_unmodified hrs hrs' - \ hrs_htd hrs \\<^sub>t (p :: ('a :: wf_type) ptr) - \ hrs_htd hrs' = hrs_htd hrs - \ typ_uinfo_t TYPE ('a) \ typ_uinfo_t TYPE (word8) - \ byte_regions_unmodified hrs - (hrs_mem_update (heap_update p v) hrs')" - apply (erule byte_regions_unmodified_trans, simp_all) - apply (clarsimp simp: byte_regions_unmodified_def hrs_mem_update - heap_update_def h_t_valid_def - valid_footprint_def Let_def) - apply (rule heap_update_nmem_same) - apply (clarsimp simp: size_of_def intvl_def) - apply (drule spec, drule(1) mp, clarsimp) - apply (cut_tac s="(typ_uinfo_t TYPE('a))" and n=k in ladder_set_self) - apply (clarsimp dest!: in_set_list_map) - apply (drule(1) map_le_trans) - apply (simp add: map_le_def) - apply metis - done - -lemma byte_regions_unmodified_hrs_mem_update2: - "byte_regions_unmodified hrs hrs' - \ hrs_htd hrs \\<^sub>t (p :: ('a :: wf_type) ptr) - \ typ_uinfo_t TYPE ('a) \ typ_uinfo_t TYPE (word8) - \ byte_regions_unmodified (hrs_mem_update (heap_update p v) hrs) hrs'" - apply (erule byte_regions_unmodified_trans[rotated], simp_all) - apply (clarsimp simp: byte_regions_unmodified_def hrs_mem_update - heap_update_def h_t_valid_def - valid_footprint_def Let_def) - apply (rule sym, rule heap_update_nmem_same) - apply (clarsimp simp: size_of_def intvl_def) - apply (drule spec, drule(1) mp, clarsimp) - apply (cut_tac s="(typ_uinfo_t TYPE('a))" and n=k in ladder_set_self) - apply (clarsimp dest!: in_set_list_map) - apply (drule(1) map_le_trans) - apply (simp add: map_le_def) - apply metis - done - -lemmas byte_regions_unmodified_hrs_mem_update - = byte_regions_unmodified_hrs_mem_update1 - byte_regions_unmodified_hrs_mem_update2 - -lemma byte_regions_unmodified_hrs_htd_update[iff]: - "byte_regions_unmodified - (hrs_htd_update h hrs) hrs" - by (clarsimp simp: byte_regions_unmodified_def) - -lemma byte_regions_unmodified_flip: - "byte_regions_unmodified (hrs_htd_update (\_. hrs_htd hrs) hrs') hrs - \ byte_regions_unmodified hrs hrs'" - by (simp add: byte_regions_unmodified_def hrs_htd_update) - -lemma mdb_node_ptr_set_mdbPrev_preserves_bytes: - "\s. \\\<^bsub>/UNIV\<^esub> {s} Call mdb_node_ptr_set_mdbPrev_'proc - {t. hrs_htd (t_hrs_' (globals t)) = hrs_htd (t_hrs_' (globals s)) - \ byte_regions_unmodified' s t}" - apply (hoare_rule HoarePartial.ProcNoRec1) - apply (rule allI, rule conseqPre, vcg) - apply (clarsimp simp: ) - apply (intro byte_regions_unmodified_hrs_mem_update byte_regions_unmodified_refl, - simp_all add: typ_heap_simps) - done - -lemma mdb_node_ptr_set_mdbNext_preserves_bytes: - "\s. \\\<^bsub>/UNIV\<^esub> {s} Call mdb_node_ptr_set_mdbNext_'proc - {t. hrs_htd (t_hrs_' (globals t)) = hrs_htd (t_hrs_' (globals s)) - \ byte_regions_unmodified' s t}" - apply (hoare_rule HoarePartial.ProcNoRec1) - apply (rule allI, rule conseqPre, vcg) - apply (clarsimp simp: ) - apply (intro byte_regions_unmodified_hrs_mem_update byte_regions_unmodified_refl, - simp_all add: typ_heap_simps) - done - lemma insertNewCap_preserves_bytes: "\s. \\\<^bsub>/UNIV\<^esub> {s} Call insertNewCap_'proc {t. hrs_htd (t_hrs_' (globals t)) = hrs_htd (t_hrs_' (globals s)) @@ -7813,20 +7758,9 @@ lemma offset_intvl_first_chunk_subsets_unat: apply (simp add: word_le_nat_alt word_less_nat_alt unat_of_nat) done -lemma byte_regions_unmodified_region_is_bytes: - "byte_regions_unmodified hrs hrs' - \ region_is_bytes' y n (hrs_htd hrs) - \ x \ {y ..+ n} - \ hrs_mem hrs' x = hrs_mem hrs x" - apply (clarsimp simp: byte_regions_unmodified_def) - apply (drule spec, erule mp) - apply (clarsimp simp: region_is_bytes'_def) - apply metis - done - lemma retype_offs_region_is_zero_bytes: - "\ ctes_of s p = Some cte; (s, s') \ rf_sr; invs' s; - cteCap cte = UntypedCap (ptr &&~~ mask sz) sz idx; + "\ ctes_of s p = Some cte; (s, s') \ rf_sr; untyped_ranges_zero' s; + cteCap cte = UntypedCap False (ptr &&~~ mask sz) sz idx; idx \ unat (ptr && mask sz); range_cover ptr sz (getObjectSize newType userSize) num_ret \ \ region_is_zero_bytes ptr @@ -7857,6 +7791,66 @@ lemma retype_offs_region_is_zero_bytes: apply simp done +lemma createNewCaps_valid_cap_hd: + "\\s. pspace_no_overlap' ptr sz s \ + valid_pspace' s \ n \ 0 \ + range_cover ptr sz (APIType_capBits ty us) n \ + (ty = APIObjectType ArchTypes_H.CapTableObject \ 0 < us) \ + (ty = APIObjectType ArchTypes_H.apiobject_type.Untyped \ 4\ us \ us \ 29) \ + ptr \ 0 \ + createNewCaps ty ptr n us dev + \\r s. s \' hd r\" + apply (cases "n = 0") + apply simp + apply (rule hoare_chain) + apply (rule hoare_vcg_conj_lift) + apply (rule createNewCaps_ret_len) + apply (rule createNewCaps_valid_cap'[where sz=sz]) + apply (clarsimp simp: range_cover_n_wb) + apply simp + done + +lemma insertNewCap_ccorres: + "ccorres dc xfdc (pspace_aligned' and valid_mdb' and cte_wp_at' (\_. True) slot + and valid_objs' and valid_cap' cap) + ({s. cap_get_tag (cap_' s) = scast cap_untyped_cap + \ (case untypedZeroRange (cap_to_H (the (cap_lift (cap_' s)))) of None \ True + | Some (a, b) \ region_is_zero_bytes a (unat ((b + 1) - a)) s)} + \ {s. ccap_relation cap (cap_' s)} \ {s. parent_' s = Ptr parent} + \ {s. slot_' s = Ptr slot}) [] + (insertNewCap parent slot cap) + (Call insertNewCap_'proc)" + (is "ccorres _ _ ?P ?P' _ _ _") + apply (rule ccorres_guard_imp2, rule insertNewCap_ccorres1) + apply (clarsimp simp: cap_get_tag_isCap) + apply (clarsimp simp: ccap_relation_def map_option_Some_eq2) + apply (simp add: untypedZeroRange_def Let_def) + done + +lemma createObject_untyped_region_is_zero_bytes: + "\\. \\\<^bsub>/UNIV\<^esub> {s. let tp = (object_type_to_H (t_' s)); + sz = APIType_capBits tp (unat (userSize_' s)) + in (\ to_bool (deviceMemory_' s) + \ region_is_zero_bytes (ptr_val (regionBase_' s)) (2 ^ sz) s) + \ is_aligned (ptr_val (regionBase_' s)) sz + \ sz < 32 \ (tp = APIObjectType ArchTypes_H.apiobject_type.Untyped \ sz \ 4)} + Call createObject_'proc + {t. cap_get_tag (ret__struct_cap_C_' t) = scast cap_untyped_cap + \ (case untypedZeroRange (cap_to_H (the (cap_lift (ret__struct_cap_C_' t)))) of None \ True + | Some (a, b) \ region_is_zero_bytes a (unat ((b + 1) - a)) t)}" + apply (rule allI, rule conseqPre, vcg exspec=copyGlobalMappings_modifies + exspec=Arch_initContext_modifies + exspec=cleanCacheRange_PoU_modifies) + apply (clarsimp simp: cap_tag_defs) + apply (simp add: cap_lift_untyped_cap cap_tag_defs cap_to_H_simps + cap_untyped_cap_lift_def object_type_from_H_def) + apply (simp add: untypedZeroRange_def split: split_if) + apply (clarsimp simp: getFreeRef_def Let_def object_type_to_H_def) + apply (simp add: is_aligned_neg_mask_eq[OF is_aligned_weaken]) + apply (simp add: APIType_capBits_def + less_mask_eq word_less_nat_alt) + done + lemma createNewObjects_ccorres: notes blah[simp del] = atLeastAtMost_iff atLeastatMost_subset_iff atLeastLessThan_iff Int_atLeastAtMost atLeastatMost_empty_iff split_paired_Ex @@ -7891,8 +7885,8 @@ shows "ccorres dc xfdc isFrameType newType) \ (unat num = length destSlots) ))) - ({s. region_is_zero_bytes ptr - (length destSlots * 2 ^ APIType_capBits newType userSize) s} + ({s. (\ isdev \ region_is_zero_bytes ptr + (length destSlots * 2 ^ APIType_capBits newType userSize) s)} \ {s. t_' s = object_type_from_H newType} \ {s. parent_' s = cte_Ptr srcSlot} \ {s. slots_' s = slot_range_C (cte_Ptr cnodeptr) start num @@ -7931,7 +7925,7 @@ shows "ccorres dc xfdc apply simp apply (clarsimp simp: whileAnno_def) apply (rule ccorres_rel_imp) - apply (rule_tac Q="{s. region_is_zero_bytes + apply (rule_tac Q="{s. \ isdev \ region_is_zero_bytes (ptr + (i_' s << APIType_capBits newType userSize)) (unat (num - i_' s) * 2 ^ APIType_capBits newType userSize) s}" in ccorres_zipWithM_x_while_genQ[where j=1, OF _ _ _ _ _ i_xf_for_sequence, simplified]) @@ -7954,25 +7948,29 @@ shows "ccorres dc xfdc rule_tac S="{ptr .. ptr + of_nat (length destSlots) * 2^ (getObjectSize newType userSize) - 1}" in ccorres_typ_region_bytes_dummy, ceqv) apply (rule ccorres_Guard_Seq)+ - apply (ctac (no_vcg) add:createObject_ccorres) - apply (rule ccorres_move_array_assertion_cnode_ctes - ccorres_move_c_guard_cte)+ - apply (rule ccorres_add_return2) - apply (ctac (no_vcg) add: insertNewCap_ccorres_with_Guard) + apply (ctac add:createObject_ccorres) apply (rule ccorres_move_array_assertion_cnode_ctes - ccorres_return_Skip')+ - apply wp - apply (clarsimp simp:createObject_def3 conj_ac) - apply (wp createNewCaps_valid_pspace_extras[where sz = sz] - createNewCaps_cte_wp_at[where sz = sz]) - apply (rule range_cover_one) - apply (rule aligned_add_aligned[OF is_aligned_shiftl_self]) - apply (simp add:range_cover.aligned) + ccorres_move_c_guard_cte)+ + apply (rule ccorres_add_return2) + apply (ctac (no_vcg) add: insertNewCap_ccorres) + apply (rule ccorres_move_array_assertion_cnode_ctes + ccorres_return_Skip')+ + apply wp + apply (clarsimp simp:createObject_def3 conj_ac) + apply (wp createNewCaps_valid_pspace_extras[where sz = sz] + createNewCaps_cte_wp_at[where sz = sz] + createNewCaps_valid_cap_hd[where sz = sz]) + apply (rule range_cover_one) + apply (rule aligned_add_aligned[OF is_aligned_shiftl_self]) + apply (simp add:range_cover.aligned) + apply (simp add:range_cover_def) apply (simp add:range_cover_def) apply (simp add:range_cover_def) - apply (simp add:range_cover_def) - apply (simp add:range_cover.sz) - apply (wp createNewCaps_1_gsCNodes_p[simplified])[1] + apply (simp add:range_cover.sz) + apply (wp createNewCaps_1_gsCNodes_p[simplified] + createNewCaps_cte_wp_at'[where sz=sz])[1] + apply clarsimp + apply (vcg exspec=createObject_untyped_region_is_zero_bytes) apply (simp add:size_of_def) apply (rule_tac P = "\s. cte_wp_at' (\cte. isUntypedCap (cteCap cte) \ {ptr .. ptr + (of_nat (length destSlots)<< APIType_capBits newType userSize) - 1} \ untypedRange (cteCap cte)) srcSlot s @@ -8030,8 +8028,14 @@ shows "ccorres dc xfdc subgoal premises prems using prems by (simp add: range_cover_sz'[where 'a=32, folded word_bits_def] range_cover.sz[where 'a=32, folded word_bits_def])+ + apply (erule caps_overlap_reserved'_subseteq) + apply (frule_tac x = "of_nat n" in range_cover_bound3) + apply (rule word_of_nat_less) + apply (simp add:range_cover.unat_of_nat_n) + apply (clarsimp simp:field_simps shiftl_t2n blah) apply (erule disjoint_subset[rotated]) apply (rule_tac p1 = n in subset_trans[OF _ range_cover_subset]) + apply (simp add: upto_intvl_eq is_aligned_add range_cover.aligned is_aligned_shiftl) apply (simp add:field_simps shiftl_t2n) apply simp+ apply (erule caps_overlap_reserved'_subseteq) @@ -8041,18 +8045,26 @@ shows "ccorres dc xfdc apply (clarsimp simp: field_simps shiftl_t2n blah) apply (clarsimp simp:createObject_c_preconds_def field_simps cong:region_is_bytes_cong) apply vcg - apply clarsimp + apply (clarsimp simp: cte_C_size conj_comms) + apply (simp cong: conj_cong) apply (intro conjI impI) - apply (clarsimp simp: typ_region_bytes_actually_is_bytes hrs_htd_update) - apply (erule heap_list_is_zero_mono) - apply (subgoal_tac "unat (num - of_nat n) \ 0") - apply simp - apply (simp only: unat_eq_0, clarsimp simp: unat_of_nat) - apply (simp add:unat_eq_def) - apply (simp add: cte_C_size) - apply (rule word_of_nat_less) - subgoal by (case_tac newType,simp_all add: objBits_simps - APIType_capBits_def range_cover_def split:apiobject_type.splits) + apply (simp add: unat_eq_def) + apply (drule range_cover_sz') + apply (simp add: unat_eq_def word_less_nat_alt) + apply clarsimp + apply (erule heap_list_is_zero_mono) + apply (subgoal_tac "unat (num - of_nat n) \ 0") + apply simp + apply (simp only: unat_eq_0, clarsimp simp: unat_of_nat) + apply (simp add: hrs_htd_update typ_region_bytes_actually_is_bytes) + apply (frule range_cover_sz') + apply (clarsimp simp: Let_def hrs_htd_update + APIType_capBits_def[where ty="APIObjectType ArchTypes_H.apiobject_type.Untyped"]) + apply (subst is_aligned_add, erule range_cover.aligned) + apply (simp add: is_aligned_shiftl)+ + apply clarsimp + apply (rule region_is_bytes_typ_region_bytes) + apply simp apply clarsimp apply (subst range_cover.unat_of_nat_n) apply (erule range_cover_le) diff --git a/proof/crefine/SR_lemmas_C.thy b/proof/crefine/SR_lemmas_C.thy index b28ee8a24..f98cf8962 100644 --- a/proof/crefine/SR_lemmas_C.thy +++ b/proof/crefine/SR_lemmas_C.thy @@ -1668,6 +1668,35 @@ lemma heap_to_device_data_cong [cong]: \ heap_to_device_data ks bhp = heap_to_device_data ks' bhp'" unfolding heap_to_device_data_def by simp +lemma map_leD: + "\ map_le m m'; m x = Some y \ \ m' x = Some y" + by (simp add: map_le_def dom_def) + +lemma region_is_bytes_disjoint: + assumes cleared: "region_is_bytes' p n (hrs_htd hrs)" + and not_byte: "typ_uinfo_t TYPE('a :: wf_type) \ typ_uinfo_t TYPE(word8)" + shows "hrs_htd hrs \\<^sub>t (p' :: 'a ptr) + \ {p ..+ n} \ {ptr_val p' ..+ size_of TYPE('a)} = {}" + apply (clarsimp simp: h_t_valid_def valid_footprint_def Let_def) + apply (clarsimp simp: set_eq_iff dest!: intvlD[where p="ptr_val p'"]) + apply (drule_tac x="of_nat k" in spec, clarsimp simp: size_of_def) + apply (cut_tac m=k in typ_slice_t_self[where td="typ_uinfo_t TYPE('a)"]) + apply (clarsimp simp: in_set_conv_nth) + apply (drule_tac x=i in map_leD, simp) + apply (simp add: cleared[unfolded region_is_bytes'_def] not_byte size_of_def) + done + +lemma zero_ranges_are_zero_update[simp]: + "h_t_valid (hrs_htd hrs) c_guard (ptr :: 'a ptr) + \ typ_uinfo_t TYPE('a :: wf_type) \ typ_uinfo_t TYPE(word8) + \ zero_ranges_are_zero rs (hrs_mem_update (heap_update ptr v) hrs) + = zero_ranges_are_zero rs hrs" + apply (clarsimp simp: zero_ranges_are_zero_def hrs_mem_update + intro!: ball_cong[OF refl] conj_cong[OF refl]) + apply (drule(2) region_is_bytes_disjoint) + apply (simp add: heap_update_def heap_list_update_disjoint_same Int_commute) + done + lemma inj_tcb_ptr_to_ctcb_ptr [simp]: "inj tcb_ptr_to_ctcb_ptr" apply (rule injI) @@ -2197,6 +2226,58 @@ lemma page_table_at_rf_sr: by (clarsimp simp: rf_sr_def cstate_relation_def Let_def cpspace_relation_def page_table_at_carray_map_relation) +lemma gsUntypedZeroRanges_rf_sr: + "\ (start, end) \ gsUntypedZeroRanges s; (s, s') \ rf_sr \ + \ region_is_zero_bytes start (unat ((end + 1) - start)) s'" + apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def + zero_ranges_are_zero_def) + apply (drule(1) bspec) + apply clarsimp + done + +lemma ctes_of_untyped_zero_rf_sr: + "\ ctes_of s p = Some cte; (s, s') \ rf_sr; + untyped_ranges_zero' s; + untypedZeroRange (cteCap cte) = Some (start, end) \ + \ region_is_zero_bytes start (unat ((end + 1) - start)) s'" + apply (erule gsUntypedZeroRanges_rf_sr[rotated]) + apply (clarsimp simp: untyped_ranges_zero_inv_def) + apply (rule_tac a=p in ranI) + apply (simp add: map_comp_def cteCaps_of_def) + done + +lemma heap_list_is_zero_mono: + "heap_list_is_zero hmem p n \ n' \ n + \ heap_list_is_zero hmem p n'" + apply (induct n arbitrary: n' p) + apply simp + apply clarsimp + apply (case_tac n', simp_all) + done + +lemma heap_list_h_eq_better: + "\p. \ x \ {p..+q}; heap_list h q p = heap_list h' q p \ + \ h x = h' x" +proof (induct q) + case 0 thus ?case by simp +next + case (Suc n) thus ?case by (force dest: intvl_neq_start) +qed + +lemma heap_list_is_zero_mono2: + "heap_list_is_zero hmem p n + \ {p' ..+ n'} \ {p ..+ n} + \ heap_list_is_zero hmem p' n'" + using heap_list_h_eq2[where h'="\_. 0"] + heap_list_h_eq_better[where h'="\_. 0"] + apply (simp(no_asm_use) add: heap_list_rpbs) + apply blast + done + +lemma invs_urz[elim!]: + "invs' s \ untyped_ranges_zero' s" + by (clarsimp simp: invs'_def valid_state'_def) + end end diff --git a/proof/crefine/Schedule_C.thy b/proof/crefine/Schedule_C.thy index 004afefad..963b6f83b 100644 --- a/proof/crefine/Schedule_C.thy +++ b/proof/crefine/Schedule_C.thy @@ -818,11 +818,7 @@ proof - apply (frule_tac t=t in valid_queues_no_bitmap_objD) apply (blast intro: cons_set_intro) apply (simp add: lookupBitmapPriority_def) - apply normalise_obj_at' - apply (rule valid_queues_queued_runnable[simplified valid_queues_def]) - apply simp - apply (simp add: invs_no_cicd'_def) - apply (clarsimp simp add: inQ_def obj_at'_def) + apply (clarsimp simp: obj_at'_def st_tcb_at'_def) apply (fold lookupBitmapPriority_def) apply (drule invs_no_cicd'_queues) diff --git a/proof/crefine/StateRelation_C.thy b/proof/crefine/StateRelation_C.thy index fdf98d925..b8369ac73 100644 --- a/proof/crefine/StateRelation_C.thy +++ b/proof/crefine/StateRelation_C.thy @@ -672,6 +672,31 @@ where end +definition + region_is_bytes' :: "word32 \ nat \ heap_typ_desc \ bool" +where + "region_is_bytes' ptr sz htd \ \z\{ptr ..+ sz}. \ td. td \ typ_uinfo_t TYPE (word8) \ + (\n b. snd (htd z) n \ Some (td, b))" + +abbreviation + region_is_bytes :: "word32 \ nat \ globals myvars \ bool" +where + "region_is_bytes ptr sz s \ region_is_bytes' ptr sz (hrs_htd (t_hrs_' (globals s)))" + +abbreviation(input) + "heap_list_is_zero hp ptr n \ heap_list hp n ptr = replicate n 0" + +abbreviation + "region_is_zero_bytes ptr n x \ region_is_bytes ptr n x + \ heap_list_is_zero (hrs_mem (t_hrs_' (globals x))) ptr n" + +definition + zero_ranges_are_zero +where + "zero_ranges_are_zero rs hrs + = (\(start, end) \ rs. region_is_bytes' start (unat ((end + 1) - start)) (hrs_htd hrs) + \ heap_list_is_zero (hrs_mem hrs) start (unat ((end + 1) - start)))" + definition (in state_rel) cstate_relation :: "KernelStateData_H.kernel_state \ globals \ bool" where @@ -682,6 +707,7 @@ where cready_queues_relation (clift cheap) (ksReadyQueues_' cstate) (ksReadyQueues astate) \ + zero_ranges_are_zero (gsUntypedZeroRanges astate) cheap \ cbitmap_L1_relation (ksReadyQueuesL1Bitmap_' cstate) (ksReadyQueuesL1Bitmap astate) \ cbitmap_L2_relation (ksReadyQueuesL2Bitmap_' cstate) (ksReadyQueuesL2Bitmap astate) \ ksCurThread_' cstate = (tcb_ptr_to_ctcb_ptr (ksCurThread astate)) \ diff --git a/proof/crefine/StoreWord_C.thy b/proof/crefine/StoreWord_C.thy index a6fac4fbb..95e4baf76 100644 --- a/proof/crefine/StoreWord_C.thy +++ b/proof/crefine/StoreWord_C.thy @@ -624,6 +624,30 @@ proof (intro allI impI) apply (simp add: typ_heap_simps clift_heap_update_same) done + have subset: "{ptr..+ 2 ^ 2} \ {ptr && ~~ mask pageBits ..+ 2 ^ pageBits}" + apply (simp only: upto_intvl_eq al is_aligned_neg_mask2) + apply (cut_tac ptr="ptr && ~~ mask pageBits" and x="ptr && mask pageBits" + in aligned_range_offset_subset, rule is_aligned_neg_mask2) + apply (rule is_aligned_andI1[OF al]) + apply (simp add: pageBits_def) + apply (rule and_mask_less', simp add: pageBits_def) + apply (erule order_trans[rotated]) + apply (simp add: mask_out_sub_mask) + done + + hence zr: "\rs. zero_ranges_are_zero rs (hrs_mem_update (heap_update ?ptr w) (t_hrs_' (globals s))) + = zero_ranges_are_zero rs (t_hrs_' (globals s))" + using page + apply (clarsimp simp: zero_ranges_are_zero_def hrs_mem_update base_def + heap_update_def + intro!: ball_cong[OF refl] conj_cong[OF refl]) + apply (frule(1) region_is_bytes_disjoint[rotated 2, OF h_t_valid_clift]) + apply simp + apply (subst heap_list_update_disjoint_same, simp_all) + apply ((subst Int_commute)?, erule disjoint_subset2[rotated]) + apply (simp add: pageBits_def) + done + have cmap_relation_heap_cong: "\as cs cs' f rel. \ cmap_relation as cs f rel; cs = cs' \ \ cmap_relation as cs' f rel" by simp @@ -664,7 +688,7 @@ proof (intro allI impI) thus ?thesis using rf apply (simp add: rf_sr_def cstate_relation_def Let_def rl' tag_disj_via_td_name) apply (simp add: carch_state_relation_def cmachine_state_relation_def carch_globals_def) - apply (simp add: rl' tag_disj_via_td_name) + apply (simp add: rl' tag_disj_via_td_name zr) done qed @@ -908,6 +932,30 @@ proof (intro allI impI) apply (simp add: typ_heap_simps clift_heap_update_same) done + have subset: "{ptr..+ 2 ^ 2} \ {ptr && ~~ mask pageBits ..+ 2 ^ pageBits}" + apply (simp only: upto_intvl_eq al is_aligned_neg_mask2) + apply (cut_tac ptr="ptr && ~~ mask pageBits" and x="ptr && mask pageBits" + in aligned_range_offset_subset, rule is_aligned_neg_mask2) + apply (rule is_aligned_andI1[OF al]) + apply (simp add: pageBits_def) + apply (rule and_mask_less', simp add: pageBits_def) + apply (erule order_trans[rotated]) + apply (simp add: mask_out_sub_mask) + done + + hence zr: "\rs. zero_ranges_are_zero rs (hrs_mem_update (heap_update ?ptr w) (t_hrs_' (globals s))) + = zero_ranges_are_zero rs (t_hrs_' (globals s))" + using page + apply (clarsimp simp: zero_ranges_are_zero_def hrs_mem_update base_def + heap_update_def + intro!: ball_cong[OF refl] conj_cong[OF refl]) + apply (frule(1) region_is_bytes_disjoint[rotated 2, OF h_t_valid_clift]) + apply simp + apply (subst heap_list_update_disjoint_same, simp_all) + apply ((subst Int_commute)?, erule disjoint_subset2[rotated]) + apply (simp add: pageBits_def) + done + have cmap_relation_heap_cong: "\as cs cs' f rel. \ cmap_relation as cs f rel; cs = cs' \ \ cmap_relation as cs' f rel" by simp @@ -948,7 +996,7 @@ proof (intro allI impI) thus ?thesis using rf apply (simp add: rf_sr_def cstate_relation_def Let_def rl' tag_disj_via_td_name) apply (simp add: carch_state_relation_def cmachine_state_relation_def carch_globals_def) - apply (simp add: rl' tag_disj_via_td_name) + apply (simp add: rl' tag_disj_via_td_name zr) done qed diff --git a/proof/crefine/Syscall_C.thy b/proof/crefine/Syscall_C.thy index 0937d80f1..42e81d4a3 100644 --- a/proof/crefine/Syscall_C.thy +++ b/proof/crefine/Syscall_C.thy @@ -574,7 +574,22 @@ lemma threadGet_tcbFaultHandler_ccorres [corres]: apply (clarsimp simp: obj_at'_def ctcb_relation_def) done +lemma rf_sr_tcb_update_twice: + "h_t_valid (hrs_htd (hrs2 (globals s') (t_hrs_' (gs2 (globals s'))))) c_guard + (ptr (t_hrs_' (gs2 (globals s'))) (globals s')) + \ ((s, globals_update (\gs. t_hrs_'_update (\ths. + hrs_mem_update (heap_update (ptr ths gs :: tcb_C ptr) (v ths gs)) + (hrs_mem_update (heap_update (ptr ths gs) (v' ths gs)) (hrs2 gs ths))) (gs2 gs)) s') \ rf_sr) + = ((s, globals_update (\gs. t_hrs_'_update (\ths. + hrs_mem_update (heap_update (ptr ths gs) (v ths gs)) (hrs2 gs ths)) (gs2 gs)) s') \ rf_sr)" + by (simp add: rf_sr_def cstate_relation_def Let_def + cpspace_relation_def typ_heap_simps' + carch_state_relation_def + cmachine_state_relation_def) +lemma hrs_mem_update_use_hrs_mem: + "hrs_mem_update f = (\hrs. (hrs_mem_update $ (\_. f (hrs_mem hrs))) hrs)" + by (simp add: hrs_mem_update_def hrs_mem_def fun_eq_iff) lemma sendFaultIPC_ccorres: "ccorres (cfault_rel2 \ dc) (liftxf errstate id (K ()) ret__unsigned_long_') @@ -635,21 +650,21 @@ lemma sendFaultIPC_ccorres: (lookup_fault_lift(original_lookup_fault_' s)))}" in threadSet_ccorres_lemma4) apply vcg - apply (clarsimp simp: typ_heap_simps') + apply (clarsimp simp: typ_heap_simps' rf_sr_tcb_update_twice) + apply (intro conjI allI impI) - apply (simp add: typ_heap_simps rf_sr_def) + apply (simp add: typ_heap_simps' rf_sr_def) apply (rule rf_sr_tcb_update_no_queue2[unfolded rf_sr_def, simplified], - assumption+) - apply (simp add: typ_heap_simps)+ + assumption+, (simp add: typ_heap_simps')+) apply (rule ball_tcb_cte_casesI, simp+) apply (simp add: ctcb_relation_def cthread_state_relation_def ) apply (case_tac "tcbState tcb", simp+) apply (simp add: rf_sr_def) apply (rule rf_sr_tcb_update_no_queue2[unfolded rf_sr_def, simplified], - assumption+) - apply (simp add: typ_heap_simps)+ + assumption+, (simp add: typ_heap_simps' | simp only: hrs_mem_update_use_hrs_mem)+) apply (rule ball_tcb_cte_casesI, simp+) - apply (simp add: ctcb_relation_def cthread_state_relation_def ) + apply (clarsimp simp: typ_heap_simps') + apply (simp add: ctcb_relation_def cthread_state_relation_def) apply (rule conjI) apply (case_tac "tcbState tcb", simp+) apply (simp add: cfault_rel_def) diff --git a/proof/crefine/TcbQueue_C.thy b/proof/crefine/TcbQueue_C.thy index 9f4676a9d..942d3c10a 100644 --- a/proof/crefine/TcbQueue_C.thy +++ b/proof/crefine/TcbQueue_C.thy @@ -1108,9 +1108,8 @@ lemma cvariable_relation_upd_const: lemma rf_sr_tcb_update_no_queue: "\ (s, s') \ rf_sr; ko_at' tcb thread s; - (cslift t :: tcb_C typ_heap) = (cslift s')(tcb_ptr_to_ctcb_ptr thread \ ctcb); - cslift_all_but_tcb_C t s'; - hrs_htd (t_hrs_' (globals t)) = hrs_htd (t_hrs_' (globals s')); + t_hrs_' (globals t) = hrs_mem_update (heap_update + (tcb_ptr_to_ctcb_ptr thread) ctcb) (t_hrs_' (globals s')); tcbEPNext_C ctcb = tcbEPNext_C (the (cslift s' (tcb_ptr_to_ctcb_ptr thread))); tcbEPPrev_C ctcb = tcbEPPrev_C (the (cslift s' (tcb_ptr_to_ctcb_ptr thread))); tcbSchedNext_C ctcb = tcbSchedNext_C (the (cslift s' (tcb_ptr_to_ctcb_ptr thread))); @@ -1125,7 +1124,8 @@ lemma rf_sr_tcb_update_no_queue: apply (frule (1) cmap_relation_ko_atD) apply (erule obj_atE') apply (clarsimp simp: projectKOs) - apply (clarsimp simp: map_comp_update projectKO_opt_tcb cvariable_relation_upd_const) + apply (clarsimp simp: map_comp_update projectKO_opt_tcb cvariable_relation_upd_const + typ_heap_simps') apply (intro conjI) subgoal by (clarsimp simp: cmap_relation_def map_comp_update projectKO_opts_defs inj_eq) apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1]) @@ -1141,7 +1141,7 @@ lemma rf_sr_tcb_update_no_queue: apply (erule cready_queues_relation_not_queue_ptrs) subgoal by (clarsimp intro!: ext) subgoal by (clarsimp intro!: ext) - subgoal by (simp add: carch_state_relation_def) + subgoal by (simp add: carch_state_relation_def typ_heap_simps') by (simp add: cmachine_state_relation_def) lemma rf_sr_tcb_update_no_queue_helper: @@ -1161,9 +1161,8 @@ lemma tcb_queue_relation_not_in_q: lemma rf_sr_tcb_update_not_in_queue: "\ (s, s') \ rf_sr; ko_at' tcb thread s; - (cslift t :: tcb_C typ_heap) = (cslift s')(tcb_ptr_to_ctcb_ptr thread \ ctcb); - cslift_all_but_tcb_C t s'; - hrs_htd (t_hrs_' (globals t)) = hrs_htd (t_hrs_' (globals s')); + t_hrs_' (globals t) = hrs_mem_update (heap_update + (tcb_ptr_to_ctcb_ptr thread) ctcb) (t_hrs_' (globals s')); \ live' (KOTCB tcb); invs' s; (\x\ran tcb_cte_cases. (\(getF, setF). getF tcb' = getF tcb) x); ctcb_relation tcb' ctcb \ @@ -1175,7 +1174,8 @@ lemma rf_sr_tcb_update_not_in_queue: apply (frule (1) cmap_relation_ko_atD) apply (erule obj_atE') apply (clarsimp simp: projectKOs) - apply (clarsimp simp: map_comp_update projectKO_opt_tcb cvariable_relation_upd_const) + apply (clarsimp simp: map_comp_update projectKO_opt_tcb cvariable_relation_upd_const + typ_heap_simps') apply (subgoal_tac "\rf. \ ko_wp_at' (\ko. rf \ refs_of' ko) thread s") prefer 2 apply (auto simp: obj_at'_def ko_wp_at'_def)[1] @@ -1205,7 +1205,8 @@ lemma rf_sr_tcb_update_not_in_queue: apply (drule valid_queues_obj_at'D, clarsimp) apply (clarsimp simp: obj_at'_def projectKOs inQ_def) subgoal by simp - subgoal by (simp add: carch_state_relation_def carch_globals_def) + subgoal by (simp add: carch_state_relation_def carch_globals_def + typ_heap_simps') by (simp add: cmachine_state_relation_def) lemmas rf_sr_tcb_update_not_in_queue2 diff --git a/proof/crefine/Tcb_C.thy b/proof/crefine/Tcb_C.thy index 1044c0d63..c0c15bde9 100644 --- a/proof/crefine/Tcb_C.thy +++ b/proof/crefine/Tcb_C.thy @@ -428,8 +428,7 @@ lemma setPriority_ccorres: apply vcg apply clarsimp apply (erule(1) rf_sr_tcb_update_no_queue2, - (simp add: typ_heap_simps)+, simp_all)[1] - apply (subst heap_update_field_hrs | fastforce intro: typ_heap_simps)+ + (simp add: typ_heap_simps')+, simp_all?)[1] apply (rule ball_tcb_cte_casesI, simp+) apply (simp add: ctcb_relation_def cast_simps) apply (clarsimp simp: down_cast_same [symmetric] ucast_up_ucast is_up is_down) @@ -484,14 +483,12 @@ lemma setMCPriority_ccorres: apply vcg apply clarsimp apply (erule(1) rf_sr_tcb_update_no_queue2, - (simp add: typ_heap_simps)+, simp_all)[1] - apply (subst heap_update_field_hrs | fastforce intro: typ_heap_simps)+ + (simp add: typ_heap_simps')+)[1] apply (rule ball_tcb_cte_casesI, simp+) apply (simp add: ctcb_relation_def cast_simps) apply (clarsimp simp: down_cast_same [symmetric] ucast_up_ucast is_up is_down) - done - + lemma ccorres_subgoal_tailE: "\ ccorres rvr xf Q Q' hs (b ()) d; ccorres rvr xf Q Q' hs (b ()) d \ ccorres rvr xf P P' hs (a >>=E b) (c ;; d) \ @@ -655,10 +652,7 @@ lemma invokeTCB_ThreadControl_ccorres: apply vcg apply clarsimp apply (erule(1) rf_sr_tcb_update_no_queue2, - (simp add: typ_heap_simps)+, simp_all)[1] - apply (subst heap_update_field_hrs - | simp add: typ_heap_simps - | fastforce intro: typ_heap_simps)+ + (simp add: typ_heap_simps')+, simp_all?)[1] apply (rule ball_tcb_cte_casesI, simp+) apply (clarsimp simp: ctcb_relation_def option_to_0_def) apply (rule ceqv_refl) @@ -943,7 +937,8 @@ lemma setupReplyMaster_ccorres: apply (rule fst_setCTE[OF ctes_of_cte_at], assumption) apply (rule rev_bexI, assumption) apply (clarsimp simp: rf_sr_def cstate_relation_def - cpspace_relation_def Let_def) + cpspace_relation_def Let_def + typ_heap_simps') apply (subst setCTE_tcb_case, assumption+) apply (rule_tac r="s'" in KernelStateData_H.kernel_state.cases) apply clarsimp @@ -960,7 +955,7 @@ lemma setupReplyMaster_ccorres: apply (simp add: true_def mask_def to_bool_def) apply simp apply (simp add: cmachine_state_relation_def - h_t_valid_clift_Some_iff + typ_heap_simps' carch_state_relation_def carch_globals_def cvariable_array_map_const_add_map_option[where f="tcb_no_ctes_proj"]) apply (wp | simp)+ @@ -1392,7 +1387,6 @@ lemma asUser_sysargs_rel: apply (wp asUser_getMRs_rel hoare_valid_ipc_buffer_ptr_typ_at'|simp)+ done - lemma invokeTCB_WriteRegisters_ccorres[where S=UNIV]: notes static_imp_wp [wp] shows @@ -1455,6 +1449,7 @@ lemma invokeTCB_WriteRegisters_ccorres[where S=UNIV]: frame_gp_registers_convs word_less_nat_alt) apply (simp add: unat_of_nat32 word_bits_def) apply arith + apply clarsimp apply (vcg exspec=setRegister_modifies exspec=getSyscallArg_modifies exspec=sanitiseRegister_modifies) apply clarsimp @@ -1484,6 +1479,7 @@ lemma invokeTCB_WriteRegisters_ccorres[where S=UNIV]: word_less_nat_alt word_bits_def less_diff_conv) apply (simp add: unat_word_ariths cong: conj_cong) + apply clarsimp apply (vcg exspec=setRegister_modifies exspec=getSyscallArg_modifies exspec=sanitiseRegister_modifies) apply clarsimp @@ -3545,7 +3541,8 @@ lemma bindNotification_ccorres: apply (clarsimp simp: setNotification_def split_def) apply (rule bexI [OF _ setObject_eq]) apply (simp add: rf_sr_def cstate_relation_def Let_def init_def - cpspace_relation_def update_ntfn_map_tos) + cpspace_relation_def update_ntfn_map_tos + typ_heap_simps') apply (elim conjE) apply (intro conjI) -- "tcb relation" @@ -3556,7 +3553,7 @@ lemma bindNotification_ccorres: apply ((clarsimp simp: option_to_ctcb_ptr_def tcb_ptr_to_ctcb_ptr_def ctcb_offset_def obj_at'_def projectKOs objBitsKO_def bindNTFN_alignment_junk)+)[4] - apply (simp add: carch_state_relation_def) + apply (simp add: carch_state_relation_def typ_heap_simps') apply (simp add: cmachine_state_relation_def) apply (simp add: h_t_valid_clift_Some_iff) apply (simp add: objBits_simps) @@ -3568,8 +3565,8 @@ lemma bindNotification_ccorres: apply (rule_tac P'=\ and P=\ in threadSet_ccorres_lemma3[unfolded dc_def]) apply vcg apply simp - apply (erule (1) rf_sr_tcb_update_no_queue2) - apply (simp add: typ_heap_simps)+ + apply (erule (1) rf_sr_tcb_update_no_queue2, + (simp add: typ_heap_simps')+, simp_all?)[1] apply (simp add: ctcb_relation_def option_to_ptr_def option_to_0_def) apply simp apply (wp get_ntfn_ko'| simp add: guard_is_UNIV_def)+ diff --git a/proof/crefine/VSpace_C.thy b/proof/crefine/VSpace_C.thy index e5eb5d648..64faa35d5 100644 --- a/proof/crefine/VSpace_C.thy +++ b/proof/crefine/VSpace_C.thy @@ -712,8 +712,8 @@ lemma generic_frame_cap_ptr_set_capFMappedAddress_spec: {t. (\cte' cap'. generic_frame_cap_set_capFMappedAddress_CL (cap_lift (the (cslift s \<^bsup>s\<^esup>cap_ptr))) \<^bsup>s\<^esup>asid \<^bsup>s\<^esup>addr = Some cap' \ cte_lift cte' = option_map (cap_CL_update (K cap')) (cte_lift (the (cslift s cte_slot))) \ - cslift t = cslift s(cte_slot \ cte')) \ cslift_all_but_cte_C t s - \ (hrs_htd \<^bsup>t\<^esup>t_hrs) = (hrs_htd \<^bsup>s\<^esup>t_hrs)}" + t_hrs_' (globals t) = hrs_mem_update (heap_update cte_slot cte') + (t_hrs_' (globals s)))}" apply vcg apply (clarsimp simp: typ_heap_simps) apply (subgoal_tac "cap_lift ret__struct_cap_C \ None") @@ -721,9 +721,8 @@ lemma generic_frame_cap_ptr_set_capFMappedAddress_spec: apply (clarsimp simp: generic_frame_cap_set_capFMappedAddress_CL_def split: cap_CL.splits) apply (clarsimp simp: clift_ptr_safe2 typ_heap_simps) apply (rule_tac x="cte_C.cap_C_update (\_. ret__struct_cap_C) y" in exI) - apply simp apply (case_tac y) - apply (clarsimp simp: cte_lift_def) + apply (clarsimp simp: cte_lift_def typ_heap_simps') done lemma lookupPDSlot_spec: diff --git a/proof/drefine/Untyped_DR.thy b/proof/drefine/Untyped_DR.thy index 417d2228c..7a64e2320 100644 --- a/proof/drefine/Untyped_DR.thy +++ b/proof/drefine/Untyped_DR.thy @@ -1256,6 +1256,33 @@ lemma mapME_x_append: apply (simp add: mapME_x_Cons bindE_assoc) done +lemma hd_filter_eq: + "P (hd xs) \ hd (filter P xs) = hd xs" + by (cases xs, simp_all) + +lemma free_range_of_untyped_subset'[rule_format]: + assumes al: "is_aligned ptr sz" + and sz: "sz < size (ptr)" + shows "a < a' \ a < 2 ^ sz \ free_range_of_untyped a' sz ptr \ free_range_of_untyped a sz ptr" +proof - + note no_ov = is_aligned_no_overflow'[OF al] + note mono = word_plus_mono_right[OF _ no_ov, simplified field_simps] + note sz_len = sz[unfolded word_size] +(* + note sz_simp[simp] = unat_2p_sub_1[OF sz_len] unat_power_lower[OF sz_len] +*) + show ?thesis using sz + apply (intro impI psubsetI) + apply (rule free_range_of_untyped_subseteq', simp_all add: al) + apply (clarsimp simp:free_range_of_untyped_def) + apply (strengthen mono word_of_nat_le)+ + apply (simp add: sz_len unat_2p_sub_1) + apply (intro conjI impI; (clarsimp dest!: le_p2_minus_1)?) + apply (drule word_unat.Abs_eqD, simp_all add: unats_def word_size + less_trans[OF _ power_strict_increasing[OF sz_len]]) + done +qed + lemma reset_untyped_cap_corres: "dcorres (dc \ dc) \ (invs and valid_etcbs and ct_active and cte_wp_at (\cap. is_untyped_cap cap \ free_index_of cap = idx) cref @@ -1271,14 +1298,28 @@ lemma reset_untyped_cap_corres: apply (rule_tac F="is_untyped_cap capa \ cap_aligned capa \ bits_of capa > 2 \ free_index_of capa \ 2 ^ bits_of capa" in corres_gen_asm2) + apply (simp add: whenE_def if_flip split del: split_if) + apply (rule corres_if) + apply (clarsimp simp: is_cap_simps free_range_of_untyped_def + cap_aligned_def free_index_of_def) + apply (simp add: word_unat.Rep_inject[symmetric]) + apply (subst unat_of_nat_eq, erule order_le_less_trans, + rule power_strict_increasing, simp_all add: word_bits_def bits_of_def)[1] + apply (rule corres_trivial, rule corres_returnOk, simp) + apply (rule_tac F="free_index_of capa \ 0" in corres_gen_asm2) apply (rule corres_split_nor) prefer 2 apply (rule delete_objects_dcorres) apply (clarsimp simp: is_cap_simps bits_of_def) apply (rule corres_if_rhs) - apply (rule_tac x=0 in select_pick_corres_asm) - apply simp - apply (simp add: mapME_x_Nil) + apply (rule_tac x="[cap_objects (transform_cap capa)]" in select_pick_corres_asm) + apply (clarsimp simp: is_cap_simps cap_aligned_def free_index_of_def) + apply (rule order_less_le_trans, rule free_range_of_untyped_subset'[where a=0], + simp_all)[1] + apply (simp add: word_size word_bits_def) + apply (simp add: free_range_of_untyped_def) + + apply (simp add: mapME_x_Nil mapME_x_Cons liftE_def bind_assoc) apply (rule corres_add_noop_lhs) apply (rule corres_split_nor) prefer 2 @@ -1288,96 +1329,80 @@ lemma reset_untyped_cap_corres: apply (clarsimp simp: is_cap_simps bits_of_def free_range_of_untyped_def) apply (clarsimp simp: is_cap_simps cap_aligned_def bits_of_def) apply (rule corres_trivial, simp) - apply (rule set_cap_corres) - apply (clarsimp simp: is_cap_simps bits_of_def free_range_of_untyped_def) - apply simp - apply (wp | simp add: unless_def)+ - apply (rule_tac x="length xs" for xs in select_pick_corres_asm) - apply simp - apply (simp add: mapME_x_upt_length_ys) - apply (rule corres_dummy_returnOk_r) - apply (rule corres_splitEE) - prefer 2 - apply (rule_tac P="\\" and P'="\_. valid_objs + apply (rule corres_split_nor) + apply (rule corres_alternate1) + apply (rule corres_trivial, simp add: returnOk_def) + apply (rule set_cap_corres) + apply (clarsimp simp: is_cap_simps bits_of_def free_range_of_untyped_def) + apply simp + apply (wp | simp add: unless_def)+ + apply (rule_tac F="\ (is_device_untyped_cap capa \ bits_of capa < reset_chunk_bits) + \ (\s. s \ capa)" + in corres_gen_asm2) + apply (rule_tac x="map (\i. free_range_of_untyped (i * 2 ^ reset_chunk_bits) + (bits_of capa) (obj_ref_of capa)) xs" for xs + in select_pick_corres_asm) + prefer 2 + apply (simp add: mapME_x_map_simp o_def) + apply (rule_tac P="\\" and P'="\_. valid_objs and valid_etcbs and pspace_no_overlap (untyped_range capa) and valid_idle and (\s. idle_thread s \ fst cref) and cte_wp_at (\cp. \idx. cp = free_index_update (\_. idx) capa) cref" in mapME_x_corres_same_xs[OF _ _ _ refl]) - apply (rule corres_guard_imp) - apply (rule corres_add_noop_lhs) - apply (rule corres_split_nor[OF _ clearMemory_corres_noop[OF refl]]) - apply (rule_tac x="available_range (transform_cap - (free_index_update (\_. x * 2 ^ reset_chunk_bits) capa))" - in select_pick_corres_asm) - apply (clarsimp simp: is_cap_simps bits_of_def free_index_of_def) - apply (strengthen free_range_of_untyped_subseteq') - apply (strengthen order_trans[OF free_range_of_untyped_subseteq'[where a=0]]) - apply (simp add: cap_aligned_def free_range_of_untyped_def - word_bits_def word_size) - apply (rule corres_split[OF _ set_cap_corres]) - apply (subst alternative_com) - apply (rule throw_or_return_preemption_corres[where P=\ and P'=\]) - apply (clarsimp simp: is_cap_simps bits_of_def) - apply simp - apply wp - apply (clarsimp simp add: is_cap_simps cap_aligned_def bits_of_def - aligned_add_aligned is_aligned_shiftl) - apply (simp add: reset_chunk_bits_def) - apply (simp add: reset_chunk_bits_def word_bits_def) - apply (wp hoare_vcg_all_lift hoare_vcg_const_imp_lift - preemption_point_inv' set_cap_no_overlap - update_untyped_cap_valid_objs set_cap_idle - | simp)+ - apply (clarsimp simp: is_cap_simps bits_of_def cap_aligned_def) - apply (erule pspace_no_overlap_subset) - apply (rule aligned_range_offset_subset, simp_all add: is_aligned_shiftl)[1] - apply (rule shiftl_less_t2n[OF word_of_nat_less]) - apply simp - apply (simp add: word_bits_def) - apply (rule hoare_pre, wp) + apply (rule corres_guard_imp) + apply (rule corres_add_noop_lhs) + apply (rule corres_split_nor[OF _ clearMemory_corres_noop[OF refl]]) + apply (rule corres_split[OF _ set_cap_corres]) + apply (subst alternative_com) + apply (rule throw_or_return_preemption_corres[where P=\ and P'=\]) + apply (clarsimp simp: is_cap_simps bits_of_def) + apply simp + apply wp + apply (clarsimp simp add: is_cap_simps cap_aligned_def bits_of_def + aligned_add_aligned is_aligned_shiftl) + apply (simp add: reset_chunk_bits_def) + apply (simp add: reset_chunk_bits_def word_bits_def) + apply (wp hoare_vcg_all_lift hoare_vcg_const_imp_lift + preemption_point_inv' set_cap_no_overlap + update_untyped_cap_valid_objs set_cap_idle + | simp)+ + apply (clarsimp simp: is_cap_simps bits_of_def cap_aligned_def) + apply (erule pspace_no_overlap_subset) + apply (rule aligned_range_offset_subset, simp_all add: is_aligned_shiftl)[1] + apply (rule shiftl_less_t2n[OF word_of_nat_less]) apply simp - apply (rule hoare_pre) - apply (wp hoare_vcg_all_lift hoare_vcg_const_imp_lift - update_untyped_cap_valid_objs set_cap_no_overlap - set_cap_idle preemption_point_inv' - set_cap_cte_wp_at - | simp)+ - apply (clarsimp simp: cte_wp_at_caps_of_state exI - is_cap_simps bits_of_def) - apply (frule(1) cte_wp_at_valid_objs_valid_cap[OF caps_of_state_cteD]) - apply (clarsimp simp: valid_cap_simps cap_aligned_def - valid_untyped_pspace_no_overlap - free_index_of_def) - apply (simp add: returnOk_liftE) - apply (rule monadic_rewrite_corres_rhs, - rule_tac cap="free_index_update (\_. 0) capa" - and p=cref in monadic_set_cap_id2) - apply (rule set_cap_corres[unfolded dc_def]) - apply (clarsimp simp: is_cap_simps free_range_of_untyped_def) + apply (simp add: word_bits_def) + apply (rule hoare_pre, wp) apply simp - apply wp - apply (rule_tac P="valid_etcbs - and valid_idle and (\s. idle_thread s \ fst cref) - and cte_wp_at (op = capa) cref" - in hoare_trivE_R) - apply (case_tac "free_index_of capa = 0") - apply (simp add: mapME_x_Nil) - apply wp - apply (clarsimp simp: cte_wp_at_caps_of_state is_cap_simps + apply (rule hoare_pre) + apply (wp hoare_vcg_all_lift hoare_vcg_const_imp_lift + update_untyped_cap_valid_objs set_cap_no_overlap + set_cap_idle preemption_point_inv' + set_cap_cte_wp_at + | simp)+ + apply (clarsimp simp: cte_wp_at_caps_of_state exI + is_cap_simps bits_of_def) + apply (frule(1) cte_wp_at_valid_objs_valid_cap[OF caps_of_state_cteD]) + apply (clarsimp simp: valid_cap_simps cap_aligned_def + valid_untyped_pspace_no_overlap free_index_of_def) - apply (subst upt_conv_Cons) - apply simp - apply (simp add: mapME_x_append mapME_x_Cons mapME_x_Nil bindE_assoc) - apply (rule hoare_pre) - apply (wp preemption_point_inv' set_cap_idle - set_cap_cte_wp_at | simp - | (wp_once mapME_x_inv_wp, rule hoare_pre))+ - apply (clarsimp simp: is_cap_simps bits_of_def cte_wp_at_caps_of_state) - apply (simp del: hoare_TrueI) + + apply (clarsimp simp: is_cap_simps bits_of_def) + apply (strengthen order_trans[OF free_range_of_untyped_subseteq'[where a=0]] + free_range_of_untyped_subset') + apply (clarsimp simp: conj_comms) + apply (clarsimp simp: filter_empty_conv bexI[where x=0] + last_rev hd_map hd_filter_eq + is_cap_simps bits_of_def cap_aligned_def + word_bits_def word_size exI + rev_map[symmetric]) + apply (clarsimp simp: free_index_of_def free_range_of_untyped_def) apply wp + apply clarsimp apply (wp add: hoare_vcg_const_imp_lift get_cap_wp | simp - | strengthen invs_valid_objs invs_valid_idle)+ + | strengthen invs_valid_objs invs_valid_idle + | rule impI)+ apply (clarsimp simp: cte_wp_at_caps_of_state descendants_range_def2) apply (cases cref, simp)