Skip to content

Commit

Permalink
refine+crefine: update for arch_requalify design spec updates
Browse files Browse the repository at this point in the history
Signed-off-by: Rafal Kolanski <[email protected]>
  • Loading branch information
Xaphiosis committed Aug 19, 2024
1 parent 7e538e4 commit e152c9e
Show file tree
Hide file tree
Showing 16 changed files with 34 additions and 34 deletions.
4 changes: 2 additions & 2 deletions proof/crefine/AARCH64/Invoke_C.thy
Original file line number Diff line number Diff line change
Expand Up @@ -2812,10 +2812,10 @@ lemma ctes_of_ex_cte_cap_to':


lemma Arch_isFrameType_spec:
"\<forall>s. \<Gamma> \<turnstile> \<lbrace>s. unat \<acute>type \<le> fromEnum (maxBound::ArchTypes_H.object_type)\<rbrace>
"\<forall>s. \<Gamma> \<turnstile> \<lbrace>s. unat \<acute>type \<le> fromEnum (maxBound::object_type)\<rbrace>
Call Arch_isFrameType_'proc
\<lbrace> \<acute>ret__unsigned_long =
from_bool (isFrameType ((toEnum (unat \<^bsup>s\<^esup> type))::ArchTypes_H.object_type))\<rbrace>"
from_bool (isFrameType ((toEnum (unat \<^bsup>s\<^esup> type))::object_type))\<rbrace>"
apply vcg
apply (simp add: toEnum_object_type_to_H)
apply (frule object_type_from_to_H)
Expand Down
6 changes: 3 additions & 3 deletions proof/crefine/AARCH64/Retype_C.thy
Original file line number Diff line number Diff line change
Expand Up @@ -5047,7 +5047,7 @@ lemma placeNewObject_user_data:
done

