Isabelle2018: new comment syntax

(result of "isabelle update_comments <dirs>")
This commit is contained in:
Gerwin Klein 2018-06-09 18:27:30 +10:00
parent b1aa74d306
commit 011e08458e
171 changed files with 2012 additions and 2001 deletions

View File

@ -75,12 +75,12 @@ text {*
*}
datatype number =
-- "High level types"
\<comment> \<open>High level types\<close>
UnsignedInteger
| Integer
| Real
| Boolean
-- "C-specific types"
\<comment> \<open>C-specific types\<close>
| uint8_t
| uint16_t
| uint32_t
@ -94,10 +94,10 @@ datatype number =
| uintptr_t
datatype textual =
-- "High level types"
\<comment> \<open>High level types\<close>
Character
| String
-- "C-specific types"
\<comment> \<open>C-specific types\<close>
| char
datatype primitive =
@ -171,8 +171,8 @@ text {*
Finally, @{text dataport}, is used to model shared memory communication.
*}
type_synonym procedure = "method list"
type_synonym event = nat -- "ID"
type_synonym dataport = "string option" -- "type"
type_synonym event = nat \<comment> \<open>ID\<close>
type_synonym dataport = "string option" \<comment> \<open>type\<close>
datatype interface =
Procedure procedure
| Event event
@ -191,9 +191,9 @@ text {*
data.
*}
datatype native_connector_type =
AsynchronousEvent -- "an asynchronous notification"
| RPC -- "a synchronous channel"
| SharedData -- "a shared memory region"
AsynchronousEvent \<comment> \<open>an asynchronous notification\<close>
| RPC \<comment> \<open>a synchronous channel\<close>
| SharedData \<comment> \<open>a shared memory region\<close>
text {*
Recalling that hardware devices are modelled as components in CAmkES, hardware
@ -202,9 +202,9 @@ text {*
corresponds to the mode of interaction with the device.
*}
datatype hardware_connector_type =
HardwareMMIO -- "memory mapped IO"
| HardwareInterrupt -- "device interrupts"
| HardwareIOPort -- "IA32 IO ports"
HardwareMMIO \<comment> \<open>memory mapped IO\<close>
| HardwareInterrupt \<comment> \<open>device interrupts\<close>
| HardwareIOPort \<comment> \<open>IA32 IO ports\<close>
text {*
Export connectors are used when

View File

@ -57,22 +57,22 @@ text {*
*}
datatype question_data
-- "Inter-component questions"
\<comment> \<open>Inter-component questions\<close>
= Call nat "variable list"
| Return "variable list"
-- "Questions from components to events"
\<comment> \<open>Questions from components to events\<close>
| Set
| Poll
-- "Questions from components to shared memory"
\<comment> \<open>Questions from components to shared memory\<close>
| Read nat
| Write nat variable
datatype answer_data
-- "Answers from events to components"
\<comment> \<open>Answers from events to components\<close>
= Pending bool
-- "Answers from shared memory to components"
\<comment> \<open>Answers from shared memory to components\<close>
| Value variable
-- "An answer that conveys no information"
\<comment> \<open>An answer that conveys no information\<close>
| Void
record ('channel) question =

View File

@ -175,8 +175,8 @@ notepad begin
fix D C
assume DC:"D \<Longrightarrow> C"
have "D \<and> D \<Longrightarrow> C \<and> C"
apply (only_asm \<open>simp\<close>) -- "stash conclusion before applying method"
apply (only_concl \<open>simp add: DC\<close>) -- "hide premises from method"
apply (only_asm \<open>simp\<close>) \<comment> \<open>stash conclusion before applying method\<close>
apply (only_concl \<open>simp add: DC\<close>) \<comment> \<open>hide premises from method\<close>
by (rule DC)
end
@ -281,12 +281,12 @@ notepad begin
by (rule A)+
have "A \<and> A" "A \<and> A" "B"
apply (find_goal \<open>fails \<open>simp\<close>\<close>) -- "find the first goal which cannot be simplified"
apply (find_goal \<open>fails \<open>simp\<close>\<close>) \<comment> \<open>find the first goal which cannot be simplified\<close>
apply (rule B)
by (simp add: A)+
have "B" "A" "A \<and> A"
apply (find_goal \<open>succeeds \<open>simp\<close>\<close>) -- "find the first goal which can be simplified (without doing so)"
apply (find_goal \<open>succeeds \<open>simp\<close>\<close>) \<comment> \<open>find the first goal which can be simplified (without doing so)\<close>
apply (rule conjI)
by (rule A B)+
@ -373,7 +373,7 @@ notepad begin
by (rule A)
have "B \<Longrightarrow> A" "A"
by (distinct_subgoals_strong \<open>assumption\<close>, rule A) -- "backtracking required here"
by (distinct_subgoals_strong \<open>assumption\<close>, rule A) \<comment> \<open>backtracking required here\<close>
}
{
@ -382,7 +382,7 @@ notepad begin
assume B: "B"
assume BC: "B \<Longrightarrow> C" "B \<Longrightarrow> A"
have "A" "B \<longrightarrow> (A \<and> C)" "B"
apply (distinct_subgoals_strong \<open>simp\<close>, rule B) -- "backtracking required here"
apply (distinct_subgoals_strong \<open>simp\<close>, rule B) \<comment> \<open>backtracking required here\<close>
by (simp add: BC)
}

View File

@ -192,7 +192,7 @@ always discards the effect of the method it is given.
\<close>
lemma "\<lbrakk> A; B; \<lbrakk> A; B; D; E \<rbrakk> \<Longrightarrow> C; D; E \<rbrakk> \<Longrightarrow> A \<and> ((B \<and> C) \<and> D) \<and> E"
apply (lift_conjunct \<open>match conclusion in C \<Rightarrow> \<open>-\<close>\<close>)
-- \<open>@{term C} as been moved to the front of the conclusion.\<close>
\<comment> \<open>@{term C} as been moved to the front of the conclusion.\<close>
apply (match conclusion in \<open>C \<and> A \<and> (B \<and> D) \<and> E\<close> \<Rightarrow> \<open>-\<close>)
oops
@ -203,10 +203,10 @@ is discarded, so the body of the @{method match} is irrelevant.
\<close>
lemma "\<lbrakk> A; B; \<lbrakk> A; B; D; E \<rbrakk> \<Longrightarrow> C; D; E \<rbrakk> \<Longrightarrow> A \<and> ((B \<and> C) \<and> D) \<and> E"
apply (extract_conjunct \<open>match conclusion in C \<Rightarrow> \<open>-\<close>\<close>)
-- \<open>@{method extract_conjunct} gives us the matched conjunct @{term C} as a separate subgoal.\<close>
\<comment> \<open>@{method extract_conjunct} gives us the matched conjunct @{term C} as a separate subgoal.\<close>
apply (match conclusion in C \<Rightarrow> \<open>-\<close>)
apply blast
-- \<open>The other subgoal contains the remaining conjuncts untouched.\<close>
\<comment> \<open>The other subgoal contains the remaining conjuncts untouched.\<close>
apply (match conclusion in \<open>A \<and> (B \<and> D) \<and> E\<close> \<Rightarrow> \<open>-\<close>)
oops
@ -216,12 +216,12 @@ to the extracted subgoal.
\<close>
lemma "\<lbrakk> A; B; \<lbrakk> A; B; D; E \<rbrakk> \<Longrightarrow> C; D; E \<rbrakk> \<Longrightarrow> A \<and> ((B \<and> C) \<and> D) \<and> E"
apply (apply_conjunct \<open>match conclusion in C \<Rightarrow> \<open>match premises in H: _ \<Rightarrow> \<open>rule H\<close>\<close>\<close>)
-- \<open>We get four subgoals from applying the given method to the matched conjunct @{term C}.\<close>
\<comment> \<open>We get four subgoals from applying the given method to the matched conjunct @{term C}.\<close>
apply (match premises in H: A \<Rightarrow> \<open>rule H\<close>)
apply (match premises in H: B \<Rightarrow> \<open>rule H\<close>)
apply (match premises in H: D \<Rightarrow> \<open>rule H\<close>)
apply (match premises in H: E \<Rightarrow> \<open>rule H\<close>)
-- \<open>The last subgoal contains the remaining conjuncts untouched.\<close>
\<comment> \<open>The last subgoal contains the remaining conjuncts untouched.\<close>
apply (match conclusion in \<open>A \<and> (B \<and> D) \<and> E\<close> \<Rightarrow> \<open>-\<close>)
oops

View File

@ -164,9 +164,9 @@ text \<open>
lemmas lexord_list_simps =
simp_thms
lexord_list.inject list.distinct list.inject -- "eq"
lexord_not_less_Nil lexord_Nil_less_Cons lexord_Cons_less_Cons -- "less"
lexord_le_Nil lexord_Nil_le_Cons lexord_Cons_le_Cons -- "le"
lexord_list.inject list.distinct list.inject \<comment> \<open>eq\<close>
lexord_not_less_Nil lexord_Nil_less_Cons lexord_Cons_less_Cons \<comment> \<open>less\<close>
lexord_le_Nil lexord_Nil_le_Cons lexord_Cons_le_Cons \<comment> \<open>le\<close>
experiment begin
lemma "lexord_list [1, 2, 3] < lexord_list [1, 3 :: nat]"

View File

@ -92,7 +92,7 @@ lemma hoare_triple_refinement:
by (simp add: refines_def hoare_triple_def) blast
-- "composing two relations"
\<comment> \<open>composing two relations\<close>
definition
rel_semi :: "('a \<times> 'b) set \<Rightarrow> ('b \<times> 'c) set \<Rightarrow> ('a \<times> 'c) set" (infixl ";;;" 65)
where

View File

@ -99,7 +99,7 @@ experiment begin
"HOL simp": \<open>simp\<close>
"l4v simp": \<open>simp cong: if_cong cong del: if_weak_cong\<close>
-- "exponential time!"
\<comment> \<open>exponential time!\<close>
)+
done

View File

@ -138,7 +138,7 @@ lemma ccorres_if_cond_throws:
assumes abs: "\<forall>s s'. (s, s') \<in> sr \<and> Q s \<and> Q' s' \<longrightarrow> P = (s' \<in> P')"
and ac: "P \<Longrightarrow> ccorres_underlying sr \<Gamma> r xf arrel axf R R' (SKIP # hs) a c"
and bd: "\<not> P \<Longrightarrow> ccorres_underlying sr \<Gamma> r xf arrel axf U U' (SKIP # hs) b d"
and cthrows: "\<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> PT' c {}, UNIV" -- "c always throws"
and cthrows: "\<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> PT' c {}, UNIV" \<comment> \<open>c always throws\<close>
shows "ccorres_underlying sr \<Gamma> r xf arrel axf
(Q and (\<lambda>s. P \<longrightarrow> R s) and (\<lambda>s. \<not> P \<longrightarrow> U s))
(Collect Q' \<inter> {s. (s \<in> P' \<longrightarrow> s \<in> R' \<inter> PT') \<and> (s \<notin> P' \<longrightarrow> s \<in> U')})
@ -184,7 +184,7 @@ lemma ccorres_if_cond_throws2:
assumes abs: "\<forall>s s'. (s, s') \<in> sr \<and> Q s \<and> Q' s' \<longrightarrow> (\<not> P) = (s' \<in> P')"
and ac: "\<not> P \<Longrightarrow> ccorres_underlying sr \<Gamma> r xf arrel axf R R' (SKIP # hs) a c"
and bd: "P \<Longrightarrow> ccorres_underlying sr \<Gamma> r xf arrel axf U U' (SKIP # hs) b d"
and cthrows: "\<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> PT' c {}, UNIV" -- "c always throws"
and cthrows: "\<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> PT' c {}, UNIV" \<comment> \<open>c always throws\<close>
shows "ccorres_underlying sr \<Gamma> r xf arrel axf
(Q and (\<lambda>s. \<not> P \<longrightarrow> R s) and (\<lambda>s. P \<longrightarrow> U s))
(Collect Q' \<inter> {s. (s \<in> P' \<longrightarrow> s \<in> R' \<inter> PT') \<and> (s \<notin> P' \<longrightarrow> s \<in> U')})
@ -245,7 +245,7 @@ lemma ccorres_split_when_throwError_cond:
R R' (SKIP # hs) (throwError e) c"
and bd: "\<not> P \<Longrightarrow> ccorres_underlying sr \<Gamma> r xf ar axf
U U' (SKIP # hs) b d"
and cthrows: "\<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> PT' c {}, UNIV" -- "c always throws"
and cthrows: "\<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> PT' c {}, UNIV" \<comment> \<open>c always throws\<close>
shows "ccorres_underlying sr \<Gamma> r xf ar axf
(Q and (\<lambda>s. P \<longrightarrow> R s) and (\<lambda>s. \<not> P \<longrightarrow> U s))
(Collect Q' \<inter> {s. (s \<in> P' \<longrightarrow> s \<in> R' \<inter> PT') \<and> (s \<notin> P' \<longrightarrow> s \<in> U')})

View File

@ -1166,7 +1166,7 @@ lemma ccorres_trim_redundant_throw':
apply (erule exec_handlers_Seq_cases')
apply simp
apply (erule disjE)
-- "Non-abrupt case"
\<comment> \<open>Non-abrupt case\<close>
apply clarsimp
apply (erule_tac x = "t'" in ccorresE [OF cc])
apply assumption
@ -1175,7 +1175,7 @@ lemma ccorres_trim_redundant_throw':
apply (erule (1) EHOther)
apply simp
apply (erule exec_Normal_elim_cases | simp)+
-- "Abrupt case"
\<comment> \<open>Abrupt case\<close>
apply clarsimp
apply (erule_tac x = "t'" in ccorresE [OF cc])
apply assumption
@ -1383,7 +1383,7 @@ lemma ccorres_split_nothrow_record_novcg:
and bd: "\<And>rv rv'. r' rv rv' \<Longrightarrow> ccorres_underlying sr \<Gamma> r xf arrel axf (Q rv) (Q' rv rv') hs (b rv) (d' (xfru (\<lambda>_. rv') oldv))"
and valid: "\<lbrace>R\<rbrace> a \<lbrace>Q\<rbrace>"
and novcg: "guard_is_UNIV r' (xfr \<circ> xf') Q'"
-- "This might cause problems \<dots> has to be preserved across c in vcg case, but we can't do that"
\<comment> \<open>This might cause problems \<dots> has to be preserved across c in vcg case, but we can't do that\<close>
and xfoldv: "\<And>s. xf' s = xfru (\<lambda>_. (xfr \<circ> xf') s) oldv"
shows "ccorres_underlying sr \<Gamma> r xf arrel axf (P and R) P' hs (a >>= (\<lambda>rv. b rv)) (c ;; d)"
apply (rule ccorres_master_split_nohs_UNIV)

View File

@ -315,15 +315,15 @@ text \<open>One layer of context around a Simpl program.
command of a @{term Seq} if the first command @{text never_continues}.\<close>
datatype ('s,'p,'f) com_ctxt
= Seq1C "('s,'p,'f) com" -- "first command of a @{term Seq}"
| Seq2C "('s,'p,'f) com" -- "second command of a @{term Seq}"
| CondTC "'s bexp" "('s,'p,'f) com" -- "first branch of a @{term Cond}"
| CondFC "'s bexp" "('s,'p,'f) com" bool -- "second branch of a @{term Cond}"
| WhileC "'s bexp" -- "body of a @{term While}"
| WhileAnnoC "'s bexp" "'s assn" "('s \<times> 's) assn" -- "body of a @{term whileAnno}"
| GuardC "'f" "'s bexp" -- "body of a @{term Guard}"
| TryC "('s,'p,'f) com" -- "body of a @{term Catch}"
| HandlerC "('s,'p,'f) com" bool -- "handler of a @{term Catch}"
= Seq1C "('s,'p,'f) com" \<comment> \<open>first command of a @{term Seq}\<close>
| Seq2C "('s,'p,'f) com" \<comment> \<open>second command of a @{term Seq}\<close>
| CondTC "'s bexp" "('s,'p,'f) com" \<comment> \<open>first branch of a @{term Cond}\<close>
| CondFC "'s bexp" "('s,'p,'f) com" bool \<comment> \<open>second branch of a @{term Cond}\<close>
| WhileC "'s bexp" \<comment> \<open>body of a @{term While}\<close>
| WhileAnnoC "'s bexp" "'s assn" "('s \<times> 's) assn" \<comment> \<open>body of a @{term whileAnno}\<close>
| GuardC "'f" "'s bexp" \<comment> \<open>body of a @{term Guard}\<close>
| TryC "('s,'p,'f) com" \<comment> \<open>body of a @{term Catch}\<close>
| HandlerC "('s,'p,'f) com" bool \<comment> \<open>handler of a @{term Catch}\<close>
text \<open>Rewrite Simpl programs under a predicate @{term P} which is preserved by @{term com_eq}.\<close>

View File

@ -454,21 +454,21 @@ end
*)
lemmas typ_heap_simps =
-- "c_guard"
\<comment> \<open>c_guard\<close>
c_guard_field
c_guard_h_t_valid
-- "h_t_valid"
\<comment> \<open>h_t_valid\<close>
h_t_valid_field
h_t_valid_clift
-- "h_val"
\<comment> \<open>h_val\<close>
h_val_field_clift'
h_val_clift'
-- "clift"
\<comment> \<open>clift\<close>
clift_field
clift_field_update
heap_update_field_hrs
heap_update_field'
clift_heap_update
clift_heap_update_same_td_name -- "Try this last (is expensive)"
clift_heap_update_same_td_name \<comment> \<open>Try this last (is expensive)\<close>
end

View File

@ -721,7 +721,7 @@ syntax (HTML output)
syntax (latex output)
"_sep_map_set_conj" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add" ("(3\<And>\<^sup>*(00\<^bsub>_\<in>_\<^esub>) _)" [0, 51, 10] 10)
translations -- {* Beware of argument permutation! *}
translations \<comment> \<open>Beware of argument permutation!\<close>
"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"

View File

@ -150,7 +150,7 @@ lemma "VARS H p q r
apply (rule_tac x = "p # qs" in exI)
apply (simp add: sep_conj_exists sep_conj_ac)
apply (rule exI)
apply (sep_rule (direct) maps_to_write) -- "note: demonstrates computation"
apply (sep_rule (direct) maps_to_write) \<comment> \<open>note: demonstrates computation\<close>
apply (sep_solve add: maps_to_maps_to_ex)
apply clarsimp
done

View File

@ -108,7 +108,7 @@ record 'a PAS =
pasASIDAbs :: "'a agent_asid_map"
pasIRQAbs :: "'a agent_irq_map"
pasPolicy :: "'a auth_graph"
pasSubject :: "'a" -- "The active label"
pasSubject :: "'a" \<comment> \<open>The active label\<close>
pasMayActivate :: "bool"
pasMayEditReadyQueues :: "bool"
pasMaySendIrqs :: "bool"

View File

@ -1114,12 +1114,12 @@ lemma decode_arch_invocation_authorised:
neq_Nil_conv invs_psp_aligned invs_vspace_objs cli_no_irqs)
apply (drule diminished_cte_wp_at_valid_cap, clarsimp+)
apply (cases cap, simp_all)
-- "asid pool"
\<comment> \<open>asid pool\<close>
apply ((clarsimp simp: split_def cap_auth_conferred_def is_page_cap_def
pas_refined_all_auth_is_owns asid_high_bits_of_add_ucast
valid_cap_simps cap_links_asid_slot_def
label_owns_asid_slot_def pas_refined_refl )+)[1]
-- "ControlCap"
\<comment> \<open>ControlCap\<close>
apply (clarsimp simp: neq_Nil_conv split_def valid_cap_simps)
apply (cases excaps, simp_all)[1]
apply (clarsimp simp: neq_Nil_conv aag_has_auth_to_Control_eq_owns)
@ -1127,12 +1127,12 @@ lemma decode_arch_invocation_authorised:
apply (clarsimp simp: valid_cap_def cap_aligned_def)
apply (clarsimp simp: is_cap_simps cap_auth_conferred_def pas_refined_all_auth_is_owns aag_cap_auth_def)
-- "PageCap"
\<comment> \<open>PageCap\<close>
apply (clarsimp simp: valid_cap_simps cli_no_irqs)
apply (cases "invocation_type label", simp_all)
apply (rename_tac archlabel)
apply (case_tac archlabel, simp_all)
-- "Map"
\<comment> \<open>Map\<close>
apply (clarsimp simp: cap_auth_conferred_def is_cap_simps is_page_cap_def pas_refined_all_auth_is_owns)
apply (rule conjI)
apply (clarsimp simp: cap_auth_conferred_def is_page_cap_def pas_refined_all_auth_is_owns aag_cap_auth_def cli_no_irqs cap_links_asid_slot_def)
@ -1144,7 +1144,7 @@ lemma decode_arch_invocation_authorised:
exI vspace_cap_rights_to_auth_def mask_vm_rights_def
validate_vm_rights_def vm_read_write_def vm_read_only_def
vm_kernel_only_def)
-- "Remap"
\<comment> \<open>Remap\<close>
apply (clarsimp simp: cap_auth_conferred_def
is_page_cap_def pas_refined_all_auth_is_owns)
apply (rule conjI, fastforce)
@ -1155,16 +1155,16 @@ lemma decode_arch_invocation_authorised:
validate_vm_rights_def vm_read_write_def vm_read_only_def
vm_kernel_only_def
split: if_split_asm)
-- "Unmap"
\<comment> \<open>Unmap\<close>
apply (simp add: aag_cap_auth_def cli_no_irqs)
-- "PageTableCap"
\<comment> \<open>PageTableCap\<close>
apply (cases "invocation_type label", simp_all)
apply (rename_tac archlabel)
apply (case_tac archlabel, simp_all)
-- "PTMap"
\<comment> \<open>PTMap\<close>
apply (clarsimp simp: aag_cap_auth_def cli_no_irqs cap_links_asid_slot_def cap_auth_conferred_def is_page_cap_def
pde_ref2_def pas_refined_all_auth_is_owns pas_refined_refl pd_shifting [folded pd_bits_14] )
-- "Unmap"
\<comment> \<open>Unmap\<close>
apply (rename_tac word option archlabel)
apply (clarsimp simp: aag_cap_auth_def cli_no_irqs cap_links_asid_slot_def cap_auth_conferred_def is_page_cap_def
pde_ref2_def pas_refined_all_auth_is_owns pas_refined_refl )

View File

@ -341,7 +341,7 @@ lemma set_cap_pas_refined [wp]:
apply (rule hoare_pre)
apply (wp set_cap_caps_of_state| wps)+
apply clarsimp
apply (intro conjI) -- "auth_graph_map"
apply (intro conjI) \<comment> \<open>auth_graph_map\<close>
apply (clarsimp dest!: auth_graph_map_memD)
apply (erule state_bits_to_policy.cases, auto simp: cap_links_asid_slot_def label_owns_asid_slot_def intro: auth_graph_map_memI state_bits_to_policy.intros
split: if_split_asm)[1]

View File

@ -183,7 +183,7 @@ lemma dmo_storeWord_respects_ipc:
apply (simp add: msg_align_bits)
apply (erule mul_word_size_lt_msg_align_bits_ofnat)
apply simp
-- "non auth case"
\<comment> \<open>non auth case\<close>
apply (rule hoare_pre)
apply (simp add: storeWord_def)
apply (wp dmo_wp)
@ -201,7 +201,7 @@ lemma dmo_storeWord_respects_ipc:
apply (simp add: msg_align_bits)
apply (erule mul_word_size_lt_msg_align_bits_ofnat)
apply simp
-- "otherwise"
\<comment> \<open>otherwise\<close>
apply (auto simp: is_aligned_mask [symmetric] intro!: trm_lrefl ptr_range_memI ptr_range_add_memI)
done
@ -593,7 +593,7 @@ lemma send_signal_respects:
apply (rule hoare_seq_ext[OF _ get_simple_ko_sp])
apply (rule hoare_name_pre_state)
apply (case_tac "ntfn_obj ntfn = IdleNtfn \<and> ntfn_bound_tcb ntfn \<noteq> None")
-- "ntfn-binding case"
\<comment> \<open>ntfn-binding case\<close>
apply (rule hoare_pre)
apply (wp set_notification_respects[where auth=Notify]
as_user_set_register_respects_indirect[where ntfnptr=ntfnptr]
@ -637,7 +637,7 @@ lemma send_signal_respects:
apply simp
apply simp
apply (intro impI conjI)
-- "st_tcb_at receive_blocked st"
\<comment> \<open>st_tcb_at receive_blocked st\<close>
apply (erule (2) integrity_receive_blocked_chain)
apply clarsimp
done
@ -1125,7 +1125,7 @@ lemma receive_ipc_base_pas_refined:
apply (clarsimp simp: tcb_at_def [symmetric] conj_ac tcb_at_st_tcb_at)
apply (rule conjI)
apply (rule impI)
-- "is_subject"
\<comment> \<open>is_subject\<close>
apply (subgoal_tac "aag_has_auth_to aag Control (hd x)")
apply (fastforce simp add: pas_refined_refl dest!: aag_Control_into_owns)
apply (rule_tac ep = "pasObjectAbs aag epptr" in aag_wellformed_grant_Control_to_send [OF _ _ pas_refined_wellformed])
@ -1410,7 +1410,7 @@ lemma receive_ipc_base_integrity:
apply safe
apply (fastforce simp: valid_objs_def valid_obj_def obj_at_def
ep_redux_simps neq_Nil_conv valid_ep_def case_list_cons_cong)
-- "is_subject"
\<comment> \<open>is_subject\<close>
apply (subgoal_tac "aag_has_auth_to aag Control (hd x)")
apply (fastforce simp add: pas_refined_refl dest!: aag_Control_into_owns)
apply (rule_tac ep = "pasObjectAbs aag epptr" in aag_wellformed_grant_Control_to_send [OF _ _ pas_refined_wellformed])
@ -1534,12 +1534,12 @@ lemma integrity_tcb_in_ipc_final:
apply (clarsimp simp: tcb_bound_notification_reset_integrity_def)
apply (rule disjI1)
apply (clarsimp simp: direct_send_def)
-- "trm"
\<comment> \<open>trm\<close>
apply clarsimp
apply (cases "is_subject aag thread")
apply (rule trm_write)
apply simp
-- "doesn't own"
\<comment> \<open>doesn't own\<close>
apply (erule tcb_in_ipc.cases, simp_all)[1]
apply clarsimp
apply (rule trm_ipc [where p' = thread and ep = epptr])
@ -1653,7 +1653,7 @@ lemma set_object_integrity_in_ipc_autarch:
apply simp
apply (erule integrity_trans)
apply (clarsimp simp: integrity_def)
-- "tii"
\<comment> \<open>tii\<close>
apply clarsimp
done
@ -2018,14 +2018,14 @@ lemma send_ipc_integrity_autarch:
apply (rule hoare_pre)
apply (wp set_endpoinintegrity set_thread_state_integrity_autarch
| wpc | simp)+
apply (fastforce simp: obj_at_def is_ep) -- "ep_at and has_auth"
-- "SendEP"
apply (fastforce simp: obj_at_def is_ep) \<comment> \<open>ep_at and has_auth\<close>
\<comment> \<open>SendEP\<close>
apply simp
apply (rule hoare_pre)
apply (wp set_endpoinintegrity set_thread_state_integrity_autarch
| wpc | simp)+
apply (fastforce simp: obj_at_def is_ep) -- "ep_at and has_auth"
-- "WaitingEP"
apply (fastforce simp: obj_at_def is_ep) \<comment> \<open>ep_at and has_auth\<close>
\<comment> \<open>WaitingEP\<close>
apply (rename_tac list)
apply simp
apply (case_tac "is_subject aag (hd list)") (* autarch or not on rec. side *)
@ -2045,12 +2045,12 @@ lemma send_ipc_integrity_autarch:
apply blast
apply (fastforce simp: valid_objs_def valid_obj_def obj_at_def valid_simple_obj_def
ep_redux_simps neq_Nil_conv valid_ep_def case_list_cons_cong)
-- "we don't own head of queue"
\<comment> \<open>we don't own head of queue\<close>
apply clarsimp
apply (rule use_spec') -- "Name initial state"
apply (simp add: spec_valid_def) -- "no imp rule?"
apply (rule use_spec') \<comment> \<open>Name initial state\<close>
apply (simp add: spec_valid_def) \<comment> \<open>no imp rule?\<close>
apply (rule_tac Q = "\<lambda>_ s'. integrity aag X st s \<and> integrity aag X s s'" in hoare_post_imp) -- "We want to apply refl later on, so use initial state"
apply (rule_tac Q = "\<lambda>_ s'. integrity aag X st s \<and> integrity aag X s s'" in hoare_post_imp) \<comment> \<open>We want to apply refl later on, so use initial state\<close>
apply (clarsimp elim!: integrity_trans)
apply (rule hoare_pre)
apply (wp set_endpoinintegrity set_thread_state_integrity_autarch setup_caller_cap_integrity_autarch
@ -2077,9 +2077,9 @@ lemma send_ipc_integrity_autarch:
apply simp
apply simp
apply (intro conjI)
-- "\<not> can_grant"
\<comment> \<open>\<not> can_grant\<close>
apply (clarsimp simp: obj_at_def)
-- "refl tcb_in_ipc"
\<comment> \<open>refl tcb_in_ipc\<close>
apply (erule (4) integrity_tcb_in_ipc_refl)
apply (fastforce simp: valid_objs_def valid_obj_def obj_at_def
ep_redux_simps neq_Nil_conv valid_ep_def case_list_cons_cong)

View File

@ -99,9 +99,9 @@ lemma perform_invocation_respects:
invoke_arch_respects invoke_irq_control_respects invoke_irq_handler_respects
| wp_once hoare_pre_cont)+
apply (clarsimp simp: authorised_invocation_def split: Invocations_A.invocation.splits)
-- "EP case"
\<comment> \<open>EP case\<close>
apply (fastforce simp: obj_at_def is_tcb split: if_split_asm)
-- "NTFN case"
\<comment> \<open>NTFN case\<close>
apply fastforce
done
@ -529,7 +529,7 @@ lemma handle_interrupt_integrity:
| wpc
| simp add: get_irq_slot_def get_irq_state_def ackInterrupt_def resetTimer_def handle_reserved_irq_def)+
apply clarsimp
apply (rule conjI, fastforce)+ -- "valid_objs etc."
apply (rule conjI, fastforce)+ \<comment> \<open>valid_objs etc.\<close>
apply (clarsimp simp: cte_wp_at_caps_of_state)
apply (rule_tac s = s in hacky_ipc_Send [where irq = irq])
apply (drule (1) cap_auth_caps_of_state)

View File

@ -82,7 +82,7 @@ definition
definition
"separate_tcb p cs \<equiv> case_option True (separate_cnode_cap cs) (cs (p, tcb_cnode_index 0))
\<and> cs (p, tcb_cnode_index 3) = Some NullCap" -- "ctable and caller cap"
\<and> cs (p, tcb_cnode_index 3) = Some NullCap" \<comment> \<open>ctable and caller cap\<close>
lemma separate_cnode_cap_rab:
"\<lbrakk> separate_cnode_cap cs cap; length cref = word_bits \<rbrakk> \<Longrightarrow>

View File

@ -288,7 +288,7 @@ lemma send_fault_ipc_bisim:
apply (wp | simp)+
apply clarsimp
apply assumption
-- "det_ont"
\<comment> \<open>det_ont\<close>
apply (simp add: Let_def cong: cap.case_cong)
apply (wp not_empty_lc)
apply (rule_tac P = "separate_cap xa" in not_empty_gen_asm)

View File

@ -1782,7 +1782,7 @@ lemma createMappingEntries_valid_pde_slots'2:
apply (rule hoare_pre, wp)
apply (clarsimp simp: vmsz_aligned'_def page_directory_at'_def lookup_pd_slot_def)
apply (rule conjI)
subgoal -- "valid_pde_mapping_offset'"
subgoal \<comment> \<open>valid_pde_mapping_offset'\<close>
apply (clarsimp simp: superSectionPDEOffsets_def length_upto_enum_step pdeBits_def)
apply (clarsimp simp: upto_enum_step_def upto_enum_def comp_def)
apply (clarsimp simp: linorder_not_less field_simps mask_add_aligned)
@ -1790,7 +1790,7 @@ lemma createMappingEntries_valid_pde_slots'2:
apply (rule word_of_nat_le, simp)
done
apply (rule conjI)
subgoal -- "pde_at'"
subgoal \<comment> \<open>pde_at'\<close>
apply (clarsimp simp: superSectionPDEOffsets_def length_upto_enum_step pdeBits_def)
apply (clarsimp simp:upto_enum_step_def upto_enum_def hd_map_simp comp_def)
apply (simp add: vaddr_segment_nonsense6)
@ -3012,7 +3012,7 @@ lemma decodeARMFrameInvocation_ccorres:
apply simp
apply (vcg exspec=getSyscallArg_modifies)
-- "PageMap"
\<comment> \<open>PageMap\<close>
apply (rule ccorres_rhs_assoc)+
apply csymbr+
apply (simp add: if_1_0_0 word_less_nat_alt del: Collect_const)
@ -4130,7 +4130,7 @@ lemma Arch_decodeInvocation_ccorres:
apply wp
apply simp
apply (vcg exspec=getSyscallArg_modifies)
apply (rule ccorres_Cond_rhs) -- "ASIDPoolCap case"
apply (rule ccorres_Cond_rhs) \<comment> \<open>ASIDPoolCap case\<close>
apply (simp add: imp_conjR[symmetric] decodeARMMMUInvocation_def
del: Collect_const)
apply (rule ccorres_rhs_assoc)+

View File

@ -207,7 +207,7 @@ lemma (* handleDoubleFault_ccorres: *)
apply ac_init
prefer 3 apply simp
apply (unfold handleDoubleFault_def handleDoubleFault'_def K_bind_def)
apply (rule corres_add_noop_rhs2) -- "split out extra haskell code"
apply (rule corres_add_noop_rhs2) \<comment> \<open>split out extra haskell code\<close>
apply (rule corres_split[OF _ setThreadState_ccorres[ac]])
apply (rule corres_symb_exec_l_no_exs)
apply simp

View File

@ -303,7 +303,7 @@ lemmas ccorres_move_array_assertion_cnode_ctes [corres_pre]
ccorres_move_Guard [OF array_assertion_abs_cnode_ctes]
lemma locateSlotCNode_ccorres [corres]:
assumes gl: "\<And>v s. globals (xfu v s) = globals s" -- "for state rel. preservation"
assumes gl: "\<And>v s. globals (xfu v s) = globals s" \<comment> \<open>for state rel. preservation\<close>
and fg: "\<And>v s. xf (xfu (\<lambda>_. v) s) = v"
shows "ccorres (\<lambda>v v'. v' = Ptr v) xf \<top> {_. cnode = cnode' \<and> offset = offset'} hs
(locateSlotCNode cnode bits offset)
@ -324,7 +324,7 @@ lemma locateSlotCNode_ccorres [corres]:
done
lemma locateSlotTCB_ccorres [corres]:
assumes gl: "\<And>v s. globals (xfu v s) = globals s" -- "for state rel. preservation"
assumes gl: "\<And>v s. globals (xfu v s) = globals s" \<comment> \<open>for state rel. preservation\<close>
and fg: "\<And>v s. xf (xfu (\<lambda>_. v) s) = v"
shows "ccorres (\<lambda>v v'. v' = Ptr v) xf \<top> {_. cnode = cnode' \<and> offset = offset'} hs
(locateSlotTCB cnode offset)
@ -340,13 +340,13 @@ lemma locateSlotTCB_ccorres [corres]:
lemma getSlotCap_h_val_ccorres [corres]:
fixes p :: "cstate \<Rightarrow> cte_C ptr"
assumes gl: "\<And>v s. globals (xfu v s) = globals s" -- "for state rel. preservation"
assumes gl: "\<And>v s. globals (xfu v s) = globals s" \<comment> \<open>for state rel. preservation\<close>
and fg: "\<And>v s. xf (xfu (\<lambda>_. v) s) = v"
shows "ccorres ccap_relation xf \<top> {s. p s = Ptr a} hs
(getSlotCap a) (Basic (\<lambda>s. xfu (\<lambda>_. h_val (hrs_mem (t_hrs_' (globals s))) (Ptr &(p s\<rightarrow>[''cap_C'']) :: cap_C ptr)) s))"
unfolding getSlotCap_def
apply (rule ccorres_add_UNIV_Int)
apply (cinitlift p) -- "EVIL!"
apply (cinitlift p) \<comment> \<open>EVIL!\<close>
apply simp
apply (rule ccorres_guard_imp2)
apply (rule ccorres_pre_getCTE)

View File

@ -40,9 +40,9 @@ lemma lookupCap_ccorres':
apply (cinit lift: cPtr_' thread_' simp: lookupCapAndSlot_def liftME_def bindE_assoc)
apply (ctac (no_vcg) add: lookupSlotForThread_ccorres')
--"case where lu_ret.status is EXCEPTION_NONE"
\<comment> \<open>case where lu_ret.status is EXCEPTION_NONE\<close>
apply (simp add: split_beta cong:call_ignore_cong)
apply csymbr --"call status_C_update"
apply csymbr \<comment> \<open>call status_C_update\<close>
apply (simp add: Collect_const[symmetric] lookupSlot_raw_rel_def liftE_def
del: Collect_const)
apply (rule ccorres_move_c_guard_cte)
@ -50,18 +50,18 @@ lemma lookupCap_ccorres':
apply (rule ccorres_return_CE [unfolded returnOk_def, simplified], simp+)[1]
apply wp
apply vcg
--"case where lu_ret.status is *not* EXCEPTION_NONE"
\<comment> \<open>case where lu_ret.status is *not* EXCEPTION_NONE\<close>
apply simp
apply (rule ccorres_split_throws)
apply (rule ccorres_rhs_assoc)+
apply csymbr
apply csymbr -- "call cap_null_cap_new_'proc"
apply csymbr \<comment> \<open>call cap_null_cap_new_'proc\<close>
apply csymbr
apply (rule ccorres_return_C_errorE, simp+)[1]
apply vcg
apply (wp | simp)+
-- "last subgoal"
\<comment> \<open>last subgoal\<close>
apply (clarsimp simp: valid_pspace_valid_objs')
apply (intro impI conjI allI)
apply (simp add: lookupSlot_raw_rel_def)
@ -91,10 +91,10 @@ lemma lookupCapAndSlot_ccorres :
apply (cinit lift: thread_' cPtr_')
apply (ctac (no_vcg))
--"case where lu_ret.status is EXCEPTION_NONE"
\<comment> \<open>case where lu_ret.status is EXCEPTION_NONE\<close>
apply (simp add: split_beta cong:call_ignore_cong)
apply csymbr --"call status_C_update"
apply csymbr --"call slot_C_update"
apply csymbr \<comment> \<open>call status_C_update\<close>
apply csymbr \<comment> \<open>call slot_C_update\<close>
apply (simp add: Collect_const[symmetric] lookupSlot_raw_rel_def liftE_bindE
del: Collect_const)
apply (rule ccorres_move_c_guard_cte)
@ -103,7 +103,7 @@ lemma lookupCapAndSlot_ccorres :
apply (clarsimp simp: Bex_def in_monad getSlotCap_def in_getCTE2 cte_wp_at_ctes_of)
apply (erule(1) cmap_relationE1[OF cmap_relation_cte])
apply (simp add: ccte_relation_ccap_relation typ_heap_simps')
--"case where lu_ret.status is *not* EXCEPTION_NONE"
\<comment> \<open>case where lu_ret.status is *not* EXCEPTION_NONE\<close>
apply simp
apply (rule ccorres_split_throws)
apply (rule ccorres_rhs_assoc)+
@ -113,7 +113,7 @@ lemma lookupCapAndSlot_ccorres :
apply vcg
apply (wp | simp)+
-- "last subgoal"
\<comment> \<open>last subgoal\<close>
apply clarsimp
apply (rule conjI, fastforce)
apply clarsimp
@ -160,14 +160,14 @@ lemma lookupSlotForCNodeOp_ccorres':
(lookupSlotForCNodeOp isSource croot capptr depth)
(Call lookupSlotForCNodeOp_'proc)"
apply (cinit lift: capptr_' isSource_' root_' depth_')
apply csymbr -- "slot_C_update"
apply csymbr -- "cap_get_capType"
apply csymbr \<comment> \<open>slot_C_update\<close>
apply csymbr \<comment> \<open>cap_get_capType\<close>
apply (rule_tac Q=\<top> and Q'=\<top> in ccorres_if_cond_throws2)
-- "correspondance of Haskell and C conditions"
\<comment> \<open>correspondance of Haskell and C conditions\<close>
apply (clarsimp simp: Collect_const_mem cap_get_tag_isCap)
-- "case where root is *not* a CNode => throw InvalidRoot"
\<comment> \<open>case where root is *not* a CNode => throw InvalidRoot\<close>
apply simp
apply (rule_tac P=\<top> and P' =UNIV in ccorres_from_vcg_throws)
apply (rule allI, rule conseqPre, vcg)
@ -177,7 +177,7 @@ lemma lookupSlotForCNodeOp_ccorres':
apply clarsimp
apply (subst syscall_error_to_H_cases(6), simp+)[1]
-- " case where root is a CNode"
\<comment> \<open>case where root is a CNode\<close>
apply (simp add: rangeCheck_def)
apply csymbr
apply (rule ccorres_Cond_rhs_Seq)
@ -194,13 +194,13 @@ lemma lookupSlotForCNodeOp_ccorres':
apply vcg
apply csymbr
apply (rule_tac Q="\<lambda>s. depth < 2 ^ word_bits" and Q'=\<top> in ccorres_split_unless_throwError_cond)
-- "correspondance of Haskell and C conditions"
\<comment> \<open>correspondance of Haskell and C conditions\<close>
apply (clarsimp simp: Collect_const_mem fromIntegral_def integral_inv
if_1_0_0)
apply (simp add: word_size unat_of_nat32 word_less_nat_alt
word_less_1[symmetric] linorder_not_le)
-- "case of RangeError"
\<comment> \<open>case of RangeError\<close>
apply (rule_tac P= \<top> and P' =UNIV in ccorres_from_vcg_throws)
apply (rule allI, rule conseqPre, vcg)
apply (clarsimp simp: throwError_def return_def syscall_error_rel_def)
@ -208,20 +208,20 @@ lemma lookupSlotForCNodeOp_ccorres':
apply (subst syscall_error_to_H_cases(4), simp+)[1]
apply (simp add: word_size word_sle_def)
-- "case where there is *no* RangeError"
\<comment> \<open>case where there is *no* RangeError\<close>
apply (rule_tac xf'=lookupSlot_xf in ccorres_rel_imp)
apply (rule_tac r="\<lambda>w w'. w'= Ptr w"
and f="\<lambda> st fl es. fl = scast EXCEPTION_SYSCALL_ERROR \<and>
syscall_error_to_H (errsyscall es) (errlookup_fault es) = Some (FailedLookup isSource st)"
in lookupErrorOnFailure_ccorres)
apply (ctac (no_vcg)) -- "resolveAddressBits"
-- " case where resolveAddressBits results in EXCEPTION_NONE"
apply (ctac (no_vcg)) \<comment> \<open>resolveAddressBits\<close>
\<comment> \<open>case where resolveAddressBits results in EXCEPTION_NONE\<close>
apply clarsimp
apply (rule_tac A=\<top> and A'=UNIV in ccorres_guard_imp2)
apply (rule_tac Q=\<top> and Q'=\<top> in ccorres_if_cond_throws2)
-- "correspondance of Haskell and C conditions"
\<comment> \<open>correspondance of Haskell and C conditions\<close>
apply (clarsimp simp: Collect_const_mem unat_gt_0)
-- " case where bits are remaining"
\<comment> \<open>case where bits are remaining\<close>
apply (rule_tac P= \<top> and P' =UNIV in ccorres_from_vcg_throws)
apply (rule allI, rule conseqPre, vcg)
apply (clarsimp simp: throwError_def return_def)
@ -230,17 +230,17 @@ lemma lookupSlotForCNodeOp_ccorres':
apply (simp add: lookup_fault_depth_mismatch_lift)
apply (erule le_32_mask_eq)
-- " case where *no* bits are remaining"
apply csymbr -- "slot_C_update"
apply csymbr -- "status_C_update"
\<comment> \<open>case where *no* bits are remaining\<close>
apply csymbr \<comment> \<open>slot_C_update\<close>
apply csymbr \<comment> \<open>status_C_update\<close>
apply (rule ccorres_return_CE, simp+)[1]
apply vcg
-- "guard_imp subgoal"
\<comment> \<open>guard_imp subgoal\<close>
apply clarsimp
-- " case where resolveAddressBits does *not* results in EXCEPTION_NONE"
\<comment> \<open>case where resolveAddressBits does *not* results in EXCEPTION_NONE\<close>
apply clarsimp
apply (rule_tac P= \<top> and P' ="\<lbrace>\<exists>v. (lookup_fault_lift (\<acute>current_lookup_fault)) = Some v
\<and> lookup_fault_to_H v = err \<rbrace>"
@ -251,7 +251,7 @@ lemma lookupSlotForCNodeOp_ccorres':
apply (subst syscall_error_to_H_cases(6), simp+)[1]
apply wp
-- "rel_imp"
\<comment> \<open>rel_imp\<close>
apply clarsimp
apply (case_tac x, clarsimp)
apply (simp add: syscall_error_rel_def errstate_def)
@ -261,7 +261,7 @@ lemma lookupSlotForCNodeOp_ccorres':
apply vcg
apply vcg
-- "last subgoal"
\<comment> \<open>last subgoal\<close>
apply (clarsimp simp: if_1_0_0 to_bool_def true_def word_size
fromIntegral_def integral_inv)
apply (case_tac "cap_get_tag root = scast cap_cnode_cap")

File diff suppressed because it is too large Load Diff

View File

@ -214,7 +214,7 @@ proof (cases "isCNodeCap cap'")
show ?thesis using False
apply (cinit' lift: nodeCap_' capptr_' n_bits_')
apply csymbr+
-- "Exception stuff"
\<comment> \<open>Exception stuff\<close>
apply (rule ccorres_split_throws)
apply (simp add: Collect_const cap_get_tag_isCap isCap_simps ccorres_cond_iffs
resolveAddressBits.simps scast_id)
@ -236,8 +236,8 @@ next
from True show ?thesis
apply -
apply (cinit' simp add: whileAnno_def ucast_id)
-- "This is done here as init lift usually throws away the relationship between nodeCap_' s and nodeCap. Normally
this OK, but the induction messes with everything :("
\<comment> \<open>This is done here as init lift usually throws away the relationship between nodeCap_' s and nodeCap. Normally
this OK, but the induction messes with everything :(\<close>
apply (rule ccorres_abstract [where xf' = nodeCap_'])
apply ceqv
apply (rename_tac "nodeCap")
@ -266,11 +266,11 @@ next
apply clarsimp
apply (rule_tac I = "{s. cap_get_tag (nodeCap_' s) = scast cap_cnode_cap}"
in HoarePartial.While [unfolded whileAnno_def, OF subset_refl])
apply (vcg strip_guards=true) -- "takes a while"
apply (vcg strip_guards=true) \<comment> \<open>takes a while\<close>
apply clarsimp
apply simp
apply (clarsimp simp: cap_get_tag_isCap to_bool_def)
-- "Main thm"
\<comment> \<open>Main thm\<close>
proof (induct cap' cptr' guard' rule: resolveAddressBits.induct [case_names ind])
case (ind cap cptr guard)
@ -476,12 +476,12 @@ next
apply (csymbr | rule iffD2 [OF ccorres_seq_skip])+
apply (rule ccorres_Guard_Seq)+
apply csymbr
-- "handle the stateAssert in locateSlotCap very carefully"
\<comment> \<open>handle the stateAssert in locateSlotCap very carefully\<close>
apply (simp(no_asm) only: liftE_bindE[where a="locateSlotCap a b" for a b])
apply (rule ccorres_locateSlotCap_push[rotated])
apply (simp add: unlessE_def)
apply (rule hoare_pre, wp, simp)
-- "Convert guardBits, radixBits and capGuard to their Haskell versions"
\<comment> \<open>Convert guardBits, radixBits and capGuard to their Haskell versions\<close>
apply (drule (2) cgD, drule (2) rbD, drule (2) gbD)
apply (elim conjE)
apply (rule ccorres_gen_asm [where P = "guard \<le> 32"])
@ -550,9 +550,9 @@ next
apply (vcg strip_guards=true)
apply (vcg strip_guards=true)
apply (rule conjI)
-- "Haskell guard"
\<comment> \<open>Haskell guard\<close>
apply (thin_tac "unat n_bits = guard")
apply (clarsimp simp del: imp_disjL) -- "take a while"
apply (clarsimp simp del: imp_disjL) \<comment> \<open>take a while\<close>
apply (intro impI conjI allI)
apply fastforce
apply clarsimp
@ -562,7 +562,7 @@ next
apply (clarsimp simp: isCap_simps valid_cap_simps' cte_level_bits_def objBits_defs
real_cte_at')
apply (clarsimp simp: isCap_simps valid_cap'_def)
-- "C guard"
\<comment> \<open>C guard\<close>
apply (frule (1) cgD [OF refl], frule (1) rbD [OF refl], frule (1) gbD [OF refl])
apply (simp add: Collect_const_mem cap_get_tag_isCap exception_defs lookup_fault_lifts
n_bits_guard mask6_eqs word_le_nat_alt word_less_nat_alt gm)

View File

@ -1766,7 +1766,7 @@ proof -
, assumption +,
simp_all add: objBits_simps' archObjSize_def pageBits_def projectKOs
pteBits_def pdeBits_def
heap_to_user_data_restrict heap_to_device_data_restrict)[1])+ -- "waiting ..."
heap_to_user_data_restrict heap_to_device_data_restrict)[1])+ \<comment> \<open>waiting ...\<close>
apply (simp add: map_to_ctes_delete' cmap_relation_restrict_both_proj
cmap_relation_restrict_both cmap_array_helper hrs_htd_update
ptBits_def pdBits_def pageBits_def cmap_array)
@ -1816,7 +1816,7 @@ proof -
apply fastforce
apply ((subst lift_t_typ_region_bytes, rule cm_disj_tcb, assumption+,
simp_all add: objBits_simps archObjSize_def pageBits_def projectKOs)[1])+
-- "waiting ..."
\<comment> \<open>waiting ...\<close>
apply (simp add: tcb_queue_relation_live_restrict
[OF D.valid_untyped tat tlive rl])
done

View File

@ -661,7 +661,7 @@ lemma doUnbindNotification_ccorres:
cpspace_relation_def update_ntfn_map_tos)
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (rule cpspace_relation_ntfn_update_ntfn, assumption+)
apply (clarsimp simp: cnotification_relation_def Let_def
mask_def [where n=2] NtfnState_Waiting_def)
@ -709,7 +709,7 @@ lemma doUnbindNotification_ccorres':
cpspace_relation_def update_ntfn_map_tos)
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (rule cpspace_relation_ntfn_update_ntfn, assumption+)
apply (clarsimp simp: cnotification_relation_def Let_def
mask_def [where n=2] NtfnState_Waiting_def)
@ -906,7 +906,7 @@ lemma finaliseCap_True_cases_ccorres:
apply (clarsimp simp add: return_def ccap_relation_NullCap_iff)
apply (clarsimp simp add: irq_opt_relation_def)
apply vcg
-- "NullCap case by exhaustion"
\<comment> \<open>NullCap case by exhaustion\<close>
apply (simp add: cap_get_tag_isCap Let_def
ccorres_cond_empty_iff ccorres_cond_univ_iff)
apply (rule ccorres_rhs_assoc2, rule ccorres_rhs_assoc2,

View File

@ -441,13 +441,13 @@ lemma invokeCNodeSaveCaller_ccorres:
Collect_False Collect_True
del: Collect_const)[1]
apply (rule ccorres_fail)+
-- "NullCap case"
\<comment> \<open>NullCap case\<close>
apply (rule ccorres_from_vcg_throws [where P=\<top> and P'=UNIV])
apply (rule allI, rule conseqPre, vcg)
apply clarsimp
apply (simp add: return_def)
apply (rule ccorres_fail)+
-- "ReplyCap case"
\<comment> \<open>ReplyCap case\<close>
apply (rule ccorres_rhs_assoc)
apply csymbr
apply (frule cap_get_tag_isCap_unfolded_H_cap)
@ -736,7 +736,7 @@ lemma decodeCNodeInvocation_ccorres:
injection_handler_returnOk Collect_const[symmetric]
cong: call_ignore_cong del: Collect_const)
apply (rule ccorres_Cond_rhs_Seq)
-- "CNodeCopy case"
\<comment> \<open>CNodeCopy case\<close>
apply (simp add: Collect_const[symmetric] del: Collect_const)
apply (rule ccorres_rhs_assoc)+
apply (rule ccorres_Cond_rhs_Seq)
@ -806,7 +806,7 @@ lemma decodeCNodeInvocation_ccorres:
apply (simp add: Collect_const_mem)
apply (vcg exspec=getSyscallArg_modifies)
apply (rule ccorres_Cond_rhs_Seq)
-- "CNodeMint case"
\<comment> \<open>CNodeMint case\<close>
apply (simp add: Collect_const[symmetric]
del: Collect_const)
apply (rule ccorres_rhs_assoc)+
@ -893,7 +893,7 @@ lemma decodeCNodeInvocation_ccorres:
apply (simp add: Collect_const_mem)
apply (vcg exspec=getSyscallArg_modifies)
apply (rule ccorres_Cond_rhs_Seq)
-- "CNodeMove case"
\<comment> \<open>CNodeMove case\<close>
apply (simp add: Collect_const[symmetric] split_def
injection_handler_returnOk whenE_def
ccorres_invocationCatch_Inr
@ -921,7 +921,7 @@ lemma decodeCNodeInvocation_ccorres:
apply vcg
apply (rule conseqPre, vcg, clarsimp)
apply (rule ccorres_Cond_rhs_Seq)
-- "CNodeMutate case"
\<comment> \<open>CNodeMutate case\<close>
apply (rule ccorres_rhs_assoc)+
apply (simp add: Collect_const[symmetric] del: Collect_const
cong: call_ignore_cong)

View File

@ -229,7 +229,7 @@ lemma cancelSignal_ccorres_helper:
apply (clarsimp simp: typ_heap_simps cong: imp_cong split del: if_split simp del: comp_def)
apply (frule null_ep_queue [simplified Fun.comp_def])
apply (intro impI conjI allI)
-- "empty case"
\<comment> \<open>empty case\<close>
apply clarsimp
apply (frule iffD1 [OF tcb_queue_head_empty_iff [OF tcb_queue_relation'_queue_rel]])
apply (rule ballI, erule bspec)
@ -243,20 +243,20 @@ lemma cancelSignal_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cendpoint_relation_ntfn_queue [OF invs_sym'], assumption+)
apply simp
apply (erule (1) map_to_ko_atI')
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (rule cpspace_relation_ntfn_update_ntfn, assumption+)
apply (simp add: cnotification_relation_def Let_def NtfnState_Idle_def)
apply (simp add: carch_state_relation_def carch_globals_def)
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (simp add: carch_state_relation_def carch_globals_def
@ -268,7 +268,7 @@ lemma cancelSignal_ccorres_helper:
apply (simp add: objBits_simps)
apply assumption
-- "non empty case"
\<comment> \<open>non empty case\<close>
apply (frule tcb_queue_head_empty_iff [OF tcb_queue_relation'_queue_rel])
apply (rule ballI, erule bspec)
apply (erule subsetD [rotated])
@ -281,10 +281,10 @@ lemma cancelSignal_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cendpoint_relation_ntfn_queue)
@ -292,7 +292,7 @@ lemma cancelSignal_ccorres_helper:
apply assumption+
apply simp
apply (erule (1) map_to_ko_atI')
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (rule cpspace_relation_ntfn_update_ntfn, assumption+)
apply (simp add: cnotification_relation_def Let_def isWaitingNtfn_def
split: ntfn.splits split del: if_split)
@ -302,7 +302,7 @@ lemma cancelSignal_ccorres_helper:
apply (clarsimp simp add: Ptr_ptr_val h_t_valid_clift_Some_iff)
apply (simp add: tcb_queue_relation'_prev_mask)
apply simp
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (simp add: carch_state_relation_def carch_globals_def
@ -740,7 +740,7 @@ lemma state_relation_queue_update_helper':
apply (clarsimp simp: inQ_def obj_at'_def)
apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def)
apply (intro conjI)
-- "cpspace_relation"
\<comment> \<open>cpspace_relation\<close>
apply (erule nonemptyE, drule(1) bspec)
apply (clarsimp simp: cpspace_relation_def)
apply (drule obj_at_ko_at', clarsimp)
@ -749,18 +749,18 @@ lemma state_relation_queue_update_helper':
apply (frule null_sched_queue)
apply (frule null_sched_epD)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (drule ctcb_relation_null_queue_ptrs,
simp_all)[1]
-- "endpoint relation"
\<comment> \<open>endpoint relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (erule cendpoint_relation_upd_tcb_no_queues, simp+)
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (erule cnotification_relation_upd_tcb_no_queues, simp+)
-- "ready queues"
\<comment> \<open>ready queues\<close>
apply (simp add: cready_queues_relation_def Let_def
cready_queues_index_to_C_in_range
seL4_MinPrio_def minDom_def)
@ -3050,7 +3050,7 @@ lemma cancelIPC_ccorres_helper:
apply (clarsimp simp: typ_heap_simps cong: imp_cong split del: if_split simp del: comp_def)
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (intro impI conjI allI)
-- "empty case"
\<comment> \<open>empty case\<close>
apply clarsimp
apply (frule iffD1 [OF tcb_queue_head_empty_iff [OF tcb_queue_relation'_queue_rel]])
apply (rule ballI, erule bspec)
@ -3063,21 +3063,21 @@ lemma cancelIPC_ccorres_helper:
cpspace_relation_def update_ep_map_tos typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
subgoal by (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (rule cpspace_relation_ep_update_ep, assumption+)
subgoal by (simp add: cendpoint_relation_def Let_def EPState_Idle_def)
subgoal by simp
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cnotification_relation_ep_queue [OF invs_sym'], assumption+)
subgoal by simp
apply (erule (1) map_to_ko_atI')
apply (simp add: heap_to_user_data_def Let_def)
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
subgoal by (clarsimp simp: comp_def)
subgoal by (simp add: carch_state_relation_def carch_globals_def
@ -3087,7 +3087,7 @@ lemma cancelIPC_ccorres_helper:
subgoal by (simp add: objBits_simps')
subgoal by (simp add: objBits_simps)
apply assumption
-- "non empty case"
\<comment> \<open>non empty case\<close>
apply clarsimp
apply (frule tcb_queue_head_empty_iff [OF tcb_queue_relation'_queue_rel])
apply (rule ballI, erule bspec)
@ -3100,28 +3100,28 @@ lemma cancelIPC_ccorres_helper:
cpspace_relation_def update_ep_map_tos typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
subgoal by (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (rule cpspace_relation_ep_update_ep, assumption+)
apply (simp add: cendpoint_relation_def Let_def isSendEP_def isRecvEP_def split: endpoint.splits split del: if_split)
-- "recv case"
\<comment> \<open>recv case\<close>
apply (clarsimp simp add: Ptr_ptr_val h_t_valid_clift_Some_iff
tcb_queue_relation'_next_mask tcb_queue_relation'_prev_mask cong: tcb_queue_relation'_cong)
subgoal by (intro impI conjI; simp)
-- "send case"
\<comment> \<open>send case\<close>
apply (clarsimp simp add: Ptr_ptr_val h_t_valid_clift_Some_iff
tcb_queue_relation'_next_mask tcb_queue_relation'_prev_mask cong: tcb_queue_relation'_cong)
subgoal by (intro impI conjI; simp)
subgoal by simp
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cnotification_relation_ep_queue [OF invs_sym'], assumption+)
apply simp
apply (erule (1) map_to_ko_atI')
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
subgoal by (clarsimp simp: comp_def)
subgoal by (simp add: carch_state_relation_def carch_globals_def
@ -3280,7 +3280,7 @@ lemma cancelIPC_ccorres1:
apply (rule_tac xf'=ret__unsigned_' in ccorres_abstract, ceqv)
apply (rule_tac P="rv' = thread_state_to_tsType rv" in ccorres_gen_asm2)
apply wpc
-- "BlockedOnReceive"
\<comment> \<open>BlockedOnReceive\<close>
apply (simp add: word_sle_def "StrictC'_thread_state_defs" ccorres_cond_iffs cong: call_ignore_cong)
apply (fold dc_def)
apply (rule ccorres_rhs_assoc)+
@ -3288,7 +3288,7 @@ lemma cancelIPC_ccorres1:
apply csymbr
apply (rule ccorres_pre_getEndpoint)
apply (rule ccorres_assert)
apply (rule ccorres_symb_exec_r) -- "ptr_get lemmas don't work so well :("
apply (rule ccorres_symb_exec_r) \<comment> \<open>ptr_get lemmas don't work so well :(\<close>
apply (rule ccorres_symb_exec_r)
apply (simp only: fun_app_def simp_list_case_return
return_bind ccorres_seq_skip)
@ -3307,7 +3307,7 @@ lemma cancelIPC_ccorres1:
apply (rule subset_refl)
apply (rule conseqPre, vcg)
apply clarsimp
-- "BlockedOnReply case"
\<comment> \<open>BlockedOnReply case\<close>
apply (simp add: "StrictC'_thread_state_defs" ccorres_cond_iffs
Collect_False Collect_True word_sle_def
cong: call_ignore_cong del: Collect_const)
@ -3370,7 +3370,7 @@ lemma cancelIPC_ccorres1:
apply (wp threadSet_invs_trivial | simp)+
apply (clarsimp simp add: guard_is_UNIV_def tcbReplySlot_def
Kernel_C.tcbReply_def tcbCNodeEntries_def)
-- "BlockedOnNotification"
\<comment> \<open>BlockedOnNotification\<close>
apply (simp add: word_sle_def "StrictC'_thread_state_defs" ccorres_cond_iffs dc_def [symmetric] cong: call_ignore_cong)
apply (rule ccorres_symb_exec_r)
apply (ctac (no_vcg))
@ -3379,18 +3379,18 @@ lemma cancelIPC_ccorres1:
apply (rule subset_refl)
apply (rule conseqPre, vcg)
apply clarsimp
-- "Running, Inactive, and Idle"
\<comment> \<open>Running, Inactive, and Idle\<close>
apply (simp add: word_sle_def "StrictC'_thread_state_defs" ccorres_cond_iffs dc_def [symmetric] cong: call_ignore_cong,
rule ccorres_return_Skip)+
-- "BlockedOnSend"
\<comment> \<open>BlockedOnSend\<close>
apply (simp add: word_sle_def "StrictC'_thread_state_defs" ccorres_cond_iffs dc_def [symmetric] cong: call_ignore_cong)
-- "clag"
\<comment> \<open>clag\<close>
apply (rule ccorres_rhs_assoc)+
apply csymbr
apply csymbr
apply (rule ccorres_pre_getEndpoint)
apply (rule ccorres_assert)
apply (rule ccorres_symb_exec_r) -- "ptr_get lemmas don't work so well :("
apply (rule ccorres_symb_exec_r) \<comment> \<open>ptr_get lemmas don't work so well :(\<close>
apply (rule ccorres_symb_exec_r)
apply (simp only: fun_app_def simp_list_case_return return_bind ccorres_seq_skip)
apply (rule ccorres_rhs_assoc2)
@ -3408,10 +3408,10 @@ lemma cancelIPC_ccorres1:
apply (rule conseqPre, vcg, rule subset_refl)
apply (rule conseqPre, vcg)
apply clarsimp
-- "Restart"
\<comment> \<open>Restart\<close>
apply (simp add: word_sle_def "StrictC'_thread_state_defs" ccorres_cond_iffs dc_def [symmetric] cong: call_ignore_cong,
rule ccorres_return_Skip)
-- "Post wp proofs"
\<comment> \<open>Post wp proofs\<close>
apply vcg
apply clarsimp
apply (rule conseqPre, vcg)

View File

@ -2523,7 +2523,7 @@ lemma getReceiveSlots_ccorres:
apply clarsimp
apply (simp add: cct_relation_def)
apply (case_tac rv, clarsimp)
apply (rule UNIV_I) -- "still a schematic here ..."
apply (rule UNIV_I) \<comment> \<open>still a schematic here ...\<close>
done
@ -2622,7 +2622,7 @@ lemma ccorres_if_cond_throws_break:
assumes abs: "\<forall>s s'. (s, s') \<in> sr \<and> Q s \<and> Q' s' \<longrightarrow> P = (s' \<in> P')"
and ac: "P \<Longrightarrow> ccorres_underlying sr \<Gamma> r xf arrel axf R R' (catchbrk_C # hs) a c"
and bd: "\<not> P \<Longrightarrow> ccorres_underlying sr \<Gamma> r xf arrel axf U U' (catchbrk_C # hs) b d"
and cthrows: "\<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> PT' c {}, UNIV" -- "c always throws"
and cthrows: "\<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> PT' c {}, UNIV" \<comment> \<open>c always throws\<close>
shows "ccorres_underlying sr \<Gamma> r xf arrel axf
(Q and (\<lambda>s. P \<longrightarrow> R s) and (\<lambda>s. \<not> P \<longrightarrow> U s))
(Collect Q' \<inter> {s. (s \<in> P' \<longrightarrow> s \<in> R' \<inter> PT') \<and> (s \<notin> P' \<longrightarrow> s \<in> U')})
@ -2670,7 +2670,7 @@ lemma ccorres_if_cond_throws_break2:
assumes abs: "\<forall>s s'. (s, s') \<in> sr \<and> Q s \<and> Q' s' \<longrightarrow> (\<not> P) = (s' \<in> P')"
and ac: "\<not> P \<Longrightarrow> ccorres_underlying sr \<Gamma> r xf arrel axf R R' (catchbrk_C # hs) a c"
and bd: "P \<Longrightarrow> ccorres_underlying sr \<Gamma> r xf arrel axf U U' (catchbrk_C # hs) b d"
and cthrows: "\<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> PT' c {}, UNIV" -- "c always throws"
and cthrows: "\<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> PT' c {}, UNIV" \<comment> \<open>c always throws\<close>
shows "ccorres_underlying sr \<Gamma> r xf arrel axf
(Q and (\<lambda>s. \<not> P \<longrightarrow> R s) and (\<lambda>s. P \<longrightarrow> U s))
(Collect Q' \<inter> {s. (s \<in> P' \<longrightarrow> s \<in> R' \<inter> PT') \<and> (s \<notin> P' \<longrightarrow> s \<in> U')})
@ -2697,7 +2697,7 @@ lemma ccorres_split_when_throwError_cond_break:
R R' (catchbrk_C # hs) (throwError e) c"
and bd: "\<not> P \<Longrightarrow> ccorres_underlying sr \<Gamma> r xf ar axf
U U' (catchbrk_C # hs) b d"
and cthrows: "\<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> PT' c {}, UNIV" -- "c always throws"
and cthrows: "\<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> PT' c {}, UNIV" \<comment> \<open>c always throws\<close>
shows "ccorres_underlying sr \<Gamma> r xf ar axf
(Q and (\<lambda>s. P \<longrightarrow> R s) and (\<lambda>s. \<not> P \<longrightarrow> U s))
(Collect Q' \<inter> {s. (s \<in> P' \<longrightarrow> s \<in> R' \<inter> PT') \<and> (s \<notin> P' \<longrightarrow> s \<in> U')})
@ -2975,7 +2975,7 @@ next
apply (clarsimp simp: Collect_const_mem)
apply (rule sym, rule from_bool_neq_0)
-- "case where a badge is sent"
\<comment> \<open>case where a badge is sent\<close>
apply (rule ccorres_rhs_assoc)+
apply csymbr
apply (simp only: Let_def liftE_bindE withoutFailure_def fun_app_def)
@ -2996,7 +2996,7 @@ next
apply (simp split del: if_split)
apply (vcg exspec=setExtraBadge_modifies)
-- "case where a cap is sent (or rather a send is attempted)"
\<comment> \<open>case where a cap is sent (or rather a send is attempted)\<close>
apply (simp add: split_def del: Collect_const
cong: call_ignore_cong split del: if_split)
apply (rule ccorres_rhs_assoc)+
@ -3013,7 +3013,7 @@ next
apply (rule ccorres_cond_false_seq)
apply (simp)
-- "case not diminish"
\<comment> \<open>case not diminish\<close>
apply (rule ccorres_split_nothrowE)
apply (rule unifyFailure_ccorres)
apply (ctac add: deriveCap_ccorres')
@ -3081,7 +3081,7 @@ next
apply (wp deriveCap_derived is_the_ep_deriveCap)
apply (vcg exspec=deriveCap_modifies)
--"remaining non ccorres subgoals"
\<comment> \<open>remaining non ccorres subgoals\<close>
apply (clarsimp simp: Collect_const_mem if_1_0_0
split del: if_split)
apply (rule_tac Q'="\<lbrace>\<acute>ret__int = from_bool (cap_get_tag cap' = scast cap_endpoint_cap
@ -4787,7 +4787,7 @@ lemma sendIPC_dequeue_ccorres_helper:
apply (intro conjI impI allI)
apply (fastforce simp: h_t_valid_clift)
apply (fastforce simp: h_t_valid_clift)
-- "empty case"
\<comment> \<open>empty case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (frule iffD1 [OF tcb_queue_head_empty_iff
@ -4804,21 +4804,21 @@ lemma sendIPC_dequeue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (rule cpspace_relation_ep_update_ep, assumption+)
apply (simp add: cendpoint_relation_def Let_def EPState_Idle_def
tcb_queue_relation'_def)
apply simp
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cnotification_relation_ep_queue [OF invs_sym'], assumption+)
apply simp
apply (erule (1) map_to_ko_atI')
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (simp add: carch_state_relation_def typ_heap_simps')
@ -4828,7 +4828,7 @@ lemma sendIPC_dequeue_ccorres_helper:
apply (simp add: objBits_simps)
apply assumption
apply (clarsimp simp: cendpoint_relation_def Let_def tcb_queue_relation'_def)
-- "non-empty case"
\<comment> \<open>non-empty case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (frule tcb_queue_head_empty_iff [OF tcb_queue_relation'_queue_rel])
@ -4844,10 +4844,10 @@ lemma sendIPC_dequeue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (rule cpspace_relation_ep_update_ep, assumption+)
apply (clarsimp simp: cendpoint_relation_def Let_def
isRecvEP_def isSendEP_def
@ -4860,13 +4860,13 @@ lemma sendIPC_dequeue_ccorres_helper:
split del: if_split)
apply (clarsimp split: if_split)
apply simp
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cnotification_relation_ep_queue [OF invs_sym'], assumption+)
apply simp
apply (erule (1) map_to_ko_atI')
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (simp add: carch_state_relation_def typ_heap_simps')
@ -5189,7 +5189,7 @@ lemma sendIPC_enqueue_ccorres_helper:
apply (fastforce simp: h_t_valid_clift)
apply (fastforce simp: h_t_valid_clift)
apply (case_tac ep, simp_all add: EPState_Idle_def EPState_Send_def)[1]
-- "IdleEP case"
\<comment> \<open>IdleEP case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (simp add: setEndpoint_def split_def)
@ -5199,16 +5199,16 @@ lemma sendIPC_enqueue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (rule cpspace_relation_ep_update_ep', assumption+)
apply (clarsimp simp: cendpoint_relation_def Let_def
mask_def [where n=2] EPState_Send_def)
apply (clarsimp simp: tcb_queue_relation'_def is_aligned_neg_mask)
apply (simp add: isSendEP_def isRecvEP_def)
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cnotification_relation_ep_queue, assumption+)
@ -5219,7 +5219,7 @@ lemma sendIPC_enqueue_ccorres_helper:
apply (simp only:projectKOs injectKO_ep objBits_simps)
apply clarsimp
apply (clarsimp simp: obj_at'_def projectKOs)
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (simp add: carch_state_relation_def typ_heap_simps')
@ -5228,7 +5228,7 @@ lemma sendIPC_enqueue_ccorres_helper:
apply (simp add: objBits_simps')
apply (simp add: objBits_simps)
apply assumption
-- "SendEP case"
\<comment> \<open>SendEP case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (simp add: setEndpoint_def split_def)
@ -5238,10 +5238,10 @@ lemma sendIPC_enqueue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (rule cpspace_relation_ep_update_ep', assumption+)
apply (clarsimp simp: cendpoint_relation_def Let_def
mask_def [where n=2] EPState_Send_def
@ -5250,7 +5250,7 @@ lemma sendIPC_enqueue_ccorres_helper:
valid_ep'_def
dest: tcb_queue_relation_next_not_NULL)
apply (simp add: isSendEP_def isRecvEP_def)
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cnotification_relation_ep_queue, assumption+)
@ -5260,7 +5260,7 @@ lemma sendIPC_enqueue_ccorres_helper:
apply (erule(2) map_to_ko_at_updI')
apply (clarsimp simp: objBitsKO_def)
apply (clarsimp simp: obj_at'_def projectKOs)
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (simp add: carch_state_relation_def typ_heap_simps')
@ -5308,7 +5308,7 @@ lemma sendIPC_ccorres [corres]:
and (\<lambda>s. \<forall>d p. thread \<notin> set (ksReadyQueues s (d, p)))"
in ccorres_guard_imp2 [where A'=UNIV])
apply wpc
-- "RecvEP case"
\<comment> \<open>RecvEP case\<close>
apply (rule ccorres_cond_false)
apply (rule ccorres_cond_true)
apply (intro ccorres_rhs_assoc)
@ -5421,15 +5421,15 @@ lemma sendIPC_ccorres [corres]:
option_to_ptr_def option_to_0_def
split: bool.split_asm)
-- "IdleEP case"
\<comment> \<open>IdleEP case\<close>
apply (rule ccorres_cond_true)
apply (rule ccorres_cond)
apply (clarsimp simp: from_bool_def split: bool.split)
-- "blocking case"
\<comment> \<open>blocking case\<close>
apply (intro ccorres_rhs_assoc)
apply csymbr
apply (simp only:)
-- "apply (ctac (trace, no_vcg,c_lines 6) add: sendIPC_block_ccorres_helper)"
\<comment> \<open>apply (ctac (trace, no_vcg,c_lines 6) add: sendIPC_block_ccorres_helper)\<close>
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
@ -5449,14 +5449,14 @@ lemma sendIPC_ccorres [corres]:
apply wp
apply (clarsimp simp: guard_is_UNIV_def)
apply (rule ccorres_return_Skip)
-- "SendEP case"
\<comment> \<open>SendEP case\<close>
apply (rule ccorres_cond_true)
apply (rule ccorres_cond)
apply (clarsimp simp: from_bool_def split: bool.split)
-- "blocking case"
\<comment> \<open>blocking case\<close>
apply (intro ccorres_rhs_assoc)
apply csymbr
-- "apply (ctac (no_vcg,c_lines 6) add: sendIPC_block_ccorres_helper)"
\<comment> \<open>apply (ctac (no_vcg,c_lines 6) add: sendIPC_block_ccorres_helper)\<close>
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
@ -5622,7 +5622,7 @@ lemma receiveIPC_enqueue_ccorres_helper:
apply (fastforce simp: h_t_valid_clift)
apply (fastforce simp: h_t_valid_clift)
apply (case_tac ep, simp_all add: EPState_Idle_def EPState_Recv_def)[1]
-- "RecvEP case"
\<comment> \<open>RecvEP case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (simp add: setEndpoint_def split_def)
@ -5632,10 +5632,10 @@ lemma receiveIPC_enqueue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (rule cpspace_relation_ep_update_ep', assumption+)
apply (clarsimp simp: cendpoint_relation_def Let_def
mask_def [where n=2] EPState_Recv_def
@ -5644,7 +5644,7 @@ lemma receiveIPC_enqueue_ccorres_helper:
valid_ep'_def
dest: tcb_queue_relation_next_not_NULL)
apply (simp add: isSendEP_def isRecvEP_def)
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cnotification_relation_ep_queue, assumption+)
@ -5654,7 +5654,7 @@ lemma receiveIPC_enqueue_ccorres_helper:
apply (erule(2) map_to_ko_at_updI')
apply (clarsimp simp: objBitsKO_def)
apply (clarsimp simp: obj_at'_def projectKOs)
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (simp add: carch_state_relation_def typ_heap_simps')
@ -5663,7 +5663,7 @@ lemma receiveIPC_enqueue_ccorres_helper:
apply (simp add: objBits_simps')
apply (simp add: objBits_simps)
apply assumption
-- "IdleEP case"
\<comment> \<open>IdleEP case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (simp add: setEndpoint_def split_def)
@ -5673,16 +5673,16 @@ lemma receiveIPC_enqueue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (rule cpspace_relation_ep_update_ep', assumption+)
apply (clarsimp simp: cendpoint_relation_def Let_def
mask_def [where n=2] EPState_Recv_def)
apply (clarsimp simp: tcb_queue_relation'_def is_aligned_neg_mask)
apply (simp add: isSendEP_def isRecvEP_def)
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cnotification_relation_ep_queue, assumption+)
@ -5692,7 +5692,7 @@ lemma receiveIPC_enqueue_ccorres_helper:
apply (erule(2) map_to_ko_at_updI')
apply (clarsimp simp: objBitsKO_def)
apply (clarsimp simp: obj_at'_def projectKOs)
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (simp add: carch_state_relation_def typ_heap_simps')
@ -5745,7 +5745,7 @@ lemma receiveIPC_dequeue_ccorres_helper:
apply (intro conjI impI allI)
apply (fastforce simp: h_t_valid_clift)
apply (fastforce simp: h_t_valid_clift)
-- "empty case"
\<comment> \<open>empty case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (frule iffD1 [OF tcb_queue_head_empty_iff
@ -5762,21 +5762,21 @@ lemma receiveIPC_dequeue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (rule cpspace_relation_ep_update_ep, assumption+)
apply (simp add: cendpoint_relation_def Let_def EPState_Idle_def
tcb_queue_relation'_def)
apply simp
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cnotification_relation_ep_queue [OF invs_sym'], assumption+)
apply simp
apply (erule (1) map_to_ko_atI')
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (simp add: carch_state_relation_def typ_heap_simps')
@ -5786,7 +5786,7 @@ lemma receiveIPC_dequeue_ccorres_helper:
apply (simp add: objBits_simps)
apply assumption
apply (clarsimp simp: cendpoint_relation_def Let_def tcb_queue_relation'_def)
-- "non-empty case"
\<comment> \<open>non-empty case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (frule tcb_queue_head_empty_iff [OF tcb_queue_relation'_queue_rel])
@ -5802,10 +5802,10 @@ lemma receiveIPC_dequeue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (rule cpspace_relation_ep_update_ep, assumption+)
apply (clarsimp simp: cendpoint_relation_def Let_def
isRecvEP_def isSendEP_def
@ -5818,13 +5818,13 @@ lemma receiveIPC_dequeue_ccorres_helper:
split del: if_split)
apply (clarsimp split: if_split)
apply simp
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cnotification_relation_ep_queue [OF invs_sym'], assumption+)
apply simp
apply (erule (1) map_to_ko_atI')
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (simp add: carch_state_relation_def typ_heap_simps')
@ -5874,12 +5874,12 @@ lemma completeSignal_ccorres:
split: Structures_H.ntfn.splits)
apply ceqv
apply wpc
-- "IdleNtfn case"
\<comment> \<open>IdleNtfn case\<close>
apply (clarsimp simp: NtfnState_Idle_def NtfnState_Active_def)
apply csymbr
apply (rule ccorres_cond_false)
apply (rule ccorres_fail)
-- "ActiveNtfn case"
\<comment> \<open>ActiveNtfn case\<close>
apply (clarsimp, csymbr, rule ccorres_cond_true)
apply (rule ccorres_rhs_assoc)+
apply (rule ccorres_rhs_assoc2)
@ -5914,7 +5914,7 @@ lemma completeSignal_ccorres:
apply (clarsimp simp: guard_is_UNIV_def ARM_H.badgeRegister_def
ARM.badgeRegister_def Kernel_C.badgeRegister_def
Kernel_C.R0_def)
-- "WaitingNtfn case"
\<comment> \<open>WaitingNtfn case\<close>
apply (clarsimp simp: NtfnState_Active_def NtfnState_Waiting_def)
apply csymbr
apply (rule ccorres_cond_false)
@ -6018,7 +6018,7 @@ lemma receiveIPC_ccorres [corres]:
and ko_at' ep (capEPPtr cap)"
in ccorres_guard_imp2 [where A'=UNIV])
apply wpc
-- "RecvEP case"
\<comment> \<open>RecvEP case\<close>
apply (rule ccorres_cond_true)
apply csymbr
apply (simp only: case_bool_If from_bool_neq_0)
@ -6047,7 +6047,7 @@ lemma receiveIPC_ccorres [corres]:
apply (clarsimp simp: guard_is_UNIV_def)
apply simp
apply (ctac add: doNBRecvFailedTransfer_ccorres[unfolded dc_def])
-- "IdleEP case"
\<comment> \<open>IdleEP case\<close>
apply (rule ccorres_cond_true)
apply csymbr
apply (simp only: case_bool_If from_bool_neq_0)
@ -6073,7 +6073,7 @@ lemma receiveIPC_ccorres [corres]:
apply (clarsimp simp: guard_is_UNIV_def)
apply simp
apply (ctac add: doNBRecvFailedTransfer_ccorres[unfolded dc_def])
-- "SendEP case"
\<comment> \<open>SendEP case\<close>
apply (thin_tac "isBlockinga = from_bool P" for P)
apply (rule ccorres_cond_false)
apply (rule ccorres_cond_true)
@ -6324,7 +6324,7 @@ lemma sendSignal_dequeue_ccorres_helper:
apply (intro conjI impI allI)
apply (fastforce simp: h_t_valid_clift)
apply (fastforce simp: h_t_valid_clift)
-- "empty case"
\<comment> \<open>empty case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (frule iffD1 [OF tcb_queue_head_empty_iff
@ -6342,21 +6342,21 @@ lemma sendSignal_dequeue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cendpoint_relation_ntfn_queue [OF invs_sym'], assumption+)
apply simp+
apply (erule (1) map_to_ko_atI')
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (rule cpspace_relation_ntfn_update_ntfn, assumption+)
apply (simp add: cnotification_relation_def Let_def NtfnState_Idle_def
tcb_queue_relation'_def)
apply simp
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (simp add: carch_state_relation_def typ_heap_simps')
@ -6367,7 +6367,7 @@ lemma sendSignal_dequeue_ccorres_helper:
apply assumption
apply (clarsimp simp: cnotification_relation_def Let_def
tcb_queue_relation'_def)
-- "non-empty case"
\<comment> \<open>non-empty case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (frule tcb_queue_head_empty_iff [OF tcb_queue_relation'_queue_rel])
@ -6384,16 +6384,16 @@ lemma sendSignal_dequeue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cendpoint_relation_ntfn_queue [OF invs_sym'], assumption+)
apply simp+
apply (erule (1) map_to_ko_atI')
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (rule cpspace_relation_ntfn_update_ntfn, assumption+)
apply (clarsimp simp: cnotification_relation_def Let_def
isWaitingNtfn_def
@ -6406,7 +6406,7 @@ lemma sendSignal_dequeue_ccorres_helper:
split del: if_split)
apply (clarsimp split: if_split)
apply simp
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (simp add: carch_state_relation_def typ_heap_simps')
@ -6481,7 +6481,7 @@ lemma sendSignal_ccorres [corres]:
option_to_ctcb_ptr (ntfnBoundTCB ntfn) \<noteq> NULL"
in ccorres_gen_asm)
apply wpc
-- "IdleNtfn case"
\<comment> \<open>IdleNtfn case\<close>
apply (rule ccorres_cond_true)
apply (rule ccorres_rhs_assoc)+
apply (rule_tac xf'=ret__unsigned_'
@ -6529,7 +6529,7 @@ lemma sendSignal_ccorres [corres]:
Collect_const_mem)
apply (case_tac ts, simp_all add: receiveBlocked_def typ_heap_simps
cthread_state_relation_def "StrictC'_thread_state_defs")[1]
-- "ActiveNtfn case"
\<comment> \<open>ActiveNtfn case\<close>
apply (rename_tac old_badge)
apply (rule ccorres_cond_false)
apply (rule ccorres_cond_false)
@ -6559,7 +6559,7 @@ lemma sendSignal_ccorres [corres]:
apply (simp add: objBits_simps')
apply (simp add: objBits_simps)
apply assumption
-- "WaitingNtfn case"
\<comment> \<open>WaitingNtfn case\<close>
apply (rule ccorres_cond_false)
apply (rule ccorres_cond_true)
apply wpc
@ -6754,7 +6754,7 @@ lemma receiveSignal_enqueue_ccorres_helper:
apply (fastforce simp: h_t_valid_clift)
apply (fastforce simp: h_t_valid_clift)
apply (case_tac "ntfnObj ntfn", simp_all add: NtfnState_Idle_def NtfnState_Waiting_def)[1]
-- "IdleNtfn case"
\<comment> \<open>IdleNtfn case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (simp add: setNotification_def split_def)
@ -6764,10 +6764,10 @@ lemma receiveSignal_enqueue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cendpoint_relation_ntfn_queue, assumption+)
@ -6777,7 +6777,7 @@ lemma receiveSignal_enqueue_ccorres_helper:
apply (erule(2) map_to_ko_at_updI')
apply (clarsimp simp: objBitsKO_def)
apply (clarsimp simp: obj_at'_def projectKOs)
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (rule cpspace_relation_ntfn_update_ntfn', assumption+)
apply (case_tac "ntfn", simp_all)[1]
apply (clarsimp simp: cnotification_relation_def Let_def
@ -6786,7 +6786,7 @@ lemma receiveSignal_enqueue_ccorres_helper:
valid_ntfn'_def
dest: tcb_queue_relation_next_not_NULL)
apply (simp add: isWaitingNtfn_def)
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
subgoal by (clarsimp simp: comp_def)
apply (simp add: carch_state_relation_def typ_heap_simps')
@ -6795,7 +6795,7 @@ lemma receiveSignal_enqueue_ccorres_helper:
apply (simp add: objBits_simps')
apply (simp add: objBits_simps)
apply assumption
-- "WaitingNtfn case"
\<comment> \<open>WaitingNtfn case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (simp add: setNotification_def split_def)
@ -6805,10 +6805,10 @@ lemma receiveSignal_enqueue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cendpoint_relation_ntfn_queue, assumption+)
@ -6818,7 +6818,7 @@ lemma receiveSignal_enqueue_ccorres_helper:
apply (erule(2) map_to_ko_at_updI')
apply (clarsimp simp: objBitsKO_def)
apply (clarsimp simp: obj_at'_def projectKOs)
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (rule cpspace_relation_ntfn_update_ntfn', assumption+)
apply (case_tac "ntfn", simp_all)[1]
apply (clarsimp simp: cnotification_relation_def Let_def
@ -6826,7 +6826,7 @@ lemma receiveSignal_enqueue_ccorres_helper:
split: if_split)
apply (fastforce simp: tcb_queue_relation'_def is_aligned_neg_mask)
apply (simp add: isWaitingNtfn_def)
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (simp add: carch_state_relation_def typ_heap_simps')
@ -6874,7 +6874,7 @@ lemma receiveSignal_ccorres [corres]:
apply ceqv
apply wpc
-- "IdleNtfn case"
\<comment> \<open>IdleNtfn case\<close>
apply (rule ccorres_cond_true)
apply csymbr
apply (simp only: case_bool_If from_bool_neq_0)
@ -6899,7 +6899,7 @@ lemma receiveSignal_ccorres [corres]:
apply (clarsimp simp: guard_is_UNIV_def)
apply simp
apply (ctac add: doNBRecvFailedTransfer_ccorres[unfolded dc_def])
-- "ActiveNtfn case"
\<comment> \<open>ActiveNtfn case\<close>
apply (rename_tac badge)
apply (rule ccorres_cond_false)
apply (rule ccorres_cond_true)
@ -6940,7 +6940,7 @@ lemma receiveSignal_ccorres [corres]:
apply wp
apply (clarsimp simp: guard_is_UNIV_def)
apply (clarsimp simp: ARM.badgeRegister_def Kernel_C.R0_def)
-- "WaitingNtfn case"
\<comment> \<open>WaitingNtfn case\<close>
apply (rename_tac list)
apply (rule ccorres_cond_true)
apply csymbr

View File

@ -248,7 +248,7 @@ lemma handleSyscall_ccorres:
apply (rule ccorres_split_nothrow_novcg)
apply wpc
prefer 3
-- "SysSend"
\<comment> \<open>SysSend\<close>
apply (clarsimp simp: syscall_from_H_def syscall_defs)
apply (rule ccorres_cond_empty |rule ccorres_cond_univ)+
apply (simp add: handleSend_def)
@ -289,7 +289,7 @@ lemma handleSyscall_ccorres:
apply clarsimp
apply (vcg exspec=handleInvocation_modifies)
prefer 3
-- "SysNBSend"
\<comment> \<open>SysNBSend\<close>
apply (clarsimp simp: syscall_from_H_def syscall_defs)
apply (rule ccorres_cond_empty |rule ccorres_cond_univ)+
apply (simp add: handleSend_def)
@ -328,7 +328,7 @@ lemma handleSyscall_ccorres:
apply (simp add: invs'_def valid_state'_def)
apply clarsimp
apply (vcg exspec=handleInvocation_modifies)
-- "SysCall"
\<comment> \<open>SysCall\<close>
apply (clarsimp simp: syscall_from_H_def syscall_defs)
apply (rule ccorres_cond_empty |rule ccorres_cond_univ)+
apply (simp add: handleCall_def)
@ -368,7 +368,7 @@ lemma handleSyscall_ccorres:
apply clarsimp
apply (vcg exspec=handleInvocation_modifies)
prefer 2
-- "SysRecv"
\<comment> \<open>SysRecv\<close>
apply (clarsimp simp: syscall_from_H_def syscall_defs)
apply (rule ccorres_cond_empty |rule ccorres_cond_univ)+
apply (simp add: liftE_bind)
@ -377,7 +377,7 @@ lemma handleSyscall_ccorres:
apply (rule ccorres_returnOk_skip[unfolded returnOk_def, simplified])
apply wp
prefer 2
-- "SysReply"
\<comment> \<open>SysReply\<close>
apply (clarsimp simp: syscall_from_H_def syscall_defs)
apply (rule ccorres_cond_empty |rule ccorres_cond_univ)+
apply (simp add: liftE_bind)
@ -385,7 +385,7 @@ lemma handleSyscall_ccorres:
apply (ctac (no_vcg) add: handleReply_ccorres)
apply (rule ccorres_returnOk_skip[unfolded returnOk_def, simplified])
apply wp
-- "SysReplyRecv"
\<comment> \<open>SysReplyRecv\<close>
apply (clarsimp simp: syscall_from_H_def syscall_defs)
apply (rule ccorres_cond_empty |rule ccorres_cond_univ)+
apply (simp add: liftE_bind bind_assoc)
@ -401,7 +401,7 @@ lemma handleSyscall_ccorres:
in hoare_post_imp)
apply (simp add: ct_in_state'_def)
apply (wp handleReply_sane handleReply_ct_not_ksQ)
-- "SysYield"
\<comment> \<open>SysYield\<close>
apply (clarsimp simp: syscall_from_H_def syscall_defs)
apply (rule ccorres_cond_empty |rule ccorres_cond_univ)+
apply (simp add: liftE_bind)
@ -409,7 +409,7 @@ lemma handleSyscall_ccorres:
apply (ctac (no_vcg) add: handleYield_ccorres)
apply (rule ccorres_returnOk_skip[unfolded returnOk_def, simplified])
apply wp
-- "SysNBRecv"
\<comment> \<open>SysNBRecv\<close>
apply (clarsimp simp: syscall_from_H_def syscall_defs)
apply (rule ccorres_cond_empty |rule ccorres_cond_univ)+
apply (simp add: liftE_bind)
@ -417,7 +417,7 @@ lemma handleSyscall_ccorres:
apply (ctac (no_vcg) add: handleRecv_ccorres)
apply (rule ccorres_returnOk_skip[unfolded returnOk_def, simplified])
apply wp
-- " rest of body"
\<comment> \<open>rest of body\<close>
apply ceqv
apply (ctac (no_vcg) add: schedule_ccorres)
apply (rule ccorres_add_return2)

View File

@ -3487,7 +3487,7 @@ proof -
note hacky_cte = retype_ctes_helper [where sz = "objBitsKO kotcb" and ko = kotcb and ptr = "ctcb_ptr_to_tcb_ptr p",
OF pal pds pno al _ _ mko, simplified new_cap_addrs_def, simplified]
-- "Ugh"
\<comment> \<open>Ugh\<close>
moreover have
"\<And>y. y \<in> ptr_val ` (CTypesDefs.ptr_add (cte_Ptr (ctcb_ptr_to_tcb_ptr p)) \<circ> of_nat) ` {k. k < 5}
= (y && ~~ mask tcbBlockSizeBits = ctcb_ptr_to_tcb_ptr p \<and> y && mask tcbBlockSizeBits \<in> dom tcb_cte_cases)"
@ -3592,7 +3592,7 @@ proof -
eval_nat_numeral ThreadState_Inactive_def)
apply (simp add: ccontext_relation_def newContext_def2 carch_tcb_relation_def)
apply rule
apply (case_tac r, simp_all add: "StrictC'_register_defs" eval_nat_numeral atcbContext_def newArchTCB_def newContext_def initContext_def)[1] -- "takes ages"
apply (case_tac r, simp_all add: "StrictC'_register_defs" eval_nat_numeral atcbContext_def newArchTCB_def newContext_def initContext_def)[1] \<comment> \<open>takes ages\<close>
apply (simp add: thread_state_lift_def eval_nat_numeral atcbContextGet_def)+
apply (simp add: timeSlice_def)
apply (simp add: cfault_rel_def seL4_Fault_lift_def seL4_Fault_get_tag_def Let_def
@ -3718,15 +3718,15 @@ proof -
apply (simp add: rl kotcb_def projectKOs rl_tcb rl_cte)
apply (elim conjE)
apply (intro conjI)
-- "cte"
\<comment> \<open>cte\<close>
apply (erule cmap_relation_retype2)
apply (simp add:ccte_relation_nullCap nullMDBNode_def nullPointer_def)
-- "tcb"
\<comment> \<open>tcb\<close>
apply (erule cmap_relation_updI2 [where dest = "ctcb_ptr_to_tcb_ptr p" and f = "tcb_ptr_to_ctcb_ptr", simplified])
apply (rule map_comp_simps)
apply (rule pks)
apply (rule tcb_rel)
-- "ep"
\<comment> \<open>ep\<close>
apply (erule iffD2 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply (simp add: cendpoint_relation_def Let_def)
apply (subst endpoint.case_cong)
@ -3735,7 +3735,7 @@ proof -
apply (simp add: tcb_queue_update_other' del: tcb_queue_relation'_empty)
apply (simp add: tcb_queue_update_other' ep2)
apply clarsimp
-- "ntfn"
\<comment> \<open>ntfn\<close>
apply (erule iffD2 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply (simp add: cnotification_relation_def Let_def)
apply (subst ntfn.case_cong)
@ -5534,7 +5534,7 @@ proof -
framesize_to_H_def cap_to_H_simps cap_small_frame_cap_lift
vmrights_to_H_def mask_def vm_rights_defs)
-- "Page objects: could possibly fix the duplication here"
\<comment> \<open>Page objects: could possibly fix the duplication here\<close>
apply (cinit' lift: t_' regionBase_' userSize_' deviceMemory_')
apply (simp add: object_type_from_H_def Kernel_C_defs)
apply (simp add: ccorres_cond_univ_iff ccorres_cond_empty_iff
@ -5616,7 +5616,7 @@ proof -
cl_valid_cap_def c_valid_cap_def
is_aligned_neg_mask_eq_concrete[THEN sym])
-- "PageTableObject"
\<comment> \<open>PageTableObject\<close>
apply (cinit' lift: t_' regionBase_' userSize_' deviceMemory_')
apply (simp add: object_type_from_H_def Kernel_C_defs)
apply (simp add: ccorres_cond_univ_iff ccorres_cond_empty_iff
@ -5649,7 +5649,7 @@ proof -
Kernel_C.VMKernelOnly_def Kernel_C.VMReadOnly_def)
apply (simp add: to_bool_def false_def isFrameType_def)
-- "PageDirectoryObject"
\<comment> \<open>PageDirectoryObject\<close>
apply (cinit' lift: t_' regionBase_' userSize_' deviceMemory_')
apply (simp add: object_type_from_H_def Kernel_C_defs)
apply (simp add: ccorres_cond_univ_iff ccorres_cond_empty_iff

View File

@ -101,7 +101,7 @@ lemma cap_get_tag_isCap0:
apply (erule ccap_relationE)
apply (simp add: cap_to_H_def cap_lift_def Let_def isArchCap_tag_def2 isArchCap_def)
apply (clarsimp simp: isCap_simps cap_tag_defs word_le_nat_alt pageSize_def Let_def
split: if_split_asm) -- "takes a while"
split: if_split_asm) \<comment> \<open>takes a while\<close>
done
@ -1499,7 +1499,7 @@ lemma cmap_relation_cong:
apply (rule Some_the [where f = cm'])
apply (erule subsetD)
apply (erule imageI)
-- "clag"
\<comment> \<open>clag\<close>
apply simp
apply (erule conjE)
apply (drule equalityD1)
@ -2059,7 +2059,7 @@ lemma cap_get_tag_isCap_ArchObject0:
apply -
apply (erule ccap_relationE)
apply (simp add: cap_to_H_def cap_lift_def Let_def isArchCap_def)
apply (clarsimp simp: isCap_simps cap_tag_defs word_le_nat_alt pageSize_def Let_def split: if_split_asm) -- "takes a while"
apply (clarsimp simp: isCap_simps cap_tag_defs word_le_nat_alt pageSize_def Let_def split: if_split_asm) \<comment> \<open>takes a while\<close>
done
lemma cap_get_tag_isCap_ArchObject:

View File

@ -266,7 +266,7 @@ proof -
apply (rule_tac R="\<lambda>s. l1 = ksReadyQueuesL1Bitmap s curdom \<and> curdom \<le> maxDomain"
in ccorres_cond)
subgoal by (fastforce dest!: rf_sr_cbitmap_L1_relation simp: cbitmap_L1_relation_def)
prefer 2 -- "switchToIdleThread"
prefer 2 \<comment> \<open>switchToIdleThread\<close>
apply (ctac(no_vcg) add: switchToIdleThread_ccorres)
apply clarsimp
apply (rule ccorres_rhs_assoc)+

View File

@ -926,7 +926,7 @@ lemma cready_queues_relation_null_queue_ptrs:
apply (clarsimp simp: tcb_null_ep_ptrs_def)
apply (case_tac z, case_tac a)
apply simp
-- "clag"
\<comment> \<open>clag\<close>
apply (rule ext)
apply (case_tac "mp' x")
apply (frule compD [OF same])

View File

@ -3874,7 +3874,7 @@ lemma bindNotification_ccorres:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (rule cpspace_relation_ntfn_update_ntfn, assumption+)
apply (clarsimp simp: cnotification_relation_def Let_def
mask_def [where n=2] NtfnState_Waiting_def)
@ -4092,7 +4092,7 @@ lemma decodeBindNotification_ccorres:
apply (rule_tac P="\<lambda>s. ksCurThread s = thread" in ccorres_cross_over_guard)
apply (simp add: bindE_bind_linearise del: Collect_const)
apply wpc
-- "IdleNtfn"
\<comment> \<open>IdleNtfn\<close>
apply (simp add: case_option_If del: Collect_const)
apply (rule ccorres_Cond_rhs_Seq)
apply (clarsimp simp: isWaitingNtfn_def from_bool_neq_0)

View File

@ -2361,7 +2361,7 @@ lemma unmapPage_ccorres:
apply (simp add: liftE_bindE Collect_False bind_bindE_assoc
del: Collect_const)
apply (rule ccorres_splitE_novcg[where r'=dc and xf'=xfdc])
-- "ARMSmallPage"
\<comment> \<open>ARMSmallPage\<close>
apply (rule ccorres_Cond_rhs)
apply (simp add: gen_framesize_to_H_def dc_def[symmetric])
apply (rule ccorres_rhs_assoc)+
@ -2403,7 +2403,7 @@ lemma unmapPage_ccorres:
apply (wp)
apply simp
apply (vcg exspec=lookupPTSlot_modifies)
-- "ARMLargePage"
\<comment> \<open>ARMLargePage\<close>
apply (rule ccorres_Cond_rhs)
apply (simp add: gen_framesize_to_H_def dc_def[symmetric]
largePagePTEOffsets_def pteBits_def)
@ -2493,7 +2493,7 @@ lemma unmapPage_ccorres:
apply (wp lookupPTSlot_inv Arch_R.lookupPTSlot_aligned
lookupPTSlot_page_table_at' | simp add: K_def)+
apply (vcg exspec=lookupPTSlot_modifies)
-- "ARMSection"
\<comment> \<open>ARMSection\<close>
apply (rule ccorres_Cond_rhs)
apply (rule ccorres_rhs_assoc)+
apply (csymbr, csymbr)
@ -2528,7 +2528,7 @@ lemma unmapPage_ccorres:
apply simp
apply wp
apply (simp add: guard_is_UNIV_def)
-- "ARMSuperSection"
\<comment> \<open>ARMSuperSection\<close>
apply (rule ccorres_Cond_rhs)
apply (rule ccorres_rhs_assoc)+
apply csymbr

View File

@ -99,11 +99,11 @@ definition
definition
ep_at_C' :: "word32 \<Rightarrow> heap_raw_state \<Rightarrow> bool"
where
"ep_at_C' p h \<equiv> Ptr p \<in> dom (clift h :: endpoint_C typ_heap)" -- "endpoint_lift is total"
"ep_at_C' p h \<equiv> Ptr p \<in> dom (clift h :: endpoint_C typ_heap)" \<comment> \<open>endpoint_lift is total\<close>
definition
ntfn_at_C' :: "word32 \<Rightarrow> heap_raw_state \<Rightarrow> bool"
where -- "notification_lift is total"
where \<comment> \<open>notification_lift is total\<close>
"ntfn_at_C' p h \<equiv> Ptr p \<in> dom (clift h :: notification_C typ_heap)"
definition
@ -209,8 +209,8 @@ abbreviation
ARMSuperSection :: "vmpage_size" where
"ARMSuperSection == ARM.ARMSuperSection"
-- "ARMSmallFrame is treated in a separate cap in C,
so needs special treatment in ccap_relation"
\<comment> \<open>ARMSmallFrame is treated in a separate cap in C,
so needs special treatment in ccap_relation\<close>
definition
framesize_to_H:: "word32 \<Rightarrow> vmpage_size" where
"framesize_to_H c \<equiv>
@ -218,7 +218,7 @@ framesize_to_H:: "word32 \<Rightarrow> vmpage_size" where
else if c = scast Kernel_C.ARMSection then ARMSection
else ARMSuperSection"
-- "Use this for results of generic_frame_cap_get_capFSize"
\<comment> \<open>Use this for results of generic_frame_cap_get_capFSize\<close>
definition
gen_framesize_to_H:: "word32 \<Rightarrow> vmpage_size" where
"gen_framesize_to_H c \<equiv>

View File

@ -1960,7 +1960,7 @@ lemma createMappingEntries_valid_pde_slots'2:
apply (rule hoare_pre, wp)
apply (clarsimp simp: vmsz_aligned'_def page_directory_at'_def lookup_pd_slot_def)
apply (rule conjI)
subgoal -- "valid_pde_mapping_offset'"
subgoal \<comment> \<open>valid_pde_mapping_offset'\<close>
apply (clarsimp simp: superSectionPDEOffsets_def length_upto_enum_step table_bits_defs)
apply (clarsimp simp: upto_enum_step_def upto_enum_def comp_def)
apply (clarsimp simp: linorder_not_less field_simps mask_add_aligned)
@ -1968,7 +1968,7 @@ lemma createMappingEntries_valid_pde_slots'2:
apply (rule word_of_nat_le, simp)
done
apply (rule conjI)
subgoal -- "pde_at'"
subgoal \<comment> \<open>pde_at'\<close>
apply (clarsimp simp: superSectionPDEOffsets_def length_upto_enum_step table_bits_defs)
apply (clarsimp simp:upto_enum_step_def upto_enum_def hd_map_simp comp_def)
apply (simp add: vaddr_segment_nonsense6)
@ -3112,7 +3112,7 @@ lemma resolveVAddr_ccorres:
apply (subgoal_tac "scast Kernel_C.ARMLargePage && mask 2 = (scast Kernel_C.ARMLargePage :: machine_word)")
prefer 2
apply (simp add: mask_def ARMLargePage_def)
-- "reduce to resolve_ret_rel goals first"
\<comment> \<open>reduce to resolve_ret_rel goals first\<close>
apply (clarsimp simp: fst_return pte_get_tag_alt true_def false_def split: pte.splits)
apply (safe ; clarsimp simp: cpte_relation_get_tag_simps c_pages_noteq)
(* 4 subgoals *)
@ -3151,7 +3151,7 @@ lemma resolveVAddr_ccorres:
ARMSection_def mask_def)
done
apply clarsimp
apply (rule conjI, fastforce elim: valid_objs_valid_pte') -- "valid_pte'"
apply (rule conjI, fastforce elim: valid_objs_valid_pte') \<comment> \<open>valid_pte'\<close>
apply (frule(1) page_directory_at_rf_sr)
apply (clarsimp simp: isPageTablePDE_def pde_get_tag_alt pde_tag_defs cpde_relation_def
typ_heap_simps pde_pde_coarse_lift_def
@ -3602,7 +3602,7 @@ lemma decodeARMFrameInvocation_ccorres:
apply simp
apply (vcg exspec=getSyscallArg_modifies)
-- "PageMap"
\<comment> \<open>PageMap\<close>
apply (rule ccorres_rhs_assoc)+
apply csymbr+
apply (simp add: if_1_0_0 word_less_nat_alt del: Collect_const)
@ -4269,7 +4269,7 @@ lemma decodeARMPageDirectoryInvocation_ccorres:
apply (clarsimp simp: invs_valid_objs' invs_sch_act_wf'
valid_tcb_state'_def invs_queues)
-- "cache flush constraints"
\<comment> \<open>cache flush constraints\<close>
subgoal for _ _ _ _ _ _ sz p
using pbfs_atleast_pageBits[simplified pageBits_def, of sz]
apply (intro conjI)
@ -4726,7 +4726,7 @@ lemma decodeARMMMUInvocation_ccorres:
apply wp
apply simp
apply (vcg exspec=getSyscallArg_modifies)
apply (rule ccorres_Cond_rhs) -- "ASIDPoolCap case"
apply (rule ccorres_Cond_rhs) \<comment> \<open>ASIDPoolCap case\<close>
apply (simp add: imp_conjR[symmetric] decodeARMMMUInvocation_def
del: Collect_const)
apply (rule ccorres_rhs_assoc)+
@ -5105,10 +5105,10 @@ lemma writeVCPUReg_ccorres:
and Q'=UNIV in ccorres_rewrite_cond_sr)
subgoal by (fastforce dest: rf_sr_ksArchState_armHSCurVCPU simp: cur_vcpu_relation_def
split: option.splits)
apply (clarsimp simp: vcpureg_eq_use_types) -- "C register comparison into reg comparison"
apply (clarsimp simp: vcpureg_eq_use_types) \<comment> \<open>C register comparison into reg comparison\<close>
apply (rule ccorres_Cond_rhs)
-- "vcpuptr is current vcpu"
\<comment> \<open>vcpuptr is current vcpu\<close>
apply clarsimp
apply (rename_tac curvcpuactive)
apply (rule ccorres_Cond_rhs ; clarsimp)
@ -5117,22 +5117,23 @@ lemma writeVCPUReg_ccorres:
and Q'=UNIV in ccorres_rewrite_cond_sr)
subgoal by (clarsimp dest!: rf_sr_ksArchState_armHSCurVCPU simp: cur_vcpu_relation_def
split: option.splits)
(* unification choking on schematics with pairs *)
apply (rule_tac A="vcpu_at' vcpuptr" and A'=UNIV in ccorres_guard_imp)
apply (rule ccorres_Cond_rhs ; clarsimp)
apply (ctac (no_vcg) add: setSCTLR_ccorres)
apply (ctac (no_vcg) add: vcpu_write_reg_ccorres)
apply fastforce
apply fastforce
apply (ctac (no_vcg) add: vcpu_hw_write_reg_ccorres)
-- "no current vcpu"
apply clarsimp
apply wpc
apply (subgoal_tac "\<not> x1")
prefer 2
apply fastforce
apply simp
apply (ctac (no_vcg) add: vcpu_write_reg_ccorres)
(* unification choking on schematics with pairs *)
apply (rule ccorres_guard_imp)
apply (rule ccorres_Cond_rhs; clarsimp)
\<comment> \<open>SCTLR to hardware\<close>
apply (ctac (no_vcg) add: setSCTLR_ccorres)
\<comment> \<open>SCTLR from vcpu\<close>
apply (rule ccorres_pre_getObject_vcpu, rename_tac vcpu)
(* 20 subgoals *)
apply (solves \<open>
match conclusion in "_ (heap_update (Ptr x))" for x \<Rightarrow> \<open>print_term x\<close>,
rule ccorres_guard_imp, rule ccorres_move_c_guard_vcpu,
rule_tac P="ko_at' vcpu vcpuptr" in setObject_ccorres_helper[where P'=UNIV],
rule conseqPre, vcg,
determ \<open>solve_rf_sr_vcpu_update\<close>,
(fastforce simp: ko_at_vcpu_at'D objBits_simps archObjSize_def machine_bits_defs)+\<close>)+
apply fastforce
apply fastforce
done
@ -5175,9 +5176,9 @@ lemma readVCPUReg_ccorres:
and Q'=UNIV in ccorres_rewrite_cond_sr)
subgoal by (fastforce dest: rf_sr_ksArchState_armHSCurVCPU simp: cur_vcpu_relation_def
split: option.splits)
apply (clarsimp simp: vcpureg_eq_use_types) -- "C register comparison into reg comparison"
apply (clarsimp simp: vcpureg_eq_use_types) \<comment> \<open>C register comparison into reg comparison\<close>
apply (rule ccorres_Cond_rhs)
-- "vcpuptr is current vcpu"
\<comment> \<open>vcpuptr is current vcpu\<close>
apply clarsimp
apply (rename_tac curvcpuactive)
apply (rule ccorres_Cond_rhs ; clarsimp)
@ -5186,11 +5187,10 @@ lemma readVCPUReg_ccorres:
and Q'=UNIV in ccorres_rewrite_cond_sr)
subgoal by (clarsimp dest!: rf_sr_ksArchState_armHSCurVCPU simp: cur_vcpu_relation_def
split: option.splits)
(* unification choking on schematics with pairs *)
apply (rule_tac A=\<top> and A'=UNIV in ccorres_guard_imp)
apply (rule ccorres_Cond_rhs ; clarsimp)
-- "SCTLR"
\<comment> \<open>SCTLR\<close>
apply (rule ccorres_add_return2)
apply (ctac (no_vcg) add: getSCTLR_ccorres)
apply (fastforce intro!: ccorres_return_C)
@ -5203,7 +5203,7 @@ lemma readVCPUReg_ccorres:
(* re-unify with ccorres_guard_imp *)
apply fastforce
apply fastforce
-- "any other hyp register"
\<comment> \<open>any other hyp register\<close>
apply (rule ccorres_add_return2)
apply (ctac (no_vcg) add: vcpu_hw_read_reg_ccorres)
apply (fastforce intro!: ccorres_return_C)
@ -5246,7 +5246,7 @@ lemma invokeVCPUReadReg_ccorres: (* styled after invokeTCB_ReadRegisters_ccorres
apply (ctac add: readVCPUReg_ccorres)
apply (rule ccorres_Cond_rhs_Seq[rotated]; clarsimp)
-- "if we are not part of a call"
\<comment> \<open>if we are not part of a call\<close>
apply (simp add: replyOnRestart_def liftE_def bind_assoc)
apply (rule getThreadState_ccorres_foo, rename_tac tstate)
apply (rule_tac P="tstate = Restart" in ccorres_gen_asm)
@ -5259,7 +5259,7 @@ lemma invokeVCPUReadReg_ccorres: (* styled after invokeTCB_ReadRegisters_ccorres
apply (clarsimp simp: return_def dc_simp)
apply (rule hoare_post_taut[of \<top>])
-- "now if we are part of a call"
\<comment> \<open>now if we are part of a call\<close>
apply (rule ccorres_rhs_assoc)+
apply (rename_tac rval)
apply (clarsimp simp: replyOnRestart_def liftE_def bind_assoc)
@ -5384,7 +5384,7 @@ lemma decodeVCPUWriteReg_ccorres:
apply (clarsimp simp: word_le_nat_alt)
apply (simp add: returnOk_bind bindE_assoc
performARMMMUInvocations performARMVCPUInvocation_def)
-- "we want the second alternative - nothing to return to user"
\<comment> \<open>we want the second alternative - nothing to return to user\<close>
apply (subst liftE_invokeVCPUWriteReg_empty_return, clarsimp)
apply (ctac add: setThreadState_ccorres)
apply csymbr
@ -5406,10 +5406,10 @@ lemma decodeVCPUWriteReg_ccorres:
invs_no_0_obj' tcb_at_invs' invs_queues invs_valid_objs' invs_sch_act_wf'
rf_sr_ksCurThread msgRegisters_unfold
valid_tcb_state'_def ThreadState_Restart_def mask_def)
apply (rule conjI; clarsimp) -- "not enough args"
apply (rule conjI; clarsimp) \<comment> \<open>not enough args\<close>
apply (clarsimp simp: isCap_simps cap_get_tag_isCap capVCPUPtr_eq)
apply (subst from_to_enum; clarsimp simp: fromEnum_maxBound_vcpureg_def)
-- "enough args"
\<comment> \<open>enough args\<close>
apply (clarsimp simp: isCap_simps cap_get_tag_isCap capVCPUPtr_eq valid_cap'_def)
apply (subgoal_tac "args \<noteq> []")
prefer 2 subgoal by (cases args; clarsimp, unat_arith?)
@ -5629,7 +5629,7 @@ lemma decodeVCPUInjectIRQ_ccorres:
apply (subst liftE_invokeVCPUInjectIRQ_empty_return)
apply clarsimp
-- "we want the second alternative - nothing to return to user"
\<comment> \<open>we want the second alternative - nothing to return to user\<close>
apply (ctac add: setThreadState_ccorres)
apply (ctac add: invokeVCPUInjectIRQ_ccorres)
apply (rule ccorres_alternative2)
@ -5761,10 +5761,10 @@ lemma decodeVCPUReadReg_ccorres:
rf_sr_ksCurThread msgRegisters_unfold
valid_tcb_state'_def ThreadState_Restart_def mask_def)
apply (rule conjI; clarsimp) -- "no args"
apply (rule conjI; clarsimp) \<comment> \<open>no args\<close>
subgoal by (clarsimp simp: isCap_simps cap_get_tag_isCap capVCPUPtr_eq)
(subst from_to_enum; clarsimp simp: fromEnum_maxBound_vcpureg_def)
-- "at least one arg"
\<comment> \<open>at least one arg\<close>
apply (clarsimp simp: isCap_simps cap_get_tag_isCap capVCPUPtr_eq valid_cap'_def)
apply (subgoal_tac "args \<noteq> []")
prefer 2 apply (cases args; clarsimp, unat_arith?)
@ -5845,7 +5845,7 @@ lemma decodeVCPUSetTCB_ccorres:
apply clarsimp
apply (simp add: returnOk_bind bindE_assoc
performARMMMUInvocations performARMVCPUInvocation_def)
-- "we want the second alternative - nothing to return to user"
\<comment> \<open>we want the second alternative - nothing to return to user\<close>
apply (subst liftE_associateVCPUTCB_empty_return, clarsimp)
apply (ctac add: setThreadState_ccorres)
apply csymbr
@ -5928,7 +5928,7 @@ lemma decodeARMVCPUInvocation_ccorres:
apply (rule ccorres_trim_returnE, simp+)
apply (rule ccorres_call[OF decodeVCPUInjectIRQ_ccorres]; simp)
-- "unknown (arch) invocation labels all throw IllegalOperation in line with the Haskell"
\<comment> \<open>unknown (arch) invocation labels all throw IllegalOperation in line with the Haskell\<close>
apply (rule ccorres_inst[where P=\<top> and P'=UNIV])
apply (rule ccorres_from_vcg_throws[where P=\<top> and P'=UNIV])
apply (intro allI, rule conseqPre, vcg)
@ -5939,7 +5939,7 @@ lemma decodeARMVCPUInvocation_ccorres:
; simp add: invocation_eq_use_types throwError_invocationCatch fst_throwError_returnOk
exception_defs syscall_error_rel_def syscall_error_to_H_cases)
done
-- "preconditions imply calculated preconditions"
\<comment> \<open>preconditions imply calculated preconditions\<close>
apply auto
done

View File

@ -207,7 +207,7 @@ lemma (* handleDoubleFault_ccorres: *)
apply ac_init
prefer 3 apply simp
apply (unfold handleDoubleFault_def handleDoubleFault'_def K_bind_def)
apply (rule corres_add_noop_rhs2) -- "split out extra haskell code"
apply (rule corres_add_noop_rhs2) \<comment> \<open>split out extra haskell code\<close>
apply (rule corres_split[OF _ setThreadState_ccorres[ac]])
apply (rule corres_symb_exec_l_no_exs)
apply simp

View File

@ -303,7 +303,7 @@ lemmas ccorres_move_array_assertion_cnode_ctes [corres_pre]
ccorres_move_Guard [OF array_assertion_abs_cnode_ctes]
lemma locateSlotCNode_ccorres [corres]:
assumes gl: "\<And>v s. globals (xfu v s) = globals s" -- "for state rel. preservation"
assumes gl: "\<And>v s. globals (xfu v s) = globals s" \<comment> \<open>for state rel. preservation\<close>
and fg: "\<And>v s. xf (xfu (\<lambda>_. v) s) = v"
shows "ccorres (\<lambda>v v'. v' = Ptr v) xf \<top> {_. cnode = cnode' \<and> offset = offset'} hs
(locateSlotCNode cnode bits offset)
@ -324,7 +324,7 @@ lemma locateSlotCNode_ccorres [corres]:
done
lemma locateSlotTCB_ccorres [corres]:
assumes gl: "\<And>v s. globals (xfu v s) = globals s" -- "for state rel. preservation"
assumes gl: "\<And>v s. globals (xfu v s) = globals s" \<comment> \<open>for state rel. preservation\<close>
and fg: "\<And>v s. xf (xfu (\<lambda>_. v) s) = v"
shows "ccorres (\<lambda>v v'. v' = Ptr v) xf \<top> {_. cnode = cnode' \<and> offset = offset'} hs
(locateSlotTCB cnode offset)
@ -340,13 +340,13 @@ lemma locateSlotTCB_ccorres [corres]:
lemma getSlotCap_h_val_ccorres [corres]:
fixes p :: "cstate \<Rightarrow> cte_C ptr"
assumes gl: "\<And>v s. globals (xfu v s) = globals s" -- "for state rel. preservation"
assumes gl: "\<And>v s. globals (xfu v s) = globals s" \<comment> \<open>for state rel. preservation\<close>
and fg: "\<And>v s. xf (xfu (\<lambda>_. v) s) = v"
shows "ccorres ccap_relation xf \<top> {s. p s = Ptr a} hs
(getSlotCap a) (Basic (\<lambda>s. xfu (\<lambda>_. h_val (hrs_mem (t_hrs_' (globals s))) (Ptr &(p s\<rightarrow>[''cap_C'']) :: cap_C ptr)) s))"
unfolding getSlotCap_def
apply (rule ccorres_add_UNIV_Int)
apply (cinitlift p) -- "EVIL!"
apply (cinitlift p) \<comment> \<open>EVIL!\<close>
apply simp
apply (rule ccorres_guard_imp2)
apply (rule ccorres_pre_getCTE)

View File

@ -40,9 +40,9 @@ lemma lookupCap_ccorres':
apply (cinit lift: cPtr_' thread_' simp: lookupCapAndSlot_def liftME_def bindE_assoc)
apply (ctac (no_vcg) add: lookupSlotForThread_ccorres')
--"case where lu_ret.status is EXCEPTION_NONE"
\<comment> \<open>case where lu_ret.status is EXCEPTION_NONE\<close>
apply (simp add: split_beta cong:call_ignore_cong)
apply csymbr --"call status_C_update"
apply csymbr \<comment> \<open>call status_C_update\<close>
apply (simp add: Collect_const[symmetric] lookupSlot_raw_rel_def liftE_def
del: Collect_const)
apply (rule ccorres_move_c_guard_cte)
@ -50,18 +50,18 @@ lemma lookupCap_ccorres':
apply (rule ccorres_return_CE [unfolded returnOk_def, simplified], simp+)[1]
apply wp
apply vcg
--"case where lu_ret.status is *not* EXCEPTION_NONE"
\<comment> \<open>case where lu_ret.status is *not* EXCEPTION_NONE\<close>
apply simp
apply (rule ccorres_split_throws)
apply (rule ccorres_rhs_assoc)+
apply csymbr
apply csymbr -- "call cap_null_cap_new_'proc"
apply csymbr \<comment> \<open>call cap_null_cap_new_'proc\<close>
apply csymbr
apply (rule ccorres_return_C_errorE, simp+)[1]
apply vcg
apply (wp | simp)+
-- "last subgoal"
\<comment> \<open>last subgoal\<close>
apply (clarsimp simp: valid_pspace_valid_objs')
apply (intro impI conjI allI)
apply (simp add: lookupSlot_raw_rel_def)
@ -91,10 +91,10 @@ lemma lookupCapAndSlot_ccorres :
apply (cinit lift: thread_' cPtr_')
apply (ctac (no_vcg))
--"case where lu_ret.status is EXCEPTION_NONE"
\<comment> \<open>case where lu_ret.status is EXCEPTION_NONE\<close>
apply (simp add: split_beta cong:call_ignore_cong)
apply csymbr --"call status_C_update"
apply csymbr --"call slot_C_update"
apply csymbr \<comment> \<open>call status_C_update\<close>
apply csymbr \<comment> \<open>call slot_C_update\<close>
apply (simp add: Collect_const[symmetric] lookupSlot_raw_rel_def liftE_bindE
del: Collect_const)
apply (rule ccorres_move_c_guard_cte)
@ -103,7 +103,7 @@ lemma lookupCapAndSlot_ccorres :
apply (clarsimp simp: Bex_def in_monad getSlotCap_def in_getCTE2 cte_wp_at_ctes_of)
apply (erule(1) cmap_relationE1[OF cmap_relation_cte])
apply (simp add: ccte_relation_ccap_relation typ_heap_simps')
--"case where lu_ret.status is *not* EXCEPTION_NONE"
\<comment> \<open>case where lu_ret.status is *not* EXCEPTION_NONE\<close>
apply simp
apply (rule ccorres_split_throws)
apply (rule ccorres_rhs_assoc)+
@ -113,7 +113,7 @@ lemma lookupCapAndSlot_ccorres :
apply vcg
apply (wp | simp)+
-- "last subgoal"
\<comment> \<open>last subgoal\<close>
apply clarsimp
apply (rule conjI, fastforce)
apply clarsimp
@ -160,14 +160,14 @@ lemma lookupSlotForCNodeOp_ccorres':
(lookupSlotForCNodeOp isSource croot capptr depth)
(Call lookupSlotForCNodeOp_'proc)"
apply (cinit lift: capptr_' isSource_' root_' depth_')
apply csymbr -- "slot_C_update"
apply csymbr -- "cap_get_capType"
apply csymbr \<comment> \<open>slot_C_update\<close>
apply csymbr \<comment> \<open>cap_get_capType\<close>
apply (rule_tac Q=\<top> and Q'=\<top> in ccorres_if_cond_throws2)
-- "correspondance of Haskell and C conditions"
\<comment> \<open>correspondance of Haskell and C conditions\<close>
apply (clarsimp simp: Collect_const_mem cap_get_tag_isCap)
-- "case where root is *not* a CNode => throw InvalidRoot"
\<comment> \<open>case where root is *not* a CNode => throw InvalidRoot\<close>
apply simp
apply (rule_tac P=\<top> and P' =UNIV in ccorres_from_vcg_throws)
apply (rule allI, rule conseqPre, vcg)
@ -177,7 +177,7 @@ lemma lookupSlotForCNodeOp_ccorres':
apply clarsimp
apply (subst syscall_error_to_H_cases(6), simp+)[1]
-- " case where root is a CNode"
\<comment> \<open>case where root is a CNode\<close>
apply (simp add: rangeCheck_def)
apply csymbr
apply (rule ccorres_Cond_rhs_Seq)
@ -194,13 +194,13 @@ lemma lookupSlotForCNodeOp_ccorres':
apply vcg
apply csymbr
apply (rule_tac Q="\<lambda>s. depth < 2 ^ word_bits" and Q'=\<top> in ccorres_split_unless_throwError_cond)
-- "correspondance of Haskell and C conditions"
\<comment> \<open>correspondance of Haskell and C conditions\<close>
apply (clarsimp simp: Collect_const_mem fromIntegral_def integral_inv
if_1_0_0)
apply (simp add: word_size unat_of_nat32 word_less_nat_alt
word_less_1[symmetric] linorder_not_le)
-- "case of RangeError"
\<comment> \<open>case of RangeError\<close>
apply (rule_tac P= \<top> and P' =UNIV in ccorres_from_vcg_throws)
apply (rule allI, rule conseqPre, vcg)
apply (clarsimp simp: throwError_def return_def syscall_error_rel_def)
@ -208,20 +208,20 @@ lemma lookupSlotForCNodeOp_ccorres':
apply (subst syscall_error_to_H_cases(4), simp+)[1]
apply (simp add: word_size word_sle_def)
-- "case where there is *no* RangeError"
\<comment> \<open>case where there is *no* RangeError\<close>
apply (rule_tac xf'=lookupSlot_xf in ccorres_rel_imp)
apply (rule_tac r="\<lambda>w w'. w'= Ptr w"
and f="\<lambda> st fl es. fl = scast EXCEPTION_SYSCALL_ERROR \<and>
syscall_error_to_H (errsyscall es) (errlookup_fault es) = Some (FailedLookup isSource st)"
in lookupErrorOnFailure_ccorres)
apply (ctac (no_vcg)) -- "resolveAddressBits"
-- " case where resolveAddressBits results in EXCEPTION_NONE"
apply (ctac (no_vcg)) \<comment> \<open>resolveAddressBits\<close>
\<comment> \<open>case where resolveAddressBits results in EXCEPTION_NONE\<close>
apply clarsimp
apply (rule_tac A=\<top> and A'=UNIV in ccorres_guard_imp2)
apply (rule_tac Q=\<top> and Q'=\<top> in ccorres_if_cond_throws2)
-- "correspondance of Haskell and C conditions"
\<comment> \<open>correspondance of Haskell and C conditions\<close>
apply (clarsimp simp: Collect_const_mem unat_gt_0)
-- " case where bits are remaining"
\<comment> \<open>case where bits are remaining\<close>
apply (rule_tac P= \<top> and P' =UNIV in ccorres_from_vcg_throws)
apply (rule allI, rule conseqPre, vcg)
apply (clarsimp simp: throwError_def return_def)
@ -230,17 +230,17 @@ lemma lookupSlotForCNodeOp_ccorres':
apply (simp add: lookup_fault_depth_mismatch_lift)
apply (erule le_32_mask_eq)
-- " case where *no* bits are remaining"
apply csymbr -- "slot_C_update"
apply csymbr -- "status_C_update"
\<comment> \<open>case where *no* bits are remaining\<close>
apply csymbr \<comment> \<open>slot_C_update\<close>
apply csymbr \<comment> \<open>status_C_update\<close>
apply (rule ccorres_return_CE, simp+)[1]
apply vcg
-- "guard_imp subgoal"
\<comment> \<open>guard_imp subgoal\<close>
apply clarsimp
-- " case where resolveAddressBits does *not* results in EXCEPTION_NONE"
\<comment> \<open>case where resolveAddressBits does *not* results in EXCEPTION_NONE\<close>
apply clarsimp
apply (rule_tac P= \<top> and P' ="\<lbrace>\<exists>v. (lookup_fault_lift (\<acute>current_lookup_fault)) = Some v
\<and> lookup_fault_to_H v = err \<rbrace>"
@ -251,7 +251,7 @@ lemma lookupSlotForCNodeOp_ccorres':
apply (subst syscall_error_to_H_cases(6), simp+)[1]
apply wp
-- "rel_imp"
\<comment> \<open>rel_imp\<close>
apply clarsimp
apply (case_tac x, clarsimp)
apply (simp add: syscall_error_rel_def errstate_def)
@ -261,7 +261,7 @@ lemma lookupSlotForCNodeOp_ccorres':
apply vcg
apply vcg
-- "last subgoal"
\<comment> \<open>last subgoal\<close>
apply (clarsimp simp: if_1_0_0 to_bool_def true_def word_size
fromIntegral_def integral_inv)
apply (case_tac "cap_get_tag root = scast cap_cnode_cap")

File diff suppressed because it is too large Load Diff

View File

@ -214,7 +214,7 @@ proof (cases "isCNodeCap cap'")
show ?thesis using False
apply (cinit' lift: nodeCap_' capptr_' n_bits_')
apply csymbr+
-- "Exception stuff"
\<comment> \<open>Exception stuff\<close>
apply (rule ccorres_split_throws)
apply (simp add: Collect_const cap_get_tag_isCap isCap_simps ccorres_cond_iffs
resolveAddressBits.simps scast_id)
@ -236,8 +236,8 @@ next
from True show ?thesis
apply -
apply (cinit' simp add: whileAnno_def ucast_id)
-- "This is done here as init lift usually throws away the relationship between nodeCap_' s and nodeCap. Normally
this OK, but the induction messes with everything :("
\<comment> \<open>This is done here as init lift usually throws away the relationship between nodeCap_' s and nodeCap. Normally
this OK, but the induction messes with everything :(\<close>
apply (rule ccorres_abstract [where xf' = nodeCap_'])
apply ceqv
apply (rename_tac "nodeCap")
@ -266,11 +266,11 @@ next
apply clarsimp
apply (rule_tac I = "{s. cap_get_tag (nodeCap_' s) = scast cap_cnode_cap}"
in HoarePartial.While [unfolded whileAnno_def, OF subset_refl])
apply (vcg strip_guards=true) -- "takes a while"
apply (vcg strip_guards=true) \<comment> \<open>takes a while\<close>
apply clarsimp
apply simp
apply (clarsimp simp: cap_get_tag_isCap to_bool_def)
-- "Main thm"
\<comment> \<open>Main thm\<close>
proof (induct cap' cptr' guard' rule: resolveAddressBits.induct [case_names ind])
case (ind cap cptr guard)
@ -476,12 +476,12 @@ next
apply (csymbr | rule iffD2 [OF ccorres_seq_skip])+
apply (rule ccorres_Guard_Seq)+
apply csymbr
-- "handle the stateAssert in locateSlotCap very carefully"
\<comment> \<open>handle the stateAssert in locateSlotCap very carefully\<close>
apply (simp(no_asm) only: liftE_bindE[where a="locateSlotCap a b" for a b])
apply (rule ccorres_locateSlotCap_push[rotated])
apply (simp add: unlessE_def)
apply (rule hoare_pre, wp, simp)
-- "Convert guardBits, radixBits and capGuard to their Haskell versions"
\<comment> \<open>Convert guardBits, radixBits and capGuard to their Haskell versions\<close>
apply (drule (2) cgD, drule (2) rbD, drule (2) gbD)
apply (elim conjE)
apply (rule ccorres_gen_asm [where P = "guard \<le> 32"])
@ -550,9 +550,9 @@ next
apply (vcg strip_guards=true)
apply (vcg strip_guards=true)
apply (rule conjI)
-- "Haskell guard"
\<comment> \<open>Haskell guard\<close>
apply (thin_tac "unat n_bits = guard")
apply (clarsimp simp del: imp_disjL) -- "take a while"
apply (clarsimp simp del: imp_disjL) \<comment> \<open>take a while\<close>
apply (intro impI conjI allI)
apply fastforce
apply clarsimp
@ -562,7 +562,7 @@ next
apply (clarsimp simp: isCap_simps valid_cap_simps' cte_level_bits_def objBits_defs
real_cte_at')
apply (clarsimp simp: isCap_simps valid_cap'_def)
-- "C guard"
\<comment> \<open>C guard\<close>
apply (frule (1) cgD [OF refl], frule (1) rbD [OF refl], frule (1) gbD [OF refl])
apply (simp add: Collect_const_mem cap_get_tag_isCap exception_defs lookup_fault_lifts
n_bits_guard mask6_eqs word_le_nat_alt word_less_nat_alt gm)

View File

@ -1764,7 +1764,7 @@ proof -
, assumption +,
simp_all add: objBits_simps' archObjSize_def projectKOs
machine_bits_defs
heap_to_user_data_restrict heap_to_device_data_restrict)[1])+ -- "waiting ..."
heap_to_user_data_restrict heap_to_device_data_restrict)[1])+ \<comment> \<open>waiting ...\<close>
apply (simp add: map_to_ctes_delete' cmap_relation_restrict_both_proj
cmap_relation_restrict_both cmap_array_helper hrs_htd_update
table_bits_defs cmap_array)
@ -1811,7 +1811,7 @@ proof -
apply fastforce
apply ((subst lift_t_typ_region_bytes, rule cm_disj_tcb, assumption+,
simp_all add: objBits_simps archObjSize_def pageBits_def projectKOs)[1])+
-- "waiting ..."
\<comment> \<open>waiting ...\<close>
apply (simp add: tcb_queue_relation_live_restrict
[OF D.valid_untyped tat tlive rl])
done

View File

@ -661,7 +661,7 @@ lemma doUnbindNotification_ccorres:
cpspace_relation_def update_ntfn_map_tos)
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (rule cpspace_relation_ntfn_update_ntfn, assumption+)
apply (clarsimp simp: cnotification_relation_def Let_def
mask_def [where n=2] NtfnState_Waiting_def)
@ -709,7 +709,7 @@ lemma doUnbindNotification_ccorres':
cpspace_relation_def update_ntfn_map_tos)
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (rule cpspace_relation_ntfn_update_ntfn, assumption+)
apply (clarsimp simp: cnotification_relation_def Let_def
mask_def [where n=2] NtfnState_Waiting_def)
@ -906,7 +906,7 @@ lemma finaliseCap_True_cases_ccorres:
apply (clarsimp simp add: return_def ccap_relation_NullCap_iff)
apply (clarsimp simp add: irq_opt_relation_def)
apply vcg
-- "NullCap case by exhaustion"
\<comment> \<open>NullCap case by exhaustion\<close>
apply (simp add: cap_get_tag_isCap Let_def
ccorres_cond_empty_iff ccorres_cond_univ_iff)
apply (rule ccorres_rhs_assoc2, rule ccorres_rhs_assoc2,

View File

@ -441,13 +441,13 @@ lemma invokeCNodeSaveCaller_ccorres:
Collect_False Collect_True
del: Collect_const)[1]
apply (rule ccorres_fail)+
-- "NullCap case"
\<comment> \<open>NullCap case\<close>
apply (rule ccorres_from_vcg_throws [where P=\<top> and P'=UNIV])
apply (rule allI, rule conseqPre, vcg)
apply clarsimp
apply (simp add: return_def)
apply (rule ccorres_fail)+
-- "ReplyCap case"
\<comment> \<open>ReplyCap case\<close>
apply (rule ccorres_rhs_assoc)
apply csymbr
apply (frule cap_get_tag_isCap_unfolded_H_cap)
@ -736,7 +736,7 @@ lemma decodeCNodeInvocation_ccorres:
injection_handler_returnOk Collect_const[symmetric]
cong: call_ignore_cong del: Collect_const)
apply (rule ccorres_Cond_rhs_Seq)
-- "CNodeCopy case"
\<comment> \<open>CNodeCopy case\<close>
apply (simp add: Collect_const[symmetric] del: Collect_const)
apply (rule ccorres_rhs_assoc)+
apply (rule ccorres_Cond_rhs_Seq)
@ -806,7 +806,7 @@ lemma decodeCNodeInvocation_ccorres:
apply (simp add: Collect_const_mem)
apply (vcg exspec=getSyscallArg_modifies)
apply (rule ccorres_Cond_rhs_Seq)
-- "CNodeMint case"
\<comment> \<open>CNodeMint case\<close>
apply (simp add: Collect_const[symmetric]
del: Collect_const)
apply (rule ccorres_rhs_assoc)+
@ -893,7 +893,7 @@ lemma decodeCNodeInvocation_ccorres:
apply (simp add: Collect_const_mem)
apply (vcg exspec=getSyscallArg_modifies)
apply (rule ccorres_Cond_rhs_Seq)
-- "CNodeMove case"
\<comment> \<open>CNodeMove case\<close>
apply (simp add: Collect_const[symmetric] split_def
injection_handler_returnOk whenE_def
ccorres_invocationCatch_Inr
@ -921,7 +921,7 @@ lemma decodeCNodeInvocation_ccorres:
apply vcg
apply (rule conseqPre, vcg, clarsimp)
apply (rule ccorres_Cond_rhs_Seq)
-- "CNodeMutate case"
\<comment> \<open>CNodeMutate case\<close>
apply (rule ccorres_rhs_assoc)+
apply (simp add: Collect_const[symmetric] del: Collect_const
cong: call_ignore_cong)

View File

@ -232,7 +232,7 @@ lemma cancelSignal_ccorres_helper:
apply (clarsimp simp: typ_heap_simps cong: imp_cong split del: if_split simp del: comp_def)
apply (frule null_ep_queue [simplified Fun.comp_def])
apply (intro impI conjI allI)
-- "empty case"
\<comment> \<open>empty case\<close>
apply clarsimp
apply (frule iffD1 [OF tcb_queue_head_empty_iff [OF tcb_queue_relation'_queue_rel]])
apply (rule ballI, erule bspec)
@ -246,20 +246,20 @@ lemma cancelSignal_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cendpoint_relation_ntfn_queue [OF invs_sym'], assumption+)
apply simp
apply (erule (1) map_to_ko_atI')
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (rule cpspace_relation_ntfn_update_ntfn, assumption+)
apply (simp add: cnotification_relation_def Let_def NtfnState_Idle_def)
apply (simp add: carch_state_relation_def carch_globals_def)
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (simp add: carch_state_relation_def carch_globals_def
@ -271,7 +271,7 @@ lemma cancelSignal_ccorres_helper:
apply (simp add: objBits_simps)
apply assumption
-- "non empty case"
\<comment> \<open>non empty case\<close>
apply (frule tcb_queue_head_empty_iff [OF tcb_queue_relation'_queue_rel])
apply (rule ballI, erule bspec)
apply (erule subsetD [rotated])
@ -284,10 +284,10 @@ lemma cancelSignal_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cendpoint_relation_ntfn_queue)
@ -295,7 +295,7 @@ lemma cancelSignal_ccorres_helper:
apply assumption+
apply simp
apply (erule (1) map_to_ko_atI')
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (rule cpspace_relation_ntfn_update_ntfn, assumption+)
apply (simp add: cnotification_relation_def Let_def isWaitingNtfn_def
split: ntfn.splits split del: if_split)
@ -305,7 +305,7 @@ lemma cancelSignal_ccorres_helper:
apply (clarsimp simp add: Ptr_ptr_val h_t_valid_clift_Some_iff)
apply (simp add: tcb_queue_relation'_prev_mask)
apply simp
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (simp add: carch_state_relation_def carch_globals_def
@ -743,7 +743,7 @@ lemma state_relation_queue_update_helper':
apply (clarsimp simp: inQ_def obj_at'_def)
apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def)
apply (intro conjI)
-- "cpspace_relation"
\<comment> \<open>cpspace_relation\<close>
apply (erule nonemptyE, drule(1) bspec)
apply (clarsimp simp: cpspace_relation_def)
apply (drule obj_at_ko_at', clarsimp)
@ -752,18 +752,18 @@ lemma state_relation_queue_update_helper':
apply (frule null_sched_queue)
apply (frule null_sched_epD)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (drule ctcb_relation_null_queue_ptrs,
simp_all)[1]
-- "endpoint relation"
\<comment> \<open>endpoint relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (erule cendpoint_relation_upd_tcb_no_queues, simp+)
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (erule cnotification_relation_upd_tcb_no_queues, simp+)
-- "ready queues"
\<comment> \<open>ready queues\<close>
apply (simp add: cready_queues_relation_def Let_def
cready_queues_index_to_C_in_range
seL4_MinPrio_def minDom_def)
@ -3063,7 +3063,7 @@ lemma cancelIPC_ccorres_helper:
apply (clarsimp simp: typ_heap_simps cong: imp_cong split del: if_split simp del: comp_def)
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (intro impI conjI allI)
-- "empty case"
\<comment> \<open>empty case\<close>
apply clarsimp
apply (frule iffD1 [OF tcb_queue_head_empty_iff [OF tcb_queue_relation'_queue_rel]])
apply (rule ballI, erule bspec)
@ -3076,21 +3076,21 @@ lemma cancelIPC_ccorres_helper:
cpspace_relation_def update_ep_map_tos typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
subgoal by (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (rule cpspace_relation_ep_update_ep, assumption+)
subgoal by (simp add: cendpoint_relation_def Let_def EPState_Idle_def)
subgoal by simp
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cnotification_relation_ep_queue [OF invs_sym'], assumption+)
subgoal by simp
apply (erule (1) map_to_ko_atI')
apply (simp add: heap_to_user_data_def Let_def)
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
subgoal by (clarsimp simp: comp_def)
subgoal by (simp add: carch_state_relation_def carch_globals_def
@ -3100,7 +3100,7 @@ lemma cancelIPC_ccorres_helper:
subgoal by (simp add: objBits_simps')
subgoal by (simp add: objBits_simps)
apply assumption
-- "non empty case"
\<comment> \<open>non empty case\<close>
apply clarsimp
apply (frule tcb_queue_head_empty_iff [OF tcb_queue_relation'_queue_rel])
apply (rule ballI, erule bspec)
@ -3113,28 +3113,28 @@ lemma cancelIPC_ccorres_helper:
cpspace_relation_def update_ep_map_tos typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
subgoal by (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (rule cpspace_relation_ep_update_ep, assumption+)
apply (simp add: cendpoint_relation_def Let_def isSendEP_def isRecvEP_def split: endpoint.splits split del: if_split)
-- "recv case"
\<comment> \<open>recv case\<close>
apply (clarsimp simp add: Ptr_ptr_val h_t_valid_clift_Some_iff
tcb_queue_relation'_next_mask tcb_queue_relation'_prev_mask cong: tcb_queue_relation'_cong)
subgoal by (intro impI conjI; simp)
-- "send case"
\<comment> \<open>send case\<close>
apply (clarsimp simp add: Ptr_ptr_val h_t_valid_clift_Some_iff
tcb_queue_relation'_next_mask tcb_queue_relation'_prev_mask cong: tcb_queue_relation'_cong)
subgoal by (intro impI conjI; simp)
subgoal by simp
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cnotification_relation_ep_queue [OF invs_sym'], assumption+)
apply simp
apply (erule (1) map_to_ko_atI')
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
subgoal by (clarsimp simp: comp_def)
subgoal by (simp add: carch_state_relation_def carch_globals_def
@ -3293,7 +3293,7 @@ lemma cancelIPC_ccorres1:
apply (rule_tac xf'=ret__unsigned_' in ccorres_abstract, ceqv)
apply (rule_tac P="rv' = thread_state_to_tsType rv" in ccorres_gen_asm2)
apply wpc
-- "BlockedOnReceive"
\<comment> \<open>BlockedOnReceive\<close>
apply (simp add: word_sle_def "StrictC'_thread_state_defs" ccorres_cond_iffs cong: call_ignore_cong)
apply (fold dc_def)
apply (rule ccorres_rhs_assoc)+
@ -3301,7 +3301,7 @@ lemma cancelIPC_ccorres1:
apply csymbr
apply (rule ccorres_pre_getEndpoint)
apply (rule ccorres_assert)
apply (rule ccorres_symb_exec_r) -- "ptr_get lemmas don't work so well :("
apply (rule ccorres_symb_exec_r) \<comment> \<open>ptr_get lemmas don't work so well :(\<close>
apply (rule ccorres_symb_exec_r)
apply (simp only: fun_app_def simp_list_case_return
return_bind ccorres_seq_skip)
@ -3320,7 +3320,7 @@ lemma cancelIPC_ccorres1:
apply (rule subset_refl)
apply (rule conseqPre, vcg)
apply clarsimp
-- "BlockedOnReply case"
\<comment> \<open>BlockedOnReply case\<close>
apply (simp add: "StrictC'_thread_state_defs" ccorres_cond_iffs
Collect_False Collect_True word_sle_def
cong: call_ignore_cong del: Collect_const)
@ -3383,7 +3383,7 @@ lemma cancelIPC_ccorres1:
apply (wp threadSet_invs_trivial | simp)+
apply (clarsimp simp add: guard_is_UNIV_def tcbReplySlot_def
Kernel_C.tcbReply_def tcbCNodeEntries_def)
-- "BlockedOnNotification"
\<comment> \<open>BlockedOnNotification\<close>
apply (simp add: word_sle_def "StrictC'_thread_state_defs" ccorres_cond_iffs dc_def [symmetric] cong: call_ignore_cong)
apply (rule ccorres_symb_exec_r)
apply (ctac (no_vcg))
@ -3392,18 +3392,18 @@ lemma cancelIPC_ccorres1:
apply (rule subset_refl)
apply (rule conseqPre, vcg)
apply clarsimp
-- "Running, Inactive, and Idle"
\<comment> \<open>Running, Inactive, and Idle\<close>
apply (simp add: word_sle_def "StrictC'_thread_state_defs" ccorres_cond_iffs dc_def [symmetric] cong: call_ignore_cong,
rule ccorres_return_Skip)+
-- "BlockedOnSend"
\<comment> \<open>BlockedOnSend\<close>
apply (simp add: word_sle_def "StrictC'_thread_state_defs" ccorres_cond_iffs dc_def [symmetric] cong: call_ignore_cong)
-- "clag"
\<comment> \<open>clag\<close>
apply (rule ccorres_rhs_assoc)+
apply csymbr
apply csymbr
apply (rule ccorres_pre_getEndpoint)
apply (rule ccorres_assert)
apply (rule ccorres_symb_exec_r) -- "ptr_get lemmas don't work so well :("
apply (rule ccorres_symb_exec_r) \<comment> \<open>ptr_get lemmas don't work so well :(\<close>
apply (rule ccorres_symb_exec_r)
apply (simp only: fun_app_def simp_list_case_return return_bind ccorres_seq_skip)
apply (rule ccorres_rhs_assoc2)
@ -3421,10 +3421,10 @@ lemma cancelIPC_ccorres1:
apply (rule conseqPre, vcg, rule subset_refl)
apply (rule conseqPre, vcg)
apply clarsimp
-- "Restart"
\<comment> \<open>Restart\<close>
apply (simp add: word_sle_def "StrictC'_thread_state_defs" ccorres_cond_iffs dc_def [symmetric] cong: call_ignore_cong,
rule ccorres_return_Skip)
-- "Post wp proofs"
\<comment> \<open>Post wp proofs\<close>
apply vcg
apply clarsimp
apply (rule conseqPre, vcg)

View File

@ -2689,7 +2689,7 @@ lemma getReceiveSlots_ccorres:
apply clarsimp
apply (simp add: cct_relation_def)
apply (case_tac rv, clarsimp)
apply (rule UNIV_I) -- "still a schematic here ..."
apply (rule UNIV_I) \<comment> \<open>still a schematic here ...\<close>
done
@ -2784,7 +2784,7 @@ lemma ccorres_if_cond_throws_break:
assumes abs: "\<forall>s s'. (s, s') \<in> sr \<and> Q s \<and> Q' s' \<longrightarrow> P = (s' \<in> P')"
and ac: "P \<Longrightarrow> ccorres_underlying sr \<Gamma> r xf arrel axf R R' (catchbrk_C # hs) a c"
and bd: "\<not> P \<Longrightarrow> ccorres_underlying sr \<Gamma> r xf arrel axf U U' (catchbrk_C # hs) b d"
and cthrows: "\<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> PT' c {}, UNIV" -- "c always throws"
and cthrows: "\<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> PT' c {}, UNIV" \<comment> \<open>c always throws\<close>
shows "ccorres_underlying sr \<Gamma> r xf arrel axf
(Q and (\<lambda>s. P \<longrightarrow> R s) and (\<lambda>s. \<not> P \<longrightarrow> U s))
(Collect Q' \<inter> {s. (s \<in> P' \<longrightarrow> s \<in> R' \<inter> PT') \<and> (s \<notin> P' \<longrightarrow> s \<in> U')})
@ -2832,7 +2832,7 @@ lemma ccorres_if_cond_throws_break2:
assumes abs: "\<forall>s s'. (s, s') \<in> sr \<and> Q s \<and> Q' s' \<longrightarrow> (\<not> P) = (s' \<in> P')"
and ac: "\<not> P \<Longrightarrow> ccorres_underlying sr \<Gamma> r xf arrel axf R R' (catchbrk_C # hs) a c"
and bd: "P \<Longrightarrow> ccorres_underlying sr \<Gamma> r xf arrel axf U U' (catchbrk_C # hs) b d"
and cthrows: "\<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> PT' c {}, UNIV" -- "c always throws"
and cthrows: "\<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> PT' c {}, UNIV" \<comment> \<open>c always throws\<close>
shows "ccorres_underlying sr \<Gamma> r xf arrel axf
(Q and (\<lambda>s. \<not> P \<longrightarrow> R s) and (\<lambda>s. P \<longrightarrow> U s))
(Collect Q' \<inter> {s. (s \<in> P' \<longrightarrow> s \<in> R' \<inter> PT') \<and> (s \<notin> P' \<longrightarrow> s \<in> U')})
@ -2859,7 +2859,7 @@ lemma ccorres_split_when_throwError_cond_break:
R R' (catchbrk_C # hs) (throwError e) c"
and bd: "\<not> P \<Longrightarrow> ccorres_underlying sr \<Gamma> r xf ar axf
U U' (catchbrk_C # hs) b d"
and cthrows: "\<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> PT' c {}, UNIV" -- "c always throws"
and cthrows: "\<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> PT' c {}, UNIV" \<comment> \<open>c always throws\<close>
shows "ccorres_underlying sr \<Gamma> r xf ar axf
(Q and (\<lambda>s. P \<longrightarrow> R s) and (\<lambda>s. \<not> P \<longrightarrow> U s))
(Collect Q' \<inter> {s. (s \<in> P' \<longrightarrow> s \<in> R' \<inter> PT') \<and> (s \<notin> P' \<longrightarrow> s \<in> U')})
@ -3137,7 +3137,7 @@ next
apply (clarsimp simp: Collect_const_mem)
apply (rule sym, rule from_bool_neq_0)
-- "case where a badge is sent"
\<comment> \<open>case where a badge is sent\<close>
apply (rule ccorres_rhs_assoc)+
apply csymbr
apply (simp only: Let_def liftE_bindE withoutFailure_def fun_app_def)
@ -3158,7 +3158,7 @@ next
apply (simp split del: if_split)
apply (vcg exspec=setExtraBadge_modifies)
-- "case where a cap is sent (or rather a send is attempted)"
\<comment> \<open>case where a cap is sent (or rather a send is attempted)\<close>
apply (simp add: split_def del: Collect_const
cong: call_ignore_cong split del: if_split)
apply (rule ccorres_rhs_assoc)+
@ -3175,7 +3175,7 @@ next
apply (rule ccorres_cond_false_seq)
apply (simp)
-- "case not diminish"
\<comment> \<open>case not diminish\<close>
apply (rule ccorres_split_nothrowE)
apply (rule unifyFailure_ccorres)
apply (ctac add: deriveCap_ccorres')
@ -3243,7 +3243,7 @@ next
apply (wp deriveCap_derived is_the_ep_deriveCap)
apply (vcg exspec=deriveCap_modifies)
--"remaining non ccorres subgoals"
\<comment> \<open>remaining non ccorres subgoals\<close>
apply (clarsimp simp: Collect_const_mem if_1_0_0
split del: if_split)
apply (rule_tac Q'="\<lbrace>\<acute>ret__int = from_bool (cap_get_tag cap' = scast cap_endpoint_cap
@ -4980,7 +4980,7 @@ lemma sendIPC_dequeue_ccorres_helper:
apply (intro conjI impI allI)
apply (fastforce simp: h_t_valid_clift)
apply (fastforce simp: h_t_valid_clift)
-- "empty case"
\<comment> \<open>empty case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (frule iffD1 [OF tcb_queue_head_empty_iff
@ -4997,21 +4997,21 @@ lemma sendIPC_dequeue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (rule cpspace_relation_ep_update_ep, assumption+)
apply (simp add: cendpoint_relation_def Let_def EPState_Idle_def
tcb_queue_relation'_def)
apply simp
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cnotification_relation_ep_queue [OF invs_sym'], assumption+)
apply simp
apply (erule (1) map_to_ko_atI')
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (simp add: carch_state_relation_def typ_heap_simps')
@ -5021,7 +5021,7 @@ lemma sendIPC_dequeue_ccorres_helper:
apply (simp add: objBits_simps)
apply assumption
apply (clarsimp simp: cendpoint_relation_def Let_def tcb_queue_relation'_def)
-- "non-empty case"
\<comment> \<open>non-empty case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (frule tcb_queue_head_empty_iff [OF tcb_queue_relation'_queue_rel])
@ -5037,10 +5037,10 @@ lemma sendIPC_dequeue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (rule cpspace_relation_ep_update_ep, assumption+)
apply (clarsimp simp: cendpoint_relation_def Let_def
isRecvEP_def isSendEP_def
@ -5053,13 +5053,13 @@ lemma sendIPC_dequeue_ccorres_helper:
split del: if_split)
apply (clarsimp split: if_split)
apply simp
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cnotification_relation_ep_queue [OF invs_sym'], assumption+)
apply simp
apply (erule (1) map_to_ko_atI')
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (simp add: carch_state_relation_def typ_heap_simps')
@ -5382,7 +5382,7 @@ lemma sendIPC_enqueue_ccorres_helper:
apply (fastforce simp: h_t_valid_clift)
apply (fastforce simp: h_t_valid_clift)
apply (case_tac ep, simp_all add: EPState_Idle_def EPState_Send_def)[1]
-- "IdleEP case"
\<comment> \<open>IdleEP case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (simp add: setEndpoint_def split_def)
@ -5392,16 +5392,16 @@ lemma sendIPC_enqueue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (rule cpspace_relation_ep_update_ep', assumption+)
apply (clarsimp simp: cendpoint_relation_def Let_def
mask_def [where n=2] EPState_Send_def)
apply (clarsimp simp: tcb_queue_relation'_def is_aligned_neg_mask)
apply (simp add: isSendEP_def isRecvEP_def)
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cnotification_relation_ep_queue, assumption+)
@ -5412,7 +5412,7 @@ lemma sendIPC_enqueue_ccorres_helper:
apply (simp only:projectKOs injectKO_ep objBits_simps)
apply clarsimp
apply (clarsimp simp: obj_at'_def projectKOs)
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (simp add: carch_state_relation_def typ_heap_simps')
@ -5421,7 +5421,7 @@ lemma sendIPC_enqueue_ccorres_helper:
apply (simp add: objBits_simps')
apply (simp add: objBits_simps)
apply assumption
-- "SendEP case"
\<comment> \<open>SendEP case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (simp add: setEndpoint_def split_def)
@ -5431,10 +5431,10 @@ lemma sendIPC_enqueue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (rule cpspace_relation_ep_update_ep', assumption+)
apply (clarsimp simp: cendpoint_relation_def Let_def
mask_def [where n=2] EPState_Send_def
@ -5443,7 +5443,7 @@ lemma sendIPC_enqueue_ccorres_helper:
valid_ep'_def
dest: tcb_queue_relation_next_not_NULL)
apply (simp add: isSendEP_def isRecvEP_def)
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cnotification_relation_ep_queue, assumption+)
@ -5453,7 +5453,7 @@ lemma sendIPC_enqueue_ccorres_helper:
apply (erule(2) map_to_ko_at_updI')
apply (clarsimp simp: objBitsKO_def)
apply (clarsimp simp: obj_at'_def projectKOs)
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (simp add: carch_state_relation_def typ_heap_simps')
@ -5501,7 +5501,7 @@ lemma sendIPC_ccorres [corres]:
and (\<lambda>s. \<forall>d p. thread \<notin> set (ksReadyQueues s (d, p)))"
in ccorres_guard_imp2 [where A'=UNIV])
apply wpc
-- "RecvEP case"
\<comment> \<open>RecvEP case\<close>
apply (rule ccorres_cond_false)
apply (rule ccorres_cond_true)
apply (intro ccorres_rhs_assoc)
@ -5614,15 +5614,15 @@ lemma sendIPC_ccorres [corres]:
option_to_ptr_def option_to_0_def
split: bool.split_asm)
-- "IdleEP case"
\<comment> \<open>IdleEP case\<close>
apply (rule ccorres_cond_true)
apply (rule ccorres_cond)
apply (clarsimp simp: from_bool_def split: bool.split)
-- "blocking case"
\<comment> \<open>blocking case\<close>
apply (intro ccorres_rhs_assoc)
apply csymbr
apply (simp only:)
-- "apply (ctac (trace, no_vcg,c_lines 6) add: sendIPC_block_ccorres_helper)"
\<comment> \<open>apply (ctac (trace, no_vcg,c_lines 6) add: sendIPC_block_ccorres_helper)\<close>
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
@ -5642,14 +5642,14 @@ lemma sendIPC_ccorres [corres]:
apply wp
apply (clarsimp simp: guard_is_UNIV_def)
apply (rule ccorres_return_Skip)
-- "SendEP case"
\<comment> \<open>SendEP case\<close>
apply (rule ccorres_cond_true)
apply (rule ccorres_cond)
apply (clarsimp simp: from_bool_def split: bool.split)
-- "blocking case"
\<comment> \<open>blocking case\<close>
apply (intro ccorres_rhs_assoc)
apply csymbr
-- "apply (ctac (no_vcg,c_lines 6) add: sendIPC_block_ccorres_helper)"
\<comment> \<open>apply (ctac (no_vcg,c_lines 6) add: sendIPC_block_ccorres_helper)\<close>
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
@ -5815,7 +5815,7 @@ lemma receiveIPC_enqueue_ccorres_helper:
apply (fastforce simp: h_t_valid_clift)
apply (fastforce simp: h_t_valid_clift)
apply (case_tac ep, simp_all add: EPState_Idle_def EPState_Recv_def)[1]
-- "RecvEP case"
\<comment> \<open>RecvEP case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (simp add: setEndpoint_def split_def)
@ -5825,10 +5825,10 @@ lemma receiveIPC_enqueue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (rule cpspace_relation_ep_update_ep', assumption+)
apply (clarsimp simp: cendpoint_relation_def Let_def
mask_def [where n=2] EPState_Recv_def
@ -5837,7 +5837,7 @@ lemma receiveIPC_enqueue_ccorres_helper:
valid_ep'_def
dest: tcb_queue_relation_next_not_NULL)
apply (simp add: isSendEP_def isRecvEP_def)
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cnotification_relation_ep_queue, assumption+)
@ -5847,7 +5847,7 @@ lemma receiveIPC_enqueue_ccorres_helper:
apply (erule(2) map_to_ko_at_updI')
apply (clarsimp simp: objBitsKO_def)
apply (clarsimp simp: obj_at'_def projectKOs)
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (simp add: carch_state_relation_def typ_heap_simps')
@ -5856,7 +5856,7 @@ lemma receiveIPC_enqueue_ccorres_helper:
apply (simp add: objBits_simps')
apply (simp add: objBits_simps)
apply assumption
-- "IdleEP case"
\<comment> \<open>IdleEP case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (simp add: setEndpoint_def split_def)
@ -5866,16 +5866,16 @@ lemma receiveIPC_enqueue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (rule cpspace_relation_ep_update_ep', assumption+)
apply (clarsimp simp: cendpoint_relation_def Let_def
mask_def [where n=2] EPState_Recv_def)
apply (clarsimp simp: tcb_queue_relation'_def is_aligned_neg_mask)
apply (simp add: isSendEP_def isRecvEP_def)
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cnotification_relation_ep_queue, assumption+)
@ -5885,7 +5885,7 @@ lemma receiveIPC_enqueue_ccorres_helper:
apply (erule(2) map_to_ko_at_updI')
apply (clarsimp simp: objBitsKO_def)
apply (clarsimp simp: obj_at'_def projectKOs)
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (simp add: carch_state_relation_def typ_heap_simps')
@ -5938,7 +5938,7 @@ lemma receiveIPC_dequeue_ccorres_helper:
apply (intro conjI impI allI)
apply (fastforce simp: h_t_valid_clift)
apply (fastforce simp: h_t_valid_clift)
-- "empty case"
\<comment> \<open>empty case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (frule iffD1 [OF tcb_queue_head_empty_iff
@ -5955,21 +5955,21 @@ lemma receiveIPC_dequeue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (rule cpspace_relation_ep_update_ep, assumption+)
apply (simp add: cendpoint_relation_def Let_def EPState_Idle_def
tcb_queue_relation'_def)
apply simp
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cnotification_relation_ep_queue [OF invs_sym'], assumption+)
apply simp
apply (erule (1) map_to_ko_atI')
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (simp add: carch_state_relation_def typ_heap_simps')
@ -5979,7 +5979,7 @@ lemma receiveIPC_dequeue_ccorres_helper:
apply (simp add: objBits_simps)
apply assumption
apply (clarsimp simp: cendpoint_relation_def Let_def tcb_queue_relation'_def)
-- "non-empty case"
\<comment> \<open>non-empty case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (frule tcb_queue_head_empty_iff [OF tcb_queue_relation'_queue_rel])
@ -5995,10 +5995,10 @@ lemma receiveIPC_dequeue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (rule cpspace_relation_ep_update_ep, assumption+)
apply (clarsimp simp: cendpoint_relation_def Let_def
isRecvEP_def isSendEP_def
@ -6011,13 +6011,13 @@ lemma receiveIPC_dequeue_ccorres_helper:
split del: if_split)
apply (clarsimp split: if_split)
apply simp
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cnotification_relation_ep_queue [OF invs_sym'], assumption+)
apply simp
apply (erule (1) map_to_ko_atI')
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (simp add: carch_state_relation_def typ_heap_simps')
@ -6067,12 +6067,12 @@ lemma completeSignal_ccorres:
split: Structures_H.ntfn.splits)
apply ceqv
apply wpc
-- "IdleNtfn case"
\<comment> \<open>IdleNtfn case\<close>
apply (clarsimp simp: NtfnState_Idle_def NtfnState_Active_def)
apply csymbr
apply (rule ccorres_cond_false)
apply (rule ccorres_fail)
-- "ActiveNtfn case"
\<comment> \<open>ActiveNtfn case\<close>
apply (clarsimp, csymbr, rule ccorres_cond_true)
apply (rule ccorres_rhs_assoc)+
apply (rule ccorres_rhs_assoc2)
@ -6107,7 +6107,7 @@ lemma completeSignal_ccorres:
apply (clarsimp simp: guard_is_UNIV_def ARM_HYP_H.badgeRegister_def
ARM_HYP.badgeRegister_def Kernel_C.badgeRegister_def
Kernel_C.R0_def)
-- "WaitingNtfn case"
\<comment> \<open>WaitingNtfn case\<close>
apply (clarsimp simp: NtfnState_Active_def NtfnState_Waiting_def)
apply csymbr
apply (rule ccorres_cond_false)
@ -6211,7 +6211,7 @@ lemma receiveIPC_ccorres [corres]:
and ko_at' ep (capEPPtr cap)"
in ccorres_guard_imp2 [where A'=UNIV])
apply wpc
-- "RecvEP case"
\<comment> \<open>RecvEP case\<close>
apply (rule ccorres_cond_true)
apply csymbr
apply (simp only: case_bool_If from_bool_neq_0)
@ -6240,7 +6240,7 @@ lemma receiveIPC_ccorres [corres]:
apply (clarsimp simp: guard_is_UNIV_def)
apply simp
apply (ctac add: doNBRecvFailedTransfer_ccorres[unfolded dc_def])
-- "IdleEP case"
\<comment> \<open>IdleEP case\<close>
apply (rule ccorres_cond_true)
apply csymbr
apply (simp only: case_bool_If from_bool_neq_0)
@ -6266,7 +6266,7 @@ lemma receiveIPC_ccorres [corres]:
apply (clarsimp simp: guard_is_UNIV_def)
apply simp
apply (ctac add: doNBRecvFailedTransfer_ccorres[unfolded dc_def])
-- "SendEP case"
\<comment> \<open>SendEP case\<close>
apply (thin_tac "isBlockinga = from_bool P" for P)
apply (rule ccorres_cond_false)
apply (rule ccorres_cond_true)
@ -6517,7 +6517,7 @@ lemma sendSignal_dequeue_ccorres_helper:
apply (intro conjI impI allI)
apply (fastforce simp: h_t_valid_clift)
apply (fastforce simp: h_t_valid_clift)
-- "empty case"
\<comment> \<open>empty case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (frule iffD1 [OF tcb_queue_head_empty_iff
@ -6535,21 +6535,21 @@ lemma sendSignal_dequeue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cendpoint_relation_ntfn_queue [OF invs_sym'], assumption+)
apply simp+
apply (erule (1) map_to_ko_atI')
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (rule cpspace_relation_ntfn_update_ntfn, assumption+)
apply (simp add: cnotification_relation_def Let_def NtfnState_Idle_def
tcb_queue_relation'_def)
apply simp
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (simp add: carch_state_relation_def typ_heap_simps')
@ -6560,7 +6560,7 @@ lemma sendSignal_dequeue_ccorres_helper:
apply assumption
apply (clarsimp simp: cnotification_relation_def Let_def
tcb_queue_relation'_def)
-- "non-empty case"
\<comment> \<open>non-empty case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (frule tcb_queue_head_empty_iff [OF tcb_queue_relation'_queue_rel])
@ -6577,16 +6577,16 @@ lemma sendSignal_dequeue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cendpoint_relation_ntfn_queue [OF invs_sym'], assumption+)
apply simp+
apply (erule (1) map_to_ko_atI')
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (rule cpspace_relation_ntfn_update_ntfn, assumption+)
apply (clarsimp simp: cnotification_relation_def Let_def
isWaitingNtfn_def
@ -6599,7 +6599,7 @@ lemma sendSignal_dequeue_ccorres_helper:
split del: if_split)
apply (clarsimp split: if_split)
apply simp
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (simp add: carch_state_relation_def typ_heap_simps')
@ -6674,7 +6674,7 @@ lemma sendSignal_ccorres [corres]:
option_to_ctcb_ptr (ntfnBoundTCB ntfn) \<noteq> NULL"
in ccorres_gen_asm)
apply wpc
-- "IdleNtfn case"
\<comment> \<open>IdleNtfn case\<close>
apply (rule ccorres_cond_true)
apply (rule ccorres_rhs_assoc)+
apply (rule_tac xf'=ret__unsigned_'
@ -6722,7 +6722,7 @@ lemma sendSignal_ccorres [corres]:
Collect_const_mem)
apply (case_tac ts, simp_all add: receiveBlocked_def typ_heap_simps
cthread_state_relation_def "StrictC'_thread_state_defs")[1]
-- "ActiveNtfn case"
\<comment> \<open>ActiveNtfn case\<close>
apply (rename_tac old_badge)
apply (rule ccorres_cond_false)
apply (rule ccorres_cond_false)
@ -6752,7 +6752,7 @@ lemma sendSignal_ccorres [corres]:
apply (simp add: objBits_simps')
apply (simp add: objBits_simps)
apply assumption
-- "WaitingNtfn case"
\<comment> \<open>WaitingNtfn case\<close>
apply (rule ccorres_cond_false)
apply (rule ccorres_cond_true)
apply wpc
@ -6947,7 +6947,7 @@ lemma receiveSignal_enqueue_ccorres_helper:
apply (fastforce simp: h_t_valid_clift)
apply (fastforce simp: h_t_valid_clift)
apply (case_tac "ntfnObj ntfn", simp_all add: NtfnState_Idle_def NtfnState_Waiting_def)[1]
-- "IdleNtfn case"
\<comment> \<open>IdleNtfn case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (simp add: setNotification_def split_def)
@ -6957,10 +6957,10 @@ lemma receiveSignal_enqueue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cendpoint_relation_ntfn_queue, assumption+)
@ -6970,7 +6970,7 @@ lemma receiveSignal_enqueue_ccorres_helper:
apply (erule(2) map_to_ko_at_updI')
apply (clarsimp simp: objBitsKO_def)
apply (clarsimp simp: obj_at'_def projectKOs)
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (rule cpspace_relation_ntfn_update_ntfn', assumption+)
apply (case_tac "ntfn", simp_all)[1]
apply (clarsimp simp: cnotification_relation_def Let_def
@ -6979,7 +6979,7 @@ lemma receiveSignal_enqueue_ccorres_helper:
valid_ntfn'_def
dest: tcb_queue_relation_next_not_NULL)
apply (simp add: isWaitingNtfn_def)
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
subgoal by (clarsimp simp: comp_def)
apply (simp add: carch_state_relation_def typ_heap_simps')
@ -6988,7 +6988,7 @@ lemma receiveSignal_enqueue_ccorres_helper:
apply (simp add: objBits_simps')
apply (simp add: objBits_simps)
apply assumption
-- "WaitingNtfn case"
\<comment> \<open>WaitingNtfn case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (simp add: setNotification_def split_def)
@ -6998,10 +6998,10 @@ lemma receiveSignal_enqueue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cendpoint_relation_ntfn_queue, assumption+)
@ -7011,7 +7011,7 @@ lemma receiveSignal_enqueue_ccorres_helper:
apply (erule(2) map_to_ko_at_updI')
apply (clarsimp simp: objBitsKO_def)
apply (clarsimp simp: obj_at'_def projectKOs)
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (rule cpspace_relation_ntfn_update_ntfn', assumption+)
apply (case_tac "ntfn", simp_all)[1]
apply (clarsimp simp: cnotification_relation_def Let_def
@ -7019,7 +7019,7 @@ lemma receiveSignal_enqueue_ccorres_helper:
split: if_split)
apply (fastforce simp: tcb_queue_relation'_def is_aligned_neg_mask)
apply (simp add: isWaitingNtfn_def)
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (simp add: carch_state_relation_def typ_heap_simps')
@ -7067,7 +7067,7 @@ lemma receiveSignal_ccorres [corres]:
apply ceqv
apply wpc
-- "IdleNtfn case"
\<comment> \<open>IdleNtfn case\<close>
apply (rule ccorres_cond_true)
apply csymbr
apply (simp only: case_bool_If from_bool_neq_0)
@ -7092,7 +7092,7 @@ lemma receiveSignal_ccorres [corres]:
apply (clarsimp simp: guard_is_UNIV_def)
apply simp
apply (ctac add: doNBRecvFailedTransfer_ccorres[unfolded dc_def])
-- "ActiveNtfn case"
\<comment> \<open>ActiveNtfn case\<close>
apply (rename_tac badge)
apply (rule ccorres_cond_false)
apply (rule ccorres_cond_true)
@ -7133,7 +7133,7 @@ lemma receiveSignal_ccorres [corres]:
apply wp
apply (clarsimp simp: guard_is_UNIV_def)
apply (clarsimp simp: ARM_HYP.badgeRegister_def Kernel_C.R0_def)
-- "WaitingNtfn case"
\<comment> \<open>WaitingNtfn case\<close>
apply (rename_tac list)
apply (rule ccorres_cond_true)
apply csymbr

View File

@ -105,7 +105,7 @@ lemma sym_refs_tcb_vcpu':
lemma ko_at'_tcb_vcpu_not_NULL:
"\<lbrakk> ko_at' (tcb::tcb) t s ; valid_objs' s ; no_0_obj' s ; atcbVCPUPtr (tcbArch tcb) = Some p \<rbrakk>
\<Longrightarrow> 0 < p"
-- "when C pointer is NULL, need this to show atcbVCPUPtr is None"
\<comment> \<open>when C pointer is NULL, need this to show atcbVCPUPtr is None\<close>
unfolding valid_pspace'_def
by (fastforce simp: valid_tcb'_def valid_arch_tcb'_def word_gt_0 typ_at'_no_0_objD
dest: valid_objs_valid_tcb')

View File

@ -266,7 +266,7 @@ lemma handleSyscall_ccorres:
apply (rule ccorres_split_nothrow_novcg)
apply wpc
prefer 3
-- "SysSend"
\<comment> \<open>SysSend\<close>
apply (clarsimp simp: syscall_from_H_def syscall_defs)
apply (rule ccorres_cond_empty |rule ccorres_cond_univ)+
apply (simp add: handleSend_def)
@ -301,7 +301,7 @@ lemma handleSyscall_ccorres:
| strengthen None_drop invs'_irq_strg | subst Ex_Some_conv)+
apply (vcg exspec=handleInvocation_modifies)
prefer 3
-- "SysNBSend"
\<comment> \<open>SysNBSend\<close>
apply (clarsimp simp: syscall_from_H_def syscall_defs)
apply (rule ccorres_cond_empty |rule ccorres_cond_univ)+
apply (simp add: handleSend_def)
@ -334,7 +334,7 @@ lemma handleSyscall_ccorres:
apply (wp dmo'_getActiveIRQ_non_kernel getActiveIRQ_neq_Some0xFF | simp
| strengthen None_drop invs'_irq_strg | subst Ex_Some_conv)+
apply (vcg exspec=handleInvocation_modifies)
-- "SysCall"
\<comment> \<open>SysCall\<close>
apply (clarsimp simp: syscall_from_H_def syscall_defs)
apply (rule ccorres_cond_empty |rule ccorres_cond_univ)+
apply (simp add: handleCall_def)
@ -368,7 +368,7 @@ lemma handleSyscall_ccorres:
| strengthen None_drop invs'_irq_strg | subst Ex_Some_conv)+
apply (vcg exspec=handleInvocation_modifies)
prefer 2
-- "SysRecv"
\<comment> \<open>SysRecv\<close>
apply (clarsimp simp: syscall_from_H_def syscall_defs)
apply (rule ccorres_cond_empty |rule ccorres_cond_univ)+
apply (simp add: liftE_bind)
@ -377,7 +377,7 @@ lemma handleSyscall_ccorres:
apply (rule ccorres_returnOk_skip[unfolded returnOk_def, simplified])
apply wp
prefer 2
-- "SysReply"
\<comment> \<open>SysReply\<close>
apply (clarsimp simp: syscall_from_H_def syscall_defs)
apply (rule ccorres_cond_empty |rule ccorres_cond_univ)+
apply (simp add: liftE_bind)
@ -385,7 +385,7 @@ lemma handleSyscall_ccorres:
apply (ctac (no_vcg) add: handleReply_ccorres)
apply (rule ccorres_returnOk_skip[unfolded returnOk_def, simplified])
apply wp
-- "SysReplyRecv"
\<comment> \<open>SysReplyRecv\<close>
apply (clarsimp simp: syscall_from_H_def syscall_defs)
apply (rule ccorres_cond_empty |rule ccorres_cond_univ)+
apply (simp add: liftE_bind bind_assoc)
@ -401,7 +401,7 @@ lemma handleSyscall_ccorres:
in hoare_post_imp)
apply (simp add: ct_in_state'_def)
apply (wp handleReply_sane handleReply_ct_not_ksQ)
-- "SysYield"
\<comment> \<open>SysYield\<close>
apply (clarsimp simp: syscall_from_H_def syscall_defs)
apply (rule ccorres_cond_empty |rule ccorres_cond_univ)+
apply (simp add: liftE_bind)
@ -409,7 +409,7 @@ lemma handleSyscall_ccorres:
apply (ctac (no_vcg) add: handleYield_ccorres)
apply (rule ccorres_returnOk_skip[unfolded returnOk_def, simplified])
apply wp
-- "SysNBRecv"
\<comment> \<open>SysNBRecv\<close>
apply (clarsimp simp: syscall_from_H_def syscall_defs)
apply (rule ccorres_cond_empty |rule ccorres_cond_univ)+
apply (simp add: liftE_bind)
@ -417,7 +417,7 @@ lemma handleSyscall_ccorres:
apply (ctac (no_vcg) add: handleRecv_ccorres)
apply (rule ccorres_returnOk_skip[unfolded returnOk_def, simplified])
apply wp
-- " rest of body"
\<comment> \<open>rest of body\<close>
apply ceqv
apply (ctac (no_vcg) add: schedule_ccorres)
apply (rule ccorres_add_return2)

View File

@ -3502,7 +3502,7 @@ proof -
note hacky_cte = retype_ctes_helper [where sz = "objBitsKO kotcb" and ko = kotcb and ptr = "ctcb_ptr_to_tcb_ptr p",
OF pal pds pno al _ _ mko, simplified new_cap_addrs_def, simplified]
-- "Ugh"
\<comment> \<open>Ugh\<close>
moreover have
"\<And>y. y \<in> ptr_val ` (CTypesDefs.ptr_add (cte_Ptr (ctcb_ptr_to_tcb_ptr p)) \<circ> of_nat) ` {k. k < 5}
= (y && ~~ mask tcbBlockSizeBits = ctcb_ptr_to_tcb_ptr p \<and> y && mask tcbBlockSizeBits \<in> dom tcb_cte_cases)"
@ -3613,7 +3613,7 @@ proof -
apply (case_tac r,
simp_all add: "StrictC'_register_defs" eval_nat_numeral
atcbContext_def newArchTCB_def newContext_def
initContext_def)[1] -- "takes ages"
initContext_def)[1] \<comment> \<open>takes ages\<close>
apply (simp add: thread_state_lift_def eval_nat_numeral atcbContextGet_def)+
apply (simp add: timeSlice_def)
apply (simp add: cfault_rel_def seL4_Fault_lift_def seL4_Fault_get_tag_def Let_def
@ -3739,15 +3739,15 @@ proof -
apply (simp add: rl kotcb_def projectKOs rl_tcb rl_cte)
apply (elim conjE)
apply (intro conjI)
-- "cte"
\<comment> \<open>cte\<close>
apply (erule cmap_relation_retype2)
apply (simp add:ccte_relation_nullCap nullMDBNode_def nullPointer_def)
-- "tcb"
\<comment> \<open>tcb\<close>
apply (erule cmap_relation_updI2 [where dest = "ctcb_ptr_to_tcb_ptr p" and f = "tcb_ptr_to_ctcb_ptr", simplified])
apply (rule map_comp_simps)
apply (rule pks)
apply (rule tcb_rel)
-- "ep"
\<comment> \<open>ep\<close>
apply (erule iffD2 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply (simp add: cendpoint_relation_def Let_def)
apply (subst endpoint.case_cong)
@ -3756,7 +3756,7 @@ proof -
apply (simp add: tcb_queue_update_other' del: tcb_queue_relation'_empty)
apply (simp add: tcb_queue_update_other' ep2)
apply clarsimp
-- "ntfn"
\<comment> \<open>ntfn\<close>
apply (erule iffD2 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply (simp add: cnotification_relation_def Let_def)
apply (subst ntfn.case_cong)
@ -5981,7 +5981,7 @@ proof -
framesize_to_H_def cap_to_H_simps cap_small_frame_cap_lift
vmrights_to_H_def mask_def vm_rights_defs)
-- "Page objects: could possibly fix the duplication here"
\<comment> \<open>Page objects: could possibly fix the duplication here\<close>
apply (cinit' lift: t_' regionBase_' userSize_' deviceMemory_')
apply (simp add: object_type_from_H_def Kernel_C_defs)
apply (simp add: ccorres_cond_univ_iff ccorres_cond_empty_iff
@ -6063,7 +6063,7 @@ proof -
cl_valid_cap_def c_valid_cap_def
is_aligned_neg_mask_eq_concrete[THEN sym])
-- "PageTableObject"
\<comment> \<open>PageTableObject\<close>
apply (cinit' lift: t_' regionBase_' userSize_' deviceMemory_')
apply (simp add: object_type_from_H_def Kernel_C_defs)
apply (simp add: ccorres_cond_univ_iff ccorres_cond_empty_iff
@ -6098,7 +6098,7 @@ proof -
apply (rule sym)
apply (simp add: is_aligned_neg_mask_eq'[symmetric] is_aligned_weaken)
-- "PageDirectoryObject"
\<comment> \<open>PageDirectoryObject\<close>
apply (cinit' lift: t_' regionBase_' userSize_' deviceMemory_')
apply (simp add: object_type_from_H_def Kernel_C_defs)
apply (simp add: ccorres_cond_univ_iff ccorres_cond_empty_iff
@ -6145,7 +6145,7 @@ proof -
apply fastforce
apply ((clarsimp simp: is_aligned_no_overflow'[where n=14, simplified]
field_simps is_aligned_mask[symmetric] mask_AND_less_0)+)[3]
-- "VCPU"
\<comment> \<open>VCPU\<close>
apply (cinit' lift: t_' regionBase_' userSize_' deviceMemory_')
apply (simp add: object_type_from_H_def Kernel_C_defs)
apply ccorres_rewrite

View File

@ -102,7 +102,7 @@ lemma cap_get_tag_isCap0:
apply (erule ccap_relationE)
apply (simp add: cap_to_H_def cap_lift_def Let_def isArchCap_tag_def2 isArchCap_def)
by (clarsimp simp: isCap_simps cap_tag_defs word_le_nat_alt pageSize_def Let_def
split: if_split_asm) -- "takes a while"
split: if_split_asm) \<comment> \<open>takes a while\<close>
lemma cap_get_tag_isCap:
@ -1541,7 +1541,7 @@ lemma cmap_relation_cong:
apply (rule Some_the [where f = cm'])
apply (erule subsetD)
apply (erule imageI)
-- "clag"
\<comment> \<open>clag\<close>
apply simp
apply (erule conjE)
apply (drule equalityD1)
@ -2124,7 +2124,7 @@ lemma cap_get_tag_isCap_ArchObject0:
apply -
apply (erule ccap_relationE)
apply (simp add: cap_to_H_def cap_lift_def Let_def isArchCap_def)
by (clarsimp simp: isCap_simps cap_tag_defs word_le_nat_alt pageSize_def Let_def split: if_split_asm) -- "takes a while"
by (clarsimp simp: isCap_simps cap_tag_defs word_le_nat_alt pageSize_def Let_def split: if_split_asm) \<comment> \<open>takes a while\<close>
lemma cap_get_tag_isCap_ArchObject:
assumes cr: "ccap_relation (capability.ArchObjectCap cap) cap'"

View File

@ -264,7 +264,7 @@ proof -
apply (rule_tac R="\<lambda>s. l1 = ksReadyQueuesL1Bitmap s curdom \<and> curdom \<le> maxDomain"
in ccorres_cond)
subgoal by (fastforce dest!: rf_sr_cbitmap_L1_relation simp: cbitmap_L1_relation_def)
prefer 2 -- "switchToIdleThread"
prefer 2 \<comment> \<open>switchToIdleThread\<close>
apply (ctac(no_vcg) add: switchToIdleThread_ccorres)
apply clarsimp
apply (rule ccorres_rhs_assoc)+

View File

@ -926,7 +926,7 @@ lemma cready_queues_relation_null_queue_ptrs:
apply (clarsimp simp: tcb_null_ep_ptrs_def)
apply (case_tac z, case_tac a)
apply simp
-- "clag"
\<comment> \<open>clag\<close>
apply (rule ext)
apply (case_tac "mp' x")
apply (frule compD [OF same])

View File

@ -3887,7 +3887,7 @@ lemma bindNotification_ccorres:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (rule cpspace_relation_ntfn_update_ntfn, assumption+)
apply (clarsimp simp: cnotification_relation_def Let_def
mask_def [where n=2] NtfnState_Waiting_def)
@ -4109,7 +4109,7 @@ lemma decodeBindNotification_ccorres:
apply (rule_tac P="\<lambda>s. ksCurThread s = thread" in ccorres_cross_over_guard)
apply (simp add: bindE_bind_linearise del: Collect_const)
apply wpc
-- "IdleNtfn"
\<comment> \<open>IdleNtfn\<close>
apply (simp add: case_option_If del: Collect_const)
apply (rule ccorres_Cond_rhs_Seq)
apply (clarsimp simp: isWaitingNtfn_def from_bool_neq_0)

View File

@ -3152,7 +3152,7 @@ lemma unmapPage_ccorres:
apply (simp add: liftE_bindE Collect_False bind_bindE_assoc
del: Collect_const)
apply (rule ccorres_splitE_novcg[where r'=dc and xf'=xfdc])
-- "ARMSmallPage"
\<comment> \<open>ARMSmallPage\<close>
apply (rule ccorres_Cond_rhs)
apply (simp add: gen_framesize_to_H_def dc_def[symmetric])
apply (rule ccorres_rhs_assoc)+
@ -3195,7 +3195,7 @@ lemma unmapPage_ccorres:
apply (wp)
apply simp
apply (vcg exspec=lookupPTSlot_modifies)
-- "ARMLargePage"
\<comment> \<open>ARMLargePage\<close>
apply (rule ccorres_Cond_rhs)
apply (simp add: gen_framesize_to_H_def dc_def[symmetric])
apply (rule ccorres_rhs_assoc)+
@ -3289,7 +3289,7 @@ lemma unmapPage_ccorres:
apply (wp lookupPTSlot_inv Arch_R.lookupPTSlot_aligned
lookupPTSlot_page_table_at' | simp add: ptBits_eq)+
apply (vcg exspec=lookupPTSlot_modifies)
-- "ARMSection"
\<comment> \<open>ARMSection\<close>
apply (rule ccorres_Cond_rhs)
apply (rule ccorres_rhs_assoc)+
apply (csymbr, csymbr)
@ -3324,7 +3324,7 @@ lemma unmapPage_ccorres:
apply simp
apply wp
apply (simp add: guard_is_UNIV_def)
-- "ARMSuperSection"
\<comment> \<open>ARMSuperSection\<close>
apply (rule ccorres_Cond_rhs)
apply (rule ccorres_rhs_assoc)+
apply csymbr

View File

@ -118,11 +118,11 @@ definition
definition
ep_at_C' :: "word32 \<Rightarrow> heap_raw_state \<Rightarrow> bool"
where
"ep_at_C' p h \<equiv> Ptr p \<in> dom (clift h :: endpoint_C typ_heap)" -- "endpoint_lift is total"
"ep_at_C' p h \<equiv> Ptr p \<in> dom (clift h :: endpoint_C typ_heap)" \<comment> \<open>endpoint_lift is total\<close>
definition
ntfn_at_C' :: "word32 \<Rightarrow> heap_raw_state \<Rightarrow> bool"
where -- "notification_lift is total"
where \<comment> \<open>notification_lift is total\<close>
"ntfn_at_C' p h \<equiv> Ptr p \<in> dom (clift h :: notification_C typ_heap)"
definition
@ -235,8 +235,8 @@ abbreviation
ARMSuperSection :: "vmpage_size" where
"ARMSuperSection == ARM_HYP.ARMSuperSection"
-- "ARMSmallFrame is treated in a separate cap in C,
so needs special treatment in ccap_relation"
\<comment> \<open>ARMSmallFrame is treated in a separate cap in C,
so needs special treatment in ccap_relation\<close>
definition
framesize_to_H:: "word32 \<Rightarrow> vmpage_size" where
"framesize_to_H c \<equiv>
@ -244,7 +244,7 @@ framesize_to_H:: "word32 \<Rightarrow> vmpage_size" where
else if c = scast Kernel_C.ARMSection then ARMSection
else ARMSuperSection"
-- "Use this for results of generic_frame_cap_get_capFSize"
\<comment> \<open>Use this for results of generic_frame_cap_get_capFSize\<close>
definition
gen_framesize_to_H:: "word32 \<Rightarrow> vmpage_size" where
"gen_framesize_to_H c \<equiv>

View File

@ -2791,7 +2791,7 @@ lemma decodeX64FrameInvocation_ccorres:
apply (rule_tac t=a and s="hd args" in ssubst, fastforce)
apply (rule_tac t=b and s="hd (tl args)" in ssubst, fastforce)
apply (vcg exspec=getSyscallArg_modifies)
-- "PageMap"
\<comment> \<open>PageMap\<close>
apply (rule ccorres_rhs_assoc)+
apply csymbr+
apply (simp add: if_1_0_0 word_less_nat_alt del: Collect_const)
@ -5797,7 +5797,7 @@ proof -
apply (clarsimp cong: StateSpace.state.fold_congs globals.fold_congs)
apply (rule ccorres_Cond_rhs) (* IN invocations *)
apply (erule ccorres_disj_division)
-- "In8"
\<comment> \<open>In8\<close>
apply (clarsimp simp: invocation_eq_use_types cong: list.case_cong)
apply (rule ccorres_rhs_assoc)+
apply (rule ccorres_Cond_rhs_Seq) (* length error *)
@ -5841,7 +5841,7 @@ proof -
apply wpsimp
apply (vcg exspec=getSyscallArg_modifies)
apply (erule ccorres_disj_division)
-- "In16"
\<comment> \<open>In16\<close>
apply (clarsimp simp: invocation_eq_use_types cong: list.case_cong)
apply (rule ccorres_rhs_assoc)+
apply (rule ccorres_Cond_rhs_Seq) (* length error *)
@ -5884,7 +5884,7 @@ proof -
apply (vcg exspec=ensurePortOperationAllowed_modifies)
apply wpsimp
apply (vcg exspec=getSyscallArg_modifies)
-- "In32"
\<comment> \<open>In32\<close>
apply (clarsimp simp: invocation_eq_use_types cong: list.case_cong)
apply (rule ccorres_rhs_assoc)+
apply (rule ccorres_Cond_rhs_Seq) (* length error *)
@ -5930,7 +5930,7 @@ proof -
apply clarsimp
apply (rule ccorres_Cond_rhs) (* OUT invocations *)
apply (erule ccorres_disj_division)
-- "Out8"
\<comment> \<open>Out8\<close>
apply (clarsimp simp: invocation_eq_use_types cong: list.case_cong)
apply (rule ccorres_rhs_assoc)+
apply (rule ccorres_Cond_rhs_Seq) (* length error *)
@ -5980,7 +5980,7 @@ proof -
apply wpsimp
apply (vcg exspec=getSyscallArg_modifies)
apply (erule ccorres_disj_division)
-- "Out16"
\<comment> \<open>Out16\<close>
apply (clarsimp simp: invocation_eq_use_types cong: list.case_cong)
apply (rule ccorres_rhs_assoc)+
apply (rule ccorres_Cond_rhs_Seq) (* length error *)
@ -6029,7 +6029,7 @@ proof -
apply (vcg exspec=getSyscallArg_modifies)
apply wpsimp
apply (vcg exspec=getSyscallArg_modifies)
-- "Out32"
\<comment> \<open>Out32\<close>
apply (clarsimp simp: invocation_eq_use_types cong: list.case_cong)
apply (rule ccorres_rhs_assoc)+
apply (rule ccorres_Cond_rhs_Seq) (* length error *)

View File

@ -302,7 +302,7 @@ lemmas ccorres_move_array_assertion_cnode_ctes [corres_pre]
ccorres_move_Guard [OF array_assertion_abs_cnode_ctes]
lemma locateSlotCNode_ccorres [corres]:
assumes gl: "\<And>v s. globals (xfu v s) = globals s" -- "for state rel. preservation"
assumes gl: "\<And>v s. globals (xfu v s) = globals s" \<comment> \<open>for state rel. preservation\<close>
and fg: "\<And>v s. xf (xfu (\<lambda>_. v) s) = v"
shows "ccorres (\<lambda>v v'. v' = Ptr v) xf \<top> {_. cnode = cnode' \<and> offset = offset'} hs
(locateSlotCNode cnode bits offset)
@ -323,7 +323,7 @@ lemma locateSlotCNode_ccorres [corres]:
done
lemma locateSlotTCB_ccorres [corres]:
assumes gl: "\<And>v s. globals (xfu v s) = globals s" -- "for state rel. preservation"
assumes gl: "\<And>v s. globals (xfu v s) = globals s" \<comment> \<open>for state rel. preservation\<close>
and fg: "\<And>v s. xf (xfu (\<lambda>_. v) s) = v"
shows "ccorres (\<lambda>v v'. v' = Ptr v) xf \<top> {_. cnode = cnode' \<and> offset = offset'} hs
(locateSlotTCB cnode offset)
@ -339,13 +339,13 @@ lemma locateSlotTCB_ccorres [corres]:
lemma getSlotCap_h_val_ccorres [corres]:
fixes p :: "cstate \<Rightarrow> cte_C ptr"
assumes gl: "\<And>v s. globals (xfu v s) = globals s" -- "for state rel. preservation"
assumes gl: "\<And>v s. globals (xfu v s) = globals s" \<comment> \<open>for state rel. preservation\<close>
and fg: "\<And>v s. xf (xfu (\<lambda>_. v) s) = v"
shows "ccorres ccap_relation xf \<top> {s. p s = Ptr a} hs
(getSlotCap a) (Basic (\<lambda>s. xfu (\<lambda>_. h_val (hrs_mem (t_hrs_' (globals s))) (Ptr &(p s\<rightarrow>[''cap_C'']) :: cap_C ptr)) s))"
unfolding getSlotCap_def
apply (rule ccorres_add_UNIV_Int)
apply (cinitlift p) -- "EVIL!"
apply (cinitlift p) \<comment> \<open>EVIL!\<close>
apply simp
apply (rule ccorres_guard_imp2)
apply (rule ccorres_pre_getCTE)

View File

@ -40,9 +40,9 @@ lemma lookupCap_ccorres':
apply (cinit lift: cPtr_' thread_' simp: lookupCapAndSlot_def liftME_def bindE_assoc)
apply (ctac (no_vcg) add: lookupSlotForThread_ccorres')
--"case where lu_ret.status is EXCEPTION_NONE"
\<comment> \<open>case where lu_ret.status is EXCEPTION_NONE\<close>
apply (simp add: split_beta cong:call_ignore_cong)
apply csymbr --"call status_C_update"
apply csymbr \<comment> \<open>call status_C_update\<close>
apply (simp add: Collect_const[symmetric] lookupSlot_raw_rel_def liftE_def
del: Collect_const)
apply (rule ccorres_move_c_guard_cte)
@ -50,18 +50,18 @@ lemma lookupCap_ccorres':
apply (rule ccorres_return_CE [unfolded returnOk_def, simplified], simp+)[1]
apply wp
apply vcg
--"case where lu_ret.status is *not* EXCEPTION_NONE"
\<comment> \<open>case where lu_ret.status is *not* EXCEPTION_NONE\<close>
apply simp
apply (rule ccorres_split_throws)
apply (rule ccorres_rhs_assoc)+
apply csymbr
apply csymbr -- "call cap_null_cap_new_'proc"
apply csymbr \<comment> \<open>call cap_null_cap_new_'proc\<close>
apply csymbr
apply (rule ccorres_return_C_errorE, simp+)[1]
apply vcg
apply (wp | simp)+
-- "last subgoal"
\<comment> \<open>last subgoal\<close>
apply (clarsimp simp: valid_pspace_valid_objs')
apply (intro impI conjI allI)
apply (simp add: lookupSlot_raw_rel_def)
@ -91,10 +91,10 @@ lemma lookupCapAndSlot_ccorres :
apply (cinit lift: thread_' cPtr_')
apply (ctac (no_vcg))
--"case where lu_ret.status is EXCEPTION_NONE"
\<comment> \<open>case where lu_ret.status is EXCEPTION_NONE\<close>
apply (simp add: split_beta cong:call_ignore_cong)
apply csymbr --"call status_C_update"
apply csymbr --"call slot_C_update"
apply csymbr \<comment> \<open>call status_C_update\<close>
apply csymbr \<comment> \<open>call slot_C_update\<close>
apply (simp add: Collect_const[symmetric] lookupSlot_raw_rel_def liftE_bindE
del: Collect_const)
apply (rule ccorres_move_c_guard_cte)
@ -103,7 +103,7 @@ lemma lookupCapAndSlot_ccorres :
apply (clarsimp simp: Bex_def in_monad getSlotCap_def in_getCTE2 cte_wp_at_ctes_of)
apply (erule(1) cmap_relationE1[OF cmap_relation_cte])
apply (simp add: ccte_relation_ccap_relation typ_heap_simps')
--"case where lu_ret.status is *not* EXCEPTION_NONE"
\<comment> \<open>case where lu_ret.status is *not* EXCEPTION_NONE\<close>
apply simp
apply (rule ccorres_split_throws)
apply (rule ccorres_rhs_assoc)+
@ -113,7 +113,7 @@ lemma lookupCapAndSlot_ccorres :
apply vcg
apply (wp | simp)+
-- "last subgoal"
\<comment> \<open>last subgoal\<close>
apply clarsimp
apply (rule conjI, fastforce)
apply clarsimp
@ -166,14 +166,14 @@ lemma lookupSlotForCNodeOp_ccorres':
(lookupSlotForCNodeOp isSource croot capptr depth)
(Call lookupSlotForCNodeOp_'proc)"
apply (cinit lift: capptr_' isSource_' root___struct_cap_C_' depth_')
apply csymbr -- "slot_C_update"
apply csymbr -- "cap_get_capType"
apply csymbr \<comment> \<open>slot_C_update\<close>
apply csymbr \<comment> \<open>cap_get_capType\<close>
apply (rule_tac Q=\<top> and Q'=\<top> in ccorres_if_cond_throws2)
-- "correspondance of Haskell and C conditions"
\<comment> \<open>correspondance of Haskell and C conditions\<close>
apply (clarsimp simp: Collect_const_mem cap_get_tag_isCap)
-- "case where root is *not* a CNode => throw InvalidRoot"
\<comment> \<open>case where root is *not* a CNode => throw InvalidRoot\<close>
apply simp
apply (rule_tac P=\<top> and P' =UNIV in ccorres_from_vcg_throws)
apply (rule allI, rule conseqPre, vcg)
@ -183,7 +183,7 @@ lemma lookupSlotForCNodeOp_ccorres':
apply clarsimp
apply (subst syscall_error_to_H_cases(6), simp+)[1]
-- " case where root is a CNode"
\<comment> \<open>case where root is a CNode\<close>
apply (simp add: rangeCheck_def)
apply csymbr
apply (rule ccorres_Cond_rhs_Seq)
@ -200,12 +200,12 @@ lemma lookupSlotForCNodeOp_ccorres':
apply vcg
apply csymbr
apply (rule_tac Q="\<lambda>s. depth < 2 ^ word_bits" and Q'=\<top> in ccorres_split_unless_throwError_cond)
-- "correspondance of Haskell and C conditions"
\<comment> \<open>correspondance of Haskell and C conditions\<close>
apply (clarsimp simp: Collect_const_mem fromIntegral_def integral_inv)
apply (simp add: word_size unat_of_nat64 word_less_nat_alt
word_less_1[symmetric] linorder_not_le)
-- "case of RangeError"
\<comment> \<open>case of RangeError\<close>
apply (rule_tac P= \<top> and P' =UNIV in ccorres_from_vcg_throws)
apply (rule allI, rule conseqPre, vcg)
apply (clarsimp simp: throwError_def return_def syscall_error_rel_def)
@ -213,20 +213,20 @@ lemma lookupSlotForCNodeOp_ccorres':
apply (subst syscall_error_to_H_cases(4), simp+)[1]
apply (simp add: word_size word_sle_def)
-- "case where there is *no* RangeError"
\<comment> \<open>case where there is *no* RangeError\<close>
apply (rule_tac xf'=lookupSlot_xf in ccorres_rel_imp)
apply (rule_tac r="\<lambda>w w'. w'= Ptr w"
and f="\<lambda> st fl es. fl = scast EXCEPTION_SYSCALL_ERROR \<and>
syscall_error_to_H (errsyscall es) (errlookup_fault es) = Some (FailedLookup isSource st)"
in lookupErrorOnFailure_ccorres)
apply (ctac (no_vcg)) -- "resolveAddressBits"
-- " case where resolveAddressBits results in EXCEPTION_NONE"
apply (ctac (no_vcg)) \<comment> \<open>resolveAddressBits\<close>
\<comment> \<open>case where resolveAddressBits results in EXCEPTION_NONE\<close>
apply clarsimp
apply (rule_tac A=\<top> and A'=UNIV in ccorres_guard_imp2)
apply (rule_tac Q=\<top> and Q'=\<top> in ccorres_if_cond_throws2)
-- "correspondance of Haskell and C conditions"
\<comment> \<open>correspondance of Haskell and C conditions\<close>
apply (clarsimp simp: Collect_const_mem unat_gt_0)
-- " case where bits are remaining"
\<comment> \<open>case where bits are remaining\<close>
apply (rule_tac P= \<top> and P' =UNIV in ccorres_from_vcg_throws)
apply (rule allI, rule conseqPre, vcg)
apply (clarsimp simp: throwError_def return_def)
@ -235,17 +235,17 @@ lemma lookupSlotForCNodeOp_ccorres':
apply (simp add: lookup_fault_depth_mismatch_lift)
apply (erule le_64_mask_eq)
-- " case where *no* bits are remaining"
apply csymbr -- "slot_C_update"
apply csymbr -- "status_C_update"
\<comment> \<open>case where *no* bits are remaining\<close>
apply csymbr \<comment> \<open>slot_C_update\<close>
apply csymbr \<comment> \<open>status_C_update\<close>
apply (rule ccorres_return_CE, simp+)[1]
apply vcg
-- "guard_imp subgoal"
\<comment> \<open>guard_imp subgoal\<close>
apply clarsimp
-- " case where resolveAddressBits does *not* results in EXCEPTION_NONE"
\<comment> \<open>case where resolveAddressBits does *not* results in EXCEPTION_NONE\<close>
apply clarsimp
apply (rule_tac P= \<top> and P' ="\<lbrace>\<exists>v. (lookup_fault_lift (\<acute>current_lookup_fault)) = Some v
\<and> lookup_fault_to_H v = err \<rbrace>"
@ -256,7 +256,7 @@ lemma lookupSlotForCNodeOp_ccorres':
apply (subst syscall_error_to_H_cases(6), simp+)[1]
apply wp
-- "rel_imp"
\<comment> \<open>rel_imp\<close>
apply clarsimp
apply (case_tac x, clarsimp)
apply (simp add: syscall_error_rel_def errstate_def)
@ -266,7 +266,7 @@ lemma lookupSlotForCNodeOp_ccorres':
apply vcg
apply vcg
-- "last subgoal"
\<comment> \<open>last subgoal\<close>
apply (clarsimp simp: if_1_0_0 to_bool_def true_def word_size
fromIntegral_def integral_inv)
apply (case_tac "cap_get_tag root___struct_cap_C = scast cap_cnode_cap")

View File

@ -515,7 +515,7 @@ lemma revokable_ccorres:
(Call isCapRevocable_'proc)"
apply (rule ccorres_gen_asm[where G=\<top>, simplified])
apply (cinit' lift: derivedCap_' srcCap_')
-- "Clear up Arch cap case"
\<comment> \<open>Clear up Arch cap case\<close>
apply csymbr
apply (clarsimp simp: cap_get_tag_isCap split del: if_splits simp del: Collect_const)
apply (rule ccorres_Cond_rhs_Seq)
@ -529,9 +529,9 @@ lemma revokable_ccorres:
apply csymbr
apply (rule_tac P'=UNIV and P=\<top> in ccorres_inst)
apply (cases cap)
-- "Uninteresting caps"
\<comment> \<open>Uninteresting caps\<close>
apply revokable'_hammer+
-- "NotificationCap"
\<comment> \<open>NotificationCap\<close>
apply (simp add: cap_get_tag_isCap isCap_simps ccorres_cond_iffs from_bool_def true_def false_def)
apply (rule ccorres_guard_imp, (rule ccorres_rhs_assoc)+, csymbr, csymbr)
apply (rule ccorres_return_C, clarsimp+)
@ -540,12 +540,12 @@ lemma revokable_ccorres:
apply (frule_tac cap'1=derivedCap in cap_get_tag_NotificationCap[THEN iffD1])
apply (clarsimp simp: cap_get_tag_isCap isCap_simps)
apply (fastforce simp: cap_get_tag_isCap isCap_simps)
-- "IRQHandlerCap"
\<comment> \<open>IRQHandlerCap\<close>
apply (simp add: cap_get_tag_isCap isCap_simps ccorres_cond_iffs from_bool_def true_def false_def)
apply (rule ccorres_guard_imp, csymbr)
apply (rule ccorres_return_C, clarsimp+)
apply (fastforce simp: cap_get_tag_isCap isCap_simps)
-- "EndpointCap"
\<comment> \<open>EndpointCap\<close>
apply (simp add: cap_get_tag_isCap isCap_simps ccorres_cond_iffs from_bool_def true_def false_def)
apply (rule ccorres_guard_imp, (rule ccorres_rhs_assoc)+, csymbr, csymbr)
apply (rule ccorres_return_C, clarsimp+)
@ -554,7 +554,7 @@ lemma revokable_ccorres:
apply (frule_tac cap'1=derivedCap in cap_get_tag_EndpointCap[THEN iffD1])
apply (clarsimp simp: cap_get_tag_isCap isCap_simps)
apply (fastforce simp: cap_get_tag_isCap isCap_simps)
-- "Other Caps"
\<comment> \<open>Other Caps\<close>
by (revokable'_hammer | fastforce simp: isCap_simps)+
end (* revokable_ccorres *)
@ -1096,7 +1096,7 @@ lemma cteInsert_ccorres:
apply vcg
apply wp
apply vcg
apply (simp add: Collect_const_mem split del: if_split) -- "Takes a while"
apply (simp add: Collect_const_mem split del: if_split) \<comment> \<open>Takes a while\<close>
apply (rule conjI)
apply (clarsimp simp: conj_comms cte_wp_at_ctes_of)
apply (intro conjI)
@ -1520,7 +1520,7 @@ lemma emptySlot_helper:
apply (case_tac "mdbNext rva \<noteq> 0")
apply (case_tac "mdbNext_CL (mdb_node_lift mdbNode) = 0", simp)
-- "case where mdbNext rva \<noteq> 0 and mdbNext_CL (mdb_node_lift mdbNode) \<noteq> 0"
\<comment> \<open>case where mdbNext rva \<noteq> 0 and mdbNext_CL (mdb_node_lift mdbNode) \<noteq> 0\<close>
apply (unfold updateMDB_def)
apply (clarsimp simp: Let_def)
apply (rule ccorres_pre_getCTE [where P = "\<lambda>cte s. ctes_of s (mdbNext rva) = Some cte" and P' = "\<lambda>_. UNIV"])
@ -1532,7 +1532,7 @@ lemma emptySlot_helper:
apply (frule(1) rf_sr_ctes_of_clift)
apply (clarsimp simp: typ_heap_simps' nextmdb_def if_1_0_0 nextcte_def)
apply (intro conjI impI allI)
-- "\<dots> \<exists>x\<in>fst \<dots>"
\<comment> \<open>\<dots> \<exists>x\<in>fst \<dots>\<close>
apply clarsimp
apply (rule fst_setCTE [OF ctes_of_cte_at], assumption )
apply (erule bexI [rotated])
@ -1563,7 +1563,7 @@ lemma emptySlot_helper:
apply (drule cteMDBNode_CL_lift [symmetric])
subgoal by (simp add: mdb_node_lift_def mask_def word_bw_assocs)
subgoal by (simp add: to_bool_def mask_def)
-- "\<dots> \<exists>x\<in>fst \<dots>"
\<comment> \<open>\<dots> \<exists>x\<in>fst \<dots>\<close>
apply clarsimp
apply (rule fst_setCTE [OF ctes_of_cte_at], assumption )
apply (erule bexI [rotated])
@ -1595,7 +1595,7 @@ lemma emptySlot_helper:
subgoal by (simp add: mdb_node_lift_def mask_def word_bw_assocs)
apply (simp add: to_bool_def mask_def split: if_split)
-- "trivial case where mdbNext rva = 0"
\<comment> \<open>trivial case where mdbNext rva = 0\<close>
apply (simp add:ccorres_cond_empty_iff)
apply (rule ccorres_guard_imp2)
apply (rule ccorres_return_Skip)
@ -2764,12 +2764,12 @@ lemma emptySlot_ccorres:
(Call emptySlot_'proc)"
apply (cinit lift: slot_' cleanupInfo_' simp: case_Null_If)
-- "--- handle the clearUntypedFreeIndex"
\<comment> \<open>--- handle the clearUntypedFreeIndex\<close>
apply (rule ccorres_split_noop_lhs, rule clearUntypedFreeIndex_noop_ccorres)
-- "--- instruction: newCTE \<leftarrow> getCTE slot; ---"
\<comment> \<open>--- instruction: newCTE \<leftarrow> getCTE slot; ---\<close>
apply (rule ccorres_pre_getCTE)
-- "--- instruction: CALL on C side"
\<comment> \<open>--- instruction: CALL on C side\<close>
apply (rule ccorres_move_c_guard_cte)
apply csymbr
apply (rule ccorres_abstract_cleanup)
@ -2778,21 +2778,21 @@ lemma emptySlot_ccorres:
= (cteCap newCTE = NullCap)" in ccorres_gen_asm2)
apply (simp del: Collect_const)
-- "--- instruction: if-then-else / IF-THEN-ELSE "
\<comment> \<open>--- instruction: if-then-else / IF-THEN-ELSE\<close>
apply (rule ccorres_cond2'[where R=\<top>])
-- "*** link between abstract and concrete conditionals ***"
\<comment> \<open>*** link between abstract and concrete conditionals ***\<close>
apply (clarsimp split: if_split)
-- "*** proof for the 'else' branch (return () and SKIP) ***"
\<comment> \<open>*** proof for the 'else' branch (return () and SKIP) ***\<close>
prefer 2
apply (ctac add: ccorres_return_Skip[unfolded dc_def])
-- "*** proof for the 'then' branch ***"
\<comment> \<open>*** proof for the 'then' branch ***\<close>
-- "---instructions: multiple on C side, including mdbNode fetch"
\<comment> \<open>---instructions: multiple on C side, including mdbNode fetch\<close>
apply (rule ccorres_rhs_assoc)+
-- "we have to do it here because the first assoc did not apply inside the then block"
\<comment> \<open>we have to do it here because the first assoc did not apply inside the then block\<close>
apply (rule ccorres_move_c_guard_cte | csymbr)+
apply (rule ccorres_symb_exec_r)
apply (rule_tac xf'="mdbNode_'" in ccorres_abstract, ceqv)
@ -2801,54 +2801,54 @@ lemma emptySlot_ccorres:
in ccorres_gen_asm2)
apply csymbr+
-- "--- instruction: updateMDB (mdbPrev rva) (mdbNext_update \<dots>) but with Ptr\<dots>\<noteq> NULL on C side"
apply (simp only:Ptr_not_null_pointer_not_zero) --"replaces Ptr p \<noteq> NULL with p\<noteq>0"
\<comment> \<open>--- instruction: updateMDB (mdbPrev rva) (mdbNext_update \<dots>) but with Ptr\<dots>\<noteq> NULL on C side\<close>
apply (simp only:Ptr_not_null_pointer_not_zero) \<comment> \<open>replaces Ptr p \<noteq> NULL with p\<noteq>0\<close>
-- "--- instruction: y \<leftarrow> updateMDB (mdbPrev rva) (mdbNext_update (\<lambda>_. mdbNext rva)) "
\<comment> \<open>--- instruction: y \<leftarrow> updateMDB (mdbPrev rva) (mdbNext_update (\<lambda>_. mdbNext rva))\<close>
apply (ctac (no_simp, no_vcg) pre:ccorres_move_guard_ptr_safe
add: updateMDB_mdbPrev_set_mdbNext)
-- "here ctac alone does not apply because the subgoal generated
by the rule are not solvable by simp"
-- "so we have to use (no_simp) (or apply (rule ccorres_split_nothrow))"
\<comment> \<open>here ctac alone does not apply because the subgoal generated
by the rule are not solvable by simp\<close>
\<comment> \<open>so we have to use (no_simp) (or apply (rule ccorres_split_nothrow))\<close>
apply (simp add: cmdbnode_relation_def)
apply assumption
-- "*** Main goal ***"
-- "--- instruction: updateMDB (mdbNext rva)
\<comment> \<open>*** Main goal ***\<close>
\<comment> \<open>--- instruction: updateMDB (mdbNext rva)
(\<lambda>mdb. mdbFirstBadged_update (\<lambda>_. mdbFirstBadged mdb \<or> mdbFirstBadged rva)
(mdbPrev_update (\<lambda>_. mdbPrev rva) mdb));"
apply (rule ccorres_rhs_assoc2 ) -- " to group the 2 first C instrutions together"
(mdbPrev_update (\<lambda>_. mdbPrev rva) mdb));\<close>
apply (rule ccorres_rhs_assoc2 ) \<comment> \<open>to group the 2 first C instrutions together\<close>
apply (ctac (no_vcg) add: emptySlot_helper)
-- "--- instruction: y \<leftarrow> updateCap slot capability.NullCap;"
\<comment> \<open>--- instruction: y \<leftarrow> updateCap slot capability.NullCap;\<close>
apply (simp del: Collect_const)
apply csymbr
apply (ctac (no_vcg) pre:ccorres_move_guard_ptr_safe)
apply csymbr
apply (rule ccorres_move_c_guard_cte)
-- "--- instruction y \<leftarrow> updateMDB slot (\<lambda>a. nullMDBNode);"
\<comment> \<open>--- instruction y \<leftarrow> updateMDB slot (\<lambda>a. nullMDBNode);\<close>
apply (ctac (no_vcg) pre: ccorres_move_guard_ptr_safe
add: ccorres_updateMDB_const [unfolded const_def])
-- "the post_cap_deletion case "
\<comment> \<open>the post_cap_deletion case\<close>
apply (ctac(no_vcg) add: postCapDeletion_ccorres [unfolded dc_def])
-- "Haskell pre/post for y \<leftarrow> updateMDB slot (\<lambda>a. nullMDBNode);"
\<comment> \<open>Haskell pre/post for y \<leftarrow> updateMDB slot (\<lambda>a. nullMDBNode);\<close>
apply wp
-- "C pre/post for y \<leftarrow> updateMDB slot (\<lambda>a. nullMDBNode);"
\<comment> \<open>C pre/post for y \<leftarrow> updateMDB slot (\<lambda>a. nullMDBNode);\<close>
apply simp
-- "C pre/post for the 2nd CALL "
-- "Haskell pre/post for y \<leftarrow> updateCap slot capability.NullCap;"
\<comment> \<open>C pre/post for the 2nd CALL\<close>
\<comment> \<open>Haskell pre/post for y \<leftarrow> updateCap slot capability.NullCap;\<close>
apply wp
-- "C pre/post for y \<leftarrow> updateCap slot capability.NullCap;"
\<comment> \<open>C pre/post for y \<leftarrow> updateCap slot capability.NullCap;\<close>
apply (simp add: Collect_const_mem cmdbnode_relation_def mdb_node_to_H_def nullMDBNode_def false_def)
-- "Haskell pre/post for the two nested updates "
\<comment> \<open>Haskell pre/post for the two nested updates\<close>
apply wp
-- "C pre/post for the two nested updates "
\<comment> \<open>C pre/post for the two nested updates\<close>
apply (simp add: Collect_const_mem ccap_relation_NullCap_iff)
-- "Haskell pre/post for (updateMDB (mdbPrev rva) (mdbNext_update (\<lambda>_. mdbNext rva)))"
\<comment> \<open>Haskell pre/post for (updateMDB (mdbPrev rva) (mdbNext_update (\<lambda>_. mdbNext rva)))\<close>
apply (simp, wp)
-- "C pre/post for (updateMDB (mdbPrev rva) (mdbNext_update (\<lambda>_. mdbNext rva)))"
\<comment> \<open>C pre/post for (updateMDB (mdbPrev rva) (mdbNext_update (\<lambda>_. mdbNext rva)))\<close>
apply simp+
apply vcg
apply (rule conseqPre, vcg)
@ -2856,15 +2856,15 @@ lemma emptySlot_ccorres:
apply simp
apply (wp hoare_vcg_all_lift hoare_vcg_imp_lift)
-- "final precondition proof"
\<comment> \<open>final precondition proof\<close>
apply (clarsimp simp: typ_heap_simps Collect_const_mem
cte_wp_at_ctes_of)
apply (rule conjI)
-- "Haskell side"
\<comment> \<open>Haskell side\<close>
apply (simp add: is_aligned_3_next canonical_address_mdbNext)
-- "C side"
\<comment> \<open>C side\<close>
apply (clarsimp simp: map_comp_Some_iff typ_heap_simps)
apply (subst cap_get_tag_isCap)
apply (rule ccte_relation_ccap_relation)
@ -2900,40 +2900,40 @@ lemma capSwapForDelete_ccorres:
(capSwapForDelete slot1 slot2)
(Call capSwapForDelete_'proc)"
apply (cinit lift: slot1_' slot2_' simp del: return_bind)
-- "***Main goal***"
-- "--- instruction: when (slot1 \<noteq> slot2) \<dots> / IF Ptr slot1 = Ptr slot2 THEN \<dots>"
\<comment> \<open>***Main goal***\<close>
\<comment> \<open>--- instruction: when (slot1 \<noteq> slot2) \<dots> / IF Ptr slot1 = Ptr slot2 THEN \<dots>\<close>
apply (simp add:when_def)
apply (rule ccorres_if_cond_throws2 [where Q = \<top> and Q' = \<top>])
apply (case_tac "slot1=slot2", simp+)
apply (rule ccorres_return_void_C [simplified dc_def])
-- "***Main goal***"
-- "--- ccorres goal with 2 affectations (cap1 and cap2) on both on Haskell and C "
-- "--- \<Longrightarrow> execute each part independently"
\<comment> \<open>***Main goal***\<close>
\<comment> \<open>--- ccorres goal with 2 affectations (cap1 and cap2) on both on Haskell and C\<close>
\<comment> \<open>--- \<Longrightarrow> execute each part independently\<close>
apply (simp add: liftM_def cong: call_ignore_cong)
apply (rule ccorres_pre_getCTE)+
apply (rule ccorres_move_c_guard_cte, rule ccorres_symb_exec_r)+
-- "***Main goal***"
\<comment> \<open>***Main goal***\<close>
apply (ctac (no_vcg) add: cteSwap_ccorres [unfolded dc_def] )
-- "C Hoare triple for \<acute>cap2 :== \<dots>"
\<comment> \<open>C Hoare triple for \<acute>cap2 :== \<dots>\<close>
apply vcg
-- "C existential Hoare triple for \<acute>cap2 :== \<dots>"
\<comment> \<open>C existential Hoare triple for \<acute>cap2 :== \<dots>\<close>
apply simp
apply (rule conseqPre)
apply vcg
apply simp
-- "C Hoare triple for \<acute>cap1 :== \<dots>"
\<comment> \<open>C Hoare triple for \<acute>cap1 :== \<dots>\<close>
apply vcg
-- "C existential Hoare triple for \<acute>cap1 :== \<dots>"
\<comment> \<open>C existential Hoare triple for \<acute>cap1 :== \<dots>\<close>
apply simp
apply (rule conseqPre)
apply vcg
apply simp
-- "Hoare triple for return_void"
\<comment> \<open>Hoare triple for return_void\<close>
apply vcg
-- "***Generalized preconditions***"
\<comment> \<open>***Generalized preconditions***\<close>
apply simp
apply (clarsimp simp: cte_wp_at_ctes_of map_comp_Some_iff
typ_heap_simps ccap_relation_def)
@ -3347,7 +3347,7 @@ lemma sameRegionAs_spec:
apply clarsimp
apply (simp add: sameRegionAs_def isArchCap_tag_def2 ccap_relation_c_valid_cap)
apply (case_tac capa, simp_all add: cap_get_tag_isCap_unfolded_H_cap isCap_simps)
-- "capa is a ThreadCap"
\<comment> \<open>capa is a ThreadCap\<close>
apply (case_tac capb, simp_all add: cap_get_tag_isCap_unfolded_H_cap
isCap_simps cap_tag_defs from_bool_def false_def)[1]
apply (frule_tac cap'=cap_a in cap_get_tag_isCap_unfolded_H_cap(1))
@ -3359,9 +3359,9 @@ lemma sameRegionAs_spec:
cong: if_cong)
apply (frule_tac cap'=cap_b in cap_get_tag_isArchCap_unfolded_H_cap)
apply (clarsimp simp: isArchCap_tag_def2)
-- "capa is a NullCap"
\<comment> \<open>capa is a NullCap\<close>
apply (simp add: cap_tag_defs from_bool_def false_def)
-- "capa is an NotificationCap"
\<comment> \<open>capa is an NotificationCap\<close>
apply (case_tac capb, simp_all add: cap_get_tag_isCap_unfolded_H_cap
isCap_simps cap_tag_defs from_bool_def false_def)[1]
apply (frule_tac cap'=cap_a in cap_get_tag_isCap_unfolded_H_cap(3))
@ -3372,7 +3372,7 @@ lemma sameRegionAs_spec:
apply (clarsimp split: if_split)
apply (frule_tac cap'=cap_b in cap_get_tag_isArchCap_unfolded_H_cap)
apply (clarsimp simp: isArchCap_tag_def2)
-- "capa is an IRQHandlerCap"
\<comment> \<open>capa is an IRQHandlerCap\<close>
apply (case_tac capb, simp_all add: cap_get_tag_isCap_unfolded_H_cap
isCap_simps cap_tag_defs from_bool_def false_def)[1]
apply (frule_tac cap'=cap_a in cap_get_tag_isCap_unfolded_H_cap(5))
@ -3387,7 +3387,7 @@ lemma sameRegionAs_spec:
| simp )
apply (frule_tac cap'=cap_b in cap_get_tag_isArchCap_unfolded_H_cap)
apply (clarsimp simp: isArchCap_tag_def2)
-- "capa is an EndpointCap"
\<comment> \<open>capa is an EndpointCap\<close>
apply (case_tac capb, simp_all add: cap_get_tag_isCap_unfolded_H_cap
isCap_simps cap_tag_defs from_bool_def false_def)[1]
apply (frule_tac cap'=cap_a in cap_get_tag_isCap_unfolded_H_cap(4))
@ -3398,23 +3398,23 @@ lemma sameRegionAs_spec:
apply (clarsimp split: if_split)
apply (frule_tac cap'=cap_b in cap_get_tag_isArchCap_unfolded_H_cap)
apply (clarsimp simp: isArchCap_tag_def2)
-- "capa is a DomainCap"
\<comment> \<open>capa is a DomainCap\<close>
apply (case_tac capb, simp_all add: cap_get_tag_isCap_unfolded_H_cap
isCap_simps cap_tag_defs from_bool_def false_def true_def)[1]
apply (frule_tac cap'=cap_b in cap_get_tag_isArchCap_unfolded_H_cap)
apply (fastforce simp: isArchCap_tag_def2 split: if_split)
-- "capa is a Zombie"
\<comment> \<open>capa is a Zombie\<close>
apply (simp add: cap_tag_defs from_bool_def false_def)
-- "capa is an Arch object cap"
\<comment> \<open>capa is an Arch object cap\<close>
apply (frule_tac cap'=cap_a in cap_get_tag_isArchCap_unfolded_H_cap)
apply (clarsimp simp: isArchCap_tag_def2 cap_tag_defs linorder_not_less [THEN sym])
apply (rule conjI, clarsimp, rule impI)+
apply (case_tac capb, simp_all add: cap_get_tag_isCap_unfolded_H_cap
isCap_simps cap_tag_defs from_bool_def false_def)[1]
-- "capb is an Arch object cap"
\<comment> \<open>capb is an Arch object cap\<close>
apply (frule_tac cap'=cap_b in cap_get_tag_isArchCap_unfolded_H_cap)
apply (fastforce simp: isArchCap_tag_def2 cap_tag_defs linorder_not_less [THEN sym])
-- "capa is a ReplyCap"
\<comment> \<open>capa is a ReplyCap\<close>
apply (case_tac capb, simp_all add: cap_get_tag_isCap_unfolded_H_cap
isCap_simps cap_tag_defs from_bool_def false_def)[1]
apply (frule_tac cap'=cap_b in cap_get_tag_isArchCap_unfolded_H_cap)
@ -3425,7 +3425,7 @@ lemma sameRegionAs_spec:
apply (simp add: cap_reply_cap_lift)
apply (simp add: cap_to_H_def ctcb_ptr_to_tcb_ptr_def)
apply (clarsimp split: if_split)
-- "capa is an UntypedCap"
\<comment> \<open>capa is an UntypedCap\<close>
apply (frule_tac cap'=cap_a in cap_get_tag_isCap_unfolded_H_cap(9))
apply (intro conjI)
apply (rule impI, intro conjI)
@ -3455,7 +3455,7 @@ lemma sameRegionAs_spec:
field_simps valid_cap'_def)+)[4]
apply (rule impI, simp add: from_bool_0 ccap_relation_get_capIsPhysical[symmetric])
apply (simp add: from_bool_def false_def)
-- "capa is a CNodeCap"
\<comment> \<open>capa is a CNodeCap\<close>
apply (case_tac capb, simp_all add: cap_get_tag_isCap_unfolded_H_cap
isCap_simps cap_tag_defs from_bool_def false_def)[1]
apply (frule_tac cap'=cap_b in cap_get_tag_isArchCap_unfolded_H_cap)
@ -3466,7 +3466,7 @@ lemma sameRegionAs_spec:
apply (simp add: cap_cnode_cap_lift)
apply (simp add: cap_to_H_def)
apply (clarsimp split: if_split bool.split)
-- "capa is an IRQControlCap"
\<comment> \<open>capa is an IRQControlCap\<close>
apply (case_tac capb, simp_all add: cap_get_tag_isCap_unfolded_H_cap
isCap_simps cap_tag_defs from_bool_def false_def true_def)[1]
apply (frule_tac cap'=cap_b in cap_get_tag_isArchCap_unfolded_H_cap)
@ -3590,21 +3590,21 @@ lemma sameObjectAs_spec:
isCap_simps cap_tag_defs
from_bool_def false_def)
apply fastforce+
-- "capa is an arch cap"
\<comment> \<open>capa is an arch cap\<close>
apply (frule cap_get_tag_isArchCap_unfolded_H_cap)
apply (simp add: isArchCap_tag_def2)
apply (rule conjI, rule impI, clarsimp, rule impI)+
apply (case_tac capb, simp_all add: cap_get_tag_isCap_unfolded_H_cap
isCap_simps cap_tag_defs)[1]
apply ((fastforce)+)[7]
-- "capb is an arch cap"
\<comment> \<open>capb is an arch cap\<close>
apply (frule_tac cap'=cap_b in cap_get_tag_isArchCap_unfolded_H_cap)
apply (fastforce simp: isArchCap_tag_def2 linorder_not_less [symmetric])+
-- "capa is an irq handler cap"
\<comment> \<open>capa is an irq handler cap\<close>
apply (case_tac capb, simp_all add: cap_get_tag_isCap_unfolded_H_cap
isCap_simps cap_tag_defs)
apply fastforce+
-- "capb is an arch cap"
\<comment> \<open>capb is an arch cap\<close>
apply (frule cap_get_tag_isArchCap_unfolded_H_cap)
apply (fastforce simp: isArchCap_tag_def2)+
done
@ -3683,24 +3683,24 @@ lemma isMDBParentOf_spec:
split: option.split_asm)
apply (rule conjI)
-- "sameRegionAs = 0"
\<comment> \<open>sameRegionAs = 0\<close>
apply (rule impI)
apply (clarsimp simp: from_bool_def false_def
split: if_split bool.splits)
-- "sameRegionAs \<noteq> 0"
\<comment> \<open>sameRegionAs \<noteq> 0\<close>
apply (clarsimp simp: from_bool_def false_def)
apply (clarsimp cong:bool.case_cong if_cong simp: typ_heap_simps)
apply (rule conjI)
--" cap_get_tag of cte_a is an endpoint"
\<comment> \<open>cap_get_tag of cte_a is an endpoint\<close>
apply clarsimp
apply (frule cap_get_tag_EndpointCap)
apply simp
apply (clarsimp simp: to_bool_def isNotificationCap_def isEndpointCap_def true_def) -- "badge of A is not 0 now"
apply (clarsimp simp: to_bool_def isNotificationCap_def isEndpointCap_def true_def) \<comment> \<open>badge of A is not 0 now\<close>
apply (subgoal_tac "cap_get_tag (cte_C.cap_C cte_b) = scast cap_endpoint_cap") --"needed also after"
apply (subgoal_tac "cap_get_tag (cte_C.cap_C cte_b) = scast cap_endpoint_cap") \<comment> \<open>needed also after\<close>
prefer 2
apply (rule sameRegionAs_EndpointCap, assumption+)
@ -3709,14 +3709,14 @@ lemma isMDBParentOf_spec:
apply (clarsimp split: if_split_asm simp: if_distrib [where f=scast])
apply (clarsimp, rule conjI)
--" cap_get_tag of cte_a is an notification"
\<comment> \<open>cap_get_tag of cte_a is an notification\<close>
apply clarsimp
apply (frule cap_get_tag_NotificationCap)
apply simp
apply (clarsimp simp: to_bool_def isNotificationCap_def isEndpointCap_def true_def) -- "badge of A is not 0 now"
apply (clarsimp simp: to_bool_def isNotificationCap_def isEndpointCap_def true_def) \<comment> \<open>badge of A is not 0 now\<close>
apply (subgoal_tac "cap_get_tag (cte_C.cap_C cte_b) = scast cap_notification_cap") --"needed also after"
apply (subgoal_tac "cap_get_tag (cte_C.cap_C cte_b) = scast cap_notification_cap") \<comment> \<open>needed also after\<close>
prefer 2
apply (rule sameRegionAs_NotificationCap, assumption+)
@ -3726,7 +3726,7 @@ lemma isMDBParentOf_spec:
apply (frule_tac cap="(cap_to_H x2c)" in cap_get_tag_NotificationCap)
apply clarsimp
-- " main goal"
\<comment> \<open>main goal\<close>
apply clarsimp
apply (simp add: to_bool_def)
apply (subgoal_tac "(\<not> (isEndpointCap (cap_to_H x2b))) \<and> ( \<not> (isNotificationCap (cap_to_H x2b)))")
@ -3746,21 +3746,21 @@ lemma updateCapData_spec:
apply (case_tac cap, simp_all add: cap_get_tag_isCap_unfolded_H_cap
isCap_simps from_bool_def isArchCap_tag_def2 cap_tag_defs Let_def)
-- "NotificationCap"
\<comment> \<open>NotificationCap\<close>
apply clarsimp
apply (frule cap_get_tag_isCap_unfolded_H_cap(3))
apply (frule (1) iffD1[OF cap_get_tag_NotificationCap])
apply clarsimp
apply (intro conjI impI)
-- "preserve is zero and capNtfnBadge_CL \<dots> = 0"
\<comment> \<open>preserve is zero and capNtfnBadge_CL \<dots> = 0\<close>
apply clarsimp
apply (clarsimp simp:cap_notification_cap_lift_def cap_lift_def cap_tag_defs)
apply (simp add: ccap_relation_def cap_lift_def cap_tag_defs cap_to_H_def)
-- "preserve is zero and capNtfnBadge_CL \<dots> \<noteq> 0"
\<comment> \<open>preserve is zero and capNtfnBadge_CL \<dots> \<noteq> 0\<close>
apply clarsimp
apply (simp add: ccap_relation_NullCap_iff cap_tag_defs)
-- "preserve is not zero"
\<comment> \<open>preserve is not zero\<close>
apply clarsimp
apply (simp add: to_bool_def)
apply (case_tac "preserve_' x = 0 \<and> capNtfnBadge_CL (cap_notification_cap_lift (cap_' x))= 0",
@ -3768,34 +3768,34 @@ lemma updateCapData_spec:
apply (simp add: if_not_P)
apply (simp add: ccap_relation_NullCap_iff cap_tag_defs)
-- "EndpointCap"
\<comment> \<open>EndpointCap\<close>
apply clarsimp
apply (frule cap_get_tag_isCap_unfolded_H_cap(4))
apply (frule (1) iffD1[OF cap_get_tag_EndpointCap])
apply clarsimp
apply (intro impI conjI)
-- "preserve is zero and capNtfnBadge_CL \<dots> = 0"
\<comment> \<open>preserve is zero and capNtfnBadge_CL \<dots> = 0\<close>
apply clarsimp
apply (clarsimp simp:cap_endpoint_cap_lift_def cap_lift_def cap_tag_defs)
apply (simp add: ccap_relation_def cap_lift_def cap_tag_defs cap_to_H_def)
-- "preserve is zero and capNtfnBadge_CL \<dots> \<noteq> 0"
\<comment> \<open>preserve is zero and capNtfnBadge_CL \<dots> \<noteq> 0\<close>
apply clarsimp
apply (simp add: ccap_relation_NullCap_iff cap_tag_defs)
-- "preserve is not zero"
\<comment> \<open>preserve is not zero\<close>
apply clarsimp
apply (simp add: to_bool_def)
apply (case_tac "preserve_' x = 0 \<and> capEPBadge_CL (cap_endpoint_cap_lift (cap_' x))= 0", clarsimp)
apply (simp add: if_not_P)
apply (simp add: ccap_relation_NullCap_iff cap_tag_defs)
-- "ArchObjectCap"
\<comment> \<open>ArchObjectCap\<close>
apply clarsimp
apply (frule cap_get_tag_isArchCap_unfolded_H_cap)
apply (simp add: isArchCap_tag_def2)
apply (simp add: X64_H.updateCapData_def)
-- "CNodeCap"
\<comment> \<open>CNodeCap\<close>
apply (clarsimp simp: cteRightsBits_def cteGuardBits_def)
apply (frule cap_get_tag_isCap_unfolded_H_cap(10))
apply (frule (1) iffD1[OF cap_get_tag_CNodeCap])
@ -3808,7 +3808,7 @@ lemma updateCapData_spec:
cong: if_cong)
apply (simp only: unat_word_ariths(1))
apply (rule ssubst [OF nat_mod_eq' [where n = "2 ^ len_of TYPE(64)"]])
-- " unat (\<dots> && 0x3F) + unat (\<dots> mod 0x40) < 2 ^ len_of TYPE(64)"
\<comment> \<open>unat (\<dots> && 0x3F) + unat (\<dots> mod 0x40) < 2 ^ len_of TYPE(64)\<close>
apply (rule order_le_less_trans, rule add_le_mono)
apply (rule word_le_nat_alt[THEN iffD1])
apply (rule word_and_le1)
@ -3854,9 +3854,9 @@ lemma ensureNoChildren_ccorres:
apply (clarsimp simp: whenE_def throwError_def return_def nullPointer_def liftE_bindE)
apply (clarsimp simp: returnOk_def return_def) -- "solve the case where mdbNext is zero"
apply (clarsimp simp: returnOk_def return_def) \<comment> \<open>solve the case where mdbNext is zero\<close>
-- "main goal"
\<comment> \<open>main goal\<close>
apply (simp add: ccte_relation_def)
apply (frule_tac cte="cte_to_H y" in valid_mdb_ctes_of_next, simp+)
apply (clarsimp simp: cte_wp_at_ctes_of)
@ -3870,7 +3870,7 @@ lemma ensureNoChildren_ccorres:
apply clarsimp
apply (rule conjI)
-- "isMDBParentOf is not zero"
\<comment> \<open>isMDBParentOf is not zero\<close>
apply clarsimp
apply (simp add: from_bool_def)
apply (case_tac "isMDBParentOf (cte_to_H y) (cte_to_H ya)", simp_all)[1]
@ -3881,7 +3881,7 @@ lemma ensureNoChildren_ccorres:
apply (simp add: cte_wp_at_ctes_of)
apply (simp add: syscall_error_rel_def EXCEPTION_NONE_def EXCEPTION_SYSCALL_ERROR_def)
apply (simp add: syscall_error_to_H_cases(9))
-- "isMDBParentOf is zero"
\<comment> \<open>isMDBParentOf is zero\<close>
apply clarsimp
apply (simp add: from_bool_def)
apply (case_tac "isMDBParentOf (cte_to_H y) (cte_to_H ya)", simp_all)[1]
@ -3891,7 +3891,7 @@ lemma ensureNoChildren_ccorres:
apply (simp add: cte_wp_at_ctes_of)
apply (simp add: returnOk_def return_def)
-- " last goal"
\<comment> \<open>last goal\<close>
apply clarsimp
apply (simp add: cte_wp_at_ctes_of)
done
@ -3960,7 +3960,7 @@ lemma Arch_deriveCap_ccorres:
cap_tag_defs cap_to_H_def to_bool_def
cap_page_directory_cap_lift_def
split: if_split_asm)
-- "PDPTCap"
\<comment> \<open>PDPTCap\<close>
apply wpc
apply (clarsimp simp: cap_get_tag_isCap_ArchObject
ccorres_cond_iffs)
@ -3990,7 +3990,7 @@ lemma Arch_deriveCap_ccorres:
errstate_def syscall_error_rel_def
syscall_error_to_H_cases
exception_defs)
-- "PML4Cap"
\<comment> \<open>PML4Cap\<close>
apply wpc
apply (clarsimp simp: cap_get_tag_isCap_ArchObject
ccorres_cond_iffs)
@ -4020,7 +4020,7 @@ lemma Arch_deriveCap_ccorres:
errstate_def syscall_error_rel_def
syscall_error_to_H_cases
exception_defs)
-- "PageCap"
\<comment> \<open>PageCap\<close>
apply wpc
apply (clarsimp simp: cap_get_tag_isCap_ArchObject
ccorres_cond_iffs)

View File

@ -215,7 +215,7 @@ proof (cases "isCNodeCap cap'")
show ?thesis using False
apply (cinit' lift: nodeCap_' capptr_' n_bits_')
apply csymbr+
-- "Exception stuff"
\<comment> \<open>Exception stuff\<close>
apply (rule ccorres_split_throws)
apply (simp add: Collect_const cap_get_tag_isCap isCap_simps ccorres_cond_iffs
resolveAddressBits.simps scast_id)
@ -237,8 +237,8 @@ next
from True show ?thesis
apply -
apply (cinit' simp add: whileAnno_def ucast_id)
-- "This is done here as init lift usually throws away the relationship between nodeCap_' s and nodeCap. Normally
this OK, but the induction messes with everything :("
\<comment> \<open>This is done here as init lift usually throws away the relationship between nodeCap_' s and nodeCap. Normally
this OK, but the induction messes with everything :(\<close>
apply (rule ccorres_abstract [where xf' = nodeCap_'])
apply ceqv
apply (rename_tac "nodeCap")
@ -267,11 +267,11 @@ next
apply clarsimp
apply (rule_tac I = "{s. cap_get_tag (nodeCap_' s) = scast cap_cnode_cap}"
in HoarePartial.While [unfolded whileAnno_def, OF subset_refl])
apply (vcg strip_guards=true) -- "takes a while"
apply (vcg strip_guards=true) \<comment> \<open>takes a while\<close>
apply clarsimp
apply simp
apply (clarsimp simp: cap_get_tag_isCap to_bool_def)
-- "Main thm"
\<comment> \<open>Main thm\<close>
proof (induct cap' cptr' guard' rule: resolveAddressBits.induct [case_names ind])
case (ind cap cptr guard)
@ -477,12 +477,12 @@ next
apply (csymbr | rule iffD2 [OF ccorres_seq_skip])+
apply (rule ccorres_Guard_Seq)+
apply csymbr
-- "handle the stateAssert in locateSlotCap very carefully"
\<comment> \<open>handle the stateAssert in locateSlotCap very carefully\<close>
apply (simp(no_asm) only: liftE_bindE[where a="locateSlotCap a b" for a b])
apply (rule ccorres_locateSlotCap_push[rotated])
apply (simp add: unlessE_def)
apply (rule hoare_pre, wp, simp)
-- "Convert guardBits, radixBits and capGuard to their Haskell versions"
\<comment> \<open>Convert guardBits, radixBits and capGuard to their Haskell versions\<close>
apply (drule (2) cgD, drule (2) rbD, drule (2) gbD)
apply (elim conjE)
apply (rule ccorres_gen_asm [where P = "guard \<le> 64"])
@ -551,9 +551,9 @@ next
apply (vcg strip_guards=true)
apply (vcg strip_guards=true)
apply (rule conjI)
-- "Haskell guard"
\<comment> \<open>Haskell guard\<close>
apply (thin_tac "unat n_bits = guard")
apply (clarsimp simp del: imp_disjL) -- "take a while"
apply (clarsimp simp del: imp_disjL) \<comment> \<open>take a while\<close>
apply (intro impI conjI allI)
apply fastforce
apply clarsimp
@ -563,7 +563,7 @@ next
apply (clarsimp simp: isCap_simps valid_cap_simps' cte_level_bits_def cteSizeBits_def
real_cte_at')
apply (clarsimp simp: isCap_simps valid_cap'_def)
-- "C guard"
\<comment> \<open>C guard\<close>
apply (frule (1) cgD [OF refl], frule (1) rbD [OF refl], frule (1) gbD [OF refl])
apply (simp add: Collect_const_mem cap_get_tag_isCap exception_defs lookup_fault_lifts
n_bits_guard mask7_eqs word_le_nat_alt word_less_nat_alt gm)

View File

@ -1763,7 +1763,7 @@ proof -
, assumption +,
simp_all add: objBits_simps' archObjSize_def projectKOs
bit_simps
heap_to_user_data_restrict heap_to_device_data_restrict)[1])+ -- "waiting ..."
heap_to_user_data_restrict heap_to_device_data_restrict)[1])+ \<comment> \<open>waiting ...\<close>
apply (simp add: map_to_ctes_delete' cmap_relation_restrict_both_proj
cmap_relation_restrict_both cmap_array_helper hrs_htd_update
bit_simps cmap_array)
@ -1810,7 +1810,7 @@ proof -
apply fastforce
apply ((subst lift_t_typ_region_bytes, rule cm_disj_tcb, assumption+,
simp_all add: objBits_simps archObjSize_def pageBits_def projectKOs)[1])+
-- "waiting ..."
\<comment> \<open>waiting ...\<close>
apply (simp add: tcb_queue_relation_live_restrict
[OF D.valid_untyped tat tlive rl])
done

View File

@ -667,7 +667,7 @@ lemma doUnbindNotification_ccorres:
cpspace_relation_def update_ntfn_map_tos)
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (rule cpspace_relation_ntfn_update_ntfn, assumption+)
apply (clarsimp simp: cnotification_relation_def Let_def
mask_def [where n=2] NtfnState_Waiting_def)
@ -717,7 +717,7 @@ lemma doUnbindNotification_ccorres':
cpspace_relation_def update_ntfn_map_tos)
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (rule cpspace_relation_ntfn_update_ntfn, assumption+)
apply (clarsimp simp: cnotification_relation_def Let_def
mask_def [where n=2] NtfnState_Waiting_def)
@ -916,7 +916,7 @@ lemma finaliseCap_True_cases_ccorres:
apply (clarsimp simp add: return_def ccap_relation_NullCap_iff)
apply (clarsimp simp add: irq_opt_relation_def)
apply vcg
-- "NullCap case by exhaustion"
\<comment> \<open>NullCap case by exhaustion\<close>
apply (simp add: cap_get_tag_isCap Let_def
ccorres_cond_empty_iff ccorres_cond_univ_iff)
apply (rule ccorres_rhs_assoc2, rule ccorres_rhs_assoc2,
@ -2214,7 +2214,7 @@ lemma Arch_finaliseCap_ccorres:
apply return_NullCap_pair_ccorres (* ASIDControlCap *)
apply return_NullCap_pair_ccorres (* IOPortCap *)
apply return_NullCap_pair_ccorres (* IOPortControlCap *)
-- "PageTableCap"
\<comment> \<open>PageTableCap\<close>
apply (subst ccorres_cond_seq2_seq[symmetric])
apply (rule ccorres_guard_imp)
apply (rule ccorres_rhs_assoc)
@ -2223,7 +2223,7 @@ lemma Arch_finaliseCap_ccorres:
apply ccorres_rewrite
apply (rule ccorres_inst[where P=\<top> and P'=UNIV])
apply (return_NullCap_pair_ccorres, simp+)
-- "PageDirectoryCap"
\<comment> \<open>PageDirectoryCap\<close>
apply (subst ccorres_cond_seq2_seq[symmetric])
apply (rule ccorres_guard_imp)
apply (rule ccorres_rhs_assoc)
@ -2232,7 +2232,7 @@ lemma Arch_finaliseCap_ccorres:
apply ccorres_rewrite
apply (rule ccorres_inst[where P=\<top> and P'=UNIV])
apply (return_NullCap_pair_ccorres, simp+)
-- "PDPointerTableCap"
\<comment> \<open>PDPointerTableCap\<close>
apply (rule ccorres_guard_imp)
apply (rule ccorres_add_return2)
apply (ctac (no_vcg) add: Mode_finaliseCap_ccorres_pdpt[where is_final=False, simplified if_False])
@ -2242,7 +2242,7 @@ lemma Arch_finaliseCap_ccorres:
apply wp
apply fastforce
apply (fastforce simp: false_def)
-- "PML4Cap"
\<comment> \<open>PML4Cap\<close>
apply (rule ccorres_guard_imp)
apply (rule ccorres_add_return2)
apply (ctac (no_vcg) add: Mode_finaliseCap_ccorres_pml4[where is_final=False, simplified if_False])
@ -2252,7 +2252,7 @@ lemma Arch_finaliseCap_ccorres:
apply wp
apply fastforce
apply (fastforce simp: false_def)
-- "PageCap"
\<comment> \<open>PageCap\<close>
apply (clarsimp simp: isCap_simps)
apply (rule ccorres_guard_imp[where A="?abstract_pre" and A'=UNIV])
apply ccorres_rewrite
@ -2267,7 +2267,7 @@ lemma Arch_finaliseCap_ccorres:
apply (fastforce simp: isCap_simps)
apply ccorres_rewrite
apply (rule ccorres_Cond_rhs_Seq)
-- "final \<and> PageDirectoryCap"
\<comment> \<open>final \<and> PageDirectoryCap\<close>
apply (clarsimp simp: isCap_simps)
apply (rule ccorres_rhs_assoc)+
apply csymbr
@ -2308,7 +2308,7 @@ lemma Arch_finaliseCap_ccorres:
split: if_split_asm)
apply simp
apply (rule ccorres_inst[where P=\<top> and P'=UNIV], return_NullCap_pair_ccorres)
-- "final \<and> PageTableCap"
\<comment> \<open>final \<and> PageTableCap\<close>
apply (rule ccorres_Cond_rhs_Seq)
apply (clarsimp simp: isCap_simps)
apply (rule ccorres_rhs_assoc)+
@ -2352,7 +2352,7 @@ lemma Arch_finaliseCap_ccorres:
split: if_split_asm)
apply simp
apply (rule ccorres_inst[where P=\<top> and P'=UNIV], return_NullCap_pair_ccorres)
-- "final \<and> ASIDPoolCap"
\<comment> \<open>final \<and> ASIDPoolCap\<close>
apply (rule ccorres_Cond_rhs_Seq)
apply (clarsimp simp: isCap_simps)
apply (rule ccorres_rhs_assoc)+
@ -2364,22 +2364,22 @@ lemma Arch_finaliseCap_ccorres:
apply (ctac (no_vcg) add: deleteASIDPool_ccorres)
apply (rule ccorres_inst[where P=\<top> and P'=UNIV], return_NullCap_pair_ccorres)
apply (rule wp_post_taut)
-- "final \<and> (ASIDControlCap \<or> IOPortControlCap"
\<comment> \<open>final \<and> (ASIDControlCap \<or> IOPortControlCap\<close>
apply (rule ccorres_Cond_rhs_Seq)
apply (clarsimp simp: isCap_simps)
apply (erule disjE; ccorres_rewrite; clarsimp; (rule ccorres_inst[where P=\<top> and P'=UNIV], return_NullCap_pair_ccorres))
-- "final \<and> IOPortCap"
\<comment> \<open>final \<and> IOPortCap\<close>
apply (rule ccorres_Cond_rhs_Seq)
apply (clarsimp simp: isCap_simps)
apply ccorres_rewrite
apply (rule ccorres_from_vcg_throws[where P=\<top> and P'=UNIV])
apply (rule allI, rule conseqPre, vcg)
apply (clarsimp simp: return_def ccap_relation_NullCap_iff)
-- "Mode_finaliseCap"
\<comment> \<open>Mode_finaliseCap\<close>
apply ccorres_rewrite
(* Winnow out the irrelevant cases *)
apply (wpc; (rule ccorres_inst[where P=\<top> and P'=UNIV], fastforce simp: isCap_simps)?)
-- "PageCap"
\<comment> \<open>PageCap\<close>
apply clarsimp
apply (rule ccorres_add_return2)
apply (ctac (no_vcg) add: Mode_finaliseCap_ccorres_page_cap)
@ -2387,14 +2387,14 @@ lemma Arch_finaliseCap_ccorres:
apply (rule allI, rule conseqPre, vcg)
apply (fastforce simp: return_def)
apply wp
-- "PDPointerTableCap"
\<comment> \<open>PDPointerTableCap\<close>
apply (rule ccorres_add_return2)
apply (ctac (no_vcg) add: Mode_finaliseCap_ccorres_pdpt[where is_final=True, simplified if_True])
apply (rule ccorres_from_vcg_throws[where P=\<top> and P'=UNIV])
apply (rule allI, rule conseqPre, vcg)
apply (fastforce simp: return_def)
apply wp
-- "PML4Cap"
\<comment> \<open>PML4Cap\<close>
apply (rule ccorres_add_return2)
apply (ctac (no_vcg) add: Mode_finaliseCap_ccorres_pml4[where is_final=True, simplified if_True])
apply (rule ccorres_from_vcg_throws[where P=\<top> and P'=UNIV])

View File

@ -434,13 +434,13 @@ lemma invokeCNodeSaveCaller_ccorres:
Collect_False Collect_True
del: Collect_const)[1]
apply (rule ccorres_fail)+
-- "NullCap case"
\<comment> \<open>NullCap case\<close>
apply (rule ccorres_from_vcg_throws [where P=\<top> and P'=UNIV])
apply (rule allI, rule conseqPre, vcg)
apply clarsimp
apply (simp add: return_def)
apply (rule ccorres_fail)+
-- "ReplyCap case"
\<comment> \<open>ReplyCap case\<close>
apply (rule ccorres_rhs_assoc)
apply csymbr
apply (frule cap_get_tag_isCap_unfolded_H_cap)
@ -726,7 +726,7 @@ lemma decodeCNodeInvocation_ccorres:
injection_handler_returnOk Collect_const[symmetric]
cong: call_ignore_cong del: Collect_const)
apply (rule ccorres_Cond_rhs_Seq)
-- "CNodeCopy case"
\<comment> \<open>CNodeCopy case\<close>
apply (simp add: Collect_const[symmetric] del: Collect_const)
apply (rule ccorres_rhs_assoc)+
apply (rule ccorres_Cond_rhs_Seq)
@ -796,7 +796,7 @@ lemma decodeCNodeInvocation_ccorres:
apply (simp add: Collect_const_mem)
apply (vcg exspec=getSyscallArg_modifies)
apply (rule ccorres_Cond_rhs_Seq)
-- "CNodeMint case"
\<comment> \<open>CNodeMint case\<close>
apply (simp add: Collect_const[symmetric]
del: Collect_const)
apply (rule ccorres_rhs_assoc)+
@ -883,7 +883,7 @@ lemma decodeCNodeInvocation_ccorres:
apply (simp add: Collect_const_mem)
apply (vcg exspec=getSyscallArg_modifies)
apply (rule ccorres_Cond_rhs_Seq)
-- "CNodeMove case"
\<comment> \<open>CNodeMove case\<close>
apply (simp add: Collect_const[symmetric] split_def
injection_handler_returnOk whenE_def
ccorres_invocationCatch_Inr
@ -911,7 +911,7 @@ lemma decodeCNodeInvocation_ccorres:
apply vcg
apply (rule conseqPre, vcg, clarsimp)
apply (rule ccorres_Cond_rhs_Seq)
-- "CNodeMutate case"
\<comment> \<open>CNodeMutate case\<close>
apply (rule ccorres_rhs_assoc)+
apply (simp add: Collect_const[symmetric] del: Collect_const
cong: call_ignore_cong)

View File

@ -237,7 +237,7 @@ lemma cancelSignal_ccorres_helper:
apply (clarsimp simp: typ_heap_simps cong: imp_cong split del: if_split)
apply (frule null_ep_queue [simplified Fun.comp_def])
apply (intro impI conjI allI)
-- "empty case"
\<comment> \<open>empty case\<close>
apply clarsimp
apply (frule iffD1 [OF tcb_queue_head_empty_iff [OF tcb_queue_relation'_queue_rel]])
apply (rule ballI, erule bspec)
@ -251,20 +251,20 @@ lemma cancelSignal_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cendpoint_relation_ntfn_queue [OF invs_sym'], assumption+)
apply simp
apply (erule (1) map_to_ko_atI')
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (rule cpspace_relation_ntfn_update_ntfn, assumption+)
apply (simp add: cnotification_relation_def Let_def NtfnState_Idle_def)
apply (simp add: carch_state_relation_def carch_globals_def)
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (clarsimp simp: carch_state_relation_def carch_globals_def
@ -276,7 +276,7 @@ lemma cancelSignal_ccorres_helper:
apply (simp add: objBits_simps')
apply (simp add: objBits_simps)
apply assumption
-- "non empty case"
\<comment> \<open>non empty case\<close>
apply (frule tcb_queue_head_empty_iff [OF tcb_queue_relation'_queue_rel])
apply (rule ballI, erule bspec)
apply (erule subsetD [rotated])
@ -289,10 +289,10 @@ lemma cancelSignal_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cendpoint_relation_ntfn_queue)
@ -300,7 +300,7 @@ lemma cancelSignal_ccorres_helper:
apply assumption+
apply simp
apply (erule (1) map_to_ko_atI')
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (rule cpspace_relation_ntfn_update_ntfn, assumption+)
apply (simp add: cnotification_relation_def Let_def isWaitingNtfn_def
split: ntfn.splits split del: if_split)
@ -314,7 +314,7 @@ lemma cancelSignal_ccorres_helper:
apply fastforce
apply simp
apply simp
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
subgoal by (clarsimp simp: carch_state_relation_def carch_globals_def
@ -753,7 +753,7 @@ lemma state_relation_queue_update_helper':
apply (clarsimp simp: inQ_def obj_at'_def)
apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def)
apply (intro conjI)
-- "cpspace_relation"
\<comment> \<open>cpspace_relation\<close>
apply (erule nonemptyE, drule(1) bspec)
apply (clarsimp simp: cpspace_relation_def)
apply (drule obj_at_ko_at', clarsimp)
@ -762,18 +762,18 @@ lemma state_relation_queue_update_helper':
apply (frule null_sched_queue)
apply (frule null_sched_epD)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (drule ctcb_relation_null_queue_ptrs,
simp_all)[1]
-- "endpoint relation"
\<comment> \<open>endpoint relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (erule cendpoint_relation_upd_tcb_no_queues, simp+)
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (erule cnotification_relation_upd_tcb_no_queues, simp+)
-- "ready queues"
\<comment> \<open>ready queues\<close>
apply (simp add: cready_queues_relation_def Let_def
cready_queues_index_to_C_in_range
seL4_MinPrio_def minDom_def)
@ -3077,7 +3077,7 @@ lemma cancelIPC_ccorres_helper:
apply (clarsimp simp: typ_heap_simps cong: imp_cong split del: if_split simp del: comp_def)
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (intro impI conjI allI)
-- "empty case"
\<comment> \<open>empty case\<close>
apply clarsimp
apply (frule iffD1 [OF tcb_queue_head_empty_iff [OF tcb_queue_relation'_queue_rel]])
apply (rule ballI, erule bspec)
@ -3090,21 +3090,21 @@ lemma cancelIPC_ccorres_helper:
cpspace_relation_def update_ep_map_tos typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
subgoal by (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (rule cpspace_relation_ep_update_ep, assumption+)
subgoal by (simp add: cendpoint_relation_def Let_def EPState_Idle_def)
subgoal by simp
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cnotification_relation_ep_queue [OF invs_sym'], assumption+)
subgoal by simp
apply (erule (1) map_to_ko_atI')
apply (simp add: heap_to_user_data_def Let_def)
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
subgoal by (clarsimp simp: comp_def)
subgoal by (clarsimp simp: carch_state_relation_def carch_globals_def
@ -3117,7 +3117,7 @@ lemma cancelIPC_ccorres_helper:
subgoal by (simp add: objBits_simps')
subgoal by (simp add: objBits_simps)
apply assumption
-- "non empty case"
\<comment> \<open>non empty case\<close>
apply clarsimp
apply (frule tcb_queue_head_empty_iff [OF tcb_queue_relation'_queue_rel])
apply (rule ballI, erule bspec)
@ -3130,13 +3130,13 @@ lemma cancelIPC_ccorres_helper:
cpspace_relation_def update_ep_map_tos typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
subgoal by (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (rule cpspace_relation_ep_update_ep, assumption+)
apply (simp add: cendpoint_relation_def Let_def isSendEP_def isRecvEP_def split: endpoint.splits split del: if_split)
-- "recv case"
\<comment> \<open>recv case\<close>
apply (subgoal_tac "pspace_canonical' \<sigma>")
prefer 2
apply fastforce
@ -3145,7 +3145,7 @@ lemma cancelIPC_ccorres_helper:
tcb_queue_relation'_next_sign tcb_queue_relation'_prev_sign
cong: tcb_queue_relation'_cong)
subgoal by (intro impI conjI; simp)
-- "send case"
\<comment> \<open>send case\<close>
apply (subgoal_tac "pspace_canonical' \<sigma>")
prefer 2
apply fastforce
@ -3155,13 +3155,13 @@ lemma cancelIPC_ccorres_helper:
cong: tcb_queue_relation'_cong)
subgoal by (intro impI conjI; simp)
subgoal by simp
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cnotification_relation_ep_queue [OF invs_sym'], assumption+)
apply simp
apply (erule (1) map_to_ko_atI')
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
subgoal by (clarsimp simp: comp_def)
subgoal by (clarsimp simp: carch_state_relation_def carch_globals_def
@ -3228,7 +3228,7 @@ lemma cancelIPC_ccorres1:
apply (rule_tac xf'=ret__unsigned_longlong_' in ccorres_abstract, ceqv)
apply (rule_tac P="rv' = thread_state_to_tsType rv" in ccorres_gen_asm2)
apply wpc
-- "BlockedOnReceive"
\<comment> \<open>BlockedOnReceive\<close>
apply (simp add: word_sle_def "StrictC'_thread_state_defs" ccorres_cond_iffs cong: call_ignore_cong)
apply (fold dc_def)
apply (rule ccorres_rhs_assoc)+
@ -3236,7 +3236,7 @@ lemma cancelIPC_ccorres1:
apply csymbr
apply (rule ccorres_pre_getEndpoint)
apply (rule ccorres_assert)
apply (rule ccorres_symb_exec_r) -- "ptr_get lemmas don't work so well :("
apply (rule ccorres_symb_exec_r) \<comment> \<open>ptr_get lemmas don't work so well :(\<close>
apply (rule ccorres_symb_exec_r)
apply (simp only: fun_app_def simp_list_case_return
return_bind ccorres_seq_skip)
@ -3255,7 +3255,7 @@ lemma cancelIPC_ccorres1:
apply (rule subset_refl)
apply (rule conseqPre, vcg)
apply clarsimp
-- "BlockedOnReply case"
\<comment> \<open>BlockedOnReply case\<close>
apply (simp add: "StrictC'_thread_state_defs" ccorres_cond_iffs
Collect_False Collect_True word_sle_def
cong: call_ignore_cong del: Collect_const)
@ -3318,7 +3318,7 @@ lemma cancelIPC_ccorres1:
apply (wp threadSet_invs_trivial | simp)+
apply (clarsimp simp add: guard_is_UNIV_def tcbReplySlot_def
Kernel_C.tcbReply_def tcbCNodeEntries_def)
-- "BlockedOnNotification"
\<comment> \<open>BlockedOnNotification\<close>
apply (simp add: word_sle_def "StrictC'_thread_state_defs" ccorres_cond_iffs dc_def [symmetric] cong: call_ignore_cong)
apply (rule ccorres_symb_exec_r)
apply (ctac (no_vcg))
@ -3327,18 +3327,18 @@ lemma cancelIPC_ccorres1:
apply (rule subset_refl)
apply (rule conseqPre, vcg)
apply clarsimp
-- "Running, Inactive, and Idle"
\<comment> \<open>Running, Inactive, and Idle\<close>
apply (simp add: word_sle_def "StrictC'_thread_state_defs" ccorres_cond_iffs dc_def [symmetric] cong: call_ignore_cong,
rule ccorres_return_Skip)+
-- "BlockedOnSend"
\<comment> \<open>BlockedOnSend\<close>
apply (simp add: word_sle_def "StrictC'_thread_state_defs" ccorres_cond_iffs dc_def [symmetric] cong: call_ignore_cong)
-- "clag"
\<comment> \<open>clag\<close>
apply (rule ccorres_rhs_assoc)+
apply csymbr
apply csymbr
apply (rule ccorres_pre_getEndpoint)
apply (rule ccorres_assert)
apply (rule ccorres_symb_exec_r) -- "ptr_get lemmas don't work so well :("
apply (rule ccorres_symb_exec_r) \<comment> \<open>ptr_get lemmas don't work so well :(\<close>
apply (rule ccorres_symb_exec_r)
apply (simp only: fun_app_def simp_list_case_return return_bind ccorres_seq_skip)
apply (rule ccorres_rhs_assoc2)
@ -3356,10 +3356,10 @@ lemma cancelIPC_ccorres1:
apply (rule conseqPre, vcg, rule subset_refl)
apply (rule conseqPre, vcg)
apply clarsimp
-- "Restart"
\<comment> \<open>Restart\<close>
apply (simp add: word_sle_def "StrictC'_thread_state_defs" ccorres_cond_iffs dc_def [symmetric] cong: call_ignore_cong,
rule ccorres_return_Skip)
-- "Post wp proofs"
\<comment> \<open>Post wp proofs\<close>
apply vcg
apply clarsimp
apply (rule conseqPre, vcg)

View File

@ -2538,7 +2538,7 @@ lemma getReceiveSlots_ccorres:
apply clarsimp
apply (simp add: cct_relation_def)
apply (case_tac rv, clarsimp)
apply (rule UNIV_I) -- "still a schematic here ..."
apply (rule UNIV_I) \<comment> \<open>still a schematic here ...\<close>
done
@ -2633,7 +2633,7 @@ lemma ccorres_if_cond_throws_break:
assumes abs: "\<forall>s s'. (s, s') \<in> sr \<and> Q s \<and> Q' s' \<longrightarrow> P = (s' \<in> P')"
and ac: "P \<Longrightarrow> ccorres_underlying sr \<Gamma> r xf arrel axf R R' (catchbrk_C # hs) a c"
and bd: "\<not> P \<Longrightarrow> ccorres_underlying sr \<Gamma> r xf arrel axf U U' (catchbrk_C # hs) b d"
and cthrows: "\<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> PT' c {}, UNIV" -- "c always throws"
and cthrows: "\<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> PT' c {}, UNIV" \<comment> \<open>c always throws\<close>
shows "ccorres_underlying sr \<Gamma> r xf arrel axf
(Q and (\<lambda>s. P \<longrightarrow> R s) and (\<lambda>s. \<not> P \<longrightarrow> U s))
(Collect Q' \<inter> {s. (s \<in> P' \<longrightarrow> s \<in> R' \<inter> PT') \<and> (s \<notin> P' \<longrightarrow> s \<in> U')})
@ -2681,7 +2681,7 @@ lemma ccorres_if_cond_throws_break2:
assumes abs: "\<forall>s s'. (s, s') \<in> sr \<and> Q s \<and> Q' s' \<longrightarrow> (\<not> P) = (s' \<in> P')"
and ac: "\<not> P \<Longrightarrow> ccorres_underlying sr \<Gamma> r xf arrel axf R R' (catchbrk_C # hs) a c"
and bd: "P \<Longrightarrow> ccorres_underlying sr \<Gamma> r xf arrel axf U U' (catchbrk_C # hs) b d"
and cthrows: "\<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> PT' c {}, UNIV" -- "c always throws"
and cthrows: "\<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> PT' c {}, UNIV" \<comment> \<open>c always throws\<close>
shows "ccorres_underlying sr \<Gamma> r xf arrel axf
(Q and (\<lambda>s. \<not> P \<longrightarrow> R s) and (\<lambda>s. P \<longrightarrow> U s))
(Collect Q' \<inter> {s. (s \<in> P' \<longrightarrow> s \<in> R' \<inter> PT') \<and> (s \<notin> P' \<longrightarrow> s \<in> U')})
@ -2708,7 +2708,7 @@ lemma ccorres_split_when_throwError_cond_break:
R R' (catchbrk_C # hs) (throwError e) c"
and bd: "\<not> P \<Longrightarrow> ccorres_underlying sr \<Gamma> r xf ar axf
U U' (catchbrk_C # hs) b d"
and cthrows: "\<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> PT' c {}, UNIV" -- "c always throws"
and cthrows: "\<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> PT' c {}, UNIV" \<comment> \<open>c always throws\<close>
shows "ccorres_underlying sr \<Gamma> r xf ar axf
(Q and (\<lambda>s. P \<longrightarrow> R s) and (\<lambda>s. \<not> P \<longrightarrow> U s))
(Collect Q' \<inter> {s. (s \<in> P' \<longrightarrow> s \<in> R' \<inter> PT') \<and> (s \<notin> P' \<longrightarrow> s \<in> U')})
@ -2993,7 +2993,7 @@ next
apply (clarsimp simp: Collect_const_mem)
apply (rule sym, rule from_bool_neq_0)
-- "case where a badge is sent"
\<comment> \<open>case where a badge is sent\<close>
apply (rule ccorres_rhs_assoc)+
apply csymbr
apply (simp only: Let_def liftE_bindE withoutFailure_def fun_app_def)
@ -3014,7 +3014,7 @@ next
apply (simp split del: if_split)
apply (vcg exspec=setExtraBadge_modifies)
-- "case where a cap is sent (or rather a send is attempted)"
\<comment> \<open>case where a cap is sent (or rather a send is attempted)\<close>
apply (simp add: split_def del: Collect_const
cong: call_ignore_cong split del: if_split)
apply (rule ccorres_rhs_assoc)+
@ -3031,7 +3031,7 @@ next
apply (rule ccorres_cond_false_seq)
apply (simp)
-- "case not diminish"
\<comment> \<open>case not diminish\<close>
apply (rule ccorres_split_nothrowE)
apply (rule unifyFailure_ccorres)
apply (ctac add: deriveCap_ccorres')
@ -3099,7 +3099,7 @@ next
apply (wp deriveCap_derived is_the_ep_deriveCap)
apply (vcg exspec=deriveCap_modifies)
--"remaining non ccorres subgoals"
\<comment> \<open>remaining non ccorres subgoals\<close>
apply (clarsimp simp: Collect_const_mem if_1_0_0
split del: if_split)
apply (rule_tac Q'="\<lbrace>\<acute>ret__int = from_bool (cap_get_tag cap' = scast cap_endpoint_cap
@ -4807,7 +4807,7 @@ lemma sendIPC_dequeue_ccorres_helper:
apply (fastforce simp: h_t_valid_clift)
apply (fastforce simp: h_t_valid_clift)
apply (fastforce simp: h_t_valid_clift)
-- "empty case"
\<comment> \<open>empty case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (frule iffD1 [OF tcb_queue_head_empty_iff
@ -4824,21 +4824,21 @@ lemma sendIPC_dequeue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (rule cpspace_relation_ep_update_ep, assumption+)
apply (simp add: cendpoint_relation_def Let_def EPState_Idle_def
tcb_queue_relation'_def)
apply simp
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cnotification_relation_ep_queue [OF invs_sym'], assumption+)
apply simp
apply (erule (1) map_to_ko_atI')
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (clarsimp simp: carch_state_relation_def packed_heap_update_collapse_hrs
@ -4851,7 +4851,7 @@ lemma sendIPC_dequeue_ccorres_helper:
apply (simp add: objBits_simps)
apply assumption
apply (clarsimp simp: cendpoint_relation_def Let_def tcb_queue_relation'_def)
-- "non-empty case"
\<comment> \<open>non-empty case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (frule tcb_queue_head_empty_iff [OF tcb_queue_relation'_queue_rel])
@ -4867,10 +4867,10 @@ lemma sendIPC_dequeue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (rule cpspace_relation_ep_update_ep, assumption+)
apply (clarsimp simp: cendpoint_relation_def Let_def
isRecvEP_def isSendEP_def
@ -4882,13 +4882,13 @@ lemma sendIPC_dequeue_ccorres_helper:
apply (simp add: objBits_simps')
apply (clarsimp split: if_split)
apply simp
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cnotification_relation_ep_queue [OF invs_sym'], assumption+)
apply simp
apply (erule (1) map_to_ko_atI')
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (clarsimp simp: carch_state_relation_def packed_heap_update_collapse_hrs
@ -5222,7 +5222,7 @@ lemma sendIPC_enqueue_ccorres_helper:
apply (fastforce simp: h_t_valid_clift)
apply (fastforce simp: h_t_valid_clift)
apply (case_tac ep, simp_all add: EPState_Idle_def EPState_Send_def)[1]
-- "IdleEP case"
\<comment> \<open>IdleEP case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (simp add: setEndpoint_def split_def)
@ -5232,10 +5232,10 @@ lemma sendIPC_enqueue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (rule cpspace_relation_ep_update_ep', assumption+)
apply (clarsimp simp: cendpoint_relation_def Let_def
mask_def [where n=3] EPState_Send_def)
@ -5246,7 +5246,7 @@ lemma sendIPC_enqueue_ccorres_helper:
intro!: tcb_and_not_mask_canonical
dest!: st_tcb_strg'[rule_format])
apply (simp add: isSendEP_def isRecvEP_def)
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cnotification_relation_ep_queue, assumption+)
@ -5257,7 +5257,7 @@ lemma sendIPC_enqueue_ccorres_helper:
apply (simp only:projectKOs injectKO_ep objBits_simps)
apply clarsimp
apply (clarsimp simp: obj_at'_def projectKOs)
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (clarsimp simp: carch_state_relation_def packed_heap_update_collapse_hrs
@ -5269,7 +5269,7 @@ lemma sendIPC_enqueue_ccorres_helper:
apply (simp add: objBits_simps')
apply (simp add: objBits_simps)
apply assumption
-- "SendEP case"
\<comment> \<open>SendEP case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (simp add: setEndpoint_def split_def)
@ -5279,10 +5279,10 @@ lemma sendIPC_enqueue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (rule cpspace_relation_ep_update_ep', assumption+)
apply (clarsimp simp: cendpoint_relation_def Let_def
mask_def [where n=3] EPState_Send_def
@ -5301,7 +5301,7 @@ lemma sendIPC_enqueue_ccorres_helper:
intro!: tcb_and_not_mask_canonical
dest!: st_tcb_strg'[rule_format])
apply (simp add: isSendEP_def isRecvEP_def)
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cnotification_relation_ep_queue, assumption+)
@ -5311,7 +5311,7 @@ lemma sendIPC_enqueue_ccorres_helper:
apply (erule(2) map_to_ko_at_updI')
apply (clarsimp simp: objBitsKO_def)
apply (clarsimp simp: obj_at'_def projectKOs)
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (clarsimp simp: carch_state_relation_def packed_heap_update_collapse_hrs
@ -5362,7 +5362,7 @@ lemma sendIPC_ccorres [corres]:
and (\<lambda>s. \<forall>d p. thread \<notin> set (ksReadyQueues s (d, p)))"
in ccorres_guard_imp2 [where A'=UNIV])
apply wpc
-- "RecvEP case"
\<comment> \<open>RecvEP case\<close>
apply (rule ccorres_cond_false)
apply (rule ccorres_cond_true)
apply (intro ccorres_rhs_assoc)
@ -5475,15 +5475,15 @@ lemma sendIPC_ccorres [corres]:
option_to_ptr_def option_to_0_def
split: bool.split_asm)
-- "IdleEP case"
\<comment> \<open>IdleEP case\<close>
apply (rule ccorres_cond_true)
apply (rule ccorres_cond)
apply (clarsimp simp: from_bool_def split: bool.split)
-- "blocking case"
\<comment> \<open>blocking case\<close>
apply (intro ccorres_rhs_assoc)
apply csymbr
apply (simp only:)
-- "apply (ctac (trace, no_vcg,c_lines 6) add: sendIPC_block_ccorres_helper)"
\<comment> \<open>apply (ctac (trace, no_vcg,c_lines 6) add: sendIPC_block_ccorres_helper)\<close>
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
@ -5503,14 +5503,14 @@ lemma sendIPC_ccorres [corres]:
apply wp
apply (clarsimp simp: guard_is_UNIV_def)
apply (rule ccorres_return_Skip)
-- "SendEP case"
\<comment> \<open>SendEP case\<close>
apply (rule ccorres_cond_true)
apply (rule ccorres_cond)
apply (clarsimp simp: from_bool_def split: bool.split)
-- "blocking case"
\<comment> \<open>blocking case\<close>
apply (intro ccorres_rhs_assoc)
apply csymbr
-- "apply (ctac (no_vcg,c_lines 6) add: sendIPC_block_ccorres_helper)"
\<comment> \<open>apply (ctac (no_vcg,c_lines 6) add: sendIPC_block_ccorres_helper)\<close>
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
apply (rule ccorres_rhs_assoc2)
@ -5681,7 +5681,7 @@ lemma receiveIPC_enqueue_ccorres_helper:
apply (fastforce simp: h_t_valid_clift)
apply (fastforce simp: h_t_valid_clift)
apply (case_tac ep, simp_all add: EPState_Idle_def EPState_Recv_def)[1]
-- "RecvEP case"
\<comment> \<open>RecvEP case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (simp add: setEndpoint_def split_def)
@ -5691,10 +5691,10 @@ lemma receiveIPC_enqueue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (rule cpspace_relation_ep_update_ep', assumption+)
apply (clarsimp simp: cendpoint_relation_def Let_def
mask_def [where n=3] EPState_Recv_def
@ -5713,7 +5713,7 @@ lemma receiveIPC_enqueue_ccorres_helper:
intro!: tcb_and_not_mask_canonical
dest!: st_tcb_strg'[rule_format])
apply (simp add: isSendEP_def isRecvEP_def)
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cnotification_relation_ep_queue, assumption+)
@ -5723,7 +5723,7 @@ lemma receiveIPC_enqueue_ccorres_helper:
apply (erule(2) map_to_ko_at_updI')
apply (clarsimp simp: objBitsKO_def)
apply (clarsimp simp: obj_at'_def projectKOs)
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (clarsimp simp: carch_state_relation_def packed_heap_update_collapse_hrs
@ -5735,7 +5735,7 @@ lemma receiveIPC_enqueue_ccorres_helper:
apply (simp add: objBits_simps')
apply (simp add: objBits_simps)
apply assumption
-- "IdleEP case"
\<comment> \<open>IdleEP case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (simp add: setEndpoint_def split_def)
@ -5745,10 +5745,10 @@ lemma receiveIPC_enqueue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (rule cpspace_relation_ep_update_ep', assumption+)
apply (clarsimp simp: cendpoint_relation_def Let_def
mask_def [where n=3] EPState_Recv_def)
@ -5759,7 +5759,7 @@ lemma receiveIPC_enqueue_ccorres_helper:
intro!: tcb_and_not_mask_canonical
dest!: st_tcb_strg'[rule_format])
apply (simp add: isSendEP_def isRecvEP_def)
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cnotification_relation_ep_queue, assumption+)
@ -5769,7 +5769,7 @@ lemma receiveIPC_enqueue_ccorres_helper:
apply (erule(2) map_to_ko_at_updI')
apply (clarsimp simp: objBitsKO_def)
apply (clarsimp simp: obj_at'_def projectKOs)
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (clarsimp simp: carch_state_relation_def packed_heap_update_collapse_hrs
@ -5826,7 +5826,7 @@ lemma receiveIPC_dequeue_ccorres_helper:
apply (fastforce simp: h_t_valid_clift)
apply (fastforce simp: h_t_valid_clift)
apply (fastforce simp: h_t_valid_clift)
-- "empty case"
\<comment> \<open>empty case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (frule iffD1 [OF tcb_queue_head_empty_iff
@ -5843,21 +5843,21 @@ lemma receiveIPC_dequeue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (rule cpspace_relation_ep_update_ep, assumption+)
apply (simp add: cendpoint_relation_def Let_def EPState_Idle_def
tcb_queue_relation'_def)
apply simp
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cnotification_relation_ep_queue [OF invs_sym'], assumption+)
apply simp
apply (erule (1) map_to_ko_atI')
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (clarsimp simp: carch_state_relation_def packed_heap_update_collapse_hrs
@ -5870,7 +5870,7 @@ lemma receiveIPC_dequeue_ccorres_helper:
apply (simp add: objBits_simps)
apply assumption
apply (clarsimp simp: cendpoint_relation_def Let_def tcb_queue_relation'_def)
-- "non-empty case"
\<comment> \<open>non-empty case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (frule tcb_queue_head_empty_iff [OF tcb_queue_relation'_queue_rel])
@ -5886,10 +5886,10 @@ lemma receiveIPC_dequeue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (rule cpspace_relation_ep_update_ep, assumption+)
apply (clarsimp simp: cendpoint_relation_def Let_def
isRecvEP_def isSendEP_def
@ -5901,13 +5901,13 @@ lemma receiveIPC_dequeue_ccorres_helper:
apply (clarsimp simp: objBits_simps')
apply (clarsimp split: if_split)
apply simp
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cnotification_relation_ep_queue [OF invs_sym'], assumption+)
apply simp
apply (erule (1) map_to_ko_atI')
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (clarsimp simp: carch_state_relation_def packed_heap_update_collapse_hrs
@ -5960,12 +5960,12 @@ lemma completeSignal_ccorres:
split: Structures_H.ntfn.splits)
apply ceqv
apply wpc
-- "IdleNtfn case"
\<comment> \<open>IdleNtfn case\<close>
apply (clarsimp simp: NtfnState_Idle_def NtfnState_Active_def)
apply csymbr
apply (rule ccorres_cond_false)
apply (rule ccorres_fail)
-- "ActiveNtfn case"
\<comment> \<open>ActiveNtfn case\<close>
apply (clarsimp, csymbr, rule ccorres_cond_true)
apply (rule ccorres_rhs_assoc)+
apply (rename_tac word)
@ -6002,7 +6002,7 @@ lemma completeSignal_ccorres:
apply (clarsimp simp: guard_is_UNIV_def X64_H.badgeRegister_def
X64.badgeRegister_def Kernel_C.badgeRegister_def
X64.capRegister_def Kernel_C.RDI_def)
-- "WaitingNtfn case"
\<comment> \<open>WaitingNtfn case\<close>
apply (clarsimp simp: NtfnState_Active_def NtfnState_Waiting_def)
apply csymbr
apply (rule ccorres_cond_false)
@ -6106,7 +6106,7 @@ lemma receiveIPC_ccorres [corres]:
and ko_at' ep (capEPPtr cap)"
in ccorres_guard_imp2 [where A'=UNIV])
apply wpc
-- "RecvEP case"
\<comment> \<open>RecvEP case\<close>
apply (rule ccorres_cond_true)
apply csymbr
apply (simp only: case_bool_If from_bool_neq_0)
@ -6135,7 +6135,7 @@ lemma receiveIPC_ccorres [corres]:
apply (clarsimp simp: guard_is_UNIV_def)
apply simp
apply (ctac add: doNBRecvFailedTransfer_ccorres[unfolded dc_def])
-- "IdleEP case"
\<comment> \<open>IdleEP case\<close>
apply (rule ccorres_cond_true)
apply csymbr
apply (simp only: case_bool_If from_bool_neq_0)
@ -6161,7 +6161,7 @@ lemma receiveIPC_ccorres [corres]:
apply (clarsimp simp: guard_is_UNIV_def)
apply simp
apply (ctac add: doNBRecvFailedTransfer_ccorres[unfolded dc_def])
-- "SendEP case"
\<comment> \<open>SendEP case\<close>
apply (thin_tac "isBlockinga = from_bool P" for P)
apply (rule ccorres_cond_false)
apply (rule ccorres_cond_true)
@ -6412,7 +6412,7 @@ lemma sendSignal_dequeue_ccorres_helper:
apply (intro conjI impI allI)
apply (fastforce simp: h_t_valid_clift)
apply (fastforce simp: h_t_valid_clift)
-- "empty case"
\<comment> \<open>empty case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (frule iffD1 [OF tcb_queue_head_empty_iff
@ -6430,21 +6430,21 @@ lemma sendSignal_dequeue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cendpoint_relation_ntfn_queue [OF invs_sym'], assumption+)
apply simp+
apply (erule (1) map_to_ko_atI')
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (rule cpspace_relation_ntfn_update_ntfn, assumption+)
apply (simp add: cnotification_relation_def Let_def NtfnState_Idle_def
tcb_queue_relation'_def)
apply simp
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (clarsimp simp: carch_state_relation_def packed_heap_update_collapse_hrs
@ -6458,7 +6458,7 @@ lemma sendSignal_dequeue_ccorres_helper:
apply assumption
apply (clarsimp simp: cnotification_relation_def Let_def
tcb_queue_relation'_def)
-- "non-empty case"
\<comment> \<open>non-empty case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (frule tcb_queue_head_empty_iff [OF tcb_queue_relation'_queue_rel])
@ -6475,16 +6475,16 @@ lemma sendSignal_dequeue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cendpoint_relation_ntfn_queue [OF invs_sym'], assumption+)
apply simp+
apply (erule (1) map_to_ko_atI')
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (rule cpspace_relation_ntfn_update_ntfn, assumption+)
apply (clarsimp simp: cnotification_relation_def Let_def
isWaitingNtfn_def
@ -6499,7 +6499,7 @@ lemma sendSignal_dequeue_ccorres_helper:
apply clarsimp
apply (clarsimp split: if_split)
apply simp
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (clarsimp simp: carch_state_relation_def
@ -6579,7 +6579,7 @@ lemma sendSignal_ccorres [corres]:
option_to_ctcb_ptr (ntfnBoundTCB ntfn) \<noteq> NULL"
in ccorres_gen_asm)
apply wpc
-- "IdleNtfn case"
\<comment> \<open>IdleNtfn case\<close>
apply (rule ccorres_cond_true)
apply (rule ccorres_rhs_assoc)+
apply (rule_tac xf'=ret__unsigned_longlong_'
@ -6628,7 +6628,7 @@ lemma sendSignal_ccorres [corres]:
Collect_const_mem)
apply (case_tac ts, simp_all add: receiveBlocked_def typ_heap_simps
cthread_state_relation_def "StrictC'_thread_state_defs")[1]
-- "ActiveNtfn case"
\<comment> \<open>ActiveNtfn case\<close>
apply (rename_tac old_badge)
apply (rule ccorres_cond_false)
apply (rule ccorres_cond_false)
@ -6660,7 +6660,7 @@ lemma sendSignal_ccorres [corres]:
apply (simp add: objBits_simps')
apply (simp add: objBits_simps)
apply assumption
-- "WaitingNtfn case"
\<comment> \<open>WaitingNtfn case\<close>
apply (rule ccorres_cond_false)
apply (rule ccorres_cond_true)
apply wpc
@ -6857,7 +6857,7 @@ lemma receiveSignal_enqueue_ccorres_helper:
apply (fastforce simp: h_t_valid_clift)
apply (fastforce simp: h_t_valid_clift)
apply (case_tac "ntfnObj ntfn", simp_all add: NtfnState_Idle_def NtfnState_Waiting_def)[1]
-- "IdleNtfn case"
\<comment> \<open>IdleNtfn case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (simp add: setNotification_def split_def)
@ -6867,10 +6867,10 @@ lemma receiveSignal_enqueue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cendpoint_relation_ntfn_queue, assumption+)
@ -6880,7 +6880,7 @@ lemma receiveSignal_enqueue_ccorres_helper:
apply (erule(2) map_to_ko_at_updI')
apply (clarsimp simp: objBitsKO_def)
apply (clarsimp simp: obj_at'_def projectKOs)
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (rule cpspace_relation_ntfn_update_ntfn', assumption+)
apply (case_tac "ntfn", simp_all)[1]
apply (clarsimp simp: cnotification_relation_def Let_def
@ -6895,7 +6895,7 @@ lemma receiveSignal_enqueue_ccorres_helper:
dest!: st_tcb_strg'[rule_format])
by clarsimp
apply (simp add: isWaitingNtfn_def)
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
subgoal by (clarsimp simp: comp_def)
apply (clarsimp simp: carch_state_relation_def packed_heap_update_collapse_hrs
@ -6907,7 +6907,7 @@ lemma receiveSignal_enqueue_ccorres_helper:
apply (simp add: objBits_simps')
apply (simp add: objBits_simps)
apply assumption
-- "WaitingNtfn case"
\<comment> \<open>WaitingNtfn case\<close>
apply clarsimp
apply (frule null_ep_queue [simplified comp_def] null_ep_queue)
apply (simp add: setNotification_def split_def)
@ -6917,10 +6917,10 @@ lemma receiveSignal_enqueue_ccorres_helper:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (erule ctcb_relation_null_queue_ptrs)
apply (clarsimp simp: comp_def)
-- "ep relation"
\<comment> \<open>ep relation\<close>
apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply simp
apply (rule cendpoint_relation_ntfn_queue, assumption+)
@ -6930,7 +6930,7 @@ lemma receiveSignal_enqueue_ccorres_helper:
apply (erule(2) map_to_ko_at_updI')
apply (clarsimp simp: objBitsKO_def)
apply (clarsimp simp: obj_at'_def projectKOs)
-- "ntfn relation"
\<comment> \<open>ntfn relation\<close>
apply (rule cpspace_relation_ntfn_update_ntfn', assumption+)
apply (case_tac "ntfn", simp_all)[1]
apply (clarsimp simp: cnotification_relation_def Let_def
@ -6956,7 +6956,7 @@ lemma receiveSignal_enqueue_ccorres_helper:
sign_extended_sign_extend)
done
apply (simp add: isWaitingNtfn_def)
-- "queue relation"
\<comment> \<open>queue relation\<close>
apply (rule cready_queues_relation_null_queue_ptrs, assumption+)
apply (clarsimp simp: comp_def)
apply (clarsimp simp: carch_state_relation_def packed_heap_update_collapse_hrs
@ -7007,7 +7007,7 @@ lemma receiveSignal_ccorres [corres]:
apply ceqv
apply wpc
-- "IdleNtfn case"
\<comment> \<open>IdleNtfn case\<close>
apply (rule ccorres_cond_true)
apply csymbr
apply (simp only: case_bool_If from_bool_neq_0)
@ -7032,7 +7032,7 @@ lemma receiveSignal_ccorres [corres]:
apply (clarsimp simp: guard_is_UNIV_def)
apply simp
apply (ctac add: doNBRecvFailedTransfer_ccorres[unfolded dc_def])
-- "ActiveNtfn case"
\<comment> \<open>ActiveNtfn case\<close>
apply (rename_tac badge)
apply (rule ccorres_cond_false)
apply (rule ccorres_cond_true)
@ -7074,7 +7074,7 @@ lemma receiveSignal_ccorres [corres]:
apply wp
apply (clarsimp simp: guard_is_UNIV_def)
apply (clarsimp simp: X64.badgeRegister_def X64.capRegister_def Kernel_C.RDI_def)
-- "WaitingNtfn case"
\<comment> \<open>WaitingNtfn case\<close>
apply (rename_tac list)
apply (rule ccorres_cond_true)
apply csymbr

View File

@ -246,7 +246,7 @@ lemma handleSyscall_ccorres:
apply (rule ccorres_split_nothrow_novcg)
apply wpc
prefer 3
-- "SysSend"
\<comment> \<open>SysSend\<close>
apply (clarsimp simp: syscall_from_H_def syscall_defs)
apply (rule ccorres_cond_empty |rule ccorres_cond_univ)+
apply (simp add: handleSend_def)
@ -285,7 +285,7 @@ lemma handleSyscall_ccorres:
apply clarsimp
apply (vcg exspec=handleInvocation_modifies)
prefer 3
-- "SysNBSend"
\<comment> \<open>SysNBSend\<close>
apply (clarsimp simp: syscall_from_H_def syscall_defs)
apply (rule ccorres_cond_empty |rule ccorres_cond_univ)+
apply (simp add: handleSend_def)
@ -322,7 +322,7 @@ lemma handleSyscall_ccorres:
apply (simp add: invs'_def valid_state'_def)
apply clarsimp
apply (vcg exspec=handleInvocation_modifies)
-- "SysCall"
\<comment> \<open>SysCall\<close>
apply (clarsimp simp: syscall_from_H_def syscall_defs)
apply (rule ccorres_cond_empty |rule ccorres_cond_univ)+
apply (simp add: handleCall_def)
@ -362,7 +362,7 @@ lemma handleSyscall_ccorres:
apply clarsimp
apply (vcg exspec=handleInvocation_modifies)
prefer 2
-- "SysRecv"
\<comment> \<open>SysRecv\<close>
apply (clarsimp simp: syscall_from_H_def syscall_defs)
apply (rule ccorres_cond_empty |rule ccorres_cond_univ)+
apply (simp add: liftE_bind)
@ -371,7 +371,7 @@ lemma handleSyscall_ccorres:
apply (rule ccorres_returnOk_skip[unfolded returnOk_def, simplified])
apply wp
prefer 2
-- "SysReply"
\<comment> \<open>SysReply\<close>
apply (clarsimp simp: syscall_from_H_def syscall_defs)
apply (rule ccorres_cond_empty |rule ccorres_cond_univ)+
apply (simp add: liftE_bind)
@ -379,7 +379,7 @@ lemma handleSyscall_ccorres:
apply (ctac (no_vcg) add: handleReply_ccorres)
apply (rule ccorres_returnOk_skip[unfolded returnOk_def, simplified])
apply wp
-- "SysReplyRecv"
\<comment> \<open>SysReplyRecv\<close>
apply (clarsimp simp: syscall_from_H_def syscall_defs)
apply (rule ccorres_cond_empty |rule ccorres_cond_univ)+
apply (simp add: liftE_bind bind_assoc)
@ -395,7 +395,7 @@ lemma handleSyscall_ccorres:
in hoare_post_imp)
apply (simp add: ct_in_state'_def)
apply (wp handleReply_sane handleReply_ct_not_ksQ)
-- "SysYield"
\<comment> \<open>SysYield\<close>
apply (clarsimp simp: syscall_from_H_def syscall_defs)
apply (rule ccorres_cond_empty |rule ccorres_cond_univ)+
apply (simp add: liftE_bind)
@ -403,7 +403,7 @@ lemma handleSyscall_ccorres:
apply (ctac (no_vcg) add: handleYield_ccorres)
apply (rule ccorres_returnOk_skip[unfolded returnOk_def, simplified])
apply wp
-- "SysNBRecv"
\<comment> \<open>SysNBRecv\<close>
apply (clarsimp simp: syscall_from_H_def syscall_defs)
apply (rule ccorres_cond_empty |rule ccorres_cond_univ)+
apply (simp add: liftE_bind)
@ -411,7 +411,7 @@ lemma handleSyscall_ccorres:
apply (ctac (no_vcg) add: handleRecv_ccorres)
apply (rule ccorres_returnOk_skip[unfolded returnOk_def, simplified])
apply wp
-- " rest of body"
\<comment> \<open>rest of body\<close>
apply ceqv
apply (ctac (no_vcg) add: schedule_ccorres)
apply (rule ccorres_add_return2)

View File

@ -3959,7 +3959,7 @@ proof -
note hacky_cte = retype_ctes_helper [where sz = "objBitsKO kotcb" and ko = kotcb and ptr = "ctcb_ptr_to_tcb_ptr p",
OF pal pds pno al _ _ mko, simplified new_cap_addrs_def, simplified]
-- "Ugh"
\<comment> \<open>Ugh\<close>
moreover have
"\<And>y. y \<in> ptr_val ` (CTypesDefs.ptr_add (cte_Ptr (ctcb_ptr_to_tcb_ptr p)) \<circ> of_nat) ` {k. k < 5}
= (y && ~~ mask tcbBlockSizeBits = ctcb_ptr_to_tcb_ptr p \<and> y && mask tcbBlockSizeBits \<in> dom tcb_cte_cases)" (is "\<And>y. ?LHS y = ?RHS y")
@ -4203,15 +4203,15 @@ proof -
apply (simp add: rl kotcb_def projectKOs rl_tcb rl_cte)
apply (elim conjE)
apply (intro conjI)
-- "cte"
\<comment> \<open>cte\<close>
apply (erule cmap_relation_retype2)
apply (simp add:ccte_relation_nullCap nullMDBNode_def nullPointer_def)
-- "tcb"
\<comment> \<open>tcb\<close>
apply (erule cmap_relation_updI2 [where dest = "ctcb_ptr_to_tcb_ptr p" and f = "tcb_ptr_to_ctcb_ptr", simplified])
apply (rule map_comp_simps)
apply (rule pks)
apply (rule tcb_rel[simplified FLAGS_default_eq, simplified])
-- "ep"
\<comment> \<open>ep\<close>
apply (erule iffD2 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply (simp add: cendpoint_relation_def Let_def)
apply (subst endpoint.case_cong)
@ -4220,7 +4220,7 @@ proof -
apply (simp add: tcb_queue_update_other' del: tcb_queue_relation'_empty)
apply (simp add: tcb_queue_update_other' ep2)
apply clarsimp
-- "ntfn"
\<comment> \<open>ntfn\<close>
apply (erule iffD2 [OF cmap_relation_cong, OF refl refl, rotated -1])
apply (simp add: cnotification_relation_def Let_def)
apply (subst ntfn.case_cong)
@ -6126,7 +6126,7 @@ proof -
vmrights_to_H_def mask_def vm_rights_defs c_valid_cap_def cl_valid_cap_def
vm_page_map_type_defs maptype_to_H_def)
-- "Page objects: could possibly fix the duplication here"
\<comment> \<open>Page objects: could possibly fix the duplication here\<close>
apply (cinit' lift: t_' regionBase_' userSize_' deviceMemory_')
apply (simp add: object_type_from_H_def Kernel_C_defs)
apply (simp add: ccorres_cond_univ_iff ccorres_cond_empty_iff
@ -6177,7 +6177,7 @@ proof -
vmrights_to_H_def mask_def vm_rights_defs c_valid_cap_def cl_valid_cap_def
vm_page_map_type_defs maptype_to_H_def)
-- "PageTableObject"
\<comment> \<open>PageTableObject\<close>
apply (cinit' lift: t_' regionBase_' userSize_' deviceMemory_')
apply (simp add: object_type_from_H_def Kernel_C_defs)
apply (simp add: ccorres_cond_univ_iff ccorres_cond_empty_iff
@ -6206,7 +6206,7 @@ proof -
vmrights_to_H_def)
apply (clarsimp simp: to_bool_def false_def isFrameType_def)
-- "PageDirectoryObject"
\<comment> \<open>PageDirectoryObject\<close>
apply (cinit' lift: t_' regionBase_' userSize_' deviceMemory_')
apply (simp add: object_type_from_H_def Kernel_C_defs)
apply (simp add: ccorres_cond_univ_iff ccorres_cond_empty_iff
@ -6235,7 +6235,7 @@ proof -
vmrights_to_H_def bit_simps)
apply (clarsimp simp: to_bool_def false_def isFrameType_def)
-- "PDPointerTableObject"
\<comment> \<open>PDPointerTableObject\<close>
apply (cinit' lift: t_' regionBase_' userSize_' deviceMemory_')
apply (simp add: object_type_from_H_def Kernel_C_defs)
apply (simp add: ccorres_cond_univ_iff ccorres_cond_empty_iff
@ -6263,7 +6263,7 @@ proof -
vmrights_to_H_def bit_simps)
apply (clarsimp simp: to_bool_def false_def isFrameType_def)
-- "PML4Object"
\<comment> \<open>PML4Object\<close>
apply (cinit' lift: t_' regionBase_' userSize_' deviceMemory_')
apply (simp add: object_type_from_H_def Kernel_C_defs)
apply (simp add: ccorres_cond_univ_iff ccorres_cond_empty_iff

View File

@ -106,7 +106,7 @@ lemma cap_get_tag_isCap0:
apply (erule ccap_relationE)
apply (simp add: cap_to_H_def cap_lift_def Let_def isArchCap_tag_def2 isArchCap_def)
by (clarsimp simp: isCap_simps cap_tag_defs word_le_nat_alt pageSize_def Let_def
split: if_split_asm) -- "takes a while"
split: if_split_asm) \<comment> \<open>takes a while\<close>
@ -1532,7 +1532,7 @@ lemma cmap_relation_cong:
apply (rule Some_the [where f = cm'])
apply (erule subsetD)
apply (erule imageI)
-- "clag"
\<comment> \<open>clag\<close>
apply simp
apply (erule conjE)
apply (drule equalityD1)
@ -2132,7 +2132,7 @@ lemma cap_get_tag_isCap_ArchObject0:
apply -
apply (erule ccap_relationE)
apply (simp add: cap_to_H_def cap_lift_def Let_def isArchCap_def)
by (clarsimp simp: isCap_simps cap_tag_defs word_le_nat_alt pageSize_def Let_def split: if_split_asm) -- "takes a while"
by (clarsimp simp: isCap_simps cap_tag_defs word_le_nat_alt pageSize_def Let_def split: if_split_asm) \<comment> \<open>takes a while\<close>
lemma cap_get_tag_isCap_ArchObject:
assumes cr: "ccap_relation (capability.ArchObjectCap cap) cap'"

View File

@ -275,7 +275,7 @@ proof -
apply (rule_tac R="\<lambda>s. l1 = ksReadyQueuesL1Bitmap s curdom \<and> curdom \<le> maxDomain"
in ccorres_cond)
subgoal by (fastforce dest!: rf_sr_cbitmap_L1_relation simp: cbitmap_L1_relation_def)
prefer 2 -- "switchToIdleThread"
prefer 2 \<comment> \<open>switchToIdleThread\<close>
apply (ctac(no_vcg) add: switchToIdleThread_ccorres)
apply clarsimp
apply (rule ccorres_rhs_assoc)+

View File

@ -1010,7 +1010,7 @@ lemma cready_queues_relation_null_queue_ptrs:
apply (clarsimp simp: tcb_null_ep_ptrs_def)
apply (case_tac z, case_tac a)
apply simp
-- "clag"
\<comment> \<open>clag\<close>
apply (rule ext)
apply (case_tac "mp' x")
apply (frule compD [OF same])
@ -1414,7 +1414,7 @@ schematic_goal user_fpu_state_C_typ_info_unfold:
lemma user_fpu_state_C_in_tcb_C_offset:
"(typ_uinfo_t TYPE(user_fpu_state_C), n) \<in> td_set (typ_uinfo_t TYPE(tcb_C)) 0 \<Longrightarrow> n = 0"
-- \<open>Examine the fields of tcb_C.\<close>
\<comment> \<open>Examine the fields of tcb_C.\<close>
apply (simp add: typ_uinfo_t_def tcb_C_typ_info_unfold td_set_export_uinfo_eq td_set_adjust_ti_eq
image_comp image_Un apfst_comp o_def[where f=export_uinfo]
del: export_uinfo_typdesc_simp)
@ -1422,20 +1422,20 @@ lemma user_fpu_state_C_in_tcb_C_offset:
apply (all \<open>drule td_set_image_field_lookup[rotated]; clarsimp\<close>)
apply (all \<open>drule arg_cong[where f=typ_name]; simp add: adjust_ti_def\<close>)
apply (all \<open>(solves \<open>drule field_lookup_td_set; drule td_set_td_names; simp\<close>)?\<close>)
-- \<open>Only the arch_tcb_C may contain user_fpu_state_C, so examine the fields of that.\<close>
\<comment> \<open>Only the arch_tcb_C may contain user_fpu_state_C, so examine the fields of that.\<close>
apply (drule field_lookup_td_set)
apply (simp add: arch_tcb_C_typ_info_unfold td_set_adjust_ti_eq)
apply (drule td_set_image_field_lookup[rotated]; clarsimp simp: adjust_ti_def)
-- \<open>Similarly for user_context_C.\<close>
\<comment> \<open>Similarly for user_context_C.\<close>
apply (drule field_lookup_td_set)
apply (simp add: user_context_C_typ_info_unfold td_set_adjust_ti_eq)
apply (elim disjE)
apply (all \<open>drule td_set_image_field_lookup[rotated]; clarsimp simp: adjust_ti_def\<close>)
apply (all \<open>(solves \<open>drule field_lookup_td_set; drule td_set_td_names; simp\<close>)?\<close>)
-- \<open>Finally, we have user_fpu_state_C.\<close>
\<comment> \<open>Finally, we have user_fpu_state_C.\<close>
apply (rename_tac fns s)
apply (case_tac fns; clarsimp)
-- \<open>But we must also show that there is no user_fpu_state_C buried within user_fpu_state_C.\<close>
\<comment> \<open>But we must also show that there is no user_fpu_state_C buried within user_fpu_state_C.\<close>
apply (drule field_lookup_td_set)
apply (simp add: user_fpu_state_C_typ_info_unfold td_set_adjust_ti_eq)
apply (elim disjE; clarsimp simp: adjust_ti_def)

View File

@ -3870,7 +3870,7 @@ lemma bindNotification_ccorres:
typ_heap_simps')
apply (elim conjE)
apply (intro conjI)
-- "tcb relation"
\<comment> \<open>tcb relation\<close>
apply (rule cpspace_relation_ntfn_update_ntfn, assumption+)
apply (clarsimp simp: cnotification_relation_def Let_def
mask_def [where n=2] NtfnState_Waiting_def)
@ -4094,7 +4094,7 @@ lemma decodeBindNotification_ccorres:
apply (rule_tac P="\<lambda>s. ksCurThread s = thread" in ccorres_cross_over_guard)
apply (simp add: bindE_bind_linearise del: Collect_const)
apply wpc
-- "IdleNtfn"
\<comment> \<open>IdleNtfn\<close>
apply (simp add: case_option_If del: Collect_const)
apply (rule ccorres_Cond_rhs_Seq)
apply (clarsimp simp: isWaitingNtfn_def from_bool_neq_0)

View File

@ -1904,7 +1904,7 @@ lemma unmapPage_ccorres:
apply (simp add: liftE_bindE Collect_False bind_bindE_assoc
del: Collect_const)
apply (rule ccorres_splitE_novcg[where r'=dc and xf'=xfdc])
-- "X64SmallPage"
\<comment> \<open>X64SmallPage\<close>
apply (rule ccorres_Cond_rhs)
apply (simp add: framesize_to_H_def dc_def[symmetric])
apply (rule ccorres_rhs_assoc)+
@ -1940,7 +1940,7 @@ lemma unmapPage_ccorres:
apply (wp)
apply simp
apply (vcg exspec=lookupPTSlot_modifies)
-- "X64LargePage"
\<comment> \<open>X64LargePage\<close>
apply (rule ccorres_Cond_rhs)
apply (simp add: framesize_to_H_def dc_def[symmetric]
del: Collect_const)
@ -1978,7 +1978,7 @@ lemma unmapPage_ccorres:
apply (wp)
apply simp
apply (vcg exspec=lookupPDSlot_modifies)
-- "X64HugePage"
\<comment> \<open>X64HugePage\<close>
apply (simp add: framesize_to_H_def dc_def[symmetric])
apply (rule ccorres_add_return2)
apply (ctac add: modeUnmapPage_ccorres)

View File

@ -121,11 +121,11 @@ definition
definition
ep_at_C' :: "word64 \<Rightarrow> heap_raw_state \<Rightarrow> bool"
where
"ep_at_C' p h \<equiv> Ptr p \<in> dom (clift h :: endpoint_C typ_heap)" -- "endpoint_lift is total"
"ep_at_C' p h \<equiv> Ptr p \<in> dom (clift h :: endpoint_C typ_heap)" \<comment> \<open>endpoint_lift is total\<close>
definition
ntfn_at_C' :: "word64 \<Rightarrow> heap_raw_state \<Rightarrow> bool"
where -- "notification_lift is total"
where \<comment> \<open>notification_lift is total\<close>
"ntfn_at_C' p h \<equiv> Ptr p \<in> dom (clift h :: notification_C typ_heap)"
definition

View File

@ -647,7 +647,7 @@ lemma autocorres_transfer_spec_no_modifies:
assumes ac_def:
"ac_f \<equiv> AC_call_L1 arg_rel globals ret_xf (L1_call_simpl check_termination \<Gamma> f_'proc)"
assumes c_spec:
"\<forall>s0. hoarep \<Gamma> {} {} (P' s0) (Call f_'proc) (Collect (Q s0)) A" -- \<open>syntax parser barfs...\<close>
"\<forall>s0. hoarep \<Gamma> {} {} (P' s0) (Call f_'proc) (Collect (Q s0)) A" \<comment> \<open>syntax parser barfs...\<close>
assumes c_modifies:
"\<forall>\<sigma>. \<Gamma>\<turnstile>\<^bsub>/UNIV\<^esub> {\<sigma>} Call f_'proc {t. t may_not_modify_globals \<sigma>}"
assumes c_spec_unify:

View File

@ -169,7 +169,7 @@ lemma ccorres_split_nothrow_novcgE:
and valid: "\<lbrace>R\<rbrace> a \<lbrace>Q\<rbrace>,\<lbrace>QE\<rbrace>"
and novcg: "guard_is_UNIV (\<lambda>rv rv'. r' rv (vf' rv'))
xf' (\<lambda>rv rv'. {s. ef' rv' = scast EXCEPTION_NONE \<longrightarrow> s \<in> Q' rv rv'})"
-- "hack"
\<comment> \<open>hack\<close>
and novcg_err: "\<And>err. guard_is_UNIV (\<lambda>rv. f' err (ef' rv)) es
(\<lambda>rv rv'. {s. ef' (xf' s) \<noteq> scast EXCEPTION_NONE \<longrightarrow> s \<in> Q'' err rv rv'})"
shows "ccorres_underlying sr \<Gamma> r xf arrel axf (P and R) P' hs (a >>=E (\<lambda>rv. b rv)) (c ;; d)"
@ -497,7 +497,7 @@ lemma ccorres_split_nothrow_call_record_novcg:
and igl: "\<And>s. globals (i s) = globals s"
and valid: "\<lbrace>R\<rbrace> a \<lbrace>Q\<rbrace>"
and novcg: "guard_is_UNIV r' (xfr \<circ> xf') Q'"
-- "This might cause problems \<dots> has to be preserved across c in vcg case, but we can't do that"
\<comment> \<open>This might cause problems \<dots> has to be preserved across c in vcg case, but we can't do that\<close>
and xfoldv: "\<And>s. xf' s = xfru (\<lambda>_. (xfr \<circ> xf') s) oldv"
shows "ccorres_underlying rf_sr \<Gamma> r xf arrel axf (P and R) ({s. i s \<in> P'}) hs (a >>= (\<lambda>rv. b rv)) (call i f (\<lambda>s t. s\<lparr>globals := globals t\<rparr>)
(\<lambda>_ t. Basic (xfu' (\<lambda>_. xfru (\<lambda>_. xf'' t) oldv))) ;; d)"

View File

@ -558,14 +558,14 @@ lemma pis_Guard:
done
lemmas push_in_stmt_rules =
-- "No ordering apart from pis_base which must be last."
\<comment> \<open>No ordering apart from pis_base which must be last.\<close>
pis_throw
pis_creturn
pis_Seq_right
pis_Cond
pis_switch_Singleton pis_switch_Cons
pis_Guard
-- "Last, just stick it where it is"
\<comment> \<open>Last, just stick it where it is\<close>
pis_base
lemma ccorres_special_trim_guard_DontReach_pis:
@ -581,7 +581,7 @@ lemma ccorres_special_trim_guard_DontReach_pis:
end
lemmas ccorres_boilerplace_simp_dels =
Collect_const -- "Avoid getting an implication due to if_split. Should probably just remove if_split"
Collect_const \<comment> \<open>Avoid getting an implication due to if_split. Should probably just remove if_split\<close>
lemma ccorres_introduce_UNIV_Int_when_needed:
"ccorres_underlying sr Gamm r xf ar axf P (UNIV \<inter> {x. Q x}) hs a c
@ -1030,7 +1030,7 @@ lemma creturn_ceqv_xf:
lemma creturn_ceqv_not_xf:
fixes \<Gamma> :: "(('a globals_scheme, 'b) myvars_scheme) c_body"
assumes rl: "\<And>t. rewrite_xf xf t v rv rv'"
and xfu: "\<And>s f. xf (xfu' f s) = xf s" -- "i.e., xf is independent of xfu"
and xfu: "\<And>s f. xf (xfu' f s) = xf s" \<comment> \<open>i.e., xf is independent of xfu\<close>
and xfg: "\<And>s f. xf (global_exn_var_'_update f s) = xf s"
shows "ceqv \<Gamma> xf v t t' (return_C xfu' rv) (return_C xfu' rv')"
unfolding creturn_def
@ -1092,14 +1092,14 @@ lemma Cond_empty_ceqv: (* Crops up occasionally *)
lemmas Guard_UNIV_ceqv = Guard_ceqv [where x = UNIV and x' = UNIV, simplified]
lemmas ceqv_rules = ceqv_refl [where xf' = xfdc] -- "Any ceqv with xfdc should be ignored"
lemmas ceqv_rules = ceqv_refl [where xf' = xfdc] \<comment> \<open>Any ceqv with xfdc should be ignored\<close>
While_ceqv [OF Collect_mem_eqv] While_ceqv [OF UNIV_mem_eqv]
Cond_ceqv [OF Collect_mem_eqv] Cond_UNIV_ceqv Cond_empty_ceqv
Guard_ceqv [OF Collect_mem_eqv] Guard_UNIV_ceqv
Seq_ceqv Seq_weak_ceqv
Basic_ceqv call_ceqv Skip_ceqv
Catch_ceqv Throw_ceqv
creturn_ceqv_xf creturn_ceqv_not_xf -- "order is important with these two, the second is more general"
creturn_ceqv_xf creturn_ceqv_not_xf \<comment> \<open>order is important with these two, the second is more general\<close>
ceqv_refl [where c = return_void_C] ceqv_refl [where c = break_C]
ceqv_refl [where c = catchbrk_C]

View File

@ -1592,7 +1592,7 @@ lemma invoke_page_corres:
apply (clarsimp simp:invs_def valid_state_def cte_wp_at_caps_of_state)
apply (frule_tac v = b in valid_idle_has_null_cap,simp+)
apply (clarsimp simp:is_arch_update_def is_arch_cap_def cap_master_cap_def split:cap.split_asm)
-- "PageRemap"
\<comment> \<open>PageRemap\<close>
apply (rename_tac sum)
apply (case_tac sum)
apply (clarsimp simp: mapM_singleton mapM_x_mapM)
@ -1631,7 +1631,7 @@ lemma invoke_page_corres:
)+
apply (clarsimp simp:invs_def valid_state_def
cte_wp_at_caps_of_state page_inv_duplicates_valid_def)
-- "PageUnmap"
\<comment> \<open>PageUnmap\<close>
apply (rule dcorres_expand_pfx)
apply (clarsimp simp: valid_page_inv_def transform_mapping_def liftM_def
split:arch_cap.splits option.splits)
@ -1674,7 +1674,7 @@ lemma invoke_page_corres:
apply (rule sym)
apply (simp add:get_cap_caps_of_state)+
-- "PageFlush"
\<comment> \<open>PageFlush\<close>
apply (clarsimp simp:invoke_page_def)
apply (clarsimp simp: when_def split: if_splits)
apply (rule corres_guard_imp)
@ -1688,7 +1688,7 @@ lemma invoke_page_corres:
apply (erule notE)+
apply (clarsimp)
apply (wp do_machine_op_wp | clarsimp)+
-- "Get Address"
\<comment> \<open>Get Address\<close>
apply (rule corres_guard_imp)
apply (rule corres_split[OF _ get_cur_thread_corres])
apply (subst duplicate_corrupt_tcb_intent[symmetric])

View File

@ -1810,7 +1810,7 @@ lemma dcorres_get_object_special:
and ap_lift: "\<And>s obj. \<lbrakk>AP ptr s = Some obj; AP_Q s\<rbrakk>
\<Longrightarrow> opt_object ptr (transform s) = Some (C (AP_LIFT obj s))"
and c: "\<And>obj. dcorres r (R obj) (R' obj) (a (C obj)) c"
-- "weak"
\<comment> \<open>weak\<close>
shows "dcorres r (\<lambda>s. (\<forall>obj. opt_object ptr s = Some (C obj) \<longrightarrow> R obj s))
(\<lambda>s. (\<forall>obj'. AP ptr s = Some obj' \<longrightarrow> R' (AP_LIFT obj' s) s) \<and> AP ptr s \<noteq> None \<and> AP_Q s)
(KHeap_D.get_object ptr >>= a) c"
@ -1826,7 +1826,7 @@ lemma dcorres_get_object_special:
apply (simp add: unc)
apply clarsimp
apply (drule (1) ap_lift [symmetric])
apply (simp, simp add: unc) -- "yuck, first simp applies unc too early"
apply (simp, simp add: unc) \<comment> \<open>yuck, first simp applies unc too early\<close>
apply clarsimp
apply (frule (1) ap_lift)
apply (simp add: unc)
@ -1839,7 +1839,7 @@ lemma dcorres_get_object_special_2:
assumes ap_lift: "\<And>s obj etcb. \<lbrakk> get_tcb ptr s = Some obj; get_etcb ptr s = Some etcb; AP_Q s\<rbrakk>
\<Longrightarrow> opt_object ptr (transform s) = Some (Tcb (AP_LIFT obj etcb s))"
and c: "\<And>obj. dcorres r (R obj) (R' obj) (a (Tcb obj)) c"
-- "weak"
\<comment> \<open>weak\<close>
shows "dcorres r (\<lambda>s. (\<forall>obj. opt_object ptr s = Some (Tcb obj) \<longrightarrow> R obj s))
(\<lambda>s. (\<forall>obj' etcb. get_tcb ptr s = Some obj' \<and> get_etcb ptr s = Some etcb \<longrightarrow> R' (AP_LIFT obj' etcb s) s) \<and>
get_tcb ptr s \<noteq> None \<and> get_etcb ptr s \<noteq> None \<and> AP_Q s)
@ -2260,7 +2260,7 @@ lemma derive_cap_dcorres:
(CSpace_A.derive_cap slot cap)"
unfolding CSpace_D.derive_cap_def derive_cap_def
apply (cases cap, simp_all add: dcorres_returnOk')
-- "Untyped"
\<comment> \<open>Untyped\<close>
apply (rule corres_guard_imp)
apply (rule corres_splitEE[where r' = "op ="])
apply (rule dcorres_returnOk)
@ -2485,7 +2485,7 @@ lemma decode_cnode_corres:
prefer 2
apply fastforce
apply (fastforce simp: valid_mdb_def mdb_cte_at_def)
-- Mint
\<comment> \<open>Mint\<close>
apply (rule corres_guard_imp)
apply (rule corres_splitEE[OF _ lookup_slot_for_cnode_op_corres])
apply simp
@ -2519,7 +2519,7 @@ lemma decode_cnode_corres:
apply (fastforce simp: valid_mdb_def mdb_cte_at_def)
apply fastforce
-- Move
\<comment> \<open>Move\<close>
apply (rule corres_guard_imp)
apply (rule corres_splitEE)
prefer 2
@ -2546,7 +2546,7 @@ lemma decode_cnode_corres:
apply (simp add: translate_cnode_invocation_def all_rights_def)
apply ((wp hoare_drop_imps get_cap_wellformed lsfco_not_idle|simp)+)
apply fastforce
-- Mutate
\<comment> \<open>Mutate\<close>
apply (rule corres_guard_imp)
apply (rule corres_splitEE[OF _ lookup_slot_for_cnode_op_corres])
apply simp

View File

@ -1781,7 +1781,7 @@ lemma dcorres_unmap_large_section:
apply (rule conjI)
apply (rule less_kernel_base_mapping_slots)
apply (simp add:pd_bits_def pageBits_def)+
apply (rule conjI,fastforce) -- valid_idle
apply (rule conjI,fastforce) \<comment> \<open>valid_idle\<close>
apply (rule conj_comms[THEN iffD1])
apply (rule context_conjI)
apply (clarsimp simp:tl_map field_simps)
@ -2391,7 +2391,7 @@ lemma dcorres_unmap_page:
apply (rule dcorres_expand_pfx)
apply (clarsimp simp:valid_cap_def)
apply (case_tac vmpage_size)
-- ARMSmallPage
\<comment> \<open>ARMSmallPage\<close>
apply (simp add:ARM_A.unmap_page_def bindE_assoc mapM_x_singleton
PageTableUnmap_D.unmap_page_def cdl_page_mapping_entries_def)
apply (rule corres_guard_imp)
@ -2420,7 +2420,7 @@ prefer 2
| wp lookup_pt_slot_inv find_pd_for_asid_kernel_mapping_help
| rule conjI | clarify)+
-- ARMLargePage
\<comment> \<open>ARMLargePage\<close>
apply (simp add:ARM_A.unmap_page_def bindE_assoc mapM_x_singleton
PageTableUnmap_D.unmap_page_def cdl_page_mapping_entries_def)
@ -2458,7 +2458,7 @@ prefer 2
find_pd_for_asid_kernel_mapping_help
| safe)+
-- Section
\<comment> \<open>Section\<close>
apply (simp add:ARM_A.unmap_page_def bindE_assoc mapM_x_singleton
PageTableUnmap_D.unmap_page_def cdl_page_mapping_entries_def)
apply (rule corres_guard_imp)
@ -2486,7 +2486,7 @@ prefer 2
| wp lookup_pt_slot_inv find_pd_for_asid_kernel_mapping_help
| safe)+
-- SuperSection
\<comment> \<open>SuperSection\<close>
apply (simp add:ARM_A.unmap_page_def bindE_assoc mapM_x_singleton
PageTableUnmap_D.unmap_page_def cdl_page_mapping_entries_def)
@ -2660,7 +2660,7 @@ lemma dcorres_finalise_cap:
apply (case_tac arch_cap)
apply (simp_all add: arch_finalise_cap_def split:arch_cap.split_asm)
apply clarsimp
-- arch_cap.ASIDPoolCap
\<comment> \<open>arch_cap.ASIDPoolCap\<close>
apply (rule corres_guard_imp)
apply (simp add:transform_asid_def)
apply (rule corres_split[OF _ dcorres_delete_asid_pool])
@ -2668,7 +2668,7 @@ lemma dcorres_finalise_cap:
apply (clarsimp simp:transform_cap_def)
apply (wp|clarsimp)+
apply (clarsimp split:option.splits | rule conjI)+
-- arch_cap.PageCap
\<comment> \<open>arch_cap.PageCap\<close>
apply (simp add:transform_mapping_def)
apply (clarsimp simp:transform_mapping_def)
apply (rule corres_guard_imp)
@ -2677,7 +2677,7 @@ lemma dcorres_finalise_cap:
apply (clarsimp simp:transform_cap_def)
apply (wp | clarsimp )+
apply simp
--arch_cap.PageTableCap
\<comment> \<open>arch_cap.PageTableCap\<close>
apply (clarsimp simp:transform_mapping_def split:option.splits)
apply (rule dcorres_expand_pfx)
apply (rule corres_guard_imp)

View File

@ -56,7 +56,7 @@ lemma decode_irq_control_corres:
apply (simp add: transform_intent_def transform_intent_issue_irq_handler_def
split: list.splits)
apply (rule conjI)
prefer 2 -- "error case"
prefer 2 \<comment> \<open>error case\<close>
apply clarsimp
apply (rule corres_guard_imp)
apply (rule dcorres_alternative_throw)
@ -142,7 +142,7 @@ lemma decode_irq_handler_corres:
apply (cases "invocation_type label' = IRQSetIRQHandler")
apply (simp add: transform_intent_def cdl_cap_irq_def)
apply (rule conjI)
prefer 2 -- "excaps' = []"
prefer 2 \<comment> \<open>excaps' = []\<close>
apply (clarsimp intro!: corres_alternate2)
apply (clarsimp simp: neq_Nil_conv)
apply (rule conjI)

View File

@ -894,7 +894,7 @@ lemma send_signal_corres:
apply (rename_tac ntfn_ext)
apply (case_tac "ntfn_obj ntfn_ext", clarsimp)
apply (case_tac "ntfn_bound_tcb ntfn_ext", clarsimp)
-- "Idle, not bound"
\<comment> \<open>Idle, not bound\<close>
apply (rule corres_alternate1)
apply (rule dcorres_absorb_get_l)
apply (clarsimp split del: if_split)
@ -904,12 +904,12 @@ lemma send_signal_corres:
apply (frule get_notification_pick,simp)
apply (clarsimp simp:ntfn_waiting_set_lift valid_state_def valid_ntfn_abstract_def none_is_waiting_ntfn_def)
apply (rule corres_guard_imp,rule corres_dummy_set_notification,simp+)[1]
-- "Idle, bound"
\<comment> \<open>Idle, bound\<close>
apply (clarsimp simp: get_thread_state_def thread_get_def gets_the_def gets_def bind_assoc split del: if_split)
apply (rule dcorres_absorb_get_r)
apply (clarsimp simp: assert_opt_def corres_free_fail split: Structures_A.kernel_object.splits option.splits)
apply (case_tac "receive_blocked (tcb_state x2)")
-- "receive_blocked"
\<comment> \<open>receive_blocked\<close>
apply (clarsimp)
apply (rule corres_alternate2)
apply (clarsimp simp: send_signal_bound_def gets_def)
@ -947,7 +947,7 @@ lemma send_signal_corres:
apply (frule get_notification_pick,simp)
apply (clarsimp simp:ntfn_waiting_set_lift valid_state_def valid_ntfn_abstract_def none_is_waiting_ntfn_def)
apply (rule corres_guard_imp,rule corres_dummy_set_notification,simp+)[1]
-- "Waiting"
\<comment> \<open>Waiting\<close>
apply (rule corres_alternate1)
apply (rule dcorres_absorb_get_l)
apply (clarsimp split del: if_split)
@ -968,7 +968,7 @@ lemma send_signal_corres:
apply (drule_tac A="ntfn_waiting_set epptr s'" and x = y in eqset_imp_iff)
apply (clarsimp simp:valid_pspace_def ntfn_waiting_set_def)
apply (clarsimp simp: pred_tcb_at_def obj_at_def valid_ntfn_def split:list.splits option.splits)
-- "Active"
\<comment> \<open>Active\<close>
apply (rule corres_alternate1)
apply (rule dcorres_absorb_get_l)
apply (clarsimp split del: if_split)
@ -2541,7 +2541,7 @@ lemma recv_sync_ipc_corres:
apply (rule dcorres_absorb_get_r)
apply (clarsimp simp: assert_opt_def corres_free_fail split: Structures_A.kernel_object.splits option.splits )
apply safe[1]
-- "not bound"
\<comment> \<open>not bound\<close>
apply (rule corres_alternate2)
apply (rule dcorres_receive_sync, simp_all)[1]
apply (simp add: get_simple_ko_def gets_def get_object_def bind_assoc)

View File

@ -740,7 +740,7 @@ proof -
apply (clarsimp simp:assert_def corres_free_fail)
apply (rename_tac obj')
apply (case_tac obj', simp_all add:corres_free_fail split del: if_split)
-- "cnode or IRQ Node case"
\<comment> \<open>cnode or IRQ Node case\<close>
apply (clarsimp simp: corres_free_fail split: if_split)
apply (rename_tac sz cn ocap)
apply (clarsimp simp: corres_underlying_def in_monad set_object_def cte_wp_at_cases caps_of_state_cte_wp_at)
@ -764,7 +764,7 @@ proof -
apply (rule ext)
apply clarsimp
apply (simp add: option_map_def restrict_map_def map_add_def split: option.splits)
-- "tcb case"
\<comment> \<open>tcb case\<close>
apply (drule(1) valid_etcbs_tcb_etcb)
apply (clarsimp simp: cdl_objects_tcb opt_object_def
assert_opt_def has_slots_def object_slots_def

View File

@ -647,7 +647,7 @@ where
map (\<lambda>(cap, slot). (transform_cap cap, transform_cslot_ptr slot))"
-- "Convert a nat into a bool list of the given size."
\<comment> \<open>Convert a nat into a bool list of the given size.\<close>
definition
nat_to_bl :: "nat \<Rightarrow> nat \<Rightarrow> bool list option"
where

View File

@ -96,7 +96,7 @@ lemma kernel_entry_if_valid_domain_time:
apply (rule hoare_pre)
apply (wp handle_interrupt_valid_domain_time
| clarsimp | wpc)+
-- "strengthen post of do_machine_op; we know interrupt occurred"
\<comment> \<open>strengthen post of do_machine_op; we know interrupt occurred\<close>
apply (rule_tac Q="\<lambda>_ s. 0 < domain_time s" in hoare_post_imp, fastforce)
apply (wp+, simp)
done
@ -1198,7 +1198,7 @@ lemma kernelEntry_if_valid_domain_time:
apply (clarsimp simp: handleEvent_def)
apply (rule hoare_pre)
apply (wp handleInterrupt_valid_domain_time | wpc | clarsimp)+
apply (rule hoare_false_imp) -- "debugPrint"
apply (rule hoare_false_imp) \<comment> \<open>debugPrint\<close>
apply (wp handleInterrupt_valid_domain_time hoare_vcg_all_lift hoare_drop_imps | simp)+
done

View File

@ -189,10 +189,10 @@ lemma handleEvent_ccorres:
apply (wpc, simp_all)[1]
apply (wpc, simp_all add: handleSyscall_C_body_if_def syscall_from_H_def liftE_def
ccorres_cond_univ_iff syscall_defs ccorres_cond_empty_iff)[1]
-- "SysCall"
\<comment> \<open>SysCall\<close>
apply (simp add: handleCall_def)
apply (ctac (no_vcg) add: handleInvocation_ccorres)
-- "SysReplyRecv"
\<comment> \<open>SysReplyRecv\<close>
apply (simp add: bind_assoc)
apply (rule ccorres_rhs_assoc)+
apply (ctac (no_vcg) add: handleReply_ccorres)
@ -208,37 +208,37 @@ lemma handleEvent_ccorres:
in hoare_post_imp)
apply (simp add: ct_in_state'_def)
apply (wp handleReply_sane handleReply_ct_not_ksQ)
-- "SysSend"
\<comment> \<open>SysSend\<close>
apply (simp add: handleSend_def)
apply (ctac (no_vcg) add: handleInvocation_ccorres)
-- "SysNBSend"
\<comment> \<open>SysNBSend\<close>
apply (simp add: handleSend_def)
apply (ctac (no_vcg) add: handleInvocation_ccorres)
-- "SysRecv"
\<comment> \<open>SysRecv\<close>
apply (ctac (no_vcg) add: handleRecv_ccorres)
apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg)
apply (rule allI, rule conseqPre, vcg)
apply (clarsimp simp: return_def)
apply wp
-- "SysReply"
\<comment> \<open>SysReply\<close>
apply (ctac (no_vcg) add: handleReply_ccorres)
apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg)
apply (rule allI, rule conseqPre, vcg)
apply (clarsimp simp: return_def)
apply wp
-- "SysYield"
\<comment> \<open>SysYield\<close>
apply (ctac (no_vcg) add: handleYield_ccorres)
apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg)
apply (rule allI, rule conseqPre, vcg)
apply (clarsimp simp: return_def)
apply wp
-- "SysNBRecv"
\<comment> \<open>SysNBRecv\<close>
apply (ctac (no_vcg) add: handleRecv_ccorres)
apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg)
apply (rule allI, rule conseqPre, vcg)
apply (clarsimp simp: return_def)
apply wp
-- "UnknownSyscall"
\<comment> \<open>UnknownSyscall\<close>
apply (simp add: liftE_def bind_assoc handleUnknownSyscall_C_body_if_def)
apply (rule ccorres_pre_getCurThread)
apply (rule ccorres_symb_exec_r)
@ -249,7 +249,7 @@ lemma handleEvent_ccorres:
apply wp
apply (clarsimp, vcg)
apply (clarsimp, rule conseqPre, vcg, clarsimp)
-- "UserLevelFault"
\<comment> \<open>UserLevelFault\<close>
apply (simp add: liftE_def bind_assoc handleUserLevelFault_C_body_if_def)
apply (rule ccorres_pre_getCurThread)
apply (rule ccorres_symb_exec_r)
@ -260,9 +260,9 @@ lemma handleEvent_ccorres:
apply wp
apply (clarsimp, vcg)
apply (clarsimp, rule conseqPre, vcg, clarsimp)
-- "Interrupt"
\<comment> \<open>Interrupt\<close>
apply (ctac add: handleInterrupt_ccorres[unfolded handleEvent_def, simplified])
-- "VMFaultEvent"
\<comment> \<open>VMFaultEvent\<close>
apply (simp add: liftE_def bind_assoc handleVMFaultEvent_C_body_if_def)
apply (rule ccorres_pre_getCurThread)
apply (simp add: catch_def)
@ -295,7 +295,7 @@ lemma handleEvent_ccorres:
is_cap_fault_def
elim: pred_tcb'_weakenE st_tcb_ex_cap''
dest: st_tcb_at_idle_thread' rf_sr_ksCurThread)
-- "HypervisorEvent"
\<comment> \<open>HypervisorEvent\<close>
apply (simp add: liftE_def bind_assoc)
apply (rule ccorres_guard_imp2)
apply (rule ccorres_symb_exec_l)

View File

@ -2254,10 +2254,10 @@ lemma decode_arch_invocation_authorised_for_globals:
neq_Nil_conv invs_psp_aligned invs_vspace_objs cli_no_irqs)
apply (drule diminished_cte_wp_at_valid_cap, clarsimp+)
apply (cases cap, simp_all)
-- "PageCap"
\<comment> \<open>PageCap\<close>
apply (clarsimp simp: valid_cap_simps cli_no_irqs)
apply (cases "invocation_type label";(rename_tac arch, case_tac arch; simp add: isPageFlushLabel_def isPDFlushLabel_def))
-- "Map"
\<comment> \<open>Map\<close>
apply (rename_tac word cap_rights vmpage_size option arch)
apply(clarsimp simp: isPageFlushLabel_def isPDFlushLabel_def | rule conjI)+
apply(drule diminished_cte_wp_at_valid_cap)
@ -2268,13 +2268,13 @@ lemma decode_arch_invocation_authorised_for_globals:
apply(insert pbfs_less_wb)
apply(clarsimp)
apply(fastforce simp: x_power_minus_1)
-- "Remap"
\<comment> \<open>Remap\<close>
apply(clarsimp)
apply(fastforce dest: diminished_cte_wp_at_valid_cap simp: invs_def valid_state_def valid_cap_def)
-- "Unmap"
\<comment> \<open>Unmap\<close>
apply(simp add: authorised_for_globals_page_inv_def)+
apply(clarsimp)
-- "PageTableCap"
\<comment> \<open>PageTableCap\<close>
apply(simp add: authorised_for_globals_page_table_inv_def)
apply(clarsimp)
apply(frule_tac vptr="msg ! 0" in pd_shifting')

View File

@ -1195,7 +1195,7 @@ lemma unbind_maybe_notification_reads_respects:
apply wp
apply (case_tac "ntfn_bound_tcb rv")
apply (clarsimp, wp)[1]
-- "interesting case, ntfn is bound"
\<comment> \<open>interesting case, ntfn is bound\<close>
apply (clarsimp)
apply ((wp set_bound_notification_none_reads_respects set_simple_ko_reads_respects
get_simple_ko_reads_respects

View File

@ -589,7 +589,7 @@ lemma blocked_cancel_ipc_nosts_reads_respects:
subgoal
proof (cases "aag_can_read_label aag (pasObjectAbs aag x) \<or> aag_can_affect aag l x")
case True thus ?thesis -- "boring case, can read or affect ep"
case True thus ?thesis \<comment> \<open>boring case, can read or affect ep\<close>
unfolding blocked_cancel_ipc_nosts_def get_blocking_object_def
apply clarsimp
apply (rule pre_ev)
@ -603,7 +603,7 @@ lemma blocked_cancel_ipc_nosts_reads_respects:
by (fastforce dest:read_sync_ep_read_receivers_strong )
next
case False thus ?thesis apply - -- "can't read or affect ep"
case False thus ?thesis apply - \<comment> \<open>can't read or affect ep\<close>
apply (rule gen_asm_ev)
apply (drule label_is_invisible[THEN iffD2])
apply clarsimp

View File

@ -1999,7 +1999,7 @@ lemma integrity_part:
(s, s') \<in> Simulation.Step (big_step_ADT_A_if utf) ();
(part s, u) \<notin> policyFlows (pasPolicy initial_aag); u \<noteq> PSched; part s \<noteq> PSched\<rbrakk> \<Longrightarrow>
(s,s') \<in> uwr u"
supply [[simp_depth_limit=0]] -- "speedup"
supply [[simp_depth_limit=0]] \<comment> \<open>speedup\<close>
apply(simp add: uwr_def_cur[where s=s])
apply(case_tac s, case_tac s', simp)
apply(case_tac a, case_tac aa, simp)
@ -2018,7 +2018,7 @@ lemma integrity_part:
apply(fastforce dest: silc_inv_initial_aag_reachable[OF reachable_Step'] simp: silc_inv_cur)
apply(rule pas_wellformed_cur)
apply(simp add: current_aag_def)
supply [[simp_depth_limit=0]] -- "speeds up the rest of the proof"
supply [[simp_depth_limit=0]] \<comment> \<open>speeds up the rest of the proof\<close>
apply(fastforce dest!: reachable_invs_if domains_distinct[THEN pas_domains_distinct_inj]
simp: invs_if_def Invs_def guarded_pas_domain_def
guarded_is_subject_cur_thread_def current_aag_def)
@ -2895,7 +2895,7 @@ lemma do_user_op_A_if_confidentiality:
xx = yy \<and>
(s', t') \<in> uwr u \<and> (s', t') \<in> uwr PSched \<and> (s', t') \<in> uwr (part s)"
including no_pre
supply [[simp_depth_limit=2]] -- "speedup"
supply [[simp_depth_limit=2]] \<comment> \<open>speedup\<close>
apply(frule (1) uwr_part_sys_mode_of_user_context_of_eq)
apply(clarsimp simp: check_active_irq_A_if_def)
apply(case_tac s, case_tac t, simp_all)
@ -2998,7 +2998,7 @@ lemma kernel_schedule_if_confidentiality:
((fst t),(),(fst t')) \<in> kernel_schedule_if;
snd s' = snd t'\<rbrakk> \<Longrightarrow>
(s', t') \<in> uwr u \<and> (s', t') \<in> uwr PSched \<and> (s', t') \<in> uwr (part s)"
supply [[simp_depth_limit=1]] -- "speedup"
supply [[simp_depth_limit=1]] \<comment> \<open>speedup\<close>
apply(frule (1) uwr_part_sys_mode_of_user_context_of_eq)
apply(frule part_not_PSched_sys_mode_of_not_KernelSchedule_True)
apply(clarsimp simp: kernel_schedule_if_def)
@ -3128,7 +3128,7 @@ lemma kernel_call_A_if_confidentiality:
snd s' = f x; snd t' = f y\<rbrakk> \<Longrightarrow>
x = y \<and>
(s', t') \<in> uwr u \<and> (s', t') \<in> uwr PSched \<and> (s', t') \<in> uwr (part s)"
supply [[simp_depth_limit=3]] -- "speedup"
supply [[simp_depth_limit=3]] \<comment> \<open>speedup\<close>
apply(frule (1) uwr_part_sys_mode_of_user_context_of_eq)
apply(frule part_not_PSched_sys_mode_of_not_KernelSchedule_True)
apply(clarsimp simp: kernel_call_A_if_def)
@ -3305,7 +3305,7 @@ lemma kernel_exit_A_if_confidentiality:
snd s' = f x; snd t' = f y\<rbrakk> \<Longrightarrow>
x = y \<and>
(s', t') \<in> uwr u \<and> (s', t') \<in> uwr PSched \<and> (s', t') \<in> uwr (part s)"
supply [[simp_depth_limit=2]] -- "speedup"
supply [[simp_depth_limit=2]] \<comment> \<open>speedup\<close>
apply(clarsimp simp: kernel_exit_A_if_def)
apply(case_tac s, case_tac t, simp_all)
apply(case_tac u, simp_all)
@ -4040,7 +4040,7 @@ lemma scheduler_step_1_confidentiality:
assumes reach_t: "system.reachable (ADT_A_if utf) s0 t"
shows "\<lbrakk>interrupted_modes (sys_mode_of s)\<rbrakk> \<Longrightarrow>
(s',t') \<in> uwr u"
supply [[simp_depth_limit=2]] -- "speedup"
supply [[simp_depth_limit=2]] \<comment> \<open>speedup\<close>
apply (insert uwr step_s step_t)
apply (cut_tac ADT_A_if_reachable_invs_if[OF reach_s])
apply (cut_tac ADT_A_if_reachable_invs_if[OF reach_t])

View File

@ -660,19 +660,19 @@ lemma invoke_tcb_reads_respects_f:
defer
apply((wp suspend_reads_respects_f[where st=st] restart_reads_respects_f[where st=st]
| simp add: authorised_tcb_inv_def )+)[2]
-- "NotificationControl"
\<comment> \<open>NotificationControl\<close>
apply (rename_tac option)
apply (case_tac option, simp_all)[1]
apply ((wp unbind_notification_is_subj_reads_respects unbind_notification_silc_inv
bind_notification_reads_respects
| clarsimp simp: authorised_tcb_inv_def
| rule_tac Q=\<top> and st=st in reads_respects_f)+)[2]
-- "SetTLSBase"
\<comment> \<open>SetTLSBase\<close>
apply (rename_tac tcb tls_base)
apply(wpsimp wp: when_ev reschedule_required_reads_respects_f
as_user_reads_respects_f hoare_drop_imps)+
apply(auto simp: det_setRegister authorised_tcb_inv_def)[1]
-- "ThreadControl"
apply(auto simp: det_setRegister authorised_tcb_inv_def)[1]
\<comment> \<open>ThreadControl\<close>
apply (simp add: split_def cong: option.case_cong)
apply (wpsimp wp: hoare_vcg_const_imp_lift_R simp: when_def | wpc)+
apply (rule conjI)

View File

@ -12,7 +12,7 @@ theory ArchInvariants_AI
imports "../InvariantsPre_AI"
begin
-- ---------------------------------------------------------------------------
\<comment> \<open>---------------------------------------------------------------------------\<close>
section "Move this up"
qualify ARM (in Arch)
@ -24,7 +24,7 @@ axiomatization
end_qualify
-- ---------------------------------------------------------------------------
\<comment> \<open>---------------------------------------------------------------------------\<close>
section "ARM-specific invariant definitions"
qualify ARM_A (in Arch)

View File

@ -4545,7 +4545,7 @@ lemma perform_page_invs [wp]:
"\<lbrace>invs and valid_page_inv pg_inv\<rbrace> perform_page_invocation pg_inv \<lbrace>\<lambda>_. invs\<rbrace>"
apply (simp add: perform_page_invocation_def)
apply (cases pg_inv, simp_all)
-- "PageMap"
\<comment> \<open>PageMap\<close>
apply (rename_tac asid cap cslot_ptr sum)
apply clarsimp
apply (rule hoare_pre)
@ -4627,7 +4627,7 @@ lemma perform_page_invs [wp]:
apply (drule_tac x=sl in imageI[where f="\<lambda>x. x && ~~ mask pd_bits"])
apply (drule (1) subsetD)
apply (clarsimp simp: cap_range_def)
-- "PageRemap"
\<comment> \<open>PageRemap\<close>
apply (rule hoare_pre)
apply (wp get_master_pte_wp get_master_pde_wp hoare_vcg_ex_lift mapM_x_swp_store_pde_invs_unmap
| wpc | simp add: pte_check_if_mapped_def pde_check_if_mapped_def
@ -4674,7 +4674,7 @@ lemma perform_page_invs [wp]:
apply simp+
apply force
-- "PageUnmap"
\<comment> \<open>PageUnmap\<close>
apply (rename_tac arch_cap cslot_ptr)
apply (rule hoare_pre)
apply (wp dmo_invs arch_update_cap_invs_unmap_page get_cap_wp
@ -4704,7 +4704,7 @@ lemma perform_page_invs [wp]:
apply (auto simp: valid_cap_def cap_aligned_def mask_def vs_cap_ref_def data_at_def
split: vmpage_size.splits option.splits if_splits)[1]
-- "PageFlush"
\<comment> \<open>PageFlush\<close>
apply (rule hoare_pre)
apply (wp dmo_invs set_vm_root_for_flush_invs
hoare_vcg_const_imp_lift hoare_vcg_all_lift

View File

@ -12,7 +12,7 @@ theory ArchInvariants_AI
imports "../InvariantsPre_AI"
begin
-- ---------------------------------------------------------------------------
\<comment> \<open>---------------------------------------------------------------------------\<close>
section "Move this up"
qualify ARM (in Arch)
@ -24,7 +24,7 @@ axiomatization
end_qualify
-- ---------------------------------------------------------------------------
\<comment> \<open>---------------------------------------------------------------------------\<close>
section "ARM-specific invariant definitions"
qualify ARM_HYP_A (in Arch)

View File

@ -5728,7 +5728,7 @@ lemma perform_page_invs [wp]:
"\<lbrace>invs and valid_page_inv pinv\<rbrace> perform_page_invocation pinv \<lbrace>\<lambda>_. invs\<rbrace>"
apply (simp add: perform_page_invocation_def)
apply (cases pinv, simp_all)
-- "PageMap"
\<comment> \<open>PageMap\<close>
apply (rename_tac asid cap cslot_ptr sum)
apply clarsimp
apply (rule hoare_pre)
@ -5810,7 +5810,7 @@ lemma perform_page_invs [wp]:
apply (drule_tac x=sl in imageI[where f="\<lambda>x. x && ~~ mask pd_bits"])
apply (drule (1) subsetD)
apply (clarsimp simp: cap_range_def)
-- "PageRemap"
\<comment> \<open>PageRemap\<close>
apply (rule hoare_pre)
apply (wp get_master_pte_wp get_master_pde_wp hoare_vcg_ex_lift mapM_x_swp_store_pde_invs_unmap
| wpc | simp add: pte_check_if_mapped_def pde_check_if_mapped_def
@ -5869,7 +5869,7 @@ lemma perform_page_invs [wp]:
valid_global_refsD[OF invs_valid_global_refs caps_of_state_cteD]) *)
apply simp+
apply force
-- "PageUnmap"
\<comment> \<open>PageUnmap\<close>
apply (rename_tac arch_cap cslot_ptr)
apply (rule hoare_pre)
apply (wp dmo_invs arch_update_cap_invs_unmap_page get_cap_wp
@ -5898,7 +5898,7 @@ lemma perform_page_invs [wp]:
using valid_validate_vm_rights[simplified valid_vm_rights_def]
apply (auto simp: valid_cap_def cap_aligned_def mask_def vs_cap_ref_def vspace_bits_defs data_at_def
split: vmpage_size.splits option.splits if_splits)[1]
-- "PageFlush"
\<comment> \<open>PageFlush\<close>
apply (rule hoare_pre)
apply (wp dmo_invs set_vm_root_for_flush_invs
hoare_vcg_const_imp_lift hoare_vcg_all_lift

View File

@ -437,7 +437,7 @@ lemma decode_cnode_inv_wf[wp]:
decode_cnode_invocation mi args cap cs
\<lbrace>valid_cnode_inv\<rbrace>,-"
apply (rule decode_cnode_cases2[where args=args and exs=cs and label=mi])
-- "Move/Insert"
\<comment> \<open>Move/Insert\<close>
apply (simp add: decode_cnode_invocation_def unlessE_whenE
split del: if_split)
apply (wp lsfco_cte_at ensure_no_children_wp whenE_throwError_wp
@ -477,17 +477,17 @@ lemma decode_cnode_inv_wf[wp]:
apply (wp get_cap_cte_wp_at ensure_empty_cte_wp_at)+
apply simp
apply (clarsimp simp: invs_def valid_state_def valid_pspace_def)
-- "Revoke"
\<comment> \<open>Revoke\<close>
apply (simp add: decode_cnode_invocation_def unlessE_whenE cong: if_cong)
apply (wp lsfco_cte_at hoare_drop_imps whenE_throwError_wp
| simp add: split_beta validE_R_def[symmetric])+
apply clarsimp
-- "Delete"
\<comment> \<open>Delete\<close>
apply (simp add: decode_cnode_invocation_def unlessE_whenE cong: if_cong)
apply (wp lsfco_cte_at hoare_drop_imps whenE_throwError_wp
| simp add: split_beta validE_R_def[symmetric])+
apply clarsimp
-- "Save"
\<comment> \<open>Save\<close>
apply (simp add: decode_cnode_invocation_def unlessE_whenE cong: if_cong)
apply (rule hoare_pre)
apply (wp ensure_empty_stronger whenE_throwError_wp
@ -496,7 +496,7 @@ lemma decode_cnode_inv_wf[wp]:
| simp add: split_beta
| wp_once hoare_drop_imps)+
apply clarsimp
-- "CancelBadgedSends"
\<comment> \<open>CancelBadgedSends\<close>
apply (simp add: decode_cnode_invocation_def
unlessE_def whenE_def
split del: if_split)
@ -504,7 +504,7 @@ lemma decode_cnode_inv_wf[wp]:
apply (rule_tac Q'="\<lambda>rv. invs and cte_wp_at (\<lambda>_. True) rv" in hoare_post_imp_R)
apply (wp lsfco_cte_at)
apply (clarsimp simp: cte_wp_valid_cap invs_valid_objs has_cancel_send_rights_ep_cap)+
-- "Rotate"
\<comment> \<open>Rotate\<close>
apply (simp add: decode_cnode_invocation_def split_def
whenE_def unlessE_def)
apply (rule hoare_pre)

View File

@ -129,7 +129,7 @@ lemmas [simp] = acap_rights_update_id state_hyp_refs_update
term "vs_lookup :: 'z::state_ext state \<Rightarrow> vs_chain set"
term "(a \<rhd> b) :: ('z:: state_ext state) \<Rightarrow> bool"
-- ---------------------------------------------------------------------------
\<comment> \<open>---------------------------------------------------------------------------\<close>
section "Invariant Definitions for Abstract Spec"
definition
@ -1057,7 +1057,7 @@ abbreviation(input)
and cur_tcb"
-- ---------------------------------------------------------------------------
\<comment> \<open>---------------------------------------------------------------------------\<close>
section "Lemmas"
lemma valid_bound_ntfn_None[simp]:
@ -1102,7 +1102,7 @@ lemma is_cap_table:
lemmas is_obj_defs = is_ep is_ntfn is_tcb is_cap_table
-- "sanity check"
\<comment> \<open>sanity check\<close>
lemma obj_at_get_object:
"obj_at P ref s \<Longrightarrow> fst (get_object ref s) \<noteq> {}"
by (auto simp: obj_at_def get_object_def gets_def get_def

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