(* * Copyright 2014, General Dynamics C4 Systems * * This software may be distributed and modified according to the terms of * the GNU General Public License version 2. Note that NO WARRANTY is provided. * See "LICENSE_GPLv2.txt" for details. * * @TAG(GD_GPL) *) theory TcbAcc_C imports "../../lib/clib/Ctac" begin context kernel begin lemma ccorres_pre_threadGet: assumes cc: "\rv. ccorres r xf (P rv) (P' rv) hs (g rv) c" shows "ccorres r xf (\s. \tcb. ko_at' tcb p s \ P (f tcb) s) ({s'. \tcb ctcb. cslift s' (tcb_ptr_to_ctcb_ptr p) = Some ctcb \ ctcb_relation tcb ctcb \ s' \ P' (f tcb)}) hs (threadGet f p >>= (\rv. g rv)) c" apply (rule ccorres_guard_imp) apply (rule ccorres_symb_exec_l) defer apply wp[1] apply (rule tg_sp') apply simp apply assumption defer apply (rule ccorres_guard_imp) apply (rule cc) apply clarsimp apply (frule obj_at_ko_at') apply clarsimp apply assumption apply clarsimp apply (frule (1) obj_at_cslift_tcb) apply clarsimp done lemma threadGet_eq: "ko_at' tcb thread s \ (f tcb, s) \ fst (threadGet f thread s)" unfolding threadGet_def apply (simp add: liftM_def in_monad) apply (rule exI [where x = tcb]) apply simp apply (subst getObject_eq) apply simp apply (simp add: objBits_simps) apply assumption apply simp done lemma get_tsType_ccorres [corres]: "ccorres (\r r'. r' = thread_state_to_tsType r) ret__unsigned_' (tcb_at' thread) (UNIV \ {s. thread_state_ptr_' s = Ptr &(tcb_ptr_to_ctcb_ptr thread\[''tcbState_C''])}) [] (getThreadState thread) (Call thread_state_ptr_get_tsType_'proc)" unfolding getThreadState_def apply (rule ccorres_from_spec_modifies) apply (rule thread_state_ptr_get_tsType_spec) apply (rule thread_state_ptr_get_tsType_modifies) apply simp apply (frule (1) obj_at_cslift_tcb) apply (clarsimp simp: typ_heap_simps) apply (frule (1) obj_at_cslift_tcb) apply (clarsimp simp: typ_heap_simps) apply (rule bexI [rotated, OF threadGet_eq], assumption) apply simp apply (erule ctcb_relation_thread_state_to_tsType) done lemma threadGet_obj_at2: "\\\ threadGet f thread \\v. obj_at' (\t. f t = v) thread\" apply (rule hoare_post_imp) prefer 2 apply (rule tg_sp') apply simp done lemma register_from_H_less: "register_from_H hr < 18" by (cases hr, simp_all add: "StrictC'_register_defs") lemma register_from_H_sless: "register_from_H hr ({s. thread_' s = tcb_ptr_to_ctcb_ptr thread} \ {s. reg_' s = register_from_H reg}) [] (asUser thread (getRegister reg)) (Call getRegister_'proc)" apply (unfold asUser_def) apply (rule ccorres_guard_imp) apply (rule ccorres_symb_exec_l [where Q="\u. obj_at' (\t. tcbContext t = u) thread" and Q'="\rv. {s. thread_' s = tcb_ptr_to_ctcb_ptr thread} \ {s. reg_' s = register_from_H reg}"]) apply (rule ccorres_from_vcg) apply (rule allI, rule conseqPre) apply vcg apply clarsimp apply (drule (1) obj_at_cslift_tcb) apply (clarsimp simp: typ_heap_simps register_from_H_less register_from_H_sless) apply (clarsimp simp: getRegister_def typ_heap_simps) apply (rule_tac x = "(tcbContext ko reg, \)" in bexI [rotated]) apply (simp add: in_monad' asUser_def select_f_def split_def) apply (subst arg_cong2 [where f = "op \"]) defer apply (rule refl) apply (erule threadSet_eq) apply (clarsimp simp: ctcb_relation_def ccontext_relation_def) apply (wp threadGet_obj_at2) apply simp apply simp apply (erule obj_atE') apply (clarsimp simp: projectKOs ) apply (subst fun_upd_idem) apply (case_tac ko) apply clarsimp apply simp done lemma getRestartPC_ccorres [corres]: "ccorres (op =) ret__unsigned_long_' \ \\thread = tcb_ptr_to_ctcb_ptr thread\ [] (asUser thread getRestartPC) (Call getRestartPC_'proc)" unfolding getRestartPC_def apply (cinit') apply (rule ccorres_add_return2, ctac) apply (rule ccorres_return_C, simp+)[1] apply wp apply vcg apply (simp add: scast_id) done lemma threadSet_corres_lemma: assumes spec: "\s. \\ \s. P s\ Call f {t. Q s t}" and mod: "modifies_heap_spec f" and rl: "\\ x t ko. \(\, x) \ rf_sr; Q x t; x \ P'; ko_at' ko thread \\ \ (\\ksPSpace := ksPSpace \(thread \ KOTCB (g ko))\, t\globals := globals x\t_hrs_' := t_hrs_' (globals t)\\) \ rf_sr" and g: "\s x. \tcb_at' thread s; x \ P'; (s, x) \ rf_sr\ \ P x" shows "ccorres dc xfdc (tcb_at' thread) P' [] (threadSet g thread) (Call f)" apply (rule ccorres_Call_call_for_vcg) apply (rule ccorres_from_vcg) apply (rule allI, rule conseqPre) apply (rule HoarePartial.ProcModifyReturnNoAbr [where return' = "\s t. t\ globals := globals s\t_hrs_' := t_hrs_' (globals t) \\"]) apply (rule HoarePartial.ProcSpecNoAbrupt [OF _ _ spec]) apply (rule subset_refl) apply vcg prefer 2 apply (rule mod) apply (clarsimp simp: mex_def meq_def) apply clarsimp apply (rule conjI) apply (erule (2) g) apply clarsimp apply (frule obj_at_ko_at') apply clarsimp apply (rule bexI [rotated]) apply (erule threadSet_eq) apply simp apply (rule_tac x1 = "t\globals := globals x\t_hrs_' := t_hrs_' (globals t)\\" in iffD1 [OF rf_sr_upd], simp_all)[1] apply (erule (3) rl) done lemma threadSet_ccorres_lemma4: "\ \s tcb. \ \ (Q s tcb) c {s'. (s \ksPSpace := ksPSpace s(thread \ injectKOS (F tcb))\, s') \ rf_sr}; \s s' tcb tcb'. \ (s, s') \ rf_sr; P tcb; ko_at' tcb thread s; cslift s' (tcb_ptr_to_ctcb_ptr thread) = Some tcb'; ctcb_relation tcb tcb'; P' s ; s' \ R\ \ s' \ Q s tcb \ \ ccorres dc xfdc (obj_at' (P :: tcb \ bool) thread and P') R hs (threadSet F thread) c" apply (rule ccorres_from_vcg) apply (rule allI) apply (case_tac "obj_at' P thread \") apply (drule obj_at_ko_at', clarsimp) apply (rule conseqPre, rule conseqPost) apply assumption apply clarsimp apply (rule rev_bexI, rule threadSet_eq) apply assumption apply simp apply simp apply clarsimp apply (drule(1) obj_at_cslift_tcb, clarsimp) apply simp apply (rule hoare_complete') apply (simp add: cnvalid_def nvalid_def) (* pretty *) done lemmas threadSet_ccorres_lemma3 = threadSet_ccorres_lemma4[where R=UNIV] lemmas threadSet_ccorres_lemma2 = threadSet_ccorres_lemma3[where P'=\] lemma is_aligned_tcb_ptr_to_ctcb_ptr: "obj_at' (P :: tcb \ bool) p s \ is_aligned (ptr_val (tcb_ptr_to_ctcb_ptr p)) 8" apply (clarsimp simp: obj_at'_def objBits_simps projectKOs tcb_ptr_to_ctcb_ptr_def ctcb_offset_def) apply (erule aligned_add_aligned, simp_all add: word_bits_conv) apply (simp add: is_aligned_def) done lemma sanitiseRegister_spec: "\s v r. \ \ ({s} \ \\v___unsigned_long = v\ \ \\reg = register_from_H r\) Call sanitiseRegister_'proc \\ret__unsigned_long = sanitiseRegister r v\" apply vcg apply (auto simp: C_register_defs ARM_H.sanitiseRegister_def ARM.sanitiseRegister_def word_0_sle_from_less split: register.split) done end end