definition
createObject_hs_preconds :: "machine_word \<Rightarrow> ArchTypes_H.object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> kernel_state \<Rightarrow> bool"
createObject_hs_preconds :: "machine_word \<Rightarrow> object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> kernel_state \<Rightarrow> bool"
where
"createObject_hs_preconds regionBase newType userSize d \<equiv>
(invs' and pspace_no_overlap' regionBase (getObjectSize newType userSize)
Expand All @@ -5070,14 +5070,14 @@ abbreviation

(* these preconds actually used throughout the proof *)
abbreviation(input)
createObject_c_preconds1 :: "machine_word \<Rightarrow> ArchTypes_H.object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> (globals myvars) set"
createObject_c_preconds1 :: "machine_word \<Rightarrow> object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> (globals myvars) set"
where
"createObject_c_preconds1 regionBase newType userSize deviceMemory \<equiv>
{s. region_actually_is_dev_bytes regionBase (2 ^ getObjectSize newType userSize) deviceMemory s}"

(* these preconds used at start of proof *)
definition
createObject_c_preconds :: "machine_word \<Rightarrow> ArchTypes_H.object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> (globals myvars) set"
createObject_c_preconds :: "machine_word \<Rightarrow> object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> (globals myvars) set"
where
"createObject_c_preconds regionBase newType userSize deviceMemory \<equiv>
(createObject_c_preconds1 regionBase newType userSize deviceMemory
Expand Down
4 changes: 2 additions & 2 deletions proof/crefine/ARM/Invoke_C.thy
Original file line number Diff line number Diff line change
Expand Up @@ -2623,10 +2623,10 @@ lemma ctes_of_ex_cte_cap_to':


lemma Arch_isFrameType_spec:
"\<forall>s. \<Gamma> \<turnstile> \<lbrace>s. unat \<acute>type \<le> fromEnum (maxBound::ArchTypes_H.object_type)\<rbrace>
"\<forall>s. \<Gamma> \<turnstile> \<lbrace>s. unat \<acute>type \<le> fromEnum (maxBound::object_type)\<rbrace>
Call Arch_isFrameType_'proc
\<lbrace> \<acute>ret__unsigned_long =
from_bool (isFrameType ((toEnum (unat \<^bsup>s\<^esup> type))::ArchTypes_H.object_type))\<rbrace>"
from_bool (isFrameType ((toEnum (unat \<^bsup>s\<^esup> type))::object_type))\<rbrace>"
apply vcg
apply (simp add: toEnum_object_type_to_H)
apply (frule object_type_from_to_H)
Expand Down
6 changes: 3 additions & 3 deletions proof/crefine/ARM/Retype_C.thy
Original file line number Diff line number Diff line change
Expand Up @@ -4350,7 +4350,7 @@ lemma placeNewObject_user_data:


definition
createObject_hs_preconds :: "word32 \<Rightarrow> ArchTypes_H.object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> kernel_state \<Rightarrow> bool"
createObject_hs_preconds :: "word32 \<Rightarrow> object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> kernel_state \<Rightarrow> bool"
where
"createObject_hs_preconds regionBase newType userSize d \<equiv>
(invs' and pspace_no_overlap' regionBase (getObjectSize newType userSize)
Expand All @@ -4373,14 +4373,14 @@ abbreviation

(* these preconds actually used throughout the proof *)
abbreviation(input)
createObject_c_preconds1 :: "word32 \<Rightarrow> ArchTypes_H.object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> (globals myvars) set"
createObject_c_preconds1 :: "word32 \<Rightarrow> object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> (globals myvars) set"
where
"createObject_c_preconds1 regionBase newType userSize deviceMemory \<equiv>
{s. region_actually_is_dev_bytes regionBase (2 ^ getObjectSize newType userSize) deviceMemory s}"

(* these preconds used at start of proof *)
definition
createObject_c_preconds :: "word32 \<Rightarrow> ArchTypes_H.object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> (globals myvars) set"
createObject_c_preconds :: "word32 \<Rightarrow> object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> (globals myvars) set"
where
"createObject_c_preconds regionBase newType userSize deviceMemory \<equiv>
(createObject_c_preconds1 regionBase newType userSize deviceMemory
Expand Down
4 changes: 2 additions & 2 deletions proof/crefine/ARM_HYP/Invoke_C.thy
Original file line number Diff line number Diff line change
Expand Up @@ -2828,10 +2828,10 @@ lemma ctes_of_ex_cte_cap_to':


lemma Arch_isFrameType_spec:
"\<forall>s. \<Gamma> \<turnstile> \<lbrace>s. unat \<acute>type \<le> fromEnum (maxBound::ArchTypes_H.object_type)\<rbrace>
"\<forall>s. \<Gamma> \<turnstile> \<lbrace>s. unat \<acute>type \<le> fromEnum (maxBound::object_type)\<rbrace>
Call Arch_isFrameType_'proc
\<lbrace> \<acute>ret__unsigned_long =
from_bool (isFrameType ((toEnum (unat \<^bsup>s\<^esup> type))::ArchTypes_H.object_type))\<rbrace>"
from_bool (isFrameType ((toEnum (unat \<^bsup>s\<^esup> type))::object_type))\<rbrace>"
apply vcg
apply (simp add: toEnum_object_type_to_H)
apply (frule object_type_from_to_H)
Expand Down
6 changes: 3 additions & 3 deletions proof/crefine/ARM_HYP/Retype_C.thy
Original file line number Diff line number Diff line change
Expand Up @@ -4919,7 +4919,7 @@ lemma placeNewObject_user_data:


definition
createObject_hs_preconds :: "word32 \<Rightarrow> ArchTypes_H.object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> kernel_state \<Rightarrow> bool"
createObject_hs_preconds :: "word32 \<Rightarrow> object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> kernel_state \<Rightarrow> bool"
where
"createObject_hs_preconds regionBase newType userSize d \<equiv>
(invs' and pspace_no_overlap' regionBase (getObjectSize newType userSize)
Expand All @@ -4942,14 +4942,14 @@ abbreviation

(* these preconds actually used throughout the proof *)
abbreviation(input)
createObject_c_preconds1 :: "word32 \<Rightarrow> ArchTypes_H.object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> (globals myvars) set"
createObject_c_preconds1 :: "word32 \<Rightarrow> object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> (globals myvars) set"
where
"createObject_c_preconds1 regionBase newType userSize deviceMemory \<equiv>
{s. region_actually_is_dev_bytes regionBase (2 ^ getObjectSize newType userSize) deviceMemory s}"

(* these preconds used at start of proof *)
definition
createObject_c_preconds :: "word32 \<Rightarrow> ArchTypes_H.object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> (globals myvars) set"
createObject_c_preconds :: "word32 \<Rightarrow> object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> (globals myvars) set"
where
"createObject_c_preconds regionBase newType userSize deviceMemory \<equiv>
(createObject_c_preconds1 regionBase newType userSize deviceMemory
Expand Down
4 changes: 2 additions & 2 deletions proof/crefine/RISCV64/Invoke_C.thy
Original file line number Diff line number Diff line change
Expand Up @@ -2774,10 +2774,10 @@ lemma ctes_of_ex_cte_cap_to':


lemma Arch_isFrameType_spec:
"\<forall>s. \<Gamma> \<turnstile> \<lbrace>s. unat \<acute>type \<le> fromEnum (maxBound::ArchTypes_H.object_type)\<rbrace>
"\<forall>s. \<Gamma> \<turnstile> \<lbrace>s. unat \<acute>type \<le> fromEnum (maxBound::object_type)\<rbrace>
Call Arch_isFrameType_'proc
\<lbrace> \<acute>ret__unsigned_long =
from_bool (isFrameType ((toEnum (unat \<^bsup>s\<^esup> type))::ArchTypes_H.object_type))\<rbrace>"
from_bool (isFrameType ((toEnum (unat \<^bsup>s\<^esup> type))::object_type))\<rbrace>"
apply vcg
apply (simp add: toEnum_object_type_to_H)
apply (frule object_type_from_to_H)
Expand Down
6 changes: 3 additions & 3 deletions proof/crefine/RISCV64/Retype_C.thy
Original file line number Diff line number Diff line change
Expand Up @@ -4796,7 +4796,7 @@ lemma placeNewObject_user_data:
done

definition
createObject_hs_preconds :: "machine_word \<Rightarrow> ArchTypes_H.object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> kernel_state \<Rightarrow> bool"
createObject_hs_preconds :: "machine_word \<Rightarrow> object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> kernel_state \<Rightarrow> bool"
where
"createObject_hs_preconds regionBase newType userSize d \<equiv>
(invs' and pspace_no_overlap' regionBase (getObjectSize newType userSize)
Expand All @@ -4819,14 +4819,14 @@ abbreviation

(* these preconds actually used throughout the proof *)
abbreviation(input)
createObject_c_preconds1 :: "machine_word \<Rightarrow> ArchTypes_H.object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> (globals myvars) set"
createObject_c_preconds1 :: "machine_word \<Rightarrow> object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> (globals myvars) set"
where
"createObject_c_preconds1 regionBase newType userSize deviceMemory \<equiv>
{s. region_actually_is_dev_bytes regionBase (2 ^ getObjectSize newType userSize) deviceMemory s}"

(* these preconds used at start of proof *)
definition
createObject_c_preconds :: "machine_word \<Rightarrow> ArchTypes_H.object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> (globals myvars) set"
createObject_c_preconds :: "machine_word \<Rightarrow> object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> (globals myvars) set"
where
"createObject_c_preconds regionBase newType userSize deviceMemory \<equiv>
(createObject_c_preconds1 regionBase newType userSize deviceMemory
Expand Down
4 changes: 2 additions & 2 deletions proof/crefine/X64/Invoke_C.thy
Original file line number Diff line number Diff line change
Expand Up @@ -2800,10 +2800,10 @@ lemma ctes_of_ex_cte_cap_to':


lemma Arch_isFrameType_spec:
"\<forall>s. \<Gamma> \<turnstile> \<lbrace>s. unat \<acute>type \<le> fromEnum (maxBound::ArchTypes_H.object_type)\<rbrace>
"\<forall>s. \<Gamma> \<turnstile> \<lbrace>s. unat \<acute>type \<le> fromEnum (maxBound::object_type)\<rbrace>
Call Arch_isFrameType_'proc
\<lbrace> \<acute>ret__unsigned_long =
from_bool (isFrameType ((toEnum (unat \<^bsup>s\<^esup> type))::ArchTypes_H.object_type))\<rbrace>"
from_bool (isFrameType ((toEnum (unat \<^bsup>s\<^esup> type))::object_type))\<rbrace>"
apply vcg
apply (simp add: toEnum_object_type_to_H)
apply (frule object_type_from_to_H)
Expand Down
6 changes: 3 additions & 3 deletions proof/crefine/X64/Retype_C.thy
Original file line number Diff line number Diff line change
Expand Up @@ -5628,7 +5628,7 @@ lemma placeNewObject_user_data:
done

definition
createObject_hs_preconds :: "machine_word \<Rightarrow> ArchTypes_H.object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> kernel_state \<Rightarrow> bool"
createObject_hs_preconds :: "machine_word \<Rightarrow> object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> kernel_state \<Rightarrow> bool"
where
"createObject_hs_preconds regionBase newType userSize d \<equiv>
(invs' and pspace_no_overlap' regionBase (getObjectSize newType userSize)
Expand All @@ -5651,14 +5651,14 @@ abbreviation

(* these preconds actually used throughout the proof *)
abbreviation(input)
createObject_c_preconds1 :: "machine_word \<Rightarrow> ArchTypes_H.object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> (globals myvars) set"
createObject_c_preconds1 :: "machine_word \<Rightarrow> object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> (globals myvars) set"
where
"createObject_c_preconds1 regionBase newType userSize deviceMemory \<equiv>
{s. region_actually_is_dev_bytes regionBase (2 ^ getObjectSize newType userSize) deviceMemory s}"

(* these preconds used at start of proof *)
definition
createObject_c_preconds :: "machine_word \<Rightarrow> ArchTypes_H.object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> (globals myvars) set"
createObject_c_preconds :: "machine_word \<Rightarrow> object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> (globals myvars) set"
where
"createObject_c_preconds regionBase newType userSize deviceMemory \<equiv>
(createObject_c_preconds1 regionBase newType userSize deviceMemory
Expand Down
2 changes: 1 addition & 1 deletion proof/refine/AARCH64/Syscall_R.thy
Original file line number Diff line number Diff line change
Expand Up @@ -535,7 +535,7 @@ crunch InterruptDecls_H.invokeIRQHandler
for typ_at'[wp]: "\<lambda>s. P (typ_at' T p s)"

lemmas invokeIRQHandler_typ_ats[wp] =
typ_at_lifts [OF InterruptDecls_H_invokeIRQHandler_typ_at']
typ_at_lifts [OF invokeIRQHandler_typ_at']

crunch setDomain
for tcb_at'[wp]: "tcb_at' tptr"
Expand Down
2 changes: 1 addition & 1 deletion proof/refine/ARM/Syscall_R.thy
Original file line number Diff line number Diff line change
Expand Up @@ -526,7 +526,7 @@ crunch InterruptDecls_H.invokeIRQHandler
for typ_at'[wp]: "\<lambda>s. P (typ_at' T p s)"

lemmas invokeIRQHandler_typ_ats[wp] =
typ_at_lifts [OF InterruptDecls_H_invokeIRQHandler_typ_at']
typ_at_lifts [OF invokeIRQHandler_typ_at']

crunch setDomain
for tcb_at'[wp]: "tcb_at' tptr"
Expand Down
8 changes: 4 additions & 4 deletions proof/refine/ARM/orphanage/Orphanage.thy
Original file line number Diff line number Diff line change
Expand Up @@ -1797,11 +1797,11 @@ lemma invokeIRQControl_no_orphans [wp]:
apply (wp | clarsimp)+
done

lemma invokeIRQHandler_no_orphans [wp]:
lemma arch_invokeIRQHandler_no_orphans[wp]:
"\<lbrace> \<lambda>s. no_orphans s \<and> invs' s \<rbrace>
invokeIRQHandler i
ARM_H.invokeIRQHandler i
\<lbrace> \<lambda>reply s. no_orphans s \<rbrace>"
apply (cases i, simp_all add: invokeIRQHandler_def)
apply (cases i, simp_all add: ARM_H.invokeIRQHandler_def)
apply (wp | clarsimp | fastforce)+
done

Expand Down Expand Up @@ -1939,7 +1939,7 @@ lemma setDomain_no_orphans [wp]:
apply (fastforce simp: tcb_at_typ_at' is_active_tcb_ptr_runnable')
done

crunch InterruptDecls_H.invokeIRQHandler
crunch invokeIRQHandler
for no_orphans[wp]: no_orphans

lemma performInvocation_no_orphans [wp]:
Expand Down
2 changes: 1 addition & 1 deletion proof/refine/ARM_HYP/Syscall_R.thy
Original file line number Diff line number Diff line change
Expand Up @@ -536,7 +536,7 @@ crunch InterruptDecls_H.invokeIRQHandler
for typ_at'[wp]: "\<lambda>s. P (typ_at' T p s)"

lemmas invokeIRQHandler_typ_ats[wp] =
typ_at_lifts [OF InterruptDecls_H_invokeIRQHandler_typ_at']
typ_at_lifts [OF invokeIRQHandler_typ_at']

crunch setDomain
for tcb_at'[wp]: "tcb_at' tptr"
Expand Down
2 changes: 1 addition & 1 deletion proof/refine/RISCV64/Syscall_R.thy
Original file line number Diff line number Diff line change
Expand Up @@ -534,7 +534,7 @@ crunch InterruptDecls_H.invokeIRQHandler
for typ_at'[wp]: "\<lambda>s. P (typ_at' T p s)"

lemmas invokeIRQHandler_typ_ats[wp] =
typ_at_lifts [OF InterruptDecls_H_invokeIRQHandler_typ_at']
typ_at_lifts [OF invokeIRQHandler_typ_at']

crunch setDomain
for tcb_at'[wp]: "tcb_at' tptr"
Expand Down
2 changes: 1 addition & 1 deletion proof/refine/X64/Syscall_R.thy
Original file line number Diff line number Diff line change
Expand Up @@ -535,7 +535,7 @@ crunch InterruptDecls_H.invokeIRQHandler
for typ_at'[wp]: "\<lambda>s. P (typ_at' T p s)"

lemmas invokeIRQHandler_typ_ats[wp] =
typ_at_lifts [OF InterruptDecls_H_invokeIRQHandler_typ_at']
typ_at_lifts [OF invokeIRQHandler_typ_at']

crunch setDomain
for tcb_at'[wp]: "tcb_at' tptr"
Expand Down

0 comments on commit e152c9e

Please sign in to comment.