Merge branch 'master' into multi_arch_refactor

This commit is contained in:
Michael Norrish 2016-05-12 15:35:25 +10:00
commit 41f12cf570
130 changed files with 1902 additions and 1734 deletions

View File

@ -17,9 +17,9 @@
format.
-->
<testsuite depends="isabelle" cpu-timeout="600">
<testsuite depends="isabelle" cpu-timeout="3600">
<test name="CamkesAdlSpec">make CamkesAdlSpec</test>
<test name="CamkesCdlRefine" depends="CamkesAdlSpec DPolicy">make CamkesCdlRefine</test>
<test name="CamkesGlueSpec">make CamkesGlueSpec</test>
<test name="CamkesGlueProofs" depends="AutoCorres" cpu-timeout="3600">make CamkesGlueProofs</test>
<test name="CamkesGlueProofs" depends="AutoCorres">make CamkesGlueProofs</test>
</testsuite>

View File

@ -71,7 +71,7 @@ fun proof_body_descend' (_,("",_,body)) = fold (append o proof_body_descend') (t
fun used_facts thm = fold (append o proof_body_descend') (thms_of (Thm.proof_body_of thm)) []
fun raw_primitive_text f = Method.Basic (fn _ => (Method.METHOD (K (fn thm => Seq.single (f thm)))))
fun raw_primitive_text f = Method.Basic (fn _ => ((K (fn (ctxt, thm) => Seq.make_results (Seq.single (ctxt, f thm))))))
(*Find local facts from new hyps*)

77
lib/BitFieldProofsLib.thy Normal file
View File

@ -0,0 +1,77 @@
theory BitFieldProofsLib
imports TypHeapLib
begin
lemmas guard_simps =
word_sle_def word_sless_def scast_id
lemmas mask_shift_simps =
ucast_def shift_over_ao_dists word_bw_assocs
word_size multi_shift_simps mask_def
word_ao_dist NOT_eq scast_id
word_and_max_word max_word_def
lemmas sep_heap_simps =
sep_app_def hrs_mem_update_def
hrs_htd_def split_def
lemma tag_eq_to_tag_masked_eq:
"tag == v ==> tag && m = v && m"
by simp
lemma clift_heap_update_footprint:
"\<forall>p' :: 'a ptr. hrs_htd hp,g \<Turnstile>\<^sub>t p' \<longrightarrow> s_footprint p \<inter> s_footprint p' = {}
\<Longrightarrow> (lift_t g (hrs_mem_update (heap_update p (v :: 'b :: wf_type)) hp)
:: ('a :: mem_type) typ_heap) = lift_t g hp"
apply (cases hp, simp add: lift_t_if fun_eq_iff hrs_mem_update_def hrs_htd_def)
apply clarsimp
apply (drule spec, drule(1) mp)
apply (simp add: h_val_def heap_update_def)
apply (subst heap_list_update_disjoint_same, simp_all)
apply (simp add: set_eq_iff s_footprint_intvl[symmetric])
done
lemma s_footprint_field_lvalue_disj_helper:
"field_lookup (typ_info_t TYPE('b::mem_type)) f 0 = Some (ti,b)
\<Longrightarrow> export_uinfo ti = typ_uinfo_t TYPE('a)
\<Longrightarrow> (\<forall>x. P x \<longrightarrow> s_footprint (p::'b ptr) \<inter> S x = {})
\<Longrightarrow> (\<forall>x. P x \<longrightarrow> s_footprint ((Ptr &(p\<rightarrow>f))::'a::mem_type ptr) \<inter> S x = {})"
apply (drule field_lookup_export_uinfo_Some)
apply simp
apply (drule field_ti_s_sub_typ)
apply blast
done
lemma s_footprint_distinct_helper:
"h_t_valid htd g (p :: 'a ptr) \<Longrightarrow> typ_uinfo_t TYPE('a :: c_type) \<bottom>\<^sub>t typ_uinfo_t TYPE('b :: c_type)
\<Longrightarrow> (\<forall>q :: 'b ptr. h_t_valid htd g' q \<longrightarrow> s_footprint p \<inter> s_footprint q = {})"
apply clarsimp
apply (drule(1) h_t_valid_neq_disjoint)
apply (clarsimp simp: sub_typ_proper_def tag_disj_def preorder_class.less_imp_le)
apply (clarsimp simp: field_of_t_def)
apply (drule field_of_sub)
apply (simp add: tag_disj_def)
apply (simp add: set_eq_iff )
apply (blast dest: s_footprintD)
done
text {* Use these handy rules to prove that clift doesn't change
over various updates. *}
method prove_one_bf_clift_invariance
= (intro clift_heap_update_footprint[THEN trans]
s_footprint_field_lvalue_disj_helper
s_footprint_distinct_helper,
simp_all only: hrs_htd_mem_update,
tactic {* distinct_subgoals_tac *},
(auto simp: typ_uinfo_t_def[symmetric] tag_disj_via_td_name ntbs
elim: h_t_valid_clift)+)[1]
text {* Select points in the goal where we need to prove that
clift doesn't change over updates, and prove them with the above. *}
method prove_bf_clift_invariance
= ((subst Eq_TrueI[where P="lift_t g h = lift_t g h'" for g h h'],
prove_one_bf_clift_invariance)+, simp)
end

View File

@ -26,7 +26,7 @@ abbreviation (input) "flip \<equiv> swp"
we can turn off this abbreviation later (with no_notation)
to avoid syntax conflicts *)
abbreviation(input) bind_drop :: "('a, 'c) nondet_monad \<Rightarrow> ('a, 'b) nondet_monad
\<Rightarrow> ('a, 'b) nondet_monad" (infixl ">>" 60)
\<Rightarrow> ('a, 'b) nondet_monad" (infixl ">>_" 60)
where "bind_drop \<equiv> (\<lambda>x y. bind x (K_bind y))"
lemma bind_drop_test:

View File

@ -14,6 +14,7 @@ imports
"../sep_algebra/Sep_Algebra_L4v"
begin
(* FIXME: needs cleanup *)
lemma hoare_eq_pre: " \<lbrakk> \<And>s. P s = G s; \<lbrace>P\<rbrace> f \<lbrace>Q\<rbrace>\<rbrakk> \<Longrightarrow> \<lbrace>G\<rbrace> f \<lbrace>Q\<rbrace>"
by (erule hoare_pre_subst [rotated], auto)
@ -22,17 +23,16 @@ lemma hoare_eq_preE: " \<lbrakk> \<And>s. P s = G s; \<lbrace>P\<rbrace> f \<lb
apply (rule hoare_weaken_preE)
apply (assumption)
apply (clarsimp)
done
done
lemma hoare_eq_preE_R: " \<lbrakk> \<And>s. P s = G s; \<lbrace>P\<rbrace> f \<lbrace>Q\<rbrace>,-\<rbrakk> \<Longrightarrow> \<lbrace>G\<rbrace> f \<lbrace>Q\<rbrace>,-"
apply (clarsimp simp: validE_R_def)
apply (erule hoare_eq_preE)
apply (clarsimp)
done
done
lemma hoare_eq_post: " \<lbrakk> \<And>rv s. Q rv s = G rv s; \<lbrace>P\<rbrace> f \<lbrace>Q\<rbrace>\<rbrakk> \<Longrightarrow> \<lbrace>P\<rbrace> f \<lbrace>G\<rbrace>"
apply (rule hoare_strengthen_post, assumption, clarsimp)
done
by (rule hoare_strengthen_post, assumption, clarsimp)
lemma hoare_eq_postE: " \<lbrakk> \<And>rv s. Q rv s = G rv s; \<lbrace>P\<rbrace> f \<lbrace>Q\<rbrace>, \<lbrace>E\<rbrace>\<rbrakk> \<Longrightarrow> \<lbrace>P\<rbrace> f \<lbrace>G\<rbrace>, \<lbrace>E\<rbrace>"
by (metis (full_types) hoare_post_impErr')
@ -370,30 +370,24 @@ done
ML {*
fun J f x = f x
handle _ => x (* FIXME *)
handle _ => x (* FIXME! exceptions *)
fun sep_wp thms ctxt =
let val thms' = map (sep_wandise_helper ctxt |> J) thms;
val wp = WeakestPre.apply_once_tac false ctxt thms' (Unsynchronized.ref [] : thm list Unsynchronized.ref)
val sep_impi = (REPEAT_ALL_NEW (sep_match_trivial_tac ctxt)) THEN' assume_tac ctxt
val schemsolve = sep_rule_tac (eresolve0_tac [@{thm boxsolve}]) ctxt
val hoare_post = (resolve0_tac [(rotate_prems ~1 @{thm hoare_strengthen_post})])
in K wp THEN' (TRY o sep_flatten ctxt) THEN' (TRY o (hoare_post THEN' (schemsolve ORELSE' sep_impi))) THEN'
(TRY o (sep_match_trivial_tac ctxt |> REPEAT_ALL_NEW)) THEN'
(TRY o sep_flatten ctxt)
end ;
fun sep_wp_parse xs =
(Attrib.thms >> curry (fn (thms, ctxt) =>
Method.SIMPLE_METHOD' (
sep_wp thms ctxt
))) xs
let
val thms' = map (sep_wandise_helper ctxt |> J) thms;
val wp = WeakestPre.apply_once_tac false ctxt thms' (Unsynchronized.ref [] : thm list Unsynchronized.ref)
val sep_impi = (REPEAT_ALL_NEW (sep_match_trivial_tac ctxt)) THEN' assume_tac ctxt
val schemsolve = sep_rule_tac (eresolve0_tac [@{thm boxsolve}]) ctxt
val hoare_post = (resolve0_tac [(rotate_prems ~1 @{thm hoare_strengthen_post})])
in
K wp THEN' (TRY o sep_flatten ctxt) THEN' (TRY o (hoare_post THEN' (schemsolve ORELSE' sep_impi))) THEN'
(TRY o (sep_match_trivial_tac ctxt |> REPEAT_ALL_NEW)) THEN'
(TRY o sep_flatten ctxt)
end
*}
method_setup sep_wp = {*
sep_wp_parse
Attrib.thms >> (fn thms => fn ctxt => Method.SIMPLE_METHOD' (sep_wp thms ctxt))
*}
end

View File

@ -1,5 +1,5 @@
(*
* @LICENSE(NICTA_BSD)
* @TAG(NICTA_BSD)
*
* Insulin.thy: regulate sugar in terms, thms and proof goals.
*

View File

@ -1,5 +1,5 @@
(*
* @LICENSE(NICTA_BSD)
* @TAG(NICTA_BSD)
*
* ShowTypes: show "hidden" type constraints in terms.
* This is a simple utility around Sledgehammer's type annotation code.

View File

@ -6816,5 +6816,49 @@ lemma unat_ucast_no_overflow_le:
thus ?thesis by (simp add:RL LR iffI)
qed
(* casting a long word to a shorter word and casting back to the long word
is equal to the original long word -- if the word is small enough.
'l is the longer word.
's is the shorter word.
*)
lemma bl_cast_long_short_long_ingoreLeadingZero_generic:
"length (dropWhile Not (to_bl w)) \<le> len_of TYPE('s) \<Longrightarrow>
len_of TYPE('s) \<le> len_of TYPE('l) \<Longrightarrow>
(of_bl:: bool list \<Rightarrow> 'l::len word) (to_bl ((of_bl:: bool list \<Rightarrow> 's::len word) (to_bl w))) = w"
apply(rule Word.word_uint_eqI)
apply(subst WordLib.uint_of_bl_is_bl_to_bin)
apply(simp; fail)
apply(subst Word.to_bl_bin)
apply(subst uint_of_bl_is_bl_to_bin_drop)
apply blast
apply(simp)
done
(*
Casting between longer and shorter word.
'l is the longer word.
's is the shorter word.
For example: 'l::len word is 128 word (full ipv6 address)
's::len word is 16 word (address piece of ipv6 address in colon-text-representation)
*)
corollary ucast_short_ucast_long_ingoreLeadingZero:
"length (dropWhile Not (to_bl w)) \<le> len_of TYPE('s) \<Longrightarrow>
len_of TYPE('s) \<le> len_of TYPE('l) \<Longrightarrow>
(ucast:: 's::len word \<Rightarrow> 'l::len word) ((ucast:: 'l::len word \<Rightarrow> 's::len word) w) = w"
apply(subst Word.ucast_bl)+
apply(rule bl_cast_long_short_long_ingoreLeadingZero_generic)
apply(simp_all)
done
lemma length_drop_mask:
fixes w::"'a::len word"
shows "length (dropWhile Not (to_bl (w AND mask n))) \<le> n"
proof -
have length_takeWhile_Not_replicate_False:
"length (takeWhile Not (replicate n False @ ls)) = n + length (takeWhile Not ls)"
for ls n by(subst takeWhile_append2) simp+
show ?thesis
unfolding Word.bl_and_mask
by (simp add: List.dropWhile_eq_drop length_takeWhile_Not_replicate_False)
qed
end

View File

@ -27,17 +27,32 @@ lemma shiftl_power:
lemmas of_bl_reasoning = to_bl_use_of_bl of_bl_append
lemma uint_of_bl_is_bl_to_bin:
"length l\<le>len_of TYPE('a) \<Longrightarrow>
(*TODO: delete, this will be in the Isabelle main distribution in future*)
lemma bl_to_bin_lt2p_drop: "bl_to_bin bs < 2 ^ length (dropWhile Not bs)"
unfolding bl_to_bin_def
proof(induction bs)
case(Cons b bs)
with bl_to_bin_lt2p_aux[where w=1] show ?case by simp
qed(simp)
lemma uint_of_bl_is_bl_to_bin_drop:
"length (dropWhile Not l) \<le> len_of TYPE('a) \<Longrightarrow>
uint ((of_bl::bool list\<Rightarrow> ('a :: len) word) l) = bl_to_bin l"
apply (simp add: of_bl_def)
apply (rule word_uint.Abs_inverse)
apply (simp add: uints_num bl_to_bin_ge0)
apply (rule order_less_le_trans, rule bl_to_bin_lt2p)
apply (rule order_trans, erule power_increasing)
apply simp_all
apply (rule order_less_le_trans)
apply (rule bl_to_bin_lt2p_drop)
apply(simp)
done
corollary uint_of_bl_is_bl_to_bin:
"length l\<le>len_of TYPE('a) \<Longrightarrow>
uint ((of_bl::bool list\<Rightarrow> ('a :: len) word) l) = bl_to_bin l"
apply(rule uint_of_bl_is_bl_to_bin_drop)
using le_trans length_dropWhile_le by blast
lemma bin_to_bl_or:
"bin_to_bl n (a OR b) = map2 (op \<or>) (bin_to_bl n a) (bin_to_bl n b)"
using bl_or_aux_bin[where n=n and v=a and w=b and bs="[]" and cs="[]"]

View File

@ -41,7 +41,7 @@ text \<open>
ML {*
structure Term_Pattern_Antiquote = struct
val quote_string = quote o String.toString
val quote_string = quote
(* typ matching; doesn't support matching on named TVars.
* This is because each TVar is likely to appear many times in the pattern. *)
@ -61,7 +61,14 @@ fun gen_term_pattern (Var (("_dummy_", _), _)) = "_"
"Term.Const (" ^ quote_string n ^ ", " ^ gen_typ_pattern typ ^ ")"
| gen_term_pattern (Free (n, typ)) =
"Term.Free (" ^ quote_string n ^ ", " ^ gen_typ_pattern typ ^ ")"
| gen_term_pattern (f $ x) = "(" ^ gen_term_pattern f ^ " $ " ^ gen_term_pattern x ^ ")"
| gen_term_pattern (t as f $ x) =
(* (read_term_pattern "_") "helpfully" generates a dummy var that is
* applied to all bound vars in scope. We go back and remove them. *)
let fun default () = "(" ^ gen_term_pattern f ^ " $ " ^ gen_term_pattern x ^ ")";
in case strip_comb t of
(h as Var (("_dummy_", _), _), bs) =>
if forall is_Bound bs then gen_term_pattern h else default ()
| _ => default () end
| gen_term_pattern (Abs (_, typ, t)) =
"Term.Abs (_, " ^ gen_typ_pattern typ ^ ", " ^ gen_term_pattern t ^ ")"
| gen_term_pattern (Bound n) = "Bound " ^ string_of_int n
@ -99,4 +106,14 @@ fun eval_num @{term_pat "numeral ?n"} = HOLogic.dest_num n
eval_num @{term "(1 + 2) * 3 - 4 div 5"}
*}
text \<open>Regression test: backslash handling\<close>
ML_val {*
val @{term_pat "\<alpha>"} = @{term "\<alpha>"}
*}
text \<open>Regression test: special-casing for dummy vars\<close>
ML_val {*
val @{term_pat "\<lambda>x y. _"} = @{term "\<lambda>x y. ()"}
*}
end

View File

@ -10,7 +10,6 @@
theory Generic_Separation_Algebras
imports Sep_Algebra_L4v
begin
@ -22,10 +21,12 @@ end
instantiation "prod" :: (stronger_sep_algebra,stronger_sep_algebra) stronger_sep_algebra
begin
definition "sep_disj_prod x y \<equiv> sep_disj (fst x) (fst y) \<and> sep_disj (snd x) (snd y)"
definition "plus_prod x y \<equiv> ( (fst x) + (fst y) , (snd x) + (snd y))"
definition "sep_disj_prod x y \<equiv> sep_disj (fst x) (fst y) \<and> sep_disj (snd x) (snd y)"
definition "plus_prod x y \<equiv> ( (fst x) + (fst y) , (snd x) + (snd y))"
instance
apply (default)
apply intro_classes
apply (metis fst_conv sep_disj_prod_def sep_disj_zero snd_conv zero_prod_def)
apply (metis (full_types) sep_disj_commuteI sep_disj_prod_def)
apply (clarsimp simp: plus_prod_def zero_prod_def)
@ -38,7 +39,8 @@ instance
apply (clarsimp simp: plus_prod_def sep_disj_prod_def)
apply (clarsimp simp: plus_prod_def sep_disj_prod_def)
apply (clarsimp simp: plus_prod_def sep_disj_prod_def)
done
done
end
@ -48,48 +50,53 @@ begin
instance ..
end
definition orelse
definition
orelse :: "'a option \<Rightarrow> 'a option \<Rightarrow> 'a option"
where
"orelse x y \<equiv> case x of Some x' => Some x' | None => y"
"orelse x y \<equiv> case x of Some x' => Some x' | None => y"
definition onlyone :: "'a option => 'a option => bool"
definition
onlyone :: "'a option => 'a option => bool"
where
"onlyone x y \<equiv> case x of Some x' => case y of Some y' => False |
_ => True
| _ => True"
"onlyone x y \<equiv> case x of
Some x' => (case y of Some y' => False | _ => True)
| None => True"
instantiation "option" :: (type) stronger_sep_algebra
begin
definition "sep_disj_option \<equiv> onlyone "
definition "plus_option \<equiv> orelse"
instance
apply (default)
apply (clarsimp simp: sep_disj_option_def zero_option_def plus_option_def onlyone_def orelse_def split: option.splits)+
done
end
definition "sep_disj_option \<equiv> onlyone "
definition "plus_option \<equiv> orelse"
instance
by intro_classes
(clarsimp simp: sep_disj_option_def zero_option_def plus_option_def onlyone_def orelse_def
split: option.splits)+
end
record 'a generic_heap_record =
heap :: "'a "
instantiation "generic_heap_record_ext" :: (stronger_sep_algebra, stronger_sep_algebra) stronger_sep_algebra
begin
definition "zero_generic_heap_record_ext \<equiv> Abs_generic_heap_record_ext 0"
definition "plus_generic_heap_record_ext h1 h2 \<equiv>
begin
definition "zero_generic_heap_record_ext \<equiv> Abs_generic_heap_record_ext 0"
definition "plus_generic_heap_record_ext h1 h2 \<equiv>
Abs_generic_heap_record_ext (Rep_generic_heap_record_ext h1 + Rep_generic_heap_record_ext h2)"
definition "sep_disj_generic_heap_record_ext h1 h2 \<equiv> Rep_generic_heap_record_ext h1 ## Rep_generic_heap_record_ext h2"
instance
apply (default)
apply (fastforce simp: sep_add_left_commute sep_disj_commute Rep_generic_heap_record_ext_inverse sep_disj_commuteI sep_add_assoc sep_add_commute Abs_generic_heap_record_ext_inverse zero_generic_heap_record_ext_def sep_disj_generic_heap_record_ext_def plus_generic_heap_record_ext_def)+
done
definition "sep_disj_generic_heap_record_ext h1 h2 \<equiv> Rep_generic_heap_record_ext h1 ## Rep_generic_heap_record_ext h2"
instance
apply intro_classes
apply (fastforce simp: sep_add_left_commute sep_disj_commute Rep_generic_heap_record_ext_inverse
sep_disj_commuteI sep_add_assoc sep_add_commute Abs_generic_heap_record_ext_inverse
zero_generic_heap_record_ext_def sep_disj_generic_heap_record_ext_def
plus_generic_heap_record_ext_def)+
done
end
end

View File

@ -13,7 +13,7 @@
Rafal Kolanski <rafal.kolanski at nicta.com.au>
*)
header {* More properties of maps plus map disjuction. *}
chapter {* More properties of maps plus map disjuction. *}
theory Map_Extra
imports "~~/src/HOL/Main"
@ -57,11 +57,11 @@ definition
subsection {* Properties of maps not related to restriction *}
lemma empty_forall_equiv: "(m = empty) = (\<forall>x. m x = None)"
by (fastforce intro!: ext)
by fastforce
lemma map_le_empty2 [simp]:
"(m \<subseteq>\<^sub>m empty) = (m = empty)"
by (auto simp: map_le_def intro: ext)
by (auto simp: map_le_def)
lemma dom_iff:
"(\<exists>y. m x = Some y) = (x \<in> dom m)"
@ -111,20 +111,18 @@ lemma map_add_right_dom_eq:
lemma map_le_same_dom_eq:
"\<lbrakk> m\<^sub>0 \<subseteq>\<^sub>m m\<^sub>1 ; dom m\<^sub>0 = dom m\<^sub>1 \<rbrakk> \<Longrightarrow> m\<^sub>0 = m\<^sub>1"
by (auto intro!: ext simp: map_le_def elim!: ballE)
by (simp add: map_le_antisym map_le_def)
subsection {* Properties of map restriction *}
lemma restrict_map_cancel:
"(m |` S = m |` T) = (dom m \<inter> S = dom m \<inter> T)"
by (fastforce intro: ext dest: fun_cong
simp: restrict_map_def None_not_eq
split: split_if_asm)
by (fastforce dest: fun_cong simp: restrict_map_def None_not_eq split: split_if_asm)
lemma map_add_restricted_self [simp]:
"m ++ m |` S = m"
by (auto intro: ext simp: restrict_map_def map_add_def split: option.splits)
by (auto simp: restrict_map_def map_add_def split: option.splits)
lemma map_add_restrict_dom_right [simp]:
"(m ++ m') |` dom m' = m'"
@ -136,19 +134,19 @@ lemma restrict_map_UNIV [simp]:
lemma restrict_map_dom:
"S = dom m \<Longrightarrow> m |` S = m"
by (auto intro!: ext simp: restrict_map_def None_not_eq)
by (rule ext, auto simp: restrict_map_def None_not_eq)
lemma restrict_map_subdom:
"dom m \<subseteq> S \<Longrightarrow> m |` S = m"
by (fastforce simp: restrict_map_def None_com intro: ext)
by (fastforce simp: restrict_map_def None_com)
lemma map_add_restrict:
"(m\<^sub>0 ++ m\<^sub>1) |` S = ((m\<^sub>0 |` S) ++ (m\<^sub>1 |` S))"
by (force simp: map_add_def restrict_map_def intro: ext)
by (force simp: map_add_def restrict_map_def)
lemma map_le_restrict:
"m \<subseteq>\<^sub>m m' \<Longrightarrow> m = m' |` dom m"
by (force simp: map_le_def restrict_map_def None_com intro: ext)
by (force simp: map_le_def restrict_map_def None_com)
lemma restrict_map_le:
"m |` S \<subseteq>\<^sub>m m"
@ -161,15 +159,15 @@ lemma restrict_map_remerge:
lemma restrict_map_empty:
"dom m \<inter> S = {} \<Longrightarrow> m |` S = empty"
by (fastforce simp: restrict_map_def intro: ext)
by (fastforce simp: restrict_map_def)
lemma map_add_restrict_comp_right [simp]:
"(m |` S ++ m |` (UNIV - S)) = m"
by (force simp: map_add_def restrict_map_def split: option.splits intro: ext)
by (force simp: map_add_def restrict_map_def split: option.splits)
lemma map_add_restrict_comp_right_dom [simp]:
"(m |` S ++ m |` (dom m - S)) = m"
by (auto simp: map_add_def restrict_map_def split: option.splits intro!: ext)
by (rule ext, auto simp: map_add_def restrict_map_def split: option.splits)
lemma map_add_restrict_comp_left [simp]:
"(m |` (UNIV - S) ++ m |` S) = m"
@ -177,7 +175,7 @@ lemma map_add_restrict_comp_left [simp]:
lemma restrict_self_UNIV:
"m |` (dom m - S) = m |` (UNIV - S)"
by (auto intro!: ext simp: restrict_map_def)
by (rule ext, auto simp: restrict_map_def)
lemma map_add_restrict_nonmember_right:
"x \<notin> dom m' \<Longrightarrow> (m ++ m') |` {x} = m |` {x}"
@ -193,7 +191,7 @@ lemma map_add_restrict_right:
lemma restrict_map_compose:
"\<lbrakk> S \<union> T = dom m ; S \<inter> T = {} \<rbrakk> \<Longrightarrow> m |` S ++ m |` T = m"
by (fastforce intro: ext simp: map_add_def restrict_map_def)
by (fastforce simp: map_add_def restrict_map_def)
lemma map_le_dom_subset_restrict:
"\<lbrakk> m' \<subseteq>\<^sub>m m; dom m' \<subseteq> S \<rbrakk> \<Longrightarrow> m' \<subseteq>\<^sub>m (m |` S)"
@ -201,22 +199,20 @@ lemma map_le_dom_subset_restrict:
lemma map_le_dom_restrict_sub_add:
"m' \<subseteq>\<^sub>m m \<Longrightarrow> m |` (dom m - dom m') ++ m' = m"
by (auto simp: None_com map_add_def restrict_map_def map_le_def
split: option.splits
intro!: ext)
(force simp: Some_com)+
by (rule ext, auto simp: None_com map_add_def restrict_map_def map_le_def
split: option.splits; force simp: Some_com)
lemma subset_map_restrict_sub_add:
"T \<subseteq> S \<Longrightarrow> m |` (S - T) ++ m |` T = m |` S"
by (auto simp: restrict_map_def map_add_def intro!: ext split: option.splits)
"T \<subseteq> S \<Longrightarrow> m |` (S - T) ++ m |` T = m |` S"
by (rule ext) (auto simp: restrict_map_def map_add_def split: option.splits)
lemma restrict_map_sub_union:
"m |` (dom m - (S \<union> T)) = (m |` (dom m - T)) |` (dom m - S)"
by (auto intro!: ext simp: restrict_map_def)
by (auto simp: restrict_map_def)
lemma prod_restrict_map_add:
"\<lbrakk> S \<union> T = U; S \<inter> T = {} \<rbrakk> \<Longrightarrow> m |` (X \<times> S) ++ m |` (X \<times> T) = m |` (X \<times> U)"
by (auto simp: map_add_def restrict_map_def intro!: ext split: option.splits)
by (auto simp: map_add_def restrict_map_def split: option.splits)
section {* Things that should not go into Map.thy (separation logic) *}
@ -240,8 +236,7 @@ lemma restrict_map_sub_disj: "h |` S \<bottom> h `- S"
lemma restrict_map_sub_add: "h |` S ++ h `- S = h"
by (fastforce simp: sub_restrict_map_def restrict_map_def map_add_def
split: option.splits split_if
intro: ext)
split: option.splits split_if)
subsection {* Properties of map disjunction *}
@ -275,7 +270,7 @@ lemma map_add_com:
lemma map_add_left_commute:
"h\<^sub>0 \<bottom> h\<^sub>1 \<Longrightarrow> h\<^sub>0 ++ (h\<^sub>1 ++ h\<^sub>2) = h\<^sub>1 ++ (h\<^sub>0 ++ h\<^sub>2)"
by (simp add: map_add_com map_disj_com map_add_assoc)
by (simp add: map_add_com map_disj_com)
lemma map_add_disj:
"h\<^sub>0 \<bottom> (h\<^sub>1 ++ h\<^sub>2) = (h\<^sub>0 \<bottom> h\<^sub>1 \<and> h\<^sub>0 \<bottom> h\<^sub>2)"
@ -340,15 +335,15 @@ lemma map_add_eval_left:
lemma map_add_eval_right:
"\<lbrakk> x \<in> dom h' \<rbrakk> \<Longrightarrow> (h ++ h') x = h' x"
by (auto elim!: map_disjD simp: map_add_comm map_add_eval_left map_disj_com)
by (rule map_add_dom_app_simps)
lemma map_add_eval_left':
"\<lbrakk> x \<notin> dom h' \<rbrakk> \<Longrightarrow> (h ++ h') x = h x"
by (clarsimp simp: map_disj_def map_add_def split: option.splits)
by (rule map_add_dom_app_simps)
lemma map_add_eval_right':
"\<lbrakk> x \<notin> dom h \<rbrakk> \<Longrightarrow> (h ++ h') x = h' x"
by (clarsimp simp: map_disj_def map_add_def split: option.splits)
by (rule map_add_dom_app_simps)
lemma map_add_left_dom_eq:
assumes eq: "h\<^sub>0 ++ h\<^sub>1 = h\<^sub>0' ++ h\<^sub>1'"
@ -366,7 +361,7 @@ lemma map_add_left_eq:
shows "h\<^sub>0 = h\<^sub>1"
proof (rule ext)
fix x
from eq have eq': "(h\<^sub>0 ++ h) x = (h\<^sub>1 ++ h) x" by (auto intro!: ext)
from eq have eq': "(h\<^sub>0 ++ h) x = (h\<^sub>1 ++ h) x" by auto
{ assume "x \<in> dom h"
hence "h\<^sub>0 x = h\<^sub>1 x" using disj by (simp add: map_disj_None_left)
} moreover {
@ -410,11 +405,10 @@ lemma map_add_left_cancel:
proof (rule iffI, rule ext)
fix x
assume "(h\<^sub>0 ++ h\<^sub>1) = (h\<^sub>0 ++ h\<^sub>1')"
hence "(h\<^sub>0 ++ h\<^sub>1) x = (h\<^sub>0 ++ h\<^sub>1') x" by (auto intro!: ext)
hence "(h\<^sub>0 ++ h\<^sub>1) x = (h\<^sub>0 ++ h\<^sub>1') x" by auto
hence "h\<^sub>1 x = h\<^sub>1' x" using disj
by - (cases "x \<in> dom h\<^sub>0",
simp_all add: map_disj_None_right map_add_eval_right')
thus "h\<^sub>1 x = h\<^sub>1' x" by (auto intro!: ext)
by (cases "x \<in> dom h\<^sub>0"; simp add: map_disj_None_right map_add_eval_right')
thus "h\<^sub>1 x = h\<^sub>1' x" by auto
qed auto
lemma map_add_lr_disj:
@ -427,7 +421,7 @@ subsection {* Map disjunction and map updates *}
lemma map_disj_update_left [simp]:
"p \<in> dom h\<^sub>1 \<Longrightarrow> h\<^sub>0 \<bottom> h\<^sub>1(p \<mapsto> v) = h\<^sub>0 \<bottom> h\<^sub>1"
by (clarsimp simp add: map_disj_def, blast)
by (clarsimp simp: map_disj_def, blast)
lemma map_disj_update_right [simp]:
"p \<in> dom h\<^sub>1 \<Longrightarrow> h\<^sub>1(p \<mapsto> v) \<bottom> h\<^sub>0 = h\<^sub>1 \<bottom> h\<^sub>0"
@ -436,12 +430,12 @@ lemma map_disj_update_right [simp]:
lemma map_add_update_left:
"\<lbrakk> h\<^sub>0 \<bottom> h\<^sub>1 ; p \<in> dom h\<^sub>0 \<rbrakk> \<Longrightarrow> (h\<^sub>0 ++ h\<^sub>1)(p \<mapsto> v) = (h\<^sub>0(p \<mapsto> v) ++ h\<^sub>1)"
by (drule (1) map_disj_None_right)
(auto intro: ext simp: map_add_def cong: option.case_cong)
(auto simp: map_add_def cong: option.case_cong)
lemma map_add_update_right:
"\<lbrakk> h\<^sub>0 \<bottom> h\<^sub>1 ; p \<in> dom h\<^sub>1 \<rbrakk> \<Longrightarrow> (h\<^sub>0 ++ h\<^sub>1)(p \<mapsto> v) = (h\<^sub>0 ++ h\<^sub>1 (p \<mapsto> v))"
by (drule (1) map_disj_None_left)
(auto intro: ext simp: map_add_def cong: option.case_cong)
(auto simp: map_add_def cong: option.case_cong)
lemma map_add3_update:
"\<lbrakk> h\<^sub>0 \<bottom> h\<^sub>1 ; h\<^sub>1 \<bottom> h\<^sub>2 ; h\<^sub>0 \<bottom> h\<^sub>2 ; p \<in> dom h\<^sub>0 \<rbrakk>
@ -499,7 +493,7 @@ lemma map_le_conv:
unfolding map_le_def map_disj_def map_add_def
by (rule iffI,
clarsimp intro!: exI[where x="\<lambda>x. if x \<notin> dom h\<^sub>0' then h\<^sub>0 x else None"])
(fastforce intro: ext intro: split: option.splits split_if_asm)+
(fastforce intro: split: option.splits split_if_asm)+
lemma map_le_conv2:
"h\<^sub>0' \<subseteq>\<^sub>m h\<^sub>0 = (\<exists>h\<^sub>1. h\<^sub>0 = h\<^sub>0' ++ h\<^sub>1 \<and> h\<^sub>0' \<bottom> h\<^sub>1)"
@ -522,7 +516,7 @@ lemma map_disj_restrict_dom [simp]:
lemma restrict_map_disj_dom_empty:
"h \<bottom> h' \<Longrightarrow> h |` dom h' = empty"
by (fastforce simp: map_disj_def restrict_map_def intro: ext)
by (fastforce simp: map_disj_def restrict_map_def)
lemma restrict_map_univ_disj_eq:
"h \<bottom> h' \<Longrightarrow> h |` (UNIV - dom h') = h"
@ -562,7 +556,7 @@ lemma restrict_map_on_disj:
lemma restrict_map_on_disj':
"h\<^sub>0 \<bottom> h\<^sub>1 \<Longrightarrow> h\<^sub>0 \<bottom> h\<^sub>1 |` S"
by (auto simp: map_disj_def map_add_def)
by (rule restrict_map_disj_right)
lemma map_le_sub_dom:
"\<lbrakk> h\<^sub>0 ++ h\<^sub>1 \<subseteq>\<^sub>m h ; h\<^sub>0 \<bottom> h\<^sub>1 \<rbrakk> \<Longrightarrow> h\<^sub>0 \<subseteq>\<^sub>m h |` (dom h - dom h\<^sub>1)"
@ -571,9 +565,9 @@ lemma map_le_sub_dom:
lemma map_submap_break:
"\<lbrakk> h \<subseteq>\<^sub>m h' \<rbrakk> \<Longrightarrow> h' = (h' |` (UNIV - dom h)) ++ h"
by (fastforce intro!: ext split: option.splits
simp: map_le_restrict restrict_map_def map_le_def map_add_def
dom_def)
by (fastforce split: option.splits
simp: map_le_restrict restrict_map_def map_le_def map_add_def
dom_def)
lemma map_add_disj_restrict_both:
"\<lbrakk> h\<^sub>0 \<bottom> h\<^sub>1; S \<inter> S' = {}; T \<inter> T' = {} \<rbrakk>

View File

@ -141,4 +141,5 @@ lemma sep_set_conj_subset_wp':
done
end
end

View File

@ -16,23 +16,23 @@ text{* Beyond the tactics above, there is also a set of attributes implemented t
things in separation logic easier. These rules should be considered internals and are not
intended for direct use. *}
lemma sep_curry_atomised: "\<lbrakk>(\<And>s. (P \<and>* Q) s \<longrightarrow> R s); P s \<rbrakk> \<Longrightarrow> (Q \<longrightarrow>* R) s"
by (clarsimp simp: sep_conj_sep_impl)
by (clarsimp simp: sep_conj_sep_impl)
lemma sep_remove_pure_imp_sep_imp: "( P \<longrightarrow>* (\<lambda>s. P' \<longrightarrow> Q s)) s \<Longrightarrow> P' \<Longrightarrow> (P \<longrightarrow>* Q) s"
by (clarsimp)
lemma sep_backward: "\<lbrakk>\<And>s. P s \<longrightarrow> (Q \<and>* T) s; (P \<and>* (Q \<longrightarrow>* R)) s \<rbrakk> \<Longrightarrow> (T \<and>* R) s"
by (metis sep_conj_commute sep_conj_impl1 sep_mp_frame)
by (metis sep_conj_commute sep_conj_impl1 sep_mp_frame)
lemma sep_remove_conj: "\<lbrakk>(P \<and>* R) s ; Q\<rbrakk> \<Longrightarrow> ((\<lambda>s. P s \<and> Q) \<and>* R) s "
apply (clarsimp)
done
apply (clarsimp)
done
lemma curry: "(P \<longrightarrow> Q \<longrightarrow> R) \<Longrightarrow> (P \<and> Q) \<longrightarrow> R"
apply (safe)
done
apply (safe)
done
ML {*
@ -41,20 +41,17 @@ local
fun setup_simpset ctxt = put_simpset HOL_basic_ss ctxt addsimps [(sym OF [@{thm sep_conj_assoc}])]
fun simp ctxt thm = simplify (setup_simpset ctxt) thm
fun REPEAT_TRYOF_N _ thm2 0 = thm2 |
REPEAT_TRYOF_N thm1 thm2 n = REPEAT_TRYOF_N thm1 (thm1 OF [thm2]) (n-1)
fun REPEAT_TRYOF_N _ thm2 0 = thm2
| REPEAT_TRYOF_N thm1 thm2 n = REPEAT_TRYOF_N thm1 (thm1 OF [thm2]) (n-1)
fun REPEAT_TRYOF'_N thm1 _ 0 = thm1
| REPEAT_TRYOF'_N thm1 thm2 n = REPEAT_TRYOF'_N (thm1 OF [thm2]) thm2 (n-1)
fun REPEAT_TRYOF'_N thm1 _ 0 = thm1 |
REPEAT_TRYOF'_N thm1 thm2 n = REPEAT_TRYOF'_N (thm1 OF [thm2]) thm2 (n-1)
fun attribute_thm ctxt thm thm' =
REPEAT_TRYOF_N @{thm sep_remove_pure_imp_sep_imp} (thm OF [atomize_thm ctxt thm']) (Thm.nprems_of thm' - 1)
fun attribute_thm ctxt thm thm' = REPEAT_TRYOF_N @{thm sep_remove_pure_imp_sep_imp}
(thm OF [atomize_thm ctxt thm']) (Thm.nprems_of thm' - 1)
fun attribute_thm' thm ctxt thm' =
thm OF [REPEAT_TRYOF_N @{thm curry} (thm' |> atomize_thm ctxt o simp ctxt) (Thm.nprems_of thm' - 1)]
fun attribute_thm' thm ctxt thm' =
thm OF [REPEAT_TRYOF_N @{thm curry} (thm' |> atomize_thm ctxt o simp ctxt) (Thm.nprems_of thm' - 1)]
in
@ -67,7 +64,8 @@ val sep_curry = Thm.rule_attribute [] (fn ctxt => sep_curry_inner (Context.proof
(*
The attribute sep_back takes a rule of the form A \<Longrightarrow> B and returns a rule (A \<and>* (B \<longrightarrow>* R)) \<Longrightarrow> R.
The R then matches with any conclusion. If the theorem is of form (A \<and>* B) \<Longrightarrow> C, it is advised to use sep_curry on the theorem first, and then sep_back. This aids sep_cancel in simplifying the result.
The R then matches with any conclusion. If the theorem is of form (A \<and>* B) \<Longrightarrow> C, it is advised to
use sep_curry on the theorem first, and then sep_back. This aids sep_cancel in simplifying the result.
*)
fun backward ctxt thm =
@ -77,11 +75,10 @@ fun backward' ctxt thm = backward (Context.proof_of ctxt) thm
val sep_backward = Thm.rule_attribute [] (backward')
end;
end
*}
attribute_setup sep_curry = {* Scan.succeed sep_curry *}
attribute_setup sep_backward = {* Scan.succeed sep_backward *}
end

View File

@ -16,59 +16,53 @@ begin
lemma sep_curry': "\<lbrakk>(P \<and>* F) s; \<And>s. (Q \<and>* P \<and>* F) s \<Longrightarrow> R s\<rbrakk> \<Longrightarrow> (Q \<longrightarrow>* R) s"
by (metis (full_types) sep.mult_commute sep_curry)
by (metis (full_types) sep.mult_commute sep_curry)
lemma sep_conj_sep_impl_safe: "(P \<longrightarrow>* P') s \<Longrightarrow> (\<And>s. ((P \<longrightarrow>* P') \<and>* Q) s \<Longrightarrow> (Q') s) \<Longrightarrow> (Q \<longrightarrow>* Q') s"
apply (erule sep_curry[rotated])
apply (clarsimp)
done
lemma sep_conj_sep_impl_safe:
"(P \<longrightarrow>* P') s \<Longrightarrow> (\<And>s. ((P \<longrightarrow>* P') \<and>* Q) s \<Longrightarrow> (Q') s) \<Longrightarrow> (Q \<longrightarrow>* Q') s"
by (rule sep_curry)
lemma sep_conj_sep_impl_safe': "P s \<Longrightarrow> (\<And>s. (P \<and>* Q) s \<Longrightarrow> (P \<and>* R) s) \<Longrightarrow> (Q \<longrightarrow>* P \<and>* R) s"
apply (erule sep_curry[rotated])
apply (clarsimp)
done
by (rule sep_curry)
lemma sep_wand_lens_simple: "(\<And>s. T s = (Q \<and>* R) s) \<Longrightarrow> (P \<longrightarrow>* T) s \<Longrightarrow> (P \<longrightarrow>* Q \<and>* R) s"
apply (clarsimp simp: sep_impl_def)
done
by (clarsimp simp: sep_impl_def)
schematic_goal schem_impAny: " (?C \<and>* B) s \<Longrightarrow> A s" by (erule sep_mp)
ML {*
fun sep_cancel_tactic ctxt concl =
let val thms = rev (SepCancel_Rules.get ctxt)
val tac = assume_tac ctxt ORELSE'
eresolve_tac ctxt [@{thm sep_mp}, @{thm sep_conj_empty}, @{thm sep_empty_conj}]
ORELSE' sep_erule_tactic ctxt thms
val direct_tac = eresolve_tac ctxt thms
val safe_sep_wand_tac = rotator' ctxt (resolve0_tac [@{thm sep_wand_lens_simple}]) (eresolve0_tac [@{thm sep_conj_sep_impl_safe'}])
fun sep_cancel_tactic_inner true = sep_erule_full_tac' tac ctxt |
sep_cancel_tactic_inner false = sep_erule_full_tac tac ctxt
in sep_cancel_tactic_inner concl ORELSE'
fun sep_cancel_tactic ctxt concl =
let val thms = rev (SepCancel_Rules.get ctxt)
val tac = assume_tac ctxt ORELSE'
eresolve_tac ctxt [@{thm sep_mp}, @{thm sep_conj_empty}, @{thm sep_empty_conj}] ORELSE'
sep_erule_tactic ctxt thms
val direct_tac = eresolve_tac ctxt thms
val safe_sep_wand_tac = rotator' ctxt (resolve0_tac [@{thm sep_wand_lens_simple}]) (eresolve0_tac [@{thm sep_conj_sep_impl_safe'}])
fun sep_cancel_tactic_inner true = sep_erule_full_tac' tac ctxt
| sep_cancel_tactic_inner false = sep_erule_full_tac tac ctxt
in sep_cancel_tactic_inner concl ORELSE'
eresolve_tac ctxt [@{thm sep_curry'}, @{thm sep_conj_sep_impl_safe}, @{thm sep_imp_empty}, @{thm sep_empty_imp'}] ORELSE'
safe_sep_wand_tac ORELSE'
direct_tac
end;
end
fun sep_cancel_tactic' ctxt concl =
let
val sep_cancel = sep_cancel_tactic ctxt
in
(sep_flatten ctxt THEN_ALL_NEW sep_cancel concl) ORELSE' sep_cancel concl
end
fun sep_cancel_tactic' ctxt concl =
let val sep_cancel = (sep_cancel_tactic ctxt)
in ((sep_flatten ctxt) THEN_ALL_NEW sep_cancel concl) ORELSE' sep_cancel concl
end;
fun sep_cancel_method (concl,_) ctxt = SIMPLE_METHOD' (sep_cancel_tactic' ctxt concl)
fun sep_cancel_method (concl,_) ctxt = SIMPLE_METHOD' (sep_cancel_tactic' ctxt concl)
val sep_cancel_syntax = Method.sections [
Args.add -- Args.colon >> K (Method.modifier SepCancel_Rules.add @{here})];
val sep_cancel_syntax' = (Scan.lift (Args.mode "concl") -- sep_cancel_syntax)
val sep_cancel_syntax =
Method.sections [Args.add -- Args.colon >> K (Method.modifier SepCancel_Rules.add @{here})];
val sep_cancel_syntax' =
Scan.lift (Args.mode "concl") -- sep_cancel_syntax
*}
method_setup sep_cancel = {*
sep_cancel_syntax' >> (sep_cancel_method)
*}{* Simple elimination of conjuncts *}
method_setup sep_cancel =
{* sep_cancel_syntax' >> sep_cancel_method *} {* Simple elimination of conjuncts *}
end

View File

@ -16,8 +16,8 @@ begin
This is fairly similar to sep_erule_full sep_cancel. *)
lemma "(A \<and>* B \<and>* C \<and>* D) s \<Longrightarrow> (C \<and>* D \<and>* A \<and>* B) s"
apply (sep_cancel)
apply (sep_cancel)
done
apply sep_cancel
apply sep_cancel
done
end

View File

@ -14,9 +14,9 @@ begin
ML {*
structure SepCancel_Rules = Named_Thms (
val name = @{binding "sep_cancel"};
val description = "sep_cancel rules";
);
val name = @{binding "sep_cancel"}
val description = "sep_cancel rules"
)
*}
setup SepCancel_Rules.setup

View File

@ -13,7 +13,7 @@
Rafal Kolanski <rafal.kolanski at nicta.com.au>
*)
header "Equivalence between Separation Algebra Formulations"
chapter "Equivalence between Separation Algebra Formulations"
theory Sep_Eq
imports
@ -29,6 +29,7 @@ text {*
*}
no_notation map_add (infixl "++" 100)
declare [[syntax_ambiguity_warning = false]]
section "Total implies Partial"

View File

@ -13,7 +13,7 @@
Rafal Kolanski <rafal.kolanski at nicta.com.au>
*)
header "Standard Heaps as an Instance of Separation Algebra"
chapter "Standard Heaps as an Instance of Separation Algebra"
theory Sep_Heap_Instance
imports Separation_Algebra
@ -52,7 +52,7 @@ definition
sep_disj_fun_def: "sep_disj m1 m2 \<equiv> domain m1 \<inter> domain m2 = {}"
instance
apply default
apply intro_classes
apply (simp add: sep_disj_fun_def domain_def zero_fun_def)
apply (fastforce simp: sep_disj_fun_def)
apply (simp add: plus_fun_def zero_fun_def)
@ -93,7 +93,7 @@ lemma
section {* @{typ unit} Instantiation *}
text {*
The @{typ unit} type also forms a seperation algebra. Although
The @{typ unit} type also forms a separation algebra. Although
typically not useful as a state space by itself, it may be
a type parameter to more complex state space.
*}
@ -103,16 +103,16 @@ begin
definition "plus_unit (a :: unit) (b :: unit) \<equiv> ()"
definition "sep_disj_unit (a :: unit) (b :: unit) \<equiv> True"
instance
apply default
apply intro_classes
apply (simp add: plus_unit_def sep_disj_unit_def)+
done
done
end
lemma unit_disj_sep_unit [simp]: "(a :: unit) ## b"
by (clarsimp simp: sep_disj_unit_def)
lemma unit_plus_unit [simp]: "(a :: unit) + b = ()"
by (clarsimp simp: plus_unit_def)
by (rule unit_eq)
section {* @{typ "'a option"} Instantiation *}
@ -120,21 +120,19 @@ text {*
The @{typ "'a option"} is a seperation algebra, with @{term None}
indicating emptyness.
*}
(* FIXME: should we be instantiating the "opt" class above instead? *)
instantiation option :: (type) stronger_sep_algebra
begin
definition "zero_option \<equiv> None"
definition "plus_option (a :: 'a option) (b :: 'a option)
\<equiv> (case b of None \<Rightarrow> a | Some x \<Rightarrow> b)"
definition "sep_disj_option (a :: 'a option) (b :: 'a option)
\<equiv> a = None \<or> b = None"
definition
"zero_option \<equiv> None"
definition
"plus_option (a :: 'a option) (b :: 'a option) \<equiv> (case b of None \<Rightarrow> a | Some x \<Rightarrow> b)"
definition
"sep_disj_option (a :: 'a option) (b :: 'a option) \<equiv> a = None \<or> b = None"
instance
apply default
apply (auto simp: zero_option_def sep_disj_option_def plus_option_def
split: option.splits)
done
by intro_classes
(auto simp: zero_option_def sep_disj_option_def plus_option_def split: option.splits)
end
lemma disj_sep_None [simp]:

View File

@ -13,15 +13,15 @@ imports Sep_Provers Sep_Cancel_Set Sep_Tactic_Helpers
begin
lemma sep_wand_lens: "(\<And>s. T s = Q s) \<Longrightarrow> ((P \<longrightarrow>* T) \<and>* R) s \<Longrightarrow> ((P \<longrightarrow>* Q) \<and>* R) s"
apply (sep_erule_full refl_imp)
apply (clarsimp simp: sep_impl_def)
done
apply (sep_erule_full refl_imp)
apply (clarsimp simp: sep_impl_def)
done
lemma sep_wand_lens': "(\<And>s. T s = Q s) \<Longrightarrow> ((T \<longrightarrow>* P) \<and>* R) s \<Longrightarrow> ((Q \<longrightarrow>* P) \<and>* R) s"
apply (sep_erule_full refl_imp, erule sep_curry[rotated])
apply (clarsimp)
apply (erule sep_mp)
done
apply (sep_erule_full refl_imp, erule sep_curry[rotated])
apply (clarsimp)
apply (erule sep_mp)
done
(* Removing wands from the conclusions *)
@ -30,54 +30,45 @@ ML {*
fun sep_wand_lens ctxt = resolve_tac ctxt[@{thm sep_wand_lens}]
fun sep_wand_lens' ctxt = resolve_tac ctxt [@{thm sep_wand_lens'}]
fun sep_wand_rule_tac tac ctxt =
let val r = rotator' ctxt in
let
val r = rotator' ctxt
in
tac |> r (sep_wand_lens' ctxt) |> r (sep_wand_lens ctxt) |> r (sep_select ctxt)
end;
end
fun sep_wand_rule_tac' thms ctxt =
let val r = rotator' ctxt in
eresolve_tac ctxt thms |> r (sep_wand_lens ctxt) |> r (sep_select ctxt) |> r (sep_asm_select ctxt)
end;
let
val r = rotator' ctxt
in
eresolve_tac ctxt thms |> r (sep_wand_lens ctxt) |> r (sep_select ctxt) |> r (sep_asm_select ctxt)
end
fun sep_wand_rule_method thms ctxt = SIMPLE_METHOD' (sep_wand_rule_tac thms ctxt)
fun sep_wand_rule_method' thms ctxt = SIMPLE_METHOD' (sep_wand_rule_tac' thms ctxt)
*}
(*
method_setup sep_wand_rule = {*
Attrib.thms >> sep_wand_rule_method
*}
method_setup sep_wand_erule = {*
Attrib.thms >> sep_wand_rule_method'
*}
*)
lemma sep_wand_match: "(\<And>s. Q s \<Longrightarrow> Q' s) \<Longrightarrow> (R \<longrightarrow>* R') s ==> (Q \<and>* R \<longrightarrow>* Q' \<and>* R') s"
lemma sep_wand_match:
"(\<And>s. Q s \<Longrightarrow> Q' s) \<Longrightarrow> (R \<longrightarrow>* R') s ==> (Q \<and>* R \<longrightarrow>* Q' \<and>* R') s"
apply (erule sep_curry[rotated])
apply (sep_select_asm 1 3)
apply (sep_drule (direct) sep_mp_frame)
apply (sep_erule_full refl_imp, clarsimp)
done
done
lemma sep_wand_trivial: "(\<And>s. Q s \<Longrightarrow> Q' s) \<Longrightarrow> R' s ==> (Q \<longrightarrow>* Q' \<and>* R') s"
apply (erule sep_curry[rotated])
apply (sep_erule_full refl_imp)
apply (clarsimp)
done
done
lemma sep_wand_collapse: "(P \<and>* Q \<longrightarrow>* R) s \<Longrightarrow> (P \<longrightarrow>* Q \<longrightarrow>* R) s "
apply (erule sep_curry[rotated])+
apply (clarsimp simp: sep_conj_assoc)
apply (erule sep_mp)
done
done
lemma sep_wand_match_less_safe:
assumes drule: " \<And>s. (Q' \<and>* R) s \<Longrightarrow> ((P \<longrightarrow>* R') \<and>* Q' \<and>* R'' ) s "
@ -89,21 +80,20 @@ lemma sep_wand_match_less_safe:
apply (sep_select_asm 1 3)
apply (sep_drule (direct) sep_mp_frame, sep_erule_full refl_imp)
apply (clarsimp)
done
done
ML {*
fun sep_match_trivial_tac ctxt =
let
fun flip f a b = f b a
val sep_cancel = flip (sep_apply_tactic ctxt) (SepCancel_Rules.get ctxt |> rev)
fun f x = x |> rotate_prems ~1 |> (fn x => [x]) |> eresolve0_tac |> sep_cancel
val sep_thms = map f [@{thm sep_wand_trivial}, @{thm sep_wand_match}]
in
sep_wand_rule_tac (resolve0_tac [@{thm sep_rule}] THEN' FIRST' sep_thms) ctxt
end
fun flip f a b = f b a
fun sep_match_trivial_tac ctxt =
let val sepcancel = (flip (sep_apply_tactic ctxt)) (SepCancel_Rules.get ctxt |> rev)
fun f x = x |> rotate_prems ~1 |> (fn x => [x]) |> eresolve0_tac |> sepcancel
val sep_thms = map f [@{thm sep_wand_trivial}, @{thm sep_wand_match}]
in sep_wand_rule_tac (resolve0_tac [@{thm sep_rule}] THEN' FIRST' sep_thms) ctxt
end;
fun sep_safe_method ctxt = SIMPLE_METHOD' (sep_match_trivial_tac ctxt)
fun sep_safe_method ctxt = SIMPLE_METHOD' (sep_match_trivial_tac ctxt)
*}
method_setup sep_safe = {*

View File

@ -13,10 +13,10 @@ imports Sep_MP
begin
lemma "((P \<and>* Q \<and>* R \<longrightarrow>* B) \<and>* P \<and>* A \<and>* R \<and>* Q) s \<Longrightarrow> (A \<and>* B) s"
(* sep_mp attempts to solve goals that could be solved by careful use of the sep_mp theorem *)
apply (sep_mp)
apply (clarsimp simp: sep_conj_ac)
done
(* sep_mp attempts to solve goals that could be solved by careful use of the sep_mp theorem *)
apply (sep_mp)
apply (clarsimp simp: sep_conj_ac)
done
(* Sep_mp uses the sep_cancel set to solve goals *)
@ -27,8 +27,8 @@ where Moo_Bar[sep_cancel] : "Moo s \<Longrightarrow> Bar s"
lemma "((Bar \<and>* Q \<and>* R \<longrightarrow>* B) \<and>* Moo \<and>* A \<and>* R \<and>* Q) s \<Longrightarrow> (A \<and>* B) s"
apply (sep_mp)
apply (clarsimp simp: sep_conj_ac)
done
apply (sep_mp)
apply (clarsimp simp: sep_conj_ac)
done
end

View File

@ -14,119 +14,106 @@ begin
(* Constrained lens for sep_erule tactic *)
lemma sep_asm_eq_erule:
"(P \<and>* R) s \<Longrightarrow> (\<And>s. T s = (P \<and>* R) s) \<Longrightarrow> (T s \<Longrightarrow> (P' \<and>* R') s)
\<Longrightarrow> (P' \<and>* R') s" by (clarsimp)
"(P \<and>* R) s \<Longrightarrow> (\<And>s. T s = (P \<and>* R) s) \<Longrightarrow> (T s \<Longrightarrow> (P' \<and>* R') s) \<Longrightarrow> (P' \<and>* R') s"
by (clarsimp)
lemma sep_rule:
"(\<And>s. T s \<Longrightarrow> P s) \<Longrightarrow> (T \<and>* R) s \<Longrightarrow> (P \<and>* R) s"
by (erule sep_conj_impl1)
by (rule sep_conj_impl1)
lemma sep_erule:
"(T \<and>* R') s \<Longrightarrow> (\<And>s. T s \<Longrightarrow> P s) \<Longrightarrow> (\<And>s. R' s \<Longrightarrow> R s) \<Longrightarrow> (P \<and>* R) s"
by (metis (full_types) sep_conj_impl)
by (rule sep_conj_impl)
(* Construct analogues to rule/drule etc, for separation logic *)
ML {*
ML {*
fun sep_select ctxt = resolve_tac ctxt [@{thm sep_eq}]
fun sep_asm_select ctxt = dresolve_tac ctxt [@{thm sep_asm_eq}]
fun sep_asm_erule_select ctxt = eresolve_tac ctxt [@{thm sep_asm_eq_erule}]
fun sep_select ctxt = (resolve_tac ctxt [@{thm sep_eq}])
fun sep_asm_select ctxt = (dresolve_tac ctxt [@{thm sep_asm_eq}])
fun sep_asm_erule_select ctxt = (eresolve_tac ctxt [@{thm sep_asm_eq_erule}])
fun sep_rule_tactic ctxt thms =
let val sep_rule = resolve_tac ctxt [@{thm sep_rule}]
in sep_apply_tactic ctxt sep_rule thms end
fun sep_drule_tactic ctxt thms =
let val sep_drule = dresolve_tac ctxt [rotate_prems ~1 @{thm sep_rule}]
in sep_apply_tactic ctxt sep_drule thms end
fun sep_rule_tactic ctxt thms =
let val sep_rule = (resolve_tac ctxt [@{thm sep_rule}]) in
sep_apply_tactic ctxt sep_rule thms
end;
fun sep_frule_tactic ctxt thms =
let val sep_frule = forward_tac ctxt [rotate_prems ~1 @{thm sep_rule}]
in sep_apply_tactic ctxt sep_frule thms end
fun sep_drule_tactic ctxt thms =
let val sep_drule = (dresolve_tac ctxt [rotate_prems ~1 @{thm sep_rule}]) in
sep_apply_tactic ctxt sep_drule thms
end;
fun sep_erule_tactic ctxt thms =
let val sep_erule = (eresolve_tac ctxt [@{thm sep_erule}])
in sep_apply_tactic ctxt sep_erule thms end
fun sep_frule_tactic ctxt thms =
let val sep_frule = (forward_tac ctxt [rotate_prems ~1 @{thm sep_rule}]) in
sep_apply_tactic ctxt sep_frule thms
end;
fun sep_rule_tac tac ctxt = rotator (sep_select ctxt) tac ctxt
fun sep_drule_tac tac ctxt = rotator (sep_asm_select ctxt) tac ctxt
fun sep_erule_tac tac ctxt = rotator (sep_asm_select ctxt) tac ctxt
fun sep_erule_concl_tac tac ctxt = rotator (sep_select ctxt) tac ctxt
fun sep_erule_tactic ctxt thms =
let val sep_erule = (eresolve_tac ctxt [@{thm sep_erule}]) in
sep_apply_tactic ctxt sep_erule thms
end;
fun sep_rule_tac tac ctxt =
rotator (sep_select ctxt) (tac) ctxt
fun sep_drule_tac tac ctxt =
rotator (sep_asm_select ctxt) tac ctxt
fun sep_erule_tac tac ctxt =
rotator (sep_asm_select ctxt) tac ctxt
fun sep_erule_concl_tac tac ctxt =
rotator (sep_select ctxt) tac ctxt
fun sep_erule_full_tac tac ctxt =
let val r = rotator' ctxt
in
tac |> r (sep_asm_erule_select ctxt) |> r (sep_select ctxt)
end;
fun sep_erule_full_tac tac ctxt =
let val r = rotator' ctxt
in
tac |> r (sep_asm_erule_select ctxt) |> r (sep_select ctxt)
end
fun sep_erule_full_tac' tac ctxt =
let val r = rotator' ctxt
in
tac |> r (sep_select ctxt) |> r (sep_asm_erule_select ctxt)
end;
fun sep_erule_full_tac' tac ctxt =
let val r = rotator' ctxt
in
tac |> r (sep_select ctxt) |> r (sep_asm_erule_select ctxt)
end
fun sep_rule_comb_tac true thms ctxt = sep_rule_tac (resolve_tac ctxt thms) ctxt
| sep_rule_comb_tac false thms ctxt = sep_rule_tac (sep_rule_tactic ctxt thms) ctxt
fun sep_rule_comb_tac true thms ctxt = (sep_rule_tac (resolve_tac ctxt thms) ctxt) |
sep_rule_comb_tac false thms ctxt = (sep_rule_tac (sep_rule_tactic ctxt thms) ctxt)
fun sep_rule_method bool thms ctxt = SIMPLE_METHOD' (sep_rule_comb_tac bool thms ctxt)
fun sep_rule_method bool thms ctxt = SIMPLE_METHOD' (sep_rule_comb_tac bool thms ctxt)
fun sep_drule_comb_tac true thms ctxt = sep_drule_tac (dresolve_tac ctxt thms) ctxt
| sep_drule_comb_tac false thms ctxt = sep_drule_tac (sep_drule_tactic ctxt thms) ctxt
fun sep_drule_comb_tac true thms ctxt = (sep_drule_tac (dresolve_tac ctxt thms) ctxt) |
sep_drule_comb_tac false thms ctxt = (sep_drule_tac (sep_drule_tactic ctxt thms) ctxt)
fun sep_drule_method bool thms ctxt = SIMPLE_METHOD' (sep_drule_comb_tac bool thms ctxt)
fun sep_drule_method bool thms ctxt = SIMPLE_METHOD' (sep_drule_comb_tac bool thms ctxt)
fun sep_frule_method true thms ctxt = SIMPLE_METHOD' (sep_drule_tac (forward_tac ctxt thms) ctxt)
| sep_frule_method false thms ctxt = SIMPLE_METHOD' (sep_drule_tac (sep_frule_tactic ctxt thms) ctxt)
fun sep_frule_method true thms ctxt = SIMPLE_METHOD' (sep_drule_tac (forward_tac ctxt thms) ctxt) |
sep_frule_method false thms ctxt = SIMPLE_METHOD' (sep_drule_tac (sep_frule_tactic ctxt thms) ctxt)
fun sep_erule_method true thms ctxt = SIMPLE_METHOD' (sep_erule_tac (eresolve_tac ctxt thms) ctxt)
| sep_erule_method false thms ctxt = SIMPLE_METHOD' (sep_erule_tac (sep_erule_tactic ctxt thms) ctxt)
fun sep_erule_method true thms ctxt = SIMPLE_METHOD' (sep_erule_tac (eresolve_tac ctxt thms) ctxt) |
sep_erule_method false thms ctxt = SIMPLE_METHOD' (sep_erule_tac (sep_erule_tactic ctxt thms) ctxt)
fun sep_erule_concl_method true thms ctxt =
SIMPLE_METHOD' (sep_erule_concl_tac (eresolve_tac ctxt thms) ctxt) |
sep_erule_concl_method false thms ctxt =
SIMPLE_METHOD' (sep_erule_concl_tac (sep_erule_tactic ctxt thms) ctxt)
fun sep_erule_concl_method true thms ctxt =
SIMPLE_METHOD' (sep_erule_concl_tac (eresolve_tac ctxt thms) ctxt)
| sep_erule_concl_method false thms ctxt =
SIMPLE_METHOD' (sep_erule_concl_tac (sep_erule_tactic ctxt thms) ctxt)
fun sep_erule_full_method true thms ctxt =
SIMPLE_METHOD' (sep_erule_full_tac (eresolve_tac ctxt thms) ctxt) |
sep_erule_full_method false thms ctxt =
SIMPLE_METHOD' (sep_erule_full_tac (sep_erule_tactic ctxt thms) ctxt)
SIMPLE_METHOD' (sep_erule_full_tac (eresolve_tac ctxt thms) ctxt)
| sep_erule_full_method false thms ctxt =
SIMPLE_METHOD' (sep_erule_full_tac (sep_erule_tactic ctxt thms) ctxt)
*}
method_setup sep_rule = {*
(Scan.lift (Args.mode "direct")) -- Attrib.thms >> uncurry sep_rule_method
Scan.lift (Args.mode "direct") -- Attrib.thms >> uncurry sep_rule_method
*}
method_setup sep_drule = {*
(Scan.lift (Args.mode "direct")) -- Attrib.thms >> uncurry sep_drule_method
Scan.lift (Args.mode "direct") -- Attrib.thms >> uncurry sep_drule_method
*}
method_setup sep_frule = {*
(Scan.lift (Args.mode "direct")) -- Attrib.thms >> uncurry sep_frule_method
Scan.lift (Args.mode "direct") -- Attrib.thms >> uncurry sep_frule_method
*}
method_setup sep_erule = {*
(Scan.lift (Args.mode "direct")) -- Attrib.thms >> uncurry sep_erule_method
Scan.lift (Args.mode "direct") -- Attrib.thms >> uncurry sep_erule_method
*}
method_setup sep_erule_concl = {*
(Scan.lift (Args.mode "direct")) -- Attrib.thms >> uncurry sep_erule_concl_method
Scan.lift (Args.mode "direct") -- Attrib.thms >> uncurry sep_erule_concl_method
*}
method_setup sep_erule_full = {*
(Scan.lift (Args.mode "direct")) -- Attrib.thms>> uncurry sep_erule_full_method
Scan.lift (Args.mode "direct") -- Attrib.thms>> uncurry sep_erule_full_method
*}
end

View File

@ -19,31 +19,30 @@ where Moo_Bar : "Moo s \<Longrightarrow> Bar s"
(* sep_rule is 'rule' with rotations of the conjuncts in the conclusions *)
lemma "(A \<and>* B \<and>* C \<and>* Bar) s"
apply (sep_rule Moo_Bar)
oops
apply (sep_rule Moo_Bar)
oops
(* sep_drule is 'drule' with rotations of the conjuncts in the assumptions *)
lemma "(A \<and>* B \<and>* C \<and>* Moo) s \<Longrightarrow> R"
apply (sep_drule Moo_Bar)
oops
apply (sep_drule Moo_Bar)
oops
(* sep_erule is 'erule' with rotations of the conjuncts in either the assumptions,
the conclusions, or both. These are sep_erule, sep_erule_concl, and sep_erule_full respectively
*)
lemma "(A \<and>* B \<and>* C \<and>* Moo) s \<Longrightarrow> (A \<and>* B \<and>* C \<and>* Bar) s"
(* In this case we need to rotate both, so we use sep_erule_full *)
apply (sep_erule_full Moo_Bar)
apply (assumption)
done
apply (sep_erule_full Moo_Bar)
apply (assumption)
done
lemma Moo_Bar_R: "(Moo \<and>* R) s \<Longrightarrow> (Bar \<and>* R) s" sorry
axiomatization where Moo_Bar_R: "(Moo \<and>* R) s \<Longrightarrow> (Bar \<and>* R) s"
(* When we have theorems with the frame explicitly mentioned, we have to invoke our tactics with
(direct) option *)
lemma "(A \<and>* B \<and>* C \<and>* Moo) s \<Longrightarrow> (A \<and>* B \<and>* C \<and>* Bar) s"
apply (sep_erule_full (direct) Moo_Bar_R)
done
apply (sep_erule_full (direct) Moo_Bar_R)
done
end

View File

@ -17,28 +17,26 @@ begin
ML {*
(* generic rotator *)
fun range lo hi =
let fun r lo = if lo > hi then [] else lo::r (lo+1)
in r lo
end;
fun range lo hi =
let
fun r lo = if lo > hi then [] else lo::r (lo+1)
in r lo end
fun rotator lens tactic ctxt i st =
let val len = case (Seq.pull ((lens THEN' (resolve0_tac [@{thm iffI}])) i st))
of NONE => 0 |
SOME (thm, _) => conj_length ctxt (Thm.cprem_of thm i)
val nums = range 1 len
val selector = sep_select_tactic lens
val tac' = map (fn x => selector [x] ctxt THEN' tactic) nums
in ((selector [1] ctxt) THEN' (FIRST' tac')) i st
end;
fun rotator lens tactic ctxt i st =
let
val len = case Seq.pull ((lens THEN' resolve0_tac [@{thm iffI}]) i st) of
NONE => 0
| SOME (thm, _) => conj_length ctxt (Thm.cprem_of thm i)
val nums = range 1 len
val selector = sep_select_tactic lens
val tac' = map (fn x => selector [x] ctxt THEN' tactic) nums
in
(selector [1] ctxt THEN' FIRST' tac') i st
end
fun rotator' ctxt lens tactic = rotator lens tactic ctxt
fun rotator' ctxt lens tactic = rotator lens tactic ctxt
fun sep_apply_tactic ctxt lens_tac thms =
lens_tac THEN'
eresolve_tac ctxt thms
fun sep_apply_tactic ctxt lens_tac thms = lens_tac THEN' eresolve_tac ctxt thms
*}
end

View File

@ -9,57 +9,58 @@
*)
theory Sep_Rule_Ext
imports Sep_Provers
Sep_Attribs
Sep_ImpI
Sep_MP
imports
Sep_Provers
Sep_Attribs
Sep_ImpI
Sep_MP
begin
ML {*
fun backwardise ctxt thm = SOME (backward ctxt thm) handle THM _ => NONE
fun sep_curry ctxt thm = SOME (sep_curry_inner ctxt thm) handle THM _ => NONE
fun backwardise ctxt thm = SOME (backward ctxt thm)
handle THM _ => NONE
fun make_sep_drule direct thms ctxt i =
let
val default = sep_drule_comb_tac direct
fun make_sep_rule_inner i thm =
let
val goal = i + Thm.nprems_of thm - 1
in
case sep_curry ctxt thm of
SOME thm' =>
(sep_drule_tac (fn i => sep_drule_tactic ctxt [thm'] i THEN
(sep_mp_solver ctxt THEN' (TRY o sep_flatten ctxt)) goal) ctxt) i
| NONE => default [thm] ctxt i
end
in
if direct then default thms ctxt i else FIRST (map (make_sep_rule_inner i) thms)
end
fun sep_curry ctxt thm = SOME (sep_curry_inner ctxt thm)
handle THM _ => NONE
fun make_sep_rule direct thms ctxt =
let
val default = sep_rule_comb_tac direct
fun make_sep_rule_inner thm =
case backwardise ctxt thm of
SOME thm' => sep_rule_comb_tac true [thm'] ctxt THEN'
REPEAT_ALL_NEW (sep_match_trivial_tac ctxt) THEN'
TRY o sep_flatten ctxt
| NONE => default [thm] ctxt
in
if direct then default thms ctxt else FIRST' (map make_sep_rule_inner thms)
end
fun make_sep_drule direct thms ctxt i =
let val default = sep_drule_comb_tac direct
fun make_sep_rule_inner i thm =
let val goal = i + Thm.nprems_of thm - 1 in
case sep_curry ctxt thm of (SOME thm') =>
(sep_drule_tac (fn i => sep_drule_tactic ctxt [thm'] i THEN
((sep_mp_solver ctxt) THEN' (TRY o sep_flatten ctxt)) goal ) ctxt) i|
NONE => default [thm] ctxt i
end;
in if direct then default thms ctxt i else FIRST (map (make_sep_rule_inner i) thms )
end;
fun make_sep_rule direct thms ctxt =
let val default = sep_rule_comb_tac direct
fun make_sep_rule_inner thm = case backwardise ctxt thm of
(SOME thm') => sep_rule_comb_tac true [thm'] ctxt THEN'
(REPEAT_ALL_NEW (sep_match_trivial_tac ctxt)) THEN'
(TRY o (sep_flatten ctxt)) |
NONE => default [thm] ctxt
in if direct then default thms ctxt else FIRST' (map make_sep_rule_inner thms)
end;
fun sep_rule_method direct thms ctxt = SIMPLE_METHOD' (make_sep_rule direct thms ctxt)
fun sep_drule_method direct thms ctxt = SIMPLE_METHOD' (make_sep_drule direct thms ctxt)
fun sep_rule_method direct thms ctxt = SIMPLE_METHOD' (make_sep_rule direct thms ctxt)
fun sep_drule_method direct thms ctxt = SIMPLE_METHOD' (make_sep_drule direct thms ctxt)
*}
method_setup sep_rule = {*
(Scan.lift (Args.mode "direct")) -- Attrib.thms >> uncurry sep_rule_method
Scan.lift (Args.mode "direct") -- Attrib.thms >> uncurry sep_rule_method
*}
method_setup sep_drule = {*
(Scan.lift (Args.mode "direct")) -- Attrib.thms >> uncurry sep_drule_method
Scan.lift (Args.mode "direct") -- Attrib.thms >> uncurry sep_drule_method
*}
end

View File

@ -15,67 +15,55 @@ begin
ML_file "sep_tactics.ML"
ML{*
structure SepSelect_Rules = Named_Thms (
val name = @{binding "sep_select"};
val description = "sep_select rules";
);
structure SepSelect_Rules = Named_Thms (
val name = @{binding "sep_select"}
val description = "sep_select rules"
)
*}
setup SepSelect_Rules.setup
ML {*
structure SepSelectAsm_Rules = Named_Thms (
val name = @{binding "sep_select_asm"};
val description = "sep_select_asm rules";
);
val name = @{binding "sep_select_asm"}
val description = "sep_select_asm rules"
)
*}
setup SepSelectAsm_Rules.setup
ML {*
fun sep_selects_tactic ns ctxt =
let val thms = SepSelect_Rules.get ctxt in
sep_select_tactic (resolve_tac ctxt thms) ns ctxt
end;
sep_select_tactic (resolve_tac ctxt (SepSelect_Rules.get ctxt)) ns ctxt
fun sep_select_asms_tactic ns ctxt =
let val thms = SepSelectAsm_Rules.get ctxt in
sep_select_tactic (dresolve_tac ctxt thms) ns ctxt
end;
fun sep_select_asms_method ns ctxt = SIMPLE_METHOD'
(sep_select_asms_tactic ns ctxt)
fun sep_selects_method ns ctxt = SIMPLE_METHOD'
(sep_selects_tactic ns ctxt)
sep_select_tactic (dresolve_tac ctxt (SepSelectAsm_Rules.get ctxt)) ns ctxt
*}
lemma sep_eq: "(\<And>s. T s = (P \<and>* R) s) \<Longrightarrow> T s \<Longrightarrow> (P \<and>* R) s" by clarsimp
lemma sep_asm_eq: "(P \<and>* R) s \<Longrightarrow> (\<And>s. T s = (P \<and>* R) s) \<Longrightarrow> T s" by clarsimp
declare sep_eq[sep_select]
declare sep_asm_eq[sep_select_asm]
method_setup sep_select_asm = {*
Scan.lift (Scan.repeat Parse.int) >> sep_select_asms_method
Scan.lift (Scan.repeat Parse.int) >>
(fn ns => fn ctxt => SIMPLE_METHOD' (sep_select_asms_tactic ns ctxt))
*} "Reorder assumptions"
method_setup sep_select = {*
Scan.lift (Scan.repeat Parse.int) >> sep_selects_method
Scan.lift (Scan.repeat Parse.int) >>
(fn ns => fn ctxt => SIMPLE_METHOD' (sep_selects_tactic ns ctxt))
*} "Reorder conclusions"
lemma sep_eq [sep_select]: "(\<And>s. T s = (P \<and>* R) s) \<Longrightarrow> T s \<Longrightarrow> (P \<and>* R) s" by clarsimp
lemma sep_asm_eq [sep_select_asm]: "(P \<and>* R) s \<Longrightarrow> (\<And>s. T s = (P \<and>* R) s) \<Longrightarrow> T s" by clarsimp
ML {*
fun sep_select_generic_method_inner ((lens, n), asm) ctxt =
let val lens_tac = (if asm then (dresolve_tac ctxt lens) else (resolve_tac ctxt lens))
in sep_select_method lens_tac n ctxt
end;
(* export method form of these two for use outisde this entry *)
fun sep_select_method lens ns ctxt =
SIMPLE_METHOD' (sep_select_tactic lens ns ctxt)
fun sep_select_generic_method asm thms ns ctxt =
sep_select_generic_method_inner ((thms,ns),asm) ctxt
sep_select_method (if asm then dresolve_tac ctxt thms else resolve_tac ctxt thms) ns ctxt
*}
method_setup sep_select_gen = {*
Attrib.thms --| (Scan.lift Args.colon) -- Scan.lift (Scan.repeat Parse.int) -- (Scan.lift (Args.mode "asm")) >>
sep_select_generic_method_inner
Attrib.thms --| Scan.lift Args.colon -- Scan.lift (Scan.repeat Parse.int) -- Scan.lift (Args.mode "asm") >>
(fn ((lens, ns), asm) => sep_select_generic_method asm lens ns)
*}
end

View File

@ -17,6 +17,6 @@ lemma sep_select: "(A \<and>* B \<and>* C \<and>* D) s \<Longrightarrow> (A \<an
apply (sep_select 2 3 4 1) (* we can reorder multiple conjuncts at once *)
apply (sep_select 2 3 4) (* the list can be partial *)
apply (sep_select_asm 3 4) (* sep_select_asm works for the assumptions *)
oops
oops
end

View File

@ -12,34 +12,37 @@ theory Sep_Solve
imports Sep_Cancel Sep_MP
begin
thm sep_conj_sep_impl
ML {*
fun sep_schem ctxt = rotator' ctxt (sep_asm_erule_select ctxt)
(SOLVED' ( (eresolve0_tac [@{thm sep_conj_sep_impl2}] THEN'
(FIRST' [assume_tac ctxt, resolve0_tac [@{thm TrueI}], sep_cancel_tactic' ctxt true] |> REPEAT_ALL_NEW)) ))
fun sep_schem ctxt =
rotator' ctxt (sep_asm_erule_select ctxt)
(SOLVED' ((eresolve0_tac [@{thm sep_conj_sep_impl2}] THEN'
(FIRST' [assume_tac ctxt, resolve0_tac [@{thm TrueI}], sep_cancel_tactic' ctxt true]
|> REPEAT_ALL_NEW))))
fun sep_solve_tactic ctxt =
let val truei = resolve0_tac [@{thm TrueI}]
fun sep_cancel_rotating i = sep_select_tactic (sep_asm_select ctxt) [1] ctxt i THEN_ELSE
(rotator' ctxt (sep_asm_select ctxt) (FIRST' [assume_tac ctxt, truei, sep_cancel_tactic' ctxt false, eresolve0_tac [@{thm sep_conj_sep_impl}]] |> REPEAT_ALL_NEW |> SOLVED' ) i,
SOLVED' (FIRST' [assume_tac ctxt, truei, sep_cancel_tactic' ctxt false, eresolve0_tac [@{thm sep_conj_sep_impl}]] |> REPEAT_ALL_NEW) i)
val sep_cancel_tac = (FIRST' [assume_tac ctxt, truei, sep_cancel_tactic' ctxt false, eresolve0_tac [@{thm sep_conj_sep_impl}]] |> REPEAT_ALL_NEW)
in (DETERM o SOLVED' (FIRST' [assume_tac ctxt,truei, sep_cancel_tac])) ORELSE'
(SOLVED' (((TRY o CHANGED_PROP o (sep_mp_solver ctxt)) THEN_ALL_NEW sep_cancel_rotating) )) |> SOLVED'
end;
fun sep_solve_method _ ctxt = SIMPLE_METHOD' ( sep_solve_tactic ctxt )
fun sep_schem_method _ ctxt = SIMPLE_METHOD' ( sep_schem ctxt )
fun sep_solve_tactic ctxt =
let
val truei = resolve0_tac [@{thm TrueI}]
fun sep_cancel_rotating i =
sep_select_tactic (sep_asm_select ctxt) [1] ctxt i THEN_ELSE
(rotator' ctxt (sep_asm_select ctxt)
(FIRST' [assume_tac ctxt, truei, sep_cancel_tactic' ctxt false, eresolve0_tac [@{thm sep_conj_sep_impl}]]
|> REPEAT_ALL_NEW |> SOLVED') i,
SOLVED' (FIRST' [assume_tac ctxt, truei, sep_cancel_tactic' ctxt false, eresolve0_tac [@{thm sep_conj_sep_impl}]]
|> REPEAT_ALL_NEW) i)
val sep_cancel_tac =
FIRST' [assume_tac ctxt, truei, sep_cancel_tactic' ctxt false, eresolve0_tac [@{thm sep_conj_sep_impl}]]
|> REPEAT_ALL_NEW
in
(DETERM o SOLVED' (FIRST' [assume_tac ctxt,truei, sep_cancel_tac])) ORELSE'
(SOLVED' ((TRY o CHANGED_PROP o sep_mp_solver ctxt) THEN_ALL_NEW sep_cancel_rotating))
|> SOLVED'
end
fun sep_solve_method _ ctxt = SIMPLE_METHOD' (sep_solve_tactic ctxt)
fun sep_schem_method _ ctxt = SIMPLE_METHOD' (sep_schem ctxt)
*}
method_setup sep_solve = {*
sep_cancel_syntax >> sep_solve_method
*}
method_setup sep_schem = {*
sep_cancel_syntax >> sep_schem_method
*}
method_setup sep_solve = {* sep_cancel_syntax >> sep_solve_method *}
method_setup sep_schem = {* sep_cancel_syntax >> sep_schem_method *}
end

View File

@ -20,12 +20,13 @@ axiomatization
where Moo_Bar[sep_cancel] : "Moo s \<Longrightarrow> Bar s"
lemma "((Bar \<and>* Q \<and>* R \<longrightarrow>* B) \<and>* Moo \<and>* A \<and>* R \<and>* Q) s \<Longrightarrow> (A \<and>* B) s"
apply (sep_solve)
done
apply (sep_solve)
done
(* encouraging better proof style with different command for schematics in assumption *)
schematic_lemma "((Bar \<and>* Q \<and>* R \<longrightarrow>* B) \<and>* Moo \<and>* A \<and>* R \<and>* ?Q) s \<Longrightarrow> (A \<and>* B) s"
schematic_goal "((Bar \<and>* Q \<and>* R \<longrightarrow>* B) \<and>* Moo \<and>* A \<and>* R \<and>* ?Q) s \<Longrightarrow> (A \<and>* B) s"
apply (sep_schem)
done
done
end

View File

@ -15,34 +15,31 @@ begin
lemmas sep_curry = sep_conj_sep_impl[rotated]
lemma sep_mp: "((Q \<longrightarrow>* R) \<and>* Q) s \<Longrightarrow> R s"
by (clarsimp simp: sep_conj_def sep_impl_def)
by (rule sep_conj_sep_impl2)
lemma sep_mp_frame: "((Q \<longrightarrow>* R) \<and>* Q \<and>* R') s \<Longrightarrow> (R \<and>* R') s"
apply (clarsimp simp: sep_conj_assoc[symmetric])
apply (erule sep_conj_impl)
apply (erule sep_mp)
apply (assumption)
done
apply (clarsimp simp: sep_conj_assoc[symmetric])
apply (erule sep_conj_impl)
apply (erule (1) sep_mp)
done
lemma sep_empty_conj: "P s \<Longrightarrow> (\<box> \<and>* P) s"
apply (clarsimp)
done
by clarsimp
lemma sep_conj_empty: "(\<box> \<and>* P) s \<Longrightarrow> P s"
apply (clarsimp)
done
by clarsimp
lemma sep_empty_imp: "(\<box> \<longrightarrow>* P) s \<Longrightarrow> P s"
apply (clarsimp simp: sep_impl_def)
apply (erule_tac x=0 in allE)
apply (clarsimp)
done
done
lemma sep_empty_imp': "(\<box> \<longrightarrow>* P) s \<Longrightarrow> (\<And>s. P s \<Longrightarrow> Q s) \<Longrightarrow> Q s"
apply (clarsimp simp: sep_impl_def)
apply (erule_tac x=0 in allE)
apply (clarsimp)
done
done
lemma sep_imp_empty: " P s \<Longrightarrow> (\<And>s. P s \<Longrightarrow> Q s) \<Longrightarrow> (\<box> \<longrightarrow>* Q) s"
by (erule sep_conj_sep_impl, clarsimp)

View File

@ -13,7 +13,7 @@
Rafal Kolanski <rafal.kolanski at nicta.com.au>
*)
header "Abstract Separation Algebra"
chapter "Abstract Separation Algebra"
theory Separation_Algebra
imports
@ -146,7 +146,7 @@ lemma sep_add_zero_sym [simp]: "0 + x = x"
by (simp add: sep_add_commute)
lemma sep_disj_addD2: "\<lbrakk> x ## y + z; y ## z \<rbrakk> \<Longrightarrow> x ## z"
by (metis sep_disj_addD1 sep_add_ac)
by (metis sep_add_commute sep_disj_addD1 sep_disj_commuteI)
lemma sep_disj_addD: "\<lbrakk> x ## y + z; y ## z \<rbrakk> \<Longrightarrow> x ## y \<and> x ## z"
by (metis sep_disj_addD1 sep_disj_addD2)
@ -156,19 +156,19 @@ lemma sep_add_disjD: "\<lbrakk> x + y ## z; x ## y \<rbrakk> \<Longrightarrow> x
lemma sep_disj_addI2:
"\<lbrakk> x ## y + z; y ## z \<rbrakk> \<Longrightarrow> x + z ## y"
by (metis sep_add_ac sep_disj_addI1)
using sep_add_commute sep_disj_addI1 sep_disj_commuteI by presburger
lemma sep_add_disjI1:
"\<lbrakk> x + y ## z; x ## y \<rbrakk> \<Longrightarrow> x + z ## y"
by (metis sep_add_ac sep_add_disjD sep_disj_addI2)
by (metis sep_add_commute sep_disj_addI1 sep_disj_commuteI sep_add_disjD)
lemma sep_add_disjI2:
"\<lbrakk> x + y ## z; x ## y \<rbrakk> \<Longrightarrow> z + y ## x"
by (metis sep_add_ac sep_add_disjD sep_disj_addI2)
by (metis sep_add_commute sep_disj_addI1 sep_disj_commuteI sep_add_disjD)
lemma sep_disj_addI3:
"x + y ## z \<Longrightarrow> x ## y \<Longrightarrow> x ## y + z"
by (metis sep_add_ac sep_add_disjD sep_add_disjI2)
"x + y ## z \<Longrightarrow> x ## y \<Longrightarrow> x ## y + z"
by (metis sep_add_commute sep_disj_addI1 sep_disj_commuteI sep_add_disjD)
lemma sep_disj_add:
"\<lbrakk> y ## z; x ## y \<rbrakk> \<Longrightarrow> x ## y + z = x + y ## z"
@ -284,11 +284,11 @@ lemma sep_conj_sep_true':
lemma sep_conj_true [simp]:
"(sep_true ** sep_true) = sep_true"
unfolding sep_conj_def
by (auto intro!: ext intro: disjoint_subheaps_exist)
by (auto intro: disjoint_subheaps_exist)
lemma sep_conj_false_right [simp]:
"(P ** sep_false) = sep_false"
by (force elim: sep_conjE intro!: ext)
by (force elim: sep_conjE)
lemma sep_conj_false_left [simp]:
"(sep_false ** P) = sep_false"
@ -323,7 +323,7 @@ lemma sep_conj_true_P [simp]:
lemma sep_conj_disj:
"((P or Q) ** R) = ((P ** R) or (Q ** R))"
by (auto simp: sep_conj_def intro!: ext)
by (rule ext, auto simp: sep_conj_def)
lemma sep_conj_sep_true_left:
"(P ** Q) h \<Longrightarrow> (sep_true ** Q) h"
@ -343,11 +343,11 @@ lemma sep_conj_conj:
lemma sep_conj_exists1:
"((EXS x. P x) ** Q) = (EXS x. (P x ** Q))"
by (force intro!: ext intro: sep_conjI elim: sep_conjE)
by (force intro: sep_conjI elim: sep_conjE)
lemma sep_conj_exists2:
"(P ** (EXS x. Q x)) = (EXS x. P ** Q x)"
by (force intro!: sep_conjI ext elim!: sep_conjE)
by (force intro!: sep_conjI elim!: sep_conjE)
lemmas sep_conj_exists = sep_conj_exists1 sep_conj_exists2
@ -378,11 +378,11 @@ lemma sep_implE:
lemma sep_impl_sep_true [simp]:
"(P \<longrightarrow>* sep_true) = sep_true"
by (force intro!: sep_implI ext)
by (force intro!: sep_implI)
lemma sep_impl_sep_false [simp]:
"(sep_false \<longrightarrow>* P) = sep_true"
by (force intro!: sep_implI ext)
by (force intro!: sep_implI)
lemma sep_impl_sep_true_P:
"(sep_true \<longrightarrow>* P) h \<Longrightarrow> P h"
@ -390,7 +390,7 @@ lemma sep_impl_sep_true_P:
lemma sep_impl_sep_true_false [simp]:
"(sep_true \<longrightarrow>* sep_false) = sep_false"
by (force intro!: ext dest: sep_impl_sep_true_P)
by (force dest: sep_impl_sep_true_P)
lemma sep_conj_sep_impl:
"\<lbrakk> P h; \<And>h. (P ** Q) h \<Longrightarrow> R h \<rbrakk> \<Longrightarrow> (Q \<longrightarrow>* R) h"
@ -427,7 +427,7 @@ lemma pure_sep_false:
lemma pure_split:
"pure P = (P = sep_true \<or> P = sep_false)"
by (force simp: pure_def intro!: ext)
by (force simp: pure_def)
lemma pure_sep_conj:
"\<lbrakk> pure P; pure Q \<rbrakk> \<Longrightarrow> pure (P \<and>* Q)"
@ -593,15 +593,14 @@ proof (rule intuitionisticI)
qed
lemma strongest_intuitionistic:
"\<not> (\<exists>Q. (\<forall>h. (Q h \<longrightarrow> (P \<and>* sep_true) h)) \<and> intuitionistic Q \<and>
Q \<noteq> (P \<and>* sep_true) \<and> (\<forall>h. P h \<longrightarrow> Q h))"
by (fastforce intro!: ext sep_substate_disj_add
dest!: sep_conjD intuitionisticD)
"\<not>(\<exists>Q. (\<forall>h. (Q h \<longrightarrow> (P \<and>* sep_true) h)) \<and> intuitionistic Q \<and> Q \<noteq> (P \<and>* sep_true) \<and> (\<forall>h. P h \<longrightarrow> Q h))"
by (fastforce intro!: ext sep_substate_disj_add dest!: sep_conjD intuitionisticD)
lemma weakest_intuitionistic:
"\<not> (\<exists>Q. (\<forall>h. ((sep_true \<longrightarrow>* P) h \<longrightarrow> Q h)) \<and> intuitionistic Q \<and>
Q \<noteq> (sep_true \<longrightarrow>* P) \<and> (\<forall>h. Q h \<longrightarrow> P h))"
apply (clarsimp intro!: ext)
apply (clarsimp)
apply (rule ext)
apply (rule iffI)
apply (rule sep_implI)
apply (drule_tac h="x" and h'="x + h'" in intuitionisticD)
@ -614,7 +613,7 @@ lemma intuitionistic_sep_conj_sep_true_P:
lemma intuitionistic_sep_conj_sep_true_simp:
"intuitionistic P \<Longrightarrow> (P \<and>* sep_true) = P"
by (fast intro!: sep_conj_sep_true ext
by (fast intro!: sep_conj_sep_true
elim: intuitionistic_sep_conj_sep_true_P)
lemma intuitionistic_sep_impl_sep_true_P:
@ -624,8 +623,7 @@ lemma intuitionistic_sep_impl_sep_true_P:
lemma intuitionistic_sep_impl_sep_true_simp:
"intuitionistic P \<Longrightarrow> (sep_true \<longrightarrow>* P) = P"
by (fast intro!: ext
elim: sep_impl_sep_true_P intuitionistic_sep_impl_sep_true_P)
by (fast elim: sep_impl_sep_true_P intuitionistic_sep_impl_sep_true_P)
subsection {* Strictly exact assertions *}
@ -666,7 +664,7 @@ begin
lemma sep_disj_add_eq [simp]: "x ## y \<Longrightarrow> x + y ## z = (x ## z \<and> y ## z)"
by (metis sep_add_disj_eq sep_disj_commute)
subclass sep_algebra by default auto
subclass sep_algebra by standard auto
end
@ -705,17 +703,14 @@ consts
notation (latex output) sep_conj_lifted ("\<And>\<^sup>* _" [60] 90)
notation (latex output) sep_map_list_conj ("\<And>\<^sup>* _" [60] 90)
adhoc_overloading
sep_conj_lifted sep_list_conj
adhoc_overloading
sep_conj_lifted sep_set_conj
adhoc_overloading sep_conj_lifted sep_list_conj
adhoc_overloading sep_conj_lifted sep_set_conj
(* FIXME. Add notation for sep_map_list_conj, and consider unifying with sep_map_set_conj. *)
text{* Now: lot's of fancy syntax. First, @{term "sep_map_set_conj (%x. g) A"} is
written @{text"\<And>+x\<in>A. g"}. *}
text{* Now: lots of fancy syntax. First, @{term "sep_map_set_conj (%x. g) A"} is written @{text"\<And>+x\<in>A. g"}. *}
(* Clagged from Big_Operators. *)
syntax
@ -731,8 +726,7 @@ translations -- {* Beware of argument permutation! *}
"SETSEPCONJ x:A. g" == "CONST sep_map_set_conj (%x. g) A"
"\<And>* x\<in>A. g" == "CONST sep_map_set_conj (%x. g) A"
text{* Instead of @{term"\<And>*x\<in>{x. P}. g"} we introduce the shorter
@{text"\<And>+x|P. g"}. *}
text{* Instead of @{term"\<And>*x\<in>{x. P}. g"} we introduce the shorter @{text"\<And>+x|P. g"}. *}
syntax
"_qsep_map_set_conj" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3SETSEPCONJ _ |/ _./ _)" [0,0,10] 10)
@ -764,18 +758,9 @@ in [(@{const_syntax sep_map_set_conj}, K setsepconj_tr')] end
*}
(* FIXME. This is not needed, but gives lemmas directly on sep_map_set_conj,
rather than needing to expand the definition and use rules about setprod.
*)
interpretation sep: comm_monoid_mult "op \<and>*" \<box>
by unfold_locales (simp add: sep.setprod_def)
interpretation sep: folding "op \<and>*" \<box>
by unfold_locales (simp add: comp_def sep_conj_ac)
print_theorems (* FIXME, remove. We get all of these lemmas. *)
lemma "\<And>* [\<box>,P] = P"
by (simp add: sep_list_conj_def)
@ -783,9 +768,7 @@ lemma "\<And>* {\<box>} = \<box>"
by (simp add: sep_set_conj_def)
lemma "\<And>* {P,\<box>} = P"
apply (cases "P = \<box>")
apply (auto simp: sep_set_conj_def)
done
by (cases "P = \<box>", auto simp: sep_set_conj_def)
lemma "(\<And>* x\<in>{0,1::nat}. if x=0 then \<box> else P) = P"
by auto
@ -799,39 +782,29 @@ lemma sep_list_conj_Nil [simp]: "\<And>* [] = \<box>"
(* apparently these two are rarely used and had to be removed from List.thy *)
lemma (in semigroup) foldl_assoc:
shows "(foldl f (f x y) zs) = (f x (foldl f y zs))"
by (induct zs arbitrary: y) (simp_all add:assoc)
"foldl f (f x y) zs = f x (foldl f y zs)"
by (induct zs arbitrary: y) (simp_all add:assoc)
lemma (in monoid) foldl_absorb1:
shows "f x (foldl f 1 zs) = foldl f x zs"
by (induct zs) (simp_all add:foldl_assoc)
"f x (foldl f z zs) = foldl f x zs"
by (induct zs) (simp_all add:foldl_assoc)
lemma sep_list_conj_Cons [simp]: "\<And>* (x#xs) = (x ** \<And>* xs)"
by (simp add: sep_list_conj_def sep.foldl_absorb1)
context comm_monoid
begin
lemma sep_list_conj_append [simp]: "\<And>* (xs @ ys) = (\<And>* xs ** \<And>* ys)"
by (simp add: sep_list_conj_def sep.foldl_absorb1)
lemma sep_list_conj_map_append:
"\<And>* map f (xs @ ys) = (\<And>* map f xs \<and>* \<And>* map f ys)"
by (metis map_append sep_list_conj_append)
lemma (in comm_monoid) foldl_map_filter:
"foldl f 1 (map P (filter t xs)) *
foldl f 1 (map P (filter (not t) xs))
= foldl f 1 (map P xs)"
lemma foldl_map_filter:
"f (foldl f z (map P (filter t xs))) (foldl f z (map P (filter (not t) xs))) = foldl f z (map P xs)"
proof (induct xs)
case Nil thus ?case by clarsimp
next
case (Cons x xs)
hence IH: "foldl f 1 (map P xs) =
foldl f 1 (map P (filter t xs)) *
foldl f 1 (map P [x\<leftarrow>xs . \<not> t x])"
by (simp only: eq_commute)
hence IH:
"foldl f z (map P xs) = f (foldl f z (map P (filter t xs))) (foldl f z (map P [x\<leftarrow>xs . \<not> t x]))"
by (simp only: eq_commute)
have foldl_Cons':
"\<And>x xs. foldl f 1 (x # xs) = f x (foldl f 1 xs)"
"\<And>x xs. foldl f z (x # xs) = f x (foldl f z xs)"
by (simp, subst foldl_absorb1[symmetric], rule refl)
{ assume "t x"
@ -843,6 +816,31 @@ next
ultimately show ?case by blast
qed
lemma foldl_map_add:
"foldl f z (map (\<lambda>x. f (P x) (Q x)) xs) = f (foldl f z (map P xs)) (foldl f z (map Q xs))"
apply (induct xs)
apply clarsimp
apply simp
by (metis (full_types) commute foldl_absorb1 foldl_assoc)
lemma foldl_map_remove1:
"x \<in> set xs \<Longrightarrow> foldl f z (map P xs) = f (P x) (foldl f z (map P (remove1 x xs)))"
apply (induction xs, simp)
apply clarsimp
by (metis foldl_absorb1 left_commute)
end
lemma sep_list_conj_Cons [simp]: "\<And>* (x#xs) = (x ** \<And>* xs)"
by (simp add: sep_list_conj_def sep.foldl_absorb1)
lemma sep_list_conj_append [simp]: "\<And>* (xs @ ys) = (\<And>* xs ** \<And>* ys)"
by (simp add: sep_list_conj_def sep.foldl_absorb1)
lemma sep_list_conj_map_append:
"\<And>* map f (xs @ ys) = (\<And>* map f xs \<and>* \<And>* map f ys)"
by (metis map_append sep_list_conj_append)
lemma sep_list_con_map_filter:
"(\<And>* map P (filter t xs) \<and>* \<And>* map P (filter (not t) xs))
= \<And>* map P xs"
@ -861,13 +859,6 @@ lemma sep_map_set_conj_restrict:
sep_map_set_conj P {x \<in> xs. \<not> t x})"
by (subst sep.setprod.union_disjoint [symmetric], (fastforce simp: union_filter)+)
lemma (in comm_monoid) foldl_map_add:
"foldl f 1 (map (\<lambda>x. P x * Q x) xs) =
(foldl f 1 (map P xs) * foldl f 1 (map Q xs))"
apply (induct xs)
apply clarsimp
apply simp
by (metis (full_types) commute foldl_absorb1 foldl_assoc)
lemma sep_list_conj_map_add:
"\<And>* map (\<lambda>x. f x \<and>* g x) xs = (\<And>* map f xs \<and>* \<And>* map g xs)"
@ -893,12 +884,6 @@ lemma remove1_filter:
apply clarsimp
done
lemma (in comm_monoid) foldl_map_remove1:
"x \<in> set xs \<Longrightarrow> foldl f 1 (map P xs) = f (P x) (foldl f 1 (map P (remove1 x xs)))"
apply (induction xs, simp)
apply clarsimp
by (metis foldl_absorb1 left_commute)
lemma sep_list_conj_map_remove1:
"x \<in> set xs \<Longrightarrow> \<And>* map P xs = (P x \<and>* \<And>* map P (remove1 x xs))"
apply (simp add: sep_list_conj_def)
@ -916,13 +901,11 @@ lemma sep_conj_map_split:
by (metis list.map(2) map_append sep_list_conj_Cons sep_list_conj_append)
(*
* Separation predicates on sets. *
*)
section "Separation predicates on sets"
lemma sep_map_set_conj_cong:
"\<lbrakk>P = Q; xs = ys\<rbrakk> \<Longrightarrow> sep_map_set_conj P xs = sep_map_set_conj Q ys"
by (simp)
by simp
lemma sep_set_conj_empty [simp]:
"sep_set_conj {} = \<box>"
@ -1017,15 +1000,8 @@ lemma sep_list_conj_filter_map:
by (clarsimp simp: sep_list_conj_def foldl_use_filter_map)
lemma sep_map_set_conj_restrict_predicate:
"finite A
\<Longrightarrow> (\<And>* x\<in>A. if T x then P x else \<box>) =
(\<And>* x\<in>(Set.filter T A). P x)"
apply (drule finite_distinct_list)
apply (clarsimp)
apply (subst filter_set)
apply (subst sep_list_conj_sep_map_set_conj [symmetric], simp)+
apply (rule sep_list_conj_filter_map)
done
"finite A \<Longrightarrow> (\<And>* x\<in>A. if T x then P x else \<box>) = (\<And>* x\<in>(Set.filter T A). P x)"
by (simp add: Set.filter_def sep.setprod.inter_filter)
lemma distinct_filters:
"\<lbrakk>distinct xs; \<And>x. (f x \<and> g x) = False\<rbrakk> \<Longrightarrow>
@ -1049,13 +1025,11 @@ lemma sep_map_set_conj_set_disjoint:
apply simp
by (metis Collect_disj_eq)
section {* Separation Algebra with a Cancellative Monoid (for completeness) *}
section {* Separation Algebra with a Cancellative Monoid *}
text {*
Separation algebra with a cancellative monoid. The results of being a precise
assertion (distributivity over separating conjunction) require this.
although we never actually use this property in our developments, we keep
it here for completeness.
*}
class cancellative_sep_algebra = sep_algebra +
assumes sep_add_cancelD: "\<lbrakk> x + z = y + z ; x ## z ; y ## z \<rbrakk> \<Longrightarrow> x = y"

View File

@ -13,7 +13,7 @@
Rafal Kolanski <rafal.kolanski at nicta.com.au>
*)
header "Abstract Separation Logic, Alternative Definition"
chapter "Abstract Separation Logic, Alternative Definition"
theory Separation_Algebra_Alt
imports "~~/src/HOL/Main"
@ -135,7 +135,7 @@ lemma sep_conj_comI:
lemma sep_conj_com:
"P ** Q = Q ** P"
by (auto intro: sep_conj_comI intro!: ext)
by (auto intro: sep_conj_comI)
lemma lift_to_add2:
"\<lbrakk>z \<oplus> q = Some s; x \<oplus> y = Some q\<rbrakk> \<Longrightarrow> Some z ++ Some x ++ Some y = Some s"
@ -186,11 +186,11 @@ lemma disjoint_submaps_exist:
lemma sep_conj_true[simp]:
"(sep_true ** sep_true) = sep_true"
unfolding sep_conj_def
by (auto intro!: ext intro: disjoint_submaps_exist)
by (auto intro: disjoint_submaps_exist)
lemma sep_conj_false_right[simp]:
"(P ** sep_false) = sep_false"
by (force elim: sep_conjE intro!: ext)
by (force elim: sep_conjE)
lemma sep_conj_false_left[simp]:
"(sep_false ** P) = sep_false"
@ -239,7 +239,7 @@ lemma sep_conj_exists1:
lemma sep_conj_exists2:
"(P ** (\<lambda>s. \<exists>x. Q x s)) = (\<lambda>s. (\<exists>x. (P ** Q x) s))"
by (force intro!: sep_conjI ext elim!: sep_conjE)
by (force intro!: sep_conjI elim!: sep_conjE)
lemmas sep_conj_exists = sep_conj_exists1 sep_conj_exists2
@ -280,11 +280,11 @@ lemma sep_implD:
lemma sep_impl_sep_true[simp]:
"(P \<longrightarrow>\<^sup>* sep_true) = sep_true"
by (force intro!: sep_implI ext)
by (force intro!: sep_implI)
lemma sep_impl_sep_false[simp]:
"(sep_false \<longrightarrow>\<^sup>* P) = sep_true"
by (force intro!: sep_implI ext)
by (force intro!: sep_implI)
lemma sep_impl_sep_true_P:
"(sep_true \<longrightarrow>\<^sup>* P) s \<Longrightarrow> P s"
@ -295,7 +295,7 @@ lemma sep_impl_sep_true_P:
lemma sep_impl_sep_true_false[simp]:
"(sep_true \<longrightarrow>\<^sup>* sep_false) = sep_false"
by (force intro!: ext dest: sep_impl_sep_true_P)
by (force dest: sep_impl_sep_true_P)
lemma sep_conj_sep_impl:
"\<lbrakk> P s; \<And>s. (P ** Q) s \<Longrightarrow> R s \<rbrakk> \<Longrightarrow> (Q \<longrightarrow>\<^sup>* R) s"

View File

@ -23,7 +23,13 @@ typedecl p
typedecl val
typedecl heap
arities heap :: sep_algebra
axiomatization where
heap_algebra: "OFCLASS(heap, sep_algebra_class)"
instantiation heap :: sep_algebra
begin
instance by (rule heap_algebra)
end
axiomatization
points_to :: "p \<Rightarrow> val \<Rightarrow> heap \<Rightarrow> bool" and
@ -34,7 +40,7 @@ where
consts
update :: "p \<Rightarrow> val \<Rightarrow> heap \<Rightarrow> heap"
(*
(* FIXME: revive
lemma
"\<lbrakk> Q2 (val h p); (K ** T ** blub ** P ** points_to p v ** P ** J) h \<rbrakk>
\<Longrightarrow> Q (val h p) (val h p)"
@ -58,22 +64,13 @@ lemma
apply simp
oops
schematic_lemma
schematic_goal
assumes a: "\<And>P. (stuff p ** P) H \<Longrightarrow> (other_stuff p v ** P) (update p v H)"
shows "(X ** Y ** other_stuff p ?v) (update p v H)"
apply (sep_rule (direct) a)
oops
text {* Example of low-level rewrites *}
lemma "\<lbrakk> unrelated s ; (P ** Q ** R) s \<rbrakk> \<Longrightarrow> (A ** B ** Q ** P) s"
apply (tactic {* dtac (mk_sep_select_rule @{context} true (3,[1])) 1 *})
apply (tactic {* rtac (mk_sep_select_rule @{context} false (4,[2])) 1 *})
(* now sep_conj_impl1 can be used *)
apply (erule (1) sep_conj_impl)
oops
text {* Conjunct selection *}
@ -92,7 +89,7 @@ section {* Test cases for @{text sep_cancel}. *}
lemma
assumes forward: "\<And>s g p v. A g p v s \<Longrightarrow> AA g p s "
shows "\<And>xv yv P s y x s. (A g x yv ** A g y yv ** P) s \<Longrightarrow> (AA g y ** sep_true) s"
shows "\<And>yv P y x s. (A g x yv ** A g y yv ** P) s \<Longrightarrow> (AA g y ** sep_true) s"
by (sep_solve add: forward)
lemma
@ -108,7 +105,7 @@ lemma "\<lbrakk> (A ** B) sa ; (A ** Y) s \<rbrakk> \<Longrightarrow> (\<lambda>
apply (sep_cancel)
oops
schematic_lemma "\<lbrakk> (B ** A ** C) s \<rbrakk> \<Longrightarrow> (\<lambda>s. (A ** ?X) s) s"
schematic_goal "\<lbrakk> (B ** A ** C) s \<rbrakk> \<Longrightarrow> (\<lambda>s. (A ** ?X) s) s"
by (sep_cancel)
(* test backtracking on premises with same state *)
@ -126,7 +123,7 @@ lemma f1:
declare sep_conj_true[sep_cancel]
lemma boxo: "P s \<Longrightarrow> (P \<and>* \<box>) s"
by (erule sep_conj_sep_emptyI)
by (erule sep_conj_sep_emptyI)
lemma f2:
assumes forward: "\<And>s. generic s \<Longrightarrow> instance s"
@ -134,5 +131,6 @@ lemma f2:
shows "generic s \<Longrightarrow> (instance2 ** \<box>) s"
apply (drule forward forward2)+
apply (sep_erule_concl (direct) boxo)
done
done
end

View File

@ -14,7 +14,7 @@
Rafal Kolanski <rafal.kolanski at nicta.com.au>
*)
header "Example from HOL/Hoare/Separation"
chapter "Example from HOL/Hoare/Separation"
theory Simple_Separation_Example
imports

View File

@ -14,7 +14,7 @@
Rafal Kolanski <rafal.kolanski at nicta.com.au>
*)
header "Separation Algebra for Virtual Memory"
chapter "Separation Algebra for Virtual Memory"
theory VM_Example
imports
@ -70,7 +70,7 @@ fun
= VMSepState (x ++ y, r \<union> r')"
instance
apply default
apply intro_classes
apply (simp add: zero_vm_sep_state_def sep_disj_vm_sep_state_def)
apply (fastforce simp: sep_disj_vm_sep_state_def map_disj_def)
apply (case_tac x, clarsimp simp: zero_vm_sep_state_def)

View File

@ -13,7 +13,7 @@
Rafal Kolanski <rafal.kolanski at nicta.com.au>
*)
header "Instantiating capDL as a separation algebra."
chapter "Instantiating capDL as a separation algebra."
theory Abstract_Separation_D
imports
@ -170,7 +170,7 @@ where
| _ \<Rightarrow> True)"
(* "Cleans" slots to conform with the compontents. *)
(* "Cleans" slots to conform with the components. *)
definition
clean_slots :: "cdl_cap_map \<Rightarrow> cdl_components \<Rightarrow> cdl_cap_map"
where
@ -215,7 +215,7 @@ where
then (update_slots (object_slots clean_obj_a ++ object_slots clean_obj_b) clean_obj_b)
else (update_slots (object_slots clean_obj_a ++ object_slots clean_obj_b) clean_obj_a)"
(* Heaps are added by adding their repsective objects.
(* Heaps are added by adding their respective objects.
* The ghost state tells us which object's fields should be taken.
* Adding objects of the same type adds their caps
* (overwrites the left with the right).
@ -224,15 +224,18 @@ definition
cdl_heap_add :: "sep_state \<Rightarrow> sep_state \<Rightarrow> cdl_heap"
where
"cdl_heap_add state_a state_b \<equiv> \<lambda>obj_id.
let heap_a = sep_heap state_a;
heap_b = sep_heap state_b;
gs_a = sep_ghost_state state_a;
gs_b = sep_ghost_state state_b
in case heap_b obj_id of
let
heap_a = sep_heap state_a;
heap_b = sep_heap state_b;
gs_a = sep_ghost_state state_a;
gs_b = sep_ghost_state state_b
in
case heap_b obj_id of
None \<Rightarrow> heap_a obj_id
| Some obj_b \<Rightarrow> case heap_a obj_id of
None \<Rightarrow> heap_b obj_id
| Some obj_a \<Rightarrow> Some (object_add obj_a obj_b (gs_a obj_id) (gs_b obj_id))"
| Some obj_b \<Rightarrow>
(case heap_a obj_id of
None \<Rightarrow> heap_b obj_id
| Some obj_a \<Rightarrow> Some (object_add obj_a obj_b (gs_a obj_id) (gs_b obj_id)))"
(* Heaps are added by adding their repsective objects.
* The ghost state tells us which object's fields should be taken.
@ -686,12 +689,12 @@ definition "(op +) \<equiv> sep_state_add"
**********************************************)
instance
apply default
apply intro_classes
(* x ## 0 *)
apply (simp add: sep_disj_sep_state_def sep_state_disj_def zero_sep_state_def)
(* x ## y \<Longrightarrow> y ## x *)
apply (clarsimp simp: not_conflicting_objects_comm sep_disj_sep_state_def sep_state_disj_def Let_unfold
map_disj_com not_conflicting_objects_comm Int_commute)
map_disj_com Int_commute)
(* x + 0 = x *)
apply (simp add: plus_sep_state_def sep_state_add_def zero_sep_state_def)
apply (case_tac x)

View File

@ -13,7 +13,7 @@
Rafal Kolanski <rafal.kolanski at nicta.com.au>
*)
header "Defining some separation logic maps-to predicates on top of the instantiation."
chapter "Defining some separation logic maps-to predicates on top of the instantiation."
theory Separation_D
imports Abstract_Separation_D

View File

@ -13,7 +13,7 @@
Rafal Kolanski <rafal.kolanski at nicta.com.au>
*)
header "A simplified version of the actual capDL specification."
chapter "A simplified version of the actual capDL specification."
theory Types_D
imports "~~/src/HOL/Word/Word"

View File

@ -16,28 +16,26 @@
(* Separating Conjunction (and Top, AKA sep_true) {{{
This defines all the constants and theorems necessary for the conjunct
This defines the constants and theorems necessary for the conjunct
selection and cancelling tactic, as well as utility functions.
*)
structure SepConj =
struct
val sep_conj_term = @{term sep_conj};
val sep_conj_str = "**";
val sep_conj_ac = @{thms sep_conj_ac};
val sep_conj_term = @{term sep_conj}
val sep_conj_str = "**"
val sep_conj_ac = @{thms sep_conj_ac}
val sep_conj_impl = @{thm sep_conj_impl}
fun is_sep_conj_const (Const (@{const_name sep_conj}, _)) = true
| is_sep_conj_const _ = false;
| is_sep_conj_const _ = false
fun is_sep_conj_term
(Const t $ _ $ _ $ _) = is_sep_conj_const (Const t)
| is_sep_conj_term _ = false;
fun is_sep_conj_term (Const t $ _ $ _ $ _) = is_sep_conj_const (Const t)
| is_sep_conj_term _ = false
fun is_sep_conj_prop
(Const _ $ t) = is_sep_conj_term t
| is_sep_conj_prop _ = false;
fun is_sep_conj_prop (Const _ $ t) = is_sep_conj_term t
| is_sep_conj_prop _ = false
fun strip_sep_conj (Const (@{const_name sep_conj},_) $ t1 $ t2 $ _) =
[t1] @ (strip_sep_conj t2)
@ -45,18 +43,18 @@ fun strip_sep_conj (Const (@{const_name sep_conj},_) $ t1 $ t2 $ _) =
[t1] @ (strip_sep_conj t2)
(* dig through eta exanded terms: *)
| strip_sep_conj (Abs (_, _, t $ Bound 0)) = strip_sep_conj t
| strip_sep_conj t = [t];
| strip_sep_conj t = [t]
fun is_sep_true_term (Abs (_, _, Const (@{const_name True}, _))) = true
| is_sep_true_term _ = false;
| is_sep_true_term _ = false
fun mk_sep_conj (t1, t2) = sep_conj_term $ t1 $ t2;
fun mk_sep_conj (t1, t2) = sep_conj_term $ t1 $ t2
(* Types of conjuncts and name of state type, for term construction *)
val sep_conj_cjt_typ = type_of sep_conj_term |> domain_type;
val sep_conj_state_typn = domain_type sep_conj_cjt_typ |> dest_TFree |> #1;
val sep_conj_cjt_typ = type_of sep_conj_term |> domain_type
val sep_conj_state_typn = domain_type sep_conj_cjt_typ |> dest_TFree |> #1
end;
end
(* }}} *)
@ -67,9 +65,9 @@ structure FunApp =
struct
(* apply a function term to a Free with given name *)
fun fun_app_free t free_name = t $ Free (free_name, type_of t |> domain_type);
fun fun_app_free t free_name = t $ Free (free_name, type_of t |> domain_type)
end; (* }}} *)
end (* }}} *)
(* Selecting Conjuncts in Premise or Conclusion {{{ *)
@ -81,102 +79,86 @@ end; (* }}} *)
(for dtac) or the premise (for rtac) of the rule.
*)
fun mk_sep_select_rule ctxt conclusion (cjt_count, cjt_selects) =
let
fun variants nctxt names = fold_map Name.variant names nctxt;
val (state, nctxt0) = Name.variant "s" (Variable.names_of ctxt);
fun mk_cjt n = Free (n, type_of SepConj.sep_conj_term |> domain_type);
fun variants nctxt names = fold_map Name.variant names nctxt
val (state, nctxt0) = Name.variant "s" (Variable.names_of ctxt)
fun mk_cjt n = Free (n, type_of SepConj.sep_conj_term |> domain_type)
fun sep_conj_prop cjts =
FunApp.fun_app_free (foldr1 SepConj.mk_sep_conj (map mk_cjt cjts)) state
|> HOLogic.mk_Trueprop;
FunApp.fun_app_free (foldr1 SepConj.mk_sep_conj (map mk_cjt cjts)) state |> HOLogic.mk_Trueprop
(* concatenate string and string of an int *)
fun conc_str_int str int = str ^ Int.toString int;
fun conc_str_int str int = str ^ Int.toString int
(* make the conjunct names *)
val (cjts, _) = 1 upto cjt_count
|> map (conc_str_int "a") |> variants nctxt0;
val (cjts, _) = 1 upto cjt_count |> map (conc_str_int "a") |> variants nctxt0
(* make normal-order separation conjunction terms *)
val orig = sep_conj_prop cjts;
val orig = sep_conj_prop cjts
(* make reordered separation conjunction terms *)
(* We gather the needed conjuncts, and then append it the original list with those conjuncts removed *)
fun dropit n (x::xs) is = if exists (fn y => y = n) is then
(dropit (n+1) xs is) else x :: (dropit (n+1) xs is)
|dropit _ [] _ = []
fun dropit n (x::xs) is = if exists (fn y => y = n) is then dropit (n+1) xs is else x :: dropit (n+1) xs is
| dropit _ [] _ = []
fun nths_to_front idxs xs = (map (nth xs) idxs) @ dropit 0 xs idxs
val reordered = sep_conj_prop (nths_to_front cjt_selects cjts);
val goal = Logic.mk_implies
(if conclusion then (orig, reordered) else (reordered, orig));
fun nths_to_front idxs xs = map (nth xs) idxs @ dropit 0 xs idxs
val reordered = sep_conj_prop (nths_to_front cjt_selects cjts)
val goal = Logic.mk_implies (if conclusion then (orig, reordered) else (reordered, orig))
(* simp add: sep_conj_ac *)
val sep_conj_ac_tac = Simplifier.asm_full_simp_tac
(put_simpset HOL_basic_ss ctxt addsimps SepConj.sep_conj_ac);
(put_simpset HOL_basic_ss ctxt addsimps SepConj.sep_conj_ac)
in
Goal.prove ctxt [] [] goal (fn _ => sep_conj_ac_tac 1)
Goal.prove ctxt [] [] goal (fn _ => sep_conj_ac_tac 1)
|> Drule.generalize ([SepConj.sep_conj_state_typn], state :: cjts)
end;
end
fun conj_length ctxt ct =
let
val ((_, ct'), _) = Variable.focus_cterm NONE ct ctxt;
val concl = ct' |> Drule.strip_imp_concl |> Thm.term_of;
in concl |> HOLogic.dest_Trueprop |> SepConj.strip_sep_conj
|> length
end;
let
val ((_, ct'), _) = Variable.focus_cterm NONE ct ctxt
val concl = ct' |> Drule.strip_imp_concl |> Thm.term_of
in
concl |> HOLogic.dest_Trueprop |> SepConj.strip_sep_conj |> length
end
local
fun all_uniq xs = forall (fn x => length (filter (fn y => x = y) xs) = 1 ) xs
fun all_uniq xs = forall (fn x => length (filter (fn y => x = y) xs) = 1 ) xs
in
fun sep_selects_tac ctxt ns =
let
fun sep_select_tac' ctxt ns (ct, i) =
let
fun th ns = mk_sep_select_rule ctxt false ((conj_length ctxt ct),ns)
in
if not (all_uniq ns) then error ("Duplicate numbers in arguments")
else resolve0_tac [th ns] i handle Subscript => no_tac
end;
let
fun th ns = mk_sep_select_rule ctxt false ((conj_length ctxt ct),ns)
in
if not (all_uniq ns)
then error ("Duplicate numbers in arguments")
else resolve0_tac [th ns] i handle Subscript => no_tac
end
in
CSUBGOAL (sep_select_tac' ctxt (map (fn m => m - 1) ns))
end;
end;
end
end
fun UNSOLVED' tac i st =
tac i st |> Seq.filter (fn st' => Thm.nprems_of st' = Thm.nprems_of st);
tac i st |> Seq.filter (fn st' => Thm.nprems_of st' = Thm.nprems_of st)
fun sep_flatten ctxt =
let fun simptac i = CHANGED_PROP (full_simp_tac
(put_simpset HOL_basic_ss ctxt addsimps [@{thm sep_conj_assoc}]) i)
in UNSOLVED' simptac
end;
let
fun simptac i =
CHANGED_PROP (full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm sep_conj_assoc}]) i)
in
UNSOLVED' simptac
end
fun sep_select_tactic lens_tac ns ctxt =
let
val sep_select = sep_selects_tac ctxt
val iffI = @{thm iffI}
val sep_conj_ac_tac = Simplifier.asm_full_simp_tac
(put_simpset HOL_basic_ss ctxt addsimps SepConj.sep_conj_ac);
in lens_tac THEN'
resolve0_tac [iffI] THEN'
(sep_select ns) THEN'
assume_tac ctxt THEN'
(sep_conj_ac_tac)
end;
fun sep_select_method lens ns ctxt =
SIMPLE_METHOD' (sep_select_tactic lens ns ctxt)
let
val sep_select = sep_selects_tac ctxt
val iffI = @{thm iffI}
val sep_conj_ac_tac =
Simplifier.asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps SepConj.sep_conj_ac)
in
lens_tac THEN'
resolve0_tac [iffI] THEN'
sep_select ns THEN'
assume_tac ctxt THEN'
sep_conj_ac_tac
end

View File

@ -97,7 +97,7 @@ class Poller(threading.Thread):
# For children that are no longer running, remember their
# most recently recorded CPU time.
reaped_cpu = 0.0
for c_id, c_t in self.current_children.iteritems():
for c_id, c_t in self.current_children.items():
if c_id not in new_current_children:
reaped_cpu += c_t
self.old_children_cpu += reaped_cpu

View File

@ -55,8 +55,8 @@
-->
<testsuite>
<!-- Build Isabelle. -->
<test name="isabelle" cwd="../../" cpu-timeout="3600">isabelle/bin/isabelle build -b -v HOL-Word</test>
<!-- Build Isabelle and common theories (Pure, HOL, HOL-Word, Word). -->
<test name="isabelle" cwd="../.." cpu-timeout="3600">isabelle/bin/isabelle build -d . -vb Word</test>
<!-- Ensure that all of our XML files are strictly correct. -->
<test name="tests-xml-correct" cwd="../../" cpu-timeout="60">

View File

@ -148,11 +148,12 @@ sub thy_dep {
# FIXME: tags and crazy names (e.g. theory 007 or theory \<AA>) not supported
sub read_thy {
my $re_word = '\b[a-zA-Z][a-zA-Z0-9_]*\b'; # a word starts with a letter
my $re_id = '(?:'.$re_word.'|"[^"]*")'; # bare word or "quoted string"
my $re_braces = '\{\*(?:[^*]|\*(?!\}))*\*\}'; # {* curly braced string *}
my $re_comment = '\(\*(?:[^*]|\*(?!\)))*\*\)'; # (* parenthesized comment *)
my $re_string = '"[^"]*"'; # "simple quoted string"
my $re_word = '(?:\b[a-zA-Z][a-zA-Z0-9_]*\b)'; # a word starts with a letter
my $re_string = '(?:".*?")'; # "simple quoted string"
my $re_id = '(?:'.$re_word.'|'.$re_string.')'; # bare word or "quoted string"
my $re_braces = '(?:\{\*.*?\*\})'; # {* curly braced string *}
my $re_comment = '(?:\(\*.*?\*\))'; # (* parenthesized comment *)
my $re_cartouche = '(?:\\\\<open>.*?\\\\<close>)'; # cartouches
my $ofn = shift; # original filename as specified by the user (used in error messages)
my $afn = shift; # absolute filename (hopefully unique) internal identifier
@ -176,6 +177,7 @@ sub read_thy {
close(FILE) || die "$!";
$_ = $content;
s/^(?:\s*(?:text|section|subsection|chapter)\s*(?:$re_word|$re_braces|$re_string|$re_cartouche))*//;
if (s/^\s*theory\s*($re_id)\s*//) {
diagnose(1, "$ofn: theory name $1 does not match file name, assuming $tn.\n")
unless $1 eq $tn;

View File

@ -22,11 +22,11 @@
#
# Caution
# ----
# This is a slow, shoddy, and possibly unsafe script that gets its completions by
# This is a slow and unsafe script that gets its completions by
# *** running isabelle and parsing the output ***
# so it may break and do strange things, especially if your Isabelle is not
# configured properly, or its version has changed since this script
# was written (Isabelle 2015).
# was updated (Isabelle 2016).
#
#
# Known problems
@ -42,7 +42,7 @@ _isabelle_completer_path=$(dirname "${(%):-%x}")
_isabelle() {
local -x _isabelle_path
_isabelle_path="${words[1]}"
_isabelle_path=${~words[1]}
if (( CURRENT == 2 )); then
# completing isabelle subcommand
# Isabelle's top-level command description has the format
@ -87,9 +87,7 @@ _isabelle() {
}
_isabelle_usage_to_completion() {
# Ugly hack that takes advantage of the fact that
# isabelle "usage" messages are fairly uniform.
# The typical usage message looks like:
# Isabelle subcommands have standard usage messages e.g.
#
# Usage: blah [OPTIONS] [ARGS]
# Options are:
@ -140,8 +138,8 @@ _isabelle_usage_to_completion() {
allopts+="$optspec"
fi
# non-OPTION arguments
if [[ ${words[CURRENT][1]} != '-' ]]; then
# non-OPTION arguments, if any
if [[ "$usage" =~ '.*Usage: isabelle '$cmd' .*' && ${words[CURRENT][1]} != '-' ]]; then
opt=( $(echo "$usage" | \
sed -e "/Usage: isabelle $cmd/ !d" \
-e 's/Usage: isabelle '$cmd' \([][A-Z_]\+ *\)*.*$/\1/' \
@ -181,18 +179,23 @@ _isabelle_usage_to_completion() {
_isabelle_parse_option() {
local opt="$1" session_dirs="${@:2}"
# Parse option:
# " -a ARG description of a (and more stuff)"
# --> [name] [arg] [ help ]
# [ shorthelp ]
# Parse a single option description:
# " -a ARG=VAL description of a (and more stuff)"
# --> [name] [arg] [ help ]
# [ shorthelp ]
# FIXME? we only parse single-letter options,
# because isabelle only uses single-letter options
local name="$(echo "$opt" | sed -e 's/^ *-\(.\).*$/\1/')"
local arg="$(echo "$opt" | sed -e 's/^ *-. \+\([A-Z_]\+\)\?.*$/\1/')"
local help="$(echo "$opt" | sed -e 's/^ *-. \+\([A-Z_]\+ \+\)\?\(.*[^ ]\) *$/\2/')"
local help="$(echo "$opt" | sed -e 's/^ *-. \+\([A-Z_=]\+ \+\)\?\(.*[^ ]\) *$/\2/')"
local optspec="-${name}[$help]"
# session dirs and runtime options can be given multiple times
local multiple=
if [[ "$help" =~ '.*(session directory|option|OPTION).*' ]]; then
multiple='*'
fi
local optspec="${multiple}-${name}[$help]"
if [[ -n "$arg" ]]; then
# short version for tab completing argument
@ -230,7 +233,7 @@ _isabelle_session_name() {
_tags isa_heaps isa_sessions
local heap_names session_names
heap_names=$("$_isabelle_path" env sh -c 'echo -n "$ISABELLE_PATH"' | \
xargs "-d:" sh -c 'find "$@" -maxdepth 2 -type f' . 2>/dev/null | \
xargs "-d:" sh -c 'find "$@" -maxdepth 2 -type f' dummy-arg 2>/dev/null | \
xargs -d'\n' -n1 basename | sort -u)
session_names=$(python "${_isabelle_completer_path}/isabelle_session_names.py" "$@" | sort -u)
_requested isa_heaps expl 'heap name' compadd -- ${=heap_names}

View File

@ -32,19 +32,16 @@ chapter "Proofs"
session Refine = BaseRefine +
description {* Refinement between Haskell and Abstract spec. *}
options [timeout=3600]
theories
"refine/Refine"
"refine/Orphanage"
session BaseRefine = AInvs +
description {* Background theory and libraries for refinement proof. *}
options [timeout=600]
theories
"refine/Include"
session AInvs = ASpec +
options [timeout=3600]
theories
"invariant-abstract/AInvs"
"invariant-abstract/EmptyFail_AI"
@ -92,7 +89,6 @@ session DPolicy = DRefine +
*)
session Access in "access-control" = AInvs +
options [timeout=3600]
theories
"Syscall_AC"
"ExampleSystem"

View File

@ -263,7 +263,7 @@ where
"tcb_st_to_auth (Structures_A.BlockedOnNotification ntfn) = {(ntfn, Receive)}"
| "tcb_st_to_auth (Structures_A.BlockedOnSend ep payl)
= {(ep, SyncSend)} \<union> (if sender_can_grant payl then {(ep, Grant)} else {})"
| "tcb_st_to_auth (Structures_A.BlockedOnReceive ep dim) = {(ep, Receive)}"
| "tcb_st_to_auth (Structures_A.BlockedOnReceive ep) = {(ep, Receive)}"
| "tcb_st_to_auth _ = {}"
definition
@ -633,7 +633,7 @@ subsection{* How kernel objects can change *}
fun
blocked_on :: "obj_ref \<Rightarrow> Structures_A.thread_state \<Rightarrow> bool"
where
"blocked_on ref (Structures_A.BlockedOnReceive ref' _) = (ref = ref')"
"blocked_on ref (Structures_A.BlockedOnReceive ref') = (ref = ref')"
| "blocked_on ref (Structures_A.BlockedOnSend ref' _) = (ref = ref')"
| "blocked_on ref (Structures_A.BlockedOnNotification ref') = (ref = ref')"
| "blocked_on _ _ = False"
@ -645,7 +645,7 @@ lemma blocked_on_def2:
fun
receive_blocked_on :: "obj_ref \<Rightarrow> Structures_A.thread_state \<Rightarrow> bool"
where
"receive_blocked_on ref (Structures_A.BlockedOnReceive ref' _) = (ref = ref')"
"receive_blocked_on ref (Structures_A.BlockedOnReceive ref') = (ref = ref')"
| "receive_blocked_on ref (Structures_A.BlockedOnNotification ref') = (ref = ref')"
| "receive_blocked_on _ _ = False"
@ -701,7 +701,7 @@ where
abbreviation ep_recv_blocked :: "obj_ref \<Rightarrow> thread_state \<Rightarrow> bool"
where
"ep_recv_blocked ep ts \<equiv> case ts of BlockedOnReceive w b \<Rightarrow> w = ep
"ep_recv_blocked ep ts \<equiv> case ts of BlockedOnReceive w \<Rightarrow> w = ep
| _ \<Rightarrow> False"
definition indirect_send :: "'a set \<Rightarrow> 'a PAS \<Rightarrow> obj_ref \<Rightarrow> obj_ref \<Rightarrow> tcb \<Rightarrow> bool"
@ -1407,7 +1407,7 @@ lemma get_endpoint_wp:
lemma ep_queued_st_tcb_at':
"\<And>P. \<lbrakk>ko_at (Endpoint ep) ptr s; (t, rt) \<in> ep_q_refs_of ep;
valid_objs s; sym_refs (state_refs_of s);
\<And>pl dim. P (Structures_A.BlockedOnSend ptr pl) \<and> P (Structures_A.BlockedOnReceive ptr dim) \<rbrakk>
\<And>pl. P (Structures_A.BlockedOnSend ptr pl) \<and> P (Structures_A.BlockedOnReceive ptr) \<rbrakk>
\<Longrightarrow> st_tcb_at P t s"
apply (case_tac ep, simp_all)
apply (frule(1) sym_refs_ko_atD, clarsimp, erule (1) my_BallE,
@ -1419,7 +1419,7 @@ lemma ep_queued_st_tcb_at':
lemma ep_rcv_queued_st_tcb_at:
"\<And>P. \<lbrakk>ko_at (Endpoint ep) epptr s; (t, EPRecv) \<in> ep_q_refs_of ep;
valid_objs s; sym_refs (state_refs_of s);
\<And>dim. P (Structures_A.BlockedOnReceive epptr dim);
P (Structures_A.BlockedOnReceive epptr);
kheap s' t = kheap s t\<rbrakk>
\<Longrightarrow> st_tcb_at P t s'"
apply (case_tac ep, simp_all)

View File

@ -999,7 +999,7 @@ lemma transfer_caps_domain_sep_inv:
cp = fst x)
(snd x) s \<and>
real_cte_at (snd x) s))\<rbrace>
transfer_caps mi caps endpoint receiver receive_buffer diminish
transfer_caps mi caps endpoint receiver receive_buffer
\<lbrace>\<lambda>_. domain_sep_inv irqs st\<rbrace>"
apply (simp add: transfer_caps_def)
apply (wpc | wp)+
@ -1013,7 +1013,7 @@ lemma transfer_caps_domain_sep_inv:
lemma do_normal_transfer_domain_sep_inv:
"\<lbrace>domain_sep_inv irqs st and valid_objs and valid_mdb\<rbrace>
do_normal_transfer sender send_buffer ep badge grant receiver recv_buffer diminish
do_normal_transfer sender send_buffer ep badge grant receiver recv_buffer
\<lbrace>\<lambda>_. domain_sep_inv irqs st\<rbrace>"
unfolding do_normal_transfer_def
apply (wp transfer_caps_domain_sep_inv hoare_vcg_ball_lift lec_valid_cap' | simp)+
@ -1023,7 +1023,7 @@ crunch domain_sep_inv[wp]: do_fault_transfer "domain_sep_inv irqs st"
lemma do_ipc_transfer_domain_sep_inv:
"\<lbrace>domain_sep_inv irqs st and valid_objs and valid_mdb\<rbrace>
do_ipc_transfer sender ep badge grant receiver diminish
do_ipc_transfer sender ep badge grant receiver
\<lbrace>\<lambda>_. domain_sep_inv irqs st\<rbrace>"
unfolding do_ipc_transfer_def
apply (wp do_normal_transfer_domain_sep_inv hoare_vcg_all_lift | wpc | wp_once hoare_drop_imps)+

View File

@ -330,7 +330,7 @@ where
tcb_reply = Structures_A.ReplyCap 3080 True, (* master reply cap to itself *)
tcb_caller = Structures_A.NullCap,
tcb_ipcframe = Structures_A.NullCap,
tcb_state = Structures_A.BlockedOnReceive 9 True,
tcb_state = Structures_A.BlockedOnReceive 9,
tcb_fault_handler = undefined,
tcb_ipc_buffer = undefined,
tcb_context = undefined,
@ -534,7 +534,7 @@ lemma Sys1_wellformed: "pas_wellformed Sys1PAS"
done
lemma tcb_states_of_state_1:
"tcb_states_of_state s1 = [0xC08 \<mapsto> Structures_A.thread_state.BlockedOnReceive 9 True, 0xC07 \<mapsto> Structures_A.thread_state.Running ]"
"tcb_states_of_state s1 = [0xC08 \<mapsto> Structures_A.thread_state.BlockedOnReceive 9, 0xC07 \<mapsto> Structures_A.thread_state.Running ]"
unfolding s1_def tcb_states_of_state_def
apply (rule ext)
apply (simp add: get_tcb_def)
@ -869,7 +869,7 @@ where
tcb_reply = Structures_A.ReplyCap 3080 True, (* master reply cap to itself *)
tcb_caller = Structures_A.NullCap,
tcb_ipcframe = Structures_A.NullCap,
tcb_state = Structures_A.BlockedOnReceive 9 True,
tcb_state = Structures_A.BlockedOnReceive 9,
tcb_fault_handler = undefined,
tcb_ipc_buffer = undefined,
tcb_context = undefined,

View File

@ -412,7 +412,7 @@ lemma drop_Suc0_iff:
by (auto simp: neq_Nil_conv)
lemma receive_blocked_on_def3:
"receive_blocked_on ref ts = ((\<exists>d. ts = Structures_A.BlockedOnReceive ref d) \<or> ts = (Structures_A.BlockedOnNotification ref))"
"receive_blocked_on ref ts = ((ts = Structures_A.BlockedOnReceive ref) \<or> ts = (Structures_A.BlockedOnNotification ref))"
by (cases ts, auto)
@ -473,7 +473,7 @@ lemma cancel_ipc_receive_blocked_respects:
apply (clarsimp simp: cancel_ipc_def bind_assoc)
apply (rule hoare_seq_ext[OF _ gts_sp])
apply (rule hoare_name_pre_state)
apply (subgoal_tac "case state of BlockedOnReceive x _\<Rightarrow> True | _ \<Rightarrow> False")
apply (subgoal_tac "case state of BlockedOnReceive x \<Rightarrow> True | _ \<Rightarrow> False")
apply (simp add: blocked_cancel_ipc_def bind_assoc set_endpoint_def set_object_def
get_ep_queue_def get_blocking_object_def
split: thread_state.splits)
@ -495,7 +495,7 @@ lemma cancel_ipc_receive_blocked_respects:
apply (rule_tac x="pasObjectAbs aag t" in exI, rule_tac x=ntfnptr in exI, simp)
apply (erule get_tcb_recv_blocked_implies_receive, erule get_tcb_rev)
apply simp
apply (rename_tac word be careful tcb your dogs finaly one)
apply (rename_tac word careful tcb your dogs finaly one)
apply (rule_tac ntfn'= "tcb_bound_notification tcb" and ep="ntfnptr" and recv=word in tro_tcb_send, simp+)
apply (rule_tac x="tcb_context tcb" in exI, simp)
apply (simp add: tcb_bound_notification_reset_integrity_def )
@ -535,7 +535,7 @@ lemma integrity_receive_blocked_chain':
"\<lbrakk> st_tcb_at receive_blocked p s; integrity aag X st s; \<not> is_subject aag p \<rbrakk> \<Longrightarrow> st_tcb_at receive_blocked p st"
apply (clarsimp simp: integrity_def st_tcb_at_tcb_states_of_state receive_blocked_def)
apply (simp split: thread_state.split_asm)
apply (rename_tac word kp)
apply (rename_tac word)
apply (drule_tac ep=word in tsos_tro [where p = p], simp+ )
done
@ -668,7 +668,7 @@ lemma transfer_caps_loop_pres_dest_aux:
(\<forall>x \<in> set caps. s \<turnstile> fst x \<and>
cte_wp_at (\<lambda>cp. fst x \<noteq> cap.NullCap \<longrightarrow> cp \<noteq> fst x \<longrightarrow> cp = masked_as_full (fst x) (fst x)) (snd x) s
\<and> real_cte_at (snd x) s))\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv. P\<rbrace>" (is "?L \<Longrightarrow> ?P n caps slots mi")
proof (induct caps arbitrary: slots n mi)
case Nil
@ -743,7 +743,7 @@ lemma transfer_caps_loop_pres_dest:
(\<forall>x \<in> set slots. real_cte_at x s) \<and>
(\<forall>x \<in> set caps. s \<turnstile> fst x \<and> cte_wp_at (\<lambda>cp. fst x \<noteq> cap.NullCap \<longrightarrow> cp \<noteq> fst x \<longrightarrow> cp = masked_as_full (fst x) (fst x)) (snd x) s
\<and> real_cte_at (snd x) s))\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv. P\<rbrace>"
apply (rule hoare_pre)
apply (rule transfer_caps_loop_pres_dest_aux [OF x eb])
@ -830,7 +830,7 @@ lemma transfer_caps_loop_pas_refined:
"\<lbrace>pas_refined aag
and (\<lambda>s. (\<forall>x \<in> set caps. valid_cap (fst x) s))
and K ((\<forall>slot \<in> set slots. is_subject aag (fst slot)) \<and> (\<forall>x \<in> set caps. is_subject aag (fst (snd x)) \<and> pas_cap_cur_auth aag (fst x)))\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv. pas_refined aag\<rbrace>"
proof (rule hoare_gen_asm, induct caps arbitrary: slots n mi)
case Nil thus ?case
@ -849,7 +849,6 @@ next
apply (wp cap_insert_pas_refined hoare_vcg_ball_lift hoare_whenE_wp hoare_drop_imps
derive_cap_aag_caps
| simp split del: split_if add: if_apply_def2)+
apply (clarsimp simp: pas_refined_refl intro!: remove_rights_cur_auth)
done
qed
@ -857,7 +856,7 @@ lemma transfer_caps_pas_refined:
"\<lbrace>pas_refined aag
and (\<lambda>s. (\<forall>x \<in> set caps. valid_cap (fst x) s))
and K (is_subject aag receiver \<and> (\<forall>x \<in> set caps. is_subject aag (fst (snd x))) \<and> (\<forall>x \<in> set caps. pas_cap_cur_auth aag (fst x))) \<rbrace>
transfer_caps info caps endpoint receiver recv_buf diminish
transfer_caps info caps endpoint receiver recv_buf
\<lbrace>\<lambda>rv. pas_refined aag\<rbrace>"
unfolding transfer_caps_def
apply (rule hoare_pre)
@ -908,7 +907,7 @@ lemma lookup_extra_caps_auth:
done
lemma transfer_caps_empty_inv:
"\<lbrace>P\<rbrace> transfer_caps mi [] endpoint receiver rbuf diminish \<lbrace>\<lambda>_. P\<rbrace>"
"\<lbrace>P\<rbrace> transfer_caps mi [] endpoint receiver rbuf \<lbrace>\<lambda>_. P\<rbrace>"
unfolding transfer_caps_def
by (wp | wpc | simp) +
@ -935,7 +934,7 @@ lemma do_normal_transfer_pas_refined:
and valid_objs
and K (grant \<longrightarrow> is_subject aag sender)
and K (grant \<longrightarrow> is_subject aag receiver)\<rbrace>
do_normal_transfer sender sbuf endpoint badge grant receiver rbuf diminish
do_normal_transfer sender sbuf endpoint badge grant receiver rbuf
\<lbrace>\<lambda>rv. pas_refined aag\<rbrace>"
proof(cases grant)
case True thus ?thesis
@ -967,7 +966,7 @@ lemma do_ipc_transfer_pas_refined:
and valid_objs
and K (grant \<longrightarrow> is_subject aag sender)
and K (grant \<longrightarrow> is_subject aag receiver)\<rbrace>
do_ipc_transfer sender ep badge grant receiver diminish
do_ipc_transfer sender ep badge grant receiver
\<lbrace>\<lambda>rv. pas_refined aag\<rbrace>"
apply (simp add: do_ipc_transfer_def)
apply (wp do_normal_transfer_pas_refined
@ -993,10 +992,16 @@ lemma send_ipc_pas_refined:
apply (simp add: send_ipc_def)
apply (rule hoare_seq_ext[OF _ get_endpoint_sp])
apply (rule hoare_pre)
apply (wp set_thread_state_pas_refined
| wpc
| simp add: hoare_if_r_and)+
apply (rename_tac list x xs recv_state diminish)
apply wpc
apply wpc
apply (wp set_thread_state_pas_refined)
apply wpc
apply (wp set_thread_state_pas_refined)
apply wpc
apply (wp set_thread_state_pas_refined)
apply (simp add: hoare_if_r_and split del:split_if)
apply (rename_tac list x xs recv_state)
apply (rule_tac Q="\<lambda>rv. pas_refined aag and K (can_grant \<longrightarrow> is_subject aag (hd list))"
in hoare_strengthen_post[rotated])
apply (clarsimp simp: cli_no_irqs pas_refined_refl aag_cap_auth_def clas_no_asid)
@ -1053,7 +1058,7 @@ abbreviation receive_ipc_base
where
"receive_ipc_base aag thread ep epptr rights is_blocking \<equiv> case ep of
IdleEP \<Rightarrow> case is_blocking of
True \<Rightarrow> do set_thread_state thread (BlockedOnReceive epptr (AllowSend \<notin> rights));
True \<Rightarrow> do set_thread_state thread (BlockedOnReceive epptr);
set_endpoint epptr (RecvEP [thread])
od
| False \<Rightarrow> do_nbrecv_failed_transfer thread
@ -1065,10 +1070,10 @@ where
sender_state \<leftarrow> get_thread_state sender;
data \<leftarrow> case sender_state of BlockedOnSend ref x \<Rightarrow> return x | _ \<Rightarrow> fail;
do_ipc_transfer sender (Some epptr) (sender_badge data)
(sender_can_grant data) thread (AllowSend \<notin> rights);
(sender_can_grant data) thread ;
fault \<leftarrow> thread_get tcb_fault sender;
if sender_is_call data \<or> fault \<noteq> None
then if sender_can_grant data \<and> \<not> AllowSend \<notin> rights
then if sender_can_grant data
then setup_caller_cap sender thread
else set_thread_state sender Inactive
else do set_thread_state sender Running;
@ -1076,7 +1081,7 @@ where
od
od
| RecvEP queue \<Rightarrow> case is_blocking of
True \<Rightarrow> do set_thread_state thread (BlockedOnReceive epptr (AllowSend \<notin> rights));
True \<Rightarrow> do set_thread_state thread (BlockedOnReceive epptr);
set_endpoint epptr (RecvEP (queue @ [thread]))
od
| False \<Rightarrow> do_nbrecv_failed_transfer thread"
@ -1199,7 +1204,7 @@ lemma transfer_caps_integrity_autarch:
real_cte_at (snd x) s))
and K (is_subject aag receiver \<and> ipc_buffer_has_auth aag receiver receive_buffer \<and>
(\<forall>x\<in>set caps. is_subject aag (fst (snd x))) \<and> length caps < 6)\<rbrace>
transfer_caps mi caps endpoint receiver receive_buffer diminish
transfer_caps mi caps endpoint receiver receive_buffer
\<lbrace>\<lambda>rv. integrity aag X st\<rbrace>"
apply (rule hoare_gen_asm)
apply (simp add: transfer_caps_def)
@ -1263,7 +1268,7 @@ lemma do_normal_transfer_send_integrity_autarch:
and K (is_subject aag receiver \<and>
ipc_buffer_has_auth aag receiver rbuf \<and>
(grant \<longrightarrow> is_subject aag sender))\<rbrace>
do_normal_transfer sender sbuf endpoint badge grant receiver rbuf diminish
do_normal_transfer sender sbuf endpoint badge grant receiver rbuf
\<lbrace>\<lambda>rv. integrity aag X st\<rbrace>"
apply (simp add: do_normal_transfer_def)
apply (wp as_user_integrity_autarch set_message_info_integrity_autarch transfer_caps_integrity_autarch
@ -1292,7 +1297,7 @@ lemma do_ipc_transfer_integrity_autarch:
and integrity aag X st
and valid_objs and valid_mdb
and K (is_subject aag receiver \<and> (grant \<longrightarrow> is_subject aag sender))\<rbrace>
do_ipc_transfer sender ep badge grant receiver diminish
do_ipc_transfer sender ep badge grant receiver
\<lbrace>\<lambda>rv. integrity aag X st\<rbrace>"
apply (simp add: do_ipc_transfer_def)
apply (wp do_normal_transfer_send_integrity_autarch do_fault_transfer_integrity_autarch
@ -1731,7 +1736,7 @@ lemma transfer_caps_loop_respects_in_ipc_autarch:
\<and> (\<not> is_subject aag receiver \<longrightarrow> auth_ipc_buffers st receiver = ptr_range buffer msg_align_bits)
\<and> is_aligned buffer msg_align_bits
\<and> n + length caps < 6 \<and> distinct slots)\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv. integrity_tcb_in_ipc aag X receiver epptr TRContext st\<rbrace>"
apply (rule hoare_gen_asm)
apply (wp transfer_caps_loop_pres_dest cap_inserintegrity_in_ipc_autarch
@ -1753,7 +1758,7 @@ lemma transfer_caps_respects_in_ipc:
\<and> (\<not> is_subject aag receiver \<longrightarrow> case_option True (\<lambda>buf'. auth_ipc_buffers st receiver = ptr_range buf' msg_align_bits) recv_buf)
\<and> (case_option True (\<lambda>buf'. is_aligned buf' msg_align_bits) recv_buf)
\<and> length caps < 6)\<rbrace>
transfer_caps mi caps endpoint receiver recv_buf diminish
transfer_caps mi caps endpoint receiver recv_buf
\<lbrace>\<lambda>rv. integrity_tcb_in_ipc aag X receiver epptr TRContext st\<rbrace>"
apply (rule hoare_gen_asm)
apply (cases recv_buf)
@ -1796,7 +1801,7 @@ lemma do_normal_transfer_respects_in_ipc:
\<and> is_subject aag receiver)
and K ((\<not> is_subject aag receiver \<longrightarrow> (case recv_buf of None \<Rightarrow> True | Some buf' \<Rightarrow> auth_ipc_buffers st receiver = ptr_range buf' msg_align_bits)) \<and>
(case recv_buf of None \<Rightarrow> True | Some buf' \<Rightarrow> is_aligned buf' msg_align_bits))\<rbrace>
do_normal_transfer sender sbuf (Some epptr) badge grant receiver recv_buf diminish
do_normal_transfer sender sbuf (Some epptr) badge grant receiver recv_buf
\<lbrace>\<lambda>rv. integrity_tcb_in_ipc aag X receiver epptr TRContext st\<rbrace>"
apply (simp add: do_normal_transfer_def)
apply (wp as_user_respects_in_ipc set_message_info_respects_in_ipc
@ -1886,7 +1891,7 @@ lemma do_ipc_transfer_respects_in_ipc:
and (\<lambda>s. grant \<longrightarrow> is_subject aag sender
\<and> is_subject aag receiver)
and K (aag_has_auth_to aag SyncSend epptr)\<rbrace>
do_ipc_transfer sender (Some epptr) badge grant receiver diminish
do_ipc_transfer sender (Some epptr) badge grant receiver
\<lbrace>\<lambda>rv. integrity_tcb_in_ipc aag X receiver epptr TRContext st\<rbrace>"
apply (simp add: do_ipc_transfer_def)
apply (wp do_normal_transfer_respects_in_ipc do_fault_transfer_respects_in_ipc
@ -1963,7 +1968,7 @@ lemma integrity_tcb_in_ipc_refl:
lemma ep_queued_st_tcb_at'':
"\<And>P. \<lbrakk>ko_at (Endpoint ep) ptr s; (t, rt) \<in> ep_q_refs_of ep;
valid_objs s; sym_refs (state_refs_of s);
\<And>pl dim. (rt = EPSend \<and> P (Structures_A.BlockedOnSend ptr pl)) \<or> (rt = EPRecv \<and> P (Structures_A.BlockedOnReceive ptr dim)) \<rbrakk>
\<And>pl. (rt = EPSend \<and> P (Structures_A.BlockedOnSend ptr pl)) \<or> (rt = EPRecv \<and> P (Structures_A.BlockedOnReceive ptr)) \<rbrakk>
\<Longrightarrow> st_tcb_at P t s"
apply (case_tac ep, simp_all)
apply (frule (1) sym_refs_ko_atD, fastforce simp: st_tcb_at_def obj_at_def refs_of_rev)+

View File

@ -491,7 +491,7 @@ lemma handle_interrupt_pas_refined:
handle_interrupt irq
\<lbrace>\<lambda>rv. pas_refined aag\<rbrace>"
apply (simp add: handle_interrupt_def)
apply (rule hoare_pre)
apply (rule conjI; rule impI;rule hoare_pre)
apply (wp send_signal_pas_refined get_cap_wp
| wpc
| simp add: get_irq_slot_def get_irq_state_def)+
@ -522,7 +522,7 @@ lemma handle_interrupt_integrity_autarch:
handle_interrupt irq
\<lbrace>\<lambda>rv. integrity aag X st\<rbrace>"
apply (simp add: handle_interrupt_def cong: irq_state.case_cong maskInterrupt_def ackInterrupt_def resetTimer_def )
apply (rule hoare_pre)
apply (rule conjI; rule impI; rule hoare_pre)
apply (wp_once send_signal_respects get_cap_auth_wp [where aag = aag] dmo_mol_respects
| simp add: get_irq_slot_def get_irq_state_def ackInterrupt_def resetTimer_def
| wp dmo_no_mem_respects
@ -546,7 +546,7 @@ lemma handle_interrupt_integrity:
handle_interrupt irq
\<lbrace>\<lambda>rv. integrity aag X st\<rbrace>"
apply (simp add: handle_interrupt_def maskInterrupt_def ackInterrupt_def resetTimer_def cong: irq_state.case_cong bind_cong)
apply (rule hoare_pre)
apply (rule conjI; rule impI; rule hoare_pre)
apply (wp_once send_signal_respects get_cap_wp dmo_mol_respects dmo_no_mem_respects
| wpc
| simp add: get_irq_slot_def get_irq_state_def ackInterrupt_def resetTimer_def)+
@ -958,6 +958,7 @@ lemma schedule_pas_refined:
lemma handle_interrupt_arch_state [wp]:
"\<lbrace>\<lambda>s :: det_ext state. P (arch_state s)\<rbrace> handle_interrupt irq \<lbrace>\<lambda>_ s. P (arch_state s)\<rbrace>"
unfolding handle_interrupt_def
apply (rule hoare_if)
apply (rule hoare_pre)
apply clarsimp
apply (wp get_cap_inv dxo_wp_weak send_signal_arch_state | wpc | simp add: get_irq_state_def)+

View File

@ -47,8 +47,6 @@ definition
destFault \<leftarrow>
unlessE (destPrio \<ge> curPrio) $ throwError ();
unlessE (capEPCanGrant epCap) $ throwError ();
destST \<leftarrow> liftE $ getThreadState dest;
unlessE (\<not> blockingIPCDiminishCaps destST) $ throwError ();
asidMap \<leftarrow> liftE $ gets $ armKSASIDMap o ksArchState;
unlessE (\<exists>v. {hwasid. (hwasid, pd) \<in> ran asidMap} = {v})
$ throwError ();
@ -132,8 +130,7 @@ definition
unlessE (callerDom = curDom) $ throwError ();
liftE $ do
threadSet (tcbState_update (\<lambda>_. BlockedOnReceive (capEPPtr epCap)
(\<not> capEPCanSend epCap))) curThread;
threadSet (tcbState_update (\<lambda>_. BlockedOnReceive (capEPPtr epCap))) curThread;
setEndpoint (capEPPtr epCap)
(case ep of IdleEP \<Rightarrow> RecvEP [curThread] | RecvEP ts \<Rightarrow> RecvEP (ts @ [curThread]));
mdbPrev \<leftarrow> liftM (mdbPrev o cteMDBNode) $ getCTE callerSlot;
@ -1240,35 +1237,6 @@ lemma thread_state_ptr_mset_blockingObject_tsType_spec:
simp_all add: mask_eq_x_eq_0)
done
lemma thread_state_ptr_set_blockingIPCDiminish_np_spec:
defines "ptr s \<equiv> cparent \<^bsup>s\<^esup>ts_ptr [''tcbState_C''] :: tcb_C ptr"
shows
"\<forall>s. \<Gamma>\<turnstile> \<lbrace>s. hrs_htd \<^bsup>s\<^esup>t_hrs \<Turnstile>\<^sub>t ptr s \<and> dim_' s && 1 = dim_' s\<rbrace>
Call thread_state_ptr_set_blockingIPCDiminish_np_'proc
{t. \<exists>tcb ts. cslift s (ptr s) = Some tcb
\<and> blockingIPCDiminishCaps_CL (thread_state_lift ts) = dim_' s
\<and> tcbQueued_CL (thread_state_lift ts)
= tcbQueued_CL (thread_state_lift (tcbState_C tcb))
\<and> tsType_CL (thread_state_lift ts)
= tsType_CL (thread_state_lift (tcbState_C tcb))
\<and> blockingObject_CL (thread_state_lift ts)
= blockingObject_CL (thread_state_lift (tcbState_C tcb))
\<and> cslift t = cslift s(ptr s \<mapsto> tcb\<lparr>tcbState_C := ts\<rparr>)
\<and> types_proofs.cslift_all_but_tcb_C t s
\<and> hrs_htd (t_hrs_' (globals t)) = hrs_htd (t_hrs_' (globals s))}"
apply (intro allI, rule conseqPre, vcg)
apply (clarsimp simp: ptr_def)
apply (frule h_t_valid_c_guard_cparent, simp+)
apply (simp add: typ_uinfo_t_def)
apply (clarsimp simp: h_t_valid_clift_Some_iff)
apply (frule clift_subtype, simp+)
apply (clarsimp simp: typ_heap_simps')
apply (subst parent_update_child, erule typ_heap_simps', simp+)
apply (clarsimp simp: typ_heap_simps' word_sle_def)
apply (rule exI, intro conjI[rotated], rule refl)
apply (simp_all add: thread_state_lift_def)
done
lemma mdb_node_ptr_mset_mdbNext_mdbRevocable_mdbFirstBadged_spec:
defines "ptr s \<equiv> cparent \<^bsup>s\<^esup>node_ptr [''cteMDBNode_C''] :: cte_C ptr"
shows
@ -2335,7 +2303,7 @@ lemma fastpath_call_ccorres:
done
show ?thesis
using [[goals_limit = 1]]
using [[goals_limit = 3]]
apply (cinit lift: cptr_' msgInfo_')
apply (simp add: catch_liftE_bindE unlessE_throw_catch_If
unifyFailure_catch_If catch_liftE
@ -2518,46 +2486,11 @@ apply (simp add: getThreadCSpaceRoot_def locateSlot_conv
del: Collect_const cong: call_ignore_cong)
apply (rule ccorres_Cond_rhs_Seq)
apply simp
apply (rule ccorres_cond_true_seq)
apply (rule ccorres_split_throws)
apply (fold dc_def)[1]
apply (rule ccorres_call_hSkip)
apply (rule slowpath_ccorres, simp+)
apply (vcg exspec=slowpath_noreturn_spec)
apply (simp del: Collect_const cong: call_ignore_cong)
apply (rule ccorres_rhs_assoc)+
apply (rule_tac xf'="ret__unsigned_'"
and r'="\<lambda>rv rv'. rv' = from_bool (blockingIPCDiminishCaps rv)"
in ccorres_split_nothrow)
apply (rule_tac P="ko_at' send_ep (capEPPtr (theRight luRet))
and (sym_refs o state_refs_of')"
in ccorres_from_vcg[where P'=UNIV])
apply (rule allI, rule conseqPre, vcg)
apply (clarsimp simp: isRecvEP_endpoint_case
neq_Nil_conv)
apply (drule(1) sym_refs_ko_atD')
apply (clarsimp simp: typ_heap_simps' ep_q_refs_of'_def
isRecvEP_endpoint_case
st_tcb_at_refs_of_rev'
st_tcb_at'_def)
apply (drule(1) obj_at_cslift_tcb)
apply (clarsimp simp: typ_heap_simps' getThreadState_def)
apply (rule rev_bexI, erule threadGet_eq)
apply (clarsimp simp: ctcb_relation_def isBlockedOnReceive_def
cthread_state_relation_def)
apply (simp add: thread_state_lift_def word_size)
apply ceqv
apply (rename_tac send_state send_state_c)
apply csymbr
apply (simp add: if_1_0_0 ccap_relation_ep_helpers from_bool_0
del: Collect_const cong: call_ignore_cong)
apply (rule ccorres_Cond_rhs_Seq)
apply simp
apply (rule ccorres_split_throws)
apply (fold dc_def)[1]
apply (rule ccorres_call_hSkip)
apply (rule slowpath_ccorres, simp+)
apply (vcg exspec=slowpath_noreturn_spec)
apply (simp add: ccap_relation_pd_helper cap_get_tag_isCap_ArchObject2
del: Collect_const WordSetup.ptr_add_def cong: call_ignore_cong)
apply csymbr
@ -2576,7 +2509,6 @@ apply (simp add: getThreadCSpaceRoot_def locateSlot_conv
apply (simp add: pde_stored_asid_def asid_map_pd_to_hwasids_def
to_bool_def
del: Collect_const cong: call_ignore_cong)
(* apply (rule ccorres_abstract_ksCurThread, ceqv) *)
apply (rule ccorres_move_c_guard_tcb ccorres_move_const_guard)+
apply (rule ccorres_symb_exec_l3[OF _ curDomain_inv _])
prefer 3
@ -2819,10 +2751,6 @@ apply (simp add: getThreadCSpaceRoot_def locateSlot_conv
apply simp
apply wp[1]
apply simp
apply (rule gts_wp')
apply (simp del: Collect_const)
apply (vcg exspec=thread_state_ptr_get_blockingIPCDiminishCaps_modifies)
apply simp
apply (rule threadGet_wp)
apply simp
apply (rule threadGet_wp)
@ -3362,7 +3290,6 @@ lemma fastpath_reply_recv_ccorres:
del: Collect_const cong: call_ignore_cong)
apply simp
apply (rule ccorres_rhs_assoc2, rule ccorres_rhs_assoc2)
apply (rule_tac xf'=xfdc and r'=dc in ccorres_split_nothrow)
apply (rule_tac P="capAligned (theRight luRet)" in ccorres_gen_asm)
apply (rule_tac P=\<top> and P'="\<lambda>s. ksCurThread s = curThread"
@ -3372,8 +3299,6 @@ lemma fastpath_reply_recv_ccorres:
h_t_valid_clift_Some_iff)
apply (clarsimp simp: capAligned_def isCap_simps objBits_simps
"StrictC'_thread_state_defs" mask_def)
apply (simp add: ccap_relation_ep_helpers)
apply (clarsimp simp: if_distrib[where f=scast] if_Const_helper[where Con="\<lambda>x. x && 1", symmetric])
apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def
typ_heap_simps')
apply (rule conjI)
@ -3529,8 +3454,7 @@ lemma fastpath_reply_recv_ccorres:
threadSet_state_refs_of' threadSet_ctes_of
valid_ep_typ_at_lift' threadSet_cte_wp_at'
| simp)+
apply (vcg exspec=thread_state_ptr_set_blockingIPCDiminish_np_modifies
exspec=thread_state_ptr_mset_blockingObject_tsType_modifies)
apply (vcg exspec=thread_state_ptr_mset_blockingObject_tsType_modifies)
apply simp
apply (rule threadGet_wp)
@ -4294,7 +4218,7 @@ lemma setVMRoot_isolatable:
done
lemma transferCaps_simple:
"transferCaps mi [] ep receiver rcvrBuf diminish =
"transferCaps mi [] ep receiver rcvrBuf =
do
getReceiveSlots receiver rcvrBuf;
return (mi\<lparr>msgExtraCaps := 0, msgCapsUnwrapped := 0\<rparr>)
@ -4308,7 +4232,7 @@ lemma transferCaps_simple:
lemma transferCaps_simple_rewrite:
"monadic_rewrite True True ((\<lambda>_. caps = []) and \<top>)
(transferCaps mi caps ep r rBuf diminish)
(transferCaps mi caps ep r rBuf)
(return (mi \<lparr> msgExtraCaps := 0, msgCapsUnwrapped := 0 \<rparr>))"
apply (rule monadic_rewrite_gen_asm)
apply (rule monadic_rewrite_imp)
@ -4333,7 +4257,7 @@ lemma doIPCTransfer_simple_rewrite:
\<le> of_nat (length State_H.msgRegisters))
and obj_at' (\<lambda>tcb. tcbFault tcb = None
\<and> tcbContext tcb msgInfoRegister = msgInfo) sender)
(doIPCTransfer sender ep badge True rcvr diminish)
(doIPCTransfer sender ep badge True rcvr)
(do rv \<leftarrow> mapM_x (\<lambda>r. do v \<leftarrow> asUser sender (getRegister r);
asUser rcvr (setRegister r v)
od)
@ -5299,10 +5223,6 @@ lemma fastpath_callKernel_SysCall_corres:
apply (rule monadic_rewrite_if_rhs[rotated])
apply (rule monadic_rewrite_alternative_l)
apply (simp add: isRight_case_sum)
apply (rule monadic_rewrite_symb_exec_r [OF gts_inv' no_fail_getThreadState])
apply (rename_tac "destState")
apply (rule monadic_rewrite_if_rhs[rotated])
apply (rule monadic_rewrite_alternative_l)
apply (rule monadic_rewrite_symb_exec_r [OF gets_inv non_fail_gets])
apply (rule monadic_rewrite_if_rhs[rotated])
apply (rule monadic_rewrite_alternative_l)
@ -5328,8 +5248,10 @@ lemma fastpath_callKernel_SysCall_corres:
apply (rule_tac P="epQueue send_ep \<noteq> []" in monadic_rewrite_gen_asm)
apply (simp add: isRecvEP_endpoint_case list_case_helper bind_assoc)
apply (rule monadic_rewrite_bind_tail)
apply (rule_tac x=destState in monadic_rewrite_symb_exec,
wp empty_fail_getThreadState)
apply (elim conjE)
apply (match premises in "isEndpointCap ep" for ep \<Rightarrow>
\<open>rule monadic_rewrite_symb_exec[where x="BlockedOnReceive (capEPPtr ep)"]\<close>,
wp empty_fail_getThreadState)
apply (rule monadic_rewrite_symb_exec2, (wp | simp)+)
apply (rule monadic_rewrite_bind)
apply (rule_tac msgInfo=msgInfo in doIPCTransfer_simple_rewrite)
@ -5367,7 +5289,7 @@ lemma fastpath_callKernel_SysCall_corres:
apply (simp add: setSchedulerAction_def)
apply wp[1]
apply (simp cong: if_cong conj_cong add: if_bool_simps)
apply (simp_all only:)[4]
apply (simp_all only:)[5]
apply ((wp setThreadState_oa_queued[of _ "\<lambda>a _ _. \<not> a"]
setThreadState_obj_at_unchanged
asUser_obj_at_unchanged mapM_x_wp'
@ -5392,6 +5314,7 @@ lemma fastpath_callKernel_SysCall_corres:
apply (rule_tac P="\<lambda>s. ksSchedulerAction s = ResumeCurrentThread
\<and> tcb_at' thread s"
and F=True and E=False in monadic_rewrite_weaken)
apply simp
apply (rule monadic_rewrite_isolate_final)
apply (simp add: isRight_case_sum cong: list.case_cong)
apply (clarsimp simp: fun_eq_iff if_flip
@ -5437,7 +5360,7 @@ lemma fastpath_callKernel_SysCall_corres:
n_msgRegisters_def order_less_imp_le
ep_q_refs_of'_def st_tcb_at_refs_of_rev'
cong: if_cong)
apply (rename_tac blockedThread ys tcba tcbb st v tcbc)
apply (rename_tac blockedThread ys tcba tcbb v tcbc)
apply (frule invs_mdb')
apply (thin_tac "Ball S P" for S P)+
apply (clarsimp simp: invs'_def valid_state'_def)
@ -5518,7 +5441,7 @@ lemma doReplyTransfer_simple:
assert (mdbPrev mdbnode \<noteq> 0 \<and> mdbNext mdbnode = 0);
parentCTE \<leftarrow> getCTE (mdbPrev mdbnode);
assert (isReplyCap (cteCap parentCTE) \<and> capReplyMaster (cteCap parentCTE));
doIPCTransfer sender Nothing 0 True receiver False;
doIPCTransfer sender Nothing 0 True receiver;
cteDeleteOne slot;
setThreadState Running receiver;
attemptSwitchTo receiver
@ -5548,7 +5471,7 @@ lemma receiveIPC_simple_rewrite:
(\<lambda>s. \<forall>ntfnptr. bound_tcb_at' (op = (Some ntfnptr)) thread s \<longrightarrow> obj_at' (Not \<circ> isActive) ntfnptr s)))
(receiveIPC thread ep_cap True)
(do
setThreadState (BlockedOnReceive (capEPPtr ep_cap) (\<not> capEPCanSend ep_cap)) thread;
setThreadState (BlockedOnReceive (capEPPtr ep_cap)) thread;
setEndpoint (capEPPtr ep_cap) (RecvEP (case ep of RecvEP q \<Rightarrow> (q @ [thread]) | _ \<Rightarrow> [thread]))
od)"
apply (rule monadic_rewrite_gen_asm)

View File

@ -2907,7 +2907,7 @@ lemma cancelIPC_ccorres_reply_helper:
done
lemma ep_blocked_in_queueD_recv:
"\<lbrakk>st_tcb_at' (op = (Structures_H.thread_state.BlockedOnReceive x xa)) thread \<sigma>; ko_at' ep' x \<sigma>; invs' \<sigma>\<rbrakk> \<Longrightarrow> thread \<in> set (epQueue ep') \<and> isRecvEP ep'"
"\<lbrakk>st_tcb_at' (op = (Structures_H.thread_state.BlockedOnReceive x)) thread \<sigma>; ko_at' ep' x \<sigma>; invs' \<sigma>\<rbrakk> \<Longrightarrow> thread \<in> set (epQueue ep') \<and> isRecvEP ep'"
apply (frule sym_refs_st_tcb_atD', clarsimp)
apply (clarsimp simp: refs_of_rev' obj_at'_def ko_wp_at'_def projectKOs)
apply (cases ep', simp_all add: isSendEP_def isRecvEP_def)[1]

View File

@ -2377,7 +2377,7 @@ lemma transferCapsLoop_ccorres:
(if index (excaprefs_C caps) (unat \<acute>i) \<noteq> cte_Ptr 0 then 1
else 0))
FI"
defines "W \<equiv> \<lambda>ep diminish caps.
defines "W \<equiv> \<lambda>ep caps.
check1;; check2 caps;;
(While \<lbrace>\<acute>ret__int \<noteq> 0\<rbrace>
(Guard ArrayBounds \<lbrace>\<acute>i < 3\<rbrace> (\<acute>slot :== index (excaprefs_C caps) (unat \<acute>i));;
@ -2401,11 +2401,7 @@ lemma transferCapsLoop_ccorres:
IF \<acute>destSlot = cte_Ptr 0 THEN
break_C
FI;;
Cond {_. diminish \<noteq> 0}
(\<acute>ret__struct_cap_rights_C :== CALL cap_rights_new(scast true,scast true,scast false);;
\<acute>ret__struct_cap_C :== CALL maskCapRights(\<acute>ret__struct_cap_rights_C,\<acute>cap);;
\<acute>dc_ret :== CALL deriveCap(\<acute>slot,\<acute>ret__struct_cap_C))
(\<acute>dc_ret :== CALL deriveCap(\<acute>slot,\<acute>cap));;
\<acute>dc_ret :== CALL deriveCap(\<acute>slot,\<acute>cap);;
IF deriveCap_ret_C.status_C \<acute>dc_ret \<noteq> scast EXCEPTION_NONE THEN
break_C
FI;;
@ -2442,8 +2438,8 @@ lemma transferCapsLoop_ccorres:
(\<lambda>_. n + length caps \<le> 3 \<and> distinct slots ))
(precond n mi slots)
[catchbrk_C]
(transferCapsToSlots ep (to_bool diminish) rcv_buffer n caps slots mi)
(W ep diminish caps')"
(transferCapsToSlots ep rcv_buffer n caps slots mi)
(W ep caps')"
unfolding W_def check1_def check2_def split_def
proof (rule ccorres_gen_asm, induct caps arbitrary: n slots mi)
note split_if[split]
@ -2617,7 +2613,7 @@ next
apply csymbr
apply wpc
apply (rule ccorres_cond_true_seq)
apply (simp add: returnOk_liftE split del: split_if)
apply (simp add: returnOk_liftE)
apply (rule ccorres_case_sum_liftE)
apply (rule ccorres_split_throws)
apply (rule_tac P=\<top> and P'="?S" in ccorres_break_return)
@ -2625,87 +2621,8 @@ next
apply simp
apply vcg
apply (rule ccorres_cond_false_seq)
apply (simp split del: split_if)
apply (rule ccorres_Cond_rhs_Seq[unfolded Collect_const])
-- " case diminish "
apply (rule ccorres_rhs_assoc)+
apply csymbr
apply (rule ccorres_return_bind_add [where f="RetypeDecls_H.maskCapRights (capAllowWrite_update \<bottom> allRights)"])
apply ctac
apply (rule_tac P = "rv = maskCapRights (capAllowWrite_update (\<lambda>_. False) allRights) (fst x)"
in ccorres_gen_asm)
apply (rule ccorres_split_nothrowE)
apply (rule unifyFailure_ccorres)
apply (ctac add: deriveCap_ccorres')
apply ceqv
apply (simp split del: split_if)
apply csymbr
apply (rule_tac Q=\<top> and Q'=\<top> in ccorres_split_when_throwError_cond_break)
apply (clarsimp simp: cap_get_tag_isCap Collect_const_mem)
apply (rule_tac P=\<top> and P'="?S" in ccorres_break)
apply clarsimp
apply simp
apply (simp add: liftE_bindE split del: split_if)
apply (ctac add: cteInsert_ccorres)
apply (rule ccorres_case_sum_liftE)
apply csymbr
apply (rule ccorres_symb_exec_r)
apply (simp add: Collect_const[symmetric] del: Collect_const)
apply (rule ccorres_rhs_assoc2)
apply (rule_tac P="ccorresG rf_sr \<Gamma> r xf Pre Pre' hs a" for r xf Pre Pre' hs a in rsubst)
apply (rule Cons.hyps)
apply (clarsimp simp: excaps_map_def dest!: drop_n_foo)
apply clarsimp
apply simp
apply vcg
apply (rule conseqPre, vcg)
apply clarsimp
apply (rule_tac P="length slots = 1" in hoare_gen_asm)
apply (simp)
apply (wp cteInsert_valid_pspace hoare_valid_ipc_buffer_ptr_typ_at'
hoare_vcg_const_Ball_lift cteInsert_cte_wp_at)
apply (clarsimp)
apply (vcg exspec=cteInsert_modifies)
apply vcg
apply (simp)
apply (rule ccorres_split_throws)
apply (rule_tac P=\<top> and P'="?S" in ccorres_break)
apply clarsimp
apply simp
apply vcg
apply wp
apply (simp add: to_bool_def cong: conj_cong)
apply (rule_tac Q' ="\<lambda>rcap s. s \<turnstile>' rcap
\<and> (rcap\<noteq> NullCap \<longrightarrow> cte_wp_at' (is_derived' (ctes_of s) (snd x) rcap \<circ> cteCap) (snd x) s)
\<and> valid_pspace' s \<and> valid_ipc_buffer_ptr' rcv_buffer s \<and> length slots = 1
\<and> \<not> is_the_ep rcap
\<and> (\<forall>x\<in>set slots. cte_wp_at' (isNullCap \<circ> cteCap) x s)
\<and> (\<forall>x\<in>set xs'. s \<turnstile>' fst x
\<and> cte_wp_at' (\<lambda>c. is_the_ep (cteCap c) \<longrightarrow> fst x = cteCap c) (snd x) s
\<and> cte_wp_at' (\<lambda>c. fst x \<noteq> NullCap \<longrightarrow> stable (fst x) (cteCap c)) (snd x) s)"
in hoare_post_imp_R)
prefer 2
apply (clarsimp simp:cte_wp_at_ctes_of valid_pspace_mdb' valid_pspace'_splits
valid_pspace_valid_objs' is_derived_capMasterCap image_def)
apply (clarsimp split:if_splits)
apply (rule conjI)
apply clarsimp+
apply (rule conjI)
apply (drule(1) bspec)+
apply (rule conjI | clarsimp)+
apply (clarsimp simp:is_the_ep_def isCap_simps stable_def)
apply (drule(1) bspec)+
apply (rule conjI | clarsimp)+
apply (clarsimp simp:is_the_ep_def stable_def split:if_splits)+
apply (case_tac "a = cteCap cteb",clarsimp)
apply (simp add:maskedAsFull_def split:if_splits)
apply (simp add:maskedAsFull_again)
apply (wp deriveCap_derived is_the_ep_deriveCap)
apply (simp split del: split_if)
apply (vcg exspec=deriveCap_modifies)
apply wp
apply (simp split del: split_if)
apply (vcg exspec= maskCapRights_modifies)
apply (simp)
-- "case not diminish"
apply (rule ccorres_split_nothrowE)
apply (rule unifyFailure_ccorres)
@ -2918,12 +2835,11 @@ lemma transferCaps_ccorres [corres]:
\<inter> \<lbrace>\<acute>receiver = tcb_ptr_to_ctcb_ptr receiver\<rbrace>
\<inter> \<lbrace> mi = message_info_to_H \<acute>info\<rbrace>
\<inter> \<lbrace>\<acute>receiveBuffer = Ptr (option_to_0 receiveBuffer)\<rbrace>
\<inter> \<lbrace>\<acute>endpoint = option_to_ptr ep\<rbrace>
\<inter> \<lbrace>to_bool \<acute>diminish = diminish\<rbrace>) []
(transferCaps mi caps ep receiver receiveBuffer diminish)
\<inter> \<lbrace>\<acute>endpoint = option_to_ptr ep\<rbrace>) []
(transferCaps mi caps ep receiver receiveBuffer)
(Call transferCaps_'proc)" (is "ccorres _ _ ?P _ _ _ _")
apply (unfold K_def, intro ccorres_gen_asm)
apply (cinit lift: caps_' receiver_' info_' receiveBuffer_' endpoint_' diminish_'
apply (cinit lift: caps_' receiver_' info_' receiveBuffer_' endpoint_'
simp: getThreadCSpaceRoot_def locateSlot_conv whileAnno_def)
apply csymbr+
apply (rule_tac P="?P" and P'="{s. info_' s = info}" in ccorres_inst)
@ -3324,10 +3240,9 @@ lemma doNormalTransfer_ccorres [corres]:
\<inter> \<lbrace>\<acute>receiveBuffer = Ptr (option_to_0 receiveBuffer)\<rbrace>
\<inter> \<lbrace>canGrant = to_bool \<acute>canGrant\<rbrace>
\<inter> \<lbrace>\<acute>badge = badge\<rbrace>
\<inter> \<lbrace>\<acute>endpoint = option_to_ptr endpoint\<rbrace>
\<inter> \<lbrace>to_bool \<acute>diminish = diminish\<rbrace>) []
\<inter> \<lbrace>\<acute>endpoint = option_to_ptr endpoint\<rbrace>) []
(doNormalTransfer sender sendBuffer endpoint badge canGrant
receiver receiveBuffer diminish)
receiver receiveBuffer)
(Call doNormalTransfer_'proc)"
proof -
note split_if[split del]
@ -3343,7 +3258,7 @@ proof -
show ?thesis
apply (cinit lift: sender_' receiver_' sendBuffer_' receiveBuffer_'
canGrant_' badge_' endpoint_' diminish_'
canGrant_' badge_' endpoint_'
cong: call_ignore_cong)
apply (clarsimp cong: call_ignore_cong simp del: dc_simp)
apply (ctac(c_lines 2, no_vcg) add: getMessageInfo_ccorres')
@ -3596,11 +3511,10 @@ lemma doIPCTransfer_ccorres [corres]:
\<inter> \<lbrace>\<acute>receiver = tcb_ptr_to_ctcb_ptr receiver\<rbrace>
\<inter> \<lbrace>canGrant = to_bool \<acute>grant\<rbrace>
\<inter> \<lbrace>\<acute>badge = badge\<rbrace>
\<inter> \<lbrace>\<acute>endpoint = option_to_ptr endpoint\<rbrace>
\<inter> \<lbrace>to_bool \<acute>diminish = diminish\<rbrace>) []
(doIPCTransfer sender endpoint badge canGrant receiver diminish)
\<inter> \<lbrace>\<acute>endpoint = option_to_ptr endpoint\<rbrace>) []
(doIPCTransfer sender endpoint badge canGrant receiver)
(Call doIPCTransfer_'proc)"
apply (cinit lift: sender_' receiver_' grant_' badge_' endpoint_' diminish_')
apply (cinit lift: sender_' receiver_' grant_' badge_' endpoint_')
apply (rule_tac xf'="\<lambda>s. ptr_coerce (receiveBuffer___ptr_to_void_' s)"
in ccorres_split_nothrow_call_novcg)
apply (rule lookupIPCBuffer_ccorres)
@ -4117,7 +4031,7 @@ lemma transferCapsToSlots_local_slots:
assumes weak: "\<And>c cap. P (maskedAsFull c cap) = P c"
shows
"\<lbrace> cte_wp_at' (\<lambda>cte. P (cteCap cte)) slot and K (slot \<notin> set destSlots) \<rbrace>
transferCapsToSlots ep dim rcvBuffer x caps destSlots mi
transferCapsToSlots ep rcvBuffer x caps destSlots mi
\<lbrace>\<lambda>tag'. cte_wp_at' (\<lambda>cte. P (cteCap cte)) slot\<rbrace>"
proof (rule hoare_gen_asm, induct caps arbitrary: x mi destSlots)
case Nil show ?case by simp
@ -4134,7 +4048,7 @@ qed
lemma transferCaps_local_slots:
assumes weak: "\<And>c cap. P (maskedAsFull c cap) = P c"
shows "\<lbrace> valid_objs' and (Not o real_cte_at' slot) and cte_wp_at' (\<lambda>cte. P (cteCap cte)) slot \<rbrace>
transferCaps tag caps ep receiver receiveBuffer dim
transferCaps tag caps ep receiver receiveBuffer
\<lbrace>\<lambda>tag'. cte_wp_at' (\<lambda>cte. P (cteCap cte)) slot\<rbrace>"
apply (simp add: transferCaps_def pred_conj_def)
apply (rule hoare_seq_ext[rotated])
@ -4150,7 +4064,7 @@ lemma doNormalTransfer_local_slots:
assumes weak: "\<And>c cap. P (maskedAsFull c cap) = P c"
shows "\<lbrace> valid_objs' and (Not o real_cte_at' slot)
and cte_wp_at' (\<lambda>cte. P (cteCap cte)) slot \<rbrace>
doNormalTransfer sender sendBuffer ep badge grant receiver receiveBuffer dim
doNormalTransfer sender sendBuffer ep badge grant receiver receiveBuffer
\<lbrace>\<lambda>rv. cte_wp_at' (\<lambda>cte. P (cteCap cte)) slot\<rbrace>"
apply (simp add: doNormalTransfer_def)
apply (wp transferCaps_local_slots weak copyMRs_typ_at'[where T=CTET, unfolded typ_at_cte]
@ -4161,7 +4075,7 @@ lemma doIPCTransfer_local_slots:
assumes weak: "\<And>c cap. P (maskedAsFull c cap) = P c"
shows "\<lbrace> valid_objs' and (Not o real_cte_at' slot)
and cte_wp_at' (\<lambda>cte. P (cteCap cte)) slot \<rbrace>
doIPCTransfer sender ep badge grant receiver dim
doIPCTransfer sender ep badge grant receiver
\<lbrace> \<lambda>rv. cte_wp_at' (\<lambda>cte. P (cteCap cte)) slot \<rbrace>"
apply (simp add: doIPCTransfer_def)
apply (wp doNormalTransfer_local_slots weak threadGet_wp | wpc)+
@ -4172,7 +4086,7 @@ lemma doIPCTransfer_reply_or_replyslot:
"\<lbrace> cte_wp_at' (\<lambda>cte. isReplyCap (cteCap cte)) slot
or (valid_objs' and (Not o real_cte_at' slot)
and cte_wp_at' (\<lambda>cte. cteCap cte = capability.NullCap \<or> isReplyCap (cteCap cte)) slot)\<rbrace>
doIPCTransfer sender ep badge grant receiver dim
doIPCTransfer sender ep badge grant receiver
\<lbrace> \<lambda>rv. cte_wp_at' (\<lambda>cte. cteCap cte = capability.NullCap \<or> isReplyCap (cteCap cte)) slot\<rbrace>"
apply (rule hoare_name_pre_state)
apply (case_tac "cte_wp_at' (\<lambda>cte. isReplyCap (cteCap cte)) slot s")
@ -4954,39 +4868,29 @@ lemma sendIPC_ccorres [corres]:
apply wpc
-- "RecvEP case"
apply (rule ccorres_cond_false)
apply (rule ccorres_cond_true)
apply (rule ccorres_cond_true)
apply (intro ccorres_rhs_assoc)
apply (csymbr, csymbr, csymbr)
apply (csymbr, csymbr)
apply wpc
apply (simp only: haskell_fail_def)
apply (rule ccorres_fail)
apply (rule ccorres_fail)
apply (rename_tac dest rest)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_split_nothrow_novcg)
apply (rule ccorres_split_nothrow_novcg)
apply (rule_tac dest=dest in sendIPC_dequeue_ccorres_helper)
apply simp
apply ceqv
apply (rename_tac dest')
apply (simp only: K_bind_def haskell_assert_def return_bind)
apply (rule ccorres_move_c_guard_tcb)
apply (rule getThreadState_ccorres_foo)
apply (rename_tac recvState)
apply (rule ccorres_assert)
apply (rule_tac val="from_bool (blockingIPCDiminishCaps recvState)"
and xf'=diminish_'
and R="\<lambda>s. \<exists>t. ko_at' t dest s \<and> tcbState t = recvState"
in ccorres_symb_exec_r_known_rv_UNIV [where R'=UNIV])
apply (vcg, clarsimp)
apply (erule(1) cmap_relation_ko_atE [OF cmap_relation_tcb])
apply (clarsimp simp: ctcb_relation_def typ_heap_simps
cthread_state_relation_def word_size
isReceive_def thread_state_lift_def
split: thread_state.splits)
apply ceqv
apply (ctac(no_vcg))
apply (ctac(no_vcg))
apply (ctac(no_vcg))
@ -5006,13 +4910,13 @@ lemma sendIPC_ccorres [corres]:
apply (ctac add: setupCallerCap_ccorres)
apply (ctac add: setThreadState_ccorres)
apply vcg
apply (rule conseqPre, vcg, clarsimp)
apply (rule conseqPre, vcg, clarsimp) (* do_call = False *)
apply (clarsimp simp: from_bool_def split del: split_if)
apply (fold dc_def)[1]
apply (rule ccorres_symb_exec_r)
apply (rule ccorres_cond_false_seq)
apply (intro ccorres_rhs_assoc)
apply (rule ccorres_move_c_guard_tcb [simplified])
apply (rule ccorres_move_c_guard_tcb [simplified])
apply (rule_tac val="case_option (scast fault_null_fault)
fault_to_fault_tag fault"
and xf'=ret__unsigned_'
@ -5047,6 +4951,7 @@ lemma sendIPC_ccorres [corres]:
guard_is_UNIV_def ThreadState_Inactive_def)
apply vcg
apply (rule conseqPre, vcg, clarsimp)
apply (rule_tac Q="\<lambda>rv. valid_queues and valid_pspace' and valid_objs'
and sch_act_not thread and tcb_at' thread
and tcb_at' dest and
@ -5059,10 +4964,6 @@ lemma sendIPC_ccorres [corres]:
attemptSwitchTo_ksQ sts_valid_queues sts_ksQ')
apply (clarsimp simp: valid_tcb_state'_def)
apply (wp weak_sch_act_wf_lift_linear tcb_in_cur_domain'_lift)
apply (clarsimp simp: guard_is_UNIV_def ThreadState_Inactive_def
ThreadState_Running_def mask_def from_bool_def
option_to_ptr_def option_to_0_def
split: bool.split_asm)
apply (rule_tac Q="\<lambda>rv. valid_queues and valid_pspace' and valid_objs'
and valid_mdb' and tcb_at' dest and cur_tcb'
and tcb_at' thread and K (dest \<noteq> thread)
@ -5073,7 +4974,11 @@ lemma sendIPC_ccorres [corres]:
apply (clarsimp simp: st_tcb_at'_def obj_at'_def is_tcb weak_sch_act_wf_def)
apply (wp setEndpoint_ksQ hoare_vcg_all_lift set_ep_valid_objs'
setEndpoint_valid_mdb')
apply (clarsimp simp: guard_is_UNIV_def)
apply (clarsimp simp: guard_is_UNIV_def ThreadState_Inactive_def
ThreadState_Running_def mask_def from_bool_def
option_to_ptr_def option_to_0_def
split: bool.split_asm)
-- "IdleEP case"
apply (rule ccorres_cond_true)
apply (rule ccorres_cond)
@ -5143,7 +5048,7 @@ lemma sendIPC_ccorres [corres]:
split: endpoint.splits)
apply (frule(1) sym_refs_obj_atD'[OF _ invs_sym'])
apply (clarsimp simp: st_tcb_at_refs_of_rev' isBlockedOnReceive_def)
apply (auto split: list.splits elim!: pred_tcb'_weakenE)[1]
apply (auto split: list.splits elim!: pred_tcb'_weakenE)[1]
apply (subgoal_tac "state_refs_of' s epptr = {}")
apply (clarsimp simp: obj_at'_def is_aligned_neg_mask objBitsKO_def
projectKOs invs'_def valid_state'_def
@ -5179,19 +5084,19 @@ lemma sendIPC_ccorres [corres]:
apply (frule simple_st_tcb_at_state_refs_ofD')
apply (case_tac ep, auto simp: st_tcb_at_refs_of_rev' st_tcb_at'_def
obj_at'_def projectKOs)[1]
apply (clarsimp simp: guard_is_UNIV_def)+
done
apply (clarsimp simp: guard_is_UNIV_def)
apply (clarsimp simp: guard_is_UNIV_def)
done
lemma receiveIPC_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
(\<lambda>s. sch_act_wf (ksSchedulerAction s) s \<and>
(\<forall>d p. thread \<notin> set (ksReadyQueues s (d, p)))) and
K (epptr = epptr && ~~ mask 4) and
K (diminish' = from_bool diminish))
K (epptr = epptr && ~~ mask 4))
UNIV hs
(setThreadState (Structures_H.thread_state.BlockedOnReceive
epptr diminish) thread)
epptr) thread)
(Guard C_Guard \<lbrace>hrs_htd \<acute>t_hrs \<Turnstile>\<^sub>t tcb_ptr_to_ctcb_ptr thread\<rbrace>
(CALL thread_state_ptr_set_tsType(Ptr
&(tcb_ptr_to_ctcb_ptr thread\<rightarrow>[''tcbState_C'']),
@ -5200,10 +5105,6 @@ lemma receiveIPC_block_ccorres_helper:
(CALL thread_state_ptr_set_blockingObject(Ptr
&(tcb_ptr_to_ctcb_ptr thread\<rightarrow>[''tcbState_C'']),
ucast (ptr_val (ep_Ptr epptr))));;
Guard C_Guard \<lbrace>hrs_htd \<acute>t_hrs \<Turnstile>\<^sub>t tcb_ptr_to_ctcb_ptr thread\<rbrace>
(CALL thread_state_ptr_set_blockingIPCDiminishCaps(Ptr
&(tcb_ptr_to_ctcb_ptr thread\<rightarrow>[''tcbState_C'']),
diminish'));;
CALL scheduleTCB(tcb_ptr_to_ctcb_ptr thread))"
unfolding K_def setThreadState_def
apply (intro ccorres_gen_asm)
@ -5222,7 +5123,6 @@ lemma receiveIPC_block_ccorres_helper:
apply (simp add: ctcb_relation_def cthread_state_relation_def
ThreadState_BlockedOnReceive_def mask_def
from_bool_def to_bool_def)
apply (clarsimp split: bool.split)
apply ceqv
apply clarsimp
apply ctac
@ -5616,16 +5516,6 @@ lemma receiveIPC_ccorres [corres]:
apply (simp add: cap_endpoint_cap_lift ccap_relation_def cap_to_H_def)
apply ceqv
apply csymbr
apply (rule_tac xf'=ret__unsigned_'
and F="\<lambda>rv. to_bool rv = capEPCanSend cap"
and R=\<top>
in ccorres_symb_exec_r_abstract_UNIV[where R'=UNIV])
apply vcg
apply (clarsimp simp: cap_get_tag_isCap isCap_simps)
apply (frule cap_get_tag_isCap_unfolded_H_cap)
apply (simp add: cap_endpoint_cap_lift ccap_relation_def cap_to_H_def)
apply ceqv
apply (csymbr, rename_tac canSend)
apply (rule ccorres_move_c_guard_tcb)
apply (rule_tac xf'=ntfnPtr_'
and r'="\<lambda>rv rv'. rv' = option_to_ptr rv \<and> rv \<noteq> Some 0"
@ -5691,7 +5581,6 @@ lemma receiveIPC_ccorres [corres]:
apply (intro ccorres_rhs_assoc)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_split_nothrow_novcg)
apply (simp split del: split_if)
apply (rule receiveIPC_block_ccorres_helper[unfolded ptr_val_def, simplified])
@ -5721,7 +5610,6 @@ lemma receiveIPC_ccorres [corres]:
apply (intro ccorres_rhs_assoc)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_split_nothrow_novcg)
apply (simp split del: split_if)
apply (rule receiveIPC_block_ccorres_helper[unfolded ptr_val_def, simplified])

View File

@ -1678,7 +1678,7 @@ where
| "thread_state_to_tsType (Structures_H.Inactive) = scast ThreadState_Inactive"
| "thread_state_to_tsType (Structures_H.IdleThreadState) = scast ThreadState_IdleThreadState"
| "thread_state_to_tsType (Structures_H.BlockedOnReply) = scast ThreadState_BlockedOnReply"
| "thread_state_to_tsType (Structures_H.BlockedOnReceive oref dimin) = scast ThreadState_BlockedOnReceive"
| "thread_state_to_tsType (Structures_H.BlockedOnReceive oref) = scast ThreadState_BlockedOnReceive"
| "thread_state_to_tsType (Structures_H.BlockedOnSend oref badge cg isc) = scast ThreadState_BlockedOnSend"
| "thread_state_to_tsType (Structures_H.BlockedOnNotification oref) = scast ThreadState_BlockedOnNotification"

View File

@ -237,10 +237,9 @@ where
= (tsType_CL (fst ts') = scast ThreadState_IdleThreadState)"
| "cthread_state_relation_lifted (Structures_H.BlockedOnReply) ts'
= (tsType_CL (fst ts') = scast ThreadState_BlockedOnReply)"
| "cthread_state_relation_lifted (Structures_H.BlockedOnReceive oref dimin) ts'
| "cthread_state_relation_lifted (Structures_H.BlockedOnReceive oref) ts'
= (tsType_CL (fst ts') = scast ThreadState_BlockedOnReceive \<and>
oref = blockingObject_CL (fst ts') \<and>
dimin = to_bool (blockingIPCDiminishCaps_CL (fst ts')))"
oref = blockingObject_CL (fst ts'))"
| "cthread_state_relation_lifted (Structures_H.BlockedOnSend oref badge cg isc) ts'
= (tsType_CL (fst ts') = scast ThreadState_BlockedOnSend
\<and> oref = blockingObject_CL (fst ts')

View File

@ -1625,14 +1625,171 @@ lemma ucast_eq_0[OF refl]:
apply (drule inj_eq[where x=x and y=0], simp)
done
lemma is_up_compose':
fixes uc :: "('a::len) word \<Rightarrow> ('b::len) word"
and uc' :: "'b word \<Rightarrow> ('c::len) sword"
shows
"\<lbrakk>is_up uc; is_up uc'\<rbrakk> \<Longrightarrow> is_up (uc' \<circ> uc)"
unfolding is_up_def by (simp add: Word.target_size Word.source_size)
lemma is_up_compose:
shows
"\<lbrakk>is_up uc; is_up uc'\<rbrakk> \<Longrightarrow> is_up (uc' \<circ> uc)"
unfolding is_up_def by (simp add: Word.target_size Word.source_size)
lemma uint_is_up_compose:
fixes uc :: "('a::len) word \<Rightarrow> ('b::len) word"
and uc' :: "'b word \<Rightarrow> ('c::len) sword"
assumes "uc = ucast"
and "uc' = ucast"
and " uuc = uc' \<circ> uc"
shows
"\<lbrakk> is_up uc; is_up uc' \<rbrakk> \<Longrightarrow> uint (uuc b) = uint b"
apply (simp add: assms)
apply (frule is_up_compose)
apply (simp_all )
apply (simp only: Word.uint_up_ucast)
done
lemma uint_is_up_compose_pred:
fixes uc :: "('a::len) word \<Rightarrow> ('b::len) word"
and uc' :: "'b word \<Rightarrow> ('c::len) sword"
assumes "uc = ucast"
and "uc' = ucast"
and " uuc = uc' \<circ> uc"
shows
"\<lbrakk> is_up uc; is_up uc' \<rbrakk> \<Longrightarrow> P (uint (uuc b)) \<longleftrightarrow> P( uint b)"
apply (simp add: assms)
apply (frule is_up_compose)
apply (simp_all )
apply (simp only: Word.uint_up_ucast)
done
lemma is_down_up_sword:
fixes uc :: "('a::len) word \<Rightarrow> ('b::len) sword"
shows "\<lbrakk>uc = ucast; len_of TYPE('a) < len_of TYPE('b) \<rbrakk> \<Longrightarrow> is_up uc = (\<not> is_down uc)"
by (simp add: target_size source_size is_up_def is_down_def )
lemma is_not_down_compose:
fixes uc :: "('a::len) word \<Rightarrow> ('b::len) word"
and uc' :: "'b word \<Rightarrow> ('c::len) sword"
shows
"\<lbrakk>uc = ucast; uc' = ucast; len_of TYPE('a) < len_of TYPE('c)\<rbrakk> \<Longrightarrow> \<not> is_down (uc' \<circ> uc) "
unfolding is_down_def
by (simp add: Word.target_size Word.source_size)
lemma sint_ucast_uint:
fixes uc :: "('a::len) word \<Rightarrow> ('b::len) word"
and uc' :: "'b word \<Rightarrow> ('c::len) sword"
assumes "uc = ucast" and " uc' = ucast" and "uuc=uc' \<circ> uc " and "len_of TYPE('a) < len_of TYPE('c signed)"
shows
"\<lbrakk> is_up uc; is_up uc'\<rbrakk> \<Longrightarrow> sint (uuc b) = uint b"
apply (simp add: assms)
apply (frule is_up_compose')
apply simp_all
apply (simp add: ucast_ucast_b )
apply (rule WordLemmaBucket.sint_ucast_eq_uint )
apply (insert assms)
apply (simp add: is_down_def target_size source_size )
done
lemma sint_ucast_uint_pred:
fixes uc :: "('a::len) word \<Rightarrow> ('b::len) word"
and uc' :: "'b word \<Rightarrow> ('c::len) sword"
and uuc :: "'a word \<Rightarrow> 'c sword"
assumes "uc = ucast" and " uc' = ucast" and "uuc=uc' \<circ> uc " and "len_of TYPE('a) < len_of TYPE('c )"
shows "\<lbrakk>is_up uc; is_up uc'\<rbrakk> \<Longrightarrow> P (uint b) \<longleftrightarrow> P (sint (uuc b))"
apply (simp add: assms )
apply (insert sint_ucast_uint[where uc=uc and uc'=uc' and uuc=uuc and b = b])
apply (simp add: assms)
done
lemma sint_uucast_uint_uucast_pred:
fixes uc :: "('a::len) word \<Rightarrow> ('b::len) word"
and uc' :: "'b word \<Rightarrow> ('c::len) sword"
assumes "uc = ucast" and " uc' = ucast" and "uuc=uc' \<circ> uc " and "len_of TYPE('a) < len_of TYPE('c )"
shows "\<lbrakk>is_up uc; is_up uc'\<rbrakk> \<Longrightarrow> P (uint(uuc b)) \<longleftrightarrow> P (sint (uuc b))"
apply (simp add: assms )
apply (insert sint_ucast_uint[where uc=uc and uc'=uc' and uuc=uuc and b = b])
apply (insert uint_is_up_compose_pred[where uc=uc and uc'=uc' and uuc=uuc and b=b])
apply (simp add: assms uint_is_up_compose_pred)
done
lemma scast_maxIRQ_is_less:
fixes uc :: "irq \<Rightarrow> 16 word"
and uc' :: "16 word\<Rightarrow> 32 sword"
and b :: irq
shows
"(Kernel_C.maxIRQ) <s (ucast \<circ> (ucast :: irq \<Rightarrow> 16 word)) b \<Longrightarrow> scast Kernel_C.maxIRQ < b"
apply (simp add: maxIRQ_def word_sless_def word_sle_def, uint_arith, clarify,simp)
apply (subgoal_tac "sint (ucast Kernel_C.maxIRQ :: 32 sword) \<le> uint b"; (simp only: Kernel_C.maxIRQ_def)?)
apply (subgoal_tac "sint (ucast Kernel_C.maxIRQ :: 32 sword) \<noteq> uint b"; (simp only: Kernel_C.maxIRQ_def)?)
apply (simp )
apply (subst uint_is_up_compose[where uc="(ucast :: irq \<Rightarrow> 16 word)" and uc' = "(ucast :: 16 word \<Rightarrow> 32 sword)",symmetric];
(simp add: is_up_def target_size source_size )?)
apply fastforce
apply (subst sint_ucast_uint_pred[where uc="(ucast :: irq \<Rightarrow> 16 word)" and uc' = "(ucast :: 16 word \<Rightarrow> 32 sword)"];
(simp add: is_up_def target_size source_size )?)
apply fastforce
done
lemma validIRQcastingLess: "Kernel_C.maxIRQ <s (ucast((ucast (b :: irq))::word16)) \<Longrightarrow> Platform.maxIRQ < b"
by (simp add: Platform_maxIRQ scast_maxIRQ_is_less is_up_def target_size source_size)
lemma scast_maxIRQ_is_not_less: "(\<not> (Kernel_C.maxIRQ) <s (ucast \<circ> (ucast :: irq \<Rightarrow> 16 word)) b) \<Longrightarrow> \<not> (scast Kernel_C.maxIRQ < b)"
apply (subgoal_tac "sint (ucast Kernel_C.maxIRQ :: 32 sword) \<ge> sint (ucast (ucast b))";
(simp only: Kernel_C.maxIRQ_def word_sless_def word_sle_def )?)
apply (simp add: maxIRQ_def word_sless_def word_sle_def, uint_arith, clarify,simp)
apply (subst (asm) sint_ucast_uint[where b=b and 'c = 32 and 'b = 16 and uc=ucast and uc' = ucast and uuc = "ucast \<circ> ucast" , simplified];
(simp add: is_up_def target_size source_size)?)
apply (subst (asm) uint_is_up_compose[where 'b = 16 and uuc="ucast \<circ> ucast", simplified];
(simp add: is_up_def target_size source_size)?)
apply (simp add: maxIRQ_def word_sless_def word_sle_def, uint_arith, clarify,simp)
apply (subst (asm) (2) sint_ucast_uint_pred[where 'a = 10 and 'b = 16 and 'c = 32 and uuc = "ucast \<circ> ucast", simplified,symmetric ];
((simp add: is_up_def target_size source_size)?))
apply (subst sint_ucast_uint_pred[where 'a = 10 and 'b = 16 and 'c = 32 and uuc = "ucast \<circ> ucast", simplified,symmetric ];
((simp add: is_up_def target_size source_size)?))
apply (subst (asm) (2) uint_is_up_compose[where 'b = 16 and uuc="ucast \<circ> ucast", simplified];
(simp add: is_up_def target_size source_size)?)
apply (uint_arith)
done
lemma handleInterrupt_ccorres:
"ccorres dc xfdc
(invs' and K (irq \<le> ucast maxIRQ))
(invs')
(UNIV \<inter> \<lbrace>\<acute>irq = ucast irq\<rbrace>)
[]
(handleInterrupt irq)
(Call handleInterrupt_'proc)"
(Call handleInterrupt_'proc)"
apply (cinit lift: irq_' cong: call_ignore_cong)
apply (rule ccorres_Cond_rhs_Seq)
apply (simp add: Platform_maxIRQ del: Collect_const)
apply (drule scast_maxIRQ_is_less[simplified])
apply (simp del: Collect_const)
apply (rule ccorres_rhs_assoc)+
apply (simp del: Collect_const)
apply (subst doMachineOp_bind)
apply (rule maskInterrupt_empty_fail)
apply (rule ackInterrupt_empty_fail)
apply (ctac add: maskInterrupt_ccorres[unfolded dc_def])
apply (subst bind_return_unit[where f="doMachineOp (ackInterrupt irq)"])
apply (ctac add: ackInterrupt_ccorres[unfolded dc_def])
apply (rule ccorres_split_throws)
apply (rule ccorres_return_void_C[unfolded dc_def])
apply vcg
apply wp
apply (vcg exspec=ackInterrupt_modifies)
apply wp
apply (vcg exspec=maskInterrupt_modifies)
apply (simp add: scast_maxIRQ_is_not_less Platform_maxIRQ del: Collect_const)
apply (rule ccorres_pre_getIRQState)
apply wpc
apply simp
@ -1644,7 +1801,7 @@ lemma handleInterrupt_ccorres:
apply csymbr
apply (rule getIRQSlot_ccorres3)
apply (rule ccorres_getSlotCap_cte_at)
apply (rule_tac P="cte_at' rva" in ccorres_cross_over_guard)
apply (rule_tac P="cte_at' rv" in ccorres_cross_over_guard)
apply (rule ptr_add_assertion_irq_guard[unfolded dc_def])
apply (rule ccorres_move_array_assertion_irq ccorres_move_c_guard_cte)+
apply ctac
@ -1669,8 +1826,7 @@ lemma handleInterrupt_ccorres:
apply (simp del: Collect_const)
apply (rule ccorres_cond_true_seq)
apply (rule ccorres_rhs_assoc)+
apply csymbr
apply csymbr
apply csymbr+
apply (rule ccorres_cond_false_seq)
apply simp
apply (ctac (no_vcg) add: maskInterrupt_ccorres)
@ -1678,16 +1834,15 @@ lemma handleInterrupt_ccorres:
apply wp
apply (rule_tac P=\<top> and P'="{s. ret__int_' s = 0 \<and> cap_get_tag cap \<noteq> scast cap_notification_cap}" in ccorres_inst)
apply (clarsimp simp: isCap_simps simp del: Collect_const)
apply (case_tac rvb, simp_all del: Collect_const)[1]
prefer 3
apply metis
apply ((rule ccorres_guard_imp2,
rule ccorres_cond_false_seq, simp,
apply (case_tac rva, simp_all del: Collect_const)[1]
prefer 3
apply metis
apply ((rule ccorres_guard_imp2,
rule ccorres_cond_false_seq, simp,
rule ccorres_cond_false_seq, simp,
ctac (no_vcg) add: maskInterrupt_ccorres,
ctac (no_vcg) add: ackInterrupt_ccorres [unfolded dc_def],
wp, simp)+)
apply simp
apply (wp getSlotCap_wp)
apply simp
apply vcg
@ -1699,37 +1854,34 @@ lemma handleInterrupt_ccorres:
apply (rule ccorres_rhs_assoc)+
apply (ctac (no_vcg) add: timerTick_ccorres)
apply (ctac (no_vcg) add: resetTimer_ccorres)
apply (ctac add: ackInterrupt_ccorres)
apply (ctac add: ackInterrupt_ccorres )
apply wp
apply (simp add: sint_ucast_eq_uint is_down uint_up_ucast is_up)
apply (simp add: sint_ucast_eq_uint is_down uint_up_ucast is_up )
apply (clarsimp simp: word_sless_alt word_less_alt word_le_def maxIRQ_def uint_up_ucast is_up_def
source_size_def target_size_def word_size
sint_ucast_eq_uint is_down is_up word_0_sle_from_less)
source_size_def target_size_def word_size
sint_ucast_eq_uint is_down is_up word_0_sle_from_less)
apply (rule conjI)
apply (clarsimp simp: cte_wp_at_ctes_of ucast_up_ucast is_up)
apply (intro conjI allI)
apply (clarsimp simp add: if_1_0_0 Collect_const_mem)
apply (clarsimp simp: cte_wp_at_ctes_of )
apply (clarsimp simp add: if_1_0_0 Collect_const_mem )
apply (clarsimp simp: Kernel_C.IRQTimer_def Kernel_C.IRQSignal_def
cte_wp_at_ctes_of ucast_ucast_b is_up)
cte_wp_at_ctes_of ucast_ucast_b is_up)
apply (intro conjI impI)
apply clarsimp
apply (erule(1) cmap_relationE1[OF cmap_relation_cte])
apply (clarsimp simp: typ_heap_simps')
apply (simp add: cap_get_tag_isCap)
apply (clarsimp simp: isCap_simps)
apply (frule cap_get_tag_isCap_unfolded_H_cap)
apply (frule cap_get_tag_to_H, assumption)
apply (clarsimp simp: to_bool_def)
apply (cut_tac un_ui_le[where b = 159 and a = irq,
simplified word_size])
apply (simp add: ucast_eq_0 is_up_def source_size_def
target_size_def word_size unat_gt_0
| subst array_assertion_abs_irq[rule_format, OF conjI])+
apply clarsimp
apply (erule(1) cmap_relationE1[OF cmap_relation_cte])
apply (clarsimp simp: typ_heap_simps')
apply (simp add: cap_get_tag_isCap)
apply (clarsimp simp: isCap_simps)
apply (frule cap_get_tag_isCap_unfolded_H_cap)
apply (frule cap_get_tag_to_H, assumption)
apply (clarsimp simp: to_bool_def)
apply (cut_tac un_ui_le[where b = 159 and a = irq,
simplified word_size])
apply (simp add: ucast_eq_0 is_up_def source_size_def
target_size_def word_size unat_gt_0
| subst array_assertion_abs_irq[rule_format, OF conjI])+
apply (erule(1) rf_sr_cte_at_valid[OF ctes_of_cte_at])
apply (clarsimp simp:nat_le_iff)
done
end
end

View File

@ -43,7 +43,7 @@ lemma tcb_cap_casesE:
\<or> (halted st \<and> (c = cap.NullCap))) \<rbrakk> \<Longrightarrow> R"
"\<lbrakk> p = tcb_cnode_index 3; gf = tcb_caller; sf = tcb_caller_update; restr =
(\<lambda>_ st. case st of
Structures_A.BlockedOnReceive e d \<Rightarrow>
Structures_A.BlockedOnReceive e \<Rightarrow>
(op = cap.NullCap)
| _ \<Rightarrow> is_reply_cap or (op = cap.NullCap)) \<rbrakk> \<Longrightarrow> R"
"\<lbrakk> p = tcb_cnode_index 4; gf = tcb_ipcframe; sf = tcb_ipcframe_update; restr =
@ -101,7 +101,7 @@ lemma tcb_cap_cases_slot_simps[simp]:
\<or> (halted st \<and> (c = cap.NullCap))))"
"tcb_cap_cases (tcb_cnode_index tcb_caller_slot) = Some (tcb_caller, tcb_caller_update,
(\<lambda>_ st. case st of
Structures_A.BlockedOnReceive e d \<Rightarrow>
Structures_A.BlockedOnReceive e \<Rightarrow>
(op = cap.NullCap)
| _ \<Rightarrow> is_reply_cap or (op = cap.NullCap)))"
"tcb_cap_cases (tcb_cnode_index tcb_ipcbuffer_slot) = Some (tcb_ipcframe, tcb_ipcframe_update,
@ -187,7 +187,7 @@ definition
generates_pending :: "Structures_A.thread_state \<Rightarrow> bool"
where
"generates_pending st \<equiv> case st of
Structures_A.BlockedOnReceive ptr diminish \<Rightarrow> True
Structures_A.BlockedOnReceive ptr \<Rightarrow> True
| Structures_A.BlockedOnSend ptr payload \<Rightarrow> True
| Structures_A.BlockedOnReply \<Rightarrow> True
| Structures_A.BlockedOnNotification ptr \<Rightarrow> True

View File

@ -281,6 +281,12 @@ lemma handle_interrupt_corres:
"dcorres dc \<top> (invs and valid_etcbs) (Interrupt_D.handle_interrupt x) (Interrupt_A.handle_interrupt x)"
apply (clarsimp simp:Interrupt_A.handle_interrupt_def)
apply (clarsimp simp:get_irq_state_def gets_def bind_assoc)
apply (rule conjI; rule impI)
apply (subst Interrupt_D.handle_interrupt_def, simp)
apply (subst Retype_AI.do_machine_op_bind)
apply (rule maskInterrupt_empty_fail)
apply (rule ackInterrupt_empty_fail)
using corres_guard2_imp handle_interrupt_corres_branch apply blast
apply (rule dcorres_absorb_get_r)+
apply (clarsimp split:irq_state.splits simp:corres_free_fail | rule conjI)+
apply (simp add:Interrupt_D.handle_interrupt_def bind_assoc)

View File

@ -55,7 +55,7 @@ lemma do_fault_transfer_cur_thread_idle_thread:
done
lemma do_normal_transfer_cur_thread_idle_thread:
"\<lbrace>\<lambda>s. P (cur_thread s) (idle_thread s) \<rbrace> Ipc_A.do_normal_transfer a b c d e f g h\<lbrace>\<lambda>rv s. P (cur_thread s) (idle_thread s)\<rbrace>"
"\<lbrace>\<lambda>s. P (cur_thread s) (idle_thread s) \<rbrace> Ipc_A.do_normal_transfer a b c d e f g \<lbrace>\<lambda>rv s. P (cur_thread s) (idle_thread s)\<rbrace>"
apply (simp add:do_normal_transfer_def set_message_info_def)
apply (wp as_user_cur_thread_idle_thread |wpc|clarsimp)+
apply (wps | wp transfer_caps_it copy_mrs_it)+
@ -63,7 +63,7 @@ lemma do_normal_transfer_cur_thread_idle_thread:
done
lemma do_ipc_transfer_cur_thread_idle_thread:
"\<lbrace>\<lambda>s. P (cur_thread s) (idle_thread s) \<rbrace> Ipc_A.do_ipc_transfer a b c d e f\<lbrace>\<lambda>rv s. P (cur_thread s) (idle_thread s)\<rbrace>"
"\<lbrace>\<lambda>s. P (cur_thread s) (idle_thread s) \<rbrace> Ipc_A.do_ipc_transfer a b c d e \<lbrace>\<lambda>rv s. P (cur_thread s) (idle_thread s)\<rbrace>"
apply (simp add:do_ipc_transfer_def)
apply (wp do_fault_transfer_cur_thread_idle_thread do_normal_transfer_cur_thread_idle_thread|wpc)+
apply (wp | simp add:thread_get_def)+
@ -441,7 +441,7 @@ lemma block_thread_on_recv_corres:
"dcorres dc \<top> (not_idle_thread thread and valid_etcbs)
(KHeap_D.set_cap (thread, tcb_pending_op_slot)
(PendingSyncRecvCap epptr False))
(set_thread_state thread (Structures_A.thread_state.BlockedOnReceive epptr x))"
(set_thread_state thread (Structures_A.thread_state.BlockedOnReceive epptr))"
apply (clarsimp simp:KHeap_D.set_cap_def set_thread_state_def)
apply (rule dcorres_gets_the, clarsimp)
apply (rule dcorres_rhs_noop_below_True[OF set_thread_state_ext_dcorres])
@ -1040,7 +1040,7 @@ lemma set_thread_state_block_on_send_corres:
lemma set_thread_state_block_on_receive_corres:
"dcorres dc \<top> (not_idle_thread thread and valid_etcbs)
(block_thread_on_ipc thread (PendingSyncRecvCap epptr False))
(set_thread_state thread (Structures_A.thread_state.BlockedOnReceive epptr x))"
(set_thread_state thread (Structures_A.thread_state.BlockedOnReceive epptr))"
apply (simp add:block_thread_on_ipc_def)
apply (rule block_thread_on_recv_corres)
done
@ -1367,7 +1367,7 @@ lemma dcorres_transfer_caps_loop:
length dests \<le> 1 \<and>
2 + msg_max_length + n + length caps' < unat max_ipc_words) and valid_etcbs)
(Endpoint_D.transfer_caps_loop ep rcv caps dest)
(Ipc_A.transfer_caps_loop ep' d rcv_buffer n caps' dests mi)"
(Ipc_A.transfer_caps_loop ep' rcv_buffer n caps' dests mi)"
proof (induct caps' arbitrary: mi n dests dest caps ep ep')
case Nil thus ?case by clarsimp
next
@ -1416,20 +1416,9 @@ next
apply (rule corres_splitEE)
prefer 2
apply (subst alternative_bindE_distrib)
apply (cases d)
prefer 2
apply simp
apply (rule corres_alternate2)
apply (rule derive_cap_dcorres [where P="\<top>"], simp)
apply (rule corres_alternate1)
apply simp
apply (rule corres_guard_imp)
apply (rule derive_cap_dcorres [where P="\<top>"])
apply (unfold remove_rights_def)[1]
apply (rule update_cap_rights_cong [symmetric])
apply (simp add: transform_cap_rights)
apply assumption
apply (simp add: swp_def)
apply (rule corres_splitEE)
prefer 2
apply (rule corres_whenE, simp)
@ -1725,7 +1714,7 @@ lemma transfer_caps_dcorres:
(Endpoint_D.transfer_caps ep
(map (\<lambda>(cap, slot). (transform_cap cap, transform_cslot_ptr slot)) caps)
sender rcv)
(Ipc_A.transfer_caps info caps ep rcv rcv_buffer d)"
(Ipc_A.transfer_caps info caps ep rcv rcv_buffer)"
unfolding Ipc_A.transfer_caps_def Endpoint_D.transfer_caps_def
apply clarsimp
apply (cases rcv_buffer)
@ -2021,7 +2010,7 @@ lemma corres_complete_ipc_transfer:
and (\<lambda>s. not_idle_thread (cur_thread s) s) and valid_irq_node
and valid_idle and not_idle_thread recv and not_idle_thread send and valid_etcbs)
(Endpoint_D.do_ipc_transfer ep' send recv can_grant)
(Ipc_A.do_ipc_transfer send ep badge can_grant recv diminish)"
(Ipc_A.do_ipc_transfer send ep badge can_grant recv)"
apply (clarsimp simp: Endpoint_D.do_ipc_transfer_def Ipc_A.do_ipc_transfer_def)
apply (rule dcorres_expand_pfx)
apply (rule dcorres_symb_exec_r_evalMonad)
@ -2413,7 +2402,7 @@ lemma dcorres_receive_sync:
(case rv of
Structures_A.endpoint.IdleEP \<Rightarrow> case is_blocking of
True \<Rightarrow> do set_thread_state thread
(Structures_A.thread_state.BlockedOnReceive word1 (AllowWrite \<notin> rights));
(Structures_A.thread_state.BlockedOnReceive word1);
set_endpoint word1 (Structures_A.endpoint.RecvEP [thread])
od
| False \<Rightarrow> do_nbrecv_failed_transfer thread
@ -2428,10 +2417,10 @@ lemma dcorres_receive_sync:
data \<leftarrow> case sender_state of
Structures_A.thread_state.BlockedOnSend ref x \<Rightarrow> return x | _ \<Rightarrow> fail;
Ipc_A.do_ipc_transfer sender (Some word1) (sender_badge data)
(sender_can_grant data) thread (AllowWrite \<notin> rights);
(sender_can_grant data) thread;
fault \<leftarrow> thread_get tcb_fault sender;
if sender_is_call data \<or> fault \<noteq> None
then if sender_can_grant data \<and> \<not> AllowWrite \<notin> rights
then if sender_can_grant data
then setup_caller_cap sender thread
else set_thread_state sender Structures_A.thread_state.Inactive
else do set_thread_state sender Structures_A.thread_state.Running;
@ -2440,7 +2429,7 @@ lemma dcorres_receive_sync:
od
| Structures_A.endpoint.RecvEP queue \<Rightarrow> case is_blocking of
True \<Rightarrow> do set_thread_state thread
(Structures_A.thread_state.BlockedOnReceive word1 (AllowWrite \<notin> rights));
(Structures_A.thread_state.BlockedOnReceive word1);
set_endpoint word1 (Structures_A.endpoint.RecvEP (queue @ [thread]))
od
| False \<Rightarrow> do_nbrecv_failed_transfer thread)"
@ -2710,27 +2699,28 @@ lemma send_sync_ipc_corres:
apply (clarsimp simp: dest!: not_empty_list_not_empty_set)
apply (rename_tac list)
apply (drule_tac s = "set list" in sym)
apply (clarsimp simp: bind_assoc neq_Nil_conv)
apply (clarsimp simp: bind_assoc neq_Nil_conv split del:split_if)
apply (rule_tac P1="\<top>" and P'="op = s'a" and x1 = y
in dcorres_absorb_pfx[OF select_pick_corres[OF dcorres_expand_pfx]])
defer
apply simp+
apply (simp+)[3]
apply (simp split del:split_if)
apply (drule_tac x1 = y in iffD2[OF eqset_imp_iff], simp)
apply (clarsimp simp:obj_at_def dc_def[symmetric])
apply (clarsimp simp:obj_at_def dc_def[symmetric] split del:split_if)
apply (subst when_def)+
apply (rule corres_guard_imp)
apply (rule dcorres_symb_exec_r)
apply (rule corres_symb_exec_r)
apply (rule corres_symb_exec_r)
apply (case_tac "recv_state"; simp add: corres_free_fail split del: split_if)
apply (rule corres_split[OF _ corres_complete_ipc_transfer])
apply (rule corres_split[OF _ set_thread_state_corres])
apply (rule dcorres_rhs_noop_above[OF attempt_switch_to_dcorres])
apply (rule_tac corres_symb_exec_r)
apply (rule dcorres_if_rhs)
apply (rule dcorres_if_rhs)
apply simp
apply (simp only:if_True)
apply (rule corres_alternate1)+
apply (rule corres_setup_caller_cap)
apply (rule corres_setup_caller_cap)
apply (clarsimp simp:ep_waiting_set_recv_def pred_tcb_at_def obj_at_def generates_pending_def)
apply (rule corres_alternate1[OF corres_alternate2])
apply (rule set_thread_state_corres)
@ -2747,11 +2737,8 @@ lemma send_sync_ipc_corres:
apply (wp hoare_vcg_conj_lift)
apply (rule hoare_disjI1)
apply (wp do_ipc_transfer_pred_tcb | wpc | simp)+
apply (simp add:return_def fail_def valid_def split:Structures_A.thread_state.splits)
apply (clarsimp simp:conj_comms not_idle_thread_def)+
apply (rule hoare_drop_imp)
apply wp
apply clarsimp
apply (wp|wps)+
apply (simp add:not_idle_thread_def)
apply (clarsimp simp:ep_waiting_set_recv_def obj_at_def st_tcb_at_def)+

View File

@ -1440,7 +1440,7 @@ lemma fast_finalise_not_idle_thread[wp]:
lemma block_lift:
"\<lbrakk>kheap b word = Some (TCB tcb_type); ekheap b word = Some etcb; transform_tcb (machine_state b) word tcb_type etcb = Tcb cdl_tcb_type\<rbrakk>
\<Longrightarrow> is_thread_blocked_on_endpoint cdl_tcb_type ep = (case tcb_state tcb_type of
Structures_A.thread_state.BlockedOnReceive p _ \<Rightarrow> ep = p
Structures_A.thread_state.BlockedOnReceive p \<Rightarrow> ep = p
| Structures_A.thread_state.BlockedOnSend p _ \<Rightarrow> ep = p
| Structures_A.thread_state.BlockedOnNotification p \<Rightarrow> ep = p
| _ \<Rightarrow> False)"
@ -1472,8 +1472,8 @@ where "none_is_sending_ep epptr s \<equiv> (ep_waiting_set_send epptr s) = {}"
definition ep_waiting_set_recv :: "obj_ref\<Rightarrow> 'z::state_ext state\<Rightarrow>obj_ref set"
where "ep_waiting_set_recv epptr s \<equiv>
{tcb. \<exists>t diminish. ((kheap s tcb) = Some (TCB t))
\<and> ((tcb_state t) = Structures_A.thread_state.BlockedOnReceive epptr diminish)}"
{tcb. \<exists>t. ((kheap s tcb) = Some (TCB t))
\<and> ((tcb_state t) = Structures_A.thread_state.BlockedOnReceive epptr)}"
definition none_is_receiving_ep:: "obj_ref \<Rightarrow> 'z::state_ext state \<Rightarrow> bool"
where "none_is_receiving_ep epptr s \<equiv> (ep_waiting_set_recv epptr s) = {}"

View File

@ -663,7 +663,7 @@ definition
where
"infer_tcb_pending_op ptr t \<equiv>
case t of
Structures_A.BlockedOnReceive ptr diminish \<Rightarrow>
Structures_A.BlockedOnReceive ptr \<Rightarrow>
PendingSyncRecvCap ptr False
|Structures_A.BlockedOnReply \<Rightarrow>

View File

@ -1132,10 +1132,9 @@ lemma handle_preemption_if_irq_masks:
"\<lbrace>(\<lambda>s. P (irq_masks_of_state s)) and domain_sep_inv False st\<rbrace>
handle_preemption_if tc
\<lbrace>\<lambda>_ s. P (irq_masks_of_state s)\<rbrace>"
apply(simp add: handle_preemption_if_def | wp handle_interrupt_irq_masks)+
apply(rule_tac Q="\<lambda>_ s. P (irq_masks_of_state s) \<and> domain_sep_inv False st s" in hoare_strengthen_post)
apply(wp | simp)+
by blast
apply(simp add: handle_preemption_if_def | wp handle_interrupt_irq_masks[where st=st])+
apply(rule_tac Q="\<lambda>rv s. P (irq_masks_of_state s) \<and> domain_sep_inv False st s \<and> (\<forall>x. rv = Some x \<longrightarrow> x \<le> maxIRQ) " in hoare_strengthen_post)
by(wp | simp)+
crunch valid_list[wp]: handle_preemption_if "valid_list"
(ignore: getActiveIRQ)
@ -1720,8 +1719,9 @@ lemma handle_interrupt_domain_time_sched_action:
"num_domains > 1 \<Longrightarrow>
\<lbrace>\<lambda>s. domain_time s > 0\<rbrace>
handle_interrupt e
\<lbrace>\<lambda>r s. domain_time s = 0 \<longrightarrow> scheduler_action s = choose_new_thread\<rbrace>"
apply(simp add: handle_interrupt_def)
\<lbrace>\<lambda>r s. domain_time s = 0 \<longrightarrow> scheduler_action s = choose_new_thread\<rbrace>"
apply(simp add: handle_interrupt_def split del: split_if)
apply (rule hoare_pre)
apply (wp)
apply(case_tac "st \<noteq> IRQTimer")
apply((wp hoare_vcg_imp_lift' | simp | wpc)+)[1]

View File

@ -36,7 +36,7 @@ crunch irq_state_of_state[wp]: set_extra_badge "\<lambda>s. P (irq_state_of_stat
lemma transfer_caps_loop_irq_state[wp]:
"\<lbrace>\<lambda>s. P (irq_state_of_state s)\<rbrace> transfer_caps_loop a b c d e f g \<lbrace>\<lambda>_ s. P (irq_state_of_state s)\<rbrace>"
"\<lbrace>\<lambda>s. P (irq_state_of_state s)\<rbrace> transfer_caps_loop a b c d e f \<lbrace>\<lambda>_ s. P (irq_state_of_state s)\<rbrace>"
apply(wp transfer_caps_loop_pres)
done

View File

@ -2679,7 +2679,7 @@ lemma transfer_caps_silc_inv:
real_cte_at (snd x) s)) and
K (is_subject aag receiver \<and> (\<forall>cap\<in>set caps.
is_subject aag (fst (snd cap))))\<rbrace>
transfer_caps mi caps endpoint receiver receive_buffer diminish
transfer_caps mi caps endpoint receiver receive_buffer
\<lbrace>\<lambda>_. silc_inv aag st\<rbrace>"
apply (rule hoare_gen_asm)
apply (simp add: transfer_caps_def)
@ -2699,7 +2699,7 @@ crunch silc_inv[wp]: copy_mrs, set_message_info "silc_inv aag st"
lemma do_normal_transfer_silc_inv:
"\<lbrace>silc_inv aag st and valid_objs and valid_mdb and pas_refined aag and
K (grant \<longrightarrow> is_subject aag sender \<and> is_subject aag receiver)\<rbrace>
do_normal_transfer sender send_buffer ep badge grant receiver recv_buffer diminish
do_normal_transfer sender send_buffer ep badge grant receiver recv_buffer
\<lbrace>\<lambda>_. silc_inv aag st\<rbrace>"
unfolding do_normal_transfer_def
apply(case_tac grant)
@ -2725,7 +2725,7 @@ lemma valid_ep_recv_dequeue':
lemma do_ipc_transfer_silc_inv:
"\<lbrace>silc_inv aag st and valid_objs and valid_mdb and pas_refined aag and
K (grant \<longrightarrow> is_subject aag sender \<and> is_subject aag receiver)\<rbrace>
do_ipc_transfer sender ep badge grant receiver diminish
do_ipc_transfer sender ep badge grant receiver
\<lbrace>\<lambda>_. silc_inv aag st\<rbrace>"
unfolding do_ipc_transfer_def
apply (wp do_normal_transfer_silc_inv hoare_vcg_all_lift | wpc | wp_once hoare_drop_imps)+
@ -2743,7 +2743,7 @@ lemma send_ipc_silc_inv:
\<lbrace>\<lambda>_. silc_inv aag st\<rbrace>"
unfolding send_ipc_def
apply (wp setup_caller_cap_silc_inv | wpc | simp)+
apply(rename_tac xs word ys recv_state diminish)
apply(rename_tac xs word ys recv_state)
apply(rule_tac Q="\<lambda> r s. (can_grant \<longrightarrow> is_subject aag thread \<and> is_subject aag (hd xs)) \<and> silc_inv aag st s" in hoare_strengthen_post)
apply simp
apply(wp do_ipc_transfer_silc_inv | wpc | simp)+
@ -3153,6 +3153,7 @@ lemma handle_recv_silc_inv:
lemma handle_interrupt_silc_inv:
"\<lbrace>silc_inv aag st\<rbrace> handle_interrupt irq \<lbrace>\<lambda>_. silc_inv aag st\<rbrace>"
unfolding handle_interrupt_def
apply (rule hoare_if)
apply(wp | wpc | simp | wp_once hoare_drop_imps)+
done

View File

@ -1272,7 +1272,7 @@ lemma get_blocking_object_reads_respects:
fun tcb_st_to_auth' where
"tcb_st_to_auth' (BlockedOnSend x xa) = SyncSend" |
"tcb_st_to_auth' (BlockedOnReceive x xa) = Receive" |
"tcb_st_to_auth' (BlockedOnReceive x) = Receive" |
"tcb_st_to_auth' (BlockedOnNotification x) = Receive"
@ -1282,7 +1282,7 @@ lemma owns_thread_blocked_reads_endpoint:
"\<lbrakk>pas_refined aag s; invs s;
st_tcb_at (\<lambda> y. y = state) tptr s;
is_subject aag tptr;
state = (BlockedOnReceive x xa) \<or> state = (BlockedOnSend x xb) \<or> state = BlockedOnNotification x\<rbrakk> \<Longrightarrow> aag_can_read aag x"
state = (BlockedOnReceive x) \<or> state = (BlockedOnSend x xb) \<or> state = BlockedOnNotification x\<rbrakk> \<Longrightarrow> aag_can_read aag x"
apply(rule_tac auth="tcb_st_to_auth' state" in reads_ep)
apply(drule sym, simp, rule pas_refined_mem)
apply(rule_tac s=s in sta_ts)
@ -1992,7 +1992,7 @@ crunch valid_ko_at_arm[wp]: set_mrs "valid_ko_at_arm" (wp: mapM_x_wp' simp: zipW
crunch globals_equiv[wp]: deleted_irq_handler "globals_equiv st"
lemma transfer_caps_valid_ko_at_arm[wp]:
"\<lbrace> valid_ko_at_arm \<rbrace> transfer_caps a b c d e f\<lbrace>\<lambda>_. valid_ko_at_arm\<rbrace>"
"\<lbrace> valid_ko_at_arm \<rbrace> transfer_caps a b c d e \<lbrace>\<lambda>_. valid_ko_at_arm\<rbrace>"
unfolding transfer_caps_def
apply (wp | wpc)+
apply (wp transfer_caps_loop_pres cap_insert_valid_ko_at_arm)

View File

@ -83,14 +83,18 @@ crunch irq_masks[wp]: machine_op_lift "\<lambda>s. P (irq_masks s)"
lemma handle_interrupt_irq_masks:
notes no_irq[wp del]
shows
"\<lbrace>(\<lambda>s. P (irq_masks_of_state s)) and domain_sep_inv False st\<rbrace>
"\<lbrace>(\<lambda>s. P (irq_masks_of_state s)) and domain_sep_inv False st and K (irq \<le> maxIRQ)\<rbrace>
handle_interrupt irq
\<lbrace>\<lambda>rv s. P (irq_masks_of_state s)\<rbrace>"
apply(simp add: handle_interrupt_def)
apply(wp dmo_wp | simp add: ackInterrupt_def maskInterrupt_def when_def
split del: split_if | wpc | simp add: get_irq_state_def |wp_once hoare_drop_imp)+
apply(fastforce simp: domain_sep_inv_def)
apply (rule hoare_gen_asm)
apply(simp add: handle_interrupt_def split del: split_if)
apply (rule hoare_pre)
apply (rule hoare_if)
apply simp
apply(wp dmo_wp | simp add: ackInterrupt_def maskInterrupt_def when_def split del: split_if | wpc | simp add: get_irq_state_def | wp_once hoare_drop_imp)+
apply (fastforce simp: domain_sep_inv_def)
done
crunch irq_masks[wp]: cap_swap_for_delete "\<lambda>s. P (irq_masks_of_state s)"
@ -396,12 +400,27 @@ crunch irq_masks[wp]: handle_vm_fault "\<lambda>s. P (irq_masks_of_state s)"
lemma dmo_getActiveIRQ_irq_masks[wp]:
"\<lbrace>(\<lambda>s. P (irq_masks_of_state s))\<rbrace>
do_machine_op getActiveIRQ
\<lbrace>\<lambda>x s. P (irq_masks_of_state s)\<rbrace>"
\<lbrace>\<lambda>rv s. P (irq_masks_of_state s) \<rbrace>"
apply(rule hoare_pre, rule dmo_wp)
apply(simp add: getActiveIRQ_def | wp | simp add: no_irq_def | clarsimp)+
done
lemma dmo_getActiveIRQ_return_axiom[wp]:
"\<lbrace>\<top>\<rbrace>
do_machine_op getActiveIRQ
\<lbrace>(\<lambda>rv s. (\<forall>x. rv = Some x \<longrightarrow> x \<le> maxIRQ)) \<rbrace>"
apply (simp add: getActiveIRQ_def)
apply(rule hoare_pre, rule dmo_wp)
apply (insert irq_oracle_max_irq)
apply (wp alternative_wp select_wp dmo_getActiveIRQ_irq_masks)
apply clarsimp
done
lemma handle_yield_irq_masks_of_state[wp]: "\<lbrace>(\<lambda>s. P (irq_masks_of_state s)) and domain_sep_inv False st and invs\<rbrace> handle_yield \<lbrace>\<lambda>rv s. P (irq_masks_of_state s)\<rbrace>"
lemma handle_yield_irq_masks_of_state[wp]:
"\<lbrace>(\<lambda>s. P (irq_masks_of_state s)) and domain_sep_inv False st and invs\<rbrace>
handle_yield
\<lbrace>\<lambda>rv s. P (irq_masks_of_state s)\<rbrace>"
apply (simp add: handle_yield_def)
apply wp
apply simp
@ -412,9 +431,20 @@ lemma handle_event_irq_masks:
handle_event ev
\<lbrace> \<lambda> rv s. P (irq_masks_of_state s) \<rbrace>"
apply(case_tac ev)
apply (rename_tac syscall)
apply(case_tac syscall)
apply(simp add: handle_send_def handle_call_def | wp handle_invocation_irq_masks[where st=st] handle_interrupt_irq_masks[where st=st] hoare_vcg_all_lift | wpc | wp_once hoare_drop_imps)+
prefer 4
apply (rule hoare_pre)
apply simp
apply (wp handle_interrupt_irq_masks[where st=st] | wpc | simp )+
apply (rule_tac Q="\<lambda>rv s. P (irq_masks_of_state s) \<and> domain_sep_inv False st s \<and> (\<forall>x. rv = Some x \<longrightarrow> x \<le> maxIRQ)" in hoare_strengthen_post)
apply (wp | clarsimp)+
apply (rename_tac syscall)
apply (case_tac syscall)
apply (simp add: handle_send_def handle_call_def
| wp handle_invocation_irq_masks[where st=st] handle_interrupt_irq_masks[where st=st] hoare_vcg_all_lift
| wpc
| wp_once hoare_drop_imps)+
done
crunch irq_masks[wp]: activate_thread "\<lambda>s. P (irq_masks_of_state s)"
@ -428,8 +458,8 @@ lemma call_kernel_irq_masks:
\<lbrace> \<lambda> rv s. P (irq_masks_of_state s) \<rbrace>"
apply(simp add: call_kernel_def)
apply (wp handle_interrupt_irq_masks[where st=st])+
apply(rule_tac Q="\<lambda>_ s. P (irq_masks_of_state s) \<and> domain_sep_inv False st s" in hoare_strengthen_post)
apply(wp | simp)+
apply (rule_tac Q="\<lambda>rv s. P (irq_masks_of_state s) \<and> domain_sep_inv False st s \<and> (\<forall>x. rv = Some x \<longrightarrow> x \<le> maxIRQ)" in hoare_strengthen_post)
apply (wp | simp)+
apply(rule_tac Q="\<lambda> x s. P (irq_masks_of_state s) \<and> domain_sep_inv False st s" and F="E" for E in hoare_post_impErr)
apply(rule valid_validE)
apply(wp handle_event_irq_masks[where st=st] valid_validE[OF handle_event_domain_sep_inv] | simp)+

View File

@ -509,7 +509,7 @@ lemma cancel_ipc_to_blocked_nosts:
apply (rule monadic_rewrite_bind_tail)
apply (rule monadic_rewrite_transverse)
apply (rename_tac state)
apply (rule_tac P="\<lambda>_. \<exists>xa d. state = BlockedOnReceive xa d" in monadic_rewrite_bind_head)
apply (rule_tac P="\<lambda>_. \<exists>xa. state = BlockedOnReceive xa" in monadic_rewrite_bind_head)
apply (rule monadic_rewrite_gen_asm[where Q=\<top>,simplified])
apply clarsimp
apply (rule monadic_rewrite_refl)
@ -552,7 +552,7 @@ lemma ev2_invisible_simple:
lemma blocked_cancel_ipc_nosts_equiv_but_for_labels:
"\<lbrace>pas_refined aag and
st_tcb_at (\<lambda>st. \<exists>xa. st = BlockedOnReceive x xa) t and
st_tcb_at (\<lambda>st. st = BlockedOnReceive x) t and
bound_tcb_at (op = (Some ntfnptr)) t and
equiv_but_for_labels aag L st and
K(pasObjectAbs aag x \<in> L) and
@ -568,7 +568,7 @@ lemma blocked_cancel_ipc_nosts_equiv_but_for_labels:
lemma blocked_cancel_ipc_nosts_reads_respects:
"reads_respects aag l (pas_refined aag
and st_tcb_at (\<lambda>st. \<exists>xa. st = (BlockedOnReceive x xa)) t
and st_tcb_at (\<lambda>st. \<exists>xa. st = (BlockedOnReceive x)) t
and bound_tcb_at (op = (Some ntfnptr)) t
and (\<lambda>s. is_subject aag (cur_thread s))
@ -652,8 +652,8 @@ crunch pas_refined[wp]: blocked_cancel_ipc_nosts "pas_refined aag"
crunch cur_thread[wp]: blocked_cancel_ipc_nosts "\<lambda>s. P (cur_thread s)"
lemma BlockedOnReceive_inj:
"x = (case (BlockedOnReceive x xa) of BlockedOnReceive x xa \<Rightarrow> x)"
by (cases "BlockedOnReceive x xa";simp)
"x = (case (BlockedOnReceive x) of BlockedOnReceive x \<Rightarrow> x)"
by (cases "BlockedOnReceive x";simp)
lemma send_signal_reads_respects:
@ -718,7 +718,7 @@ lemma send_signal_reads_respects:
apply (clarsimp simp: pred_tcb_at_def get_tcb_def obj_at_def)
apply (rule context_conjI)
apply (fastforce simp: receive_blocked_def intro!: BlockedOnReceive_inj split:thread_state.splits)
apply (frule_tac t=x and tcb=tcb and ep = "case (tcb_state tcb) of BlockedOnReceive a xb \<Rightarrow> a"
apply (frule_tac t=x and tcb=tcb and ep = "case (tcb_state tcb) of BlockedOnReceive a \<Rightarrow> a"
in get_tcb_recv_blocked_implies_receive)
apply (fastforce simp: pred_tcb_at_def get_tcb_def obj_at_def)
apply (fastforce simp: receive_blocked_def split:thread_state.splits)
@ -1098,7 +1098,7 @@ lemma transfer_caps_loop_reads_respects:
K ((\<forall>cap\<in>set caps. is_subject aag (fst (snd cap)) \<and>
pas_cap_cur_auth aag (fst cap)) \<and>
(\<forall>slot\<in>set slots. is_subject aag (fst slot))))
(transfer_caps_loop ep diminish rcv_buffer n caps slots mi)"
(transfer_caps_loop ep rcv_buffer n caps slots mi)"
apply(induct caps arbitrary: slots n mi)
apply simp
apply(rule return_ev_pre)
@ -1409,7 +1409,7 @@ lemma transfer_caps_reads_respects:
ipc_buffer_has_read_auth aag (pasSubject aag) receive_buffer \<and>
(\<forall>cap\<in>set caps.
is_subject aag (fst (snd cap)) \<and> pas_cap_cur_auth aag (fst cap))))
(transfer_caps mi caps endpoint receiver receive_buffer diminish)"
(transfer_caps mi caps endpoint receiver receive_buffer)"
unfolding transfer_caps_def fun_app_def
apply(wp transfer_caps_loop_reads_respects get_receive_slots_rev get_receive_slots_authorised
hoare_vcg_all_lift static_imp_wp
@ -1585,7 +1585,7 @@ lemma do_normal_transfer_reads_respects:
ipc_buffer_has_read_auth aag (pasObjectAbs aag sender) sbuf \<and>
ipc_buffer_has_read_auth aag (pasObjectAbs aag receiver) rbuf \<and>
(grant \<longrightarrow> (is_subject aag sender \<and> is_subject aag receiver))))
(do_normal_transfer sender sbuf endpoint badge grant receiver rbuf diminish)"
(do_normal_transfer sender sbuf endpoint badge grant receiver rbuf)"
apply(case_tac grant)
apply(rule gen_asm_ev)
apply(simp add: do_normal_transfer_def)
@ -1759,7 +1759,7 @@ lemma do_ipc_transfer_reads_respects:
aag_can_read_or_affect aag l sender \<and>
aag_can_read_or_affect aag l receiver
))
(do_ipc_transfer sender ep badge grant receiver diminish)"
(do_ipc_transfer sender ep badge grant receiver)"
unfolding do_ipc_transfer_def
apply (wp do_normal_transfer_reads_respects lookup_ipc_buffer_reads_respects
lookup_ipc_buffer_has_read_auth do_fault_transfer_reads_respects
@ -1955,7 +1955,7 @@ lemma send_ipc_reads_respects:
apply (wp set_endpoint_reads_respects set_thread_state_reads_respects
when_ev setup_caller_cap_reads_respects thread_get_reads_respects
| wpc | simp split del: split_if)+
apply(rename_tac list word list' rvb rvc)
apply(rename_tac list word list' rvb)
apply(rule_tac Q="\<lambda>r s. is_subject aag (cur_thread s) \<and>
(can_grant \<longrightarrow> is_subject aag (hd list))"
in hoare_strengthen_post)
@ -1969,7 +1969,6 @@ lemma send_ipc_reads_respects:
do_ipc_transfer_pas_refined
| wpc
| simp add: get_thread_state_def thread_get_def)+
apply (clarsimp simp: conj_comms)
apply (rule conjI)
apply(fastforce dest: reads_ep)
apply clarsimp
@ -2171,7 +2170,7 @@ lemma set_extra_badge_globals_equiv:
lemma transfer_caps_loop_globals_equiv:
"\<lbrace>globals_equiv st and valid_ko_at_arm and valid_global_objs and (\<lambda>sa. \<forall>x<length caps. ptr_range (rcv_buffer + (of_nat buffer_cptr_index + of_nat (x + n)) * of_nat word_size) 2 \<inter> range_of_arm_globals_frame sa = {})\<rbrace>
transfer_caps_loop ep diminish rcv_buffer n caps slots mi
transfer_caps_loop ep rcv_buffer n caps slots mi
\<lbrace>\<lambda>_. globals_equiv st\<rbrace>"
proof (induct caps arbitrary: slots n mi)
case Nil
@ -2197,8 +2196,7 @@ next
apply(simp add: whenE_def, rule conjI)
apply(rule impI, wp)+
apply(simp)+
apply(rule conjI)
apply(rule impI, wp)+
apply wp
apply(rule conjI)
apply(clarsimp)
apply(rule conjI)
@ -2219,7 +2217,7 @@ qed
lemma transfer_caps_globals_equiv:
"\<lbrace>globals_equiv st and valid_ko_at_arm and valid_global_objs and (\<lambda>sa. \<forall>rb. recv_buffer = Some rb \<longrightarrow> (\<forall>x<length caps.
ptr_range (rb + (of_nat buffer_cptr_index + of_nat x) * of_nat word_size) 2 \<inter> range_of_arm_globals_frame sa = {}))\<rbrace>
transfer_caps info caps endpoint receiver recv_buffer diminish
transfer_caps info caps endpoint receiver recv_buffer
\<lbrace>\<lambda>_. globals_equiv st\<rbrace>"
unfolding transfer_caps_def
apply(wp transfer_caps_loop_globals_equiv | wpc | simp)+
@ -2260,7 +2258,7 @@ lemma do_normal_transfer_globals_equiv:
(x\<in>set [length msg_registers + 1..< (2 ^ (msg_align_bits - 2))]) \<longrightarrow>
ptr_range (rb + of_nat x * of_nat word_size) 2 \<inter>
range_of_arm_globals_frame sa = {}))\<rbrace>
do_normal_transfer sender sbuf endpoint badge grant receiver rbuf diminish
do_normal_transfer sender sbuf endpoint badge grant receiver rbuf
\<lbrace>\<lambda>_. globals_equiv st\<rbrace>"
unfolding do_normal_transfer_def
apply(wp as_user_globals_equiv set_message_info_globals_equiv transfer_caps_globals_equiv)
@ -2388,7 +2386,7 @@ lemma auth_ipc_buffers_do_not_overlap_arm_globals_frame:
lemma do_ipc_transfer_globals_equiv:
"\<lbrace>globals_equiv st and valid_ko_at_arm and valid_objs and valid_arch_state and valid_global_refs and pspace_distinct and valid_global_objs and (\<lambda>s. receiver \<noteq> idle_thread s)\<rbrace>
do_ipc_transfer sender ep badge grant receiver diminish
do_ipc_transfer sender ep badge grant receiver
\<lbrace>\<lambda>_. globals_equiv st\<rbrace>"
unfolding do_ipc_transfer_def
apply(wp do_normal_transfer_globals_equiv do_fault_transfer_globals_equiv | wpc)+
@ -2467,9 +2465,8 @@ lemma send_ipc_globals_equiv:
apply(rule thread_get_inv)
apply(fastforce)
apply(wp set_thread_state_globals_equiv dxo_wp_weak | simp)+
apply wpc
apply(wp do_ipc_transfer_globals_equiv)
apply(wpc)
apply(rule fail_wp | rule return_wp)+
apply(clarsimp)
apply(rule hoare_drop_imps)
apply(wp set_endpoint_globals_equiv)
@ -2668,7 +2665,7 @@ lemma send_ipc_valid_global_objs:
apply(rule_tac Q="\<lambda>_. valid_global_objs" in hoare_strengthen_post)
apply(wp, simp, (wp dxo_wp_weak |simp)+)
apply(wpc)
apply(rule fail_wp | rule return_wp)+
apply(rule fail_wp | rule return_wp | wp)+
apply(simp)
apply(rule hoare_drop_imps)
apply(wp)

View File

@ -3225,7 +3225,18 @@ lemma confidentiality_part_not_PSched:
apply(fastforce dest: non_PSched_steps_run_in_lock_step)
done
lemma getActiveIRQ_ret_no_dmo[wp]: "\<lbrace>\<lambda>_. True\<rbrace> getActiveIRQ \<lbrace>\<lambda>rv s. \<forall>x. rv = Some x \<longrightarrow> x \<le> maxIRQ\<rbrace>"
apply (simp add: getActiveIRQ_def)
apply(rule hoare_pre)
apply (insert irq_oracle_max_irq)
apply (wp alternative_wp select_wp dmo_getActiveIRQ_irq_masks)
apply clarsimp
done
lemma try_some_magic: "(\<forall>x. y = Some x \<longrightarrow> P x) = ((\<exists>x. y = Some x) \<longrightarrow> P (the y))"
by auto
lemma preemption_interrupt_scheduler_invisible:
"equiv_valid_2 (scheduler_equiv aag) (scheduler_affects_equiv aag l) (scheduler_affects_equiv aag l) (\<lambda>r r'. r = uc \<and> snd r' = uc')
(einvs and pas_refined aag and guarded_pas_domain aag and domain_sep_inv False st and silc_inv aag st' and (\<lambda>s. irq_masks_of_state st = irq_masks_of_state s) and (\<lambda>s. ct_idle s \<longrightarrow> uc = idle_context s)
@ -3237,22 +3248,19 @@ lemma preemption_interrupt_scheduler_invisible:
apply (rule equiv_valid_2_bind_right)
apply (rule equiv_valid_2_bind_right)
apply (simp add: liftE_def bind_assoc)
apply (simp only: option.case_eq_if)
apply (rule equiv_valid_2_bind_pre[where R'="op ="])
apply simp
apply (case_tac "rv'")
prefer 2
apply simp
apply (rule equiv_valid_2_bind_pre[where R'="op =" and Q="\<top>\<top>" and Q'="\<top>\<top>"])
apply (rule return_ev2)
apply simp
apply (rule equiv_valid_2)
apply (rule handle_interrupt_reads_respects_scheduler)
apply (wp | simp)+
apply (rule return_ev2)
apply simp
apply (simp add: when_def split del: split_if)
apply (subst if_swap)
apply (simp split del: split_if)
apply (rule equiv_valid_2_bind_pre[where R'="op =" and Q="\<top>\<top>" and Q'="\<top>\<top>"])
apply (rule return_ev2)
apply simp
apply (rule equiv_valid_2)
apply (wp handle_interrupt_reads_respects_scheduler[where st=st] | simp)+
apply (rule equiv_valid_2)
apply (rule dmo_getActive_IRQ_reads_respect_scheduler)
apply (wp | simp | elim conjE | intro conjI)+
apply (wp dmo_getActiveIRQ_return_axiom[simplified try_some_magic] | simp add: imp_conjR | elim conjE | intro conjI | wp_once hoare_drop_imps)+
apply (subst thread_set_as_user)
apply (wp guarded_pas_domain_lift)
apply ((simp | wp | force)+)[7]
@ -3275,8 +3283,8 @@ lemma handle_preemption_agnostic_ret: "\<lbrace>\<top>\<rbrace> handle_preemptio
lemma handle_preemption_reads_respects_scheduler:
"reads_respects_scheduler aag l (einvs and pas_refined aag and guarded_pas_domain aag and domain_sep_inv False st and silc_inv aag st' and (\<lambda>s. irq_masks_of_state st = irq_masks_of_state s)) (handle_preemption_if uc)"
apply (simp add: handle_preemption_if_def)
apply (wp when_ev handle_interrupt_reads_respects_scheduler
dmo_getActive_IRQ_reads_respect_scheduler | simp | wp_once hoare_drop_imps)+
apply (wp when_ev handle_interrupt_reads_respects_scheduler dmo_getActiveIRQ_return_axiom[simplified try_some_magic]
dmo_getActive_IRQ_reads_respect_scheduler | simp add: imp_conjR| wp_once hoare_drop_imps)+
apply force
done
@ -3298,8 +3306,8 @@ lemma kernel_entry_scheduler_equiv_2:
apply (rule equiv_valid_2_bind_pre[where R'="op ="])
apply (rule equiv_valid_2)
apply simp
apply (wp del: no_irq add: handle_interrupt_reads_respects_scheduler
dmo_getActive_IRQ_reads_respect_scheduler hoare_vcg_all_lift | wpc | simp | wp_once hoare_drop_imps)+
apply (wp del: no_irq add: handle_interrupt_reads_respects_scheduler[where st=st]
dmo_getActive_IRQ_reads_respect_scheduler | wpc | simp add: imp_conjR all_conj_distrib | wp_once hoare_drop_imps)+
apply (rule context_update_cur_thread_snippit)
apply (wp thread_set_invs_trivial guarded_pas_domain_lift
thread_set_pas_refined thread_set_not_state_valid_sched | simp add: tcb_cap_cases_def)+
@ -3315,15 +3323,15 @@ lemma kernel_entry_if_reads_respects_scheduler:
apply (rule bind_ev_pre)
apply wp_once
apply (rule bind_ev_pre)
apply ((wp del: no_irq add: when_ev handle_interrupt_reads_respects_scheduler dmo_getActive_IRQ_reads_respect_scheduler hoare_vcg_all_lift liftE_ev | simp | wpc | wp_once hoare_drop_imps)+)[1]
apply ((wp del: no_irq add: when_ev handle_interrupt_reads_respects_scheduler[where st=st] dmo_getActive_IRQ_reads_respect_scheduler liftE_ev | simp add: imp_conjR all_conj_distrib | wpc | wp_once hoare_drop_imps)+)[1]
apply (rule reads_respects_scheduler_cases')
prefer 3
apply (rule reads_respects_scheduler_unobservable'')
apply ((wp thread_set_scheduler_equiv | simp | elim conjE)+)[3]
apply ((wp | simp | elim conjE)+)[2]
apply (clarsimp simp: guarded_pas_domain_def)
apply (wp thread_set_invs_trivial guarded_pas_domain_lift
thread_set_pas_refined thread_set_not_state_valid_sched | simp add: tcb_cap_cases_def)+
apply ((wp thread_set_invs_trivial guarded_pas_domain_lift hoare_vcg_all_lift
thread_set_pas_refined thread_set_not_state_valid_sched | simp add: tcb_cap_cases_def)+)
apply (clarsimp simp: cur_thread_idle cur_thread_not_SilcLabel)
apply force
done

View File

@ -2167,16 +2167,22 @@ lemma ackInterrupt_reads_respects_scheduler:
done
lemma handle_interrupt_reads_respects_scheduler:
"reads_respects_scheduler aag l (invs and guarded_pas_domain aag and pas_refined aag and valid_sched and domain_sep_inv False st) (handle_interrupt irq)"
apply (simp add: handle_interrupt_def)
apply (wp | rule ackInterrupt_reads_respects_scheduler)+
apply (rule_tac Q="rv = IRQTimer \<or> rv = IRQInactive" in gen_asm_ev(2))
apply (elim disjE)
apply (wp timer_tick_reads_respects_scheduler
"reads_respects_scheduler aag l (invs and guarded_pas_domain aag and pas_refined aag and valid_sched and domain_sep_inv False st and K (irq \<le> maxIRQ)) (handle_interrupt irq)"
apply (simp add: handle_interrupt_def )
apply (rule conjI; rule impI )
apply (rule gen_asm_ev)
apply simp
apply (wp modify_wp | simp )+
apply (rule ackInterrupt_reads_respects_scheduler)
apply (rule_tac Q="rv = IRQTimer \<or> rv = IRQInactive" in gen_asm_ev(2))
apply (elim disjE)
apply (wp timer_tick_reads_respects_scheduler
ackInterrupt_reads_respects_scheduler
dmo_resetTimer_reads_respects_scheduler
get_irq_state_reads_respects_scheduler_trivial
fail_ev irq_inactive_or_timer | simp add: | wpc)+
fail_ev irq_inactive_or_timer | simp )+
apply force
done

View File

@ -907,13 +907,20 @@ declare gts_st_tcb_at[wp del]
lemma handle_interrupt_globals_equiv:
"\<lbrace>globals_equiv (st :: det_ext state) and invs\<rbrace> handle_interrupt irq \<lbrace>\<lambda>r. globals_equiv st\<rbrace>"
unfolding handle_interrupt_def
apply (wp dmo_maskInterrupt_globals_equiv
apply (rule hoare_if)
apply (wp dmo_maskInterrupt_globals_equiv
dmo_return_globals_equiv
send_signal_globals_equiv
VSpace_AI.dmo_ackInterrupt
hoare_vcg_if_lift2
hoare_drop_imps
dxo_wp_weak
| wpc | simp add: ackInterrupt_def resetTimer_def invs_imps invs_valid_idle)+
Retype_IF.dmo_mol_globals_equiv
NonDetMonadLemmaBucket.no_fail_bind
NonDetMonadLemmaBucket.bind_known_operation_eq
Retype_IF.dmo_mol_globals_equiv
| wpc | simp add: Interrupt_AI.empty_fail_ackInterrupt dmo_bind_valid ackInterrupt_def resetTimer_def invs_imps invs_valid_idle)+
done

View File

@ -405,8 +405,9 @@ lemma get_receive_slots_bcorres[wp]: "bcorres (get_receive_slots a b) (get_recei
crunch (bcorres)bcorres[wp]: set_extra_badge,derive_cap truncate_state (ignore: storeWord)
lemma transfer_caps_loop_bcorres[wp]: "bcorres (transfer_caps_loop ep diminish buffer n caps slots mi) (transfer_caps_loop ep diminish buffer n caps slots mi)"
apply (induct caps arbitrary: slots n mi ep diminish)
lemma transfer_caps_loop_bcorres[wp]:
"bcorres (transfer_caps_loop ep buffer n caps slots mi) (transfer_caps_loop ep buffer n caps slots mi)"
apply (induct caps arbitrary: slots n mi ep)
apply simp
apply wp
apply (case_tac a)

View File

@ -322,7 +322,7 @@ lemma valid_tcb_state_update:
"\<lbrakk> valid_tcb p t s; valid_tcb_state st s;
case st of
Structures_A.Inactive \<Rightarrow> True
| Structures_A.BlockedOnReceive e d \<Rightarrow>
| Structures_A.BlockedOnReceive e \<Rightarrow>
tcb_caller t = cap.NullCap
\<and> is_master_reply_cap (tcb_reply t)
\<and> obj_ref_of (tcb_reply t) = p

View File

@ -1755,7 +1755,7 @@ crunch not_cur_thread[wp]: cap_insert, set_extra_badge "not_cur_thread t"
(wp: hoare_drop_imps)
lemma transfer_caps_not_cur_thread[wp]:
"\<lbrace>not_cur_thread t\<rbrace> transfer_caps info caps ep recv recv_buf dim
"\<lbrace>not_cur_thread t\<rbrace> transfer_caps info caps ep recv recv_buf
\<lbrace>\<lambda>rv. not_cur_thread t\<rbrace>"
by (simp add: transfer_caps_def | wp transfer_caps_loop_pres | wpc)+
@ -2083,7 +2083,7 @@ lemma invoke_cnode_valid_sched:
crunch valid_sched[wp]: set_extra_badge valid_sched (wp: crunch_wps)
lemma transfer_caps_valid_sched:
"\<lbrace>valid_sched\<rbrace> transfer_caps info caps ep recv recv_buf dim \<lbrace>\<lambda>rv. valid_sched\<rbrace>"
"\<lbrace>valid_sched\<rbrace> transfer_caps info caps ep recv recv_buf \<lbrace>\<lambda>rv. valid_sched\<rbrace>"
apply (simp add: transfer_caps_def | wp transfer_caps_loop_pres | wpc)+
done
@ -2235,7 +2235,8 @@ lemma send_ipc_valid_sched:
"\<lbrace>valid_sched and st_tcb_at active thread and scheduler_act_not thread and not_queued thread and invs\<rbrace>
send_ipc block call badge can_grant thread epptr \<lbrace>\<lambda>rv. valid_sched\<rbrace>"
apply (simp add: send_ipc_def)
apply (wp set_thread_state_sched_act_not_valid_sched | wpc)+
apply (wp set_thread_state_sched_act_not_valid_sched | wpc )+
apply ((wp set_thread_state_sched_act_not_valid_sched
setup_caller_cap_sched_act_not_valid_sched
attempt_switch_to_valid_sched'
@ -2245,9 +2246,8 @@ lemma send_ipc_valid_sched:
set_thread_state_valid_blocked_except sts_st_tcb_at')[1]
apply simp
apply (rule_tac Q="\<lambda>_. valid_sched and scheduler_act_not thread and not_queued thread and (\<lambda>s. x21 \<noteq> idle_thread s \<and> x21 \<noteq> thread)" in hoare_strengthen_post)
apply wp
apply ((wp|wpc)+)[1]
apply (clarsimp simp: valid_sched_def)
apply ((wp | wpc)+)[1]
apply (simp | wp gts_wp hoare_vcg_all_lift)+
apply (wp hoare_vcg_imp_lift)
apply ((simp add: set_endpoint_def set_object_def | wp hoare_drop_imps | wpc)+)[1]
@ -2857,7 +2857,7 @@ crunch ct_not_queued[wp]: do_machine_op, cap_insert, set_extra_badge "\<lambda>s
lemma transfer_caps_ct_not_queued[wp]:
"\<lbrace>\<lambda>s. not_queued (cur_thread s) s\<rbrace>
transfer_caps info caps ep recv recv_buf dim
transfer_caps info caps ep recv recv_buf
\<lbrace>\<lambda>rv s. not_queued (cur_thread s) s\<rbrace>"
by (simp add: transfer_caps_def | wp transfer_caps_loop_pres | wpc)+

View File

@ -4117,10 +4117,10 @@ interpretation timer_tick_extended: is_extended "timer_tick"
done
lemma handle_interrupt_valid_list[wp]:
"\<lbrace>valid_list\<rbrace> handle_interrupt irq \<lbrace>\<lambda>_.valid_list\<rbrace>"
unfolding handle_interrupt_def
apply (wp get_cap_wp | wpc | simp add: get_irq_slot_def get_irq_state_def)+
done
"\<lbrace>valid_list \<rbrace> handle_interrupt irq \<lbrace>\<lambda>_.valid_list\<rbrace>"
unfolding handle_interrupt_def ackInterrupt_def
apply (rule hoare_pre)
by (wp get_cap_wp do_machine_op_valid_list | wpc | simp add: get_irq_slot_def | wp_once hoare_drop_imps)+
crunch valid_list[wp]: handle_send,handle_reply valid_list

View File

@ -203,13 +203,15 @@ lemma get_irq_slot_ex_cte:
apply clarsimp
done
lemma maskInterrupt_invs:
"\<lbrace>invs and (\<lambda>s. interrupt_states s irq \<noteq> IRQInactive)\<rbrace>
"\<lbrace>invs and (\<lambda>s. \<not>b \<longrightarrow> interrupt_states s irq \<noteq> IRQInactive)\<rbrace>
do_machine_op (maskInterrupt b irq)
\<lbrace>\<lambda>rv. invs\<rbrace>"
apply (simp add: do_machine_op_def split_def maskInterrupt_def)
apply wp
apply (clarsimp simp: in_monad invs_def valid_state_def all_invs_but_valid_irq_states_for_def valid_irq_states_but_def valid_irq_masks_but_def valid_machine_state_def cur_tcb_def valid_irq_states_def valid_irq_masks_def)
apply (clarsimp simp: in_monad invs_def valid_state_def all_invs_but_valid_irq_states_for_def
valid_irq_states_but_def valid_irq_masks_but_def valid_machine_state_def cur_tcb_def valid_irq_states_def valid_irq_masks_def)
done
crunch pspace_aligned[wp]: set_irq_state "pspace_aligned"
@ -370,16 +372,23 @@ lemma send_signal_interrupt_states[wp_unsafe]:
apply (auto simp: pred_tcb_at_def obj_at_def receive_blocked_def)
done
lemma handle_interrupt_invs[wp]:
lemma empty_fail_ackInterrupt[simp, intro!]: "empty_fail (ackInterrupt irq)"
by (wp | simp add: ackInterrupt_def)+
lemma empty_fail_maskInterrupt[simp, intro!]: "empty_fail (maskInterrupt f irq)"
by (wp | simp add: maskInterrupt_def)+
lemma handle_interrupt_invs[wp]:
"\<lbrace>invs\<rbrace> handle_interrupt irq \<lbrace>\<lambda>_. invs\<rbrace>"
apply (simp add: handle_interrupt_def)
apply (wp maskInterrupt_invs | wpc)+
apply (simp add: handle_interrupt_def )
apply (rule conjI; rule impI)
apply (simp add: do_machine_op_bind)
apply (wp dmo_maskInterrupt_invs maskInterrupt_invs dmo_ackInterrupt | wpc | simp)+
apply (wp get_cap_wp send_signal_interrupt_states )
apply (rule_tac Q="\<lambda>rv. invs and (\<lambda>s. st = interrupt_states s irq)" in hoare_post_imp)
apply (clarsimp simp: ex_nonz_cap_to_def invs_valid_objs)
apply (intro allI exI, erule cte_wp_at_weakenE)
apply (clarsimp simp: is_cap_simps)
apply (wp hoare_drop_imps | simp add: get_irq_state_def)+
done
end
done
end

View File

@ -407,7 +407,7 @@ definition
valid_tcb_state :: "Structures_A.thread_state \<Rightarrow> 'z::state_ext state \<Rightarrow> bool"
where
"valid_tcb_state ts s \<equiv> case ts of
Structures_A.BlockedOnReceive ref d \<Rightarrow> ep_at ref s
Structures_A.BlockedOnReceive ref \<Rightarrow> ep_at ref s
| Structures_A.BlockedOnSend ref sp \<Rightarrow> ep_at ref s
| Structures_A.BlockedOnNotification ref \<Rightarrow> ntfn_at ref s
| _ \<Rightarrow> True"
@ -435,7 +435,7 @@ where
\<or> (halted st \<and> (c = cap.NullCap)))),
tcb_cnode_index 3 \<mapsto> (tcb_caller, tcb_caller_update,
(\<lambda>_ st. case st of
Structures_A.BlockedOnReceive e d \<Rightarrow>
Structures_A.BlockedOnReceive e \<Rightarrow>
(op = cap.NullCap)
| _ \<Rightarrow> is_reply_cap or (op = cap.NullCap))),
tcb_cnode_index 4 \<mapsto> (tcb_ipcframe, tcb_ipcframe_update,
@ -942,7 +942,7 @@ where
| (Structures_A.Restart) => {}
| (Structures_A.BlockedOnReply) => {}
| (Structures_A.IdleThreadState) => {}
| (Structures_A.BlockedOnReceive x b) => {(x, TCBBlockedRecv)}
| (Structures_A.BlockedOnReceive x) => {(x, TCBBlockedRecv)}
| (Structures_A.BlockedOnSend x payl) => {(x, TCBBlockedSend)}
| (Structures_A.BlockedOnNotification x) => {(x, TCBSignal)}"
@ -1784,7 +1784,7 @@ lemma tcb_cap_cases_simps[simp]:
"tcb_cap_cases (tcb_cnode_index 3) =
Some (tcb_caller, tcb_caller_update,
(\<lambda>_ st. case st of
Structures_A.BlockedOnReceive e d \<Rightarrow> (op = cap.NullCap)
Structures_A.BlockedOnReceive e \<Rightarrow> (op = cap.NullCap)
| _ \<Rightarrow> is_reply_cap or (op = cap.NullCap)))"
"tcb_cap_cases (tcb_cnode_index 4) =
Some (tcb_ipcframe, tcb_ipcframe_update,
@ -1799,7 +1799,7 @@ lemma ran_tcb_cap_cases:
(is_master_reply_cap c \<and> obj_ref_of c = t)
\<or> (halted st \<and> (c = cap.NullCap)))),
(tcb_caller, tcb_caller_update, (\<lambda>_ st. case st of
Structures_A.BlockedOnReceive e d \<Rightarrow>
Structures_A.BlockedOnReceive e \<Rightarrow>
(op = cap.NullCap)
| _ \<Rightarrow> is_reply_cap or (op = cap.NullCap))),
(tcb_ipcframe, tcb_ipcframe_update, (\<lambda>_ _. is_arch_cap or (op = cap.NullCap)))}"
@ -1985,7 +1985,7 @@ lemma tcb_st_refs_of_simps[simp]:
"tcb_st_refs_of (Structures_A.Restart) = {}"
"tcb_st_refs_of (Structures_A.BlockedOnReply) = {}"
"tcb_st_refs_of (Structures_A.IdleThreadState) = {}"
"\<And>x. tcb_st_refs_of (Structures_A.BlockedOnReceive x b) = {(x, TCBBlockedRecv)}"
"\<And>x. tcb_st_refs_of (Structures_A.BlockedOnReceive x) = {(x, TCBBlockedRecv)}"
"\<And>x. tcb_st_refs_of (Structures_A.BlockedOnSend x payl) = {(x, TCBBlockedSend)}"
"\<And>x. tcb_st_refs_of (Structures_A.BlockedOnNotification x) = {(x, TCBSignal)}"
by (auto simp: tcb_st_refs_of_def)
@ -2027,7 +2027,7 @@ lemma refs_of_simps[simp]:
lemma refs_of_rev:
"(x, TCBBlockedRecv) \<in> refs_of ko =
(\<exists>tcb. ko = TCB tcb \<and> (\<exists>b. tcb_state tcb = Structures_A.BlockedOnReceive x b))"
(\<exists>tcb. ko = TCB tcb \<and> (tcb_state tcb = Structures_A.BlockedOnReceive x))"
"(x, TCBBlockedSend) \<in> refs_of ko =
(\<exists>tcb. ko = TCB tcb \<and> (\<exists>pl. tcb_state tcb = Structures_A.BlockedOnSend x pl))"
"(x, TCBSignal) \<in> refs_of ko =
@ -2056,7 +2056,7 @@ lemma refs_of_rev:
lemma st_tcb_at_refs_of_rev:
"obj_at (\<lambda>ko. (x, TCBBlockedRecv) \<in> refs_of ko) t s
= st_tcb_at (\<lambda>ts. \<exists>b. ts = Structures_A.BlockedOnReceive x b ) t s"
= st_tcb_at (\<lambda>ts. ts = Structures_A.BlockedOnReceive x) t s"
"obj_at (\<lambda>ko. (x, TCBBlockedSend) \<in> refs_of ko) t s
= st_tcb_at (\<lambda>ts. \<exists>pl. ts = Structures_A.BlockedOnSend x pl ) t s"
"obj_at (\<lambda>ko. (x, TCBSignal) \<in> refs_of ko) t s
@ -5114,7 +5114,7 @@ locale invs_locale =
assumes sts_ex_inv[wp]: "\<And>a b. \<lbrace>ex_inv\<rbrace> set_thread_state a b \<lbrace>\<lambda>_.ex_inv\<rbrace>"
assumes setup_caller_cap_ex_inv[wp]: "\<And>send receive. \<lbrace>ex_inv and valid_mdb\<rbrace> setup_caller_cap send receive \<lbrace>\<lambda>_.ex_inv\<rbrace>"
assumes do_ipc_transfer_ex_inv[wp]: "\<And>a b c d e f. \<lbrace>ex_inv and valid_objs and valid_mdb\<rbrace> do_ipc_transfer a b c d e f \<lbrace>\<lambda>_.ex_inv\<rbrace>"
assumes do_ipc_transfer_ex_inv[wp]: "\<And>a b c d e. \<lbrace>ex_inv and valid_objs and valid_mdb\<rbrace> do_ipc_transfer a b c d e \<lbrace>\<lambda>_.ex_inv\<rbrace>"
assumes thread_set_ex_inv[wp]: "\<And>a b. \<lbrace>ex_inv\<rbrace> thread_set a b \<lbrace>\<lambda>_.ex_inv\<rbrace>"

View File

@ -219,7 +219,7 @@ lemma cancel_ipc_tcb [wp]:
lemma gbep_ret:
"\<lbrakk> st = Structures_A.BlockedOnReceive epPtr x \<or>
"\<lbrakk> st = Structures_A.BlockedOnReceive epPtr \<or>
st = Structures_A.BlockedOnSend epPtr p \<rbrakk> \<Longrightarrow>
get_blocking_object st = return epPtr"
by (auto simp add: get_blocking_object_def)
@ -301,7 +301,7 @@ lemma ep_redux_simps2:
lemma gbi_ep_sp:
"\<lbrace>P\<rbrace>
get_blocking_object st
\<lbrace>\<lambda>ep. P and K ((\<exists>d. st = Structures_A.BlockedOnReceive ep d)
\<lbrace>\<lambda>ep. P and K ((st = Structures_A.BlockedOnReceive ep)
\<or> (\<exists>d. st = Structures_A.BlockedOnSend ep d))\<rbrace>"
apply (cases st, simp_all add: get_blocking_object_def)
apply (wp | simp)+

View File

@ -605,10 +605,9 @@ lemma masked_as_full_null_cap[simp]:
"(cap.NullCap = masked_as_full x x) = (x = cap.NullCap)"
by (case_tac x,simp_all add:masked_as_full_def)+
lemma transfer_caps_loop_mi_label[wp]:
"\<lbrace>\<lambda>s. P (mi_label mi)\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>mi' s. P (mi_label mi')\<rbrace>"
apply (induct caps arbitrary: n slots mi)
apply simp
@ -672,7 +671,7 @@ lemma transfer_caps_loop_presM:
cte_wp_at (\<lambda>cp. fst x \<noteq> cap.NullCap \<longrightarrow> cp \<noteq> fst x \<longrightarrow> cp = masked_as_full (fst x) (fst x)) (snd x) s
\<and> real_cte_at (snd x) s))
\<and> (ex \<longrightarrow> (\<forall>x \<in> set slots. ex_cte_cap_wp_to is_cnode_cap x s))\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv. P\<rbrace>"
apply (induct caps arbitrary: slots n mi)
apply (simp, wp, simp)
@ -732,19 +731,19 @@ lemmas transfer_caps_loop_pres =
lemma transfer_caps_loop_typ_at[wp]:
"\<lbrace>\<lambda>s. P (typ_at T p s)\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv s. P (typ_at T p s)\<rbrace>"
by (wp transfer_caps_loop_pres)
lemma transfer_loop_aligned[wp]:
"\<lbrace>pspace_aligned\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv. pspace_aligned\<rbrace>"
by (wp transfer_caps_loop_pres)
lemma transfer_loop_distinct[wp]:
"\<lbrace>pspace_distinct\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv. pspace_distinct\<rbrace>"
by (wp transfer_caps_loop_pres)
@ -755,7 +754,7 @@ lemma invs_valid_objs2:
lemma transfer_caps_loop_valid_objs[wp]:
"\<lbrace>valid_objs and valid_mdb and (\<lambda>s. \<forall>slot \<in> set slots. real_cte_at slot s \<and> cte_wp_at (\<lambda>cap. cap = cap.NullCap) slot s)
and transfer_caps_srcs caps and K (distinct slots)\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv. valid_objs\<rbrace>"
apply (rule hoare_pre)
apply (rule transfer_caps_loop_presM[where vo=True and em=False and ex=False])
@ -771,7 +770,7 @@ lemma transfer_caps_loop_valid_mdb[wp]:
"\<lbrace>\<lambda>s. valid_mdb s \<and> valid_objs s \<and> pspace_aligned s \<and> pspace_distinct s
\<and> (\<forall>slot \<in> set slots. real_cte_at slot s \<and> cte_wp_at (\<lambda>cap. cap = cap.NullCap) slot s)
\<and> transfer_caps_srcs caps s \<and> distinct slots\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv. valid_mdb\<rbrace>"
apply (rule hoare_pre)
apply (rule transfer_caps_loop_presM[where vo=True and em=True and ex=False])
@ -790,7 +789,7 @@ crunch state_refs_of [wp]: set_extra_badge "\<lambda>s. P (state_refs_of s)"
lemma tcl_state_refs_of[wp]:
"\<lbrace>\<lambda>s. P (state_refs_of s)\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv s. P (state_refs_of s)\<rbrace>"
by (wp transfer_caps_loop_pres)
@ -798,7 +797,7 @@ crunch if_live [wp]: set_extra_badge if_live_then_nonz_cap
lemma tcl_iflive[wp]:
"\<lbrace>if_live_then_nonz_cap\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv. if_live_then_nonz_cap\<rbrace>"
by (wp transfer_caps_loop_pres cap_insert_iflive)
@ -806,7 +805,7 @@ crunch if_unsafe [wp]: set_extra_badge if_unsafe_then_cap
lemma tcl_ifunsafe[wp]:
"\<lbrace>\<lambda>s. if_unsafe_then_cap s \<and> (\<forall>x\<in>set slots. ex_cte_cap_wp_to is_cnode_cap x s)\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv. if_unsafe_then_cap\<rbrace>"
by (wp transfer_caps_loop_presM[where vo=False and em=False and ex=True, simplified]
cap_insert_ifunsafe | simp)+
@ -841,20 +840,20 @@ crunch pred_tcb_at [wp]: set_extra_badge "\<lambda>s. pred_tcb_at proj P p s"
crunch idle [wp]: set_extra_badge "\<lambda>s. P (idle_thread s)"
lemma tcl_idle[wp]:
"\<lbrace>valid_idle\<rbrace> transfer_caps_loop ep diminish buffer n caps slots mi \<lbrace>\<lambda>_. valid_idle\<rbrace>"
"\<lbrace>valid_idle\<rbrace> transfer_caps_loop ep buffer n caps slots mi \<lbrace>\<lambda>_. valid_idle\<rbrace>"
by (wp transfer_caps_loop_pres cap_insert_idle valid_idle_lift)
crunch cur_tcb [wp]: set_extra_badge cur_tcb
lemma tcl_ct[wp]:
"\<lbrace>cur_tcb\<rbrace> transfer_caps_loop ep diminish buffer n caps slots mi \<lbrace>\<lambda>rv. cur_tcb\<rbrace>"
"\<lbrace>cur_tcb\<rbrace> transfer_caps_loop ep buffer n caps slots mi \<lbrace>\<lambda>rv. cur_tcb\<rbrace>"
by (wp transfer_caps_loop_pres)
crunch it[wp]: cap_insert "\<lambda>s. P (idle_thread s)"
(wp: crunch_wps simp: crunch_simps)
lemma tcl_it[wp]:
"\<lbrace>\<lambda>s. P (idle_thread s)\<rbrace> transfer_caps_loop ep diminish buffer n caps slots mi
"\<lbrace>\<lambda>s. P (idle_thread s)\<rbrace> transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv s. P (idle_thread s)\<rbrace>"
by (wp transfer_caps_loop_pres)
@ -905,7 +904,7 @@ lemma tcl_zombies[wp]:
"\<lbrace>zombies_final and valid_objs and valid_mdb and K (distinct slots)
and (\<lambda>s. \<forall>slot \<in> set slots. real_cte_at slot s \<and> cte_wp_at (\<lambda>cap. cap = cap.NullCap) slot s )
and transfer_caps_srcs caps\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv. zombies_final\<rbrace>"
apply (rule hoare_pre)
apply (rule transfer_caps_loop_presM[where vo=True and em=False and ex=False])
@ -948,7 +947,7 @@ lemma transfer_caps_loop_valid_globals [wp]:
"\<lbrace>valid_global_refs and valid_objs and valid_mdb and K (distinct slots)
and (\<lambda>s. \<forall>slot \<in> set slots. real_cte_at slot s \<and> cte_wp_at (\<lambda>cap. cap = cap.NullCap) slot s)
and transfer_caps_srcs caps\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv. valid_global_refs\<rbrace>"
apply (rule hoare_pre)
apply (rule transfer_caps_loop_presM[where em=False and ex=False and vo=True])
@ -966,12 +965,12 @@ lemma transfer_caps_loop_valid_globals [wp]:
lemma transfer_caps_loop_arch[wp]:
"\<lbrace>\<lambda>s. P (arch_state s)\<rbrace> transfer_caps_loop ep diminish buffer n caps slots mi \<lbrace>\<lambda>rv s. P (arch_state s)\<rbrace>"
"\<lbrace>\<lambda>s. P (arch_state s)\<rbrace> transfer_caps_loop ep buffer n caps slots mi \<lbrace>\<lambda>rv s. P (arch_state s)\<rbrace>"
by (rule transfer_caps_loop_pres) wp
lemma transfer_caps_loop_valid_arch[wp]:
"\<lbrace>valid_arch_state\<rbrace> transfer_caps_loop ep diminish buffer n caps slots mi \<lbrace>\<lambda>rv. valid_arch_state\<rbrace>"
"\<lbrace>valid_arch_state\<rbrace> transfer_caps_loop ep buffer n caps slots mi \<lbrace>\<lambda>rv. valid_arch_state\<rbrace>"
by (rule valid_arch_state_lift) wp
@ -987,7 +986,7 @@ lemma tcl_reply':
"\<lbrace>valid_reply_caps and valid_reply_masters and valid_objs and valid_mdb and K(distinct slots)
and (\<lambda>s. \<forall>x \<in> set slots. real_cte_at x s \<and> cte_wp_at (\<lambda>cap. cap = cap.NullCap) x s)
and transfer_caps_srcs caps\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv. valid_reply_caps and valid_reply_masters\<rbrace>"
apply (rule hoare_pre)
apply (rule transfer_caps_loop_presM[where vo=True and em=False and ex=False])
@ -1022,7 +1021,7 @@ lemmas tcl_reply_masters[wp] = tcl_reply' [THEN hoare_strengthen_post
lemma transfer_caps_loop_irq_node[wp]:
"\<lbrace>\<lambda>s. P (interrupt_irq_node s)\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv s. P (interrupt_irq_node s)\<rbrace>"
by (rule transfer_caps_loop_pres) wp
@ -1041,7 +1040,7 @@ lemma transfer_caps_loop_irq_handlers[wp]:
"\<lbrace>valid_irq_handlers and valid_objs and valid_mdb and K (distinct slots)
and (\<lambda>s. \<forall>x \<in> set slots. real_cte_at x s \<and> cte_wp_at (\<lambda>cap. cap = cap.NullCap) x s)
and transfer_caps_srcs caps\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv. valid_irq_handlers\<rbrace>"
apply (rule hoare_pre)
apply (rule transfer_caps_loop_presM[where vo=True and em=False and ex=False])
@ -1064,7 +1063,7 @@ crunch valid_arch_objs [wp]: set_extra_badge valid_arch_objs
lemma transfer_caps_loop_arch_objs[wp]:
"\<lbrace>valid_arch_objs\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv. valid_arch_objs\<rbrace>"
by (rule transfer_caps_loop_pres) wp
@ -1075,7 +1074,7 @@ lemma transfer_caps_loop_valid_arch_caps[wp]:
"\<lbrace>valid_arch_caps and valid_objs and valid_mdb and K(distinct slots)
and (\<lambda>s. \<forall>x \<in> set slots. real_cte_at x s \<and> cte_wp_at (\<lambda>cap. cap = cap.NullCap) x s)
and transfer_caps_srcs caps\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv. valid_arch_caps\<rbrace>"
apply (wp transfer_caps_loop_presM[where vo=True and em=False and ex=False]
cap_insert_valid_arch_caps)
@ -1095,7 +1094,7 @@ crunch valid_global_objs [wp]: set_extra_badge valid_global_objs
lemma transfer_caps_loop_valid_global_objs[wp]:
"\<lbrace>valid_global_objs\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv. valid_global_objs\<rbrace>"
by (wp transfer_caps_loop_pres cap_insert_valid_global_objs)
@ -1105,7 +1104,7 @@ crunch valid_kernel_mappings [wp]: set_extra_badge valid_kernel_mappings
lemma transfer_caps_loop_v_ker_map[wp]:
"\<lbrace>valid_kernel_mappings\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv. valid_kernel_mappings\<rbrace>"
by (wp transfer_caps_loop_pres)
@ -1115,7 +1114,7 @@ crunch equal_kernel_mappings [wp]: set_extra_badge equal_kernel_mappings
lemma transfer_caps_loop_eq_ker_map[wp]:
"\<lbrace>equal_kernel_mappings\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv. equal_kernel_mappings\<rbrace>"
by (wp transfer_caps_loop_pres)
@ -1125,7 +1124,7 @@ crunch valid_asid_map [wp]: set_extra_badge valid_asid_map
lemma transfer_caps_loop_asid_map[wp]:
"\<lbrace>valid_asid_map\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv. valid_asid_map\<rbrace>"
by (wp transfer_caps_loop_pres | simp)+
@ -1135,7 +1134,7 @@ crunch only_idle [wp]: set_extra_badge only_idle
lemma transfer_caps_loop_only_idle[wp]:
"\<lbrace>only_idle\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv. only_idle\<rbrace>"
by (wp transfer_caps_loop_pres | simp)+
@ -1145,7 +1144,7 @@ crunch valid_global_pd_mappings [wp]: set_extra_badge valid_global_pd_mappings
lemma transfer_caps_loop_valid_global_pd_mappings[wp]:
"\<lbrace>valid_global_pd_mappings\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv. valid_global_pd_mappings\<rbrace>"
by (wp transfer_caps_loop_pres)
@ -1155,7 +1154,7 @@ crunch pspace_in_kernel_window [wp]: set_extra_badge pspace_in_kernel_window
lemma transfer_caps_loop_pspace_in_kernel_window[wp]:
"\<lbrace>pspace_in_kernel_window\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv. pspace_in_kernel_window\<rbrace>"
by (wp transfer_caps_loop_pres)
@ -1166,7 +1165,7 @@ lemma transfer_caps_loop_cap_refs_in_kernel_window [wp]:
"\<lbrace>cap_refs_in_kernel_window and valid_objs and valid_mdb and K (distinct slots)
and (\<lambda>s. \<forall>slot \<in> set slots. real_cte_at slot s \<and> cte_wp_at (\<lambda>cap. cap = cap.NullCap) slot s )
and transfer_caps_srcs caps\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv. cap_refs_in_kernel_window\<rbrace>"
apply (rule hoare_pre)
apply (rule transfer_caps_loop_presM[where em=False and ex=False and vo=True])
@ -1187,7 +1186,7 @@ crunch valid_ioc[wp]: store_word_offs valid_ioc
lemma transfer_caps_loop_valid_ioc[wp]:
"\<lbrace>\<lambda>s. valid_ioc s\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>_. valid_ioc\<rbrace>"
by (wp transfer_caps_loop_pres | simp add: set_extra_badge_def)+
@ -1248,7 +1247,7 @@ by (simp add: set_extra_badge_def) wp
lemma transfer_caps_loop_vms[wp]:
"\<lbrace>\<lambda>s. valid_machine_state s\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>_. valid_machine_state\<rbrace>"
by (wp transfer_caps_loop_pres)
@ -1257,7 +1256,7 @@ crunch valid_irq_states[wp]: set_extra_badge "valid_irq_states"
lemma transfer_caps_loop_valid_irq_states[wp]:
"\<lbrace>\<lambda>s. valid_irq_states s\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>_. valid_irq_states\<rbrace>"
apply(wp transfer_caps_loop_pres)
done
@ -1267,7 +1266,7 @@ lemma transfer_caps_loop_invs[wp]:
\<and> (\<forall>x \<in> set slots. ex_cte_cap_wp_to is_cnode_cap x s) \<and> distinct slots
\<and> (\<forall>x \<in> set slots. real_cte_at x s \<and> cte_wp_at (\<lambda>cap. cap = cap.NullCap) x s)
\<and> transfer_caps_srcs caps s\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv. invs\<rbrace>"
apply (simp add: invs_def valid_state_def valid_pspace_def)
apply (rule hoare_pre)
@ -1309,7 +1308,7 @@ lemma grs_distinct[wp]:
lemma transfer_caps_mi_label[wp]:
"\<lbrace>\<lambda>s. P (mi_label mi)\<rbrace>
transfer_caps mi caps ep receiver recv_buf diminish
transfer_caps mi caps ep receiver recv_buf
\<lbrace>\<lambda>mi' s. P (mi_label mi')\<rbrace>"
apply (simp add: transfer_caps_def)
apply (wp | wpc)+
@ -1319,7 +1318,7 @@ lemma transfer_caps_mi_label[wp]:
lemma transfer_cap_typ_at[wp]:
"\<lbrace>\<lambda>s. P (typ_at T p s)\<rbrace>
transfer_caps mi caps ep receiver recv_buf diminish
transfer_caps mi caps ep receiver recv_buf
\<lbrace>\<lambda>rv s. P (typ_at T p s)\<rbrace>"
apply (simp add: transfer_caps_def split_def split del: split_if |
wp cap_insert_typ_at hoare_drop_imps|wpc)+
@ -1328,7 +1327,7 @@ lemma transfer_cap_typ_at[wp]:
lemma transfer_cap_tcb[wp]:
"\<lbrace>tcb_at t\<rbrace>
transfer_caps mi caps ep receiver recv_buf diminish
transfer_caps mi caps ep receiver recv_buf
\<lbrace>\<lambda>rv. tcb_at t\<rbrace>"
by (simp add: tcb_at_typ, wp)
@ -1662,7 +1661,7 @@ crunch typ_at[wp]: do_normal_transfer "\<lambda>s. P (typ_at T p s)"
lemma do_normal_tcb[wp]:
"\<lbrace>tcb_at t\<rbrace>
do_normal_transfer sender send_buf ep badge
can_grant receiver recv_buf diminish
can_grant receiver recv_buf
\<lbrace>\<lambda>rv. tcb_at t\<rbrace>"
by (simp add: tcb_at_typ, wp)
@ -1859,7 +1858,7 @@ crunch typ_at[wp]: do_ipc_transfer "\<lambda>s. P (typ_at T p s)"
lemma do_ipc_transfer_valid_arch[wp]:
"\<lbrace>valid_arch_state\<rbrace> do_ipc_transfer s ep bg grt r dim \<lbrace>\<lambda>rv. valid_arch_state\<rbrace>"
"\<lbrace>valid_arch_state\<rbrace> do_ipc_transfer s ep bg grt r \<lbrace>\<lambda>rv. valid_arch_state\<rbrace>"
by (rule valid_arch_state_lift) wp
@ -2081,7 +2080,7 @@ crunch vms[wp]: do_ipc_transfer valid_machine_state (wp: mapM_UNIV_wp)
lemma do_ipc_transfer_invs[wp]:
"\<lbrace>invs and tcb_at r and tcb_at s\<rbrace>
do_ipc_transfer s ep bg grt r dim
do_ipc_transfer s ep bg grt r
\<lbrace>\<lambda>rv. invs\<rbrace>"
apply (simp add: do_ipc_transfer_def)
apply (wp|wpc)+
@ -2099,12 +2098,12 @@ lemma do_ipc_transfer_invs[wp]:
done
lemma dit_tcb_at [wp]:
"\<lbrace>tcb_at t\<rbrace> do_ipc_transfer s ep bg grt r dim \<lbrace>\<lambda>rv. tcb_at t\<rbrace>"
"\<lbrace>tcb_at t\<rbrace> do_ipc_transfer s ep bg grt r \<lbrace>\<lambda>rv. tcb_at t\<rbrace>"
by (simp add: tcb_at_typ) wp
lemma dit_cte_at [wp]:
"\<lbrace>cte_at t\<rbrace> do_ipc_transfer s ep bg grt r dim \<lbrace>\<lambda>rv. cte_at t\<rbrace>"
"\<lbrace>cte_at t\<rbrace> do_ipc_transfer s ep bg grt r \<lbrace>\<lambda>rv. cte_at t\<rbrace>"
by (wp valid_cte_at_typ)
@ -2137,7 +2136,7 @@ lemmas cap_insert_typ_ats [wp] = abs_typ_at_lifts [OF cap_insert_typ_at]
lemma transfer_caps_loop_cte_wp_at:
assumes imp: "\<And>cap. P cap \<Longrightarrow> \<not> is_untyped_cap cap"
shows "\<lbrace>cte_wp_at P sl and K (sl \<notin> set slots) and (\<lambda>s. \<forall>x \<in> set slots. cte_at x s)\<rbrace>
transfer_caps_loop ep diminish buffer n caps slots mi
transfer_caps_loop ep buffer n caps slots mi
\<lbrace>\<lambda>rv. cte_wp_at P sl\<rbrace>"
apply (induct caps arbitrary: slots n mi)
apply (simp, wp, simp)
@ -2163,7 +2162,7 @@ lemma transfer_caps_loop_cte_wp_at:
lemma transfer_caps_tcb_caps:
assumes imp: "\<And>c. P c \<Longrightarrow> \<not> is_untyped_cap c"
shows "\<lbrace>valid_objs and cte_wp_at P (t, ref) and tcb_at t\<rbrace>
transfer_caps mi caps ep receiver recv_buf diminish
transfer_caps mi caps ep receiver recv_buf
\<lbrace>\<lambda>rv. cte_wp_at P (t, ref)\<rbrace>"
apply (simp add: transfer_caps_def)
apply (wp hoare_vcg_const_Ball_lift hoare_vcg_const_imp_lift
@ -2194,7 +2193,7 @@ crunch cte_wp_at[wp]: do_fault_transfer "cte_wp_at P p"
lemma transfer_caps_non_null_cte_wp_at:
assumes imp: "\<And>c. P c \<Longrightarrow> \<not> is_untyped_cap c"
shows "\<lbrace>valid_objs and cte_wp_at (P and (op \<noteq> cap.NullCap)) ptr\<rbrace>
transfer_caps mi caps ep receiver recv_buf diminish
transfer_caps mi caps ep receiver recv_buf
\<lbrace>\<lambda>_. cte_wp_at (P and (op \<noteq> cap.NullCap)) ptr\<rbrace>"
unfolding transfer_caps_def
apply simp
@ -2221,7 +2220,7 @@ lemma transfer_caps_non_null_cte_wp_at:
lemma do_normal_transfer_non_null_cte_wp_at:
assumes imp: "\<And>c. P c \<Longrightarrow> \<not> is_untyped_cap c"
shows "\<lbrace>valid_objs and cte_wp_at (P and (op \<noteq> cap.NullCap)) ptr\<rbrace>
do_normal_transfer st send_buffer ep b gr rt recv_buffer dim
do_normal_transfer st send_buffer ep b gr rt recv_buffer
\<lbrace>\<lambda>_. cte_wp_at (P and (op \<noteq> cap.NullCap)) ptr\<rbrace>"
unfolding do_normal_transfer_def
apply simp
@ -2234,7 +2233,7 @@ lemma do_ipc_transfer_non_null_cte_wp_at:
assumes imp: "\<And>c. P c \<Longrightarrow> \<not> is_untyped_cap c"
shows
"\<lbrace>valid_objs and cte_wp_at (P and (op \<noteq> cap.NullCap)) ptr\<rbrace>
do_ipc_transfer st ep b gr rt dim
do_ipc_transfer st ep b gr rt
\<lbrace>\<lambda>_. cte_wp_at (P and (op \<noteq> cap.NullCap)) ptr\<rbrace>"
unfolding do_ipc_transfer_def
apply (wp do_normal_transfer_non_null_cte_wp_at hoare_drop_imp hoare_allI
@ -2293,7 +2292,7 @@ lemma do_normal_transfer_tcb_caps:
assumes imp: "\<And>c. P c \<Longrightarrow> \<not> is_untyped_cap c"
shows
"\<lbrace>valid_objs and cte_wp_at P (t, ref) and tcb_at t\<rbrace>
do_normal_transfer st sb ep badge grant rt rb dim
do_normal_transfer st sb ep badge grant rt rb
\<lbrace>\<lambda>rv. cte_wp_at P (t, ref)\<rbrace>"
apply (simp add: do_normal_transfer_def)
apply (rule hoare_pre)
@ -2306,7 +2305,7 @@ lemma do_ipc_transfer_tcb_caps:
assumes imp: "\<And>c. P c \<Longrightarrow> \<not> is_untyped_cap c"
shows
"\<lbrace>valid_objs and cte_wp_at P (t, ref) and tcb_at t\<rbrace>
do_ipc_transfer st ep b gr rt dim
do_ipc_transfer st ep b gr rt
\<lbrace>\<lambda>rv. cte_wp_at P (t, ref)\<rbrace>"
apply (simp add: do_ipc_transfer_def)
apply (rule hoare_pre)
@ -2885,7 +2884,7 @@ lemma ri_invs':
assumes sts_Q[wp]: "\<And>a b. \<lbrace>Q\<rbrace> set_thread_state a b \<lbrace>\<lambda>_.Q\<rbrace>"
assumes ext_Q[wp]: "\<And>a (s::'a::state_ext state). \<lbrace>Q and valid_objs\<rbrace> do_extended_op (switch_if_required_to a) \<lbrace>\<lambda>_.Q\<rbrace>"
assumes scc_Q[wp]: "\<And>a b. \<lbrace>valid_mdb and Q\<rbrace> setup_caller_cap a b \<lbrace>\<lambda>_.Q\<rbrace>"
assumes dit_Q[wp]: "\<And>a b c d e f. \<lbrace>valid_mdb and valid_objs and Q\<rbrace> do_ipc_transfer a b c d e f \<lbrace>\<lambda>_.Q\<rbrace>"
assumes dit_Q[wp]: "\<And>a b c d e. \<lbrace>valid_mdb and valid_objs and Q\<rbrace> do_ipc_transfer a b c d e \<lbrace>\<lambda>_.Q\<rbrace>"
assumes failed_transfer_Q[wp]: "\<And>a. \<lbrace>Q\<rbrace> do_nbrecv_failed_transfer a \<lbrace>\<lambda>_. Q\<rbrace>"
notes dxo_wp_weak[wp del]
shows
@ -3213,7 +3212,7 @@ lemma si_invs':
assumes ext_Q[wp]: "\<And>a. \<lbrace>Q and valid_objs\<rbrace> do_extended_op (attempt_switch_to a) \<lbrace>\<lambda>_. Q\<rbrace>"
assumes sts_Q[wp]: "\<And>a b. \<lbrace>Q\<rbrace> set_thread_state a b \<lbrace>\<lambda>_.Q\<rbrace>"
assumes setup_caller_cap_Q[wp]: "\<And>send receive. \<lbrace>Q and valid_mdb\<rbrace> setup_caller_cap send receive \<lbrace>\<lambda>_.Q\<rbrace>"
assumes do_ipc_transfer_Q[wp]: "\<And>a b c d e f. \<lbrace>Q and valid_objs and valid_mdb\<rbrace> do_ipc_transfer a b c d e f \<lbrace>\<lambda>_.Q\<rbrace>"
assumes do_ipc_transfer_Q[wp]: "\<And>a b c d e. \<lbrace>Q and valid_objs and valid_mdb\<rbrace> do_ipc_transfer a b c d e \<lbrace>\<lambda>_.Q\<rbrace>"
notes dxo_wp_weak[wp del]
shows
"\<lbrace>invs and Q and st_tcb_at active t
@ -3270,8 +3269,7 @@ lemma si_invs':
apply (drule(1) bspec, clarsimp simp: pred_tcb_at_def obj_at_def)
apply (wp, simp)
apply (rename_tac list)
apply (case_tac list, simp_all)
apply (simp add: invs_def valid_state_def valid_pspace_def)
apply (case_tac list, simp_all add: invs_def valid_state_def valid_pspace_def split del:split_if)
apply (rule hoare_pre)
apply (wp valid_irq_node_typ)
apply (simp add: if_apply_def2)
@ -3322,7 +3320,7 @@ lemma hf_invs':
assumes sts_Q[wp]: "\<And>a b. \<lbrace>Q\<rbrace> set_thread_state a b \<lbrace>\<lambda>_.Q\<rbrace>"
assumes ext_Q[wp]: "\<And>a. \<lbrace>Q and valid_objs\<rbrace> do_extended_op (attempt_switch_to a) \<lbrace>\<lambda>_.Q\<rbrace>"
assumes setup_caller_cap_Q[wp]: "\<And>send receive. \<lbrace>Q and valid_mdb\<rbrace> setup_caller_cap send receive \<lbrace>\<lambda>_.Q\<rbrace>"
assumes do_ipc_transfer_Q[wp]: "\<And>a b c d e f. \<lbrace>Q and valid_objs and valid_mdb\<rbrace> do_ipc_transfer a b c d e f \<lbrace>\<lambda>_.Q\<rbrace>"
assumes do_ipc_transfer_Q[wp]: "\<And>a b c d e. \<lbrace>Q and valid_objs and valid_mdb\<rbrace> do_ipc_transfer a b c d e \<lbrace>\<lambda>_.Q\<rbrace>"
assumes thread_set_Q[wp]: "\<And>a b. \<lbrace>Q\<rbrace> thread_set a b \<lbrace>\<lambda>_.Q\<rbrace>"
notes si_invs''[wp] = si_invs'[where Q=Q]
shows
@ -3439,7 +3437,7 @@ lemma si_blk_makes_simple:
apply clarsimp
apply (rule hoare_gen_asm[simplified])
apply (rename_tac list)
apply (case_tac list, simp_all)
apply (case_tac list, simp_all split del:split_if)
apply (rule hoare_seq_ext [OF _ set_ep_pred_tcb_at])
apply (rule hoare_seq_ext [OF _ gts_sp])
apply (case_tac recv_state, simp_all split del: split_if)

View File

@ -135,11 +135,11 @@ lemma st_tcb_at_eq:
by (clarsimp simp add: pred_tcb_at_def obj_at_def)
lemma do_ipc_transfer_tcb_at [wp]:
"\<lbrace>\<lambda>s. P (tcb_at t s)\<rbrace> do_ipc_transfer s ep bg grt r dim \<lbrace>\<lambda>rv s. P (tcb_at t s)\<rbrace>"
"\<lbrace>\<lambda>s. P (tcb_at t s)\<rbrace> do_ipc_transfer s ep bg grt r \<lbrace>\<lambda>rv s. P (tcb_at t s)\<rbrace>"
by (simp add: tcb_at_typ) wp
lemma do_ipc_transfer_emptyable[wp]:
"\<lbrace>emptyable sl\<rbrace> do_ipc_transfer sender ep badge grant receiver diminis \<lbrace>\<lambda>_. emptyable sl\<rbrace>"
"\<lbrace>emptyable sl\<rbrace> do_ipc_transfer sender ep badge grant receiver \<lbrace>\<lambda>_. emptyable sl\<rbrace>"
apply (clarsimp simp add: emptyable_def)
apply (wp hoare_convert_imp | clarsimp)+
done
@ -150,7 +150,7 @@ lemma do_ipc_transfer_non_null_cte_wp_at2:
fixes P
assumes PNN: "\<And>cap. P cap \<Longrightarrow> cap \<noteq> cap.NullCap"
assumes PUC: "\<And>cap. P cap \<Longrightarrow> \<not> is_untyped_cap cap"
shows "\<lbrace>valid_objs and cte_wp_at P ptr\<rbrace> do_ipc_transfer st ep b gr rt dim \<lbrace>\<lambda>_. cte_wp_at P ptr\<rbrace>"
shows "\<lbrace>valid_objs and cte_wp_at P ptr\<rbrace> do_ipc_transfer st ep b gr rt \<lbrace>\<lambda>_. cte_wp_at P ptr\<rbrace>"
proof -
have PimpQ: "\<And>P Q ptr s. \<lbrakk> cte_wp_at P ptr s; \<And>cap. P cap \<Longrightarrow> Q cap \<rbrakk> \<Longrightarrow> cte_wp_at (P and Q) ptr s"
by (erule cte_wp_at_weakenE, clarsimp)
@ -1253,7 +1253,7 @@ lemma he_invs[wp]:
apply (case_tac e, simp_all)
apply (rename_tac syscall)
apply (case_tac syscall, simp_all)
apply ((rule hoare_pre, wp hvmf_active hr_invs hy_inv) |
apply ((rule hoare_pre, wp hvmf_active hr_invs hy_inv ) |
wpc | wp hoare_drop_imps hoare_vcg_all_lift |
fastforce simp: tcb_at_invs ct_in_state_def valid_fault_def
elim!: st_tcb_ex_cap)+

View File

@ -77,7 +77,7 @@ lemma ball_tcb_cap_casesI:
\<and> obj_ref_of c = t)
\<or> (halted st \<and> (c = cap.NullCap))));
P (tcb_caller, tcb_caller_update, (\<lambda>_ st. case st of
Structures_A.BlockedOnReceive e d \<Rightarrow>
Structures_A.BlockedOnReceive e \<Rightarrow>
(op = cap.NullCap)
| _ \<Rightarrow> is_reply_cap or (op = cap.NullCap)));
P (tcb_ipcframe, tcb_ipcframe_update, (\<lambda>_ _. is_arch_cap or (op = cap.NullCap))) \<rbrakk>
@ -787,7 +787,7 @@ lemma idle_thread_idle[wp]:
lemma set_thread_state_valid_objs[wp]:
"\<lbrace>valid_objs and valid_tcb_state st and
(\<lambda>s. (\<forall>a b. st = Structures_A.BlockedOnReceive a b \<longrightarrow>
(\<lambda>s. (\<forall>a. st = Structures_A.BlockedOnReceive a \<longrightarrow>
cte_wp_at (op = cap.NullCap) (thread, tcb_cnode_index 3) s) \<and>
(st_tcb_at (\<lambda>st. \<not> halted st) thread s \<or> halted st \<or>
cte_wp_at (\<lambda>c. is_master_reply_cap c \<and> obj_ref_of c = thread)
@ -1450,7 +1450,7 @@ lemma st_tcb_reply_cap_valid:
split: Structures_A.thread_state.split_asm)
lemma st_tcb_caller_cap_null:
"\<And>ep. \<forall>cap. (st_tcb_at (\<lambda>st. \<exists>b. st = Structures_A.BlockedOnReceive ep b) t s \<and>
"\<And>ep. \<forall>cap. (st_tcb_at (\<lambda>st. st = Structures_A.BlockedOnReceive ep) t s \<and>
tcb_cap_valid cap (t, tcb_cnode_index 3) s) \<longrightarrow>
cap = cap.NullCap"
by (clarsimp simp: tcb_cap_valid_def st_tcb_at_tcb_at st_tcb_def2)
@ -1633,7 +1633,7 @@ lemma set_bound_notification_vms[wp]:
lemma sts_invs_minor:
"\<lbrace>st_tcb_at (\<lambda>st'. tcb_st_refs_of st' = tcb_st_refs_of st) t
and (\<lambda>s. \<not> halted st \<longrightarrow> ex_nonz_cap_to t s)
and (\<lambda>s. \<forall>a b. st = Structures_A.BlockedOnReceive a b \<longrightarrow>
and (\<lambda>s. \<forall>a. st = Structures_A.BlockedOnReceive a \<longrightarrow>
cte_wp_at (op = cap.NullCap) (t, tcb_cnode_index 3) s)
and (\<lambda>s. t \<noteq> idle_thread s)
and (\<lambda>s. st_tcb_at (\<lambda>st. \<not> halted st) t s \<or> halted st \<or>
@ -1668,7 +1668,7 @@ lemma sts_invs_minor2:
"\<lbrace>st_tcb_at (\<lambda>st'. tcb_st_refs_of st' = tcb_st_refs_of st \<and> \<not> awaiting_reply st') t
and invs and ex_nonz_cap_to t and (\<lambda>s. t \<noteq> idle_thread s)
and K (\<not> awaiting_reply st \<and> \<not>idle st)
and (\<lambda>s. \<forall>a b. st = Structures_A.BlockedOnReceive a b \<longrightarrow>
and (\<lambda>s. \<forall>a. st = Structures_A.BlockedOnReceive a \<longrightarrow>
cte_wp_at (op = cap.NullCap) (t, tcb_cnode_index 3) s)
and (\<lambda>s. st_tcb_at (\<lambda>st. \<not> halted st) t s \<or> halted st \<or>
cte_wp_at (\<lambda>c. is_master_reply_cap c \<and> obj_ref_of c = t)
@ -1841,8 +1841,8 @@ lemma ntfn_queued_st_tcb_at:
lemma ep_queued_st_tcb_at:
"\<And>P. \<lbrakk>ko_at (Endpoint ep) ptr s; (t, rt) \<in> ep_q_refs_of ep;
valid_objs s; sym_refs (state_refs_of s);
\<And>ref pl dim. P (Structures_A.BlockedOnSend ref pl) \<and>
P (Structures_A.BlockedOnReceive ref dim) \<rbrakk>
\<And>ref pl. P (Structures_A.BlockedOnSend ref pl) \<and>
P (Structures_A.BlockedOnReceive ref) \<rbrakk>
\<Longrightarrow> st_tcb_at P t s"
apply (case_tac ep, simp_all)
apply (frule(1) sym_refs_ko_atD, clarsimp, erule (1) my_BallE,

View File

@ -326,7 +326,7 @@ lemma hf_cte_at[wp]:
lemma do_ipc_transfer_cte_at[wp]:
"\<lbrace>cte_at p\<rbrace> do_ipc_transfer s ep b g r d \<lbrace>\<lambda>_. cte_at p\<rbrace>"
"\<lbrace>cte_at p\<rbrace> do_ipc_transfer s ep b g r \<lbrace>\<lambda>_. cte_at p\<rbrace>"
by (wp valid_cte_at_typ)

View File

@ -356,8 +356,8 @@ where
Structures_A.thread_state.IdleThreadState"
| "ThStateMap Structures_H.thread_state.BlockedOnReply =
Structures_A.thread_state.BlockedOnReply"
| "ThStateMap (Structures_H.thread_state.BlockedOnReceive oref dimin) =
Structures_A.thread_state.BlockedOnReceive oref dimin"
| "ThStateMap (Structures_H.thread_state.BlockedOnReceive oref) =
Structures_A.thread_state.BlockedOnReceive oref"
| "ThStateMap (Structures_H.thread_state.BlockedOnSend oref badge grant call) =
Structures_A.thread_state.BlockedOnSend oref
\<lparr> sender_badge = badge,

View File

@ -131,7 +131,7 @@ lemma deriveCap_empty_fail[intro!, wp, simp]:
crunch (empty_fail) empty_fail[intro!, wp, simp]: setExtraBadge, cteInsert
lemma transferCapsToSlots_empty_fail[intro!, wp, simp]:
"empty_fail (transferCapsToSlots ep dim buffer n caps slots mi)"
"empty_fail (transferCapsToSlots ep buffer n caps slots mi)"
apply (induct caps arbitrary: slots n mi)
apply simp
apply (simp add: Let_def split_def

View File

@ -550,24 +550,29 @@ lemma timerTick_corres:
apply simp
done
lemmas corres_eq_trivial = corres_Id[where f = h and g = h for h, simplified]
thm corres_Id
lemma handle_interrupt_corres:
"corres dc
(einvs) (invs' and (\<lambda>s. intStateIRQTable (ksInterruptState s) irq \<noteq> IRQInactive))
(handle_interrupt irq) (handleInterrupt irq)"
(is "corres dc (einvs) ?P' ?f ?g")
apply (simp add: handle_interrupt_def handleInterrupt_def)
apply (simp add: handle_interrupt_def handleInterrupt_def )
apply (rule conjI[rotated]; rule impI)
apply (rule corres_guard_imp)
apply (rule corres_split [OF _ get_irq_state_corres,
where R="\<lambda>rv. einvs"
and R'="\<lambda>rv. invs' and (\<lambda>s. rv \<noteq> IRQInactive)"])
and R'="\<lambda>rv. invs' and (\<lambda>s. rv \<noteq> IRQInactive)"])
defer
apply (wp getIRQState_prop getIRQState_inv)
apply simp+
apply (rule corres_gen_asm2)
apply (rule corres_split'[where r'=dc, rotated])
apply (rule corres_machine_op)
apply (rule corres_Id, simp+)
apply wp
apply (wp getIRQState_prop getIRQState_inv do_machine_op_bind doMachineOp_bind | simp add: do_machine_op_bind doMachineOp_bind )+
apply (rule corres_guard_imp)
apply (rule corres_split)
apply (rule corres_machine_op, rule corres_eq_trivial ; (simp add: dc_def no_fail_maskInterrupt no_fail_bind no_fail_ackInterrupt)+)+
apply ((wp | simp)+)[4]
apply (rule corres_gen_asm2)
apply (case_tac st, simp_all add: irq_state_relation_def split: irqstate.split_asm)
apply (simp add: getSlotCap_def bind_assoc)
apply (rule corres_guard_imp)
@ -577,24 +582,29 @@ lemma handle_interrupt_corres:
where R="\<lambda>rv. einvs and valid_cap rv"
and R'="\<lambda>rv. invs' and valid_cap' (cteCap rv)"])
apply (rule corres_split'[where r'=dc])
apply (case_tac cap, simp_all add: when_False when_True doMachineOp_return)[1]
apply (case_tac xb, simp_all add: when_False doMachineOp_return )[1]
apply (clarsimp simp add: when_def doMachineOp_return)
apply (rule corres_guard_imp, rule send_signal_corres)
apply (clarsimp simp: valid_cap_def)
apply (clarsimp simp: valid_cap'_def)
apply (clarsimp simp: doMachineOp_return)
apply (rule corres_machine_op)
apply (rule corres_Id, simp+)
apply (wp | simp)+
apply (clarsimp simp: valid_cap_def valid_cap'_def do_machine_op_bind doMachineOp_bind)+
apply ( rule corres_split)
apply (rule corres_machine_op, rule corres_eq_trivial ; (simp add: no_fail_maskInterrupt no_fail_bind no_fail_ackInterrupt)+)+
apply ((wp |simp)+)
apply clarsimp
apply fastforce
apply (rule corres_guard_imp)
apply (rule corres_split)
apply simp
apply (rule corres_split [OF corres_machine_op timerTick_corres])
apply (rule corres_eq_trivial, simp+)
apply (rule corres_machine_op)
apply (rule corres_eq_trivial, (simp add: no_fail_ackInterrupt)+)
apply wp
apply clarsimp
apply clarsimp
apply fastforce
apply (rule corres_guard_imp)
apply (rule corres_split [OF corres_machine_op timerTick_corres])
apply (rule corres_Id, simp+)
apply wp
apply (clarsimp simp: invs_def valid_state_def valid_sched_def
valid_sched_action_def weak_valid_sched_action_def)+
done
lemma invs_ChooseNewThread:
@ -847,16 +857,20 @@ lemma dmo_ackInterrupt[wp]:
lemma hint_invs[wp]:
"\<lbrace>invs'\<rbrace> handleInterrupt irq \<lbrace>\<lambda>rv. invs'\<rbrace>"
apply (simp add: handleInterrupt_def getSlotCap_def
cong: irqstate.case_cong)
apply (wp dmo_maskInterrupt_True getCTE_wp'
| wpc | simp)+
cong: irqstate.case_cong)
apply (rule conjI; rule impI)
apply (wp dmo_maskInterrupt_True getCTE_wp'
| wpc | simp add: doMachineOp_bind )+
apply (rule_tac Q="\<lambda>rv. invs'" in hoare_post_imp)
apply (clarsimp simp: cte_wp_at_ctes_of ex_nonz_cap_to'_def)
apply fastforce
apply (wp threadSet_invs_trivial | simp add: inQ_def)+
apply (rule hoare_strengthen_post, rule getIRQState_inv)
apply (auto simp add: tcb_at_invs')
apply (wp hoare_post_comb_imp_conj hoare_drop_imp getIRQState_inv)
apply (assumption)
done
crunch st_tcb_at'[wp]: timerTick "st_tcb_at' P t"
(wp: threadSet_pred_tcb_no_state)
@ -864,7 +878,7 @@ crunch st_tcb_at'[wp]: timerTick "st_tcb_at' P t"
lemma handleInterrupt_runnable:
"\<lbrace>st_tcb_at' runnable' t\<rbrace> handleInterrupt irq \<lbrace>\<lambda>_. st_tcb_at' runnable' t\<rbrace>"
apply (simp add: handleInterrupt_def)
apply (rule hoare_pre)
apply (rule conjI; rule impI)
apply (wp sai_st_tcb' hoare_vcg_all_lift hoare_drop_imps
threadSet_pred_tcb_no_state getIRQState_inv haskell_fail_wp
|wpc|simp)+

View File

@ -137,7 +137,7 @@ where
"tcb_st_refs_of' z \<equiv> case z of (Running) => {}
| (Inactive) => {}
| (Restart) => {}
| (BlockedOnReceive x b) => {(x, TCBBlockedRecv)}
| (BlockedOnReceive x) => {(x, TCBBlockedRecv)}
| (BlockedOnSend x a b c) => {(x, TCBBlockedSend)}
| (BlockedOnNotification x) => {(x, TCBSignal)}
| (BlockedOnReply) => {}
@ -399,7 +399,7 @@ definition
valid_tcb_state' :: "Structures_H.thread_state \<Rightarrow> kernel_state \<Rightarrow> bool"
where
"valid_tcb_state' ts s \<equiv> case ts of
Structures_H.BlockedOnReceive ref d \<Rightarrow> ep_at' ref s
Structures_H.BlockedOnReceive ref \<Rightarrow> ep_at' ref s
| Structures_H.BlockedOnSend ref b d c \<Rightarrow> ep_at' ref s
| Structures_H.BlockedOnNotification ref \<Rightarrow> ntfn_at' ref s
| _ \<Rightarrow> True"
@ -799,7 +799,7 @@ where
| "runnable' (Structures_H.Inactive) = False"
| "runnable' (Structures_H.Restart) = True"
| "runnable' (Structures_H.IdleThreadState) = False"
| "runnable' (Structures_H.BlockedOnReceive a b) = False"
| "runnable' (Structures_H.BlockedOnReceive a) = False"
| "runnable' (Structures_H.BlockedOnReply) = False"
| "runnable' (Structures_H.BlockedOnSend a b c d) = False"
| "runnable' (Structures_H.BlockedOnNotification x) = False"
@ -1303,7 +1303,7 @@ lemma capability_splits[split]:
lemma thread_state_splits[split]:
" P (case thread_state of
Structures_H.thread_state.BlockedOnReceive x xa \<Rightarrow> f1 x xa
Structures_H.thread_state.BlockedOnReceive x \<Rightarrow> f1 x
| Structures_H.thread_state.BlockedOnReply \<Rightarrow> f2
| Structures_H.thread_state.BlockedOnNotification x \<Rightarrow> f3 x
| Structures_H.thread_state.Running \<Rightarrow> f4
@ -1312,10 +1312,10 @@ lemma thread_state_splits[split]:
| Structures_H.thread_state.BlockedOnSend x xa xb xc \<Rightarrow>
f7 x xa xb xc
| Structures_H.thread_state.Restart \<Rightarrow> f8) =
((\<forall>x11 x12.
((\<forall>x11.
thread_state =
Structures_H.thread_state.BlockedOnReceive x11 x12 \<longrightarrow>
P (f1 x11 x12)) \<and>
Structures_H.thread_state.BlockedOnReceive x11 \<longrightarrow>
P (f1 x11)) \<and>
(awaiting_reply' thread_state \<longrightarrow> P f2) \<and>
(\<forall>x3. thread_state =
Structures_H.thread_state.BlockedOnNotification x3 \<longrightarrow>
@ -1332,7 +1332,7 @@ lemma thread_state_splits[split]:
P (f7 x71 x72 x73 x74)) \<and>
(thread_state = Structures_H.thread_state.Restart \<longrightarrow> P f8))"
"P (case thread_state of
Structures_H.thread_state.BlockedOnReceive x xa \<Rightarrow> f1 x xa
Structures_H.thread_state.BlockedOnReceive x \<Rightarrow> f1 x
| Structures_H.thread_state.BlockedOnReply \<Rightarrow> f2
| Structures_H.thread_state.BlockedOnNotification x \<Rightarrow> f3 x
| Structures_H.thread_state.Running \<Rightarrow> f4
@ -1341,10 +1341,10 @@ lemma thread_state_splits[split]:
| Structures_H.thread_state.BlockedOnSend x xa xb xc \<Rightarrow>
f7 x xa xb xc
| Structures_H.thread_state.Restart \<Rightarrow> f8) =
(\<not> ((\<exists>x11 x12.
(\<not> ((\<exists>x11.
thread_state =
Structures_H.thread_state.BlockedOnReceive x11 x12 \<and>
\<not> P (f1 x11 x12)) \<or>
Structures_H.thread_state.BlockedOnReceive x11 \<and>
\<not> P (f1 x11)) \<or>
awaiting_reply' thread_state \<and> \<not> P f2 \<or>
(\<exists>x3. thread_state =
Structures_H.thread_state.BlockedOnNotification
@ -1527,7 +1527,7 @@ lemma tcb_st_refs_of'_simps[simp]:
"tcb_st_refs_of' (Running) = {}"
"tcb_st_refs_of' (Inactive) = {}"
"tcb_st_refs_of' (Restart) = {}"
"\<And>x b. tcb_st_refs_of' (BlockedOnReceive x b) = {(x, TCBBlockedRecv)}"
"\<And>x. tcb_st_refs_of' (BlockedOnReceive x) = {(x, TCBBlockedRecv)}"
"\<And>x c. tcb_st_refs_of' (BlockedOnSend x a b c) = {(x, TCBBlockedSend)}"
"\<And>x. tcb_st_refs_of' (BlockedOnNotification x) = {(x, TCBSignal)}"
"tcb_st_refs_of' (BlockedOnReply) = {}"
@ -1558,7 +1558,7 @@ lemma tcb_bound_refs'_simps[simp]:
lemma refs_of_rev':
"(x, TCBBlockedRecv) \<in> refs_of' ko =
(\<exists>tcb. ko = KOTCB tcb \<and> (\<exists>b. tcbState tcb = BlockedOnReceive x b))"
(\<exists>tcb. ko = KOTCB tcb \<and> tcbState tcb = BlockedOnReceive x)"
"(x, TCBBlockedSend) \<in> refs_of' ko =
(\<exists>tcb. ko = KOTCB tcb \<and> (\<exists>a b c. tcbState tcb = BlockedOnSend x a b c))"
"(x, TCBSignal) \<in> refs_of' ko =
@ -1595,7 +1595,7 @@ lemma projectKO_opt_tcbD:
lemma st_tcb_at_refs_of_rev':
"ko_wp_at' (\<lambda>ko. (x, TCBBlockedRecv) \<in> refs_of' ko) t s
= st_tcb_at' (\<lambda>ts. \<exists>b. ts = BlockedOnReceive x b ) t s"
= st_tcb_at' (\<lambda>ts. ts = BlockedOnReceive x ) t s"
"ko_wp_at' (\<lambda>ko. (x, TCBBlockedSend) \<in> refs_of' ko) t s
= st_tcb_at' (\<lambda>ts. \<exists>a b c. ts = BlockedOnSend x a b c) t s"
"ko_wp_at' (\<lambda>ko. (x, TCBSignal) \<in> refs_of' ko) t s

View File

@ -160,7 +160,7 @@ crunch tcb_at[wp]: set_endpoint "tcb_at t"
crunch tcb_at'[wp]: setEndpoint "tcb_at' t"
lemma blocked_cancel_ipc_corres:
"\<lbrakk> st = Structures_A.BlockedOnReceive epPtr x \<or>
"\<lbrakk> st = Structures_A.BlockedOnReceive epPtr \<or>
st = Structures_A.BlockedOnSend epPtr p; thread_state_relation st st' \<rbrakk> \<Longrightarrow>
corres dc (invs and st_tcb_at (op = st) t) (invs' and st_tcb_at' (op = st') t)
(blocked_cancel_ipc st t)
@ -803,7 +803,7 @@ proof -
apply (wp, simp)
done
have Q:
"\<And>epptr. \<lbrace>st_tcb_at' (\<lambda>st. (\<exists>x. st = BlockedOnReceive epptr x)
"\<And>epptr. \<lbrace>st_tcb_at' (\<lambda>st. (st = BlockedOnReceive epptr )
\<or> (\<exists>a b c. st = BlockedOnSend epptr a b c)) t
and invs'\<rbrace>
do ep \<leftarrow> getEndpoint epptr;
@ -1870,8 +1870,8 @@ lemma (in delete_one_conc_pre) cancelIPC_queues[wp]:
elim!: pred_tcb'_weakenE)+
apply (safe)
apply (drule_tac t=t in valid_queues_not_runnable'_not_ksQ)
apply (erule pred_tcb'_weakenE, simp)
apply (drule_tac x=xc in spec, simp)
apply (erule pred_tcb'_weakenE, simp)
apply (drule_tac x=xb in spec, simp)
apply (erule pred_tcb'_weakenE, simp)
apply (drule_tac t=t in valid_queues_not_runnable'_not_ksQ)
apply (erule pred_tcb'_weakenE, simp)

View File

@ -296,11 +296,6 @@ lemma getSlotCap_cte_wp_at_rv:
apply (clarsimp simp: cte_wp_at_ctes_of)
done
lemma valid_cap_Write_update_ext [simp]:
"valid_cap' (if diminish then capAllowWrite_update \<bottom> allRights `~RetypeDecls_H.maskCapRights~` cap else cap) =
valid_cap' cap"
by (rule ext) clarsimp
lemma deriveCap_derived:
"\<lbrace>\<lambda>s. c'\<noteq> capability.NullCap \<longrightarrow> cte_wp_at' (\<lambda>cte. badge_derived' c' (cteCap cte)
\<and> capASID c' = capASID (cteCap cte)
@ -535,8 +530,8 @@ lemma tc_loop_corres:
\<and> (\<forall>(cap, slot) \<in> set caps'. valid_cap' cap s \<and>
cte_wp_at' (\<lambda>cte. cap \<noteq> NullCap \<longrightarrow> cteCap cte \<noteq> cap \<longrightarrow> cteCap cte = maskedAsFull cap cap) slot s)
\<and> 2 + msg_max_length + n + length caps' < unat max_ipc_words)
(transfer_caps_loop ep diminish buffer n caps slots mi)
(transferCapsToSlots ep diminish buffer n caps'
(transfer_caps_loop ep buffer n caps slots mi)
(transferCapsToSlots ep buffer n caps'
(map cte_map slots) mi')"
(is "\<lbrakk> list_all2 ?P caps caps'; ?v \<rbrakk> \<Longrightarrow> ?corres")
proof (induct caps caps' arbitrary: slots n mi mi' rule: list_all2_induct)
@ -640,10 +635,7 @@ next
apply (clarsimp simp:masked_as_full_def is_cap_simps cap_master_cap_simps)
apply (clarsimp split del: split_if)
apply (intro conjI)
prefer 2
apply (clarsimp simp:neq_Nil_conv)
apply (erule mp)
apply (clarsimp split:if_splits)
apply (drule hd_in_set)
apply (drule(1) bspec)
apply (clarsimp split:split_if_asm)
@ -672,7 +664,6 @@ next
apply (case_tac "fst y = capability.NullCap")
apply (clarsimp simp: neq_Nil_conv split del: split_if)+
apply (intro allI impI conjI)
apply (erule mp)
apply (clarsimp split:if_splits)
apply (clarsimp simp:image_def)+
apply (thin_tac "\<forall>x\<in>set ys. Q x" for Q)
@ -690,7 +681,7 @@ declare constOnFailure_wp [wp]
lemma transferCapsToSlots_pres1:
assumes x: "\<And>cap src dest. \<lbrace>P\<rbrace> cteInsert cap src dest \<lbrace>\<lambda>rv. P\<rbrace>"
assumes eb: "\<And>b n. \<lbrace>P\<rbrace> setExtraBadge buffer b n \<lbrace>\<lambda>_. P\<rbrace>"
shows "\<lbrace>P\<rbrace> transferCapsToSlots ep dim buffer n caps slots mi \<lbrace>\<lambda>rv. P\<rbrace>"
shows "\<lbrace>P\<rbrace> transferCapsToSlots ep buffer n caps slots mi \<lbrace>\<lambda>rv. P\<rbrace>"
apply (induct caps arbitrary: slots n mi)
apply simp
apply (simp add: Let_def split_def whenE_def
@ -725,18 +716,8 @@ lemma cteInsert_cte_cap_to':
apply clarsimp+
done
lemma maskCapRights_If_valid[simp]:
"s \<turnstile>' (if diminish then RetypeDecls_H.maskCapRights rs cap else cap)
= s \<turnstile>' cap"
by simp
declare maskCapRights_eq_null[simp]
lemma maskCapRights_if_Null[simp]:
"((if diminish then RetypeDecls_H.maskCapRights rs cap else cap)
= NullCap) = (cap = NullCap)"
by simp
lemma lsfco_cte_wp_at_univ':
"\<lbrace>valid_objs' and valid_cap' cap and K (\<forall>cte rv. P cte rv)\<rbrace>
lookupSlotForCNodeOp f cap idx depth
@ -800,7 +781,7 @@ lemma transferCapsToSlots_presM:
\<and> (pad \<longrightarrow> pspace_aligned' s \<and> pspace_distinct' s)
\<and> (drv \<longrightarrow> vo \<and> pspace_aligned' s \<and> pspace_distinct' s \<and> valid_mdb' s \<and> (\<forall>x \<in> set caps. s \<turnstile>' fst x \<and> cte_wp_at' (\<lambda>cte. fst x \<noteq> NullCap \<longrightarrow> cteCap cte \<noteq> fst x
\<longrightarrow> cteCap cte = maskedAsFull (fst x) (fst x)) (snd x) s))\<rbrace>
transferCapsToSlots ep diminish buffer n caps slots mi
transferCapsToSlots ep buffer n caps slots mi
\<lbrace>\<lambda>rv. P\<rbrace>"
apply (induct caps arbitrary: slots n mi)
apply (simp, wp, simp)
@ -831,7 +812,7 @@ lemma transferCapsToSlots_presM:
apply (clarsimp simp: cte_wp_at_ctes_of remove_rights_def
real_cte_tcb_valid if_apply_def2
split del: split_if)
apply (rule conjI, (clarsimp split del: split_if)+)
apply (rule conjI)
apply (clarsimp simp:cte_wp_at_ctes_of)
apply (case_tac "a = cteCap cte",clarsimp)
apply (clarsimp simp:maskedAsFull_def isCap_simps badge_derived'_def)
@ -853,26 +834,26 @@ lemmas transferCapsToSlots_pres2
lemma transferCapsToSlots_aligned'[wp]:
"\<lbrace>pspace_aligned'\<rbrace>
transferCapsToSlots ep diminish buffer n caps slots mi
transferCapsToSlots ep buffer n caps slots mi
\<lbrace>\<lambda>rv. pspace_aligned'\<rbrace>"
by (wp transferCapsToSlots_pres1)
lemma transferCapsToSlots_distinct'[wp]:
"\<lbrace>pspace_distinct'\<rbrace>
transferCapsToSlots ep diminish buffer n caps slots mi
transferCapsToSlots ep buffer n caps slots mi
\<lbrace>\<lambda>rv. pspace_distinct'\<rbrace>"
by (wp transferCapsToSlots_pres1)
lemma transferCapsToSlots_typ_at'[wp]:
"\<lbrace>\<lambda>s. P (typ_at' T p s)\<rbrace>
transferCapsToSlots ep diminish buffer n caps slots mi
transferCapsToSlots ep buffer n caps slots mi
\<lbrace>\<lambda>rv s. P (typ_at' T p s)\<rbrace>"
by (wp transferCapsToSlots_pres1 setExtraBadge_typ_at')
lemma transferCapsToSlots_valid_objs[wp]:
"\<lbrace>valid_objs' and valid_mdb' and (\<lambda>s. \<forall>x \<in> set slots. real_cte_at' x s \<and> cte_wp_at' (\<lambda>cte. cteCap cte = capability.NullCap) x s)
and (\<lambda>s. \<forall>x \<in> set caps. s \<turnstile>' fst x) and K(distinct slots)\<rbrace>
transferCapsToSlots ep diminish buffer n caps slots mi
transferCapsToSlots ep buffer n caps slots mi
\<lbrace>\<lambda>rv. valid_objs'\<rbrace>"
apply (rule hoare_pre)
apply (rule transferCapsToSlots_presM[where vo=True and emx=False and drv=False and pad=False])
@ -887,7 +868,7 @@ lemma transferCapsToSlots_mdb[wp]:
\<and> (\<forall>x \<in> set slots. ex_cte_cap_to' x s \<and> cte_wp_at' (\<lambda>cte. cteCap cte = capability.NullCap) x s)
\<and> (\<forall>x \<in> set slots. real_cte_at' x s)
\<and> transferCaps_srcs caps s\<rbrace>
transferCapsToSlots ep diminish buffer n caps slots mi
transferCapsToSlots ep buffer n caps slots mi
\<lbrace>\<lambda>rv. valid_mdb'\<rbrace>"
apply (wp transferCapsToSlots_presM[where drv=True and vo=True and emx=True and pad=True])
apply clarsimp
@ -909,7 +890,7 @@ lemma transferCapsToSlots_mdb[wp]:
crunch no_0' [wp]: setExtraBadge no_0_obj'
lemma transferCapsToSlots_no_0_obj' [wp]:
"\<lbrace>no_0_obj'\<rbrace> transferCapsToSlots ep diminish buffer n caps slots mi \<lbrace>\<lambda>rv. no_0_obj'\<rbrace>"
"\<lbrace>no_0_obj'\<rbrace> transferCapsToSlots ep buffer n caps slots mi \<lbrace>\<lambda>rv. no_0_obj'\<rbrace>"
by (wp transferCapsToSlots_pres1)
lemma transferCapsToSlots_vp[wp]:
@ -917,7 +898,7 @@ lemma transferCapsToSlots_vp[wp]:
\<and> (\<forall>x \<in> set slots. ex_cte_cap_to' x s \<and> cte_wp_at' (\<lambda>cte. cteCap cte = capability.NullCap) x s)
\<and> (\<forall>x \<in> set slots. real_cte_at' x s)
\<and> transferCaps_srcs caps s\<rbrace>
transferCapsToSlots ep diminish buffer n caps slots mi
transferCapsToSlots ep buffer n caps slots mi
\<lbrace>\<lambda>rv. valid_pspace'\<rbrace>"
apply (rule hoare_pre)
apply (simp add: valid_pspace'_def | wp)+
@ -936,23 +917,23 @@ crunch queuesL2 [wp]: setExtraBadge "\<lambda>s. P (ksReadyQueuesL2Bitmap s)"
lemma tcts_sch_act[wp]:
"\<lbrace>\<lambda>s. sch_act_wf (ksSchedulerAction s) s\<rbrace>
transferCapsToSlots ep diminish buffer n caps slots mi
transferCapsToSlots ep buffer n caps slots mi
\<lbrace>\<lambda>rv s. sch_act_wf (ksSchedulerAction s) s\<rbrace>"
by (wp sch_act_wf_lift tcb_in_cur_domain'_lift transferCapsToSlots_pres1)
lemma tcts_vq[wp]:
"\<lbrace>Invariants_H.valid_queues\<rbrace> transferCapsToSlots ep diminish buffer n caps slots mi \<lbrace>\<lambda>rv. Invariants_H.valid_queues\<rbrace>"
"\<lbrace>Invariants_H.valid_queues\<rbrace> transferCapsToSlots ep buffer n caps slots mi \<lbrace>\<lambda>rv. Invariants_H.valid_queues\<rbrace>"
by (wp valid_queues_lift transferCapsToSlots_pres1)
lemma tcts_vq'[wp]:
"\<lbrace>valid_queues'\<rbrace> transferCapsToSlots ep diminish buffer n caps slots mi \<lbrace>\<lambda>rv. valid_queues'\<rbrace>"
"\<lbrace>valid_queues'\<rbrace> transferCapsToSlots ep buffer n caps slots mi \<lbrace>\<lambda>rv. valid_queues'\<rbrace>"
by (wp valid_queues_lift' transferCapsToSlots_pres1)
crunch state_refs_of' [wp]: setExtraBadge "\<lambda>s. P (state_refs_of' s)"
lemma tcts_state_refs_of'[wp]:
"\<lbrace>\<lambda>s. P (state_refs_of' s)\<rbrace>
transferCapsToSlots ep diminish buffer n caps slots mi
transferCapsToSlots ep buffer n caps slots mi
\<lbrace>\<lambda>rv s. P (state_refs_of' s)\<rbrace>"
by (wp transferCapsToSlots_pres1)
@ -962,7 +943,7 @@ lemma tcts_iflive[wp]:
"\<lbrace>\<lambda>s. if_live_then_nonz_cap' s \<and> distinct slots \<and>
(\<forall>x\<in>set slots.
ex_cte_cap_to' x s \<and> cte_wp_at' (\<lambda>cte. cteCap cte = capability.NullCap) x s)\<rbrace>
transferCapsToSlots ep diminish buffer n caps slots mi
transferCapsToSlots ep buffer n caps slots mi
\<lbrace>\<lambda>rv. if_live_then_nonz_cap'\<rbrace>"
by (wp transferCapsToSlots_pres2 | simp)+
@ -971,7 +952,7 @@ crunch if_unsafe' [wp]: setExtraBadge if_unsafe_then_cap'
lemma tcts_ifunsafe[wp]:
"\<lbrace>\<lambda>s. if_unsafe_then_cap' s \<and> distinct slots \<and>
(\<forall>x\<in>set slots. cte_wp_at' (\<lambda>cte. cteCap cte = capability.NullCap) x s \<and>
ex_cte_cap_to' x s)\<rbrace> transferCapsToSlots ep diminish buffer n caps slots mi
ex_cte_cap_to' x s)\<rbrace> transferCapsToSlots ep buffer n caps slots mi
\<lbrace>\<lambda>rv. if_unsafe_then_cap'\<rbrace>"
by (wp transferCapsToSlots_pres2 | simp)+
@ -1010,7 +991,7 @@ lemma getSlotCap_not_idle [wp]:
crunch valid_idle' [wp]: setExtraBadge valid_idle'
lemma tcts_idle'[wp]:
"\<lbrace>\<lambda>s. valid_idle' s\<rbrace> transferCapsToSlots ep diminish buffer n caps slots mi
"\<lbrace>\<lambda>s. valid_idle' s\<rbrace> transferCapsToSlots ep buffer n caps slots mi
\<lbrace>\<lambda>rv. valid_idle'\<rbrace>"
apply (rule hoare_pre)
apply (wp transferCapsToSlots_pres1)
@ -1018,13 +999,13 @@ lemma tcts_idle'[wp]:
done
lemma tcts_ct[wp]:
"\<lbrace>cur_tcb'\<rbrace> transferCapsToSlots ep diminish buffer n caps slots mi \<lbrace>\<lambda>rv. cur_tcb'\<rbrace>"
"\<lbrace>cur_tcb'\<rbrace> transferCapsToSlots ep buffer n caps slots mi \<lbrace>\<lambda>rv. cur_tcb'\<rbrace>"
by (wp transferCapsToSlots_pres1 cur_tcb_lift)
crunch valid_arch_state' [wp]: setExtraBadge valid_arch_state'
lemma transferCapsToSlots_valid_arch [wp]:
"\<lbrace>valid_arch_state'\<rbrace> transferCapsToSlots ep diminish buffer n caps slots mi \<lbrace>\<lambda>rv. valid_arch_state'\<rbrace>"
"\<lbrace>valid_arch_state'\<rbrace> transferCapsToSlots ep buffer n caps slots mi \<lbrace>\<lambda>rv. valid_arch_state'\<rbrace>"
by (rule transferCapsToSlots_pres1) wp
crunch valid_global_refs' [wp]: setExtraBadge valid_global_refs'
@ -1033,7 +1014,7 @@ lemma transferCapsToSlots_valid_globals [wp]:
"\<lbrace>valid_global_refs' and valid_objs' and valid_mdb' and pspace_distinct' and pspace_aligned' and K (distinct slots)
and (\<lambda>s. \<forall>x \<in> set slots. real_cte_at' x s \<and> cte_wp_at' (\<lambda>cte. cteCap cte = capability.NullCap) x s)
and transferCaps_srcs caps\<rbrace>
transferCapsToSlots ep diminish buffer n caps slots mi
transferCapsToSlots ep buffer n caps slots mi
\<lbrace>\<lambda>rv. valid_global_refs'\<rbrace>"
apply (wp transferCapsToSlots_presM[where vo=True and emx=False and drv=True and pad=True] | clarsimp)+
apply (clarsimp simp:cte_wp_at_ctes_of)
@ -1053,7 +1034,7 @@ lemma transferCapsToSlots_valid_globals [wp]:
crunch irq_node' [wp]: setExtraBadge "\<lambda>s. P (irq_node' s)"
lemma transferCapsToSlots_irq_node'[wp]:
"\<lbrace>\<lambda>s. P (irq_node' s)\<rbrace> transferCapsToSlots ep diminish buffer n caps slots mi \<lbrace>\<lambda>rv s. P (irq_node' s)\<rbrace>"
"\<lbrace>\<lambda>s. P (irq_node' s)\<rbrace> transferCapsToSlots ep buffer n caps slots mi \<lbrace>\<lambda>rv s. P (irq_node' s)\<rbrace>"
by (wp transferCapsToSlots_pres1)
lemma valid_irq_handlers_ctes_ofD:
@ -1068,7 +1049,7 @@ lemma transferCapsToSlots_irq_handlers[wp]:
and K(distinct slots)
and (\<lambda>s. \<forall>x \<in> set slots. real_cte_at' x s \<and> cte_wp_at' (\<lambda>cte. cteCap cte = capability.NullCap) x s)
and transferCaps_srcs caps\<rbrace>
transferCapsToSlots ep diminish buffer n caps slots mi
transferCapsToSlots ep buffer n caps slots mi
\<lbrace>\<lambda>rv. valid_irq_handlers'\<rbrace>"
apply (wp transferCapsToSlots_presM[where vo=True and emx=False and drv=True and pad=False])
apply (clarsimp simp: is_derived'_def cte_wp_at_ctes_of badge_derived'_def)
@ -1097,17 +1078,17 @@ lemma setExtraBadge_irq_states'[wp]:
done
lemma transferCapsToSlots_irq_states' [wp]:
"\<lbrace>valid_irq_states'\<rbrace> transferCapsToSlots ep diminish buffer n caps slots mi \<lbrace>\<lambda>_. valid_irq_states'\<rbrace>"
"\<lbrace>valid_irq_states'\<rbrace> transferCapsToSlots ep buffer n caps slots mi \<lbrace>\<lambda>_. valid_irq_states'\<rbrace>"
by (wp transferCapsToSlots_pres1)
crunch valid_pde_mappings' [wp]: setExtraBadge valid_pde_mappings'
lemma transferCapsToSlots_pde_mappings'[wp]:
"\<lbrace>valid_pde_mappings'\<rbrace> transferCapsToSlots ep diminish buffer n caps slots mi \<lbrace>\<lambda>rv. valid_pde_mappings'\<rbrace>"
"\<lbrace>valid_pde_mappings'\<rbrace> transferCapsToSlots ep buffer n caps slots mi \<lbrace>\<lambda>rv. valid_pde_mappings'\<rbrace>"
by (wp transferCapsToSlots_pres1)
lemma transferCapsToSlots_irqs_masked'[wp]:
"\<lbrace>irqs_masked'\<rbrace> transferCapsToSlots ep diminish buffer n caps slots mi \<lbrace>\<lambda>rv. irqs_masked'\<rbrace>"
"\<lbrace>irqs_masked'\<rbrace> transferCapsToSlots ep buffer n caps slots mi \<lbrace>\<lambda>rv. irqs_masked'\<rbrace>"
by (wp transferCapsToSlots_pres1 irqs_masked_lift)
lemma storeWordUser_vms'[wp]:
@ -1152,7 +1133,7 @@ by (simp add: setExtraBadge_def) wp
lemma transferCapsToSlots_vms[wp]:
"\<lbrace>\<lambda>s. valid_machine_state' s\<rbrace>
transferCapsToSlots ep diminish buffer n caps slots mi
transferCapsToSlots ep buffer n caps slots mi
\<lbrace>\<lambda>_ s. valid_machine_state' s\<rbrace>"
by (wp transferCapsToSlots_pres1)
@ -1164,7 +1145,7 @@ crunch ct_not_inQ[wp]: setExtraBadge "ct_not_inQ"
lemma tcts_ct_not_inQ[wp]:
"\<lbrace>ct_not_inQ\<rbrace>
transferCapsToSlots ep diminish buffer n caps slots mi
transferCapsToSlots ep buffer n caps slots mi
\<lbrace>\<lambda>_. ct_not_inQ\<rbrace>"
by (wp transferCapsToSlots_pres1)
@ -1187,7 +1168,7 @@ lemma transferCapsToSlots_invs[wp]:
\<and> (\<forall>x \<in> set slots. ex_cte_cap_to' x s)
\<and> (\<forall>x \<in> set slots. real_cte_at' x s)
\<and> transferCaps_srcs caps s\<rbrace>
transferCapsToSlots ep diminish buffer n caps slots mi
transferCapsToSlots ep buffer n caps slots mi
\<lbrace>\<lambda>rv. invs'\<rbrace>"
apply (simp add: invs'_def valid_state'_def)
apply (rule hoare_pre)
@ -1252,8 +1233,8 @@ lemma tc_corres:
and case_option \<top> valid_ipc_buffer_ptr' recv_buf
and transferCaps_srcs caps'
and (\<lambda>s. length caps' \<le> msgMaxExtraCaps))
(transfer_caps info caps ep receiver recv_buf diminish)
(transferCaps info' caps' ep receiver recv_buf diminish)"
(transfer_caps info caps ep receiver recv_buf)
(transferCaps info' caps' ep receiver recv_buf)"
apply (simp add: transfer_caps_def transferCaps_def
getThreadCSpaceRoot)
apply (rule corres_assume_pre)
@ -1417,7 +1398,7 @@ lemma sameObjectAs_maskCapRights[simp]:
lemma transferCaps_invs' [wp]:
"\<lbrace>invs' and transferCaps_srcs caps\<rbrace>
transferCaps mi caps ep receiver recv_buf diminish
transferCaps mi caps ep receiver recv_buf
\<lbrace>\<lambda>rv. invs'\<rbrace>"
apply (simp add: transferCaps_def Let_def split_def)
apply (wp get_rs_cte_at' hoare_vcg_const_Ball_lift
@ -1679,8 +1660,8 @@ lemma do_normal_transfer_corres:
and (\<lambda>s. case ep of Some x \<Rightarrow> ep_at' x s | _ \<Rightarrow> True)
and case_option \<top> valid_ipc_buffer_ptr' send_buf
and case_option \<top> valid_ipc_buffer_ptr' recv_buf)
(do_normal_transfer sender send_buf ep badge can_grant receiver recv_buf diminish)
(doNormalTransfer sender send_buf ep badge can_grant receiver recv_buf diminish)"
(do_normal_transfer sender send_buf ep badge can_grant receiver recv_buf)
(doNormalTransfer sender send_buf ep badge can_grant receiver recv_buf)"
apply (simp add: do_normal_transfer_def doNormalTransfer_def)
apply (rule corres_guard_imp)
@ -1737,7 +1718,7 @@ lemmas doNormal_lifts[wp] = typ_at_lifts [OF doNormalTransfer_typ_at']
lemma doNormal_invs'[wp]:
"\<lbrace>tcb_at' sender and tcb_at' receiver and invs'\<rbrace>
doNormalTransfer sender send_buf ep badge
can_grant receiver recv_buf diminish \<lbrace>\<lambda>r. invs'\<rbrace>"
can_grant receiver recv_buf \<lbrace>\<lambda>r. invs'\<rbrace>"
apply (simp add: doNormalTransfer_def)
apply (wp hoare_vcg_const_Ball_lift | simp)+
done
@ -1922,8 +1903,8 @@ lemma dit_corres:
and (\<lambda>s. case ep of Some x \<Rightarrow> ep_at x s | _ \<Rightarrow> True))
(tcb_at' s and tcb_at' r and valid_pspace' and cur_tcb'
and (\<lambda>s. case ep of Some x \<Rightarrow> ep_at' x s | _ \<Rightarrow> True))
(do_ipc_transfer s ep bg grt r dim)
(doIPCTransfer s ep bg grt r dim)"
(do_ipc_transfer s ep bg grt r)
(doIPCTransfer s ep bg grt r)"
apply (simp add: do_ipc_transfer_def doIPCTransfer_def)
apply (rule_tac Q="%receiveBuffer sa. tcb_at s sa \<and> valid_objs sa \<and>
pspace_aligned sa \<and> tcb_at r sa \<and>
@ -2037,7 +2018,7 @@ crunch irqs_masked'[wp]: doIPCTransfer "irqs_masked'"
lemma doIPCTransfer_invs[wp]:
"\<lbrace>invs' and tcb_at' s and tcb_at' r\<rbrace>
doIPCTransfer s ep bg grt r dim
doIPCTransfer s ep bg grt r
\<lbrace>\<lambda>rv. invs'\<rbrace>" (is "valid ?P _ _")
apply (simp add: doIPCTransfer_def)
apply (wp |wpc)+
@ -2101,7 +2082,7 @@ crunch valid_pspace'[wp]: attemptSwitchTo valid_pspace'
crunch ct'[wp]: handleFaultReply "\<lambda>s. P (ksCurThread s)"
lemma doIPCTransfer_sch_act_simple [wp]:
"\<lbrace>sch_act_simple\<rbrace> doIPCTransfer sender endpoint badge grant receiver diminish \<lbrace>\<lambda>_. sch_act_simple\<rbrace>"
"\<lbrace>sch_act_simple\<rbrace> doIPCTransfer sender endpoint badge grant receiver \<lbrace>\<lambda>_. sch_act_simple\<rbrace>"
by (simp add: sch_act_simple_def, wp)
lemma possibleSwitchTo_invs'[wp]:
@ -2715,10 +2696,10 @@ proof -
apply (clarsimp split del: split_if)
apply (rule corres_guard_imp)
apply (rule corres_split [OF _ set_ep_corres])
apply (simp add: isReceive_def)
apply (simp add: isReceive_def split del:split_if)
apply (rule corres_split [OF _ gts_corres])
apply (rule_tac
F="\<exists>x. recv_state = Structures_A.BlockedOnReceive ep x"
F="recv_state = Structures_A.BlockedOnReceive ep"
in corres_gen_asm)
apply (clarsimp simp: case_bool_If case_option_If if3_fold
simp del: dc_simp split del: split_if cong: if_cong)
@ -2799,7 +2780,7 @@ proof -
apply (rule corres_split [OF _ set_ep_corres])
apply (rule corres_split [OF _ gts_corres])
apply (rule_tac
F="\<exists>x. recv_state = Structures_A.BlockedOnReceive ep x"
F="recv_state = Structures_A.BlockedOnReceive ep"
in corres_gen_asm)
apply (clarsimp simp: isReceive_def case_bool_If
split del: split_if cong: if_cong)

View File

@ -144,7 +144,7 @@ crunch valid_duplicates'[wp]: getReceiveSlots "\<lambda>s. vs_valid_duplicates'
lemma transferCapsToSlots_duplicates'[wp]:
"\<lbrace>\<lambda>s. vs_valid_duplicates' (ksPSpace s)\<rbrace>
transferCapsToSlots ep diminish buffer n caps slots mi
transferCapsToSlots ep buffer n caps slots mi
\<lbrace>\<lambda>rv s. vs_valid_duplicates' (ksPSpace s)\<rbrace>"
by (rule transferCapsToSlots_pres1,wp)
@ -2519,7 +2519,7 @@ lemma handleInterrupt_valid_duplicates'[wp]:
"\<lbrace>\<lambda>s. vs_valid_duplicates' (ksPSpace s)\<rbrace>
handleInterrupt irq \<lbrace>\<lambda>r s. vs_valid_duplicates' (ksPSpace s)\<rbrace>"
apply (simp add: handleInterrupt_def)
apply (rule hoare_pre)
apply (rule conjI; rule impI)
apply (wp sai_st_tcb' hoare_vcg_all_lift hoare_drop_imps
threadSet_pred_tcb_no_state getIRQState_inv haskell_fail_wp
|wpc|simp)+

View File

@ -157,8 +157,8 @@ where
= (ts' = Structures_H.IdleThreadState)"
| "thread_state_relation (Structures_A.BlockedOnReply) ts'
= (ts' = Structures_H.BlockedOnReply)"
| "thread_state_relation (Structures_A.BlockedOnReceive oref dimin) ts'
= (ts' = Structures_H.BlockedOnReceive oref dimin)"
| "thread_state_relation (Structures_A.BlockedOnReceive oref) ts'
= (ts' = Structures_H.BlockedOnReceive oref)"
| "thread_state_relation (Structures_A.BlockedOnSend oref sp) ts'
= (ts' = Structures_H.BlockedOnSend oref (sender_badge sp)
(sender_can_grant sp) (sender_is_call sp))"

View File

@ -773,7 +773,7 @@ crunch sch_act_simple[wp]: handleFaultReply sch_act_simple
lemma transferCaps_non_null_cte_wp_at':
assumes PUC: "\<And>cap. P cap \<Longrightarrow> \<not> isUntypedCap cap"
shows "\<lbrace>cte_wp_at' (\<lambda>cte. P (cteCap cte) \<and> cteCap cte \<noteq> capability.NullCap) ptr\<rbrace>
transferCaps info caps ep rcvr rcvBuf dim
transferCaps info caps ep rcvr rcvBuf
\<lbrace>\<lambda>_. cte_wp_at' (\<lambda>cte. P (cteCap cte) \<and> cteCap cte \<noteq> capability.NullCap) ptr\<rbrace>"
proof -
have CTEF: "\<And>P p s. \<lbrakk> cte_wp_at' P p s; \<And>cte. P cte \<Longrightarrow> False \<rbrakk> \<Longrightarrow> False"
@ -801,7 +801,7 @@ lemma doNormalTransfer_non_null_cte_wp_at':
assumes PUC: "\<And>cap. P cap \<Longrightarrow> \<not> isUntypedCap cap"
shows
"\<lbrace>cte_wp_at' (\<lambda>cte. P (cteCap cte) \<and> cteCap cte \<noteq> capability.NullCap) ptr\<rbrace>
doNormalTransfer st send_buffer ep b gr rt recv_buffer dim
doNormalTransfer st send_buffer ep b gr rt recv_buffer
\<lbrace>\<lambda>_. cte_wp_at' (\<lambda>cte. P (cteCap cte) \<and> cteCap cte \<noteq> capability.NullCap) ptr\<rbrace>"
unfolding doNormalTransfer_def
apply (wp transferCaps_non_null_cte_wp_at' | simp add:PUC)+
@ -823,7 +823,7 @@ lemma doIPCTransfer_non_null_cte_wp_at':
assumes PUC: "\<And>cap. P cap \<Longrightarrow> \<not> isUntypedCap cap"
shows
"\<lbrace>cte_wp_at' (\<lambda>cte. P (cteCap cte) \<and> cteCap cte \<noteq> capability.NullCap) ptr\<rbrace>
doIPCTransfer sender endpoint badge grant receiver dim
doIPCTransfer sender endpoint badge grant receiver
\<lbrace>\<lambda>_. cte_wp_at' (\<lambda>cte. P (cteCap cte) \<and> cteCap cte \<noteq> capability.NullCap) ptr\<rbrace>"
unfolding doIPCTransfer_def
apply (wp doNormalTransfer_non_null_cte_wp_at' hoare_drop_imp hoare_allI | wpc | clarsimp simp:PUC)+
@ -834,7 +834,7 @@ lemma doIPCTransfer_non_null_cte_wp_at2':
assumes PNN: "\<And>cte. P (cteCap cte) \<Longrightarrow> cteCap cte \<noteq> capability.NullCap"
and PUC: "\<And>cap. P cap \<Longrightarrow> \<not> isUntypedCap cap"
shows "\<lbrace>cte_wp_at' (\<lambda>cte. P (cteCap cte)) ptr\<rbrace>
doIPCTransfer sender endpoint badge grant receiver dim
doIPCTransfer sender endpoint badge grant receiver
\<lbrace>\<lambda>_. cte_wp_at' (\<lambda>cte. P (cteCap cte)) ptr\<rbrace>"
proof -
have PimpQ: "\<And>P Q ptr s. \<lbrakk> cte_wp_at' (\<lambda>cte. P (cteCap cte)) ptr s;

View File

@ -23,15 +23,15 @@
<sequence depends="ASpec ExecSpec">
<test name="AInvs" cpu-timeout="7200">make AInvs</test>
<test name="BaseRefine">make BaseRefine</test>
<test name="Refine" cpu-timeout="10800">make Refine</test>
<test name="Refine" cpu-timeout="14400">make Refine</test>
</sequence>
<!-- CRefine -->
<sequence depends="CParser">
<test name="CKernel" cpu-timeout="7200">make CKernel</test>
<test name="CSpec">make CSpec</test>
<test name="CBaseRefine" depends="Refine" cpu-timeout="21600">make CBaseRefine</test>
<test name="CRefine" cpu-timeout="25200">make CRefine</test>
<test name="CSpec" cpu-timeout="7200">make CSpec</test>
<test name="CBaseRefine" depends="Refine" cpu-timeout="28800">make CBaseRefine</test>
<test name="CRefine" cpu-timeout="28800">make CRefine</test>
</sequence>
<!-- DSpec / DRefine -->
@ -44,7 +44,7 @@
<sequence depends="AInvs">
<test name="Access">make Access</test>
<test name="InfoFlow" cpu-timeout="7200">make InfoFlow</test>
<test name="InfoFlowC" depends="CRefine" cpu-timeout="16200">make InfoFlowC</test>
<test name="InfoFlowC" depends="CRefine" cpu-timeout="14400">make InfoFlowC</test>
</sequence>
<!-- Access Control and capDL -->
@ -59,7 +59,7 @@
<!-- Binary Verification -->
<sequence depends="CSpec">
<test name="SimplExportAndRefine" cpu-timeout="14400">make SimplExportAndRefine</test>
<test name="SimplExportAndRefine" cpu-timeout="21600">make SimplExportAndRefine</test>
</sequence>
<!-- Separation Logic. -->

View File

@ -31,7 +31,7 @@ chapter "Specifications"
*)
session ASpec in "abstract" = Word +
options [document=pdf, timeout=600]
options [document=pdf]
theories [document = false]
"../../lib/NICTACompat"
"../../lib/Lib"
@ -63,7 +63,7 @@ session ASpec in "abstract" = Word +
*)
session ExecSpec = Word +
options [document = false, timeout=600]
options [document = false]
theories
"design/API_H"
"design/$L4V_ARCH/Intermediate_H"
@ -118,7 +118,7 @@ session TakeGrant in "take-grant" = "HOL-Word" +
*)
session ASepSpec = ASpec +
options [document = false, timeout=600]
options [document = false]
theories
"sep-abstract/Syscall_SA"

View File

@ -82,22 +82,27 @@ definition timer_tick :: "unit det_ext_monad" where
definition
handle_interrupt :: "irq \<Rightarrow> (unit,'z::state_ext) s_monad" where
"handle_interrupt irq \<equiv> do
st \<leftarrow> get_irq_state irq;
case st of
IRQSignal \<Rightarrow> do
slot \<leftarrow> get_irq_slot irq;
cap \<leftarrow> get_cap slot;
when (is_ntfn_cap cap \<and> AllowSend \<in> cap_rights cap)
$ send_signal (obj_ref_of cap) (cap_ep_badge cap);
do_machine_op $ maskInterrupt True irq
"handle_interrupt irq \<equiv>
if (irq > maxIRQ) then do_machine_op $ do
maskInterrupt True irq;
ackInterrupt irq
od
| IRQTimer \<Rightarrow> do
do_extended_op timer_tick;
do_machine_op resetTimer
od
| IRQInactive \<Rightarrow> fail (* not meant to be able to get IRQs from inactive lines *);
do_machine_op $ ackInterrupt irq
od"
else do
st \<leftarrow> get_irq_state irq;
case st of
IRQSignal \<Rightarrow> do
slot \<leftarrow> get_irq_slot irq;
cap \<leftarrow> get_cap slot;
when (is_ntfn_cap cap \<and> AllowSend \<in> cap_rights cap)
$ send_signal (obj_ref_of cap) (cap_ep_badge cap);
do_machine_op $ maskInterrupt True irq
od
| IRQTimer \<Rightarrow> do
do_extended_op timer_tick;
do_machine_op resetTimer
od
| IRQInactive \<Rightarrow> fail (* not meant to be able to get IRQs from inactive lines *);
do_machine_op $ ackInterrupt irq
od"
end

View File

@ -154,7 +154,7 @@ definition
get_blocking_object :: "thread_state \<Rightarrow> (obj_ref,'z::state_ext) s_monad"
where
"get_blocking_object state \<equiv>
case state of BlockedOnReceive epptr d \<Rightarrow> return epptr
case state of BlockedOnReceive epptr \<Rightarrow> return epptr
| BlockedOnSend epptr x \<Rightarrow> return epptr
| _ \<Rightarrow> fail"
@ -311,7 +311,7 @@ where
case state
of
BlockedOnSend x y \<Rightarrow> blocked_cancel_ipc state tptr
| BlockedOnReceive x d \<Rightarrow> blocked_cancel_ipc state tptr
| BlockedOnReceive x \<Rightarrow> blocked_cancel_ipc state tptr
| BlockedOnNotification event \<Rightarrow> cancel_signal tptr event
| BlockedOnReply \<Rightarrow> reply_cancel_ipc tptr
| _ \<Rightarrow> return ()

View File

@ -83,18 +83,18 @@ where
type_synonym transfer_caps_data = "(cap \<times> cslot_ptr) list \<times> cslot_ptr list"
fun
transfer_caps_loop :: "obj_ref option \<Rightarrow> bool \<Rightarrow> obj_ref \<Rightarrow> nat
transfer_caps_loop :: "obj_ref option \<Rightarrow> obj_ref \<Rightarrow> nat
\<Rightarrow> (cap \<times> cslot_ptr) list \<Rightarrow> cslot_ptr list
\<Rightarrow> message_info \<Rightarrow> (message_info,'z::state_ext) s_monad"
where
"transfer_caps_loop ep diminish rcv_buffer n [] slots
"transfer_caps_loop ep rcv_buffer n [] slots
mi = return (MI (mi_length mi) (of_nat n) (mi_caps_unwrapped mi)
(mi_label mi))"
| "transfer_caps_loop ep diminish rcv_buffer n ((cap, slot) # morecaps)
| "transfer_caps_loop ep rcv_buffer n ((cap, slot) # morecaps)
slots mi =
const_on_failure (MI (mi_length mi) (of_nat n) (mi_caps_unwrapped mi)
(mi_label mi)) (doE
transfer_rest \<leftarrow> returnOk $ transfer_caps_loop ep diminish
transfer_rest \<leftarrow> returnOk $ transfer_caps_loop ep
rcv_buffer (n + 1) morecaps;
if (is_ep_cap cap \<and> ep = Some (obj_ref_of cap))
then doE
@ -104,7 +104,7 @@ where
odE
else if slots \<noteq> []
then doE
cap' \<leftarrow> derive_cap slot (if diminish then remove_rights {AllowWrite} cap else cap);
cap' \<leftarrow> derive_cap slot cap;
whenE (cap' = NullCap) $ throwError undefined;
liftE $ cap_insert cap' slot (hd slots);
liftE $ transfer_rest (tl slots) mi
@ -115,16 +115,16 @@ where
definition
transfer_caps :: "message_info \<Rightarrow> (cap \<times> cslot_ptr) list \<Rightarrow>
obj_ref option \<Rightarrow> obj_ref \<Rightarrow> obj_ref option \<Rightarrow> bool \<Rightarrow>
obj_ref option \<Rightarrow> obj_ref \<Rightarrow> obj_ref option \<Rightarrow>
(message_info,'z::state_ext) s_monad"
where
"transfer_caps info caps endpoint receiver recv_buffer diminish \<equiv> do
"transfer_caps info caps endpoint receiver recv_buffer \<equiv> do
dest_slots \<leftarrow> get_receive_slots receiver recv_buffer;
mi' \<leftarrow> return $ MI (mi_length info) 0 0 (mi_label info);
case recv_buffer of
None \<Rightarrow> return mi'
| Some receive_buffer \<Rightarrow>
transfer_caps_loop endpoint diminish receive_buffer 0 caps dest_slots mi'
transfer_caps_loop endpoint receive_buffer 0 caps dest_slots mi'
od"
section {* Fault Handling *}
@ -202,16 +202,16 @@ definition
do_normal_transfer :: "obj_ref \<Rightarrow> obj_ref option \<Rightarrow> obj_ref option
\<Rightarrow> data \<Rightarrow> bool \<Rightarrow> obj_ref
\<Rightarrow> obj_ref option
\<Rightarrow> bool \<Rightarrow> (unit,'z::state_ext) s_monad"
\<Rightarrow> (unit,'z::state_ext) s_monad"
where
"do_normal_transfer sender sbuf endpoint badge grant
receiver rbuf diminish \<equiv>
receiver rbuf \<equiv>
do
mi \<leftarrow> get_message_info sender;
caps \<leftarrow> if grant then lookup_extra_caps sender sbuf mi <catch> K (return [])
else return [];
mrs_transferred \<leftarrow> copy_mrs sender sbuf receiver rbuf (mi_length mi);
mi' \<leftarrow> transfer_caps mi caps endpoint receiver rbuf diminish;
mi' \<leftarrow> transfer_caps mi caps endpoint receiver rbuf;
set_message_info receiver $ MI mrs_transferred (mi_extra_caps mi')
(mi_caps_unwrapped mi') (mi_label mi);
as_user receiver $ set_register badge_register badge
@ -220,10 +220,10 @@ where
text {* Transfer a message either involving a fault or not. *}
definition
do_ipc_transfer :: "obj_ref \<Rightarrow> obj_ref option \<Rightarrow>
badge \<Rightarrow> bool \<Rightarrow> obj_ref \<Rightarrow> bool \<Rightarrow> (unit,'z::state_ext) s_monad"
badge \<Rightarrow> bool \<Rightarrow> obj_ref \<Rightarrow> (unit,'z::state_ext) s_monad"
where
"do_ipc_transfer sender ep badge grant
receiver diminish \<equiv> do
receiver \<equiv> do
recv_buffer \<leftarrow> lookup_ipc_buffer True receiver;
fault \<leftarrow> thread_get tcb_fault sender;
@ -232,7 +232,7 @@ where
of None \<Rightarrow> do
send_buffer \<leftarrow> lookup_ipc_buffer False sender;
do_normal_transfer sender send_buffer ep badge grant
receiver recv_buffer diminish
receiver recv_buffer
od
| Some f \<Rightarrow> do_fault_transfer badge sender receiver recv_buffer
od"
@ -247,7 +247,7 @@ where
fault \<leftarrow> thread_get tcb_fault receiver;
case fault of
None \<Rightarrow> do
do_ipc_transfer sender None 0 True receiver False;
do_ipc_transfer sender None 0 True receiver;
cap_delete_one slot;
set_thread_state receiver Running;
do_extended_op (attempt_switch_to receiver)
@ -315,22 +315,22 @@ where
| (IdleEP, False) \<Rightarrow> return ()
| (SendEP queue, False) \<Rightarrow> return ()
| (RecvEP (dest # queue), _) \<Rightarrow> do
set_endpoint epptr $ (case queue of [] \<Rightarrow> IdleEP
| _ \<Rightarrow> RecvEP queue);
recv_state \<leftarrow> get_thread_state dest;
diminish \<leftarrow> (case recv_state
of BlockedOnReceive x d \<Rightarrow> return d
| _ \<Rightarrow> fail);
do_ipc_transfer thread (Some epptr) badge
can_grant dest diminish;
set_thread_state dest Running;
do_extended_op (attempt_switch_to dest);
fault \<leftarrow> thread_get tcb_fault thread;
when (call \<or> fault \<noteq> None) $
if can_grant \<and> \<not> diminish
then setup_caller_cap thread dest
else set_thread_state thread Inactive
od
set_endpoint epptr $ (case queue of [] \<Rightarrow> IdleEP
| _ \<Rightarrow> RecvEP queue);
recv_state \<leftarrow> get_thread_state dest;
case recv_state
of (BlockedOnReceive x) \<Rightarrow>
do_ipc_transfer thread (Some epptr) badge
can_grant dest
| _ \<Rightarrow> fail;
set_thread_state dest Running;
do_extended_op (attempt_switch_to dest);
fault \<leftarrow> thread_get tcb_fault thread;
when (call \<or> fault \<noteq> None) $
if can_grant
then setup_caller_cap thread dest
else set_thread_state thread Inactive
od
| (RecvEP [], _) \<Rightarrow> fail
od"
@ -372,7 +372,6 @@ where
(epptr,rights) \<leftarrow> (case cap
of EndpointCap ref badge rights \<Rightarrow> return (ref,rights)
| _ \<Rightarrow> fail);
diminish \<leftarrow> return (AllowSend \<notin> rights);
ep \<leftarrow> get_endpoint epptr;
ntfnptr \<leftarrow> get_bound_notification thread;
ntfn \<leftarrow> case_option (return default_notification) get_notification ntfnptr;
@ -383,13 +382,13 @@ where
case ep
of IdleEP \<Rightarrow> case is_blocking of
True \<Rightarrow> do
set_thread_state thread (BlockedOnReceive epptr diminish);
set_thread_state thread (BlockedOnReceive epptr);
set_endpoint epptr (RecvEP [thread])
od
| False \<Rightarrow> do_nbrecv_failed_transfer thread
| RecvEP queue \<Rightarrow> case is_blocking of
True \<Rightarrow> do
set_thread_state thread (BlockedOnReceive epptr diminish);
set_thread_state thread (BlockedOnReceive epptr);
set_endpoint epptr (RecvEP (queue @ [thread]))
od
| False \<Rightarrow> do_nbrecv_failed_transfer thread
@ -405,11 +404,11 @@ where
| _ \<Rightarrow> fail);
do_ipc_transfer sender (Some epptr)
(sender_badge data) (sender_can_grant data)
thread diminish;
thread;
fault \<leftarrow> thread_get tcb_fault sender;
if ((sender_is_call data) \<or> (fault \<noteq> None))
then
if sender_can_grant data \<and> \<not> diminish
if sender_can_grant data
then setup_caller_cap sender thread
else set_thread_state sender Inactive
else do
@ -448,7 +447,7 @@ definition
receive_blocked :: "thread_state \<Rightarrow> bool"
where
"receive_blocked st \<equiv> case st of
BlockedOnReceive _ _ \<Rightarrow> True
BlockedOnReceive _ \<Rightarrow> True
| _ \<Rightarrow> False"
definition

View File

@ -341,7 +341,7 @@ datatype thread_state
= Running
| Inactive
| Restart
| BlockedOnReceive obj_ref bool
| BlockedOnReceive obj_ref
| BlockedOnSend obj_ref sender_payload
| BlockedOnReply
| BlockedOnNotification obj_ref
@ -368,7 +368,7 @@ where
"runnable (Running) = True"
| "runnable (Inactive) = False"
| "runnable (Restart) = True"
| "runnable (BlockedOnReceive x y) = False"
| "runnable (BlockedOnReceive x) = False"
| "runnable (BlockedOnSend x y) = False"
| "runnable (BlockedOnNotification x) = False"
| "runnable (IdleThreadState) = False"

View File

@ -13,7 +13,7 @@
*)
theory Interrupt_D
imports Endpoint_D
imports Endpoint_D "../machine/$L4V_ARCH/Platform"
begin
(* Return the currently pending IRQ. *)
@ -113,7 +113,7 @@ where
definition
handle_interrupt :: "cdl_irq \<Rightarrow> unit k_monad"
where
"handle_interrupt irq \<equiv>
"handle_interrupt irq \<equiv> if (irq > maxIRQ) then return () else
do
irq_slot \<leftarrow> gets $ get_irq_slot irq;
c \<leftarrow> gets $ opt_cap irq_slot;

View File

@ -12,6 +12,7 @@ theory KernelState_C
imports
"../../lib/WordLib"
"../../lib/TypHeapLib"
"../../lib/BitFieldProofsLib"
Kernel_C
Substitute
begin

View File

@ -251,7 +251,7 @@ od))
state \<leftarrow> getThreadState tptr;
(case state of
BlockedOnSend _ _ _ _ \<Rightarrow> blockedIPCCancel state
| BlockedOnReceive _ _ \<Rightarrow> blockedIPCCancel state
| BlockedOnReceive _ \<Rightarrow> blockedIPCCancel state
| BlockedOnNotification _ \<Rightarrow> cancelSignal tptr (waitingOnNotification state)
| BlockedOnReply \<Rightarrow> replyIPCCancel
| _ \<Rightarrow> return ()

View File

@ -50,13 +50,12 @@ defs sendIPC_def:
recvState \<leftarrow> getThreadState dest;
haskell_assert (isReceive recvState)
[];
diminish \<leftarrow> return ( blockingIPCDiminishCaps recvState);
doIPCTransfer thread (Just epptr) badge canGrant
dest diminish;
dest;
setThreadState Running dest;
attemptSwitchTo dest;
fault \<leftarrow> threadGet tcbFault thread;
(case (call, fault, canGrant \<and> Not diminish) of
(case (call, fault, canGrant) of
(False, None, _) \<Rightarrow> return ()
| (_, _, True) \<Rightarrow> setupCallerCap thread dest
| _ \<Rightarrow> setThreadState Inactive thread
@ -79,7 +78,6 @@ defs receiveIPC_def:
then (do
epptr \<leftarrow> return ( capEPPtr cap);
ep \<leftarrow> getEndpoint epptr;
diminish \<leftarrow> return ( Not $ capEPCanSend cap);
ntfnPtr \<leftarrow> getBoundNotification thread;
ntfn \<leftarrow> maybe (return $ NTFN IdleNtfn Nothing) (getNotification) ntfnPtr;
if (isJust ntfnPtr \<and> isActive ntfn)
@ -88,8 +86,7 @@ defs receiveIPC_def:
IdleEP \<Rightarrow> (case isBlocking of
True \<Rightarrow> (do
setThreadState (BlockedOnReceive_ \<lparr>
blockingObject= epptr,
blockingIPCDiminishCaps= diminish \<rparr>) thread;
blockingObject= epptr \<rparr>) thread;
setEndpoint epptr $ RecvEP [thread]
od)
| False \<Rightarrow> doNBRecvFailedTransfer thread
@ -97,8 +94,7 @@ defs receiveIPC_def:
| RecvEP queue \<Rightarrow> (case isBlocking of
True \<Rightarrow> (do
setThreadState (BlockedOnReceive_ \<lparr>
blockingObject= epptr,
blockingIPCDiminishCaps= diminish \<rparr>) thread;
blockingObject= epptr \<rparr>) thread;
setEndpoint epptr $ RecvEP $ queue @ [thread]
od)
| False \<Rightarrow> doNBRecvFailedTransfer thread
@ -114,10 +110,10 @@ defs receiveIPC_def:
badge \<leftarrow> return ( blockingIPCBadge senderState);
canGrant \<leftarrow> return ( blockingIPCCanGrant senderState);
doIPCTransfer sender (Just epptr) badge canGrant
thread diminish;
thread;
call \<leftarrow> return ( blockingIPCIsCall senderState);
fault \<leftarrow> threadGet tcbFault sender;
(case (call, fault, canGrant \<and> Not diminish) of
(case (call, fault, canGrant) of
(False, None, _) \<Rightarrow> (do
setThreadState Running sender;
switchIfRequiredTo sender
@ -183,7 +179,7 @@ defs cancelIPC_def:
state \<leftarrow> getThreadState tptr;
(case state of
BlockedOnSend _ _ _ _ \<Rightarrow> blockedIPCCancel state
| BlockedOnReceive _ _ \<Rightarrow> blockedIPCCancel state
| BlockedOnReceive _ \<Rightarrow> blockedIPCCancel state
| BlockedOnNotification _ \<Rightarrow> cancelSignal tptr (waitingOnNotification state)
| BlockedOnReply \<Rightarrow> replyIPCCancel
| _ \<Rightarrow> return ()

View File

@ -128,28 +128,35 @@ defs initInterruptController_def:
odE)"
defs handleInterrupt_def:
"handleInterrupt irq\<equiv> (do
st \<leftarrow> getIRQState irq;
(case st of
IRQSignal \<Rightarrow> (do
slot \<leftarrow> getIRQSlot irq;
cap \<leftarrow> getSlotCap slot;
(case cap of
NotificationCap _ _ True _ \<Rightarrow>
sendSignal (capNtfnPtr cap) (capNtfnBadge cap)
| _ \<Rightarrow> doMachineOp $ debugPrint $
[] @ show irq
);
doMachineOp $ maskInterrupt True irq
"handleInterrupt irq\<equiv> (
if (irq > maxIRQ) then doMachineOp $ ((do
maskInterrupt True irq;
ackInterrupt irq
od)
)
else (do
st \<leftarrow> getIRQState irq;
(case st of
IRQSignal \<Rightarrow> (do
slot \<leftarrow> getIRQSlot irq;
cap \<leftarrow> getSlotCap slot;
(case cap of
NotificationCap _ _ True _ \<Rightarrow>
sendSignal (capNtfnPtr cap) (capNtfnBadge cap)
| _ \<Rightarrow> doMachineOp $ debugPrint $
[] @ show irq
);
doMachineOp $ maskInterrupt True irq
od)
| IRQTimer \<Rightarrow> (do
timerTick;
doMachineOp resetTimer
od)
| IRQTimer \<Rightarrow> (do
timerTick;
doMachineOp resetTimer
od)
| IRQInactive \<Rightarrow> haskell_fail $ [] @ show irq
);
doMachineOp $ ackInterrupt irq
od)"
| IRQInactive \<Rightarrow> haskell_fail $ [] @ show irq
);
doMachineOp $ ackInterrupt irq
od)
)"
defs isIRQActive_def:
"isIRQActive irq\<equiv> liftM (\<lambda>x. x \<noteq>IRQInactive) $ getIRQState irq"

Some files were not shown because too many files have changed in this diff Show More