mirror of https://github.com/seL4/l4v.git
release cleanup
This commit is contained in:
parent
2a03e81df4
commit
84595f4233
|
@ -0,0 +1,28 @@
|
|||
*.pyc
|
||||
*.o
|
||||
|
||||
spec/cspec/c/api/
|
||||
spec/cspec/c/arch/
|
||||
spec/cspec/c/kernel_all.c
|
||||
spec/cspec/c/kernel_all.c_pp
|
||||
spec/cspec/c/parsetab.py
|
||||
spec/cspec/c/plat/
|
||||
spec/cspec/c/sources_list_updated
|
||||
spec/umm_types.txt
|
||||
|
||||
tools/autocorres/doc/quickstart/umm_types.txt
|
||||
tools/autocorres/sel4.txt
|
||||
tools/autocorres/tests/ROOT
|
||||
tools/autocorres/tests/parse-tests/*.thy
|
||||
tools/autocorres/tests/umm_types.txt
|
||||
|
||||
tools/c-parser/standalone-parser/table.ML
|
||||
tools/c-parser/testfiles/ROOT
|
||||
tools/c-parser/testfiles/umm_types.txt
|
||||
|
||||
tools/haskell-translator/caseconvs-useful
|
||||
|
||||
camkes/adl-spec/camkes.ML
|
||||
camkes/glue-proofs/umm_types.txt
|
||||
|
||||
internal
|
|
@ -9,20 +9,19 @@
|
|||
#
|
||||
|
||||
## Targets
|
||||
images: Camkes CamkesARMModel
|
||||
images: Camkes
|
||||
default: images test
|
||||
test:
|
||||
all: images test
|
||||
report-regression:
|
||||
@echo Camkes CamkesARMModel
|
||||
@echo Camkes
|
||||
|
||||
#
|
||||
# Setup heaps.
|
||||
#
|
||||
|
||||
# CAmkES
|
||||
HEAPS += CamkesAdlSpec CamkesAutomaton CamkesGlueSpec CamkesCode CamkesARMModel \
|
||||
CamkesGlueProofs
|
||||
HEAPS += CamkesAdlSpec CamkesGlueSpec CamkesGlueProofs
|
||||
GROUPS += Camkes
|
||||
|
||||
# Clean
|
||||
|
|
|
@ -12,20 +12,5 @@ CAmkES is a component platform for seL4. This directory contains files related
|
|||
to a formal Isabelle model of CAmkES.
|
||||
|
||||
adl-spec/ - Architectural model.
|
||||
code/ - AutoCorres-based work (bottom-up approach to glue code).
|
||||
glue-proofs/ - AutoCorres-based work (bottom-up approach to glue code).
|
||||
glue-spec/ - Behavioural model (top-down approach to glue code).
|
||||
misc/ - What it says on the box.
|
||||
|
||||
FIXME:
|
||||
- I removed CAmkES exceptions from the model. Not sure if they really make
|
||||
sense, but if they do they'll need to be re-introduced.
|
||||
- The explicit concept of a connector is absent from the types. It may
|
||||
help to add this later.
|
||||
- Unidirectional shared memory communication and more than two parties should
|
||||
be supported.
|
||||
- Components should be allowed to have a configuration as well.
|
||||
- Attributes need some work. They basically have no type information at the
|
||||
implementation level which is a bit of a minefield.
|
||||
- There needs to be a nice way of hiding arch-specific connectors. Perhaps
|
||||
this should be generalised into a way of fail-stopping assemblies that
|
||||
involve platform connectors that are unavailable.
|
||||
|
|
|
@ -58,13 +58,6 @@ lemma "(CommunicationBi s f t \<longrightarrow> \<not> (CommunicationNone s f t)
|
|||
(CommunicationNone s f t \<longrightarrow> \<not> (CommunicationBi s f t))"
|
||||
by (clarsimp simp:CommunicationBi_def CommunicationNone_def)
|
||||
|
||||
(* FIXME: Even with this lemma, showing information flow properties
|
||||
* on a concrete input specification is tantalisingly out of reach
|
||||
* of eval. I'm not entirely sure why as eval can operate over all
|
||||
* the defs, but still can't show something that can be done with
|
||||
* unfolding and the axioms. Perhaps the axiom gets in the way
|
||||
* despite this lemma's attempt to conceal it.
|
||||
*)
|
||||
lemma blah[code]:"communication_mode = (\<lambda>s. if s = seL4Asynch then Outgoing else communication_mode s)"
|
||||
by (rule ext, clarsimp simp:seL4AsynchMode)
|
||||
|
||||
|
@ -73,4 +66,4 @@ code_const communication_mode
|
|||
(SML "!(raise/ Fail/ \"undefined\")")
|
||||
export_code Communication in SML module_name "Camkes" file "/dev/null"
|
||||
|
||||
end
|
||||
end
|
||||
|
|
|
@ -9,19 +9,14 @@
|
|||
%
|
||||
|
||||
\chapter{Introduction}
|
||||
% FIXME: This needs a cleanup and extension to further describe CAmkES. Lift
|
||||
% from CAmkES paper/intro?
|
||||
CAmkES is a component platform for embedded microkernel-based systems, offering
|
||||
many of the standard features available in component platforms.
|
||||
% TODO: Maybe reference that component book here.
|
||||
Some relevant features of CAmkES that are not common to all component platforms
|
||||
are:
|
||||
\begin{itemize}
|
||||
\item \emph{Explicit composite components.} CAmkES components can be
|
||||
assembled to form a re-usable composite, that can then be referenced within
|
||||
a containing system.
|
||||
% FIXME: This is listed in future work, but also discussed in relation to export
|
||||
% connectors. Needs to be made consistent.
|
||||
\item \emph{Multiple instantiation of a single component.} Multiple copies of
|
||||
a component can exist within a system, distinguished by different
|
||||
identifiers.
|
||||
|
|
|
@ -118,30 +118,8 @@
|
|||
\renewcommand{\isamarkupsubsubsection}[1]{\subsection{#1}}
|
||||
|
||||
\newcommand{\nictafundingacknowledgement}{%
|
||||
% Lifted from http://wiki.inside.nicta.com.au/display/CHNOPSR/Funding+Acknowledgement 28-10-2013
|
||||
NICTA is funded by the Australian Government through the Department of Communications and the Australian Research Council through the ICT Centre of Excellence Program. NICTA is also funded and supported by the Australian Capital Territory, the New South Wales, Queensland and Victorian Governments, the Australian National University, the University of New South Wales, the University of Melbourne, the University of Queensland, the University of Sydney, Griffith University, Queensland University of Technology, Monash University and other university partners.}
|
||||
|
||||
\newcommand{\govrights}{%
|
||||
The Government may use, modify, reproduce, release, perform, display or
|
||||
disclose these data within the Government without restriction, and may release
|
||||
or disclose outside the Government and authorize persons to whom such release
|
||||
or disclosure has been made to use, modify, reproduce, release, perform,
|
||||
display or disclose that data for United States Government purposes, including
|
||||
competitive procurement.}
|
||||
|
||||
\newcommand{\ABN}{ABN 62 102 206 173}
|
||||
\newcommand{\agreement}{Agreement No.: FA8750-12-9-0179}
|
||||
\newcommand{\addr}{Level 5, 13 Garden Street, Eveleigh, New South Wales,
|
||||
Australia}
|
||||
|
||||
\newcommand{\cpright}{Copyright \copyright\ 2013 NICTA, \ABN. All rights reserved except those specified herein.}
|
||||
\newcommand{\disclaimer}{%
|
||||
\cpright\\
|
||||
Government Purpose Rights\\
|
||||
\agreement;
|
||||
Recipient's Name: NICTA;
|
||||
Recipient's Address: \addr.\\
|
||||
\govrights}
|
||||
|
||||
\newcommand{\trdisclaimer}{%
|
||||
This material is based on research sponsored by Air Force Research Laboratory
|
||||
|
@ -157,12 +135,10 @@ the Defense Advanced Research Projects Agency or the U.S.Government.}
|
|||
|
||||
\newcommand{\smalldisclaimer}{}
|
||||
\newcommand{\bigdisclaimer}{%
|
||||
\nictafundingacknowledgement\\
|
||||
|
||||
\cpright\\
|
||||
\nictafundingacknowledgement
|
||||
|
||||
\vspace{2ex}
|
||||
\noindent\trdisclaimer}
|
||||
\trdisclaimer}
|
||||
|
||||
|
||||
\newcommand{\pgstyle}{%
|
||||
|
|
|
@ -121,50 +121,22 @@
|
|||
\newcommand{\nictafundingacknowledgement}{%
|
||||
NICTA is funded by the Australian Government through the Department of Communications and the Australian Research Council through the ICT Centre of Excellence Program. NICTA is also funded and supported by the Australian Capital Territory, the New South Wales, Queensland and Victorian Governments, the Australian National University, the University of New South Wales, the University of Melbourne, the University of Queensland, the University of Sydney, Griffith University, Queensland University of Technology, Monash University and other university partners.}
|
||||
|
||||
\newcommand{\govrights}{%
|
||||
The Government may use, modify, reproduce, release, perform, display or
|
||||
disclose these data within the Government without restriction, and may release
|
||||
or disclose outside the Government and authorize persons to whom such release
|
||||
or disclosure has been made to use, modify, reproduce, release, perform,
|
||||
display or disclose that data for United States Government purposes, including
|
||||
competitive procurement.}
|
||||
|
||||
\newcommand{\ABN}{ABN 62 102 206 173}
|
||||
\newcommand{\agreement}{Agreement No.: FA8750-12-9-0179}
|
||||
\newcommand{\addr}{Level 5, 13 Garden Street, Eveleigh, New South Wales,
|
||||
Australia}
|
||||
|
||||
\newcommand{\cpright}{Copyright \copyright\ 2014 NICTA, \ABN. All rights reserved except those specified herein.}
|
||||
\newcommand{\disclaimer}{%
|
||||
\cpright\\
|
||||
Government Purpose Rights\\
|
||||
\agreement;
|
||||
Recipient's Name: NICTA;
|
||||
Recipient's Address: \addr.\\
|
||||
\govrights}
|
||||
\newcommand{\cpright}{Copyright \copyright\ 2014 NICTA, \ABN. All rights reserved.}
|
||||
|
||||
\newcommand{\smalldisclaimer}{\parbox{0.9\textwidth}{\tiny\disclaimer}}
|
||||
\newcommand{\bigdisclaimer}{%
|
||||
\nictafundingacknowledgement\\
|
||||
|
||||
\cpright
|
||||
|
||||
\vspace{2ex}
|
||||
Government Purpose Rights
|
||||
|
||||
\vspace{2ex}
|
||||
\agreement\\
|
||||
Recipient's Name: NICTA\\
|
||||
Recipient's Address: \addr
|
||||
|
||||
\vspace{2ex}
|
||||
\noindent\govrights}
|
||||
\cpright}
|
||||
|
||||
\newcommand{\pgstyle}{%
|
||||
\fancyhf{}%
|
||||
\renewcommand{\headrulewidth}{0pt}%
|
||||
\fancyfoot[C]{}%
|
||||
\fancyfoot[L]{\smalldisclaimer}%
|
||||
\fancyfoot[L]{}%
|
||||
\fancyfoot[R]{\sl\thepage}}
|
||||
|
||||
\fancypagestyle{plain}{\pgstyle}
|
||||
|
@ -194,8 +166,7 @@ Recipient's Address: \addr
|
|||
{\large \authors
|
||||
|
||||
\vspace{2ex}
|
||||
May 2014 \\
|
||||
SMACCM Deliverable 4.2.2.1.1}
|
||||
May 2014}
|
||||
|
||||
\vfill
|
||||
{\small
|
||||
|
|
|
@ -1,12 +0,0 @@
|
|||
#
|
||||
# Copyright 2014, NICTA
|
||||
#
|
||||
# This software may be distributed and modified according to the terms of
|
||||
# the GNU General Public License version 2. Note that NO WARRANTY is provided.
|
||||
# See "LICENSE_GPLv2.txt" for details.
|
||||
#
|
||||
# @TAG(NICTA_GPL)
|
||||
#
|
||||
|
||||
This directory contains theories related to modelling an abstraction of the
|
||||
CAmkES glue code in a message passing semantics.
|
|
@ -1,26 +0,0 @@
|
|||
(*
|
||||
* Copyright 2014, NICTA
|
||||
*
|
||||
* This software may be distributed and modified according to the terms of
|
||||
* the GNU General Public License version 2. Note that NO WARRANTY is provided.
|
||||
* See "LICENSE_GPLv2.txt" for details.
|
||||
*
|
||||
* @TAG(NICTA_GPL)
|
||||
*)
|
||||
|
||||
(* This ROOT file is not intended for use in l4.verified. It has been extracted
|
||||
* from the general CAmkES ROOT for the SMACCM project.
|
||||
*)
|
||||
|
||||
session CamkesGlueModel = HOL +
|
||||
theories
|
||||
Abbreviations
|
||||
CIMP
|
||||
Connector
|
||||
Types
|
||||
UserStubs
|
||||
"example-procedure/GenSimpleSystem"
|
||||
"example-event/GenEventSystem"
|
||||
"example-dataport/GenDataportSystem"
|
||||
"example-untrusted/EgTop"
|
||||
"example-trusted/EgTop2"
|
|
@ -121,27 +121,11 @@
|
|||
% Lifted from http://wiki.inside.nicta.com.au/display/CHNOPSR/Funding+Acknowledgement 28-10-2013
|
||||
NICTA is funded by the Australian Government through the Department of Communications and the Australian Research Council through the ICT Centre of Excellence Program. NICTA is also funded and supported by the Australian Capital Territory, the New South Wales, Queensland and Victorian Governments, the Australian National University, the University of New South Wales, the University of Melbourne, the University of Queensland, the University of Sydney, Griffith University, Queensland University of Technology, Monash University and other university partners.}
|
||||
|
||||
\newcommand{\govrights}{%
|
||||
The Government may use, modify, reproduce, release, perform, display or
|
||||
disclose these data within the Government without restriction, and may release
|
||||
or disclose outside the Government and authorize persons to whom such release
|
||||
or disclosure has been made to use, modify, reproduce, release, perform,
|
||||
display or disclose that data for United States Government purposes, including
|
||||
competitive procurement.}
|
||||
|
||||
\newcommand{\ABN}{ABN 62 102 206 173}
|
||||
\newcommand{\agreement}{Agreement No.: FA8750-12-9-0179}
|
||||
\newcommand{\addr}{Level 5, 13 Garden Street, Eveleigh, New South Wales,
|
||||
Australia}
|
||||
|
||||
\newcommand{\cpright}{Copyright \copyright\ 2013 NICTA, \ABN. All rights reserved except those specified herein.}
|
||||
\newcommand{\cpright}{Copyright \copyright\ 2013 NICTA, \ABN. All rights reserved.}
|
||||
\newcommand{\disclaimer}{%
|
||||
\cpright\\
|
||||
Government Purpose Rights\\
|
||||
\agreement;
|
||||
Recipient's Name: NICTA;
|
||||
Recipient's Address: \addr.\\
|
||||
\govrights}
|
||||
\cpright}
|
||||
|
||||
\newcommand{\trdisclaimer}{%
|
||||
This material is based on research sponsored by Air Force Research Laboratory
|
||||
|
|
|
@ -28,6 +28,16 @@ where
|
|||
|
||||
text {* Base case facts about correspondence *}
|
||||
|
||||
lemma corres_underlyingD:
|
||||
"\<lbrakk> corres_underlying R z rs P P' f f'; (s,s') \<in> R; P s; P' s' \<rbrakk>
|
||||
\<Longrightarrow> (\<forall>(r',t')\<in>fst (f' s'). \<exists>(r,t)\<in>fst (f s). (t, t') \<in> R \<and> rs r r') \<and> (z \<longrightarrow> \<not> snd (f' s'))"
|
||||
by (fastforce simp: corres_underlying_def)
|
||||
|
||||
lemma corres_underlyingD2:
|
||||
"\<lbrakk> corres_underlying R z rs P P' f f'; (s,s') \<in> R; P s; P' s'; (r',t')\<in>fst (f' s') \<rbrakk>
|
||||
\<Longrightarrow> \<exists>(r,t)\<in>fst (f s). (t, t') \<in> R \<and> rs r r'"
|
||||
by (fastforce dest: corres_underlyingD)
|
||||
|
||||
lemma propagate_no_fail:
|
||||
"\<lbrakk> corres_underlying S True R P P' f f';
|
||||
no_fail P f; \<forall>s'. P' s' \<longrightarrow> (\<exists>s. P s \<and> (s,s') \<in> S) \<rbrakk>
|
||||
|
@ -472,6 +482,14 @@ lemma corres_weaker_disj_division:
|
|||
apply simp+
|
||||
done
|
||||
|
||||
lemma corres_symmetric_bool_cases:
|
||||
"\<lbrakk> P = P'; \<lbrakk> P; P' \<rbrakk> \<Longrightarrow> corres_underlying srel nf r Q Q' f g;
|
||||
\<lbrakk> \<not> P; \<not> P' \<rbrakk> \<Longrightarrow> corres_underlying srel nf r R R' f g \<rbrakk>
|
||||
\<Longrightarrow> corres_underlying srel nf r (\<lambda>s. (P \<longrightarrow> Q s) \<and> (\<not> P \<longrightarrow> R s))
|
||||
(\<lambda>s. (P' \<longrightarrow> Q' s) \<and> (\<not> P' \<longrightarrow> R' s))
|
||||
f g"
|
||||
by (cases P, simp_all)
|
||||
|
||||
text {* Support for symbolically executing into the guards
|
||||
and manipulating them *}
|
||||
|
||||
|
@ -630,6 +648,13 @@ lemma corres_assert_assume:
|
|||
by (auto simp: bind_def assert_def fail_def return_def
|
||||
corres_underlying_def)
|
||||
|
||||
lemma corres_state_assert:
|
||||
"corres_underlying sr nf rr P Q f (g ()) \<Longrightarrow>
|
||||
(\<And>s. Q s \<Longrightarrow> R s) \<Longrightarrow>
|
||||
corres_underlying sr nf rr P Q f (state_assert R >>= g)"
|
||||
by (clarsimp simp: corres_underlying_def state_assert_def get_def assert_def
|
||||
return_def bind_def)
|
||||
|
||||
lemma corres_stateAssert_assume:
|
||||
"\<lbrakk> corres_underlying sr nf r P Q f (g ()); \<And>s. Q s \<Longrightarrow> P' s \<rbrakk> \<Longrightarrow>
|
||||
corres_underlying sr nf r P Q f (stateAssert P' [] >>= g)"
|
||||
|
@ -1001,6 +1026,25 @@ lemma corres_either_alternate:
|
|||
apply simp
|
||||
done
|
||||
|
||||
lemma corres_either_alternate2:
|
||||
"\<lbrakk> corres_underlying sr nf r P R a c; corres_underlying sr nf r Q R b c \<rbrakk>
|
||||
\<Longrightarrow> corres_underlying sr nf r (P or Q) R (a \<sqinter> b) c"
|
||||
apply (simp add: corres_underlying_def alternative_def)
|
||||
apply clarsimp
|
||||
apply (drule (1) bspec, clarsimp)+
|
||||
apply (erule disjE)
|
||||
apply clarsimp
|
||||
apply (drule(1) bspec, clarsimp)
|
||||
apply (rule rev_bexI)
|
||||
apply (erule UnI1)
|
||||
apply simp
|
||||
apply clarsimp
|
||||
apply (drule(1) bspec, clarsimp)
|
||||
apply (rule rev_bexI)
|
||||
apply (erule UnI2)
|
||||
apply simp
|
||||
done
|
||||
|
||||
lemma option_corres:
|
||||
assumes "x = None \<Longrightarrow> corres_underlying sr nf r P P' (A None) (C None)"
|
||||
assumes "\<And>z. x = Some z \<Longrightarrow> corres_underlying sr nf r (Q z) (Q' z) (A (Some z)) (C (Some z))"
|
||||
|
@ -1009,6 +1053,11 @@ lemma option_corres:
|
|||
(A x) (C x)"
|
||||
by (cases x) (auto simp: assms)
|
||||
|
||||
lemma corres_bind_return:
|
||||
"corres_underlying sr nf r P P' (f >>= return) g \<Longrightarrow>
|
||||
corres_underlying sr nf r P P' f g"
|
||||
by (simp add: corres_underlying_def)
|
||||
|
||||
lemma corres_bind_return2:
|
||||
"corres_underlying sr nf r P P' f (g >>= return) \<Longrightarrow> corres_underlying sr nf r P P' f g"
|
||||
by simp
|
||||
|
|
|
@ -169,8 +169,8 @@ val crunchP =
|
|||
(fn lthy =>
|
||||
(case get_crunch_instance crunch_instance (Proof_Context.theory_of lthy) of
|
||||
NONE => error ("Crunch has not been defined for " ^ crunch_instance)
|
||||
| SOME (crunch_x, _) => crunch_x)
|
||||
att_srcs extra prp_name wpigs consts lthy)));
|
||||
| SOME (crunch_x, _) =>
|
||||
crunch_x att_srcs extra prp_name wpigs consts lthy))));
|
||||
|
||||
val add_sect = "add";
|
||||
val del_sect = "del";
|
||||
|
@ -189,12 +189,12 @@ val crunch_ignoreP =
|
|||
|> map (const_name o #2);
|
||||
val del = wpigs |> filter (fn (s,_) => s = del_sect)
|
||||
|> map (const_name o #2);
|
||||
in
|
||||
(case get_crunch_instance crunch_instance (Proof_Context.theory_of lthy) of
|
||||
val crunch_ignore_add_del = (case get_crunch_instance crunch_instance (Proof_Context.theory_of lthy) of
|
||||
NONE => error ("Crunch has not been defined for " ^ crunch_instance)
|
||||
| SOME (_, crunch_ignore_add_del) =>
|
||||
Local_Theory.raw_theory (crunch_ignore_add_del add del) lthy
|
||||
|> (fn lthy => Named_Target.reinit lthy lthy))
|
||||
| SOME x => snd x);
|
||||
in
|
||||
Local_Theory.raw_theory (crunch_ignore_add_del add del) lthy
|
||||
|> (fn lthy => Named_Target.reinit lthy lthy)
|
||||
end));
|
||||
|
||||
end;
|
||||
|
|
|
@ -219,6 +219,10 @@ lemma orthD1:
|
|||
lemma orthD2:
|
||||
"\<lbrakk> S \<inter> S' = {}; x \<in> S'\<rbrakk> \<Longrightarrow> x \<notin> S" by auto
|
||||
|
||||
lemma distinct_element:
|
||||
"\<lbrakk> b \<inter> d = {}; a \<in> b; c \<in> d\<rbrakk>\<Longrightarrow> a \<noteq> c"
|
||||
by auto
|
||||
|
||||
lemma ranE:
|
||||
"\<lbrakk> v \<in> ran f; \<And>x. f x = Some v \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
|
||||
by (auto simp: ran_def)
|
||||
|
@ -598,6 +602,10 @@ lemma union_trans:
|
|||
apply (blast intro: in_rtrancl_UnI rtrancl_trans)
|
||||
done
|
||||
|
||||
lemma trancl_trancl:
|
||||
"(R\<^sup>+)\<^sup>+ = R\<^sup>+"
|
||||
by auto
|
||||
|
||||
lemma if_1_0_0:
|
||||
"((if P then 1 else 0) = (0 :: ('a :: zero_neq_one))) = (\<not> P)"
|
||||
by (simp split: split_if)
|
||||
|
@ -616,6 +624,10 @@ lemma in_singleton:
|
|||
"S = {x} \<Longrightarrow> x \<in> S"
|
||||
by simp
|
||||
|
||||
lemma singleton_set:
|
||||
"x \<in> set [a] \<Longrightarrow> x = a"
|
||||
by auto
|
||||
|
||||
lemma take_drop_eqI:
|
||||
assumes t: "take n xs = take n ys"
|
||||
assumes d: "drop n xs = drop n ys"
|
||||
|
|
|
@ -435,4 +435,15 @@ next
|
|||
by (fastforce simp: corres_underlying_def split_def)
|
||||
qed
|
||||
|
||||
lemma wpc_helper_monadic_rewrite:
|
||||
"monadic_rewrite F E Q' m m'
|
||||
\<Longrightarrow> wpc_helper (P, P') (Q, {s. Q' s}) (monadic_rewrite F E (\<lambda>s. s \<in> P') m m')"
|
||||
apply (clarsimp simp: wpc_helper_def)
|
||||
apply (erule monadic_rewrite_imp)
|
||||
apply auto
|
||||
done
|
||||
|
||||
wpc_setup "\<lambda>m. monadic_rewrite F E Q' m m'" wpc_helper_monadic_rewrite
|
||||
wpc_setup "\<lambda>m. monadic_rewrite F E Q' (m >>= c) m'" wpc_helper_monadic_rewrite
|
||||
|
||||
end
|
||||
|
|
|
@ -206,6 +206,11 @@ lemma sequence_x_Cons: "\<And>x xs. sequence_x (x # xs) = (x >>= (\<lambda>_. se
|
|||
lemma mapM_Cons: "mapM m (x # xs) = (do y \<leftarrow> m x; ys \<leftarrow> (mapM m xs); return (y # ys) od)"
|
||||
by (simp add: mapM_def sequence_def Let_def)
|
||||
|
||||
lemma mapM_simps:
|
||||
"mapM m [] = return []"
|
||||
"mapM m (x#xs) = do r \<leftarrow> m x; rs \<leftarrow> (mapM m xs); return (r#rs) od"
|
||||
by (simp_all add: mapM_def sequence_def)
|
||||
|
||||
lemma zipWithM_x_mapM:
|
||||
"zipWithM_x f as bs = (mapM (split f) (zip as bs) >>= (\<lambda>_. return ()))"
|
||||
apply (simp add: zipWithM_x_def zipWith_def)
|
||||
|
@ -621,6 +626,14 @@ lemma no_fail_liftE [wp]:
|
|||
apply simp
|
||||
done
|
||||
|
||||
lemma bind_return_eq:
|
||||
"(a >>= return) = (b >>= return) \<Longrightarrow> a = b"
|
||||
apply (clarsimp simp:bind_def)
|
||||
apply (rule ext)
|
||||
apply (drule_tac x= x in fun_cong)
|
||||
apply (auto simp:return_def split_def)
|
||||
done
|
||||
|
||||
lemma bindE_bind_linearise:
|
||||
"((f >>=E g) >>= h) =
|
||||
(f >>= sum_case (h o Inl) (\<lambda>rv. g rv >>= h))"
|
||||
|
@ -899,6 +912,12 @@ lemma alternative_refl:
|
|||
"(a \<sqinter> a) = a"
|
||||
by (rule ext, simp add: alternative_def)
|
||||
|
||||
lemma alternative_com:
|
||||
"(f \<sqinter> g) = (g \<sqinter> f)"
|
||||
apply (rule ext)
|
||||
apply (auto simp: alternative_def)
|
||||
done
|
||||
|
||||
lemma liftE_alternative:
|
||||
"liftE (a \<sqinter> b) = (liftE a \<sqinter> liftE b)"
|
||||
by (simp add: liftE_def alternative_bind)
|
||||
|
|
|
@ -221,17 +221,6 @@ lemma neq_into_nprefixeq:
|
|||
"\<lbrakk> x \<noteq> take (length x) y \<rbrakk> \<Longrightarrow> \<not> x \<le> y"
|
||||
by (clarsimp simp: prefixeq_def less_eq_list_def)
|
||||
|
||||
lemma distinct_suffixeq:
|
||||
assumes dx: "distinct xs"
|
||||
and pf: "suffixeq ys xs"
|
||||
shows "distinct ys"
|
||||
using dx pf by (clarsimp elim!: suffixeqE)
|
||||
|
||||
lemma suffixeq_map:
|
||||
assumes pf: "suffixeq ys xs"
|
||||
shows "suffixeq (map f ys) (map f xs)"
|
||||
using pf by (auto elim!: suffixeqE intro: suffixeqI)
|
||||
|
||||
lemma suffixeq_drop [simp]:
|
||||
"suffixeq (drop n as) as"
|
||||
unfolding suffixeq_def
|
||||
|
@ -239,10 +228,6 @@ lemma suffixeq_drop [simp]:
|
|||
apply simp
|
||||
done
|
||||
|
||||
lemma suffixeq_take:
|
||||
"suffixeq ys xs \<Longrightarrow> xs = take (length xs - length ys) xs @ ys"
|
||||
by (clarsimp elim!: suffixeqE)
|
||||
|
||||
lemma suffixeq_eqI:
|
||||
"\<lbrakk> suffixeq xs as; suffixeq xs bs; length as = length bs;
|
||||
take (length as - length xs) as \<le> take (length bs - length xs) bs\<rbrakk> \<Longrightarrow> as = bs"
|
||||
|
@ -254,6 +239,10 @@ lemma suffixeq_Cons_mem:
|
|||
apply simp
|
||||
done
|
||||
|
||||
lemma distinct_imply_not_in_tail:
|
||||
"\<lbrakk> distinct list; suffixeq (y # ys) list\<rbrakk> \<Longrightarrow> y \<notin> set ys"
|
||||
by (clarsimp simp:suffixeq_def)
|
||||
|
||||
lemma list_induct_suffixeq [case_names Nil Cons]:
|
||||
assumes nilr: "P []"
|
||||
and consr: "\<And>x xs. \<lbrakk>P xs; suffixeq (x # xs) as \<rbrakk> \<Longrightarrow> P (x # xs)"
|
||||
|
@ -2105,6 +2094,12 @@ lemma is_aligned_neg_mask_eq:
|
|||
apply fastforce
|
||||
done
|
||||
|
||||
lemma is_aligned_shiftr_shiftl:
|
||||
"is_aligned w n \<Longrightarrow> w >> n << n = w"
|
||||
apply (simp add: shiftr_shiftl1)
|
||||
apply (erule is_aligned_neg_mask_eq)
|
||||
done
|
||||
|
||||
lemma rtrancl_insert:
|
||||
assumes x_new: "\<And>y. (x,y) \<notin> R"
|
||||
shows "R^* `` insert x S = insert x (R^* `` S)"
|
||||
|
@ -5893,4 +5888,170 @@ lemma word_rsplit_upt:
|
|||
apply (metis mult_commute given_quot_alt word_size word_size_gt_0)
|
||||
done
|
||||
|
||||
lemma aligned_shift:
|
||||
"\<lbrakk>x < 2 ^ n; is_aligned (y :: 'a :: len word) n;n \<le> len_of TYPE('a)\<rbrakk>
|
||||
\<Longrightarrow> x + y >> n = y >> n"
|
||||
apply (subst word_plus_and_or_coroll)
|
||||
apply (rule word_eqI)
|
||||
apply (clarsimp simp: is_aligned_nth)
|
||||
apply (drule(1) nth_bounded)
|
||||
apply simp
|
||||
apply simp
|
||||
apply (rule word_eqI)
|
||||
apply (simp add: nth_shiftr)
|
||||
apply safe
|
||||
apply (drule(1) nth_bounded)
|
||||
apply simp+
|
||||
done
|
||||
|
||||
lemma aligned_shift':
|
||||
"\<lbrakk>x < 2 ^ n; is_aligned (y :: 'a :: len word) n;n \<le> len_of TYPE('a)\<rbrakk>
|
||||
\<Longrightarrow> y + x >> n = y >> n"
|
||||
apply (subst word_plus_and_or_coroll)
|
||||
apply (rule word_eqI)
|
||||
apply (clarsimp simp: is_aligned_nth)
|
||||
apply (drule(1) nth_bounded)
|
||||
apply simp
|
||||
apply simp
|
||||
apply (rule word_eqI)
|
||||
apply (simp add: nth_shiftr)
|
||||
apply safe
|
||||
apply (drule(1) nth_bounded)
|
||||
apply simp+
|
||||
done
|
||||
|
||||
lemma neg_mask_add_mask:
|
||||
"((x:: 'a :: len word) && ~~ mask n) + (2 ^ n - 1) = x || mask n"
|
||||
apply (simp add:mask_2pm1[symmetric])
|
||||
apply (rule word_eqI)
|
||||
apply (rule iffI)
|
||||
apply (clarsimp simp:word_size not_less)
|
||||
apply (cut_tac w = "((x && ~~ mask n) + mask n)" and
|
||||
m = n and n = "na - n" in nth_shiftr[symmetric])
|
||||
apply clarsimp
|
||||
apply (subst (asm) aligned_shift')
|
||||
apply (simp add:mask_lt_2pn nth_shiftr is_aligned_neg_mask word_size word_bits_def )+
|
||||
apply (case_tac "na<n")
|
||||
apply clarsimp
|
||||
apply (subst word_plus_and_or_coroll)
|
||||
apply (rule iffD1[OF is_aligned_mask])
|
||||
apply (simp add:is_aligned_neg_mask word_or_nth word_size not_less)+
|
||||
apply (cut_tac w = "((x && ~~ mask n) + mask n)" and
|
||||
m = n and n = "na - n" in nth_shiftr[symmetric])
|
||||
apply clarsimp
|
||||
apply (subst (asm) aligned_shift')
|
||||
apply (simp add:mask_lt_2pn is_aligned_neg_mask word_bits_def nth_shiftr neg_mask_bang)+
|
||||
done
|
||||
|
||||
lemma subtract_mask:
|
||||
"p - (p && mask n) = (p && ~~ mask n)"
|
||||
"p - (p && ~~ mask n) = (p && mask n)"
|
||||
by (simp add: field_simps word_plus_and_or_coroll2)+
|
||||
|
||||
lemma and_neg_mask_plus_mask_mono: "(p && ~~ mask n) + mask n \<ge> p"
|
||||
apply (subst add_commute)
|
||||
apply (rule word_le_minus_cancel[where x = "p && ~~ mask n"])
|
||||
apply (clarsimp simp: subtract_mask)
|
||||
using word_and_le1[where a = "mask n" and y = p]
|
||||
apply (clarsimp simp: mask_def word_le_less_eq)
|
||||
apply (subst add_commute)
|
||||
apply (rule is_aligned_no_overflow'[folded mask_2pm1])
|
||||
apply (clarsimp simp: is_aligned_neg_mask)
|
||||
done
|
||||
|
||||
lemma word_neg_and_le:
|
||||
"ptr \<le> (ptr && ~~ mask n) + (2 ^ n - 1)"
|
||||
by (simp add: and_neg_mask_plus_mask_mono mask_2pm1[symmetric])
|
||||
|
||||
lemma aligned_less_plus_1:
|
||||
"\<lbrakk> is_aligned x n; n > 0 \<rbrakk> \<Longrightarrow> x < x + 1"
|
||||
apply (rule plus_one_helper2)
|
||||
apply (rule order_refl)
|
||||
apply (clarsimp simp: field_simps)
|
||||
apply (drule arg_cong[where f="\<lambda>x. x - 1"])
|
||||
apply (clarsimp simp: is_aligned_mask)
|
||||
apply (drule word_eqD[where x=0])
|
||||
apply simp
|
||||
done
|
||||
|
||||
lemma aligned_add_offset_less:
|
||||
"\<lbrakk>is_aligned x n; is_aligned y n; x < y; z < 2 ^ n\<rbrakk> \<Longrightarrow> x + z < y"
|
||||
apply (cases "y = 0")
|
||||
apply simp
|
||||
apply (erule is_aligned_get_word_bits[where p=y], simp_all)
|
||||
apply (cases "z = 0", simp_all)
|
||||
apply (drule(2) aligned_at_least_t2n_diff[rotated -1])
|
||||
apply (drule plus_one_helper2)
|
||||
apply (rule less_is_non_zero_p1)
|
||||
apply (rule aligned_less_plus_1)
|
||||
apply (erule aligned_sub_aligned[OF _ _ order_refl],
|
||||
simp_all add: is_aligned_triv)[1]
|
||||
apply (cases n, simp_all)[1]
|
||||
apply (simp only: trans[OF diff_add_eq diff_diff_eq2[symmetric]])
|
||||
apply (drule word_less_add_right)
|
||||
apply (rule ccontr, simp add: linorder_not_le)
|
||||
apply (drule aligned_small_is_0, erule order_less_trans)
|
||||
apply (clarsimp simp: power_overflow)
|
||||
apply simp
|
||||
apply (erule order_le_less_trans[rotated],
|
||||
rule word_plus_mono_right)
|
||||
apply (erule minus_one_helper3)
|
||||
apply (simp add: is_aligned_no_wrap' is_aligned_no_overflow field_simps)
|
||||
done
|
||||
|
||||
lemma is_aligned_add_helper:
|
||||
"\<lbrakk> is_aligned p n; d < 2 ^ n \<rbrakk>
|
||||
\<Longrightarrow> (p + d && mask n = d) \<and> (p + d && (~~ mask n) = p)"
|
||||
apply (subst(asm) is_aligned_mask)
|
||||
apply (drule less_mask_eq)
|
||||
apply (rule context_conjI)
|
||||
apply (subst word_plus_and_or_coroll)
|
||||
apply (rule word_eqI)
|
||||
apply (drule_tac x=na in word_eqD)+
|
||||
apply (simp add: word_size)
|
||||
apply blast
|
||||
apply (rule word_eqI)
|
||||
apply (drule_tac x=na in word_eqD)+
|
||||
apply (simp add: word_ops_nth_size word_size)
|
||||
apply blast
|
||||
apply (insert word_plus_and_or_coroll2[where x="p + d" and w="mask n"])
|
||||
apply simp
|
||||
done
|
||||
|
||||
lemma is_aligned_sub_helper:
|
||||
"\<lbrakk> is_aligned (p - d) n; d < 2 ^ n \<rbrakk>
|
||||
\<Longrightarrow> (p && mask n = d) \<and> (p && (~~ mask n) = p - d)"
|
||||
by (drule(1) is_aligned_add_helper, simp)
|
||||
|
||||
lemma mask_twice:
|
||||
"(x && mask n) && mask m = x && mask (min m n)"
|
||||
apply (rule word_eqI)
|
||||
apply (simp add: word_size conj_ac)
|
||||
done
|
||||
|
||||
lemma is_aligned_after_mask:
|
||||
"\<lbrakk>is_aligned k m;m\<le> n\<rbrakk> \<Longrightarrow> is_aligned (k && mask n) m"
|
||||
by (metis is_aligned_andI1)
|
||||
|
||||
lemma and_mask_plus:
|
||||
"\<lbrakk>is_aligned ptr m; m \<le> n; n < 32; a < 2 ^ m\<rbrakk>
|
||||
\<Longrightarrow> (ptr) + a && mask n = (ptr && mask n) + a"
|
||||
apply (rule mask_eqI[where n = m])
|
||||
apply (simp add:mask_twice min_def)
|
||||
apply (simp add:is_aligned_add_helper)
|
||||
apply (subst is_aligned_add_helper[THEN conjunct1])
|
||||
apply (erule is_aligned_after_mask)
|
||||
apply simp
|
||||
apply simp
|
||||
apply simp
|
||||
apply (subgoal_tac "(ptr + a && mask n) && ~~ mask m
|
||||
= (ptr + a && ~~ mask m ) && mask n")
|
||||
apply (simp add:is_aligned_add_helper)
|
||||
apply (subst is_aligned_add_helper[THEN conjunct2])
|
||||
apply (simp add:is_aligned_after_mask)
|
||||
apply simp
|
||||
apply simp
|
||||
apply (simp add:word_bw_comms word_bw_lcs)
|
||||
done
|
||||
|
||||
end
|
||||
|
|
|
@ -966,17 +966,27 @@ lemma sep_list_conj_eq:
|
|||
apply simp
|
||||
done
|
||||
|
||||
lemma sep_list_conj_impl:
|
||||
"\<lbrakk> list_all2 (\<lambda>x y. \<forall>s. x s \<longrightarrow> y s) xs ys; (\<And>* xs) s \<rbrakk> \<Longrightarrow> (\<And>* ys) s"
|
||||
apply (induct arbitrary: s rule: list_all2_induct)
|
||||
apply simp
|
||||
apply simp
|
||||
apply (erule sep_conj_impl, simp_all)
|
||||
done
|
||||
|
||||
lemma sep_list_conj_exists:
|
||||
"(\<exists>x. (\<And>* map (\<lambda>y s. P x y s) ys) s) \<Longrightarrow> ((\<And>* map (\<lambda>y s. \<exists>x. P x y s) ys) s)"
|
||||
apply clarsimp
|
||||
apply (erule sep_list_conj_impl[rotated])
|
||||
apply (rule list_all2I, simp_all)
|
||||
by (fastforce simp: in_set_zip)
|
||||
|
||||
lemma sep_list_conj_map_impl:
|
||||
"\<lbrakk>\<And>s x. \<lbrakk>x \<in> set xs; P x s\<rbrakk> \<Longrightarrow> Q x s; (\<And>* map P xs) s\<rbrakk>
|
||||
\<Longrightarrow> (\<And>* map Q xs) s"
|
||||
apply (clarsimp simp: sep_list_conj_def)
|
||||
apply (induct xs arbitrary: s)
|
||||
apply clarsimp
|
||||
apply clarsimp
|
||||
apply (subst sep.foldl_absorb1[symmetric])
|
||||
apply (subst (asm) sep.foldl_absorb1[symmetric],
|
||||
erule sep_conj_impl, simp+)
|
||||
done
|
||||
apply (erule sep_list_conj_impl[rotated])
|
||||
apply (rule list_all2I, simp_all)
|
||||
by (fastforce simp: in_set_zip)
|
||||
|
||||
lemma sep_map_set_conj_impl:
|
||||
"\<lbrakk>sep_map_set_conj P A s; \<And>s x. \<lbrakk>x \<in> A; P x s\<rbrakk> \<Longrightarrow> Q x s; finite A\<rbrakk>
|
||||
|
|
|
@ -692,11 +692,16 @@ lemma gets_the_in_monad:
|
|||
"((v, s') \<in> fst (gets_the f s)) = (s' = s \<and> f s = Some v)"
|
||||
by (auto simp: gets_the_def in_bind in_gets in_assert_opt split: option.split)
|
||||
|
||||
lemma in_alternative:
|
||||
"(r,s') \<in> fst ((f \<sqinter> g) s) = ((r,s') \<in> fst (f s) \<or> (r,s') \<in> fst (g s))"
|
||||
by (simp add: alternative_def)
|
||||
|
||||
lemmas in_monad = inl_whenE in_whenE in_liftE in_bind in_bindE_L
|
||||
in_bindE_R in_returnOk in_throwError in_fail
|
||||
in_assertE in_assert in_return in_assert_opt
|
||||
in_get in_gets in_put in_when unlessE_whenE
|
||||
unless_when in_modify gets_the_in_monad
|
||||
in_alternative
|
||||
|
||||
subsection "Non-Failure"
|
||||
|
||||
|
|
|
@ -248,9 +248,9 @@ definition
|
|||
authorised_page_inv :: "'a PAS \<Rightarrow> ArchInvocation_A.page_invocation \<Rightarrow> bool"
|
||||
where
|
||||
"authorised_page_inv aag pi \<equiv> case pi of
|
||||
ArchInvocation_A.PageMap cap ptr slots \<Rightarrow>
|
||||
ArchInvocation_A.PageMap asid cap ptr slots \<Rightarrow>
|
||||
pas_cap_cur_auth aag cap \<and> is_subject aag (fst ptr) \<and> authorised_slots aag slots
|
||||
| ArchInvocation_A.PageRemap slots \<Rightarrow> authorised_slots aag slots
|
||||
| ArchInvocation_A.PageRemap asid slots \<Rightarrow> authorised_slots aag slots
|
||||
| ArchInvocation_A.PageUnmap cap ptr \<Rightarrow> pas_cap_cur_auth aag (Structures_A.ArchObjectCap cap) \<and> is_subject aag (fst ptr)
|
||||
| ArchInvocation_A.PageFlush typ start end pstart pd asid \<Rightarrow> True"
|
||||
|
||||
|
@ -476,6 +476,17 @@ lemma do_flush_respects[wp]:
|
|||
|
||||
crunch pspace_aligned[wp]: flush_page "pspace_aligned"
|
||||
|
||||
lemma invalidate_tlb_by_asid_respects[wp]:
|
||||
"\<lbrace>integrity aag X st\<rbrace> invalidate_tlb_by_asid asid
|
||||
\<lbrace>\<lambda>rv. integrity aag X st\<rbrace>"
|
||||
apply (simp add: invalidate_tlb_by_asid_def)
|
||||
apply (wp dmo_no_mem_respects | wpc | simp add: invalidateTLB_ASID_def )+
|
||||
done
|
||||
|
||||
lemma invalidate_tlb_by_asid_pas_refined[wp]:
|
||||
"\<lbrace>pas_refined aag\<rbrace> invalidate_tlb_by_asid asid \<lbrace>\<lambda>rv. pas_refined aag\<rbrace>"
|
||||
by (wp dmo_no_mem_respects | wpc | simp add: invalidate_tlb_by_asid_def invalidateTLB_ASID_def)+
|
||||
|
||||
lemma perform_page_invocation_respects:
|
||||
"\<lbrace>integrity aag X st and pas_refined aag and K (authorised_page_inv aag pi) and valid_page_inv pi and valid_arch_objs and pspace_aligned\<rbrace>
|
||||
perform_page_invocation pi
|
||||
|
@ -490,16 +501,17 @@ proof -
|
|||
|
||||
show ?thesis
|
||||
apply (simp add: perform_page_invocation_def mapM_discarded swp_def
|
||||
valid_page_inv_def valid_unmap_def
|
||||
authorised_page_inv_def authorised_slots_def
|
||||
valid_page_inv_def valid_unmap_def
|
||||
authorised_page_inv_def authorised_slots_def
|
||||
split: ArchInvocation_A.page_invocation.split sum.split
|
||||
arch_cap.split option.split,
|
||||
safe)
|
||||
apply (wp set_cap_integrity_autarch unmap_page_respects
|
||||
apply (wp set_cap_integrity_autarch unmap_page_respects
|
||||
mapM_x_and_const_wp[OF store_pte_respects] store_pte_respects
|
||||
mapM_x_and_const_wp[OF store_pde_respects] store_pde_respects
|
||||
| elim conjE hd_valid_slots[THEN bspec[rotated]]
|
||||
| clarsimp dest!:set_tl_subset_mp )+
|
||||
| clarsimp dest!:set_tl_subset_mp
|
||||
| wpc )+
|
||||
apply (clarsimp simp: cte_wp_at_caps_of_state is_arch_diminished_def
|
||||
cap_auth_conferred_def cap_rights_update_def
|
||||
acap_rights_update_def update_map_data_def is_pg_cap_def
|
||||
|
|
|
@ -778,7 +778,7 @@ lemma perform_page_invocation_domain_sep_inv:
|
|||
|
||||
| simp add: perform_page_invocation_def o_def | wpc)+
|
||||
apply(clarsimp simp: valid_page_inv_def)
|
||||
apply(case_tac x, simp_all add: domain_sep_inv_cap_def is_pg_cap_def)
|
||||
apply(case_tac xa, simp_all add: domain_sep_inv_cap_def is_pg_cap_def)
|
||||
done
|
||||
|
||||
lemma perform_page_table_invocation_domain_sep_inv:
|
||||
|
|
File diff suppressed because it is too large
Load Diff
File diff suppressed because it is too large
Load Diff
File diff suppressed because it is too large
Load Diff
|
@ -1,696 +0,0 @@
|
|||
(*
|
||||
* Copyright 2014, NICTA
|
||||
*
|
||||
* This software may be distributed and modified according to the terms of
|
||||
* the GNU General Public License version 2. Note that NO WARRANTY is provided.
|
||||
* See "LICENSE_GPLv2.txt" for details.
|
||||
*
|
||||
* @TAG(NICTA_GPL)
|
||||
*)
|
||||
|
||||
theory sac_dump_simplified
|
||||
imports "../../spec/capDL/Types_D"
|
||||
begin
|
||||
|
||||
text {* the AEP between SAC-C and Timer *}
|
||||
definition obj0 :: cdl_object where
|
||||
"obj0 \<equiv> Types_D.Endpoint"
|
||||
|
||||
text {* the AEP between RM and Timer *}
|
||||
definition obj1 :: cdl_object where
|
||||
"obj1 \<equiv> Types_D.Endpoint"
|
||||
|
||||
text {* Timer's asid pool *}
|
||||
definition obj2 :: cdl_object where
|
||||
"obj2 \<equiv> Types_D.AsidPool"
|
||||
|
||||
text {* RM's asid pool *}
|
||||
definition obj3 :: cdl_object where
|
||||
"obj3 \<equiv> Types_D.AsidPool"
|
||||
|
||||
text {* SAC-C's asid pool *}
|
||||
definition obj4 :: cdl_object where
|
||||
"obj4 \<equiv> Types_D.AsidPool"
|
||||
|
||||
text {* SAC-C's CSpace:
|
||||
Memory:
|
||||
UntypedMem for ranges
|
||||
3082_3089, 3118_3122, 3137_3142, 3151_3153, 3155_3193,
|
||||
3258_3372 3410_3505,
|
||||
have been collapsed into one (UntypedMem 3151 undefined).
|
||||
|
||||
nic-C's mapped frames
|
||||
|
||||
FrameCap (Standard i) for i=10..42 have been collapsed in
|
||||
FrameCap (Standard 10)
|
||||
*}
|
||||
|
||||
definition caps5 :: cdl_cap_map where
|
||||
"caps5 \<equiv> [1 \<mapsto> Types_D.TcbCap (Standard 3079),
|
||||
2 \<mapsto> Types_D.CNodeCap (Standard 6) 0 0,
|
||||
3 \<mapsto> Types_D.PageDirectoryCap (Standard 3063),
|
||||
6 \<mapsto> Types_D.AsidPoolCap (Standard 4),
|
||||
11 \<mapsto> Types_D.UntypedCap (UntypedMem 3151 undefined),
|
||||
283 \<mapsto> Types_D.IrqHandlerCap 27,
|
||||
284 \<mapsto> Types_D.FrameCap (Standard 10) {} 12 True,
|
||||
317 \<mapsto> Types_D.IOSpaceCap (Standard 3057),
|
||||
318 \<mapsto> Types_D.EndpointCap (Standard 9) 0 {Write},
|
||||
319 \<mapsto> Types_D.AsyncEndpointCap (Standard 0) 0 {Read}]"
|
||||
|
||||
definition obj5 :: cdl_object where
|
||||
"obj5 \<equiv> Types_D.CNode \<lparr> cdl_cnode_caps = caps5, cdl_cnode_size_bits = 10 \<rparr>"
|
||||
|
||||
definition caps6 :: cdl_cap_map where
|
||||
"caps6 \<equiv> [0 \<mapsto> Types_D.CNodeCap (Standard 5) 0 0]"
|
||||
|
||||
definition obj6 :: cdl_object where
|
||||
"obj6 \<equiv> Types_D.CNode \<lparr> cdl_cnode_caps = caps6, cdl_cnode_size_bits = 10 \<rparr>"
|
||||
|
||||
|
||||
|
||||
text {* RM's Cspace
|
||||
|
||||
UntypedMem for ranges
|
||||
3090_3117, 3124_3136, 3143_3150, 3154, 3194_3257, 3373_3409
|
||||
have been collapsed into one (UntypedMem 3124 undefined).
|
||||
|
||||
FrameCap (Standard i) for i=43..106 have been collapsed in
|
||||
FrameCap (Standard 43)
|
||||
|
||||
FrameCap (Standard i) for i=107..170 have been collapsed in
|
||||
FrameCap (Standard 107)
|
||||
|
||||
FrameCap (Standard i) for i=171..174 have been collapsed in
|
||||
FrameCap (Standard 171)
|
||||
|
||||
*}
|
||||
|
||||
definition caps7 :: cdl_cap_map where
|
||||
"caps7 \<equiv> [1 \<mapsto> Types_D.TcbCap (Standard 3080),
|
||||
2 \<mapsto> Types_D.CNodeCap (Standard 7) 0 0,
|
||||
3 \<mapsto> Types_D.PageDirectoryCap (Standard 3065),
|
||||
6 \<mapsto> Types_D.AsidPoolCap (Standard 3),
|
||||
11 \<mapsto> Types_D.PageDirectoryCap (Standard 3064),
|
||||
12 \<mapsto> Types_D.UntypedCap (UntypedMem 3124 undefined),
|
||||
163 \<mapsto> Types_D.IrqHandlerCap 25,
|
||||
164 \<mapsto> Types_D.FrameCap (Standard 43) {} 12 True,
|
||||
228 \<mapsto> Types_D.IOSpaceCap (Standard 3054),
|
||||
229 \<mapsto> Types_D.IrqHandlerCap 26,
|
||||
230 \<mapsto> Types_D.FrameCap (Standard 107) {} 12 True,
|
||||
294 \<mapsto> Types_D.IOSpaceCap (Standard 3055),
|
||||
295 \<mapsto> Types_D.IrqHandlerCap 28,
|
||||
296 \<mapsto> Types_D.FrameCap (Standard 171) {} 12 True,
|
||||
300 \<mapsto> Types_D.IOSpaceCap (Standard 3056),
|
||||
301 \<mapsto> Types_D.IOPortsCap undefined undefined,
|
||||
302 \<mapsto> Types_D.EndpointCap (Standard 9) 0 {Read},
|
||||
303 \<mapsto> Types_D.AsyncEndpointCap (Standard 1) 0 {Read}]"
|
||||
|
||||
definition obj7 :: cdl_object where
|
||||
"obj7 \<equiv> Types_D.CNode \<lparr> cdl_cnode_caps = caps7, cdl_cnode_size_bits = 10 \<rparr>"
|
||||
|
||||
|
||||
text {* Timer's Cspace *}
|
||||
|
||||
definition caps8 :: cdl_cap_map where
|
||||
"caps8 \<equiv> [1 \<mapsto> Types_D.TcbCap (Standard 3081),
|
||||
2 \<mapsto> Types_D.CNodeCap (Standard 8) 0 0,
|
||||
3 \<mapsto> Types_D.PageDirectoryCap (Standard 3066),
|
||||
6 \<mapsto> Types_D.AsidPoolCap (Standard 2),
|
||||
11 \<mapsto> Types_D.UntypedCap (UntypedMem 3123 undefined),
|
||||
12 \<mapsto> Types_D.IrqHandlerCap 0,
|
||||
13 \<mapsto> Types_D.IOPortsCap undefined undefined,
|
||||
14 \<mapsto> Types_D.AsyncEndpointCap (Standard 0) 4 {Write},
|
||||
15 \<mapsto> Types_D.AsyncEndpointCap (Standard 1) 16 {Write}]"
|
||||
|
||||
definition obj8 :: cdl_object where
|
||||
"obj8 \<equiv> Types_D.CNode \<lparr> cdl_cnode_caps = caps8, cdl_cnode_size_bits = 10 \<rparr>"
|
||||
|
||||
|
||||
|
||||
text {* endpoint between SAC-C and RM *}
|
||||
definition obj9 :: cdl_object where
|
||||
"obj9 \<equiv> Types_D.Endpoint"
|
||||
|
||||
|
||||
|
||||
text {* nic-C's mapped frames
|
||||
|
||||
FrameCap (Standard i) for i=10..42 have been collapsed in
|
||||
FrameCap (Standard 10)
|
||||
*}
|
||||
|
||||
definition obj10 :: cdl_object where
|
||||
"obj10 \<equiv> Types_D.Frame \<lparr> cdl_frame_size_bits = 12 \<rparr>"
|
||||
|
||||
|
||||
text {* Some Nic (say A) mapped frames:
|
||||
|
||||
FrameCap (Standard i) for i=43..106 have been collapsed in
|
||||
FrameCap (Standard 43)
|
||||
*}
|
||||
|
||||
definition obj43 :: cdl_object where
|
||||
"obj43 \<equiv> Types_D.Frame \<lparr> cdl_frame_size_bits = 12 \<rparr>"
|
||||
|
||||
|
||||
text {* Some Nic (say B) mapped frames:
|
||||
|
||||
FrameCap (Standard i) for i=107..170 have been collapsed in
|
||||
FrameCap (Standard 107)
|
||||
|
||||
*}
|
||||
|
||||
|
||||
definition obj107 :: cdl_object where
|
||||
"obj107 \<equiv> Types_D.Frame \<lparr> cdl_frame_size_bits = 12 \<rparr>"
|
||||
|
||||
|
||||
text {* Some Nic (say D) mapped frames:
|
||||
|
||||
FrameCap (Standard i) for i=171..174 have been collapsed in
|
||||
FrameCap (Standard 171)
|
||||
|
||||
*}
|
||||
|
||||
definition obj171 :: cdl_object where
|
||||
"obj171 \<equiv> Types_D.Frame \<lparr> cdl_frame_size_bits = 12 \<rparr>"
|
||||
|
||||
|
||||
|
||||
|
||||
text {* R's PageTable 1 of 3
|
||||
|
||||
FrameCap (Standard i) for i=175_280, 303_1198, 1449_1454 have been collapsed in
|
||||
FrameCap (Standard 280) {Read}
|
||||
|
||||
*}
|
||||
|
||||
definition obj280 :: cdl_object where
|
||||
"obj280 \<equiv> Types_D.Frame \<lparr> cdl_frame_size_bits = 12 \<rparr>"
|
||||
|
||||
|
||||
|
||||
text {* RM's PageTable 1 of 3
|
||||
|
||||
FrameCap (Standard i) for i=281_287 have been collapsed in
|
||||
FrameCap (Standard 287) {Read}
|
||||
|
||||
*}
|
||||
|
||||
definition obj287 :: cdl_object where
|
||||
"obj287 \<equiv> Types_D.Frame \<lparr> cdl_frame_size_bits = 12 \<rparr>"
|
||||
|
||||
|
||||
text {* RM's PageTable 3 of 3
|
||||
|
||||
FrameCap (Standard i) for i=288, 2994, 2995 have been collapsed in
|
||||
FrameCap (Standard 288) {Read, Write}
|
||||
|
||||
*}
|
||||
definition obj288 :: cdl_object where
|
||||
"obj288 \<equiv> Types_D.Frame \<lparr> cdl_frame_size_bits = 12 \<rparr>"
|
||||
|
||||
|
||||
|
||||
text {* Timer's PageTable 1 of 2
|
||||
|
||||
FrameCap (Standard i) for i=289_301 have been collapsed in
|
||||
FrameCap (Standard 301) {Read}
|
||||
|
||||
*}
|
||||
|
||||
definition obj301 :: cdl_object where
|
||||
"obj301 \<equiv> Types_D.Frame \<lparr> cdl_frame_size_bits = 12 \<rparr>"
|
||||
|
||||
|
||||
|
||||
text {* Timer's PageTable 2 of 2
|
||||
|
||||
FrameCap (Standard i) for i=302, 2996, 2997 have been collapsed in
|
||||
FrameCap (Standard 302) {Read, Write}
|
||||
|
||||
*}
|
||||
definition obj302 :: cdl_object where
|
||||
"obj302 \<equiv> Types_D.Frame \<lparr> cdl_frame_size_bits = 12 \<rparr>"
|
||||
|
||||
|
||||
text {* R's PageTable 2 of 3
|
||||
|
||||
FrameCap (Standard i) for i=1199_1448,1614_1710 have been collapsed in
|
||||
FrameCap (Standard 1448) {Read}
|
||||
|
||||
*}
|
||||
|
||||
definition obj1448 :: cdl_object where
|
||||
"obj1448 \<equiv> Types_D.Frame \<lparr> cdl_frame_size_bits = 12 \<rparr>"
|
||||
|
||||
|
||||
|
||||
text {* SAC-C's PageTable 1 of 4
|
||||
|
||||
FrameCap (Standard i) for i=1455_1611, 1711_1966, 2086-2222 have been collapsed in
|
||||
FrameCap (Standard 1455) {Read}
|
||||
|
||||
FrameCap (Standard i) for i=1967_2085, 2332-2478 have been collapsed in
|
||||
FrameCap (Standard 2332) {Read, Write}
|
||||
|
||||
*}
|
||||
|
||||
definition obj1455 :: cdl_object where
|
||||
"obj1455 \<equiv> Types_D.Frame \<lparr> cdl_frame_size_bits = 12 \<rparr>"
|
||||
|
||||
|
||||
|
||||
text {* SAC-C's PageTable 4 of 4
|
||||
|
||||
FrameCap (Standard i) for i=1612, 2991, 2992 have been collapsed in
|
||||
FrameCap (Standard 1612) {Read, Write}
|
||||
|
||||
*}
|
||||
definition obj1612 :: cdl_object where
|
||||
"obj1612 \<equiv> Types_D.Frame \<lparr> cdl_frame_size_bits = 12 \<rparr>"
|
||||
|
||||
|
||||
|
||||
text {* R's PageTable 1 of 3
|
||||
|
||||
FrameCap (Standard i) for i=1613 and 2993 have been collapsed in
|
||||
FrameCap (Standard 1613) {Read}
|
||||
|
||||
*}
|
||||
definition obj1613 :: cdl_object where
|
||||
"obj1613 \<equiv> Types_D.Frame \<lparr> cdl_frame_size_bits = 12 \<rparr>"
|
||||
|
||||
|
||||
|
||||
text {* SAC-C's PageTable 2 of 4
|
||||
|
||||
FrameCap (Standard i) for i=2223_2331, 2479_2990, 2998_3052 have been collapsed in
|
||||
FrameCap (Standard 2223) {Read, Write}
|
||||
|
||||
*}
|
||||
|
||||
definition obj2223 :: cdl_object where
|
||||
"obj2223 \<equiv> Types_D.Frame \<lparr> cdl_frame_size_bits = 12 \<rparr>"
|
||||
|
||||
|
||||
text {* SAC-C's PageTable 1 of 4
|
||||
|
||||
FrameCap (Standard i) for i=1455_1611, 1711_1966, 2086-2222 have been collapsed in
|
||||
FrameCap (Standard 1455) {Read}
|
||||
|
||||
FrameCap (Standard i) for i=1967_2085, 2332-2478 have been collapsed in
|
||||
FrameCap (Standard 2332) {Read, Write}
|
||||
|
||||
*}
|
||||
definition obj2332 :: cdl_object where
|
||||
"obj2332 \<equiv> Types_D.Frame \<lparr> cdl_frame_size_bits = 12 \<rparr>"
|
||||
|
||||
|
||||
text {* IOPorts and devices *}
|
||||
definition obj3053 :: cdl_object where
|
||||
"obj3053 \<equiv> Types_D.TODO (* IOPorts *)"
|
||||
|
||||
definition caps3054 :: cdl_cap_map where
|
||||
"caps3054 \<equiv> empty"
|
||||
|
||||
definition obj3054 :: cdl_object where
|
||||
"obj3054 \<equiv> Types_D.TODO (* IODevice *)"
|
||||
|
||||
definition caps3055 :: cdl_cap_map where
|
||||
"caps3055 \<equiv> empty"
|
||||
|
||||
definition obj3055 :: cdl_object where
|
||||
"obj3055 \<equiv> Types_D.TODO (* IODevice *)"
|
||||
|
||||
definition caps3056 :: cdl_cap_map where
|
||||
"caps3056 \<equiv> empty"
|
||||
|
||||
definition obj3056 :: cdl_object where
|
||||
"obj3056 \<equiv> Types_D.TODO (* IODevice *)"
|
||||
|
||||
definition caps3057 :: cdl_cap_map where
|
||||
"caps3057 \<equiv> empty"
|
||||
|
||||
definition obj3057 :: cdl_object where
|
||||
"obj3057 \<equiv> Types_D.TODO (* IODevice *)"
|
||||
|
||||
text {* only keeping only empty cnode 3058
|
||||
(removing 3059_3062 and 35_073756) *}
|
||||
|
||||
definition caps3058 :: cdl_cap_map where
|
||||
"caps3058 \<equiv> empty"
|
||||
|
||||
definition obj3058 :: cdl_object where
|
||||
"obj3058 \<equiv> Types_D.CNode \<lparr> cdl_cnode_caps = caps3058, cdl_cnode_size_bits = 0 \<rparr>"
|
||||
|
||||
|
||||
|
||||
text {* SAC-C's VSpace (PageDirectory)*}
|
||||
|
||||
definition caps3063 :: cdl_cap_map where
|
||||
"caps3063 \<equiv> [0 \<mapsto> Types_D.PageTableCap (Standard 3072) False,
|
||||
1 \<mapsto> Types_D.PageTableCap (Standard 3071) False,
|
||||
768 \<mapsto> Types_D.PageTableCap (Standard 3070) False,
|
||||
832 \<mapsto> Types_D.PageTableCap (Standard 3067) False]"
|
||||
|
||||
definition obj3063 :: cdl_object where
|
||||
"obj3063 \<equiv> Types_D.PageDirectory \<lparr> cdl_page_directory_caps = caps3063 \<rparr>"
|
||||
|
||||
|
||||
text {* R's VSpace (PageDirectory)*}
|
||||
|
||||
definition caps3064 :: cdl_cap_map where
|
||||
"caps3064 \<equiv> [0 \<mapsto> Types_D.PageTableCap (Standard 3076) False,
|
||||
1 \<mapsto> Types_D.PageTableCap (Standard 3075) False,
|
||||
832 \<mapsto> Types_D.PageTableCap (Standard 3074) False]"
|
||||
|
||||
definition obj3064 :: cdl_object where
|
||||
"obj3064 \<equiv> Types_D.PageDirectory \<lparr> cdl_page_directory_caps = caps3064 \<rparr>"
|
||||
|
||||
|
||||
text {* RM's VSpace (PageDirectory)*}
|
||||
|
||||
definition caps3065 :: cdl_cap_map where
|
||||
"caps3065 \<equiv> [0 \<mapsto> Types_D.PageTableCap (Standard 3077) False,
|
||||
768 \<mapsto> Types_D.PageTableCap (Standard 3073) False,
|
||||
832 \<mapsto> Types_D.PageTableCap (Standard 3068) False]"
|
||||
|
||||
definition obj3065 :: cdl_object where
|
||||
"obj3065 \<equiv> Types_D.PageDirectory \<lparr> cdl_page_directory_caps = caps3065 \<rparr>"
|
||||
|
||||
|
||||
|
||||
text {* Timer's VSpace (PageDirectory)*}
|
||||
|
||||
definition caps3066 :: cdl_cap_map where
|
||||
"caps3066 \<equiv> [0 \<mapsto> Types_D.PageTableCap (Standard 3078) False,
|
||||
832 \<mapsto> Types_D.PageTableCap (Standard 3069) False]"
|
||||
|
||||
definition obj3066 :: cdl_object where
|
||||
"obj3066 \<equiv> Types_D.PageDirectory \<lparr> cdl_page_directory_caps = caps3066 \<rparr>"
|
||||
|
||||
|
||||
text {* SAC-C's PageTable 4 of 4
|
||||
|
||||
FrameCap (Standard i) for i=1612, 2991, 2992 have been collapsed in
|
||||
FrameCap (Standard 1612) {Read, Write}
|
||||
|
||||
*}
|
||||
definition caps3067 :: cdl_cap_map where
|
||||
"caps3067 \<equiv> [0 \<mapsto> Types_D.FrameCap (Standard 1612) {Read, Write} 12 False]"
|
||||
|
||||
|
||||
definition obj3067 :: cdl_object where
|
||||
"obj3067 \<equiv> Types_D.PageTable \<lparr> cdl_page_table_caps = caps3067 \<rparr>"
|
||||
|
||||
|
||||
text {* RM's PageTable 3 of 3
|
||||
|
||||
FrameCap (Standard i) for i=288, 2994, 2995 have been collapsed in
|
||||
FrameCap (Standard 288) {Read, Write}
|
||||
|
||||
*}
|
||||
definition caps3068 :: cdl_cap_map where
|
||||
"caps3068 \<equiv> [0 \<mapsto> Types_D.FrameCap (Standard 288) {Read, Write} 12 False]"
|
||||
|
||||
definition obj3068 :: cdl_object where
|
||||
"obj3068 \<equiv> Types_D.PageTable \<lparr> cdl_page_table_caps = caps3068 \<rparr>"
|
||||
|
||||
|
||||
text {* Timer's PageTable 2 of 2
|
||||
|
||||
FrameCap (Standard i) for i=302, 2996, 2997 have been collapsed in
|
||||
FrameCap (Standard 302) {Read, Write}
|
||||
|
||||
*}
|
||||
|
||||
definition caps3069 :: cdl_cap_map where
|
||||
"caps3069 \<equiv> [0 \<mapsto> Types_D.FrameCap (Standard 302) {Read, Write} 12 False]"
|
||||
|
||||
definition obj3069 :: cdl_object where
|
||||
"obj3069 \<equiv> Types_D.PageTable \<lparr> cdl_page_table_caps = caps3069 \<rparr>"
|
||||
|
||||
text {* SAC-C's PageTable 3 of 4
|
||||
|
||||
FrameCap (Standard i) for i=10_42 have been collapsed in
|
||||
FrameCap (Standard 10) {Read, Write}
|
||||
|
||||
*}
|
||||
|
||||
definition caps3070 :: cdl_cap_map where
|
||||
"caps3070 \<equiv> [384 \<mapsto> Types_D.FrameCap (Standard 10) {Read, Write} 12 False]"
|
||||
|
||||
definition obj3070 :: cdl_object where
|
||||
"obj3070 \<equiv> Types_D.PageTable \<lparr> cdl_page_table_caps = caps3070 \<rparr>"
|
||||
|
||||
|
||||
|
||||
text {* SAC-C's PageTable 2 of 4
|
||||
|
||||
FrameCap (Standard i) for i=2223_2331, 2479_2990, 2998_3052 have been collapsed in
|
||||
FrameCap (Standard 2223) {Read, Write}
|
||||
|
||||
*}
|
||||
|
||||
definition caps3071 :: cdl_cap_map where
|
||||
"caps3071 \<equiv> [0 \<mapsto> Types_D.FrameCap (Standard 2223) {Read, Write} 12 False]"
|
||||
|
||||
definition obj3071 :: cdl_object where
|
||||
"obj3071 \<equiv> Types_D.PageTable \<lparr> cdl_page_table_caps = caps3071 \<rparr>"
|
||||
|
||||
|
||||
text {* SAC-C's PageTable 1 of 4
|
||||
|
||||
FrameCap (Standard i) for i=1455_1611, 1711_1967, 2086-2222 have been collapsed in
|
||||
FrameCap (Standard 1455) {Read}
|
||||
|
||||
FrameCap (Standard i) for i=1967_2085, 2332-2478 have been collapsed in
|
||||
FrameCap (Standard 2332) {Read, Write}
|
||||
|
||||
*}
|
||||
|
||||
definition caps3072 :: cdl_cap_map where
|
||||
"caps3072 \<equiv> [208 \<mapsto> Types_D.FrameCap (Standard 1455) {Read} 12 False,
|
||||
877 \<mapsto> Types_D.FrameCap (Standard 2332) {Read, Write} 12 False]"
|
||||
|
||||
definition obj3072 :: cdl_object where
|
||||
"obj3072 \<equiv> Types_D.PageTable \<lparr> cdl_page_table_caps = caps3072 \<rparr>"
|
||||
|
||||
|
||||
text {* RM's PageTable 2 of 3 (mapped frames for Nic A, B, D)
|
||||
|
||||
FrameCap (Standard i) for i=43..106 have been collapsed in
|
||||
FrameCap (Standard 43)
|
||||
|
||||
FrameCap (Standard i) for i=107..170 have been collapsed in
|
||||
FrameCap (Standard 107)
|
||||
|
||||
FrameCap (Standard i) for i=171..174 have been collapsed in
|
||||
FrameCap (Standard 171)
|
||||
*}
|
||||
|
||||
|
||||
definition caps3073 :: cdl_cap_map where
|
||||
"caps3073 \<equiv> [512 \<mapsto> Types_D.FrameCap (Standard 43) {Read, Write} 12 False,
|
||||
576 \<mapsto> Types_D.FrameCap (Standard 107) {Read, Write} 12 False,
|
||||
768 \<mapsto> Types_D.FrameCap (Standard 171) {Read, Write} 12 False]"
|
||||
|
||||
definition obj3073 :: cdl_object where
|
||||
"obj3073 \<equiv> Types_D.PageTable \<lparr> cdl_page_table_caps = caps3073 \<rparr>"
|
||||
|
||||
|
||||
text {* R's PageTable 1 of 3
|
||||
|
||||
FrameCap (Standard i) for i=1613 and 2993 have been collapsed in
|
||||
FrameCap (Standard 1613) {Read}
|
||||
|
||||
*}
|
||||
|
||||
definition caps3074 :: cdl_cap_map where
|
||||
"caps3074 \<equiv> [0 \<mapsto> Types_D.FrameCap (Standard 1613) {Read} 12 False,
|
||||
1 \<mapsto> Types_D.FrameCap (Standard 2993) {Read} 12 False]"
|
||||
|
||||
definition obj3074 :: cdl_object where
|
||||
"obj3074 \<equiv> Types_D.PageTable \<lparr> cdl_page_table_caps = caps3074 \<rparr>"
|
||||
|
||||
|
||||
text {* R's PageTable 2 of 3
|
||||
|
||||
FrameCap (Standard i) for i=1199_1448,1614_1710 have been collapsed in
|
||||
FrameCap (Standard 1448) {Read}
|
||||
|
||||
*}
|
||||
definition caps3075 :: cdl_cap_map where
|
||||
"caps3075 \<equiv> [0 \<mapsto> Types_D.FrameCap (Standard 1448) {Read} 12 False]"
|
||||
|
||||
definition obj3075 :: cdl_object where
|
||||
"obj3075 \<equiv> Types_D.PageTable \<lparr> cdl_page_table_caps = caps3075 \<rparr>"
|
||||
|
||||
|
||||
text {* R's PageTable 1 of 3
|
||||
|
||||
FrameCap (Standard i) for i=175_280, 303_1198, 1449_1454 have been collapsed in
|
||||
FrameCap (Standard 280) {Read}
|
||||
|
||||
*}
|
||||
|
||||
definition caps3076 :: cdl_cap_map where
|
||||
"caps3076 \<equiv> [16 \<mapsto> Types_D.FrameCap (Standard 280) {Read} 12 False]"
|
||||
|
||||
definition obj3076 :: cdl_object where
|
||||
"obj3076 \<equiv> Types_D.PageTable \<lparr> cdl_page_table_caps = caps3076 \<rparr>"
|
||||
|
||||
|
||||
text {* RM's PageTable 1 of 3
|
||||
|
||||
FrameCap (Standard i) for i=281_287 have been collapsed in
|
||||
FrameCap (Standard 287) {Read}
|
||||
|
||||
*}
|
||||
|
||||
definition caps3077 :: cdl_cap_map where
|
||||
"caps3077 \<equiv> [16 \<mapsto> Types_D.FrameCap (Standard 287) {Read} 12 False]"
|
||||
|
||||
definition obj3077 :: cdl_object where
|
||||
"obj3077 \<equiv> Types_D.PageTable \<lparr> cdl_page_table_caps = caps3077 \<rparr>"
|
||||
|
||||
text {* Timer's PageTable 1 of 2
|
||||
|
||||
FrameCap (Standard i) for i=289_301 have been collapsed in
|
||||
FrameCap (Standard 301) {Read}
|
||||
|
||||
*}
|
||||
|
||||
definition caps3078 :: cdl_cap_map where
|
||||
"caps3078 \<equiv> [16 \<mapsto> Types_D.FrameCap (Standard 301) {Read} 12 False]"
|
||||
|
||||
definition obj3078 :: cdl_object where
|
||||
"obj3078 \<equiv> Types_D.PageTable \<lparr> cdl_page_table_caps = caps3078 \<rparr>"
|
||||
|
||||
|
||||
text {* SAC-C's tcb *}
|
||||
|
||||
definition caps3079 :: cdl_cap_map where
|
||||
"caps3079 \<equiv> [0 \<mapsto> Types_D.CNodeCap (Standard 6) 0 0,
|
||||
1 \<mapsto> Types_D.PageDirectoryCap (Standard 3063),
|
||||
2 \<mapsto> Types_D.ReplyCap (Standard 3079),
|
||||
4 \<mapsto> Types_D.FrameCap (Standard 1612) {} 12 True]"
|
||||
|
||||
definition obj3079 :: cdl_object where
|
||||
"obj3079 \<equiv> Types_D.Tcb \<lparr> cdl_tcb_caps = caps3079, cdl_tcb_fault_endpoint = undefined, cdl_tcb_intent = undefined \<rparr>"
|
||||
|
||||
|
||||
text {*RM's tcb *}
|
||||
definition caps3080 :: cdl_cap_map where
|
||||
"caps3080 \<equiv> [0 \<mapsto> Types_D.CNodeCap (Standard 7) 0 0,
|
||||
1 \<mapsto> Types_D.PageDirectoryCap (Standard 3065),
|
||||
2 \<mapsto> Types_D.ReplyCap (Standard 3080),
|
||||
4 \<mapsto> Types_D.FrameCap (Standard 288) {} 12 True]"
|
||||
|
||||
definition obj3080 :: cdl_object where
|
||||
"obj3080 \<equiv> Types_D.Tcb \<lparr> cdl_tcb_caps = caps3080, cdl_tcb_fault_endpoint = undefined, cdl_tcb_intent = undefined \<rparr>"
|
||||
|
||||
|
||||
text {* Timer's tcb *}
|
||||
|
||||
definition caps3081 :: cdl_cap_map where
|
||||
"caps3081 \<equiv> [0 \<mapsto> Types_D.CNodeCap (Standard 8) 0 0,
|
||||
1 \<mapsto> Types_D.PageDirectoryCap (Standard 3066),
|
||||
2 \<mapsto> Types_D.ReplyCap (Standard 3081),
|
||||
4 \<mapsto> Types_D.FrameCap (Standard 302) {} 12 True]"
|
||||
|
||||
definition obj3081 :: cdl_object where
|
||||
"obj3081 \<equiv> Types_D.Tcb \<lparr> cdl_tcb_caps = caps3081, cdl_tcb_fault_endpoint = undefined, cdl_tcb_intent = undefined \<rparr>"
|
||||
|
||||
|
||||
text {* one untyped for the Timer *}
|
||||
|
||||
definition obj3123 :: cdl_object where
|
||||
"obj3123 \<equiv> Types_D.Untyped"
|
||||
|
||||
|
||||
text {* RM's Memory
|
||||
|
||||
UntypedMem for ranges
|
||||
3090_3117, 3124_3136, 3143_3150, 3154, 3194_3257, 3373_3409
|
||||
have been collapsed into one (UntypedMem 3124 undefined).
|
||||
|
||||
*}
|
||||
|
||||
definition obj3124 :: cdl_object where
|
||||
"obj3124 \<equiv> Types_D.Untyped"
|
||||
|
||||
|
||||
text {* SAC-C's memory
|
||||
|
||||
UntypedMem for ranges
|
||||
3082_3089, 3118_3122, 3137_3142, 3151_3153, 3155_3193,
|
||||
3258_3372 3410_3505,
|
||||
have been collapsed into one (UntypedMem 3151 undefined) *}
|
||||
|
||||
definition obj3151 :: cdl_object where
|
||||
"obj3151 \<equiv> Types_D.Untyped"
|
||||
|
||||
|
||||
|
||||
|
||||
|
||||
definition objects :: "cdl_object_id \<Rightarrow> cdl_object option" where
|
||||
"objects \<equiv> [(Standard 0) \<mapsto> obj0,
|
||||
(Standard 1) \<mapsto> obj1, (Standard 2) \<mapsto> obj2,
|
||||
(Standard 3) \<mapsto> obj3, (Standard 4) \<mapsto> obj4,
|
||||
(Standard 5) \<mapsto> obj5, (Standard 6) \<mapsto> obj6,
|
||||
(Standard 7) \<mapsto> obj7, (Standard 8) \<mapsto> obj8,
|
||||
(Standard 9) \<mapsto> obj9, (Standard 10) \<mapsto> obj10,
|
||||
(Standard 43) \<mapsto> obj43,
|
||||
(Standard 107) \<mapsto> obj107,
|
||||
(Standard 171) \<mapsto> obj171,
|
||||
(Standard 280) \<mapsto> obj280,
|
||||
(Standard 287) \<mapsto> obj287,
|
||||
(Standard 288) \<mapsto> obj288,
|
||||
(Standard 301) \<mapsto> obj301,
|
||||
(Standard 302) \<mapsto> obj302,
|
||||
(Standard 1448) \<mapsto> obj1448,
|
||||
(Standard 1455) \<mapsto> obj1455,
|
||||
(Standard 1612) \<mapsto> obj1612,
|
||||
(Standard 1613) \<mapsto> obj1613,
|
||||
(Standard 2223) \<mapsto> obj2223,
|
||||
(Standard 2332) \<mapsto> obj2332,
|
||||
(Standard 3053) \<mapsto> obj3053,
|
||||
(Standard 3054) \<mapsto> obj3054,
|
||||
(Standard 3055) \<mapsto> obj3055,
|
||||
(Standard 3056) \<mapsto> obj3056,
|
||||
(Standard 3057) \<mapsto> obj3057,
|
||||
(Standard 3058) \<mapsto> obj3058,
|
||||
(Standard 3063) \<mapsto> obj3063,
|
||||
(Standard 3064) \<mapsto> obj3064,
|
||||
(Standard 3065) \<mapsto> obj3065,
|
||||
(Standard 3066) \<mapsto> obj3066,
|
||||
(Standard 3067) \<mapsto> obj3067,
|
||||
(Standard 3068) \<mapsto> obj3068,
|
||||
(Standard 3069) \<mapsto> obj3069,
|
||||
(Standard 3070) \<mapsto> obj3070,
|
||||
(Standard 3071) \<mapsto> obj3071,
|
||||
(Standard 3072) \<mapsto> obj3072,
|
||||
(Standard 3073) \<mapsto> obj3073,
|
||||
(Standard 3074) \<mapsto> obj3074,
|
||||
(Standard 3075) \<mapsto> obj3075,
|
||||
(Standard 3076) \<mapsto> obj3076,
|
||||
(Standard 3077) \<mapsto> obj3077,
|
||||
(Standard 3078) \<mapsto> obj3078,
|
||||
(Standard 3079) \<mapsto> obj3079,
|
||||
(Standard 3080) \<mapsto> obj3080,
|
||||
(Standard 3081) \<mapsto> obj3081,
|
||||
(UntypedMem 3123 undefined) \<mapsto> obj3123,
|
||||
(UntypedMem 3124 undefined) \<mapsto> obj3124,
|
||||
(UntypedMem 3151 undefined) \<mapsto> obj3151]"
|
||||
|
||||
|
||||
definition irqs :: "cdl_irq \<Rightarrow> cdl_object_id" where
|
||||
"irqs \<equiv> undefined (0 := Standard 3058)"
|
||||
|
||||
definition state :: cdl_state where
|
||||
"state \<equiv> \<lparr> cdl_arch = IA32, cdl_objects = objects, cdl_cdt = undefined, cdl_current_thread = undefined, cdl_irq_node = irqs, cdl_untyped_covers = undefined \<rparr>"
|
||||
|
||||
end
|
|
@ -134,8 +134,8 @@ lemmas globals_list_mems = kernel_all_global_addresses.global_data_mems
|
|||
ML {*
|
||||
val globals_swap_rewrites = @{thms globals_list_mems[unfolded global_data_defs]}
|
||||
RL @{thms
|
||||
globals_swap_update_mem[OF _ global_acc_valid globals_list_valid]
|
||||
globals_swap_access_mem[OF _ global_acc_valid globals_list_valid]}
|
||||
globals_swap_update_mem2[OF _ global_acc_valid globals_list_valid]
|
||||
globals_swap_access_mem2[OF _ global_acc_valid globals_list_valid]}
|
||||
*}
|
||||
|
||||
end
|
||||
|
|
|
@ -60,11 +60,6 @@ val hints = UseHints.mk_var_deps_hints funs @{context} @{typ "globals myvars"} n
|
|||
*}
|
||||
|
||||
ML {* init_graph_refines_proof funs nm @{context} *}
|
||||
thm c_guard_to_word_ineq
|
||||
find_theorems name: twice
|
||||
find_theorems globals_list_valid
|
||||
|
||||
find_theorems name: ptr_inverse
|
||||
|
||||
ML {*
|
||||
|
||||
|
@ -93,6 +88,12 @@ schematic_lemma "PROP ?P"
|
|||
apply (tactic {* ALLGOALS (nth (graph_refine_proof_tacs @{context}) 0) *})
|
||||
apply (tactic {* ALLGOALS (nth (graph_refine_proof_tacs @{context}) 1) *})
|
||||
apply (tactic {* ALLGOALS (nth (graph_refine_proof_tacs @{context}) 2) *})
|
||||
|
||||
apply (tactic {* ALLGOALS (simp_tac (@{context} addsimps @{thms
|
||||
hrs_mem_update
|
||||
hrs_htd_globals_swap mex_def meq_def}
|
||||
addsimps globals_swap_rewrites2)) *})
|
||||
|
||||
apply (tactic {* ALLGOALS (nth (graph_refine_proof_tacs @{context}) 3) *})
|
||||
apply (tactic {* ALLGOALS (nth (graph_refine_proof_tacs @{context}) 4) *})
|
||||
apply (tactic {* ALLGOALS (nth (graph_refine_proof_tacs @{context}) 5) *})
|
||||
|
@ -101,9 +102,12 @@ schematic_lemma "PROP ?P"
|
|||
apply (simp_all add: field_h_val_rewrites field_to_bytes_rewrites heap_update_def
|
||||
to_bytes_array upt_rec take_heap_list_min drop_heap_list_general
|
||||
heap_update_list_append heap_list_update_ptr heap_list_update_word32
|
||||
store_store_word32_commute_offset
|
||||
store_store_word32_commute_offset field_simps
|
||||
heap_access_Array_element h_val_word32 h_val_ptr
|
||||
field_lvalue_offset_eq)
|
||||
|
||||
|
||||
|
||||
apply (auto simp: mex_def meq_def)
|
||||
done
|
||||
|
||||
|
|
|
@ -50,9 +50,9 @@ lemma intent_reset_projection[simp]:
|
|||
|
||||
lemma asid_reset_intent_reset_switch:
|
||||
"asid_reset (intent_reset x) = intent_reset (asid_reset x)"
|
||||
apply (cut_tac object_lift_def2[unfolded Fun.comp_def,simplified,symmetric])
|
||||
apply (cut_tac object_clean_def2[unfolded Fun.comp_def,simplified,symmetric])
|
||||
apply (drule_tac x = x in fun_cong)
|
||||
apply (simp add:object_lift_def)
|
||||
apply (simp add:object_clean_def)
|
||||
done
|
||||
|
||||
lemma reset_cap_twice[simp]:
|
||||
|
@ -69,16 +69,16 @@ lemma asid_reset_twice[simp]:
|
|||
simp_all add:asid_reset_def
|
||||
update_slots_def object_slots_def)
|
||||
|
||||
lemma object_lift_twice[simp]:
|
||||
"object_lift (object_lift x) = object_lift x"
|
||||
by (clarsimp simp:object_lift_def asid_reset_intent_reset_switch)
|
||||
lemma object_clean_twice[simp]:
|
||||
"object_clean (object_clean x) = object_clean x"
|
||||
by (clarsimp simp:object_clean_def asid_reset_intent_reset_switch)
|
||||
|
||||
|
||||
(* The following should be correct once we rule the intent out of our lift function *)
|
||||
lemma sep_nonimpact_valid_lift:
|
||||
assumes non_intent_impact:
|
||||
"\<And>ptr P A. \<lbrace>\<lambda>s. A (object_at (\<lambda>obj. P (object_lift obj)) ptr s)\<rbrace>
|
||||
f \<lbrace>\<lambda>rv s. A (object_at (\<lambda>obj. P (object_lift obj)) ptr s)\<rbrace>"
|
||||
"\<And>ptr P A. \<lbrace>\<lambda>s. A (object_at (\<lambda>obj. P (object_clean obj)) ptr s)\<rbrace>
|
||||
f \<lbrace>\<lambda>rv s. A (object_at (\<lambda>obj. P (object_clean obj)) ptr s)\<rbrace>"
|
||||
assumes non_irq_node_impact:
|
||||
"\<And>ptr P A. \<lbrace>\<lambda>s. A (cdl_irq_node s)\<rbrace> f \<lbrace>\<lambda>rv s. A (cdl_irq_node s)\<rbrace>"
|
||||
shows
|
||||
|
@ -97,10 +97,10 @@ lemma sep_nonimpact_valid_lift:
|
|||
apply (simp add:object_at_def opt_object_def)+
|
||||
apply (drule_tac P1 ="\<lambda>x. True" and A1 = "\<lambda>x. x" in use_valid[OF _ non_intent_impact])
|
||||
apply (fastforce simp:object_at_def opt_object_def)+
|
||||
apply (drule_tac P1 ="\<lambda>x. object_lift x = object_lift ab" and
|
||||
apply (drule_tac P1 ="\<lambda>x. object_clean x = object_clean ab" and
|
||||
A1 = "\<lambda>x. x" in use_valid[OF _ non_intent_impact])
|
||||
apply (fastforce simp:object_at_def opt_object_def)
|
||||
apply (clarsimp simp: object_at_def opt_object_def object_slots_object_lift
|
||||
apply (clarsimp simp: object_at_def opt_object_def object_slots_object_clean
|
||||
object_project_def
|
||||
split: cdl_component.splits)
|
||||
apply (rule ext)+
|
||||
|
@ -118,19 +118,19 @@ lemma mark_tcb_intent_error_sep_inv[wp]:
|
|||
apply (simp add:mark_tcb_intent_error_def update_thread_def
|
||||
set_object_def)
|
||||
apply (wp | wpc | simp add:set_object_def)+
|
||||
apply (clarsimp simp: object_at_def opt_object_def object_lift_def intent_reset_def
|
||||
apply (clarsimp simp: object_at_def opt_object_def object_clean_def intent_reset_def
|
||||
object_slots_def asid_reset_def update_slots_def)
|
||||
apply wp
|
||||
done
|
||||
|
||||
lemma corrupt_tcb_intent_sep_helper[wp]:
|
||||
"\<lbrace>\<lambda>s. A (object_at (\<lambda>obj. P (object_lift obj)) ptr s)\<rbrace> corrupt_tcb_intent thread
|
||||
\<lbrace>\<lambda>rv s. A (object_at (\<lambda>obj. P (object_lift obj)) ptr s)\<rbrace>"
|
||||
"\<lbrace>\<lambda>s. A (object_at (\<lambda>obj. P (object_clean obj)) ptr s)\<rbrace> corrupt_tcb_intent thread
|
||||
\<lbrace>\<lambda>rv s. A (object_at (\<lambda>obj. P (object_clean obj)) ptr s)\<rbrace>"
|
||||
apply (simp add:corrupt_tcb_intent_def update_thread_def
|
||||
set_object_def)
|
||||
apply (wp select_wp | wpc | simp add:set_object_def)+
|
||||
apply (clarsimp simp:object_at_def opt_object_def)
|
||||
apply (simp add:object_lift_def intent_reset_def
|
||||
apply (simp add:object_clean_def intent_reset_def
|
||||
object_slots_def asid_reset_def update_slots_def)
|
||||
done
|
||||
|
||||
|
@ -144,13 +144,13 @@ lemma corrupt_tcb_intent_sep_inv[wp]:
|
|||
done
|
||||
|
||||
lemma corrupt_frame_sep_helper[wp]:
|
||||
"\<lbrace>\<lambda>s. A (object_at (\<lambda>obj. P (object_lift obj)) ptr s)\<rbrace>
|
||||
corrupt_frame a \<lbrace>\<lambda>rv s. A (object_at (\<lambda>obj. P (object_lift obj)) ptr s)\<rbrace>"
|
||||
"\<lbrace>\<lambda>s. A (object_at (\<lambda>obj. P (object_clean obj)) ptr s)\<rbrace>
|
||||
corrupt_frame a \<lbrace>\<lambda>rv s. A (object_at (\<lambda>obj. P (object_clean obj)) ptr s)\<rbrace>"
|
||||
apply (simp add:corrupt_frame_def)
|
||||
apply (wp select_wp)
|
||||
apply (clarsimp simp:corrupt_intents_def object_at_def
|
||||
opt_object_def map_add_def split:option.splits cdl_object.splits)
|
||||
apply (simp add:object_lift_def intent_reset_def
|
||||
apply (simp add:object_clean_def intent_reset_def
|
||||
object_slots_def asid_reset_def update_slots_def)
|
||||
done
|
||||
|
||||
|
@ -174,7 +174,7 @@ lemma update_thread_intent_update:
|
|||
apply (simp add:update_thread_def set_object_def get_thread_def)
|
||||
apply (wp | wpc | simp add:object_at_def opt_object_def
|
||||
intent_reset_def set_object_def)+
|
||||
apply (simp add:object_lift_def intent_reset_def
|
||||
apply (simp add:object_clean_def intent_reset_def
|
||||
object_slots_def asid_reset_def update_slots_def)
|
||||
apply fastforce
|
||||
apply wp
|
||||
|
@ -573,18 +573,6 @@ lemma call_kernel_with_intent_no_fault_helper:
|
|||
apply auto
|
||||
done
|
||||
|
||||
lemma opt_cap_Some_sep:
|
||||
"opt_cap dest_slot s = Some cap
|
||||
\<Longrightarrow> object_at_heap
|
||||
(\<lambda>obj. \<exists>cap. object_slots (object_clean obj {Slot (snd dest_slot)}) = [snd dest_slot \<mapsto> cap])
|
||||
(fst dest_slot) (cdl_objects s)"
|
||||
apply (clarsimp simp:opt_cap_def split_def restrict_map_def
|
||||
object_at_heap_def opt_object_def clean_slots_def
|
||||
slots_of_def split:option.splits)
|
||||
apply (rule_tac x = cap in exI)
|
||||
apply auto
|
||||
done
|
||||
|
||||
|
||||
lemma invoke_cnode_insert_cap:
|
||||
"\<lbrace>\<lambda>s. < dest_slot \<mapsto>c - \<and>* R> s \<and> \<not> is_untyped_cap cap' \<rbrace>
|
||||
|
@ -602,14 +590,6 @@ lemma invoke_cnode_insert_cap:
|
|||
apply simp
|
||||
done
|
||||
|
||||
lemma object_clean_slots_CNode:
|
||||
"(object_clean_slots (object_lift (CNode fs)) {})
|
||||
= CNode (fs\<lparr>cdl_cnode_caps := Map.empty\<rparr>)"
|
||||
by (clarsimp simp:object_lift_def intent_reset_def
|
||||
update_slots_def object_clean_slots_def
|
||||
clean_slots_def
|
||||
asid_reset_def object_slots_def)
|
||||
|
||||
lemma invoke_cnode_move_wp:
|
||||
"\<lbrace><dest \<mapsto>c - \<and>* src \<mapsto>c cap \<and>* R>\<rbrace> invoke_cnode (MoveCall cap' src dest)
|
||||
\<lbrace>\<lambda>_. <dest \<mapsto>c cap' \<and>* src \<mapsto>c NullCap \<and>* R>\<rbrace>"
|
||||
|
@ -1104,11 +1084,11 @@ lemma invoke_cnode_insert_cap':
|
|||
done
|
||||
|
||||
lemma object_to_sep_state_slot:
|
||||
"obj_to_sep_state (fst ptr) {Slot (snd ptr)} obj = [(fst ptr, Slot (snd ptr)) \<mapsto>
|
||||
"obj_to_sep_state (fst ptr) obj {Slot (snd ptr)} = [(fst ptr, Slot (snd ptr)) \<mapsto>
|
||||
(CDL_Cap ((reset_cap_asid \<circ>\<^sub>M object_slots obj) (snd ptr)))]"
|
||||
apply (rule ext)
|
||||
apply (clarsimp simp:obj_to_sep_state_def
|
||||
object_project_slot object_slots_object_lift
|
||||
object_project_slot object_slots_object_clean
|
||||
split:if_splits)
|
||||
done
|
||||
|
||||
|
|
|
@ -159,7 +159,7 @@ lemma sep_any_map_c_imp: "(dest \<mapsto>c cap) s \<Longrightarrow> (dest \<maps
|
|||
|
||||
lemma obj_exists_map_i:
|
||||
"<obj_id \<mapsto>o obj \<and>* R> s \<Longrightarrow> \<exists>obj'. (opt_object obj_id s = Some obj'
|
||||
\<and> object_lift obj = object_lift obj')"
|
||||
\<and> object_clean obj = object_clean obj')"
|
||||
apply (clarsimp simp: opt_object_def sep_map_o_conj)
|
||||
apply (case_tac "cdl_objects s obj_id")
|
||||
apply (drule_tac x = "(obj_id,Fields)" in fun_cong)
|
||||
|
@ -172,8 +172,6 @@ lemma obj_exists_map_i:
|
|||
apply (clarsimp simp: lift_def obj_to_sep_state_def object_project_def
|
||||
object_at_heap_def state_sep_projection_def
|
||||
Let_unfold)
|
||||
apply (drule_tac f = object_fields in arg_cong)
|
||||
apply simp
|
||||
apply (rule ext)
|
||||
apply (drule_tac x= "(obj_id,Slot x)" in fun_cong)
|
||||
apply (clarsimp simp: lift_def obj_to_sep_state_def
|
||||
|
@ -197,10 +195,6 @@ lemma obj_exists_map_f:
|
|||
apply simp
|
||||
done
|
||||
|
||||
lemma object_clean_slots_no_slots [simp]:
|
||||
"\<not> has_slots obj \<Longrightarrow> object_clean_slots obj cmp = obj"
|
||||
by (clarsimp simp: object_clean_slots_def)
|
||||
|
||||
lemma object_slots_asid_reset:
|
||||
"object_slots (asid_reset obj) = reset_cap_asid \<circ>\<^sub>M (object_slots obj)"
|
||||
by (clarsimp simp: object_slots_def asid_reset_def update_slots_def
|
||||
|
@ -210,28 +204,6 @@ lemma reset_cap_asid_idem [simp]:
|
|||
"reset_cap_asid (reset_cap_asid cap) = reset_cap_asid cap"
|
||||
by (simp add: reset_cap_asid_def split: cdl_cap.splits)
|
||||
|
||||
lemma opt_cap_sep_imp_helper:
|
||||
"\<lbrakk>object_slots obj = [slot \<mapsto> cap];
|
||||
heap_lift (cdl_objects s) obj_id = Some obj';
|
||||
object_slots obj' slot = object_slots (object_lift (object_clean obj {Slot slot})) slot;
|
||||
object_type obj = object_type obj'\<rbrakk>
|
||||
\<Longrightarrow> \<exists>cap'. opt_cap (obj_id, slot) s = Some cap' \<and>
|
||||
reset_cap_asid cap' = reset_cap_asid cap"
|
||||
apply (case_tac "has_slots obj", simp_all)
|
||||
apply (clarsimp simp: heap_lift_def object_lift_def)
|
||||
apply (rename_tac obj')
|
||||
apply (rule_tac x="the (object_slots obj' slot)" in exI)
|
||||
apply (clarsimp simp: opt_cap_def slots_of_def opt_object_def)
|
||||
apply (clarsimp simp: object_type_def
|
||||
split: cdl_object.splits)
|
||||
apply (clarsimp simp: object_clean_def object_clean_fields_def object_clean_slots_def
|
||||
update_slots_def object_slots_def intent_reset_def asid_reset_def
|
||||
clean_slots_def cdl_tcb.splits map_comp'_def restrict_map_def
|
||||
has_slots_def
|
||||
split: option.splits)+
|
||||
done
|
||||
|
||||
|
||||
lemma opt_cap_sep_imp:
|
||||
"\<lbrakk><p \<mapsto>c cap \<and>* R> s\<rbrakk>
|
||||
\<Longrightarrow> \<exists>cap'. opt_cap p s = Some cap' \<and> reset_cap_asid cap' = reset_cap_asid cap"
|
||||
|
@ -240,7 +212,7 @@ lemma opt_cap_sep_imp:
|
|||
opt_object_def split_def
|
||||
sep_any_def sep_map_general_def slots_of_def
|
||||
state_sep_projection_def object_project_def
|
||||
object_slots_object_lift
|
||||
object_slots_object_clean
|
||||
Let_unfold split:sep_state.splits option.splits)
|
||||
done
|
||||
|
||||
|
@ -249,7 +221,7 @@ lemma opt_cap_sep_any_imp:
|
|||
apply (clarsimp simp: sep_any_exist
|
||||
opt_cap_def sep_map_c_conj Let_def)
|
||||
apply (clarsimp simp: sep_map_c_def lift_def
|
||||
opt_object_def split_def object_slots_object_lift
|
||||
opt_object_def split_def object_slots_object_clean
|
||||
sep_any_def sep_map_general_def slots_of_def
|
||||
state_sep_projection_def object_project_def
|
||||
Let_unfold split:sep_state.splits option.splits)
|
||||
|
@ -263,8 +235,8 @@ lemma sep_f_size_opt_cnode:
|
|||
apply (case_tac obj)
|
||||
apply (auto simp: intent_reset_def empty_cnode_def
|
||||
opt_cnode_def update_slots_def state_sep_projection_def
|
||||
opt_object_def object_clean_def object_clean_slots_def
|
||||
object_project_def object_lift_def asid_reset_def
|
||||
opt_object_def object_wipe_slots_def
|
||||
object_project_def object_clean_def asid_reset_def
|
||||
split:cdl_cap.splits cdl_object.splits)
|
||||
done
|
||||
|
||||
|
|
|
@ -261,7 +261,7 @@ lemma sep_spec_simps:
|
|||
apply (case_tac s')
|
||||
apply (clarsimp simp:obj_to_sep_state_def)
|
||||
apply (rule ext)
|
||||
apply (clarsimp simp: object_project_def object_slots_object_lift
|
||||
apply (clarsimp simp: object_project_def object_slots_object_clean
|
||||
split: split_if_asm)
|
||||
done
|
||||
|
||||
|
@ -409,16 +409,16 @@ lemma all_scheduable_tcbs_rewrite:
|
|||
apply (clarsimp simp:all_scheduable_tcbs_def state_sep_projection_def
|
||||
sep_all_scheduable_tcbs_def object_at_heap_def object_project_def
|
||||
is_tcb_obj_type)
|
||||
apply (clarsimp simp:object_type_def object_slots_object_lift
|
||||
apply (clarsimp simp:object_type_def object_slots_object_clean
|
||||
tcb_scheduable_def object_slots_def scheduable_cap_def)
|
||||
apply (fastforce simp:object_lift_def asid_reset_def update_slots_def
|
||||
apply (fastforce simp:object_clean_def asid_reset_def update_slots_def
|
||||
reset_cap_asid_def intent_reset_def object_slots_def
|
||||
split:if_splits)
|
||||
apply (clarsimp simp:all_scheduable_tcbs_def state_sep_projection_def
|
||||
sep_all_scheduable_tcbs_def object_at_heap_def object_project_def
|
||||
is_tcb_obj_type split:option.splits)
|
||||
apply (clarsimp simp:object_type_def tcb_scheduable_def
|
||||
scheduable_cap_def object_slots_def object_lift_def asid_reset_def
|
||||
scheduable_cap_def object_slots_def object_clean_def asid_reset_def
|
||||
update_slots_def intent_reset_def reset_cap_asid_def
|
||||
split:cdl_object.splits cdl_cap.splits option.splits)
|
||||
done
|
||||
|
@ -429,12 +429,6 @@ lemma update_slots_rev:
|
|||
by (clarsimp simp:update_slots_def object_slots_def
|
||||
split:cdl_object.splits)
|
||||
|
||||
lemma object_clean_rev:
|
||||
"\<lbrakk>object_clean obj cmps = obj'; Fields \<in> cmps\<rbrakk> \<Longrightarrow>
|
||||
\<exists>slots. obj = update_slots (object_slots obj) obj' "
|
||||
by (clarsimp simp:object_clean_def)
|
||||
|
||||
|
||||
lemma all_scheduable_tcbsD:
|
||||
"ptr \<in> all_scheduable_tcbs (cdl_objects s)
|
||||
\<Longrightarrow> tcb_at_heap tcb_scheduable ptr (cdl_objects s)"
|
||||
|
@ -470,7 +464,7 @@ lemma set_cap_all_scheduable_tcbs:
|
|||
apply (clarsimp simp:object_at_heap_def tcb_scheduable_def
|
||||
state_sep_projection_def object_project_def)
|
||||
apply (rule conjI)
|
||||
apply (clarsimp simp:object_slots_def object_lift_def
|
||||
apply (clarsimp simp:object_slots_def object_clean_def
|
||||
update_slots_def intent_reset_def asid_reset_def
|
||||
split:option.splits)
|
||||
apply fastforce+
|
||||
|
|
|
@ -215,16 +215,16 @@ lemma dummy_detype_if_untyped:
|
|||
apply (drule_tac x = "(x,Fields)" in fun_cong)+
|
||||
apply (case_tac xa,case_tac y)
|
||||
apply (clarsimp simp:plus_sep_state_def sep_state_add_def asid_reset_def
|
||||
intent_reset_def update_slots_def object_clean_slots_def sep_disj_sep_state_def
|
||||
obj_to_sep_state_def object_project_def object_lift_def
|
||||
intent_reset_def update_slots_def object_wipe_slots_def sep_disj_sep_state_def
|
||||
obj_to_sep_state_def object_project_def object_clean_def
|
||||
sep_state_disj_def)
|
||||
apply (drule map_disj_None_right'[rotated])
|
||||
apply simp
|
||||
apply (clarsimp simp: map_add_def split:option.splits)
|
||||
apply (case_tac z)
|
||||
apply (clarsimp simp:plus_sep_state_def sep_state_add_def asid_reset_def
|
||||
intent_reset_def update_slots_def object_clean_slots_def sep_disj_sep_state_def
|
||||
obj_to_sep_state_def object_project_def object_lift_def
|
||||
intent_reset_def update_slots_def object_wipe_slots_def sep_disj_sep_state_def
|
||||
obj_to_sep_state_def object_project_def object_clean_def
|
||||
sep_state_disj_def)+
|
||||
done
|
||||
|
||||
|
|
|
@ -63,8 +63,7 @@ done
|
|||
lemma example_drule:
|
||||
"(ptr \<mapsto>o obj) s
|
||||
\<Longrightarrow> (ptr \<mapsto>S obj \<and>* ptr \<mapsto>f obj) s"
|
||||
by (metis sep_conj_commute sep_conj_left_commute
|
||||
sep_map_o_decomp)
|
||||
by (metis sep_conj_commute sep_map_o_decomp)
|
||||
|
||||
lemma sep_drule_example:
|
||||
"(ptr \<mapsto>o obj \<and>* A \<and>* B ) s
|
||||
|
@ -78,8 +77,7 @@ done
|
|||
lemma example_rule:
|
||||
"(ptr \<mapsto>f obj \<and>* ptr \<mapsto>S obj) s
|
||||
\<Longrightarrow> (ptr \<mapsto>o obj) s"
|
||||
by (metis sep_conj_commute sep_conj_left_commute
|
||||
sep_map_o_decomp)
|
||||
by (metis sep_map_o_decomp)
|
||||
|
||||
lemma sep_rule_example: "(ptr \<mapsto>f obj \<and>* A \<and>* B \<and>* ptr \<mapsto>S obj ) s \<Longrightarrow> (ptr \<mapsto>o obj \<and>* A \<and>* B) s"
|
||||
apply (sep_rule example_rule)
|
||||
|
@ -277,7 +275,7 @@ lemma sep_rule_double_conjunct_example:
|
|||
"\<lbrakk>((obj_id, slot) \<mapsto>c cap \<and>* obj_id \<mapsto>f obj) s;
|
||||
object_slots obj slot = Some cap\<rbrakk>
|
||||
\<Longrightarrow> ((obj_id, slot) \<mapsto>s obj \<and>* obj_id \<mapsto>f obj) s"
|
||||
apply (sep_drule sep_map_s_sep_map_c')
|
||||
apply (sep_drule sep_map_s_sep_map_c)
|
||||
apply assumption
|
||||
apply (sep_cancel)+
|
||||
done
|
||||
|
|
|
@ -1231,7 +1231,7 @@ lemma (in kernel_m) cDomScheduleIdx_to_H_correct:
|
|||
assumes ms: "cstate_to_machine_H cs = ksMachineState as"
|
||||
shows "unat (ksDomScheduleIdx_' cs) = ksDomScheduleIdx as"
|
||||
using assms
|
||||
by (clarsimp simp: cstate_relation_def Let_def valid_state'_def newKernelState_def unat_of_nat_eq)
|
||||
by (clarsimp simp: cstate_relation_def Let_def valid_state'_def newKernelState_def unat_of_nat_eq cdom_schedule_relation_def)
|
||||
|
||||
definition cDomSchedule_to_H :: "(dschedule_C['b :: finite]) \<Rightarrow> (8 word \<times> 32 word) list" where
|
||||
"cDomSchedule_to_H cs \<equiv> THE as. cdom_schedule_relation as cs"
|
||||
|
|
|
@ -1359,92 +1359,150 @@ lemma word_add_format:
|
|||
"(-1::32 word) + b + c = b + (c - 1)"
|
||||
by simp
|
||||
|
||||
thm pteCheckIfMapped_body_def
|
||||
|
||||
lemma pteCheckIfMapped_ccorres:
|
||||
"ccorres (\<lambda>rv rv'. rv = to_bool rv') ret__unsigned_long_' \<top>
|
||||
(UNIV \<inter> {s. pte___ptr_to_struct_pte_C_' s = Ptr slot}) []
|
||||
(pteCheckIfMapped slot)
|
||||
(Call pteCheckIfMapped_'proc)"
|
||||
apply (cinit lift: pte___ptr_to_struct_pte_C_')
|
||||
apply (rule ccorres_pre_getObject_pte)
|
||||
apply (rule_tac P'="{s. \<exists>pte'. cslift s (pte_Ptr slot) = Some pte' \<and> cpte_relation rv pte'}"
|
||||
in ccorres_from_vcg_throws[where P="\<lambda>s. True"])
|
||||
apply simp_all
|
||||
apply clarsimp
|
||||
apply (rule conseqPre, vcg)
|
||||
apply (clarsimp simp: typ_heap_simps' return_def)
|
||||
apply (case_tac rv, simp_all add: to_bool_def cpte_relation_invalid isInvalidPTE_def
|
||||
split: split_if)
|
||||
done
|
||||
|
||||
lemma cpde_relation_invalid:
|
||||
"cpde_relation pdea pde \<Longrightarrow> (pde_get_tag pde = scast pde_pde_invalid) = isInvalidPDE pdea"
|
||||
apply (simp add: cpde_relation_def Let_def)
|
||||
apply (simp add: pde_pde_invalid_lift)
|
||||
apply (case_tac pdea, simp_all add: isInvalidPDE_def) [1]
|
||||
apply clarsimp
|
||||
apply (simp add: pde_pde_invalid_lift_def)
|
||||
done
|
||||
|
||||
lemma pdeCheckIfMapped_ccorres:
|
||||
"ccorres (\<lambda>rv rv'. rv = to_bool rv') ret__unsigned_long_' \<top>
|
||||
(UNIV \<inter> {s. pde___ptr_to_struct_pde_C_' s = Ptr slot}) []
|
||||
(pdeCheckIfMapped slot)
|
||||
(Call pdeCheckIfMapped_'proc)"
|
||||
apply (cinit lift: pde___ptr_to_struct_pde_C_')
|
||||
apply (rule ccorres_pre_getObject_pde)
|
||||
apply (rule_tac P'="{s. \<exists>pde'. cslift s (pde_Ptr slot) = Some pde' \<and> cpde_relation rv pde'}"
|
||||
in ccorres_from_vcg_throws[where P="\<lambda>s. True"])
|
||||
apply simp_all
|
||||
apply clarsimp
|
||||
apply (rule conseqPre, vcg)
|
||||
apply (clarsimp simp: typ_heap_simps' return_def)
|
||||
apply (case_tac rv, simp_all add: to_bool_def cpde_relation_invalid isInvalidPDE_def
|
||||
split: split_if)
|
||||
done
|
||||
|
||||
(* FIXME: move *)
|
||||
lemma performPageInvocationMapPTE_ccorres:
|
||||
"ccorres (K (K \<bottom>) \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
|
||||
(invs' and cte_at' slot and (\<lambda>_. valid_pte_slots'2 mapping))
|
||||
(invs' and cte_at' slot and (\<lambda>_. valid_pte_slots'2 mapping \<and> asid \<le> mask asid_bits))
|
||||
(UNIV \<inter> {s. pte_range_relation (snd (theLeft mapping)) (pte_entries_' s)}
|
||||
\<inter> {s. cpte_relation (fst (theLeft mapping)) (pte_' s)}
|
||||
\<inter> {s. ccap_relation cap (cap_' s)}
|
||||
\<inter> {s. ctSlot_' s = cte_Ptr slot}
|
||||
\<inter> {s. asid_' s = asid}
|
||||
\<inter> {s. isLeft mapping}) []
|
||||
(liftE (performPageInvocation (PageMap cap slot mapping)))
|
||||
(liftE (performPageInvocation (PageMap asid cap slot mapping)))
|
||||
(Call performPageInvocationMapPTE_'proc)"
|
||||
apply (rule ccorres_gen_asm2)
|
||||
apply (rule ccorres_gen_asm)
|
||||
apply (simp only: liftE_liftM ccorres_liftM_simp)
|
||||
apply (cinit lift: pte_entries_' pte_' cap_' ctSlot_' )
|
||||
apply (rule ccorres_gen_asm [where P ="snd (theLeft mapping)\<noteq>[]"])
|
||||
apply (cinit lift: pte_entries_' pte_' cap_' ctSlot_' asid_' )
|
||||
apply (rule ccorres_gen_asm [where P ="snd (theLeft mapping)\<noteq>[]"])
|
||||
apply (clarsimp simp: isLeft_def simp del: Collect_const)
|
||||
apply (ctac (no_vcg))
|
||||
apply csymbr
|
||||
apply (simp add: mapM_discarded whileAnno_def
|
||||
Collect_False
|
||||
del: Collect_const)
|
||||
apply (rule ccorres_split_nothrow_novcg)
|
||||
apply (rule_tac F="\<top>\<top>" in ccorres_mapM_x_while' [where i=0])
|
||||
apply clarsimp
|
||||
apply (rule ccorres_guard_imp2)
|
||||
apply (rule storePTE_Basic_ccorres', simp)
|
||||
apply clarsimp
|
||||
apply (ctac (no_vcg) add: pteCheckIfMapped_ccorres)
|
||||
apply csymbr
|
||||
apply (simp add: mapM_discarded whileAnno_def
|
||||
Collect_False
|
||||
del: Collect_const)
|
||||
apply (rule ccorres_split_nothrow_novcg)
|
||||
apply (rule_tac F="\<top>\<top>" in ccorres_mapM_x_while' [where i=0])
|
||||
apply clarsimp
|
||||
apply (rule ccorres_guard_imp2)
|
||||
apply (rule storePTE_Basic_ccorres', simp)
|
||||
apply clarsimp
|
||||
apply (clarsimp simp: pte_range_relation_def ptr_range_to_list_def)
|
||||
apply (clarsimp simp: pte_range_relation_def ptr_range_to_list_def)
|
||||
apply (clarsimp simp: pte_range_relation_def ptr_range_to_list_def)
|
||||
apply vcg
|
||||
apply vcg
|
||||
apply simp
|
||||
apply wp
|
||||
apply clarsimp
|
||||
apply (clarsimp simp: pte_range_relation_def ptr_range_to_list_def)
|
||||
apply (rule order_less_le_trans)
|
||||
apply (rule unat_lt2p)
|
||||
apply (simp add: word_bits_def)
|
||||
apply ceqv
|
||||
apply csymbr
|
||||
apply (rule ccorres_Guard_Seq)
|
||||
apply (rule ccorres_Guard_Seq)
|
||||
apply (rule ccorres_move_c_guard_pte)+
|
||||
apply (ctac (no_vcg) add: cleanCacheRange_PoU_ccorres)
|
||||
apply (simp only: when_def)
|
||||
apply (rule ccorres_Cond_rhs_Seq)
|
||||
apply (simp add: to_bool_def)
|
||||
apply (rule ccorres_add_return2)
|
||||
apply (ctac (no_vcg) add: invalidateTLBByASID_ccorres)
|
||||
apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg_throws)
|
||||
apply (rule allI, rule conseqPre, vcg)
|
||||
apply (clarsimp simp:return_def)
|
||||
apply (rule wp_post_taut)
|
||||
apply (simp add: to_bool_def)
|
||||
apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg_throws)
|
||||
apply (rule allI, rule conseqPre, vcg)
|
||||
apply (clarsimp simp:return_def)
|
||||
apply (wp hoare_vcg_const_imp_lift) [1]
|
||||
apply (clarsimp simp: to_bool_def)
|
||||
apply (rule hoare_strengthen_post)
|
||||
apply (rule hoare_vcg_conj_lift[where Q'="\<lambda>rv s. valid_pde_mappings' s"])
|
||||
apply (rule mapM_x_accumulate_checks)
|
||||
apply (simp add: storePTE_def)
|
||||
apply (rule obj_at_setObject3)
|
||||
apply simp
|
||||
apply wp
|
||||
apply clarsimp
|
||||
apply (clarsimp simp: pte_range_relation_def ptr_range_to_list_def)
|
||||
apply (rule order_less_le_trans)
|
||||
apply (rule unat_lt2p)
|
||||
apply (simp add: word_bits_def)
|
||||
apply ceqv
|
||||
apply (rule ccorres_add_return2)
|
||||
apply csymbr
|
||||
apply (rule ccorres_Guard_Seq)
|
||||
apply (rule ccorres_Guard_Seq)
|
||||
apply (rule ccorres_move_c_guard_pte)+
|
||||
apply (ctac (no_vcg) add: cleanCacheRange_PoU_ccorres)
|
||||
apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg_throws)
|
||||
apply (rule allI, rule conseqPre, vcg)
|
||||
apply (clarsimp simp:return_def)
|
||||
apply wp [1]
|
||||
apply (rule hoare_strengthen_post)
|
||||
apply (rule mapM_x_accumulate_checks)
|
||||
apply (simp add: storePTE_def)
|
||||
apply (rule obj_at_setObject3)
|
||||
apply (simp add: objBits_simps archObjSize_def)
|
||||
apply (simp add: typ_at_to_obj_at_arches[symmetric])
|
||||
apply (wp mapM_x_wp_inv)
|
||||
apply clarsimp
|
||||
apply (simp add: typ_at_to_obj_at_arches)
|
||||
apply (frule bspec, erule hd_in_set)
|
||||
apply (drule bspec, erule last_in_set)
|
||||
apply (simp add: hd_conv_nth last_conv_nth)
|
||||
apply (rule iffD1[OF conj_assoc])+
|
||||
apply (rule conjI)
|
||||
(* the inequality first *)
|
||||
apply (clarsimp simp:valid_pte_slots'2_def
|
||||
objBits_simps archObjSize_def hd_conv_nth)
|
||||
apply (clarsimp simp:pte_range_relation_def ptr_range_to_list_def ptr_add_def)
|
||||
apply (simp add:add_assoc[symmetric] add_commute)
|
||||
apply (simp add: word_add_format)
|
||||
apply (frule is_aligned_addrFromPPtr_n,simp)
|
||||
apply (cut_tac n = "sz+2" in power_not_zero[where 'a="32"])
|
||||
apply simp
|
||||
apply (simp add: objBits_simps archObjSize_def)
|
||||
apply (simp add: typ_at_to_obj_at_arches[symmetric])
|
||||
apply wp
|
||||
apply clarsimp
|
||||
apply (simp add: typ_at_to_obj_at_arches)
|
||||
apply (frule bspec, erule hd_in_set)
|
||||
apply (drule bspec, erule last_in_set)
|
||||
apply (simp add: hd_conv_nth last_conv_nth)
|
||||
apply (rule iffD1[OF conj_assoc])+
|
||||
apply (rule conjI)
|
||||
(* the inequality first *)
|
||||
apply (clarsimp simp:valid_pte_slots'2_def
|
||||
objBits_simps archObjSize_def hd_conv_nth)
|
||||
apply (clarsimp simp:pte_range_relation_def ptr_range_to_list_def ptr_add_def)
|
||||
apply (simp add:add_assoc[symmetric] add_commute)
|
||||
apply (simp add: word_add_format)
|
||||
apply (frule is_aligned_addrFromPPtr_n,simp)
|
||||
apply (cut_tac n = "sz+2" in power_not_zero[where 'a="32"])
|
||||
apply simp
|
||||
apply (rule conjI)
|
||||
apply (erule is_aligned_no_wrap')
|
||||
apply (simp add:mult_commute power_not_zero)
|
||||
apply (rule conjI)
|
||||
apply (erule is_aligned_no_wrap')
|
||||
apply (simp add:mult_commute power_not_zero)
|
||||
apply (simp add:addrFromPPtr_mask_5)
|
||||
apply (clarsimp simp:pte_range_relation_def ptr_add_def ptr_range_to_list_def)
|
||||
apply (clarsimp simp: guard_is_UNIV_def Collect_const_mem hd_conv_nth last_conv_nth ucast_minus)
|
||||
apply (clarsimp simp: pte_range_relation_def ptr_range_to_list_def objBits_simps archObjSize_def)
|
||||
apply (simp add: CTypesDefs.ptr_add_def ucast_nat_def word_0_sle_from_less)
|
||||
apply wp
|
||||
apply simp
|
||||
apply (rule conjI)
|
||||
apply (erule is_aligned_no_wrap')
|
||||
apply (simp add:mult_commute power_not_zero)
|
||||
apply (rule conjI)
|
||||
apply (erule is_aligned_no_wrap')
|
||||
apply (simp add:mult_commute power_not_zero)
|
||||
apply (simp add:addrFromPPtr_mask_5)
|
||||
apply (clarsimp simp:pte_range_relation_def ptr_add_def ptr_range_to_list_def)
|
||||
apply (clarsimp simp: guard_is_UNIV_def Collect_const_mem hd_conv_nth last_conv_nth ucast_minus)
|
||||
apply (clarsimp simp: pte_range_relation_def ptr_range_to_list_def objBits_simps archObjSize_def)
|
||||
apply (simp add: CTypesDefs.ptr_add_def ucast_nat_def word_0_sle_from_less)
|
||||
apply ((wp | simp add: pteCheckIfMapped_def)+)
|
||||
apply (clarsimp simp: pte_range_relation_def map_is_Nil_conv hd_map_simp ptr_range_to_list_def)
|
||||
apply (clarsimp simp: pte_range_relation_def ptr_range_to_list_def unat_1_0)
|
||||
done
|
||||
|
||||
|
@ -1566,261 +1624,316 @@ lemma of_nat_ucast:
|
|||
apply (simp add: is_down_def target_size_def source_size_def word_size)
|
||||
done
|
||||
|
||||
lemma mapM_x_storePDE_pde_mappings':
|
||||
"\<lbrakk>valid_pde_slots'2 (Inr (a, b)); b \<noteq> []\<rbrakk> \<Longrightarrow> \<lbrace>valid_pde_mappings'\<rbrace>
|
||||
mapM_x (\<lambda>b. storePDE b a) b \<lbrace>\<lambda>rv. valid_pde_mappings'\<rbrace>"
|
||||
apply (wp mapM_x_wp')
|
||||
apply (auto simp: valid_pde_slots'2_def valid_pde_mapping'_def)
|
||||
done
|
||||
|
||||
|
||||
lemma performPageInvocationMapPDE_ccorres:
|
||||
"ccorres (K (K \<bottom>) \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
|
||||
(invs' and cte_at' slot and (\<lambda>_. valid_pde_slots'2 mapping))
|
||||
(invs' and cte_at' slot and (\<lambda>_. valid_pde_slots'2 mapping \<and> asid \<le> mask asid_bits))
|
||||
(UNIV \<inter> {s. pde_range_relation (snd (theRight mapping)) (pde_entries_' s)}
|
||||
\<inter> {s. cpde_relation (fst (theRight mapping)) (pde_' s)}
|
||||
\<inter> {s. ccap_relation cap (cap_' s)}
|
||||
\<inter> {s. ctSlot_' s = cte_Ptr slot}
|
||||
\<inter> {s. asid_' s = asid}
|
||||
\<inter> {s. isRight mapping}) []
|
||||
(liftE (performPageInvocation (PageMap cap slot mapping)))
|
||||
(liftE (performPageInvocation (PageMap asid cap slot mapping)))
|
||||
(Call performPageInvocationMapPDE_'proc)"
|
||||
apply (rule ccorres_gen_asm)
|
||||
apply (rule ccorres_gen_asm2)
|
||||
apply (simp only: liftE_liftM ccorres_liftM_simp)
|
||||
apply (cinit lift: pde_entries_' pde_' cap_' ctSlot_' )
|
||||
apply (rule ccorres_gen_asm [where P ="snd (theRight mapping)\<noteq>[]"])
|
||||
apply (cinit lift: pde_entries_' pde_' cap_' ctSlot_' asid_' )
|
||||
apply (rule ccorres_gen_asm [where P ="snd (theRight mapping)\<noteq>[]"])
|
||||
apply (clarsimp simp: isRight_def simp del: Collect_const)
|
||||
apply (ctac (no_vcg))
|
||||
apply csymbr
|
||||
apply (simp add: mapM_discarded whileAnno_def
|
||||
Collect_False
|
||||
del: Collect_const)
|
||||
apply (rule ccorres_split_nothrow_novcg)
|
||||
apply (rule_tac F="\<top>\<top>" in ccorres_mapM_x_while' [where i=0])
|
||||
apply clarsimp
|
||||
apply (rule ccorres_guard_imp2)
|
||||
apply (rule storePDE_Basic_ccorres', simp)
|
||||
apply clarsimp
|
||||
apply (clarsimp simp: pde_range_relation_def ptr_range_to_list_def
|
||||
valid_pde_slots'2_def ucast_nat_def ucast_minus
|
||||
of_nat_ucast)
|
||||
apply (clarsimp simp: pde_range_relation_def ptr_range_to_list_def)
|
||||
apply vcg
|
||||
apply (ctac (no_vcg) add: pdeCheckIfMapped_ccorres)
|
||||
apply csymbr
|
||||
apply (simp add: mapM_discarded whileAnno_def
|
||||
Collect_False
|
||||
del: Collect_const)
|
||||
apply (rule ccorres_split_nothrow_novcg)
|
||||
apply (rule_tac F="\<top>\<top>" in ccorres_mapM_x_while' [where i=0])
|
||||
apply clarsimp
|
||||
apply (rule ccorres_guard_imp2)
|
||||
apply (rule storePDE_Basic_ccorres', simp)
|
||||
apply clarsimp
|
||||
apply (clarsimp simp: pde_range_relation_def ptr_range_to_list_def
|
||||
valid_pde_slots'2_def ucast_nat_def ucast_minus
|
||||
of_nat_ucast)
|
||||
apply (clarsimp simp: pde_range_relation_def ptr_range_to_list_def)
|
||||
apply vcg
|
||||
apply simp
|
||||
apply wp
|
||||
apply clarsimp
|
||||
apply (clarsimp simp: pde_range_relation_def ptr_range_to_list_def)
|
||||
apply (rule order_less_le_trans)
|
||||
apply (rule unat_lt2p)
|
||||
apply (simp add: word_bits_def)
|
||||
apply ceqv
|
||||
apply csymbr
|
||||
apply (rule ccorres_Guard_Seq,rule ccorres_Guard_Seq)
|
||||
apply (rule ccorres_move_c_guard_pde)+
|
||||
apply (ctac (no_vcg) add: cleanCacheRange_PoU_ccorres)
|
||||
apply (simp only: when_def)
|
||||
apply (rule ccorres_Cond_rhs_Seq)
|
||||
apply (simp add: to_bool_def)
|
||||
apply (rule ccorres_add_return2)
|
||||
apply (ctac (no_vcg) add: invalidateTLBByASID_ccorres)
|
||||
apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg_throws)
|
||||
apply (rule allI, rule conseqPre, vcg)
|
||||
apply (clarsimp simp:return_def)
|
||||
apply wp [1]
|
||||
apply (simp add: to_bool_def)
|
||||
apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg_throws)
|
||||
apply (rule allI, rule conseqPre, vcg)
|
||||
apply (clarsimp simp:return_def)
|
||||
apply (wp hoare_vcg_const_imp_lift) [1]
|
||||
apply (clarsimp simp: to_bool_def)
|
||||
apply (rule hoare_strengthen_post)
|
||||
apply (rule hoare_vcg_conj_lift[where Q'="\<lambda>rv s. valid_pde_mappings' s"])
|
||||
apply (rule mapM_x_accumulate_checks)
|
||||
apply (simp add: storePDE_def)
|
||||
apply (rule obj_at_setObject3)
|
||||
apply simp
|
||||
apply wp
|
||||
apply clarsimp
|
||||
apply (clarsimp simp: pde_range_relation_def ptr_range_to_list_def)
|
||||
apply (rule order_less_le_trans)
|
||||
apply (rule unat_lt2p)
|
||||
apply (simp add: word_bits_def)
|
||||
apply ceqv
|
||||
apply (rule ccorres_add_return2)
|
||||
apply csymbr
|
||||
apply (rule ccorres_Guard_Seq,rule ccorres_Guard_Seq)
|
||||
apply (rule ccorres_move_c_guard_pde)+
|
||||
apply (ctac (no_vcg) add: cleanCacheRange_PoU_ccorres)
|
||||
apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg_throws)
|
||||
apply (rule allI, rule conseqPre, vcg)
|
||||
apply (clarsimp simp:return_def)
|
||||
apply wp [1]
|
||||
apply (rule hoare_strengthen_post)
|
||||
apply (rule mapM_x_accumulate_checks)
|
||||
apply (simp add: storePDE_def)
|
||||
apply (rule obj_at_setObject3)
|
||||
apply (simp add: objBits_simps archObjSize_def)
|
||||
apply (simp add: typ_at_to_obj_at_arches[symmetric])
|
||||
apply (wp mapM_x_storePDE_pde_mappings')
|
||||
apply simp
|
||||
apply (simp add: objBits_simps archObjSize_def)
|
||||
apply (simp add: typ_at_to_obj_at_arches[symmetric])
|
||||
apply wp
|
||||
apply clarsimp
|
||||
apply (simp add: typ_at_to_obj_at_arches)
|
||||
apply (frule bspec, erule hd_in_set)
|
||||
apply (drule bspec, erule last_in_set)
|
||||
apply (simp add: hd_conv_nth last_conv_nth)
|
||||
apply (rule conj_assoc[THEN iffD1])+
|
||||
apply (rule conjI)
|
||||
(* the inequality first *)
|
||||
apply (clarsimp simp:valid_pde_slots'2_def
|
||||
objBits_simps archObjSize_def hd_conv_nth)
|
||||
apply (clarsimp simp:pde_range_relation_def ptr_range_to_list_def ptr_add_def)
|
||||
apply (simp add:add_assoc[symmetric] add_commute)
|
||||
apply (simp add: word_add_format)
|
||||
apply (frule is_aligned_addrFromPPtr_n,simp)
|
||||
apply (cut_tac n = "sz+2" in power_not_zero[where 'a="32"])
|
||||
apply simp
|
||||
apply clarsimp
|
||||
apply (simp add: typ_at_to_obj_at_arches)
|
||||
apply (frule bspec, erule hd_in_set)
|
||||
apply (drule bspec, erule last_in_set)
|
||||
apply (simp add: hd_conv_nth last_conv_nth)
|
||||
apply (rule conj_assoc[THEN iffD1])+
|
||||
apply (rule conjI)
|
||||
apply (erule is_aligned_no_wrap')
|
||||
apply (simp add:mult_commute power_not_zero)
|
||||
apply (rule conjI)
|
||||
apply (erule is_aligned_no_wrap')
|
||||
apply (simp add:mult_commute power_not_zero)
|
||||
apply (simp add:addrFromPPtr_mask_5)
|
||||
apply (clarsimp simp: pde_range_relation_def ptr_range_to_list_def CTypesDefs.ptr_add_def)
|
||||
apply (clarsimp simp: guard_is_UNIV_def Collect_const_mem hd_conv_nth last_conv_nth)
|
||||
apply (clarsimp simp: pde_range_relation_def ptr_range_to_list_def)
|
||||
apply (simp add: CTypesDefs.ptr_add_def ucast_nat_def ucast_minus objBits_simps archObjSize_def
|
||||
word_0_sle_from_less)
|
||||
apply wp
|
||||
apply simp
|
||||
(* the inequality first *)
|
||||
apply (clarsimp simp:valid_pde_slots'2_def
|
||||
objBits_simps archObjSize_def hd_conv_nth)
|
||||
apply (clarsimp simp:pde_range_relation_def ptr_range_to_list_def ptr_add_def)
|
||||
apply (simp add:add_assoc[symmetric] add_commute)
|
||||
apply (simp add: word_add_format)
|
||||
apply (frule is_aligned_addrFromPPtr_n,simp)
|
||||
apply (cut_tac n = "sz+2" in power_not_zero[where 'a="32"])
|
||||
apply simp
|
||||
apply (rule conjI)
|
||||
apply (erule is_aligned_no_wrap')
|
||||
apply (simp add:mult_commute power_not_zero)
|
||||
apply (rule conjI)
|
||||
apply (erule is_aligned_no_wrap')
|
||||
apply (simp add:mult_commute power_not_zero)
|
||||
apply (simp add:addrFromPPtr_mask_5)
|
||||
apply (clarsimp simp: pde_range_relation_def ptr_range_to_list_def CTypesDefs.ptr_add_def)
|
||||
apply (clarsimp simp: guard_is_UNIV_def Collect_const_mem hd_conv_nth last_conv_nth)
|
||||
apply (clarsimp simp: pde_range_relation_def ptr_range_to_list_def)
|
||||
apply (simp add: CTypesDefs.ptr_add_def ucast_nat_def ucast_minus objBits_simps archObjSize_def
|
||||
word_0_sle_from_less)
|
||||
apply (wp | simp add: pdeCheckIfMapped_def)+
|
||||
apply (clarsimp simp: pde_range_relation_def map_is_Nil_conv hd_map_simp ptr_range_to_list_def)
|
||||
apply (clarsimp simp: pde_range_relation_def ptr_range_to_list_def unat_1_0)
|
||||
done
|
||||
|
||||
lemma performPageInvocationRemapPDE_ccorres:
|
||||
"ccorres (K (K \<bottom>) \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
|
||||
(invs' and (\<lambda>_. valid_pde_slots'2 mapping))
|
||||
(invs' and (\<lambda>_. valid_pde_slots'2 mapping \<and> asid \<le> mask asid_bits))
|
||||
(UNIV \<inter> {s. pde_range_relation (snd (theRight mapping)) (pde_entries_' s)}
|
||||
\<inter> {s. cpde_relation (fst (theRight mapping)) (pde_' s)}
|
||||
\<inter> {s. asid_' s = asid}
|
||||
\<inter> {s. isRight mapping}) []
|
||||
(liftE (performPageInvocation (PageRemap mapping)))
|
||||
(liftE (performPageInvocation (PageRemap asid mapping)))
|
||||
(Call performPageInvocationRemapPDE_'proc)"
|
||||
apply (rule ccorres_gen_asm)
|
||||
apply (rule ccorres_gen_asm2)
|
||||
apply (simp only: liftE_liftM ccorres_liftM_simp)
|
||||
apply (cinit lift: pde_entries_' pde_')
|
||||
apply (rule ccorres_gen_asm [where P ="snd (theRight mapping)\<noteq>[]"])
|
||||
apply (clarsimp simp: isRight_def)
|
||||
apply (cinit lift: pde_entries_' pde_' asid_')
|
||||
apply (rule_tac P="b\<noteq>[]" in ccorres_gen_asm)
|
||||
apply (clarsimp simp: isRight_def simp del: Collect_const)
|
||||
apply csymbr
|
||||
apply (simp add: mapM_discarded whileAnno_def
|
||||
Collect_False
|
||||
del: Collect_const)
|
||||
apply (rule ccorres_split_nothrow_novcg)
|
||||
apply (rule_tac F="\<top>\<top>" in ccorres_mapM_x_while' [where i=0])
|
||||
apply clarsimp
|
||||
apply (rule ccorres_guard_imp2)
|
||||
apply (rule storePDE_Basic_ccorres', simp)
|
||||
apply clarsimp
|
||||
apply (clarsimp simp: pde_range_relation_def ptr_range_to_list_def
|
||||
valid_pde_slots'2_def)
|
||||
apply (clarsimp simp: pde_range_relation_def ptr_range_to_list_def)
|
||||
apply vcg
|
||||
apply simp
|
||||
apply wp
|
||||
apply clarsimp
|
||||
apply (clarsimp simp: pde_range_relation_def ptr_range_to_list_def)
|
||||
apply (rule order_less_le_trans)
|
||||
apply (rule unat_lt2p)
|
||||
apply (simp add: word_bits_def)
|
||||
apply ceqv
|
||||
apply (rule ccorres_add_return2)
|
||||
apply csymbr
|
||||
apply (rule ccorres_Guard_Seq,rule ccorres_Guard_Seq)
|
||||
apply (rule ccorres_move_c_guard_pde)+
|
||||
apply (ctac (no_vcg) add: cleanCacheRange_PoU_ccorres)
|
||||
apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg_throws)
|
||||
apply (rule allI, rule conseqPre, vcg)
|
||||
apply (clarsimp simp:return_def)
|
||||
apply wp [1]
|
||||
apply (rule hoare_strengthen_post)
|
||||
apply (ctac(no_vcg) add: pdeCheckIfMapped_ccorres)
|
||||
apply csymbr
|
||||
apply (simp add: mapM_discarded whileAnno_def
|
||||
Collect_False
|
||||
del: Collect_const)
|
||||
apply (rule ccorres_split_nothrow_novcg)
|
||||
apply (rule_tac F="\<top>\<top>" in ccorres_mapM_x_while' [where i=0])
|
||||
apply clarsimp
|
||||
apply (rule ccorres_guard_imp2)
|
||||
apply (rule storePDE_Basic_ccorres', simp)
|
||||
apply clarsimp
|
||||
apply (clarsimp simp: pde_range_relation_def ptr_range_to_list_def
|
||||
valid_pde_slots'2_def)
|
||||
apply (clarsimp simp: pde_range_relation_def ptr_range_to_list_def)
|
||||
apply vcg
|
||||
apply simp
|
||||
apply wp
|
||||
apply clarsimp
|
||||
apply (clarsimp simp: pde_range_relation_def ptr_range_to_list_def)
|
||||
apply (rule order_less_le_trans)
|
||||
apply (rule unat_lt2p)
|
||||
apply (simp add: word_bits_def)
|
||||
apply ceqv
|
||||
apply csymbr
|
||||
apply (rule ccorres_Guard_Seq,rule ccorres_Guard_Seq)
|
||||
apply (rule ccorres_move_c_guard_pde)+
|
||||
apply (ctac (no_vcg) add: cleanCacheRange_PoU_ccorres)
|
||||
apply (simp only: when_def)
|
||||
apply (rule ccorres_Cond_rhs_Seq)
|
||||
apply (simp add: to_bool_def)
|
||||
apply (rule ccorres_add_return2)
|
||||
apply (ctac(no_vcg) add: invalidateTLBByASID_ccorres)
|
||||
apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg_throws)
|
||||
apply (rule allI, rule conseqPre, vcg)
|
||||
apply (clarsimp simp:return_def)
|
||||
apply wp [1]
|
||||
apply (simp add: to_bool_def)
|
||||
apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg_throws)
|
||||
apply (rule allI, rule conseqPre, vcg)
|
||||
apply (clarsimp simp:return_def)
|
||||
apply (wp hoare_vcg_const_imp_lift) [1]
|
||||
apply (clarsimp simp: to_bool_def)
|
||||
apply (rule hoare_strengthen_post)
|
||||
apply (rule hoare_vcg_conj_lift[where Q'="\<lambda>rv s. valid_pde_mappings' s"])
|
||||
apply (rule mapM_x_accumulate_checks)
|
||||
apply (simp add: storePDE_def)
|
||||
apply (rule obj_at_setObject3)
|
||||
apply simp
|
||||
apply (simp add: objBits_simps archObjSize_def)
|
||||
apply (simp add: typ_at_to_obj_at_arches[symmetric])
|
||||
apply wp
|
||||
apply clarsimp
|
||||
apply (simp add: typ_at_to_obj_at_arches)
|
||||
apply (frule bspec, erule hd_in_set)
|
||||
apply (drule bspec, erule last_in_set)
|
||||
apply (simp add: hd_conv_nth last_conv_nth)
|
||||
apply (rule conj_assoc[THEN iffD1])+
|
||||
apply (rule conjI)
|
||||
apply (wp mapM_x_storePDE_pde_mappings')
|
||||
apply simp
|
||||
apply (simp)
|
||||
apply clarsimp
|
||||
apply (simp add: typ_at_to_obj_at_arches)
|
||||
apply (frule bspec, erule hd_in_set)
|
||||
apply (drule bspec, erule last_in_set)
|
||||
apply (simp add: hd_conv_nth last_conv_nth)
|
||||
apply (rule conj_assoc[THEN iffD1])+
|
||||
apply (rule conjI)
|
||||
(* the inequality first *)
|
||||
apply (clarsimp simp:valid_pde_slots'2_def
|
||||
objBits_simps archObjSize_def hd_conv_nth)
|
||||
apply (clarsimp simp:pde_range_relation_def ptr_range_to_list_def ptr_add_def)
|
||||
apply (simp add:add_assoc[symmetric] add_commute)
|
||||
apply (simp add: word_add_format)
|
||||
apply (frule is_aligned_addrFromPPtr_n,simp)
|
||||
apply (cut_tac n = "sz+2" in power_not_zero[where 'a="32"])
|
||||
apply simp
|
||||
apply (rule conjI)
|
||||
apply (erule is_aligned_no_wrap')
|
||||
apply (simp add:mult_commute power_not_zero)
|
||||
apply (rule conjI)
|
||||
apply (erule is_aligned_no_wrap')
|
||||
apply (simp add:mult_commute power_not_zero)
|
||||
apply (simp add:addrFromPPtr_mask_5)
|
||||
apply (clarsimp simp: pde_range_relation_def ptr_range_to_list_def CTypesDefs.ptr_add_def)
|
||||
apply (clarsimp simp: guard_is_UNIV_def Collect_const_mem hd_conv_nth last_conv_nth)
|
||||
apply (clarsimp simp: pde_range_relation_def ptr_range_to_list_def)
|
||||
apply (simp add: CTypesDefs.ptr_add_def ucast_minus ucast_nat_def
|
||||
objBits_simps archObjSize_def word_0_sle_from_less)
|
||||
apply (clarsimp simp: pde_range_relation_def ptr_range_to_list_def unat_1_0 )
|
||||
apply (clarsimp simp:valid_pde_slots'2_def
|
||||
objBits_simps archObjSize_def hd_conv_nth)
|
||||
apply (clarsimp simp:pde_range_relation_def ptr_range_to_list_def ptr_add_def)
|
||||
apply (simp add:add_assoc[symmetric] add_commute)
|
||||
apply (simp add: word_add_format)
|
||||
apply (frule is_aligned_addrFromPPtr_n,simp)
|
||||
apply (cut_tac n = "sz+2" in power_not_zero[where 'a="32"])
|
||||
apply simp
|
||||
apply (rule conjI)
|
||||
apply (erule is_aligned_no_wrap')
|
||||
apply (simp add:mult_commute power_not_zero)
|
||||
apply (rule conjI)
|
||||
apply (erule is_aligned_no_wrap')
|
||||
apply (simp add:mult_commute power_not_zero)
|
||||
apply (simp add:addrFromPPtr_mask_5)
|
||||
apply (clarsimp simp: pde_range_relation_def ptr_range_to_list_def CTypesDefs.ptr_add_def)
|
||||
apply (clarsimp simp: guard_is_UNIV_def Collect_const_mem hd_conv_nth last_conv_nth)
|
||||
apply (clarsimp simp: pde_range_relation_def ptr_range_to_list_def)
|
||||
apply (simp add: CTypesDefs.ptr_add_def ucast_minus ucast_nat_def
|
||||
objBits_simps archObjSize_def word_0_sle_from_less)
|
||||
apply (wp | simp add: pdeCheckIfMapped_def)+
|
||||
apply (clarsimp simp: pde_range_relation_def ptr_range_to_list_def unat_1_0 hd_map_simp)
|
||||
done
|
||||
|
||||
lemma performPageInvocationRemapPTE_ccorres:
|
||||
"ccorres (K (K \<bottom>) \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
|
||||
(invs' and (\<lambda>_. valid_pte_slots'2 mapping))
|
||||
(invs' and (\<lambda>_. valid_pte_slots'2 mapping \<and> asid \<le> mask asid_bits))
|
||||
(UNIV \<inter> {s. pte_range_relation (snd (theLeft mapping)) (pte_entries_' s)}
|
||||
\<inter> {s. cpte_relation (fst (theLeft mapping)) (pte_' s)}
|
||||
\<inter> {s. asid_' s = asid}
|
||||
\<inter> {s. isLeft mapping}) []
|
||||
(liftE (performPageInvocation (PageRemap mapping)))
|
||||
(liftE (performPageInvocation (PageRemap asid mapping)))
|
||||
(Call performPageInvocationRemapPTE_'proc)"
|
||||
apply (rule ccorres_gen_asm2)
|
||||
apply (rule ccorres_gen_asm)
|
||||
apply (simp only: liftE_liftM ccorres_liftM_simp)
|
||||
apply (cinit lift: pte_entries_' pte_')
|
||||
apply (rule ccorres_gen_asm [where P ="snd (theLeft mapping)\<noteq>[]"])
|
||||
apply (clarsimp simp: isLeft_def)
|
||||
apply (cinit lift: pte_entries_' pte_' asid_')
|
||||
apply (rule_tac P="b \<noteq> []" in ccorres_gen_asm)
|
||||
apply (clarsimp simp: isLeft_def simp del: Collect_const)
|
||||
apply csymbr
|
||||
apply (simp add: mapM_discarded whileAnno_def
|
||||
Collect_False
|
||||
del: Collect_const)
|
||||
apply (rule ccorres_split_nothrow_novcg)
|
||||
apply (rule_tac F="\<top>\<top>" in ccorres_mapM_x_while' [where i=0])
|
||||
apply clarsimp
|
||||
apply (rule ccorres_guard_imp2)
|
||||
apply (rule storePTE_Basic_ccorres', simp)
|
||||
apply clarsimp
|
||||
apply (clarsimp simp: pte_range_relation_def ptr_range_to_list_def)
|
||||
apply (ctac (no_vcg) add: pteCheckIfMapped_ccorres)
|
||||
apply csymbr
|
||||
apply (simp add: mapM_discarded whileAnno_def
|
||||
Collect_False
|
||||
del: Collect_const)
|
||||
apply (rule ccorres_split_nothrow_novcg)
|
||||
apply (rule_tac F="\<top>\<top>" in ccorres_mapM_x_while' [where i=0])
|
||||
apply clarsimp
|
||||
apply (rule ccorres_guard_imp2)
|
||||
apply (rule storePTE_Basic_ccorres', simp)
|
||||
apply clarsimp
|
||||
apply (clarsimp simp: pte_range_relation_def ptr_range_to_list_def)
|
||||
apply vcg
|
||||
apply simp
|
||||
apply wp
|
||||
apply clarsimp
|
||||
apply (clarsimp simp: pte_range_relation_def ptr_range_to_list_def)
|
||||
apply (rule order_less_le_trans)
|
||||
apply (rule unat_lt2p)
|
||||
apply (simp add: word_bits_def)
|
||||
apply ceqv
|
||||
apply (rule ccorres_add_return2)
|
||||
apply csymbr
|
||||
apply (rule ccorres_Guard_Seq,rule ccorres_Guard_Seq)
|
||||
apply (rule ccorres_move_c_guard_pte)+
|
||||
apply (ctac (no_vcg) add: cleanCacheRange_PoU_ccorres)
|
||||
apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg_throws)
|
||||
apply (rule allI, rule conseqPre, vcg)
|
||||
apply (clarsimp simp:return_def)
|
||||
apply wp [1]
|
||||
apply (rule hoare_strengthen_post)
|
||||
apply (clarsimp simp: pte_range_relation_def ptr_range_to_list_def)
|
||||
apply vcg
|
||||
apply simp
|
||||
apply wp
|
||||
apply clarsimp
|
||||
apply (clarsimp simp: pte_range_relation_def ptr_range_to_list_def)
|
||||
apply (rule order_less_le_trans)
|
||||
apply (rule unat_lt2p)
|
||||
apply (simp add: word_bits_def)
|
||||
apply ceqv
|
||||
apply csymbr
|
||||
apply (rule ccorres_Guard_Seq,rule ccorres_Guard_Seq)
|
||||
apply (rule ccorres_move_c_guard_pte)+
|
||||
apply (ctac (no_vcg) add: cleanCacheRange_PoU_ccorres)
|
||||
apply (simp only: when_def)
|
||||
apply (rule ccorres_Cond_rhs_Seq)
|
||||
apply (simp add: to_bool_def)
|
||||
apply (rule ccorres_add_return2)
|
||||
apply (ctac (no_vcg) add: invalidateTLBByASID_ccorres)
|
||||
apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg_throws)
|
||||
apply (rule allI, rule conseqPre, vcg)
|
||||
apply (clarsimp simp:return_def)
|
||||
apply (rule wp_post_taut)
|
||||
apply (simp add: to_bool_def)
|
||||
apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg_throws)
|
||||
apply (rule allI, rule conseqPre, vcg)
|
||||
apply (clarsimp simp:return_def)
|
||||
apply (wp hoare_vcg_const_imp_lift) [1]
|
||||
apply (clarsimp simp: to_bool_def)
|
||||
apply (rule hoare_strengthen_post)
|
||||
apply (rule hoare_vcg_conj_lift[where Q'="\<lambda>rv s. valid_pde_mappings' s"])
|
||||
apply (rule mapM_x_accumulate_checks)
|
||||
apply (simp add: storePTE_def)
|
||||
apply (rule obj_at_setObject3)
|
||||
apply simp
|
||||
apply (simp add: objBits_simps archObjSize_def)
|
||||
apply (simp add: typ_at_to_obj_at_arches[symmetric])
|
||||
apply wp
|
||||
apply clarsimp
|
||||
apply (simp add: typ_at_to_obj_at_arches)
|
||||
apply (frule bspec, erule hd_in_set)
|
||||
apply (drule bspec, erule last_in_set)
|
||||
apply (simp add: hd_conv_nth last_conv_nth)
|
||||
apply (rule iffD1[OF conj_assoc])+
|
||||
apply (rule conjI)
|
||||
apply (wp mapM_x_wp_inv)
|
||||
apply clarsimp
|
||||
apply (simp add: typ_at_to_obj_at_arches)
|
||||
apply (frule bspec, erule hd_in_set)
|
||||
apply (drule bspec, erule last_in_set)
|
||||
apply (simp add: hd_conv_nth last_conv_nth)
|
||||
apply (rule iffD1[OF conj_assoc])+
|
||||
apply (rule conjI)
|
||||
(* the inequality first *)
|
||||
apply (clarsimp simp:valid_pte_slots'2_def
|
||||
objBits_simps archObjSize_def hd_conv_nth)
|
||||
apply (clarsimp simp:pte_range_relation_def ptr_range_to_list_def ptr_add_def)
|
||||
apply (simp add:add_assoc[symmetric] add_commute)
|
||||
apply (simp add: word_add_format)
|
||||
apply (frule is_aligned_addrFromPPtr_n,simp)
|
||||
apply (cut_tac n = "sz+2" in power_not_zero[where 'a="32"])
|
||||
apply simp
|
||||
apply (rule conjI)
|
||||
apply (erule is_aligned_no_wrap')
|
||||
apply (simp add:mult_commute power_not_zero)
|
||||
apply (rule conjI)
|
||||
apply (erule is_aligned_no_wrap')
|
||||
apply (simp add:mult_commute power_not_zero)
|
||||
apply (simp add:addrFromPPtr_mask_5)
|
||||
apply (clarsimp simp: pte_range_relation_def ptr_range_to_list_def CTypesDefs.ptr_add_def)
|
||||
apply (clarsimp simp: guard_is_UNIV_def Collect_const_mem hd_conv_nth last_conv_nth ucast_minus)
|
||||
apply (clarsimp simp: pte_range_relation_def ptr_range_to_list_def)
|
||||
apply (simp add: CTypesDefs.ptr_add_def ucast_minus ucast_nat_def
|
||||
archObjSize_def objBits_simps word_0_sle_from_less)
|
||||
apply (clarsimp simp: pte_range_relation_def ptr_range_to_list_def unat_1_0)
|
||||
apply (clarsimp simp:valid_pte_slots'2_def
|
||||
objBits_simps archObjSize_def hd_conv_nth)
|
||||
apply (clarsimp simp:pte_range_relation_def ptr_range_to_list_def ptr_add_def)
|
||||
apply (simp add:add_assoc[symmetric] add_commute)
|
||||
apply (simp add: word_add_format)
|
||||
apply (frule is_aligned_addrFromPPtr_n,simp)
|
||||
apply (cut_tac n = "sz+2" in power_not_zero[where 'a="32"])
|
||||
apply simp
|
||||
apply (rule conjI)
|
||||
apply (erule is_aligned_no_wrap')
|
||||
apply (simp add:mult_commute power_not_zero)
|
||||
apply (rule conjI)
|
||||
apply (erule is_aligned_no_wrap')
|
||||
apply (simp add:mult_commute power_not_zero)
|
||||
apply (simp add:addrFromPPtr_mask_5)
|
||||
apply (clarsimp simp: pte_range_relation_def ptr_range_to_list_def CTypesDefs.ptr_add_def)
|
||||
apply (clarsimp simp: guard_is_UNIV_def Collect_const_mem hd_conv_nth last_conv_nth ucast_minus)
|
||||
apply (clarsimp simp: pte_range_relation_def ptr_range_to_list_def)
|
||||
apply (simp add: CTypesDefs.ptr_add_def ucast_minus ucast_nat_def
|
||||
archObjSize_def objBits_simps word_0_sle_from_less)
|
||||
apply (wp | simp add: pteCheckIfMapped_def)+
|
||||
apply (clarsimp simp: pte_range_relation_def hd_map_simp ptr_range_to_list_def unat_1_0)
|
||||
done
|
||||
|
||||
lemma vmsz_aligned_addrFromPPtr':
|
||||
|
@ -2448,7 +2561,6 @@ lemma decodeARMFrameInvocation_ccorres:
|
|||
apply csymbr
|
||||
apply csymbr
|
||||
apply csymbr
|
||||
apply csymbr
|
||||
apply (ctac add: ccorres_injection_handler_csum1
|
||||
[OF ccorres_injection_handler_csum1,
|
||||
OF findPDForASID_ccorres])
|
||||
|
@ -2884,8 +2996,13 @@ lemma decodeARMFrameInvocation_ccorres:
|
|||
apply (intro conjI impI)
|
||||
apply (clarsimp simp: word_less_nat_alt framesize_from_H_eq_eqs
|
||||
vm_page_size_defs gen_framesize_to_H_def
|
||||
attribsFromWord_def vm_attribs_relation_def
|
||||
attribsFromWord_def vm_attribs_relation_def isCap_simps
|
||||
of_bool_nth[simplified of_bool_from_bool])
|
||||
apply (frule interpret_excaps_eq[rule_format, where n=0], simp)
|
||||
apply (erule_tac c="ArchObjectCap (PageDirectoryCap ?a ?b)" in ccap_relationE)
|
||||
apply (clarsimp simp: cap_lift_page_directory_cap to_bool_def
|
||||
cap_page_directory_cap_lift_def
|
||||
cap_to_H_def[split_simps cap_CL.split] valid_cap'_def)
|
||||
apply (clarsimp split:if_splits)
|
||||
apply (clarsimp simp:
|
||||
unat_less_helper isPageFlush_def InvocationLabels_H.isPageFlush_def
|
||||
|
|
|
@ -524,6 +524,16 @@ lemma chooseThread_ccorres:
|
|||
apply (clarsimp simp: all_invs_but_ct_idle_or_in_cur_domain'_def newKernelState_def)
|
||||
done
|
||||
|
||||
lemma ksDomSched_length_relation[simp]:
|
||||
"\<lbrakk>cstate_relation s s'\<rbrakk> \<Longrightarrow> length (kernel_state.ksDomSchedule s) = unat (ksDomScheduleLength)"
|
||||
apply (auto simp: cstate_relation_def cdom_schedule_relation_def Let_def ksDomScheduleLength_def)
|
||||
done
|
||||
|
||||
lemma ksDomSched_length_dom_relation[simp]:
|
||||
"\<lbrakk>cdom_schedule_relation (kernel_state.ksDomSchedule s) kernel_all_global_addresses.ksDomSchedule \<rbrakk> \<Longrightarrow> length (kernel_state.ksDomSchedule s) = unat (ksDomScheduleLength)"
|
||||
apply (auto simp: cstate_relation_def cdom_schedule_relation_def Let_def ksDomScheduleLength_def)
|
||||
done
|
||||
|
||||
lemma nextDomain_ccorres:
|
||||
"ccorres dc xfdc invs' UNIV [] nextDomain (Call nextDomain_'proc)"
|
||||
apply (cinit)
|
||||
|
@ -536,24 +546,29 @@ lemma nextDomain_ccorres:
|
|||
apply (rule conjI)
|
||||
apply clarsimp
|
||||
apply (subgoal_tac "ksDomScheduleIdx \<sigma> = unat (ksDomScheduleLength - 1)")
|
||||
apply (fastforce simp: cdom_schedule_relation_def dom_schedule_entry_relation_def dschDomain_def dschLength_def ksDomScheduleLength_def sdiv_word_def sdiv_int_def)
|
||||
apply (simp add: ksDomScheduleLength_def)
|
||||
apply (fastforce simp add: cdom_schedule_relation_def dom_schedule_entry_relation_def dschDomain_def dschLength_def ksDomScheduleLength_def sdiv_word_def sdiv_int_def simp del: ksDomSched_length_dom_relation)
|
||||
apply (simp add: ksDomScheduleLength_def)
|
||||
apply (frule invs'_ksDomScheduleIdx)
|
||||
apply (simp add: invs'_ksDomSchedule newKernelState_def)
|
||||
apply (simp only: Abs_fnat_hom_1 Abs_fnat_hom_add)
|
||||
apply (drule unat_le_helper)
|
||||
apply (simp add: sdiv_int_def sdiv_word_def)
|
||||
apply clarsimp
|
||||
apply (clarsimp simp: cdom_schedule_relation_def)
|
||||
apply (simp only: Abs_fnat_hom_1 Abs_fnat_hom_add word_not_le)
|
||||
apply clarsimp
|
||||
apply (subst (asm) of_nat_Suc[symmetric])
|
||||
apply (drule iffD1[OF of_nat_mono_maybe'[where X=3, simplified, symmetric], rotated 2])
|
||||
apply simp
|
||||
apply (drule invs'_ksDomScheduleIdx)
|
||||
apply (frule invs'_ksDomScheduleIdx)
|
||||
apply (simp add: invs'_ksDomSchedule newKernelState_def)
|
||||
apply (clarsimp simp: cdom_schedule_relation_def)
|
||||
apply (clarsimp simp: ksDomScheduleLength_def)
|
||||
apply (subst of_nat_Suc[symmetric])+
|
||||
apply (subst unat_of_nat32)
|
||||
apply (simp add: word_bits_def)
|
||||
apply (subst unat_of_nat32)
|
||||
apply (simp add: word_bits_def)
|
||||
apply (fastforce simp: cdom_schedule_relation_def dom_schedule_entry_relation_def dschDomain_def dschLength_def)
|
||||
apply (fastforce simp add: cdom_schedule_relation_def dom_schedule_entry_relation_def dschDomain_def dschLength_def simp del: ksDomSched_length_dom_relation)
|
||||
apply simp
|
||||
done
|
||||
|
||||
|
|
|
@ -1631,7 +1631,8 @@ lemma setVMRoot_ccorres:
|
|||
(* FIXME: move *)
|
||||
lemma invs'_invs_no_cicd:
|
||||
"invs' s \<Longrightarrow> all_invs_but_ct_idle_or_in_cur_domain' s"
|
||||
by (simp add: invs'_def all_invs_but_ct_idle_or_in_cur_domain'_def valid_state'_def newKernelState_def)
|
||||
by (clarsimp simp add: invs'_def all_invs_but_ct_idle_or_in_cur_domain'_def valid_state'_def newKernelState_def)
|
||||
|
||||
|
||||
lemma setVMRootForFlush_ccorres:
|
||||
"ccorres (\<lambda>rv rv'. rv' = from_bool rv) ret__unsigned_long_'
|
||||
|
|
|
@ -57,13 +57,13 @@ where "transform_page_dir_inv invok \<equiv> case invok of
|
|||
|
||||
definition transform_page_inv :: "ArchInvocation_A.page_invocation \<Rightarrow> cdl_page_invocation option"
|
||||
where "transform_page_inv invok \<equiv> case invok of
|
||||
ArchInvocation_A.page_invocation.PageMap cap ct_slot entries \<Rightarrow>
|
||||
ArchInvocation_A.page_invocation.PageMap asid cap ct_slot entries \<Rightarrow>
|
||||
Some (cdl_page_invocation.PageMap (transform_cap cap)
|
||||
(sum_case (transform_pte \<circ> fst) (transform_pde \<circ> fst) entries)
|
||||
(transform_cslot_ptr ct_slot)
|
||||
(sum_case (\<lambda>pair. [ (transform_pt_slot_ref \<circ> hd \<circ> snd) pair])
|
||||
(\<lambda>pair. [(transform_pd_slot_ref \<circ> hd \<circ> snd) pair]) entries))
|
||||
| ArchInvocation_A.page_invocation.PageRemap entries \<Rightarrow>
|
||||
| ArchInvocation_A.page_invocation.PageRemap asid entries \<Rightarrow>
|
||||
Some (cdl_page_invocation.PageRemap
|
||||
(sum_case (transform_pte \<circ> fst) (transform_pde \<circ> fst) entries)
|
||||
(sum_case (\<lambda>pair. [ (transform_pt_slot_ref \<circ> hd \<circ> snd) pair])
|
||||
|
@ -1639,13 +1639,56 @@ lemma invoke_page_directory_corres:
|
|||
apply (clarsimp simp: perform_page_directory_invocation_def)
|
||||
done
|
||||
|
||||
(* NOONE EVER SEES THIS OK *)
|
||||
lemma pte_check_if_mapped_corres:
|
||||
"dcorres dc \<top> \<top> (return a) (pte_check_if_mapped pte)"
|
||||
apply (clarsimp simp add: pte_check_if_mapped_def get_master_pte_def get_pte_def get_pt_def bind_assoc in_monad get_object_def corres_underlying_def)
|
||||
apply (case_tac y, simp_all add: in_monad)
|
||||
apply (case_tac arch_kernel_obj, simp_all add: in_monad)
|
||||
apply (clarsimp split: ARM_Structs_A.pte.splits)
|
||||
apply (simp_all add: get_pte_def get_pt_def get_object_def in_monad bind_assoc split: kernel_object.splits arch_kernel_obj.splits)
|
||||
apply clarsimp
|
||||
apply (case_tac y, simp_all add: in_monad)
|
||||
apply (case_tac arch_kernel_obj, simp_all add: in_monad)
|
||||
apply (auto simp: in_monad)[2]
|
||||
apply clarsimp
|
||||
apply (case_tac y, simp_all add: in_monad, case_tac arch_kernel_obj, simp_all add: in_monad, auto simp: in_monad)
|
||||
done
|
||||
|
||||
lemma pde_check_if_mapped_corres:
|
||||
"dcorres dc \<top> \<top> (return a) (pde_check_if_mapped pde)"
|
||||
apply (clarsimp simp add: pde_check_if_mapped_def get_master_pde_def get_pde_def get_pd_def bind_assoc in_monad get_object_def corres_underlying_def)
|
||||
apply (case_tac y, simp_all add: in_monad)
|
||||
apply (case_tac arch_kernel_obj, simp_all add: in_monad)
|
||||
apply (clarsimp split: ARM_Structs_A.pde.splits)
|
||||
apply (simp_all add: get_pde_def get_pd_def get_object_def in_monad bind_assoc)
|
||||
apply clarsimp
|
||||
apply (case_tac y, simp_all add: in_monad)
|
||||
apply (case_tac arch_kernel_obj, simp_all add: in_monad)
|
||||
apply (auto simp: in_monad)[1]
|
||||
apply clarsimp
|
||||
apply (case_tac y, simp_all add: in_monad)
|
||||
apply (case_tac arch_kernel_obj, simp_all add: in_monad)
|
||||
apply (auto simp: in_monad)[1]
|
||||
apply clarsimp
|
||||
apply (case_tac y, simp_all add: in_monad, case_tac arch_kernel_obj, simp_all add: in_monad, auto simp: in_monad)
|
||||
done
|
||||
|
||||
lemma if_invalidate_equiv_return:
|
||||
"dcorres dc \<top> \<top> (return a) (if b then invalidate_tlb_by_asid asid else return ())"
|
||||
apply (cases b, simp_all)
|
||||
apply (rule wp_to_dcorres)
|
||||
apply (wp invalidate_tlb_by_asid_dwp)
|
||||
apply clarsimp
|
||||
done
|
||||
|
||||
lemma invoke_page_corres:
|
||||
"transform_page_inv ip' = Some ip \<Longrightarrow>
|
||||
dcorres dc \<top> (valid_page_inv ip' and invs and page_inv_duplicates_valid ip' and valid_pdpt_objs and valid_etcbs)
|
||||
(invoke_page ip) (perform_page_invocation ip')"
|
||||
apply (clarsimp simp:invoke_page_def)
|
||||
apply (case_tac ip')
|
||||
apply (simp_all add:perform_page_invocation_def)
|
||||
apply (simp_all add:perform_page_invocation_def)
|
||||
apply (simp_all add:perform_page_invocation_def transform_page_inv_def)
|
||||
apply (rule dcorres_expand_pfx)
|
||||
apply (clarsimp simp:valid_page_inv_def)
|
||||
|
@ -1656,28 +1699,38 @@ lemma invoke_page_corres:
|
|||
split:if_splits)
|
||||
apply (rule corres_guard_imp)
|
||||
apply (rule corres_split [OF _ set_cap_corres])
|
||||
apply (rule corres_dummy_return_l)
|
||||
apply (rule corres_split[OF _ store_pte_set_cap_corres])
|
||||
apply (rule corres_dummy_return_l)
|
||||
apply (rule_tac corres_split[OF _ dcorres_store_invalid_pte_tail_large_page])
|
||||
apply (rule wp_to_dcorres[where Q=\<top>])
|
||||
apply (wp do_machine_op_wp mapM_wp' set_cap_idle
|
||||
store_pte_page_inv_entries_safe set_cap_page_inv_entries_safe
|
||||
| clarsimp simp:cleanCacheRange_PoU_def)+
|
||||
apply (rule corres_dummy_return_pl[where b ="()"])
|
||||
apply (rule corres_split[OF _ pte_check_if_mapped_corres])
|
||||
apply (simp split del: split_if)
|
||||
apply (rule corres_dummy_return_l)
|
||||
apply (rule corres_split[OF _ store_pte_set_cap_corres])
|
||||
apply (rule corres_dummy_return_l)
|
||||
apply (rule_tac corres_split[OF _ dcorres_store_invalid_pte_tail_large_page])
|
||||
apply (rule corres_dummy_return_l)
|
||||
apply (rule corres_split[OF if_invalidate_equiv_return])
|
||||
apply (rule wp_to_dcorres[where Q=\<top>])
|
||||
apply (wp do_machine_op_wp mapM_wp' set_cap_idle
|
||||
store_pte_page_inv_entries_safe set_cap_page_inv_entries_safe
|
||||
| clarsimp simp:cleanCacheRange_PoU_def)+
|
||||
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)
|
||||
apply (clarsimp simp:mapM_x_singleton)
|
||||
apply (rule corres_guard_imp)
|
||||
apply (rule corres_split [OF _ set_cap_corres])
|
||||
apply (rule corres_dummy_return_l)
|
||||
apply (rule corres_split[OF _ store_pde_set_cap_corres])
|
||||
apply (rule corres_dummy_return_l)
|
||||
apply (rule_tac corres_split[OF _ dcorres_store_invalid_pde_tail_super_section])
|
||||
apply (rule wp_to_dcorres[where Q=\<top>])
|
||||
apply (wp do_machine_op_wp mapM_wp' set_cap_idle
|
||||
set_cap_page_inv_entries_safe store_pde_page_inv_entries_safe
|
||||
| clarsimp simp:cleanCacheRange_PoU_def valid_slots_def)+
|
||||
apply (rule corres_dummy_return_pl[where b="()"])
|
||||
apply (rule corres_split[OF _ pde_check_if_mapped_corres])
|
||||
apply (simp split del: split_if)
|
||||
apply (rule corres_dummy_return_l)
|
||||
apply (rule corres_split[OF _ store_pde_set_cap_corres])
|
||||
apply (rule corres_dummy_return_l)
|
||||
apply (rule_tac corres_split[OF _ dcorres_store_invalid_pde_tail_super_section])
|
||||
apply (rule corres_dummy_return_l)
|
||||
apply (rule corres_split[OF if_invalidate_equiv_return])
|
||||
apply (rule wp_to_dcorres[where Q=\<top>])
|
||||
apply (wp do_machine_op_wp mapM_wp' set_cap_idle
|
||||
set_cap_page_inv_entries_safe store_pde_page_inv_entries_safe
|
||||
| clarsimp simp:cleanCacheRange_PoU_def valid_slots_def)+
|
||||
apply (simp add:page_inv_duplicates_valid_def valid_slots_def
|
||||
page_inv_entries_safe_def split:if_splits)
|
||||
apply (clarsimp simp:invs_def valid_state_def cte_wp_at_caps_of_state)
|
||||
|
@ -1686,46 +1739,58 @@ 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"
|
||||
apply (case_tac sum)
|
||||
apply (clarsimp simp: mapM_singleton mapM_x_mapM)
|
||||
apply (simp add:page_inv_duplicates_valid_def
|
||||
split:if_splits)
|
||||
apply (rule corres_guard_imp)
|
||||
apply (rule corres_dummy_return_l)
|
||||
apply (rule corres_split[OF _ store_pte_set_cap_corres])
|
||||
apply (rule corres_dummy_return_l)
|
||||
apply (rule_tac corres_split[OF _ dcorres_store_invalid_pte_tail_large_page])
|
||||
apply (rule wp_to_dcorres[where Q=\<top>])
|
||||
apply (wp do_machine_op_wp mapM_wp' set_cap_idle
|
||||
store_pte_page_inv_entries_safe set_cap_page_inv_entries_safe
|
||||
| clarsimp simp:cleanCacheRange_PoU_def)+
|
||||
apply (rule corres_dummy_return_pl[where b="()"])
|
||||
apply (rule corres_split[OF _ pte_check_if_mapped_corres])
|
||||
apply (simp split del: split_if)
|
||||
apply (rule corres_dummy_return_l)
|
||||
apply (rule corres_split[OF _ store_pte_set_cap_corres])
|
||||
apply (rule corres_dummy_return_l)
|
||||
apply (rule_tac corres_split[OF _ dcorres_store_invalid_pte_tail_large_page])
|
||||
apply (rule corres_dummy_return_l)
|
||||
apply (rule corres_split[OF if_invalidate_equiv_return])
|
||||
apply (rule wp_to_dcorres[where Q=\<top>])
|
||||
apply (wp do_machine_op_wp mapM_wp' set_cap_idle
|
||||
store_pte_page_inv_entries_safe set_cap_page_inv_entries_safe
|
||||
| clarsimp simp:cleanCacheRange_PoU_def)+
|
||||
apply (rule dcorres_expand_pfx)
|
||||
apply (clarsimp simp:mapM_singleton mapM_x_mapM valid_page_inv_def)
|
||||
apply (rule corres_guard_imp)
|
||||
apply (rule corres_dummy_return_l)
|
||||
apply (rule corres_split[OF _ store_pde_set_cap_corres])
|
||||
apply (rule corres_dummy_return_l)
|
||||
apply (rule_tac corres_split[OF _ dcorres_store_invalid_pde_tail_super_section])
|
||||
apply (rule wp_to_dcorres[where Q=\<top>])
|
||||
apply (wp do_machine_op_wp mapM_wp' set_cap_idle
|
||||
set_cap_page_inv_entries_safe store_pde_page_inv_entries_safe
|
||||
| clarsimp simp:cleanCacheRange_PoU_def valid_slots_def
|
||||
)+
|
||||
apply (rule corres_dummy_return_pl[where b="()"])
|
||||
apply (rule corres_split[OF _ pde_check_if_mapped_corres])
|
||||
apply (simp split del: split_if)
|
||||
apply (rule corres_dummy_return_l)
|
||||
apply (rule corres_split[OF _ store_pde_set_cap_corres])
|
||||
apply (rule corres_dummy_return_l)
|
||||
apply (rule_tac corres_split[OF _ dcorres_store_invalid_pde_tail_super_section])
|
||||
apply (rule corres_dummy_return_l)
|
||||
apply (rule corres_split[OF if_invalidate_equiv_return])
|
||||
apply (rule wp_to_dcorres[where Q=\<top>])
|
||||
apply (wp do_machine_op_wp mapM_wp' set_cap_idle
|
||||
set_cap_page_inv_entries_safe store_pde_page_inv_entries_safe
|
||||
| clarsimp simp:cleanCacheRange_PoU_def valid_slots_def
|
||||
)+
|
||||
apply (clarsimp simp:invs_def valid_state_def
|
||||
cte_wp_at_caps_of_state page_inv_duplicates_valid_def)
|
||||
-- "PageUnmap"
|
||||
apply (rule dcorres_expand_pfx)
|
||||
apply (clarsimp simp: valid_page_inv_def liftM_def
|
||||
split:arch_cap.splits option.splits)
|
||||
apply (simp add:alternative_bind_distrib)
|
||||
apply (rule corres_alternate2)
|
||||
apply (rule corres_guard_imp)
|
||||
apply (rule corres_split[OF _ get_cap_corres])
|
||||
apply (rule_tac P="\<lambda>y s. cte_wp_at (op = x) (a,b) s \<and> s = s'" in set_cap_corres_stronger)
|
||||
apply clarsimp
|
||||
apply (drule cte_wp_at_eqD2, simp)
|
||||
apply (clarsimp simp:is_arch_diminished_def transform_mapping_def update_map_data_def
|
||||
dest!:diminished_page_is_page)
|
||||
apply (wp get_cap_cte_wp_at_rv | clarsimp)+
|
||||
apply (simp add:alternative_bind_distrib)
|
||||
apply (rule corres_alternate2)
|
||||
apply (rule corres_guard_imp)
|
||||
apply (rule corres_split[OF _ get_cap_corres])
|
||||
apply (rule_tac P="\<lambda>y s. cte_wp_at (op = x) (a,b) s \<and> s = s'" in set_cap_corres_stronger)
|
||||
apply clarsimp
|
||||
apply (drule cte_wp_at_eqD2, simp)
|
||||
apply (clarsimp simp:is_arch_diminished_def transform_mapping_def update_map_data_def
|
||||
dest!:diminished_page_is_page)
|
||||
apply (wp get_cap_cte_wp_at_rv | clarsimp)+
|
||||
apply (clarsimp simp:cte_wp_at_def is_arch_diminished_def is_arch_cap_def is_pt_cap_def
|
||||
dest!:diminished_page_is_page)
|
||||
apply (clarsimp simp:invs_def valid_state_def not_idle_thread_def)
|
||||
|
@ -1750,7 +1815,7 @@ lemma invoke_page_corres:
|
|||
dest!:diminished_page_is_page)
|
||||
apply (rule conjI, simp)
|
||||
apply (rule conjI, simp add:invs_def valid_state_def valid_idle_def)
|
||||
apply (clarsimp simp:invs_def valid_state_def not_idle_thread_def)
|
||||
apply (clarsimp simp:invs_def valid_state_def not_idle_thread_def)
|
||||
apply (clarsimp simp:invs_def valid_state_def not_idle_thread_def valid_idle_def)
|
||||
apply (frule valid_idle_has_null_cap, (simp add: valid_idle_def)+)
|
||||
apply (rule sym)
|
||||
|
@ -1758,15 +1823,15 @@ lemma invoke_page_corres:
|
|||
apply (clarsimp simp:invoke_page_def)
|
||||
apply (clarsimp simp: when_def split: if_splits)
|
||||
apply (rule corres_guard_imp)
|
||||
apply (rule dcorres_symb_exec_r)+
|
||||
apply (simp only: split_if_asm)
|
||||
apply (safe)
|
||||
apply (erule notE)
|
||||
apply (rule dcorres_symb_exec_r)+
|
||||
apply (rule dcorres_set_vm_root)
|
||||
apply (wp)
|
||||
apply (erule notE)+
|
||||
apply (clarsimp)
|
||||
apply (rule dcorres_symb_exec_r)+
|
||||
apply (simp only: split_if_asm)
|
||||
apply (safe)
|
||||
apply (erule notE)
|
||||
apply (rule dcorres_symb_exec_r)+
|
||||
apply (rule dcorres_set_vm_root)
|
||||
apply (wp)
|
||||
apply (erule notE)+
|
||||
apply (clarsimp)
|
||||
apply (wp do_machine_op_wp | clarsimp)+
|
||||
done
|
||||
|
||||
|
@ -1983,12 +2048,6 @@ lemma dcorres_assert_assume:
|
|||
apply (rule corres_free_fail)
|
||||
done
|
||||
|
||||
lemma option_map_cong: "\<lbrakk> x \<noteq> None \<Longrightarrow> a = b \<rbrakk> \<Longrightarrow> Option.map a x = Option.map b x"
|
||||
apply (case_tac x)
|
||||
apply clarsimp
|
||||
apply clarsimp
|
||||
done
|
||||
|
||||
lemma invoke_asid_pool_corres:
|
||||
"arch_invocation_relation (InvokeAsidPool ap_inv)
|
||||
(arch_invocation.InvokeASIDPool ap_inv') \<Longrightarrow>
|
||||
|
|
|
@ -319,7 +319,7 @@ lemma insert_cap_child_corres:
|
|||
apply (rule stronger_corres_guard_imp)
|
||||
apply (rule corres_split[OF _ get_cap_corres])+
|
||||
apply (rule_tac P="old_cap \<noteq> cdl_cap.NullCap" and P'="rv' \<noteq> cap.NullCap"
|
||||
in dcorres_symmetric_bool_cases)
|
||||
in corres_symmetric_bool_cases)
|
||||
apply (clarsimp simp :transform_cap_def split:cap.splits arch_cap.splits)
|
||||
apply (simp add:assert_def)
|
||||
apply (rule corres_trivial)
|
||||
|
@ -761,7 +761,6 @@ lemma cap_revoke_corres_helper:
|
|||
apply (rule_tac Q'="\<lambda>cap. op = sfix and cte_wp_at (\<lambda>x. x = cap) slot'" in corres_symb_exec_r)
|
||||
apply (rule dcorres_expand_pfx)
|
||||
apply (rule_tac Q'="\<lambda>cdt' s. s = sfix \<and> cdt' = cdt sfix" in corres_symb_exec_r)
|
||||
(* apply (rule_tac Q'="\<lambda>ext s. s = sfix \<and> rva = cdt sfix \<and> ext = (exst sfix)" in corres_symb_exec_r) *)
|
||||
apply clarsimp
|
||||
apply (rule dcorres_expand_pfx)
|
||||
apply clarsimp
|
||||
|
@ -1478,23 +1477,11 @@ lemma arch_recycle_cap_idle[wp]:
|
|||
lemma arch_recycle_cap_st_tcb_at[wp]:
|
||||
"\<lbrace>st_tcb_at P thread\<rbrace> arch_recycle_cap is_final arch_cap \<lbrace>\<lambda>r. st_tcb_at P thread\<rbrace>"
|
||||
apply (case_tac arch_cap)
|
||||
apply (simp_all add:arch_recycle_cap_def)
|
||||
apply (wp mapM_x_wp set_asid_pool_st_tcb_at static_imp_wp
|
||||
| wpc | clarsimp| intro conjI impI | force)+
|
||||
apply (simp_all add:arch_recycle_cap_def split del: split_if)
|
||||
apply (wp mapM_x_wp set_asid_pool_st_tcb_at static_imp_wp invalidate_tlb_by_asid_st_tcb_at
|
||||
| wpc | clarsimp| intro conjI impI | fastforce)+
|
||||
done
|
||||
|
||||
(*
|
||||
lemma arch_recycle_cap_r[wp]:
|
||||
"\<lbrace>\<lambda>s. P cap\<rbrace> arch_recycle_cap is_final cap \<lbrace>\<lambda>r s. P r\<rbrace>"
|
||||
apply (clarsimp simp:arch_recycle_cap_def)
|
||||
apply (case_tac cap)
|
||||
apply simp_all
|
||||
apply (wp | clarsimp)+
|
||||
apply (simp add:no_irq_clearMemory)
|
||||
apply wp
|
||||
done
|
||||
*)
|
||||
|
||||
lemma dcorres_storeWord_mapM_x_cvt:
|
||||
"\<forall>x\<in>set ls. within_page buf x sz
|
||||
\<Longrightarrow> dcorres dc (\<lambda>_. True) (ko_at (ArchObj (DataPage sz)) buf and valid_objs and pspace_distinct and pspace_aligned and valid_etcbs)
|
||||
|
@ -1532,50 +1519,6 @@ lemma dcorres_storeWord_mapM_x_cvt:
|
|||
done
|
||||
qed
|
||||
|
||||
(* MOVE *)
|
||||
lemma cte_wp_at_cap_aligned:
|
||||
"\<lbrakk>cte_wp_at P p s; invs s\<rbrakk> \<Longrightarrow> \<exists>c. P c \<and> cap_aligned c"
|
||||
apply (drule (1) cte_wp_at_valid_objs_valid_cap [OF _ invs_valid_objs])
|
||||
apply (fastforce simp: valid_cap_def)
|
||||
done
|
||||
|
||||
lemma cte_wp_at_cap_aligned':
|
||||
"\<lbrakk>cte_wp_at (op = cap) p s; invs s\<rbrakk> \<Longrightarrow> cap_aligned cap"
|
||||
apply (drule (1) cte_wp_at_valid_objs_valid_cap [OF _ invs_valid_objs])
|
||||
apply (fastforce simp: valid_cap_def)
|
||||
done
|
||||
|
||||
lemma typ_at_pg:
|
||||
"typ_at (AArch (AIntData sz)) buf s = ko_at (ArchObj (DataPage sz)) buf s"
|
||||
unfolding obj_at_def
|
||||
by (auto simp: a_type_def split: Structures_A.kernel_object.split_asm arch_kernel_obj.split_asm split_if_asm)
|
||||
|
||||
lemmas dmo_dwp = do_machine_op_wp [OF allI]
|
||||
|
||||
(* FIXME: move *)
|
||||
lemma dmo_mapM_storeWord_0_invs[wp]:
|
||||
"valid invs (do_machine_op (mapM (\<lambda>p. storeWord p 0) S)) (\<lambda>_. invs)"
|
||||
apply (simp add: dom_mapM ef_storeWord)
|
||||
apply (rule mapM_UNIV_wp)
|
||||
apply (simp add: do_machine_op_def split_def)
|
||||
apply wp
|
||||
apply (clarsimp simp: invs_def valid_state_def cur_tcb_def
|
||||
valid_machine_state_def)
|
||||
apply (rule conjI)
|
||||
apply(erule use_valid[OF _ storeWord_valid_irq_states], simp)
|
||||
apply (clarsimp simp: disj_commute[of "in_user_frame p s", standard])
|
||||
apply (drule_tac x=p in spec, simp)
|
||||
apply (erule use_valid)
|
||||
apply (simp add: storeWord_def word_rsplit_0)
|
||||
apply wp
|
||||
apply simp
|
||||
done
|
||||
|
||||
lemma cleanCacheRange_PoU_underlying_memory[wp]:
|
||||
"\<lbrace>\<lambda>ms. underlying_memory ms = m\<rbrace> cleanCacheRange_PoU buf (buf + 2 ^ pageBitsForSize sz - 1) (addrFromPPtr buf) \<lbrace>\<lambda>rv ms. underlying_memory ms = m\<rbrace>"
|
||||
apply (clarsimp simp: cleanCacheRange_PoU_def, wp)
|
||||
done
|
||||
|
||||
lemma dcorres_arch_recycle_cap_page_cap:
|
||||
notes CSpace_D.finalise_cap.simps [simp del]
|
||||
assumes cap: "cap = (arch_cap.PageCap buf fun sz option)"
|
||||
|
@ -2319,6 +2262,11 @@ lemma get_cnode'D:
|
|||
"get_cnode' ptr s = Some (sz,cn) \<Longrightarrow> kheap s ptr = Some (Structures_A.CNode sz cn)"
|
||||
unfolding get_cnode'_def by (clarsimp split: Structures_A.kernel_object.splits option.splits)
|
||||
|
||||
lemma zombie_get_cnode:
|
||||
"\<lbrakk>cte_wp_at (op = (cap.Zombie x (Some xc) xb)) slot s; invs s\<rbrakk> \<Longrightarrow> get_cnode' x s \<noteq> None"
|
||||
by (clarsimp dest!: cte_wp_at_valid_objs_valid_cap [OF _ invs_valid_objs] simp: valid_cap_simps get_cnode'_def obj_at_def
|
||||
elim!: is_cap_tableE)
|
||||
|
||||
(* MOVE *)
|
||||
definition
|
||||
"transform_cnode sz cn \<equiv> Types_D.CNode \<lparr>
|
||||
|
@ -2379,25 +2327,6 @@ lemma recycle_cap_idle_thread: "\<lbrace> \<lambda>s. P (idle_thread s) \<rbrace
|
|||
apply (wp get_thread_state_it dxo_wp_weak | rule hoare_drop_imps | simp)+
|
||||
done
|
||||
|
||||
|
||||
|
||||
(* MOVE *)
|
||||
lemma cte_wp_at_zombie_not_idle:
|
||||
"\<lbrakk>cte_wp_at (op = (cap.Zombie ptr' zbits n)) ptr s; invs s\<rbrakk> \<Longrightarrow> not_idle_thread (fst ptr) s"
|
||||
"\<lbrakk>cte_wp_at (op = (cap.Zombie ptr' zbits n)) ptr s; invs s\<rbrakk> \<Longrightarrow> not_idle_thread ptr' s"
|
||||
by (auto dest!: zombie_cap_two_nonidles simp: cte_wp_at_caps_of_state not_idle_thread_def)
|
||||
|
||||
(* MOVE *)
|
||||
lemma is_cap_tableE:
|
||||
"\<lbrakk> is_cap_table sz ko; \<And>cs. \<lbrakk> ko = kernel_object.CNode sz cs; well_formed_cnode_n sz cs\<rbrakk> \<Longrightarrow> P \<rbrakk> \<Longrightarrow> P"
|
||||
unfolding is_cap_table_def
|
||||
by (auto split: Structures_A.kernel_object.split_asm)
|
||||
|
||||
lemma zombie_get_cnode:
|
||||
"\<lbrakk>cte_wp_at (op = (cap.Zombie x (Some xc) xb)) slot s; invs s\<rbrakk> \<Longrightarrow> get_cnode' x s \<noteq> None"
|
||||
by (clarsimp dest!: cte_wp_at_valid_objs_valid_cap [OF _ invs_valid_objs] simp: valid_cap_simps get_cnode'_def obj_at_def
|
||||
elim!: is_cap_tableE)
|
||||
|
||||
lemma recycle_cap_corres_pre:
|
||||
"dcorres (\<lambda>x x'. x = transform_cap x') \<top>
|
||||
(\<lambda>s. invs s \<and> cte_wp_at (op = cap) ptr s \<and> valid_pdpt_objs s \<and> valid_etcbs s)
|
||||
|
@ -2752,20 +2681,6 @@ lemma ensure_no_children_dummy:
|
|||
apply fastforce
|
||||
done
|
||||
|
||||
(* FIXME: move *)
|
||||
lemma corres_dummy_skip:
|
||||
"\<lbrakk> corres_underlying sr nf dc \<top> \<top> (return x) f;
|
||||
\<And>x. corres_underlying sr nf r P (P' x) g (g' x); \<lbrace>Q'\<rbrace> f \<lbrace>P'\<rbrace> \<rbrakk> \<Longrightarrow>
|
||||
corres_underlying sr nf r P Q' g (f >>= g')"
|
||||
apply (subst return_bind [of x, symmetric])
|
||||
apply (rule corres_guard_imp)
|
||||
apply (rule corres_split, assumption, assumption)
|
||||
apply wp
|
||||
apply assumption
|
||||
apply simp
|
||||
apply simp
|
||||
done
|
||||
|
||||
lemma derive_cap_dummy:
|
||||
"dcorres dc \<top> \<top> (return x) (derive_cap slot cap)"
|
||||
apply (simp add: derive_cap_def)
|
||||
|
@ -2995,20 +2910,6 @@ lemma corres_bindE_throwError:
|
|||
apply (clarsimp simp: bindE_def bind_def throwError_def return_def lift_def split: sum.splits)
|
||||
done
|
||||
|
||||
(* FIXME: move *)
|
||||
lemma get_cap_wellformed:
|
||||
"\<lbrace>valid_objs\<rbrace> get_cap slot \<lbrace>\<lambda>cap s. wellformed_cap cap\<rbrace>"
|
||||
apply (rule hoare_strengthen_post, rule get_cap_valid)
|
||||
apply (simp add: valid_cap_def2)
|
||||
done
|
||||
|
||||
(*
|
||||
lemma update_cap_rights_same[simp]:
|
||||
"update_cap_rights (Structures_A.cap_rights cap) (transform_cap cap) = transform_cap cap"
|
||||
apply (case_tac cap,simp_all add:update_cap_rights_def)
|
||||
*)
|
||||
|
||||
|
||||
lemma decode_cnode_corres:
|
||||
"\<lbrakk> Some (CNodeIntent ui) = transform_intent (invocation_type label') args';
|
||||
cap = transform_cap cap';
|
||||
|
@ -3405,10 +3306,6 @@ lemma decode_cnode_corres:
|
|||
apply (rule corres_bindE_throwError, wp, simp)
|
||||
done
|
||||
|
||||
lemma singleton_set:
|
||||
"x\<in> set[a] \<Longrightarrow> x = a"
|
||||
by auto
|
||||
|
||||
lemma decode_cnode_label_not_match:
|
||||
"\<lbrakk>Some intent = transform_intent (invocation_type label) args; \<forall>ui. intent \<noteq> CNodeIntent ui\<rbrakk>
|
||||
\<Longrightarrow> \<lbrace>op = s\<rbrace> Decode_A.decode_cnode_invocation label args (cap.CNodeCap a b c) (e) \<lbrace>\<lambda>r. \<bottom>\<rbrace>, \<lbrace>\<lambda>e. op = s\<rbrace>"
|
||||
|
|
|
@ -46,7 +46,6 @@ lemma OR_choiceE_OR[simp]: "(OR_choiceE c (f :: ('a + 'b,unit) s_monad) g) = (f
|
|||
done
|
||||
|
||||
(* state relation as set, in (simp split_def) normal form *)
|
||||
|
||||
abbreviation
|
||||
dcorres ::
|
||||
"('a\<Colon>type \<Rightarrow> 'b\<Colon>type \<Rightarrow> bool)
|
||||
|
@ -60,11 +59,6 @@ dcorres ::
|
|||
where
|
||||
"dcorres\<equiv> corres_underlying {ss'. transform (snd ss') = fst ss'} False"
|
||||
|
||||
(* FIXME: need to move all these corres lemmas from everywhere into
|
||||
one theory that is included here *)
|
||||
|
||||
(* Some lemma that trying not involving too much trivial wp rules *)
|
||||
|
||||
(* Some obvious corres lemmas *)
|
||||
lemma corres_group_bind_rhs:
|
||||
"corres_underlying sr nf rvr P P' a (do y\<leftarrow>(do f; g od); h y od)
|
||||
|
@ -114,6 +108,10 @@ lemma corres_dummy_return_pr:
|
|||
"dcorres c P P' f (do return b; g od) \<Longrightarrow> dcorres c P P' f g"
|
||||
by (fastforce simp: corres_underlying_def bind_def return_def)
|
||||
|
||||
lemma corres_dummy_returnOk_r:
|
||||
"dcorres c P P' f (g >>=E returnOk) \<Longrightarrow> dcorres c P P' f g"
|
||||
by simp
|
||||
|
||||
lemma corres_dummy_get_pr:
|
||||
"dcorres c P P' f (do s\<leftarrow>get;g od)\<Longrightarrow> dcorres c P P' f g"
|
||||
by (fastforce simp: corres_underlying_def bind_def get_def)
|
||||
|
@ -138,7 +136,6 @@ lemma absorb_imp:"B \<and> A \<Longrightarrow> (a\<longrightarrow>A) \<and> B "
|
|||
|
||||
(* This lemma is convienent if you want keep the prefix while split *)
|
||||
|
||||
|
||||
lemma corres_split_keep_pfx:
|
||||
assumes x: "corres_underlying sr nf r' P P' a c"
|
||||
assumes y: "\<And>rv rv'. r' rv rv' \<Longrightarrow> corres_underlying sr nf r (P and (Q rv)) (P' and (Q' rv')) (b rv) (d rv')"
|
||||
|
@ -146,11 +143,6 @@ lemma corres_split_keep_pfx:
|
|||
shows "corres_underlying sr nf r P P' (a >>= (\<lambda>rv. b rv)) (c >>= (\<lambda>rv'. d rv'))"
|
||||
using assms by (rule corres_split')
|
||||
|
||||
lemma dcorres_test:
|
||||
"\<lbrakk>dcorres r P (P' and G) a b ; G x\<rbrakk>\<Longrightarrow> dcorres r P (P' and (op = x)) a b"
|
||||
by (fastforce simp:corres_underlying_def)
|
||||
|
||||
|
||||
(* Following 2 lemmas allows you to get rid of the get function and move the prefix outside *)
|
||||
|
||||
lemma dcorres_absorb_pfx:
|
||||
|
@ -257,18 +249,6 @@ lemma wpc_helper_dcorres:
|
|||
wpc_setup "\<lambda>m. dcorres r P P' a m" wpc_helper_dcorres
|
||||
wpc_setup "\<lambda>m. dcorres r P P' a (m >>= c)" wpc_helper_dcorres
|
||||
|
||||
lemma UNIV_apply: "x \<in> UNIV"
|
||||
by (simp add:top_fun_def)
|
||||
|
||||
lemma dcorres_wpc_test2:
|
||||
"dcorres dc P P' (return ()) ((case c of ThreadCap p \<Rightarrow> return () | _ \<Rightarrow> return ()))"
|
||||
apply (rule corres_guard_imp)
|
||||
apply wpc
|
||||
apply (rule corres_free_return [where P = "\<lambda>_. True" and P' = "\<lambda>_. True"])+
|
||||
apply simp
|
||||
apply (simp add: )
|
||||
done
|
||||
|
||||
(* Shorthand to say that a TCB is at the given location in the given state. *)
|
||||
definition "cdl_tcb_at x \<equiv> \<lambda>s. \<exists>tcb . cdl_objects s x = Some (Tcb tcb)"
|
||||
|
||||
|
@ -479,6 +459,14 @@ lemma dcorres_to_wp:
|
|||
"dcorres dc \<top> Q (return x) g \<Longrightarrow> \<lbrace>\<lambda>s. Q s \<and> transform s = cs\<rbrace>g\<lbrace>\<lambda>r s. transform s = cs\<rbrace>"
|
||||
by (fastforce simp:corres_underlying_def valid_def return_def)
|
||||
|
||||
lemma wp_to_dcorres:
|
||||
"(\<And>cs. \<lbrace>\<lambda>s. Q s \<and> transform s = cs\<rbrace> g \<lbrace>\<lambda>r s. transform s = cs\<rbrace>) \<Longrightarrow> dcorres dc (\<lambda>_. True) Q (return x) g"
|
||||
apply (clarsimp simp:corres_underlying_def valid_def return_def)
|
||||
apply (drule_tac x = "transform b" in meta_spec)
|
||||
apply (drule_tac x = b in spec)
|
||||
apply fastforce
|
||||
done
|
||||
|
||||
lemma dcorres_symb_exec_rE:
|
||||
"\<lbrakk>\<And>rv. dcorres r P (Q' rv) h (g rv); \<lbrace>P'\<rbrace> f \<lbrace>Q'\<rbrace>; \<And>cs. \<lbrace>\<lambda>ps. transform ps = cs\<rbrace> f \<lbrace>\<lambda>r s. transform s = cs\<rbrace>\<rbrakk>
|
||||
\<Longrightarrow> dcorres r P P' h (liftE f >>=E g)"
|
||||
|
@ -527,26 +515,6 @@ lemma corres_handle2':
|
|||
apply (fastforce simp: return_def throwError_def split: sum.splits)
|
||||
done
|
||||
|
||||
|
||||
(* FIXME: move *)
|
||||
lemma corres_underlyingD:
|
||||
"\<lbrakk> corres_underlying R z rs P P' f f'; (s,s') \<in> R; P s; P' s' \<rbrakk>
|
||||
\<Longrightarrow> (\<forall>(r',t')\<in>fst (f' s'). \<exists>(r,t)\<in>fst (f s). (t, t') \<in> R \<and> rs r r') \<and> (z \<longrightarrow> \<not> snd (f' s'))"
|
||||
by (fastforce simp: corres_underlying_def)
|
||||
|
||||
(* FIXME: move *)
|
||||
lemma corres_underlyingD2:
|
||||
"\<lbrakk> corres_underlying R z rs P P' f f'; (s,s') \<in> R; P s; P' s'; (r',t')\<in>fst (f' s') \<rbrakk>
|
||||
\<Longrightarrow> \<exists>(r,t)\<in>fst (f s). (t, t') \<in> R \<and> rs r r'"
|
||||
by (fastforce dest: corres_underlyingD)
|
||||
|
||||
|
||||
(* FIXME: move *)
|
||||
lemma in_alternative:
|
||||
"(r,s') \<in> fst ((f \<sqinter> g) s) = ((r,s') \<in> fst (f s) \<or> (r,s') \<in> fst (g s))"
|
||||
by (simp add: alternative_def)
|
||||
|
||||
|
||||
lemma corres_alternative_throw_splitE:
|
||||
assumes a: "corres_underlying R z (dc \<oplus> r') P P' (f \<sqinter> Monads_D.throw) f'"
|
||||
assumes b: "\<And>x x'. r' x x' \<Longrightarrow> corres_underlying R z (dc \<oplus> r) (Q x) (Q' x') (g x \<sqinter> Monads_D.throw) (g' x')"
|
||||
|
@ -617,7 +585,6 @@ lemma corres_alternative_throw_splitE:
|
|||
apply (simp add: alternative_def throwError_def return_def)
|
||||
done
|
||||
|
||||
|
||||
lemma corres_throw_skip_r:
|
||||
assumes c: "corres_underlying R z (dc \<oplus> r) P P' (f \<sqinter> Monads_D.throw) g'"
|
||||
assumes eq: "\<And>s. \<lbrace>op = s\<rbrace> f' \<lbrace>\<lambda>_. op = s\<rbrace>"
|
||||
|
@ -666,4 +633,31 @@ lemmas dcorres_rhs_noop_above_True = dcorres_rhs_noop_above[OF _ _ hoare_TrueI h
|
|||
|
||||
declare hoare_TrueI[simp]
|
||||
|
||||
lemma dcorres_dc_rhs_noop_below_gen:
|
||||
"\<lbrakk> \<forall>rv'. dcorres dc (Q ()) (Q' rv') (return ()) (m rv');
|
||||
dcorres dc P P' f g;
|
||||
\<lbrace> P \<rbrace> f \<lbrace> Q \<rbrace>; \<lbrace> P' \<rbrace> g \<lbrace> Q' \<rbrace> \<rbrakk>
|
||||
\<Longrightarrow> dcorres dc P P' f (g >>= m)"
|
||||
apply (rule corres_add_noop_lhs2)
|
||||
apply (rule corres_underlying_split)
|
||||
apply (assumption | clarsimp)+
|
||||
done
|
||||
|
||||
lemma dcorres_dc_rhs_noop_below_2: "\<lbrakk> \<forall>rv'. dcorres dc (Q ()) (Q' rv') (return ()) m;
|
||||
dcorres dc P P' f (g >>= h);
|
||||
\<lbrace> P \<rbrace> f \<lbrace> Q \<rbrace>;
|
||||
\<lbrace> P' \<rbrace> g \<lbrace> R'\<rbrace>;
|
||||
(\<And>y. \<lbrace> R' y \<rbrace> h y \<lbrace> Q' \<rbrace>)
|
||||
\<rbrakk>
|
||||
\<Longrightarrow> dcorres dc P P' f (do x \<leftarrow> g;
|
||||
_ \<leftarrow> h x;
|
||||
m
|
||||
od)"
|
||||
apply (simp add: bind_assoc[symmetric])
|
||||
apply (rule dcorres_dc_rhs_noop_below_gen)
|
||||
apply (wp | simp | assumption)+
|
||||
done
|
||||
|
||||
lemmas dcorres_dc_rhs_noop_below_2_True = dcorres_dc_rhs_noop_below_2[OF _ _ hoare_TrueI hoare_TrueI hoare_TrueI]
|
||||
|
||||
end
|
||||
|
|
|
@ -12,6 +12,7 @@ theory Finalise_DR
|
|||
imports
|
||||
KHeap_DR
|
||||
"../invariant-abstract/PDPTEntries_AI"
|
||||
"../../lib/Apply_Trace"
|
||||
begin
|
||||
|
||||
declare dxo_wp_weak[wp del]
|
||||
|
@ -103,33 +104,6 @@ lemma dcorres_revoke_cap_no_descendants:
|
|||
apply (clarsimp simp:descendants_of_eqv valid_mdb_def)
|
||||
done
|
||||
|
||||
(* FIXME: replace version in Ipc_R with this *)
|
||||
lemma not_waiting_reply_slot_no_descendants:
|
||||
"\<lbrakk> st_tcb_at (Not \<circ> awaiting_reply) t s;
|
||||
valid_reply_caps s; valid_objs s; valid_mdb s \<rbrakk>
|
||||
\<Longrightarrow> descendants_of (t, tcb_cnode_index 2) (cdt s) = {}"
|
||||
apply (rule ccontr, erule nonemptyE)
|
||||
apply (clarsimp simp: valid_mdb_def reply_mdb_def reply_masters_mdb_def)
|
||||
apply (frule_tac ref="tcb_cnode_index 2" in tcb_at_cte_at[OF st_tcb_at_tcb_at])
|
||||
apply (simp add: domI)
|
||||
apply (clarsimp simp: cte_wp_at_caps_of_state)
|
||||
apply (frule(1) tcb_cap_valid_caps_of_stateD)
|
||||
apply (clarsimp simp: tcb_cap_valid_def st_tcb_at_tcb_at)
|
||||
apply (clarsimp simp: st_tcb_def2)
|
||||
apply (erule disjE)
|
||||
apply (clarsimp simp: cte_wp_at_caps_of_state is_cap_simps)
|
||||
apply (elim allE, drule(1) mp, clarsimp)
|
||||
apply (drule(1) bspec)
|
||||
apply (drule has_reply_cap_cte_wpD[OF caps_of_state_cteD])
|
||||
apply (erule notE[rotated], strengthen reply_cap_doesnt_exist_strg)
|
||||
apply (simp add: st_tcb_def2)
|
||||
apply clarsimp
|
||||
apply (frule mdb_Null_descendants[OF caps_of_state_cteD])
|
||||
apply (simp add: valid_mdb_def reply_mdb_def reply_masters_mdb_def)
|
||||
apply simp
|
||||
done
|
||||
|
||||
|
||||
lemma dcorres_revoke_cap_unnecessary:
|
||||
"dcorres dc \<top> (valid_reply_caps and valid_objs and only_idle and valid_mdb and st_tcb_at (Not \<circ> awaiting_reply) ptr
|
||||
and cte_at (ptr,tcb_cnode_index 2) and not_idle_thread ptr and valid_idle)
|
||||
|
@ -512,6 +486,8 @@ lemma do_machine_op_wp:
|
|||
apply (simp add: transform_objects_def2)
|
||||
done
|
||||
|
||||
lemmas dmo_dwp = do_machine_op_wp [OF allI]
|
||||
|
||||
lemma machine_op_lift[wp]:
|
||||
"\<lbrace>\<lambda>ms. underlying_memory ms = m\<rbrace> machine_op_lift x \<lbrace>\<lambda>rv ms. underlying_memory ms = m\<rbrace>"
|
||||
apply (clarsimp simp:machine_rest_lift_def ignore_failure_def machine_op_lift_def)
|
||||
|
@ -522,17 +498,6 @@ lemma machine_op_lift[wp]:
|
|||
apply clarsimp
|
||||
done
|
||||
|
||||
(*
|
||||
|
||||
lemma cleanCache_unerlying_memory[wp]:
|
||||
"\<lbrace>\<lambda>ms. underlying_memory ms = m\<rbrace> cleanCache \<lbrace>\<lambda>rv ms. underlying_memory ms = m\<rbrace>"
|
||||
by (clarsimp simp:cleanCache_def,wp)
|
||||
|
||||
lemma invalidateHWASID_underlying_memory[wp]:
|
||||
"\<lbrace>\<lambda>ms. underlying_memory ms = m\<rbrace> invalidateHWASID a \<lbrace>\<lambda>rv ms. underlying_memory ms = m\<rbrace>"
|
||||
by (clarsimp simp:invalidateHWASID_def,wp)
|
||||
*)
|
||||
|
||||
lemma invalidateTLB_ASID_underlying_memory[wp]:
|
||||
"\<lbrace>\<lambda>ms. underlying_memory ms = m\<rbrace> invalidateTLB_ASID a \<lbrace>\<lambda>rv ms. underlying_memory ms = m\<rbrace>"
|
||||
apply (clarsimp simp: invalidateTLB_ASID_def, wp)
|
||||
|
@ -619,15 +584,6 @@ lemma set_hardware_asid_dwp[wp]:
|
|||
" \<lbrace>\<lambda>ms. underlying_memory ms = m\<rbrace> setHardwareASID hw_asid \<lbrace>\<lambda>rv ms. underlying_memory ms = m\<rbrace>"
|
||||
by (clarsimp simp:setHardwareASID_def,wp)
|
||||
|
||||
(*
|
||||
|
||||
lemma invalidateMVA_dwp[wp]:
|
||||
"\<lbrace>\<lambda> ms. underlying_memory ms = m\<rbrace> invalidateMVA x \<lbrace>\<lambda>rv ms. underlying_memory ms = m\<rbrace>"
|
||||
by (clarsimp simp:invalidateMVA_def,wp)
|
||||
|
||||
*)
|
||||
|
||||
|
||||
lemma store_hardware_asid_dwp[wp]:
|
||||
"\<lbrace>\<lambda>s. transform s = cs\<rbrace> store_hw_asid a xa \<lbrace>\<lambda>xb a. transform a = cs\<rbrace>"
|
||||
apply (clarsimp simp:store_hw_asid_def)
|
||||
|
@ -639,21 +595,6 @@ lemma store_hardware_asid_dwp[wp]:
|
|||
apply (clarsimp simp:transform_def transform_objects_def2 transform_current_thread_def transform_cdt_def transform_asid_table_def)
|
||||
done
|
||||
|
||||
(*
|
||||
|
||||
lemma invalidateHWASID_dwp[wp]:
|
||||
"\<lbrace>\<lambda>ms. underlying_memory ms = m\<rbrace> invalidateHWASID next_asid \<lbrace>\<lambda>rv ms. underlying_memory ms = m\<rbrace>"
|
||||
apply (clarsimp simp:invalidateHWASID_def)
|
||||
apply (clarsimp simp:machine_op_lift_def machine_rest_lift_def)
|
||||
apply wp
|
||||
apply (clarsimp simp:valid_def in_monad)
|
||||
apply (assumption)
|
||||
apply wp
|
||||
apply (simp add:ignore_failure_def)
|
||||
done
|
||||
|
||||
*)
|
||||
|
||||
lemma transform_arch_state:
|
||||
"\<lbrakk> arm_globals_frame (arch_state a) = arm_globals_frame x;
|
||||
arm_asid_table (arch_state a) = arm_asid_table x \<rbrakk>
|
||||
|
@ -1507,10 +1448,6 @@ lemma mask_compare_imply:
|
|||
apply (auto simp:word_size)
|
||||
done
|
||||
|
||||
lemma distinct_imply_not_in_tail:
|
||||
"\<lbrakk>distinct list;suffixeq (y # ys) list\<rbrakk> \<Longrightarrow>y\<notin> set ys "
|
||||
by (clarsimp simp:suffixeq_def)
|
||||
|
||||
lemma aligned_in_step_up_to:
|
||||
"\<lbrakk>x\<in> set (map (\<lambda>x. x + ptr) [0 , (2^t) .e. up]);t< WordSetup.word_bits; is_aligned ptr t\<rbrakk>
|
||||
\<Longrightarrow> is_aligned x t"
|
||||
|
@ -1621,83 +1558,6 @@ lemma remain_pd_either_section_relation:
|
|||
apply fastforce+
|
||||
done
|
||||
|
||||
(* Should this be in VSpace_AI.thy *)
|
||||
lemmas set_asid_pool_cte_wp_at1[wp]
|
||||
= hoare_cte_wp_caps_of_state_lift [OF set_asid_pool_caps_of_state]
|
||||
|
||||
lemma mdb_cte_at_set_asid_pool[wp]:
|
||||
"\<lbrace>\<lambda>s. mdb_cte_at (swp (cte_wp_at (op \<noteq> cap.NullCap)) s) (cdt s)\<rbrace>
|
||||
set_asid_pool y pool
|
||||
\<lbrace>\<lambda>r s. mdb_cte_at (swp (cte_wp_at (op \<noteq> cap.NullCap)) s) (cdt s)\<rbrace>"
|
||||
apply (clarsimp simp:mdb_cte_at_def)
|
||||
apply (simp only: imp_conv_disj)
|
||||
apply (wp hoare_vcg_disj_lift hoare_vcg_all_lift)
|
||||
done
|
||||
|
||||
lemma mdb_cte_at_store_pte[wp]:
|
||||
"\<lbrace>\<lambda>s. mdb_cte_at (swp (cte_wp_at (op \<noteq> cap.NullCap)) s) (cdt s)\<rbrace>
|
||||
store_pte y pte
|
||||
\<lbrace>\<lambda>r s. mdb_cte_at (swp (cte_wp_at (op \<noteq> cap.NullCap)) s) (cdt s)\<rbrace>"
|
||||
apply (clarsimp simp:mdb_cte_at_def)
|
||||
apply (simp only: imp_conv_disj)
|
||||
apply (wp hoare_vcg_disj_lift hoare_vcg_all_lift)
|
||||
apply (simp add:store_pte_def set_pt_def)
|
||||
apply wp
|
||||
apply (rule hoare_drop_imp)
|
||||
apply (wp|simp)+
|
||||
done
|
||||
|
||||
lemma mdb_cte_at_store_pde[wp]:
|
||||
"\<lbrace>\<lambda>s. mdb_cte_at (swp (cte_wp_at (op \<noteq> cap.NullCap)) s) (cdt s)\<rbrace>
|
||||
store_pde y pde
|
||||
\<lbrace>\<lambda>r s. mdb_cte_at (swp (cte_wp_at (op \<noteq> cap.NullCap)) s) (cdt s)\<rbrace>"
|
||||
apply (clarsimp simp:mdb_cte_at_def)
|
||||
apply (simp only: imp_conv_disj)
|
||||
apply (wp hoare_vcg_disj_lift hoare_vcg_all_lift)
|
||||
done
|
||||
|
||||
lemma valid_idle_store_pte[wp]:
|
||||
"\<lbrace>valid_idle\<rbrace> store_pte y pte \<lbrace>\<lambda>rv. valid_idle\<rbrace>"
|
||||
apply (simp add:store_pte_def)
|
||||
apply wp
|
||||
apply (rule hoare_vcg_precond_imp[where Q="valid_idle"])
|
||||
apply (simp add:set_pt_def)
|
||||
apply wp
|
||||
apply (simp add:get_object_def)
|
||||
apply wp
|
||||
apply (clarsimp simp:obj_at_def
|
||||
split:Structures_A.kernel_object.splits arch_kernel_obj.splits)
|
||||
apply (fastforce simp:is_tcb_def)
|
||||
apply (assumption)
|
||||
apply (wp|simp)+
|
||||
done
|
||||
|
||||
(* FIXME: move *)
|
||||
lemma aligned_add_offset_less:
|
||||
"\<lbrakk>is_aligned x n; is_aligned y n; x < y; z < 2 ^ n\<rbrakk> \<Longrightarrow> x + z < y"
|
||||
apply (cases "y = 0")
|
||||
apply simp
|
||||
apply (erule is_aligned_get_word_bits[where p=y], simp_all)
|
||||
apply (cases "z = 0", simp_all)
|
||||
apply (drule(2) aligned_at_least_t2n_diff[rotated -1])
|
||||
apply (drule plus_one_helper2)
|
||||
apply (rule less_is_non_zero_p1)
|
||||
apply (rule aligned_less_plus_1)
|
||||
apply (erule aligned_sub_aligned[OF _ _ order_refl],
|
||||
simp_all add: is_aligned_triv)[1]
|
||||
apply (cases n, simp_all)[1]
|
||||
apply (simp only: trans[OF diff_add_eq diff_diff_eq2[symmetric]])
|
||||
apply (drule word_less_add_right)
|
||||
apply (rule ccontr, simp add: linorder_not_le)
|
||||
apply (drule aligned_small_is_0, erule order_less_trans)
|
||||
apply (clarsimp simp: power_overflow)
|
||||
apply simp
|
||||
apply (erule order_le_less_trans[rotated],
|
||||
rule word_plus_mono_right)
|
||||
apply (erule minus_one_helper3)
|
||||
apply (simp add: is_aligned_no_wrap' field_simps)
|
||||
done
|
||||
|
||||
lemma is_aligned_less_kernel_base_helper:
|
||||
"\<lbrakk>is_aligned (ptr :: word32) 6;
|
||||
ucast (ptr && mask pd_bits >> 2) \<notin> kernel_mapping_slots;
|
||||
|
@ -1821,9 +1681,6 @@ lemma dcorres_store_pte_non_sense:
|
|||
not_idle_thread_def obj_at_def)
|
||||
done
|
||||
|
||||
(* It is disappointing that I need to produce similar proof
|
||||
for store_pde and store_pte seperately, Xin. *)
|
||||
|
||||
lemma store_pde_non_sense_wp:
|
||||
"\<lbrace>\<lambda>s. (\<exists>f. ko_at (ArchObj (arch_kernel_obj.PageDirectory f)) (slot && ~~ mask pd_bits) s
|
||||
\<and> (\<forall>slot\<in>set xs. f (ucast (slot && mask pd_bits >> 2)) = ARM_Structs_A.pde.InvalidPDE)) \<rbrace>
|
||||
|
@ -1908,28 +1765,6 @@ lemma dcorres_store_invalid_pte_tail_large_page:
|
|||
done
|
||||
qed
|
||||
|
||||
|
||||
lemma and_mask_plus:
|
||||
"\<lbrakk>is_aligned ptr m; m \<le> n; n < 32; a < 2 ^m\<rbrakk>
|
||||
\<Longrightarrow> (ptr::word32) + a && mask n = (ptr && mask n) + a"
|
||||
apply (rule mask_eqI[where n = m])
|
||||
apply (simp add:mask_twice min_def)
|
||||
apply (simp add:is_aligned_add_helper)
|
||||
apply (subst is_aligned_add_helper[THEN conjunct1])
|
||||
apply (erule is_aligned_after_mask)
|
||||
apply simp
|
||||
apply simp
|
||||
apply simp
|
||||
apply (subgoal_tac "(ptr + a && mask n) && ~~ mask m
|
||||
= (ptr + a && ~~ mask m ) && mask n")
|
||||
apply (simp add:is_aligned_add_helper)
|
||||
apply (subst is_aligned_add_helper[THEN conjunct2])
|
||||
apply (simp add:is_aligned_after_mask)
|
||||
apply simp
|
||||
apply simp
|
||||
apply (simp add:word_bw_comms word_bw_lcs)
|
||||
done
|
||||
|
||||
lemma dcorres_unmap_sections:
|
||||
"\<lbrakk>is_aligned ptr 6;ucast (ptr && mask pd_bits >> 2) \<notin> kernel_mapping_slots\<rbrakk>
|
||||
\<Longrightarrow> dcorres dc \<top> (invs and valid_pdpt_objs
|
||||
|
@ -2041,7 +1876,7 @@ lemma dcorres_unmap_sections:
|
|||
apply (rule shiftl_less_t2n[where m = 6,simplified])
|
||||
apply (simp add:word_of_nat_less)
|
||||
apply simp+
|
||||
done
|
||||
done
|
||||
|
||||
lemma pt_page_relation_weaken:
|
||||
"\<lbrakk> pt_page_relation a b c S s; S \<subseteq> T \<rbrakk> \<Longrightarrow> pt_page_relation a b c T s"
|
||||
|
@ -2067,33 +1902,32 @@ lemma dcorres_unmap_pages:
|
|||
apply (simp add:mapM_x_singleton)
|
||||
apply (rule corres_dummy_return_l)
|
||||
apply (rule corres_guard_imp)
|
||||
apply (rule corres_split[OF _ dcorres_store_invalid_pte])
|
||||
apply(rule corres_dummy_return_l)
|
||||
apply (rule_tac r'=dc
|
||||
in corres_split[OF corres_free_return[where P=\<top> and P'=\<top>]])
|
||||
apply (rule corres_split[OF _ dcorres_store_invalid_pte])
|
||||
apply(rule corres_dummy_return_l)
|
||||
apply (rule_tac r'=dc
|
||||
in corres_split[OF corres_free_return[where P=\<top> and P'=\<top>]])
|
||||
apply (rule dcorres_store_invalid_pte_tail_large_page[where slot = ptr])
|
||||
apply wp
|
||||
apply (wp store_pte_non_sense_wp)
|
||||
apply simp
|
||||
apply simp+
|
||||
apply (simp add: hd_map_simp upto_enum_step_def upto_enum_def)
|
||||
apply (clarsimp simp:unat_def)
|
||||
apply (rule slot_with_pt_frame_relation)
|
||||
apply (simp add:invs_valid_idle)
|
||||
apply (fastforce simp:pt_page_relation_def)
|
||||
apply simp
|
||||
apply (simp add:obj_at_def hd_map_simp upto_enum_step_def upto_enum_def
|
||||
pt_page_relation_def)
|
||||
apply (rule conjI,simp)
|
||||
apply (rule conjI,fastforce)
|
||||
(* FIXME : clag from store_pde_super_section *)
|
||||
apply (subst conj_ac)
|
||||
apply (rule context_conjI)
|
||||
apply (clarsimp simp:tl_map drop_map)
|
||||
apply (simp add:field_simps)
|
||||
apply (subst mask_lower_twice[symmetric,where n = 6])
|
||||
apply (simp add:pt_bits_def pageBits_def)
|
||||
apply (subst is_aligned_add_helper)
|
||||
apply (simp add: hd_map_simp upto_enum_step_def upto_enum_def)
|
||||
apply (clarsimp simp:unat_def)
|
||||
apply (rule slot_with_pt_frame_relation)
|
||||
apply (simp add:invs_valid_idle)
|
||||
apply (fastforce simp:pt_page_relation_def)
|
||||
apply simp
|
||||
apply (simp add:obj_at_def hd_map_simp upto_enum_step_def upto_enum_def
|
||||
pt_page_relation_def)
|
||||
apply (rule conjI,simp)
|
||||
apply (rule conjI,fastforce)
|
||||
apply (subst conj_ac)
|
||||
apply (rule context_conjI)
|
||||
apply (clarsimp simp:tl_map drop_map)
|
||||
apply (simp add:field_simps)
|
||||
apply (subst mask_lower_twice[symmetric,where n = 6])
|
||||
apply (simp add:pt_bits_def pageBits_def)
|
||||
apply (subst is_aligned_add_helper)
|
||||
apply simp
|
||||
apply (simp add:upto_0_to_n)
|
||||
apply (rule word_less_power_trans_ofnat[where k = 2 and m = 6,simplified])
|
||||
|
@ -2187,21 +2021,6 @@ lemma dcorres_unmap_page_table_store_pde:
|
|||
apply(clarsimp simp:pd_pt_relation_def unat_def obj_at_def)+
|
||||
done
|
||||
|
||||
(*
|
||||
|
||||
lemma cleanCacheMVA_underlying_memory[wp]:
|
||||
" \<lbrace>\<lambda>ms. underlying_memory ms = m\<rbrace> cleanCacheMVA c
|
||||
\<lbrace>\<lambda>rv ms. underlying_memory ms = m\<rbrace>"
|
||||
apply (simp add:cleanCacheMVA_def)
|
||||
apply (simp add:machine_op_lift_def machine_rest_lift_def)
|
||||
apply wp
|
||||
apply (clarsimp simp: valid_def simpler_modify_def)
|
||||
apply (assumption)
|
||||
apply wp
|
||||
apply (simp add:ignore_failure_def)
|
||||
done
|
||||
*)
|
||||
|
||||
lemma (in pspace_update_eq) pd_pt_relation_update[iff]:
|
||||
"pd_pt_relation a b c (f s) = pd_pt_relation a b c s"
|
||||
by (simp add: pd_pt_relation_def pspace)
|
||||
|
@ -3072,16 +2891,6 @@ lemma swap_for_delete_corres:
|
|||
apply (clarsimp simp: cte_wp_at_caps_of_state)
|
||||
done
|
||||
|
||||
|
||||
lemma invs_valid_irq_node[elim!]:
|
||||
"invs s \<Longrightarrow> valid_irq_node s"
|
||||
by (simp add: invs_def valid_state_def)
|
||||
|
||||
lemma corres_assert_rhs_both_sides:
|
||||
"(F \<Longrightarrow> corres_underlying sr False r P P' f (g ()))
|
||||
\<Longrightarrow> corres_underlying sr False r (\<lambda>s. F \<longrightarrow> P s) (\<lambda>s. F \<longrightarrow> P' s) f (assert F >>= g)"
|
||||
by (cases F, simp_all add: corres_fail)
|
||||
|
||||
definition
|
||||
"arch_page_vmpage_size cap \<equiv>
|
||||
case cap of cap.ArchObjectCap (arch_cap.PageCap _ _ sz _) \<Rightarrow> sz
|
||||
|
@ -3122,14 +2931,6 @@ lemma set_cap_noop_dcorres3:
|
|||
simp_all add: get_ipc_buffer_words_def)
|
||||
done
|
||||
|
||||
lemma zombie_cap_two_nonidles:
|
||||
"\<lbrakk> caps_of_state s ptr = Some (cap.Zombie ptr' zbits n); invs s \<rbrakk>
|
||||
\<Longrightarrow> fst ptr \<noteq> idle_thread s \<and> ptr' \<noteq> idle_thread s"
|
||||
apply (frule valid_global_refsD2, clarsimp+)
|
||||
apply (simp add: cap_range_def global_refs_def)
|
||||
apply (cases ptr, auto dest: valid_idle_has_null_cap[rotated -1])[1]
|
||||
done
|
||||
|
||||
lemma finalise_zombie:
|
||||
"is_zombie cap \<Longrightarrow> finalise_cap cap final
|
||||
= do assert final; return (cap, None) od"
|
||||
|
@ -3276,7 +3077,6 @@ lemma throw_or_return_preemption_corres:
|
|||
update_work_units_def wrap_ext_bool_det_ext_ext_def work_units_limit_def
|
||||
work_units_limit_reached_def OR_choiceE_def reset_work_units_def mk_ef_def
|
||||
split: option.splits kernel_object.splits)
|
||||
(* sseefried: Wait, what? Who says we don't do automatic proofs! *)
|
||||
done
|
||||
|
||||
lemma cutMon_fail:
|
||||
|
|
|
@ -16,23 +16,16 @@ begin
|
|||
Lift the property from abstract spec to capdl model
|
||||
*)
|
||||
|
||||
(* MOVE to top level *)
|
||||
(* FIXME: MOVE to top level *)
|
||||
|
||||
declare fun_upd_restrict_conv[simp del]
|
||||
|
||||
|
||||
definition not_idle_thread:: "obj_ref \<Rightarrow> 'z::state_ext state \<Rightarrow> bool"
|
||||
where "not_idle_thread x \<equiv> (%s. x\<noteq> idle_thread s)"
|
||||
where "not_idle_thread x \<equiv> (\<lambda>s. x \<noteq> idle_thread s)"
|
||||
|
||||
(*Some trivial lemmas rule out irq_node and idle_thread*)
|
||||
|
||||
lemma get_tcb_rev:
|
||||
"kheap s p = Some (TCB t)\<Longrightarrow> get_tcb p s = Some t"
|
||||
by (clarsimp simp:get_tcb_def)
|
||||
|
||||
lemma get_etcb_rev:
|
||||
"ekheap s p = Some etcb \<Longrightarrow> get_etcb p s = Some etcb"
|
||||
by (clarsimp simp: get_etcb_def)
|
||||
|
||||
lemma ep_not_idle:
|
||||
"\<lbrakk>valid_idle s;obj_at is_ep epptr s\<rbrakk> \<Longrightarrow> not_idle_thread epptr s"
|
||||
by (clarsimp simp:valid_idle_def obj_at_def is_cap_table_def st_tcb_at_def is_ep_def not_idle_thread_def)
|
||||
|
@ -41,6 +34,11 @@ lemma aep_not_idle:
|
|||
"\<lbrakk>valid_idle s;obj_at is_aep epptr s\<rbrakk> \<Longrightarrow> not_idle_thread epptr s"
|
||||
by (clarsimp simp:valid_idle_def obj_at_def is_cap_table_def st_tcb_at_def is_aep_def not_idle_thread_def)
|
||||
|
||||
lemma cte_wp_at_zombie_not_idle:
|
||||
"\<lbrakk>cte_wp_at (op = (cap.Zombie ptr' zbits n)) ptr s; invs s\<rbrakk> \<Longrightarrow> not_idle_thread (fst ptr) s"
|
||||
"\<lbrakk>cte_wp_at (op = (cap.Zombie ptr' zbits n)) ptr s; invs s\<rbrakk> \<Longrightarrow> not_idle_thread ptr' s"
|
||||
by (auto dest!: zombie_cap_two_nonidles simp: cte_wp_at_caps_of_state not_idle_thread_def)
|
||||
|
||||
lemmas tcb_slots = Types_D.tcb_caller_slot_def Types_D.tcb_cspace_slot_def Types_D.tcb_ipcbuffer_slot_def
|
||||
Types_D.tcb_pending_op_slot_def Types_D.tcb_replycap_slot_def Types_D.tcb_vspace_slot_def
|
||||
|
||||
|
@ -85,11 +83,6 @@ lemma transform_objects_kheap:
|
|||
unfolding transform_objects_def
|
||||
by (simp)
|
||||
|
||||
(* FIXME: move *)
|
||||
lemma get_etcb_SomeD: "get_etcb ptr s = Some v \<Longrightarrow> ekheap s ptr = Some v"
|
||||
apply (case_tac "ekheap s ptr", simp_all add: get_etcb_def)
|
||||
done
|
||||
|
||||
lemma transform_objects_tcb:
|
||||
"\<lbrakk> get_tcb ptr s = Some tcb; get_etcb ptr s = Some etcb; ptr \<noteq> idle_thread s\<rbrakk>
|
||||
\<Longrightarrow> transform_objects s ptr = Some (transform_tcb (machine_state s) ptr tcb etcb)"
|
||||
|
@ -101,23 +94,6 @@ lemma opt_object_tcb:
|
|||
opt_object ptr (transform s) = Some (transform_tcb (machine_state s) ptr tcb etcb)"
|
||||
by (clarsimp simp: opt_object_def transform_def transform_objects_tcb dest!: get_tcb_SomeD)
|
||||
|
||||
(* MOVE *)
|
||||
lemma inj_on_tcb_cnode_index:
|
||||
"inj_on tcb_cnode_index {n. n < 8}"
|
||||
proof (rule inj_onI)
|
||||
fix x y
|
||||
|
||||
note [simp del] = bin_to_bl_def
|
||||
|
||||
assume "x \<in> {n. n < 8}" and "y \<in> {n. n < 8}" and tci: "tcb_cnode_index x = tcb_cnode_index y"
|
||||
hence "x < 8" and "y < 8" by auto
|
||||
thus "x = y" using tci
|
||||
apply (simp add: tcb_cnode_index_def)
|
||||
apply (erule inj_onD [OF word_unat.Abs_inj_on])
|
||||
apply (simp_all add: unats_def)
|
||||
done
|
||||
qed
|
||||
|
||||
abbreviation
|
||||
"tcb_abstract_slots \<equiv> {tcb_caller_slot, tcb_cspace_slot, tcb_ipcbuffer_slot, tcb_replycap_slot, tcb_vspace_slot}"
|
||||
|
||||
|
@ -186,18 +162,6 @@ lemma caps_of_object_update_context [simp]:
|
|||
apply (simp add: tcb_cap_cases_def split: split_if)
|
||||
done
|
||||
|
||||
|
||||
(* MOVE *)
|
||||
(* The equality is here so that any transform_object rewrites can fire and are confluent
|
||||
FIXME: remove the transform_obj_ptr?
|
||||
*)
|
||||
lemma transform_objects_update_same:
|
||||
"\<lbrakk> kheap s ptr = Some ko; transform_object (machine_state s) ptr (ekheap s ptr) ko = ko'; ptr \<noteq> idle_thread s \<rbrakk>
|
||||
\<Longrightarrow> (transform_objects s)(ptr \<mapsto> ko') = transform_objects s"
|
||||
unfolding transform_objects_def
|
||||
by (rule ext) (simp)
|
||||
|
||||
(* FIXME: add simp rules *)
|
||||
definition
|
||||
generates_pending :: "Structures_A.thread_state \<Rightarrow> bool"
|
||||
where
|
||||
|
@ -278,26 +242,11 @@ lemma duplicate_corrupt_tcb_intent:
|
|||
split:option.splits cdl_object.splits)
|
||||
done
|
||||
|
||||
|
||||
(* FIXME: Move *)
|
||||
lemma valid_etcbs_tcb_etcb: "\<And>s. \<lbrakk> valid_etcbs s; kheap s ptr = Some (TCB tcb) \<rbrakk> \<Longrightarrow> \<exists>etcb. ekheap s ptr = Some etcb"
|
||||
apply (clarsimp simp: valid_etcbs_def valid_etcbs_def st_tcb_at_def obj_at_def is_etcb_at_def)
|
||||
apply (erule_tac x=ptr in allE)
|
||||
apply clarsimp
|
||||
done
|
||||
|
||||
lemma valid_etcbs_get_tcb_get_etcb: "\<And>s. \<lbrakk> valid_etcbs s; get_tcb ptr s = Some tcb \<rbrakk> \<Longrightarrow> \<exists>etcb. get_etcb ptr s = Some etcb"
|
||||
apply (clarsimp simp: valid_etcbs_def valid_etcbs_def st_tcb_at_def obj_at_def is_etcb_at_def get_etcb_def get_tcb_def split: option.splits split_if)
|
||||
apply (erule_tac x=ptr in allE)
|
||||
apply (case_tac a)
|
||||
apply (clarsimp simp: get_etcb_def split: option.splits kernel_object.splits)+
|
||||
done
|
||||
|
||||
lemma valid_etcbs_ko_etcb: "\<And>s. \<lbrakk> valid_etcbs s; kheap s ptr = Some ko \<rbrakk> \<Longrightarrow> \<exists>tcb. (ko = TCB tcb = (\<exists>etcb. ekheap s ptr = Some etcb))"
|
||||
apply (clarsimp simp: valid_etcbs_def valid_etcbs_def st_tcb_at_def obj_at_def is_etcb_at_def)
|
||||
apply (erule_tac x="ptr" in allE)
|
||||
apply auto
|
||||
done
|
||||
(* Corrupting a register several times is the same as corrupting it once. *)
|
||||
lemma corres_corrupt_tcb_intent_dupl:
|
||||
"\<lbrakk> dcorres dc P P' (do corrupt_tcb_intent x; corrupt_tcb_intent x od) g \<rbrakk> \<Longrightarrow>
|
||||
dcorres dc P P' (corrupt_tcb_intent x) g"
|
||||
by (subst duplicate_corrupt_tcb_intent[symmetric], simp)
|
||||
|
||||
(*
|
||||
* Doing nothing at the abstract level corresponds to corrupting a TCB intent.
|
||||
|
@ -1327,23 +1276,6 @@ lemma ipc_frame_ptr_at_frame_at:
|
|||
apply (clarsimp split:Structures_A.kernel_object.splits if_splits arch_kernel_obj.splits)
|
||||
done
|
||||
|
||||
lemma distinct_element:
|
||||
"\<lbrakk>b\<inter>d = {};a\<in> b;c\<in>d\<rbrakk>\<Longrightarrow> a\<noteq>c"
|
||||
by auto
|
||||
|
||||
(* FIXME: lemma and_neg_mask_plus_mask_mono generalizes this lemma. *)
|
||||
lemma word_neg_and_le:
|
||||
shows "n<32 \<Longrightarrow> (ptr::word32) \<le> (ptr && ~~ mask n) + ((2::word32) ^ n - 1)"
|
||||
apply (subst add_commute)
|
||||
apply (rule word_le_minus_cancel[where x = "ptr && ~~ mask n"])
|
||||
apply (clarsimp simp:subtract_mask)
|
||||
using word_and_le1[where a = "mask n" and y = ptr]
|
||||
apply (clarsimp simp:mask_def word_le_less_eq plus_minus_one_rewrite32)
|
||||
apply (subst add_commute)
|
||||
apply (rule is_aligned_no_overflow')
|
||||
apply (clarsimp simp: is_aligned_neg_mask)
|
||||
done
|
||||
|
||||
lemma ipc_buffer_within_frame:
|
||||
"\<lbrakk>
|
||||
pspace_aligned s;
|
||||
|
@ -1356,35 +1288,26 @@ lemma ipc_buffer_within_frame:
|
|||
\<Longrightarrow> ptr \<noteq> ptr'"
|
||||
apply (simp add:pspace_aligned_def obj_at_def)
|
||||
apply (frule_tac x = buf in bspec)
|
||||
apply (simp add:domI obj_at_def)
|
||||
apply (simp add:domI obj_at_def)
|
||||
apply (drule_tac x = buf' in bspec)
|
||||
apply (clarsimp simp:domI obj_at_def)
|
||||
apply (clarsimp simp:domI obj_at_def)
|
||||
apply (unfold pspace_distinct_def)
|
||||
apply (drule_tac x = buf in spec)
|
||||
apply (drule_tac x = buf' in spec)
|
||||
apply (drule_tac x = "ArchObj (DataPage sz)" in spec)
|
||||
apply (drule_tac x = "ArchObj (DataPage sz')" in spec)
|
||||
apply (rule distinct_element)
|
||||
apply (rule mp)
|
||||
apply (assumption)
|
||||
apply (rule mp)
|
||||
apply (assumption)
|
||||
apply simp
|
||||
apply (clarsimp simp:is_aligned_def)
|
||||
apply (simp add:word_and_le2)
|
||||
apply (rule word_neg_and_le)
|
||||
apply (simp add:pageBitsForSize_def,case_tac sz,simp+)
|
||||
apply (clarsimp simp:word_and_le2)
|
||||
apply (rule word_neg_and_le)
|
||||
apply (simp add:pageBitsForSize_def,case_tac sz',clarsimp+)
|
||||
apply (clarsimp simp:is_aligned_def word_and_le2 word_neg_and_le)
|
||||
apply (clarsimp simp:word_and_le2 word_neg_and_le)
|
||||
done
|
||||
|
||||
definition
|
||||
within_page :: "word32 \<Rightarrow> word32 \<Rightarrow> vmpage_size \<Rightarrow> bool"
|
||||
where "within_page buf ptr sz \<equiv> (ptr && ~~ mask (pageBitsForSize sz) = buf)"
|
||||
|
||||
lemma is_aligned_after_mask:
|
||||
"\<lbrakk>is_aligned k m;m\<le> n\<rbrakk> \<Longrightarrow> is_aligned (k && mask n) m"
|
||||
by (metis is_aligned_andI1)
|
||||
|
||||
lemma valid_tcb_obj_ipc_align_etc:
|
||||
"\<lbrakk>valid_objs s;pspace_aligned s;ipc_frame_ptr_at buf thread s;ipc_frame_sz_at pgsz thread s;
|
||||
get_tcb thread s = Some tcb\<rbrakk>
|
||||
|
@ -1861,14 +1784,6 @@ lemma store_word_offs_ipc_frame_wp:
|
|||
"\<lbrace>ipc_frame_wp_at P s_id\<rbrace> store_word_offs base a aa \<lbrace>\<lambda>rv. ipc_frame_wp_at P s_id\<rbrace>"
|
||||
by (wp | simp add: ipc_frame_wp_at_def store_word_offs_def)+
|
||||
|
||||
(* FIXME: move *)
|
||||
lemma corres_state_assert:
|
||||
"corres_underlying sr nf rr P Q f (g ()) \<Longrightarrow>
|
||||
(\<And>s. Q s \<Longrightarrow> R s) \<Longrightarrow>
|
||||
corres_underlying sr nf rr P Q f (state_assert R >>= g)"
|
||||
by (clarsimp simp: corres_underlying_def state_assert_def get_def assert_def
|
||||
return_def bind_def)
|
||||
|
||||
lemma zip_cpy_word_corres:
|
||||
"\<forall>x\<in> set xs. within_page buf (base + of_nat(x*word_size)) sz
|
||||
\<Longrightarrow> dcorres dc \<top>
|
||||
|
|
|
@ -93,129 +93,8 @@ lemma handle_reply_cur_thread_idle_thread:
|
|||
apply simp+
|
||||
done
|
||||
|
||||
lemma word_of_nat_less:
|
||||
"\<lbrakk> n < unat x \<rbrakk> \<Longrightarrow> of_nat n < x"
|
||||
apply (simp add: word_less_nat_alt)
|
||||
apply (erule order_le_less_trans[rotated])
|
||||
apply (simp add: unat_of_nat)
|
||||
done
|
||||
|
||||
lemma word_shift:
|
||||
"\<lbrakk>is_aligned (ptr :: word32) 2; s \<le> mask 2\<rbrakk>\<Longrightarrow>((ptr + s) >> 2) = (ptr >> 2)"
|
||||
apply (simp add: le_mask_iff_lt_2n[THEN iffD1])
|
||||
apply (rule sym)
|
||||
apply (rule subst[OF mask_shift])
|
||||
apply (simp add: is_aligned_add_helper[THEN conjunct2,where d1=s and n1=2 and p1=ptr,simplified])
|
||||
done
|
||||
|
||||
|
||||
(* MOVE *)
|
||||
lemma is_aligned_shiftr_shiftl:
|
||||
"is_aligned w n \<Longrightarrow> w >> n << n = w"
|
||||
apply (simp add: shiftr_shiftl1)
|
||||
apply (erule is_aligned_neg_mask_eq)
|
||||
done
|
||||
|
||||
declare pbfs_atleast_pageBits [simp]
|
||||
|
||||
(* MOVE *)
|
||||
lemma corres_symb_gets:
|
||||
"(\<And>v. corres_underlying sr nf r P (Q v) a (g v)) \<Longrightarrow>
|
||||
corres_underlying sr nf r P (\<lambda>s. Q (f s) s) a (gets f >>= g)"
|
||||
apply (rule corres_symb_exec_r)
|
||||
apply assumption
|
||||
apply (wp)
|
||||
apply (rule no_fail_pre [OF non_fail_gets])
|
||||
apply simp
|
||||
done
|
||||
|
||||
lemma dcorres_from_vcg:
|
||||
assumes v: "\<And>s. \<lbrace>(op = s) and P' and (P o transform)\<rbrace> f' \<lbrace>\<lambda>rv' s'. \<exists>rv. (rv, transform s') \<in> fst (f (transform s)) \<and> r rv rv'\<rbrace>"
|
||||
shows "dcorres r P P' f f'"
|
||||
unfolding corres_underlying_def
|
||||
apply (clarsimp simp: split_def)
|
||||
apply (drule use_valid [OF _ v])
|
||||
apply simp
|
||||
apply (clarsimp elim!: bexI [rotated])
|
||||
done
|
||||
|
||||
lemma dcorres_from_vcg_dc:
|
||||
assumes v: "\<And>s. \<lbrace>(op = s) and P' and (P o transform)\<rbrace> f' \<lbrace>\<lambda>_ s'. transform s' \<in> snd ` fst (f (transform s))\<rbrace>"
|
||||
shows "dcorres dc P P' f f'"
|
||||
apply (rule dcorres_from_vcg)
|
||||
apply (rule hoare_strengthen_post [OF v])
|
||||
apply fastforce
|
||||
done
|
||||
|
||||
lemma corres_underlying_split2:
|
||||
assumes ac: "corres_underlying s nf r' G G' a c"
|
||||
assumes valid: "\<lbrace>G\<rbrace> a \<lbrace>\<lambda>rv s. \<forall>rv'. r' rv rv' \<longrightarrow> P rv rv' s\<rbrace>" "\<lbrace>G'\<rbrace> c \<lbrace>\<lambda>rv' s. \<forall>rv. r' rv rv' \<longrightarrow> P' rv rv' s\<rbrace>"
|
||||
assumes bd: "\<forall>rv rv'. r' rv rv' \<longrightarrow>
|
||||
corres_underlying s nf r (P rv rv') (P' rv rv') (b rv) (d rv')"
|
||||
shows "corres_underlying s nf r G G' (a >>= (\<lambda>rv. b rv)) (c >>= (\<lambda>rv'. d rv'))"
|
||||
using ac bd valid
|
||||
apply (clarsimp simp: corres_underlying_def bind_def)
|
||||
apply (drule (1) bspec, clarsimp simp: split_def)
|
||||
apply (rule conjI)
|
||||
apply clarsimp
|
||||
apply (drule (1) bspec, erule bexE)
|
||||
apply clarsimp
|
||||
apply (erule allE)+
|
||||
apply (erule (1) impE)
|
||||
apply (subgoal_tac "P ac aaa bc \<and> P' ac aaa baa")
|
||||
prefer 2
|
||||
apply (simp add: valid_def)
|
||||
apply (erule allE)+
|
||||
apply (erule (1) impE)+
|
||||
apply blast
|
||||
apply (drule_tac x="(bc, baa)" in bspec, assumption)
|
||||
apply simp
|
||||
apply (clarsimp simp: split_def)
|
||||
apply (rule bexI)
|
||||
prefer 2
|
||||
apply assumption
|
||||
apply clarsimp
|
||||
apply (drule (1) bspec, erule bexE)
|
||||
apply auto[1]
|
||||
apply clarsimp
|
||||
apply (drule (1) bspec, clarsimp)
|
||||
apply (erule allE)+
|
||||
apply (erule (1) impE)
|
||||
apply (clarsimp simp: valid_def)
|
||||
apply force
|
||||
done
|
||||
|
||||
text {* Derivative splitting rules *}
|
||||
|
||||
lemma corres_split2:
|
||||
assumes y: "\<And>rv rv'. r' rv rv' \<Longrightarrow> corres_underlying sr nf r (R rv rv') (R' rv rv') (b rv) (d rv')"
|
||||
assumes x: "corres_underlying sr nf r' P P' a c"
|
||||
assumes valid: "\<lbrace>P\<rbrace> a \<lbrace>\<lambda>rv s. \<forall>rv'. r' rv rv' \<longrightarrow> R rv rv' s\<rbrace>" "\<lbrace>P'\<rbrace> c \<lbrace>\<lambda>rv' s. \<forall>rv. r' rv rv' \<longrightarrow> R' rv rv' s\<rbrace>"
|
||||
shows "corres_underlying sr nf r (P and Q) (P' and Q') (a >>= (\<lambda>rv. b rv)) (c >>= (\<lambda>rv'. d rv'))"
|
||||
using assms
|
||||
apply -
|
||||
apply (rule corres_underlying_split2)
|
||||
apply (rule corres_guard_imp, rule x, simp_all)
|
||||
apply (rule hoare_weaken_pre, assumption)
|
||||
apply simp
|
||||
apply (rule hoare_weaken_pre, assumption)
|
||||
apply simp
|
||||
apply clarsimp
|
||||
done
|
||||
|
||||
lemma corres_underlying_sndE:
|
||||
assumes cul: "corres_underlying R F r P P' a c"
|
||||
and sr: "(s, s') \<in> R" "P s" "P' s'"
|
||||
and rl: "\<lbrakk>F \<longrightarrow> \<not> snd (c s')\<rbrakk> \<Longrightarrow> Q"
|
||||
shows "Q"
|
||||
using cul sr
|
||||
unfolding corres_underlying_def
|
||||
apply clarsimp
|
||||
apply (rule rl)
|
||||
apply (drule (1) bspec)
|
||||
apply clarsimp
|
||||
done
|
||||
|
||||
lemma dcorres_move_guard:
|
||||
"dcorres r P P' a c \<Longrightarrow> dcorres r \<top> (P' and (P o transform)) a c"
|
||||
apply (rule corres_underlyingI)
|
||||
|
@ -238,11 +117,6 @@ lemma dcorres_move_guard_iff:
|
|||
apply (erule dcorres_move_guard_back)
|
||||
done
|
||||
|
||||
(* MOVE *)
|
||||
lemma option_map_cong:
|
||||
"\<lbrakk>(\<And>x. f x = f' x); v = v' \<rbrakk> \<Longrightarrow> option_map f v = option_map f' v'"
|
||||
unfolding option_map_def by (cases v, auto)
|
||||
|
||||
lemma transform_cdt_cong:
|
||||
"\<lbrakk> cdt s = cdt s'; interrupt_irq_node s = interrupt_irq_node s' \<rbrakk> \<Longrightarrow> transform_cdt s = transform_cdt s'"
|
||||
unfolding transform_cdt_def by (simp)
|
||||
|
@ -382,8 +256,6 @@ lemma evalMonad_singleton [simp]:
|
|||
lemma mi_length_dest : "of_nat (unat (mi_length rva)) = mi_length rva"
|
||||
by (rule word_unat.Rep_inverse)
|
||||
|
||||
(* Some map we need to talk about the retrun value and pre condition have been moved to KHeap_DR *)
|
||||
|
||||
(* CORRES RULES *)
|
||||
|
||||
lemma no_pending_cap_when_active_corres:
|
||||
|
@ -470,18 +342,6 @@ lemma meqv_sym:
|
|||
unfolding monadic_rewrite_def
|
||||
by fastforce
|
||||
|
||||
(* MOVE *)
|
||||
lemma wpc_helper_monadic_rewrite:
|
||||
"monadic_rewrite F E Q' m m'
|
||||
\<Longrightarrow> wpc_helper (P, P') (Q, {s. Q' s}) (monadic_rewrite F E (\<lambda>s. s \<in> P') m m')"
|
||||
apply (clarsimp simp: wpc_helper_def)
|
||||
apply (erule monadic_rewrite_imp)
|
||||
apply auto
|
||||
done
|
||||
|
||||
wpc_setup "\<lambda>m. monadic_rewrite F E Q' m m'" wpc_helper_monadic_rewrite
|
||||
wpc_setup "\<lambda>m. monadic_rewrite F E Q' (m >>= c) m'" wpc_helper_monadic_rewrite
|
||||
|
||||
lemma monadic_rewrite_modify_remove_stateful:
|
||||
assumes "\<And>v. meqv (P v) m (modify (\<lambda>s. f v s))"
|
||||
shows "meqv (\<lambda>s. P (g s) s) m (modify (\<lambda>s. f (g s) s))"
|
||||
|
@ -631,22 +491,6 @@ done
|
|||
lemma corres_ignore_ret_lhs: "dcorres rv P P' (do f;g od) f' \<Longrightarrow> dcorres rv P P' (do a\<leftarrow>f;g od) f'"
|
||||
by (clarsimp simp:corres_underlying_def)
|
||||
|
||||
(* FIXME: Move to Intent_DR
|
||||
* Corrupting a register several times is the same as corrupting it once.
|
||||
*)
|
||||
lemma corres_corrupt_tcb_intent_dupl:
|
||||
"\<lbrakk> dcorres dc P P' (do corrupt_tcb_intent x; corrupt_tcb_intent x od) g \<rbrakk> \<Longrightarrow>
|
||||
dcorres dc P P' (corrupt_tcb_intent x) g"
|
||||
by (subst duplicate_corrupt_tcb_intent[symmetric], simp)
|
||||
|
||||
(* FIXME: sseefried: Move to Intent_DR.thy *)
|
||||
lemma valid_etcbs_tcb_at_is_etcb_at: "\<And>s. \<lbrakk> valid_etcbs s; tcb_at ptr s \<rbrakk> \<Longrightarrow> is_etcb_at ptr s"
|
||||
apply (clarsimp simp: valid_etcbs_def valid_etcbs_def st_tcb_at_def obj_at_def is_etcb_at_def is_tcb_def)
|
||||
apply (erule_tac x=ptr in allE)
|
||||
apply (case_tac "ko")
|
||||
apply auto
|
||||
done
|
||||
|
||||
lemma dcorres_do_async_transfer:
|
||||
"dcorres dc \<top>
|
||||
(valid_objs and pspace_aligned and pspace_distinct and valid_idle and tcb_at oid and not_idle_thread oid and valid_etcbs)
|
||||
|
@ -658,7 +502,7 @@ lemma dcorres_do_async_transfer:
|
|||
in corres_symb_exec_r)
|
||||
apply (rule dcorres_expand_pfx)
|
||||
apply (clarsimp)
|
||||
apply (frule(1) valid_etcbs_tcb_at_is_etcb_at)
|
||||
apply (frule(1) tcb_at_is_etcb_at)
|
||||
apply (case_tac rv)
|
||||
apply (rule corres_symb_exec_l)
|
||||
apply (rule_tac F="rva = None" in corres_gen_asm)
|
||||
|
@ -781,36 +625,6 @@ lemma attempt_switch_to_dcorres: "dcorres dc P P' (return ()) (attempt_switch_to
|
|||
apply (rule possible_switch_to_dcorres)
|
||||
done
|
||||
|
||||
lemma dcorres_dc_rhs_noop_below_gen:
|
||||
"\<lbrakk> \<forall>rv'. dcorres dc (Q ()) (Q' rv') (return ()) (m rv');
|
||||
dcorres dc P P' f g;
|
||||
\<lbrace> P \<rbrace> f \<lbrace> Q \<rbrace>; \<lbrace> P' \<rbrace> g \<lbrace> Q' \<rbrace> \<rbrakk>
|
||||
\<Longrightarrow> dcorres dc P P' f (g >>= m)"
|
||||
apply (rule corres_add_noop_lhs2)
|
||||
apply (rule corres_underlying_split)
|
||||
apply (assumption | clarsimp)+
|
||||
done
|
||||
|
||||
(* FIXME: sseefried: Move to Corres_D.thy *)
|
||||
|
||||
|
||||
lemma dcorres_dc_rhs_noop_below_2: "\<lbrakk> \<forall>rv'. dcorres dc (Q ()) (Q' rv') (return ()) m;
|
||||
dcorres dc P P' f (g >>= h);
|
||||
\<lbrace> P \<rbrace> f \<lbrace> Q \<rbrace>;
|
||||
\<lbrace> P' \<rbrace> g \<lbrace> R'\<rbrace>;
|
||||
(\<And>y. \<lbrace> R' y \<rbrace> h y \<lbrace> Q' \<rbrace>)
|
||||
\<rbrakk>
|
||||
\<Longrightarrow> dcorres dc P P' f (do x \<leftarrow> g;
|
||||
_ \<leftarrow> h x;
|
||||
m
|
||||
od)"
|
||||
apply (simp add: bind_assoc[symmetric])
|
||||
apply (rule dcorres_dc_rhs_noop_below_gen)
|
||||
apply (wp | simp | assumption)+
|
||||
done
|
||||
|
||||
lemmas dcorres_dc_rhs_noop_below_2_True = dcorres_dc_rhs_noop_below_2[OF _ _ hoare_TrueI hoare_TrueI hoare_TrueI]
|
||||
|
||||
lemma corres_update_waiting_aep_do_async_transfer:
|
||||
"dcorres dc \<top>
|
||||
(pspace_aligned and valid_objs and valid_mdb and pspace_distinct and valid_idle and valid_etcbs and
|
||||
|
@ -1524,7 +1338,6 @@ lemma dcorres_return:
|
|||
by (clarsimp simp:return_def corres_underlying_def)
|
||||
|
||||
|
||||
(* will need more preconds *)
|
||||
lemma get_receive_slot_dcorres:
|
||||
"dcorres (\<lambda>d d'. d = dest_of d') \<top>
|
||||
((\<lambda>s. evalMonad (lookup_ipc_buffer True t) s = Some buffer) and not_idle_thread t and valid_objs
|
||||
|
@ -1799,7 +1612,7 @@ lemma dcorres_copy_mrs':
|
|||
apply (wp|clarsimp simp:option.cases split:option.splits)+
|
||||
apply (clarsimp simp:get_ipc_buffer_def gets_the_def exs_valid_def gets_def
|
||||
get_def bind_def return_def assert_opt_def fail_def split:option.splits | rule conjI)+
|
||||
apply (frule(1) valid_etcbs_tcb_at_is_etcb_at, clarsimp simp: is_etcb_at_def, fold get_etcb_def)
|
||||
apply (frule(1) tcb_at_is_etcb_at, clarsimp simp: is_etcb_at_def, fold get_etcb_def)
|
||||
apply (clarsimp simp:not_idle_thread_def tcb_at_def opt_cap_tcb)+
|
||||
apply (simp split:cdl_cap.splits)
|
||||
apply (wp cdl_get_ipc_buffer_None)
|
||||
|
|
|
@ -462,21 +462,6 @@ lemma eq_singleton_set: "\<lbrakk>A = f` B; \<forall>x\<in>B. \<forall>y\<in> B.
|
|||
apply clarsimp
|
||||
done
|
||||
|
||||
lemma valid_idle_has_null_cap:
|
||||
"\<lbrakk> if_unsafe_then_cap s; valid_global_refs s;valid_idle s;valid_irq_node s\<rbrakk>
|
||||
\<Longrightarrow> caps_of_state s (idle_thread s, v) = Some cap
|
||||
\<Longrightarrow> cap = cap.NullCap"
|
||||
apply (rule ccontr)
|
||||
apply (drule(1) if_unsafe_then_capD[OF caps_of_state_cteD])
|
||||
apply clarsimp
|
||||
apply (clarsimp simp: ex_cte_cap_wp_to_def cte_wp_at_caps_of_state)
|
||||
apply (frule(1) valid_global_refsD2)
|
||||
apply (case_tac capa, simp_all add: cap_range_def global_refs_def)[1]
|
||||
apply (clarsimp simp: valid_irq_node_def valid_idle_def st_tcb_at_def
|
||||
obj_at_def is_cap_table_def)
|
||||
apply (drule_tac x=word in spec, simp)
|
||||
done
|
||||
|
||||
lemma final_cap_set_map:
|
||||
"\<lbrakk>valid_idle s'; valid_irq_node s';valid_objs s';if_unsafe_then_cap s'; valid_global_refs s'; cap_counts (transform_cap cap); valid_etcbs s'\<rbrakk>
|
||||
\<Longrightarrow> {cref. opt_cap_wp_at (\<lambda>cap'. cap_object (transform_cap cap) = cap_object cap'
|
||||
|
@ -739,7 +724,6 @@ lemma transform_full_intent_caps_cong_weak:
|
|||
transform_full_intent ms p tcb = transform_full_intent ms p tcb'"
|
||||
by (rule transform_full_intent_caps_cong) auto
|
||||
|
||||
(* FIXME: sseefried: Not sure whether I need tcb_domain etcb = tcb_domain etcb'. If not then make both etcb and etcb' the same *)
|
||||
lemma transform_full_intent_same_cap:
|
||||
"\<lbrakk> transform_cap (tcb_ipcframe tcb) = transform_cap cap' \<rbrakk>
|
||||
\<Longrightarrow> transform_full_intent ms p' (tcb\<lparr>tcb_ipcframe := cap'\<rparr>) =
|
||||
|
@ -1277,15 +1261,6 @@ lemma transform_cdt_single_remove_helper:
|
|||
"s'= transform s \<Longrightarrow> cdl_cdt (cdl_cdt_single_remove s' a) = transform_cdt (abs_cdt_single_remove s a') \<Longrightarrow> (cdl_cdt_single_remove (transform s) a) = transform (abs_cdt_single_remove s a')"
|
||||
by (clarsimp simp:cdl_cdt_single_remove_def abs_cdt_single_remove_def transform_def transform_current_thread_def transform_asid_table_def)
|
||||
|
||||
|
||||
(* FIXME: sseefried: This just isn't true anymore.
|
||||
trans_state is used to translate extended state. Now that 'transform' depends on extended state this lemma is no longer true *)
|
||||
(*lemma transform_extended:
|
||||
"transform (trans_state ba b) = transform b"
|
||||
apply (simp add: transform_def transform_objects_def transform_cdt_def
|
||||
transform_current_thread_def transform_asid_table_def)
|
||||
done *)
|
||||
|
||||
lemma remove_parent_corres:
|
||||
"dcorres dc \<top> (cte_at slot and weak_valid_mdb)
|
||||
(remove_parent (transform_cslot_ptr slot))
|
||||
|
@ -1344,19 +1319,6 @@ defer
|
|||
apply simp+
|
||||
done
|
||||
|
||||
|
||||
(*FIXME: Move up.*)
|
||||
lemma
|
||||
corres_bind_return_r2:
|
||||
"corres_underlying S nf dc P Q f (do x \<leftarrow> ga; (fa x) od) \<Longrightarrow>
|
||||
corres_underlying S nf dc P Q f (do x \<leftarrow> ga;
|
||||
y \<leftarrow> fa x;
|
||||
return (h x)
|
||||
od)"
|
||||
apply (fastforce simp: corres_underlying_def bind_def return_def)
|
||||
done
|
||||
|
||||
|
||||
lemma dmo_maskIRQ_dcorres:
|
||||
"dcorres dc \<top> \<top> (return ()) (do_machine_op (maskInterrupt b st))"
|
||||
apply (clarsimp simp: do_machine_op_def corres_underlying_def return_def select_f_def in_monad)
|
||||
|
@ -2493,14 +2455,6 @@ lemma set_parent_corres:
|
|||
apply (clarsimp split:if_splits)
|
||||
done
|
||||
|
||||
(* FIXME: move *)
|
||||
lemma dcorres_symmetric_bool_cases:
|
||||
"\<lbrakk>P = P'; \<lbrakk>P; P'\<rbrakk> \<Longrightarrow> dcorres r Q Q' f g; \<lbrakk>\<not> P; \<not> P'\<rbrakk> \<Longrightarrow> dcorres r R R' f g\<rbrakk>
|
||||
\<Longrightarrow> dcorres r (\<lambda>s. (P \<longrightarrow> Q s) \<and> (\<not> P \<longrightarrow> R s)) (\<lambda>s. (P' \<longrightarrow> Q' s) \<and> (\<not> P' \<longrightarrow> R' s)) f g"
|
||||
apply (cases P)
|
||||
apply simp+
|
||||
done
|
||||
|
||||
lemma set_tcb_capslot_weak_valid_mdb:
|
||||
"\<lbrace>weak_valid_mdb and cte_wp_at (op=cap.NullCap) slot\<rbrace> set_cap cap slot \<lbrace>\<lambda>r s. weak_valid_mdb s\<rbrace> "
|
||||
apply (simp add: weak_valid_mdb_def cte_wp_at_caps_of_state swp_def)
|
||||
|
@ -2588,15 +2542,6 @@ lemma tcb_reply_cap_cte_wp_at:
|
|||
apply (clarsimp simp:tcb_cap_cases_def split:Structures_A.thread_state.splits)
|
||||
done
|
||||
|
||||
(* FIX ME: move to dcorres_D *)
|
||||
lemma wp_to_dcorres:
|
||||
"(\<And>cs. \<lbrace>\<lambda>s. Q s \<and> transform s = cs\<rbrace> g \<lbrace>\<lambda>r s. transform s = cs\<rbrace>) \<Longrightarrow> dcorres dc (\<lambda>_. True) Q (return x) g"
|
||||
apply (clarsimp simp:corres_underlying_def valid_def return_def)
|
||||
apply (drule_tac x = "transform b" in meta_spec)
|
||||
apply (drule_tac x = b in spec)
|
||||
apply fastforce
|
||||
done
|
||||
|
||||
lemma transform_objects_update_kheap_simp:
|
||||
"\<lbrakk>kheap s ptr = Some ko; ekheap s ptr = opt_etcb\<rbrakk>
|
||||
\<Longrightarrow> transform_objects (update_kheap (kheap s(ptr \<mapsto> obj)) s) =
|
||||
|
@ -3192,11 +3137,6 @@ lemma cdl_current_thread:
|
|||
"(cdl_current_thread (transform s')) = transform_current_thread s'"
|
||||
by (clarsimp simp:transform_def )
|
||||
|
||||
(* FIXME: move *)
|
||||
lemma corres_dummy_returnOk_r:
|
||||
"dcorres c P P' f (g >>=E returnOk) \<Longrightarrow> dcorres c P P' f g"
|
||||
by simp
|
||||
|
||||
lemma get_cap_get_tcb_dcorres:
|
||||
"dcorres (\<lambda>r t. r = transform_cap (tcb_ctable t)) \<top> (not_idle_thread thread and valid_etcbs)
|
||||
(KHeap_D.get_cap (thread, tcb_cspace_slot))
|
||||
|
|
|
@ -28,10 +28,16 @@ begin
|
|||
no_notation bind_drop (infixl ">>" 60)
|
||||
|
||||
(* FIXME: move *)
|
||||
lemma alternative_com:
|
||||
"(f \<sqinter> g) = (g \<sqinter> f)"
|
||||
apply (rule ext)
|
||||
apply (auto simp: alternative_def)
|
||||
lemma nonempty_pick_in:
|
||||
"a \<noteq> {} \<Longrightarrow> pick a \<in> a"
|
||||
by (metis all_not_in_conv someI_ex)
|
||||
|
||||
lemma pick_singleton[simp]:
|
||||
"pick {a} = a"
|
||||
apply (rule ccontr)
|
||||
apply (cut_tac nonempty_pick_in)
|
||||
apply fastforce
|
||||
apply fastforce
|
||||
done
|
||||
|
||||
end
|
||||
|
|
|
@ -219,26 +219,6 @@ lemma set_scheduler_action_dcorres:
|
|||
"dcorres dc \<top> \<top> (return ()) (set_scheduler_action sa)"
|
||||
by (clarsimp simp: corres_underlying_def set_scheduler_action_def modify_def get_def put_def bind_def return_def)
|
||||
|
||||
(* FIXME: Move to lib/Corres_UL.thy : Hard to believe this rule doesn't exist already *)
|
||||
lemma corres_either_alternate2:
|
||||
"\<lbrakk> corres_underlying sr nf r P R a c; corres_underlying sr nf r Q R b c \<rbrakk>
|
||||
\<Longrightarrow> corres_underlying sr nf r (P or Q) R (a \<sqinter> b) c"
|
||||
apply (simp add: corres_underlying_def alternative_def)
|
||||
apply clarsimp
|
||||
apply (drule (1) bspec, clarsimp)+
|
||||
apply (erule disjE)
|
||||
apply clarsimp
|
||||
apply (drule(1) bspec, clarsimp)
|
||||
apply (rule rev_bexI)
|
||||
apply (erule UnI1)
|
||||
apply simp
|
||||
apply clarsimp
|
||||
apply (drule(1) bspec, clarsimp)
|
||||
apply (rule rev_bexI)
|
||||
apply (erule UnI2)
|
||||
apply simp
|
||||
done
|
||||
|
||||
lemma switch_to_thread_None_dcorres_L:
|
||||
"dcorres dc (\<lambda>s. cdl_current_thread s = None) \<top>
|
||||
(do _ \<leftarrow> change_current_domain;
|
||||
|
@ -340,14 +320,6 @@ lemma schedule_switch_thread_helper:
|
|||
split: option.splits split_if Structures_A.kernel_object.splits Structures_A.thread_state.splits)
|
||||
done
|
||||
|
||||
(* FIXME: Move *)
|
||||
lemma idle_thread_idle[wp]:
|
||||
"\<lbrace>\<lambda>s. valid_idle s \<and> t = idle_thread s\<rbrace> get_thread_state t \<lbrace>\<lambda>r s. idle r\<rbrace>"
|
||||
apply (clarsimp simp: valid_def get_thread_state_def thread_get_def bind_def return_def gets_the_def gets_def get_def assert_opt_def get_tcb_def
|
||||
fail_def valid_idle_def st_tcb_at_def obj_at_def
|
||||
split: option.splits Structures_A.kernel_object.splits)
|
||||
done
|
||||
|
||||
lemma schedule_switch_thread_dcorres:
|
||||
"\<And>cur_ts. dcorres dc \<top>
|
||||
(\<lambda>s. cur = cur_thread s \<and> st_tcb_at (op = cur_ts) cur s \<and> valid_etcbs s \<and> valid_sched s \<and> invs s \<and> scheduler_action s = switch_thread t)
|
||||
|
|
|
@ -70,8 +70,7 @@ abbreviation
|
|||
abbreviation
|
||||
"dupdate_tcb_intent intent tcb\<equiv> tcb \<lparr>cdl_tcb_intent := intent\<rparr>"
|
||||
|
||||
(* FIXME: name *)
|
||||
lemma dummy_update_state[simp]:
|
||||
lemma update_kheap_triv[simp]:
|
||||
"kheap s y = Some obj\<Longrightarrow> update_kheap ((kheap s)(y \<mapsto> obj)) s = s"
|
||||
apply (case_tac s,clarsimp)
|
||||
apply (rule ext,clarsimp)
|
||||
|
@ -107,38 +106,11 @@ lemmas register_overlap_check = msg_info_badge_register_no_overlap
|
|||
cap_msg_info_register_no_overlap
|
||||
badge_cap_register_overlap
|
||||
|
||||
(* These should be cong rules ...
|
||||
lemma get_tcb_message_info_update_heap [simp]:
|
||||
"get_tcb_message_info (update_kheap ptr obj s) = get_tcb_message_info s"
|
||||
by (rule ext) (simp add: get_tcb_message_info_def)
|
||||
|
||||
lemma get_tcb_mrs_update_heap [simp]:
|
||||
"get_tcb_mrs (update_kheap ptr obj' s) = get_tcb_mrs s"
|
||||
by (rule ext) (simp add: Let_def get_tcb_mrs_def get_tcb_mem_mrs_def)
|
||||
|
||||
lemma transform_full_intent_update_heap [simp]:
|
||||
"transform_full_intent (update_kheap ptr obj s) = transform_full_intent s"
|
||||
by (rule ext)+ (simp add: transform_full_intent_def)
|
||||
*)
|
||||
|
||||
lemma transform_full_intent_cong:
|
||||
"\<lbrakk>ms = ms'; ptr = ptr'; tcb_context tcb = tcb_context tcb'; tcb_ipc_buffer tcb = tcb_ipc_buffer tcb'; tcb_ipcframe tcb = tcb_ipcframe tcb'\<rbrakk>
|
||||
\<Longrightarrow> transform_full_intent ms ptr tcb = transform_full_intent ms' ptr' tcb'"
|
||||
by (simp add: transform_full_intent_def get_tcb_message_info_def get_tcb_mrs_def Suc_le_eq get_ipc_buffer_words_def)
|
||||
|
||||
|
||||
(* FIXME: Move *)
|
||||
lemma set_neg_inter: "- A \<inter> - B = - A - B"
|
||||
by auto
|
||||
|
||||
(* FIXME: Move *)
|
||||
lemma restrict_restrict2: "x \<noteq> b \<Longrightarrow> (m |` (- A - {b})) x = (m |` (- A)) x"
|
||||
apply (cut_tac m="(m |` (- A))" and x=x and A = "-{b}" in restrict_in)
|
||||
apply simp
|
||||
apply (cut_tac m=m and A="- A" and B = "-{b}" in restrict_restrict)
|
||||
apply (fastforce simp: set_neg_inter)
|
||||
done
|
||||
|
||||
lemma caps_of_state_eq_lift:
|
||||
"\<forall>cap. cte_wp_at (op=cap) p s = cte_wp_at (op=cap) p s' \<Longrightarrow> caps_of_state s p = caps_of_state s' p"
|
||||
apply (simp add:cte_wp_at_def caps_of_state_def)
|
||||
|
@ -197,8 +169,6 @@ lemma caps_of_state_update_tcb:
|
|||
|
||||
lemmas caps_of_state_upds = caps_of_state_update_tcb caps_of_state_update_same_caps
|
||||
|
||||
(* weak because we don't look at the contents *)
|
||||
|
||||
lemma transform_cdt_kheap_update [simp]:
|
||||
"transform_cdt (kheap_update f s) = transform_cdt s"
|
||||
by (clarsimp simp: transform_cdt_def cong: if_cong)
|
||||
|
@ -235,6 +205,12 @@ lemma transform_objects_update_kheap_same_caps:
|
|||
apply (simp add: option_map_def restrict_map_def map_add_def )
|
||||
done
|
||||
|
||||
lemma transform_objects_update_same:
|
||||
"\<lbrakk> kheap s ptr = Some ko; transform_object (machine_state s) ptr (ekheap s ptr) ko = ko'; ptr \<noteq> idle_thread s \<rbrakk>
|
||||
\<Longrightarrow> (transform_objects s)(ptr \<mapsto> ko') = transform_objects s"
|
||||
unfolding transform_objects_def
|
||||
by (rule ext) (simp)
|
||||
|
||||
text {* Facts about map_lift_over *}
|
||||
|
||||
lemma map_lift_over_eq_Some:
|
||||
|
|
|
@ -77,7 +77,7 @@ where
|
|||
None)
|
||||
| _ \<Rightarrow> None)"
|
||||
|
||||
(* FIXME: arch flags always set to 0 here for now. *)
|
||||
(* Arch flags always set to 0 here as they have no meaning on ARM. *)
|
||||
definition
|
||||
transform_intent_tcb_read_registers :: "word32 list \<Rightarrow> cdl_tcb_intent option"
|
||||
where
|
||||
|
@ -86,7 +86,7 @@ where
|
|||
Some (TcbReadRegistersIntent (flags !! 0) 0 n)
|
||||
| _ \<Rightarrow> None)"
|
||||
|
||||
(* FIXME: arch flags always set to 0 here for now. *)
|
||||
(* Arch flags always set to 0 here as they have no meaning on ARM. *)
|
||||
definition
|
||||
transform_intent_tcb_write_registers :: "word32 list \<Rightarrow> cdl_tcb_intent option"
|
||||
where
|
||||
|
@ -95,8 +95,7 @@ where
|
|||
Some (TcbWriteRegistersIntent (flags !! 0) 0 n values)
|
||||
| _ \<Rightarrow> None)"
|
||||
|
||||
|
||||
(* FIXME: arch flags always set to 0 here for now. *)
|
||||
(* Arch flags always set to 0 here as they have no meaning on ARM. *)
|
||||
definition
|
||||
transform_intent_tcb_copy_registers :: "word32 list \<Rightarrow> cdl_tcb_intent option"
|
||||
where
|
||||
|
@ -105,11 +104,8 @@ where
|
|||
Some (TcbCopyRegistersIntent (flags !! 0) (flags !! 1) (flags !! 2) (flags !! 3) 0)
|
||||
| _ \<Rightarrow> None)"
|
||||
|
||||
|
||||
|
||||
(* FIXME: priority set to 0 here always for now. Seems to be ignored even in the
|
||||
* abstract spec, so whatever we choose is irrelevant. However, it would
|
||||
* surely be nicer to reflect the
|
||||
(* Priority always set to 0 here. This should change if priorities
|
||||
* are ever added to the capDL spec.
|
||||
*)
|
||||
definition
|
||||
transform_priority :: "word32 \<Rightarrow> word8"
|
||||
|
@ -626,7 +622,7 @@ lemma nat_to_bl_id [simp]: "nat_to_bl (size (x :: (('a::len) word))) (unat x) =
|
|||
apply (auto simp: uint_nat le_def word_size)
|
||||
done
|
||||
|
||||
(* MOVE *)
|
||||
(* FIXME: MOVE *)
|
||||
definition
|
||||
option_join :: "'a option option \<Rightarrow> 'a option"
|
||||
where
|
||||
|
@ -634,7 +630,6 @@ where
|
|||
Some (Some y) \<Rightarrow> Some y
|
||||
| _ \<Rightarrow> None"
|
||||
|
||||
(* MOVE *)
|
||||
definition
|
||||
option_map_join :: "('a \<Rightarrow> 'b option) \<Rightarrow> 'a option \<Rightarrow> 'b option"
|
||||
where
|
||||
|
@ -681,8 +676,7 @@ where
|
|||
| _ \<Rightarrow> Types_D.NullCap
|
||||
"
|
||||
|
||||
|
||||
(* MOVE *)
|
||||
(* FIXME: MOVE *)
|
||||
definition
|
||||
evalMonad :: "('s, 'a) nondet_monad \<Rightarrow> 's \<Rightarrow> 'a option"
|
||||
where
|
||||
|
|
|
@ -936,7 +936,7 @@ lemma dcorres_set_intent_error:
|
|||
apply (rule dcorres_absorb_get_l)
|
||||
apply (clarsimp simp: not_idle_thread_def)
|
||||
apply (frule ko_at_tcb_at)
|
||||
apply (frule(1) valid_etcbs_tcb_at_is_etcb_at)
|
||||
apply (frule(1) tcb_at_is_etcb_at)
|
||||
apply (clarsimp simp:tcb_at_def is_etcb_at_def, fold get_etcb_def)
|
||||
apply (clarsimp simp:opt_object_tcb assert_opt_def transform_tcb_def
|
||||
KHeap_D.set_object_def simpler_modify_def corres_underlying_def)
|
||||
|
|
|
@ -12,21 +12,6 @@ theory Untyped_DR
|
|||
imports CNode_DR
|
||||
begin
|
||||
|
||||
lemma is_None_eq_split:
|
||||
"\<lbrakk> (opt = None) = (opt' = None);
|
||||
opt \<noteq> None \<and> opt' \<noteq> None \<longrightarrow> the opt = the opt' \<rbrakk>
|
||||
\<Longrightarrow> opt = opt'"
|
||||
by (cases opt, auto)
|
||||
|
||||
(* FIXME: move? *)
|
||||
lemma trancl_trancl:
|
||||
"(R\<^sup>+)\<^sup>+ = R\<^sup>+"
|
||||
by auto
|
||||
|
||||
lemma minus_Id_Image_singleton:
|
||||
"((S - Id) `` {x}) = ((S `` {x}) - {x})"
|
||||
by blast
|
||||
|
||||
lemma detype_dcorres:
|
||||
"S = {ptr..ptr + 2 ^ sz - 1}
|
||||
\<Longrightarrow> dcorres dc \<top> (\<lambda>s. invs s \<and> (\<exists>cref. cte_wp_at (op = (cap.UntypedCap ptr sz idx)) cref s) \<and> valid_etcbs s)
|
||||
|
@ -48,12 +33,6 @@ lemma detype_dcorres:
|
|||
apply (clarsimp simp: global_refs_def transform_object_def detype_ext_def)
|
||||
done
|
||||
|
||||
(* FIXME: move? *)
|
||||
lemma mapM_simps:
|
||||
"mapM m [] = return []"
|
||||
"mapM m (x#xs) = do r \<leftarrow> m x; rs \<leftarrow> (mapM m xs); return (r#rs) od"
|
||||
by (simp_all add: mapM_def sequence_def)
|
||||
|
||||
(* FIXME: move? *)
|
||||
lemma evalMonad_loadWords:
|
||||
"evalMonad (mapM loadWord xs) ms =
|
||||
|
@ -329,7 +308,6 @@ lemma freeMemory_dcorres:
|
|||
apply (rule_tac y = p in order_trans)
|
||||
apply simp
|
||||
apply (cut_tac ptr = p and n = "pageBitsForSize sz" in word_neg_and_le)
|
||||
apply (simp add:pbfs_less_wb'[unfolded word_bits_def,simplified])
|
||||
apply (simp add:mask_def[unfolded shiftl_t2n,simplified,symmetric] p_assoc_help)
|
||||
apply (thin_tac "?x\<noteq>?y")
|
||||
apply (erule notE)
|
||||
|
@ -338,12 +316,6 @@ lemma freeMemory_dcorres:
|
|||
apply (simp add:mask_def not_le pbfs_less_wb'[unfolded word_bits_def, simplified])
|
||||
done
|
||||
|
||||
(* FIXME: move *)
|
||||
lemma corres_bind_return:
|
||||
"corres_underlying sr nf r P P' (f >>= return) g \<Longrightarrow>
|
||||
corres_underlying sr nf r P P' f g"
|
||||
by (simp add: corres_underlying_def)
|
||||
|
||||
(* FIXME: strictly speaking, we would not need ct_active, here. Proving that,
|
||||
however, requires a stronger version of lemma detype_invariants. *)
|
||||
lemma delete_objects_dcorres:
|
||||
|
@ -494,9 +466,6 @@ lemma obj_bits_bound4:
|
|||
default_arch_object_def pageBits_def)
|
||||
done
|
||||
|
||||
(* FIXME: replacee of_nat_inj32 in Detype_R with this one *)
|
||||
lemmas of_nat_inj32 = of_nat_inj[where 'a=32, folded word_bits_def]
|
||||
|
||||
lemma distinct_retype_addrs:
|
||||
"\<lbrakk>type = Invariants_AI.Untyped \<longrightarrow> 4 \<le> us;
|
||||
range_cover ptr sz (obj_bits_api type us) n\<rbrakk>
|
||||
|
@ -547,7 +516,7 @@ lemma transform_none_to_untyped:
|
|||
done
|
||||
|
||||
|
||||
lemma retype_transform_obj_ref_avaiable:
|
||||
lemma retype_transform_obj_ref_available:
|
||||
notes blah[simp del] = untyped_range.simps usable_untyped_range.simps atLeastAtMost_iff atLeastatMost_subset_iff atLeastLessThan_iff
|
||||
Int_atLeastAtMost atLeastatMost_empty_iff split_paired_Ex
|
||||
shows
|
||||
|
@ -872,16 +841,6 @@ lemma init_arch_objects_corres_noop:
|
|||
apply (wp mapM_wp' dmo_dwp | simp)+
|
||||
done
|
||||
|
||||
|
||||
(* FIXME: move to Detype *)
|
||||
lemma monad_commute_pre:
|
||||
"(a >>= return ) = ( b >>= return ) \<Longrightarrow> a = b"
|
||||
apply (clarsimp simp:bind_def)
|
||||
apply (rule ext)
|
||||
apply (drule_tac x= x in fun_cong)
|
||||
apply (auto simp:return_def split_def)
|
||||
done
|
||||
|
||||
lemma monad_commute_set_cap_cdt:
|
||||
"monad_commute \<top> (KHeap_D.set_cap ptr cap) (modify (\<lambda>s. s\<lparr>cdl_cdt := cdl_cdt s(ptr2 \<mapsto> ptr3)\<rparr>))"
|
||||
apply (clarsimp simp:monad_commute_def)
|
||||
|
@ -934,7 +893,7 @@ lemma monad_commute_set_cap_gets_cdt:
|
|||
lemma set_cap_set_parent_swap:
|
||||
"do _ \<leftarrow> KHeap_D.set_cap ptr cap; set_parent ptr2 ptr3 od
|
||||
= do _ \<leftarrow> set_parent ptr2 ptr3; KHeap_D.set_cap ptr cap od"
|
||||
apply (rule monad_commute_pre)
|
||||
apply (rule bind_return_eq)
|
||||
apply (subst bind_assoc)+
|
||||
apply (rule ext)
|
||||
apply (subst monad_commute_simple)
|
||||
|
@ -1041,20 +1000,6 @@ lemma neg_mask_add_2p_helper:
|
|||
apply (simp add:word_power_less_1)
|
||||
done
|
||||
|
||||
|
||||
(* FIXME: move *)
|
||||
lemma nonempty_pick_in:
|
||||
"a\<noteq>{} \<Longrightarrow> pick a \<in> a"
|
||||
by (metis all_not_in_conv someI_ex)
|
||||
|
||||
lemma pick_singleton[simp]:
|
||||
"pick {a} = a"
|
||||
apply (rule ccontr)
|
||||
apply (cut_tac nonempty_pick_in)
|
||||
apply fastforce
|
||||
apply fastforce
|
||||
done
|
||||
|
||||
lemma retype_transform_ref_subseteq_strong:
|
||||
notes blah[simp del] = atLeastAtMost_iff atLeastatMost_subset_iff atLeastLessThan_iff
|
||||
Int_atLeastAtMost atLeastatMost_empty_iff split_paired_Ex
|
||||
|
@ -1201,7 +1146,7 @@ lemma generate_object_ids_exec:
|
|||
apply (erule range_cover.sz)
|
||||
apply (simp add:obj_bits_api_def)
|
||||
apply (erule(2) subset_trans[OF retype_transform_ref_subseteq_strong])
|
||||
apply (rule retype_transform_obj_ref_avaiable)
|
||||
apply (rule retype_transform_obj_ref_available)
|
||||
apply simp+
|
||||
apply (clarsimp simp:retype_transform_obj_ref_def translate_object_type_def)
|
||||
apply simp
|
||||
|
@ -1257,7 +1202,7 @@ lemma update_available_range_dcorres:
|
|||
apply (rule corres_guard_imp)
|
||||
apply (rule select_pick_corres)
|
||||
apply (rule set_cap_corres)
|
||||
apply (clarsimp simp:set_avaiable_range_def)
|
||||
apply clarsimp
|
||||
apply simp+
|
||||
done
|
||||
|
||||
|
|
|
@ -976,6 +976,42 @@ lemma tl_subseteq:
|
|||
"set (tl xs) \<subseteq> set xs"
|
||||
by (induct xs, auto)
|
||||
|
||||
|
||||
crunch states_equiv_for: invalidate_tlb_by_asid "states_equiv_for P Q R S X st"
|
||||
(wp: do_machine_op_mol_states_equiv_for ignore: do_machine_op simp: invalidateTLB_ASID_def)
|
||||
|
||||
|
||||
crunch cur_thread[wp]: invalidate_tlb_by_asid "\<lambda>s. P (cur_thread s)"
|
||||
crunch cur_domain[wp]: invalidate_tlb_by_asid "\<lambda>s. P (cur_domain s)"
|
||||
crunch sched_act[wp]: invalidate_tlb_by_asid "\<lambda>s. P (scheduler_action s)"
|
||||
crunch wuc[wp]: invalidate_tlb_by_asid "\<lambda>s. P (work_units_completed s)"
|
||||
|
||||
lemma invalidate_tlb_by_asid_reads_respects:
|
||||
"reads_respects aag l (\<lambda>_. True) (invalidate_tlb_by_asid asid)"
|
||||
apply(rule reads_respects_unobservable_unit_return)
|
||||
apply (rule invalidate_tlb_by_asid_states_equiv_for)
|
||||
apply wp
|
||||
done
|
||||
|
||||
|
||||
|
||||
lemma get_master_pte_reads_respects:
|
||||
"reads_respects aag l (K (is_subject aag (p && ~~ mask pt_bits))) (get_master_pte p)"
|
||||
unfolding get_master_pte_def
|
||||
apply(wp get_pte_reads_respects | wpc | simp
|
||||
| wp_once hoare_drop_imps)+
|
||||
apply(fastforce simp: pt_bits_def pageBits_def mask_lower_twice)
|
||||
done
|
||||
|
||||
|
||||
lemma get_master_pde_reads_respects:
|
||||
"reads_respects aag l (K (is_subject aag (x && ~~ mask pd_bits))) (get_master_pde x)"
|
||||
unfolding get_master_pde_def
|
||||
apply(wp get_pde_rev | wpc | simp
|
||||
| wp_once hoare_drop_imps)+
|
||||
apply(fastforce simp: pd_bits_def pageBits_def mask_lower_twice)
|
||||
done
|
||||
|
||||
lemma perform_page_invocation_reads_respects:
|
||||
"reads_respects aag l (pas_refined aag and K (authorised_page_inv aag pi) and valid_page_inv pi and valid_arch_objs and pspace_aligned and is_subject aag \<circ> cur_thread) (perform_page_invocation pi)"
|
||||
unfolding perform_page_invocation_def fun_app_def when_def cleanCacheRange_PoU_def
|
||||
|
@ -987,8 +1023,9 @@ lemma perform_page_invocation_reads_respects:
|
|||
set_cap_reads_respects mapM_ev'' store_pde_reads_respects
|
||||
unmap_page_reads_respects set_vm_root_reads_respects
|
||||
dmo_mol_2_reads_respects set_vm_root_for_flush_reads_respects get_cap_rev
|
||||
do_flush_reads_respects
|
||||
| simp add: cleanByVA_PoU_def | wpc | wp_once hoare_drop_imps[where R="\<lambda> r s. r"])+
|
||||
do_flush_reads_respects invalidate_tlb_by_asid_reads_respects
|
||||
get_master_pte_reads_respects get_master_pde_reads_respects
|
||||
| simp add: cleanByVA_PoU_def pte_check_if_mapped_def pde_check_if_mapped_def | wpc | wp_once hoare_drop_imps[where R="\<lambda> r s. r"])+
|
||||
apply(clarsimp simp: authorised_page_inv_def valid_page_inv_def)
|
||||
apply (auto simp: cte_wp_at_caps_of_state is_arch_diminished_def valid_slots_def
|
||||
cap_auth_conferred_def cap_rights_update_def acap_rights_update_def
|
||||
|
@ -1853,8 +1890,8 @@ done
|
|||
|
||||
definition authorised_for_globals_page_inv :: "page_invocation \<Rightarrow> 'z::state_ext state \<Rightarrow> bool"
|
||||
where "authorised_for_globals_page_inv pi \<equiv>
|
||||
\<lambda>s. case pi of PageMap cap ptr m \<Rightarrow>
|
||||
\<exists>slot. cte_wp_at (parent_for_refs m) slot s | PageRemap m \<Rightarrow>
|
||||
\<lambda>s. case pi of PageMap asid cap ptr m \<Rightarrow>
|
||||
\<exists>slot. cte_wp_at (parent_for_refs m) slot s | PageRemap asid m \<Rightarrow>
|
||||
\<exists>slot. cte_wp_at (parent_for_refs m) slot s | _ \<Rightarrow> True"
|
||||
|
||||
lemma set_cap_valid_ko_at_arm[wp]:
|
||||
|
@ -2360,19 +2397,5 @@ lemma mapM_x_swp_store_pde_pas_refined_simple:
|
|||
apply (wp store_pde_pas_refined_simple)
|
||||
done
|
||||
|
||||
crunch states_equiv_for: invalidate_tlb_by_asid "states_equiv_for P Q R S X st"
|
||||
(wp: do_machine_op_mol_states_equiv_for ignore: do_machine_op simp: invalidateTLB_ASID_def)
|
||||
|
||||
crunch cur_thread[wp]: invalidate_tlb_by_asid "\<lambda>s. P (cur_thread s)"
|
||||
crunch cur_domain[wp]: invalidate_tlb_by_asid "\<lambda>s. P (cur_domain s)"
|
||||
crunch sched_act[wp]: invalidate_tlb_by_asid "\<lambda>s. P (scheduler_action s)"
|
||||
crunch wuc[wp]: invalidate_tlb_by_asid "\<lambda>s. P (work_units_completed s)"
|
||||
|
||||
lemma invalidate_tlb_by_asid_reads_respects:
|
||||
"reads_respects aag l (\<lambda>_. True) (invalidate_tlb_by_asid asid)"
|
||||
apply(rule reads_respects_unobservable_unit_return)
|
||||
apply (rule invalidate_tlb_by_asid_states_equiv_for)
|
||||
apply wp
|
||||
done
|
||||
|
||||
end
|
||||
|
|
|
@ -288,22 +288,6 @@ lemma decode_irq_handler_invocation_rev:
|
|||
done
|
||||
|
||||
|
||||
lemma get_master_pte_reads_respects:
|
||||
"reads_respects aag l (K (is_subject aag (p && ~~ mask pt_bits))) (get_master_pte p)"
|
||||
unfolding get_master_pte_def
|
||||
apply(wp get_pte_reads_respects | wpc | simp
|
||||
| wp_once hoare_drop_imps)+
|
||||
apply(fastforce simp: pt_bits_def pageBits_def mask_lower_twice)
|
||||
done
|
||||
|
||||
|
||||
lemma get_master_pde_reads_respects:
|
||||
"reads_respects aag l (K (is_subject aag (x && ~~ mask pd_bits))) (get_master_pde x)"
|
||||
unfolding get_master_pde_def
|
||||
apply(wp get_pde_rev | wpc | simp
|
||||
| wp_once hoare_drop_imps)+
|
||||
apply(fastforce simp: pd_bits_def pageBits_def mask_lower_twice)
|
||||
done
|
||||
|
||||
|
||||
(* FIXME: cleanup this proof *)
|
||||
|
|
|
@ -2755,6 +2755,18 @@ lemma s0H_valid_pspace':
|
|||
apply ((clarsimp split: split_if_asm)+)[3]
|
||||
done
|
||||
|
||||
(* Instantiate the current, abstract domain scheduler into the
|
||||
concrete scheduler required for this example *)
|
||||
axiomatization newKSDomSchedInst where
|
||||
newKSDomSched: "newKSDomSchedule = [(0,0xA), (1, 0xA)]"
|
||||
|
||||
(* kernel_data_refs is an undefined constant at the moment, and therefore
|
||||
cannot be referred to in valid_global_refs' and pspace_domain_valid.
|
||||
We use an axiomatization for the moment. *)
|
||||
axiomatization kernel_data_refs_valid where
|
||||
kdr_valid_global_refs': "valid_global_refs' s0H_internal" and
|
||||
kdr_pspace_domain_valid: "pspace_domain_valid s0H_internal"
|
||||
|
||||
lemma s0H_invs:
|
||||
"invs' s0H_internal"
|
||||
apply (clarsimp simp: invs'_def valid_state'_def s0H_valid_pspace')
|
||||
|
@ -2768,22 +2780,22 @@ lemma s0H_invs:
|
|||
apply (clarsimp simp: sym_refs_def state_refs_of'_def refs_of'_def split: option.splits)
|
||||
apply (frule kh0H_SomeD)
|
||||
apply (elim disjE, simp_all)[1]
|
||||
apply (clarsimp simp: tcb_st_refs_of'_def idle_tcbH_def)
|
||||
apply (clarsimp simp: tcb_st_refs_of'_def High_tcbH_def)
|
||||
apply (rule conjI)
|
||||
apply (clarsimp simp: aepH_def)
|
||||
apply (clarsimp simp: objBitsKO_def aepH_def)
|
||||
apply (erule impE, simp add: is_aligned_def s0_ptr_defs)
|
||||
apply (erule notE, rule pspace_distinctD''[OF _ s0H_pspace_distinct'])
|
||||
apply (simp add: objBitsKO_def aepH_def)
|
||||
apply (clarsimp simp: tcb_st_refs_of'_def Low_tcbH_def)
|
||||
apply (clarsimp simp: aepH_def aep_q_refs_of'_def)
|
||||
apply (rule conjI)
|
||||
apply (clarsimp simp: tcb_st_refs_of'_def idle_tcbH_def)
|
||||
apply (clarsimp simp: tcb_st_refs_of'_def High_tcbH_def)
|
||||
apply (clarsimp simp: objBitsKO_def s0_ptrs_aligned)
|
||||
apply (erule notE, rule pspace_distinctD''[OF _ s0H_pspace_distinct'])
|
||||
apply (simp add: objBitsKO_def)
|
||||
apply (clarsimp simp: irq_cte_def)
|
||||
apply (rule conjI)
|
||||
apply (clarsimp simp: aepH_def)
|
||||
apply (clarsimp simp: objBitsKO_def aepH_def)
|
||||
apply (erule impE, simp add: is_aligned_def s0_ptr_defs)
|
||||
apply (erule notE, rule pspace_distinctD''[OF _ s0H_pspace_distinct'])
|
||||
apply (simp add: objBitsKO_def aepH_def)
|
||||
apply (clarsimp simp: tcb_st_refs_of'_def Low_tcbH_def)
|
||||
apply (clarsimp simp: aepH_def aep_q_refs_of'_def)
|
||||
apply (rule conjI)
|
||||
apply (clarsimp simp: tcb_st_refs_of'_def High_tcbH_def)
|
||||
apply (clarsimp simp: objBitsKO_def s0_ptrs_aligned)
|
||||
apply (erule notE, rule pspace_distinctD''[OF _ s0H_pspace_distinct'])
|
||||
apply (simp add: objBitsKO_def)
|
||||
apply (clarsimp simp: irq_cte_def)
|
||||
apply (clarsimp simp: Low_cte_def Low_cte'_def split: split_if_asm)
|
||||
apply (clarsimp simp: High_cte_def High_cte'_def split: split_if_asm)
|
||||
apply (clarsimp simp: Silc_cte_def Silc_cte'_def split: split_if_asm)
|
||||
|
@ -2874,7 +2886,7 @@ lemma s0H_invs:
|
|||
apply (rule pspace_distinctD''[OF _ s0H_pspace_distinct', simplified s0H_internal_def])
|
||||
apply (simp add: objBitsKO_def)
|
||||
apply (rule conjI)
|
||||
defer
|
||||
apply (clarsimp simp: kdr_valid_global_refs') (* use axiomatization for now *)
|
||||
apply (rule conjI)
|
||||
apply (clarsimp simp: valid_arch_state'_def)
|
||||
apply (intro conjI)
|
||||
|
@ -2986,7 +2998,7 @@ lemma s0H_invs:
|
|||
apply (clarsimp split: split_if_asm)
|
||||
apply (clarsimp split: split_if_asm)
|
||||
apply (rule conjI)
|
||||
defer
|
||||
apply (clarsimp simp: kdr_pspace_domain_valid) (* use axiomatization for now *)
|
||||
apply (rule conjI)
|
||||
apply (clarsimp simp: s0H_internal_def)
|
||||
apply (rule conjI)
|
||||
|
@ -2994,16 +3006,13 @@ lemma s0H_invs:
|
|||
apply (rule conjI)
|
||||
apply (clarsimp simp: s0H_internal_def maxDomain_def numDomains_def dschDomain_def dschLength_def)
|
||||
apply (rule conjI)
|
||||
defer
|
||||
apply (clarsimp simp: s0H_internal_def newKernelState_def newKSDomSched)
|
||||
apply (rule conjI)
|
||||
apply (clarsimp simp: s0H_internal_def newKernelState_def)
|
||||
apply (clarsimp simp: s0H_internal_def newKernelState_def newKSDomSched)
|
||||
apply (clarsimp simp: cur_tcb'_def obj_at'_def projectKO_eq project_inject s0H_internal_def objBitsKO_def s0_ptrs_aligned)
|
||||
apply (rule pspace_distinctD''[OF _ s0H_pspace_distinct', simplified s0H_internal_def])
|
||||
apply (simp add: objBitsKO_def kh0H_simps[simplified cte_level_bits_def])
|
||||
(* kernel_data_refs is an undefined const, can't talk about this in
|
||||
valid_global_refs' and pspace_domain_valid.
|
||||
Invariant ksDomSchedule s0H = ksDomSchedule (newKernelState) is False *)
|
||||
sorry
|
||||
done
|
||||
|
||||
lemma kh0_pspace_dom:
|
||||
"pspace_dom kh0 = {init_globals_frame, idle_tcb_ptr, High_tcb_ptr, Low_tcb_ptr,
|
||||
|
|
|
@ -11,6 +11,7 @@
|
|||
theory IRQMasks_IF
|
||||
imports "../access-control/DomainSepInv"
|
||||
begin
|
||||
|
||||
abbreviation irq_masks_of_state :: "det_ext state \<Rightarrow> irq \<Rightarrow> bool" where
|
||||
"irq_masks_of_state s \<equiv> irq_masks (machine_state s)"
|
||||
|
||||
|
@ -181,7 +182,7 @@ lemma no_irq_do_flush:
|
|||
done
|
||||
|
||||
crunch irq_masks[wp]: arch_perform_invocation "\<lambda>s. P (irq_masks_of_state s)"
|
||||
(wp: dmo_wp crunch_wps simp: crunch_simps no_irq_cleanByVA_PoU no_irq_do_flush)
|
||||
(wp: dmo_wp crunch_wps simp: crunch_simps no_irq_cleanByVA_PoU no_irq_invalidateTLB_ASID no_irq_do_flush)
|
||||
|
||||
crunch irq_masks[wp]: restart "\<lambda>s. P (irq_masks_of_state s)"
|
||||
|
||||
|
|
|
@ -2362,6 +2362,19 @@ lemma set_asid_pool_invs_restrict:
|
|||
done
|
||||
|
||||
|
||||
lemmas set_asid_pool_cte_wp_at1[wp]
|
||||
= hoare_cte_wp_caps_of_state_lift [OF set_asid_pool_caps_of_state]
|
||||
|
||||
|
||||
lemma mdb_cte_at_set_asid_pool[wp]:
|
||||
"\<lbrace>\<lambda>s. mdb_cte_at (swp (cte_wp_at (op \<noteq> cap.NullCap)) s) (cdt s)\<rbrace>
|
||||
set_asid_pool y pool
|
||||
\<lbrace>\<lambda>r s. mdb_cte_at (swp (cte_wp_at (op \<noteq> cap.NullCap)) s) (cdt s)\<rbrace>"
|
||||
apply (clarsimp simp:mdb_cte_at_def)
|
||||
apply (simp only: imp_conv_disj)
|
||||
apply (wp hoare_vcg_disj_lift hoare_vcg_all_lift)
|
||||
done
|
||||
|
||||
lemma set_asid_pool_invs_unmap:
|
||||
"\<lbrace>invs and ko_at (ArchObj (ARM_Structs_A.ASIDPool ap)) p and
|
||||
(\<lambda>s. \<forall>asid. asid \<le> mask asid_bits \<longrightarrow> ucast asid = x \<longrightarrow>
|
||||
|
|
|
@ -1428,6 +1428,13 @@ lemma XNever_attribs_from_word[simp]:
|
|||
"XNever \<notin> attribs_from_word w"
|
||||
by (simp add: attribs_from_word_def)
|
||||
|
||||
lemma cte_wp_at_page_cap_weaken:
|
||||
"cte_wp_at (diminished (ArchObjectCap (PageCap word seta vmpage_size None))) slot s \<Longrightarrow>
|
||||
cte_wp_at (\<lambda>a. \<exists>p R sz m. a = ArchObjectCap (PageCap p R sz m)) slot s"
|
||||
apply (clarsimp simp: cte_wp_at_def diminished_def mask_cap_def cap_rights_update_def)
|
||||
apply (clarsimp simp: acap_rights_update_def split: cap.splits arch_cap.splits)
|
||||
done
|
||||
|
||||
lemma arch_decode_inv_wf[wp]:
|
||||
"\<lbrace>invs and valid_cap (cap.ArchObjectCap arch_cap) and
|
||||
cte_wp_at (diminished (cap.ArchObjectCap arch_cap)) slot and
|
||||
|
@ -1523,21 +1530,44 @@ lemma arch_decode_inv_wf[wp]:
|
|||
apply (drule (1) caps_of_state_valid[rotated])+
|
||||
apply (drule (1) diminished_is_update)+
|
||||
apply (clarsimp simp: cap_rights_update_def)
|
||||
apply (clarsimp simp:diminished_def)
|
||||
apply (clarsimp simp:diminished_def)
|
||||
apply (simp add: arch_decode_invocation_def Let_def split_def
|
||||
cong: if_cong split del: split_if)
|
||||
apply (cases "invocation_type label = ARMPageMap")
|
||||
apply (simp split del: split_if)
|
||||
apply (rule hoare_pre)
|
||||
apply ((wp whenE_throwError_wp check_vp_wpR hoare_vcg_const_imp_lift_R
|
||||
create_mapping_entries_parent_for_refs find_pd_for_asid_inv
|
||||
create_mapping_entries_parent_for_refs find_pd_for_asid_pd_at_asid
|
||||
create_mapping_entries_valid_slots create_mapping_entries_same_refs_ex
|
||||
| wpc
|
||||
| simp add: valid_arch_inv_def valid_page_inv_def is_pg_cap_def
|
||||
| rule_tac x="fst p" in hoare_imp_eq_substR)+)[1]
|
||||
apply (rule hoare_vcg_conj_liftE_R)
|
||||
apply (simp add: find_pd_for_asid_def)
|
||||
apply (wp assertE_wp whenE_throwError_wp | wpc)+
|
||||
apply (rule hoare_vcg_conj_liftE_R)
|
||||
apply (simp add: find_pd_for_asid_def)
|
||||
apply ((wp_trace assertE_wp whenE_throwError_wp find_pd_for_asid_pd_at_asid | wpc)+)[1]
|
||||
apply (rule_tac Q'="\<lambda>rv s. page_directory_at rv s \<and>
|
||||
args ! 0 < kernel_base \<and>
|
||||
valid_arch_state s \<and>
|
||||
valid_arch_objs s \<and>
|
||||
equal_kernel_mappings s \<and>
|
||||
pspace_aligned s \<and>
|
||||
valid_global_objs s \<and>
|
||||
(\<exists>\<rhd> rv) s \<and>
|
||||
page_directory_at rv s \<and>
|
||||
typ_at (AArch (AIntData vmpage_size)) word s \<and>
|
||||
is_aligned (addrFromPPtr word) pageBits \<and>
|
||||
args ! 0 < kernel_base \<and>
|
||||
s \<turnstile> ArchObjectCap (PageCap word set vmpage_size (Some (snd p, args ! 0))) \<and>
|
||||
snd p \<le> mask asid_bits \<and>
|
||||
snd p \<noteq> 0 \<and>
|
||||
invs s \<and>
|
||||
(\<exists>\<rhd> rv) s \<and>
|
||||
page_directory_at rv s \<and>
|
||||
is_aligned rv pd_bits \<and>
|
||||
args ! 0 < kernel_base \<and> (pd_at_asid (snd p) rv s)" in hoare_post_imp_R)
|
||||
prefer 2
|
||||
apply auto[1]
|
||||
apply (wp assertE_wp whenE_throwError_wp find_pd_for_asid_pd_at_asid | wpc)+
|
||||
apply (clarsimp simp: neq_Nil_conv)
|
||||
apply (rule conjI)
|
||||
apply (clarsimp simp add: cte_wp_at_caps_of_state)
|
||||
|
@ -1547,7 +1577,7 @@ lemma arch_decode_inv_wf[wp]:
|
|||
apply (simp add: cap_master_cap_def vs_cap_ref_def
|
||||
cap_rights_update_def acap_rights_update_def
|
||||
split:option.split vmpage_size.split)
|
||||
apply (rule conjI, fastforce)+
|
||||
apply (rule conjI, fastforce dest!: cte_wp_at_page_cap_weaken)+
|
||||
apply (frule diminished_cte_wp_at_valid_cap[where p="(a, b)", standard],
|
||||
clarsimp+)
|
||||
apply (clarsimp simp: valid_cap_def
|
||||
|
@ -1571,6 +1601,7 @@ lemma arch_decode_inv_wf[wp]:
|
|||
apply (fastforce simp: diminished_pd_self)
|
||||
apply (simp add: vmsz_aligned_def split: vmpage_size.splits)
|
||||
apply (simp add: conj_ac)
|
||||
apply (rule conjI, simp, rule conjI, simp add: mask_def)
|
||||
apply (rule conjI)
|
||||
apply (rule is_aligned_addrFromPPtr)
|
||||
apply (erule is_aligned_weaken)
|
||||
|
@ -1593,20 +1624,66 @@ lemma arch_decode_inv_wf[wp]:
|
|||
| simp add: valid_arch_inv_def valid_page_inv_def
|
||||
| (simp add: cte_wp_at_caps_of_state,
|
||||
wp create_mapping_entries_same_refs_ex hoare_vcg_ex_lift_R))+)[1]
|
||||
apply (rule_tac Q'="\<lambda>rv s. fst p = rv \<and> fst pa = snd p \<longrightarrow>
|
||||
pd_at_asid (fst pa) rv s \<and>
|
||||
valid_arch_state s \<and>
|
||||
valid_arch_objs s \<and>
|
||||
equal_kernel_mappings s \<and>
|
||||
pspace_aligned s \<and>
|
||||
valid_global_objs s \<and>
|
||||
(\<exists>\<rhd> fst p) s \<and>
|
||||
page_directory_at (fst p) s \<and>
|
||||
typ_at (AArch (AIntData vmpage_size)) word s \<and>
|
||||
is_aligned (addrFromPPtr word) pageBits \<and>
|
||||
vmsz_aligned (snd pa) vmpage_size \<and>
|
||||
snd pa < kernel_base \<and>
|
||||
snd p \<le> mask asid_bits \<and>
|
||||
snd p \<noteq> 0 \<and>
|
||||
invs s \<and>
|
||||
(\<exists>\<rhd> fst p) s \<and>
|
||||
page_directory_at (fst p) s \<and>
|
||||
is_aligned (fst p) pd_bits \<and>
|
||||
vmsz_aligned (snd pa) vmpage_size \<and>
|
||||
snd pa < kernel_base \<and>
|
||||
(\<exists>v va vb.
|
||||
caps_of_state s (v, va) = Some vb \<and>
|
||||
valid_arch_state s \<and>
|
||||
valid_arch_objs s \<and>
|
||||
valid_vs_lookup s \<and>
|
||||
unique_table_refs (caps_of_state s) \<and>
|
||||
pspace_aligned s \<and>
|
||||
valid_objs s \<and>
|
||||
valid_kernel_mappings s \<and>
|
||||
(\<exists>\<rhd> fst p) s \<and>
|
||||
(\<exists>pd_cap.
|
||||
(\<exists>a b. cte_wp_at (diminished pd_cap) (a, b) s) \<and>
|
||||
(\<exists>asid. pd_cap = ArchObjectCap (PageDirectoryCap (fst p) (Some asid)) \<and>
|
||||
page_directory_at (fst p) s \<and>
|
||||
snd pa < kernel_base \<and>
|
||||
(\<exists>rights'.
|
||||
vb =
|
||||
ArchObjectCap
|
||||
(PageCap word rights' vmpage_size
|
||||
(Some (asid, snd pa)))))))" in hoare_post_imp_R)
|
||||
prefer 2
|
||||
apply clarsimp
|
||||
apply (rule conjI[rotated], auto)[1]
|
||||
apply (rule_tac x=v in exI, rule_tac x=va in exI)
|
||||
apply (rule_tac x="(ArchObjectCap (PageCap word rights' vmpage_size (Some (asid, ba))))" in exI)
|
||||
apply auto[1]
|
||||
apply (rule mp [OF strengthen_validE_R_cong])
|
||||
apply (rule impI)
|
||||
apply (subst eq_commute [where b = "fst p"])
|
||||
apply (subst eq_commute[where b="snd p"])
|
||||
apply assumption
|
||||
apply (simp add: conj_ac)
|
||||
apply ((wp whenE_throwError_wp check_vp_wpR hoare_vcg_const_imp_lift_R
|
||||
hoare_drop_impE_R
|
||||
create_mapping_entries_parent_for_refs |
|
||||
wpc|
|
||||
simp add: valid_arch_inv_def valid_page_inv_def)+)[4]
|
||||
apply ((wp_trace whenE_throwError_wp check_vp_wpR hoare_vcg_const_imp_lift_R
|
||||
hoare_drop_impE_R create_mapping_entries_parent_for_refs
|
||||
| wpc
|
||||
| simp add: valid_arch_inv_def valid_page_inv_def)+)[4]
|
||||
apply (simp add: find_pd_for_asid_def)
|
||||
apply (wp assertE_wp whenE_throwError_wp | wpc)+
|
||||
apply (clarsimp simp: valid_cap_def cap_aligned_def)
|
||||
apply (auto simp: vmsz_aligned_def invs_def valid_state_def
|
||||
apply (auto simp: vmsz_aligned_def invs_def valid_state_def mask_def
|
||||
valid_arch_caps_def cte_wp_at_caps_of_state
|
||||
elim: is_aligned_weaken
|
||||
intro!: is_aligned_addrFromPPtr pbfs_atleast_pageBits)[1]
|
||||
|
@ -1647,8 +1724,8 @@ lemma arch_decode_inv_wf[wp]:
|
|||
apply (simp split del: split_if)
|
||||
apply (rule hoare_pre)
|
||||
apply (wp whenE_throwError_wp static_imp_wp hoare_drop_imps)
|
||||
apply (simp add: valid_arch_inv_def valid_page_inv_def)
|
||||
apply (wp find_pd_for_asid_pd_at_asid | wpc)+
|
||||
apply (simp add: valid_arch_inv_def valid_page_inv_def)
|
||||
apply (wp find_pd_for_asid_pd_at_asid | wpc)+
|
||||
apply (clarsimp simp: valid_cap_def mask_def)
|
||||
apply simp
|
||||
apply (rule hoare_pre, wp)
|
||||
|
|
|
@ -481,16 +481,6 @@ lemma create_mapping_entries_bcorres[wp]: "bcorres (create_mapping_entries a b c
|
|||
apply (wp | simp)+
|
||||
done
|
||||
|
||||
lemma get_master_pte_bcorres[wp]: "bcorres (get_master_pte x) (get_master_pte x)"
|
||||
apply (simp add:get_master_pte_def)
|
||||
apply (wp | wpc | simp)+
|
||||
done
|
||||
|
||||
lemma get_master_pde_bcorres[wp]: "bcorres (get_master_pde x) (get_master_pde x)"
|
||||
apply (simp add:get_master_pde_def)
|
||||
apply (wp | wpc | simp)+
|
||||
done
|
||||
|
||||
lemma ensure_safe_mapping_bcorres[wp]: "bcorres (ensure_safe_mapping a) (ensure_safe_mapping a)"
|
||||
apply (induct rule: ensure_safe_mapping.induct)
|
||||
apply (wp | wpc | simp)+
|
||||
|
|
|
@ -2510,11 +2510,6 @@ lemma final_zombie_not_live:
|
|||
done
|
||||
|
||||
|
||||
lemma invs_ifunsafe[elim!]:
|
||||
"invs s \<Longrightarrow> if_unsafe_then_cap s"
|
||||
by (simp add: invs_def valid_state_def valid_pspace_def)
|
||||
|
||||
|
||||
lemma suspend_ex_cte_cap[wp]:
|
||||
"\<lbrace>ex_cte_cap_wp_to P p\<rbrace> IpcCancel_A.suspend t \<lbrace>\<lambda>rv. ex_cte_cap_wp_to P p\<rbrace>"
|
||||
apply (simp add: ex_cte_cap_wp_to_def cte_wp_at_caps_of_state
|
||||
|
|
|
@ -107,13 +107,17 @@ lemma valid_cap_machine_state [iff]:
|
|||
"machine_state_update f s \<turnstile> c = s \<turnstile> c"
|
||||
by (fastforce intro: valid_cap_pspaceI)
|
||||
|
||||
|
||||
lemma get_cap_valid [wp]:
|
||||
"\<lbrace> valid_objs \<rbrace> get_cap addr \<lbrace> valid_cap \<rbrace>"
|
||||
apply (wp get_cap_wp)
|
||||
apply (auto dest: cte_wp_at_valid_objs_valid_cap)
|
||||
done
|
||||
|
||||
lemma get_cap_wellformed:
|
||||
"\<lbrace>valid_objs\<rbrace> get_cap slot \<lbrace>\<lambda>cap s. wellformed_cap cap\<rbrace>"
|
||||
apply (rule hoare_strengthen_post, rule get_cap_valid)
|
||||
apply (simp add: valid_cap_def2)
|
||||
done
|
||||
|
||||
lemma update_cdt_cdt:
|
||||
"\<lbrace>\<lambda>s. valid_mdb (cdt_update (\<lambda>_. (m (cdt s))) s)\<rbrace> update_cdt m \<lbrace>\<lambda>_. valid_mdb\<rbrace>"
|
||||
|
|
|
@ -12,6 +12,14 @@ theory DetSchedInvs_AI
|
|||
imports Deterministic_AI
|
||||
begin
|
||||
|
||||
lemma get_etcb_rev:
|
||||
"ekheap s p = Some etcb \<Longrightarrow> get_etcb p s = Some etcb"
|
||||
by (clarsimp simp: get_etcb_def)
|
||||
|
||||
lemma get_etcb_SomeD: "get_etcb ptr s = Some v \<Longrightarrow> ekheap s ptr = Some v"
|
||||
apply (case_tac "ekheap s ptr", simp_all add: get_etcb_def)
|
||||
done
|
||||
|
||||
definition obj_at_kh where
|
||||
"obj_at_kh P ref kh \<equiv> obj_at P ref ((undefined :: det_ext state)\<lparr>kheap := kh\<rparr>)"
|
||||
|
||||
|
@ -393,4 +401,40 @@ lemma valid_sched_lift:
|
|||
apply (wp valid_etcbs_lift valid_queues_lift ct_not_in_q_lift ct_in_cur_domain_lift valid_sched_action_lift valid_blocked_lift a b c d e f g h i)
|
||||
done
|
||||
|
||||
lemma valid_etcbs_tcb_etcb:
|
||||
"\<lbrakk> valid_etcbs s; kheap s ptr = Some (TCB tcb) \<rbrakk> \<Longrightarrow> \<exists>etcb. ekheap s ptr = Some etcb"
|
||||
by (force simp: valid_etcbs_def is_etcb_at_def st_tcb_at_def obj_at_def)
|
||||
|
||||
lemma valid_etcbs_get_tcb_get_etcb:
|
||||
"\<lbrakk> valid_etcbs s; get_tcb ptr s = Some tcb \<rbrakk> \<Longrightarrow> \<exists>etcb. get_etcb ptr s = Some etcb"
|
||||
apply (clarsimp simp: valid_etcbs_def valid_etcbs_def st_tcb_at_def obj_at_def is_etcb_at_def get_etcb_def get_tcb_def split: option.splits split_if)
|
||||
apply (erule_tac x=ptr in allE)
|
||||
apply (case_tac a)
|
||||
apply (clarsimp simp: get_etcb_def split: option.splits kernel_object.splits)+
|
||||
done
|
||||
|
||||
lemma valid_etcbs_ko_etcb:
|
||||
"\<lbrakk> valid_etcbs s; kheap s ptr = Some ko \<rbrakk> \<Longrightarrow> \<exists>tcb. (ko = TCB tcb = (\<exists>etcb. ekheap s ptr = Some etcb))"
|
||||
apply (clarsimp simp: valid_etcbs_def valid_etcbs_def st_tcb_at_def obj_at_def is_etcb_at_def)
|
||||
apply (erule_tac x="ptr" in allE)
|
||||
apply auto
|
||||
done
|
||||
|
||||
lemma ekheap_tcb_at:
|
||||
"\<lbrakk>ekheap s x = Some y; valid_etcbs s\<rbrakk> \<Longrightarrow> tcb_at x s"
|
||||
by (fastforce simp: valid_etcbs_def is_etcb_at_def st_tcb_at_def obj_at_def is_tcb_def)
|
||||
|
||||
lemma tcb_at_is_etcb_at:
|
||||
"\<lbrakk>tcb_at t s; valid_etcbs s\<rbrakk> \<Longrightarrow> is_etcb_at t s"
|
||||
by (simp add: valid_etcbs_def tcb_at_st_tcb_at)
|
||||
|
||||
lemma tcb_at_ekheap_dom:
|
||||
"\<lbrakk>tcb_at x s; valid_etcbs s\<rbrakk> \<Longrightarrow> (\<exists>etcb. ekheap s x = Some etcb)"
|
||||
by (auto simp: is_etcb_at_def dest: tcb_at_is_etcb_at)
|
||||
|
||||
lemma ekheap_kheap_dom:
|
||||
"\<lbrakk>ekheap s x = Some etcb; valid_etcbs s\<rbrakk>
|
||||
\<Longrightarrow> \<exists>tcb. kheap s x = Some (TCB tcb)"
|
||||
by (fastforce simp: valid_etcbs_def st_tcb_at_def obj_at_def is_etcb_at_def)
|
||||
|
||||
end
|
||||
|
|
|
@ -3163,7 +3163,7 @@ crunch aligned[wp]: invalidate_tlb_by_asid "pspace_aligned"
|
|||
|
||||
crunch valid_arch_state[wp]: invalidate_tlb_by_asid "valid_arch_state"
|
||||
|
||||
|
||||
(*FIXME: move *)
|
||||
lemma corres_option_split:
|
||||
"\<lbrakk>v = v'; corres_underlying sr nf r P P' a c; (\<And>x. v = Some x \<Longrightarrow> corres_underlying sr nf r (Q x) (Q' x) (b x) (d x))\<rbrakk>
|
||||
\<Longrightarrow> corres_underlying sr nf r (option_case P Q v) (option_case P' Q' v') (option_case a b v) (option_case c d v')"
|
||||
|
@ -3194,6 +3194,34 @@ lemma zombie_not_ex_cap_to:
|
|||
apply fastforce
|
||||
done
|
||||
|
||||
lemma valid_idle_has_null_cap:
|
||||
"\<lbrakk> if_unsafe_then_cap s; valid_global_refs s;valid_idle s;valid_irq_node s\<rbrakk>
|
||||
\<Longrightarrow> caps_of_state s (idle_thread s, v) = Some cap
|
||||
\<Longrightarrow> cap = cap.NullCap"
|
||||
apply (rule ccontr)
|
||||
apply (drule(1) if_unsafe_then_capD[OF caps_of_state_cteD])
|
||||
apply clarsimp
|
||||
apply (clarsimp simp: ex_cte_cap_wp_to_def cte_wp_at_caps_of_state)
|
||||
apply (frule(1) valid_global_refsD2)
|
||||
apply (case_tac capa, simp_all add: cap_range_def global_refs_def)[1]
|
||||
apply (clarsimp simp: valid_irq_node_def valid_idle_def st_tcb_at_def
|
||||
obj_at_def is_cap_table_def)
|
||||
apply (drule_tac x=word in spec, simp)
|
||||
done
|
||||
|
||||
lemma zombie_cap_two_nonidles:
|
||||
"\<lbrakk> caps_of_state s ptr = Some (cap.Zombie ptr' zbits n); invs s \<rbrakk>
|
||||
\<Longrightarrow> fst ptr \<noteq> idle_thread s \<and> ptr' \<noteq> idle_thread s"
|
||||
apply (frule valid_global_refsD2, clarsimp+)
|
||||
apply (simp add: cap_range_def global_refs_def)
|
||||
apply (cases ptr, auto dest: valid_idle_has_null_cap[rotated -1])[1]
|
||||
done
|
||||
|
||||
lemma is_cap_tableE:
|
||||
"\<lbrakk> is_cap_table sz ko; \<And>cs. \<lbrakk> ko = kernel_object.CNode sz cs; well_formed_cnode_n sz cs\<rbrakk> \<Longrightarrow> P \<rbrakk> \<Longrightarrow> P"
|
||||
unfolding is_cap_table_def
|
||||
by (auto split: Structures_A.kernel_object.split_asm)
|
||||
|
||||
lemma recycle_cap_Null[wp]: "\<lbrace>\<top>\<rbrace> recycle_cap is_final cap \<lbrace>\<lambda>rv s. rv \<noteq> cap.NullCap\<rbrace>"
|
||||
apply (simp add: recycle_cap_def)
|
||||
apply (rule hoare_pre)
|
||||
|
|
|
@ -1900,6 +1900,11 @@ lemma asid_pool_at_ko:
|
|||
apply (case_tac arch_kernel_obj, auto)
|
||||
done
|
||||
|
||||
lemma typ_at_pg:
|
||||
"typ_at (AArch (AIntData sz)) buf s = ko_at (ArchObj (DataPage sz)) buf s"
|
||||
unfolding obj_at_def
|
||||
by (auto simp: a_type_def split: Structures_A.kernel_object.split_asm arch_kernel_obj.split_asm split_if_asm)
|
||||
|
||||
lemma symreftype_inverse[simp]:
|
||||
"symreftype (symreftype t) = t"
|
||||
by (cases t, simp+)
|
||||
|
@ -4949,6 +4954,26 @@ lemma invs_equal_kernel_mappings[elim!]:
|
|||
"invs s \<Longrightarrow> equal_kernel_mappings s"
|
||||
by (simp add:invs_def valid_state_def)
|
||||
|
||||
lemma invs_valid_irq_node[elim!]:
|
||||
"invs s \<Longrightarrow> valid_irq_node s"
|
||||
by (simp add: invs_def valid_state_def)
|
||||
|
||||
lemma invs_ifunsafe[elim!]:
|
||||
"invs s \<Longrightarrow> if_unsafe_then_cap s"
|
||||
by (simp add: invs_def valid_state_def valid_pspace_def)
|
||||
|
||||
lemma cte_wp_at_cap_aligned:
|
||||
"\<lbrakk>cte_wp_at P p s; invs s\<rbrakk> \<Longrightarrow> \<exists>c. P c \<and> cap_aligned c"
|
||||
apply (drule (1) cte_wp_at_valid_objs_valid_cap [OF _ invs_valid_objs])
|
||||
apply (fastforce simp: valid_cap_def)
|
||||
done
|
||||
|
||||
lemma cte_wp_at_cap_aligned':
|
||||
"\<lbrakk>cte_wp_at (op = cap) p s; invs s\<rbrakk> \<Longrightarrow> cap_aligned cap"
|
||||
apply (drule (1) cte_wp_at_valid_objs_valid_cap [OF _ invs_valid_objs])
|
||||
apply (fastforce simp: valid_cap_def)
|
||||
done
|
||||
|
||||
locale invs_locale =
|
||||
fixes ex_inv :: "'z::state_ext state \<Rightarrow> bool"
|
||||
assumes dmo_ex_inv[wp]: "\<And>f. \<lbrace>invs and ex_inv\<rbrace> do_machine_op f \<lbrace>\<lambda>rv::unit. ex_inv\<rbrace>"
|
||||
|
|
|
@ -1261,19 +1261,6 @@ lemma transfer_caps_loop_vms[wp]:
|
|||
\<lbrace>\<lambda>_. valid_machine_state\<rbrace>"
|
||||
by (wp transfer_caps_loop_pres)
|
||||
|
||||
lemma storeWord_valid_irq_states:
|
||||
"\<lbrace>\<lambda>m. valid_irq_states (s\<lparr>machine_state := m\<rparr>)\<rbrace> storeWord x y
|
||||
\<lbrace>\<lambda>a b. valid_irq_states (s\<lparr>machine_state := b\<rparr>)\<rbrace>"
|
||||
apply(simp add: valid_irq_states_def | wp | simp add: no_irq_storeWord)+
|
||||
done
|
||||
|
||||
lemma dmo_storeWord_valid_irq_states[wp]:
|
||||
"\<lbrace>valid_irq_states\<rbrace> do_machine_op (storeWord x y) \<lbrace>\<lambda>_. valid_irq_states\<rbrace>"
|
||||
apply(simp add: do_machine_op_def | wp | wpc)+
|
||||
apply clarsimp
|
||||
apply(erule use_valid[OF _ storeWord_valid_irq_states])
|
||||
by simp
|
||||
|
||||
crunch valid_irq_states[wp]: set_extra_badge "valid_irq_states"
|
||||
(ignore: do_machine_op)
|
||||
|
||||
|
@ -2375,22 +2362,6 @@ lemma is_derived_ReplyCap [simp]:
|
|||
done
|
||||
|
||||
|
||||
lemma not_waiting_reply_slot_no_descendants:
|
||||
"\<lbrakk> st_tcb_at (Not \<circ> awaiting_reply) t s;
|
||||
st_tcb_at (Not \<circ> halted) t s;
|
||||
valid_reply_caps s; valid_objs s; valid_mdb s \<rbrakk>
|
||||
\<Longrightarrow> descendants_of (t, tcb_cnode_index 2) (cdt s) = {}"
|
||||
apply (rule ccontr, erule nonemptyE)
|
||||
apply (clarsimp simp: valid_mdb_def reply_mdb_def reply_masters_mdb_def)
|
||||
apply (frule(1) st_tcb_at_reply_cap_valid[where P="Not \<circ> halted"], clarsimp+)
|
||||
apply (clarsimp simp: cte_wp_at_caps_of_state is_cap_simps)
|
||||
apply (elim allE, drule(1) mp, clarsimp)
|
||||
apply (drule(1) bspec)
|
||||
apply (drule has_reply_cap_cte_wpD[OF caps_of_state_cteD])
|
||||
apply (erule notE[rotated], strengthen reply_cap_doesnt_exist_strg)
|
||||
apply simp
|
||||
done
|
||||
|
||||
|
||||
lemma do_normal_transfer_tcb_caps:
|
||||
assumes imp: "\<And>c. P c \<Longrightarrow> \<not> is_untyped_cap c"
|
||||
|
|
|
@ -24,6 +24,11 @@ lemma get_object_inv [wp]: "\<lbrace>P\<rbrace> get_object t \<lbrace>\<lambda>r
|
|||
by (wp get_object_wp) simp
|
||||
|
||||
|
||||
lemma get_tcb_rev:
|
||||
"kheap s p = Some (TCB t)\<Longrightarrow> get_tcb p s = Some t"
|
||||
by (clarsimp simp:get_tcb_def)
|
||||
|
||||
|
||||
lemma get_tcb_SomeD: "get_tcb t s = Some v \<Longrightarrow> kheap s t = Some (TCB v)"
|
||||
apply (case_tac "kheap s t", simp_all add: get_tcb_def)
|
||||
apply (case_tac a, simp_all)
|
||||
|
@ -826,47 +831,6 @@ lemma get_object_ret:
|
|||
by (wp, clarsimp elim!: obj_atE)+
|
||||
|
||||
|
||||
lemma aligned_less_plus_1:
|
||||
"\<lbrakk> is_aligned x n; n > 0 \<rbrakk> \<Longrightarrow> x < x + 1"
|
||||
apply (rule plus_one_helper2)
|
||||
apply (rule order_refl)
|
||||
apply (clarsimp simp: field_simps)
|
||||
apply (drule arg_cong[where f="\<lambda>x. x - 1"])
|
||||
apply (clarsimp simp: is_aligned_mask)
|
||||
apply (drule word_eqD[where x=0])
|
||||
apply simp
|
||||
done
|
||||
|
||||
|
||||
lemma is_aligned_add_helper:
|
||||
"\<lbrakk> is_aligned p n; d < 2 ^ n \<rbrakk>
|
||||
\<Longrightarrow> (p + d && mask n = d) \<and> (p + d && (~~ mask n) = p)"
|
||||
apply (subst(asm) is_aligned_mask)
|
||||
apply (drule less_mask_eq)
|
||||
apply (rule context_conjI)
|
||||
apply (subst word_plus_and_or_coroll)
|
||||
apply (rule word_eqI)
|
||||
apply (drule_tac x=na in word_eqD)+
|
||||
apply (simp add: word_size)
|
||||
apply blast
|
||||
apply (rule word_eqI)
|
||||
apply (drule_tac x=na in word_eqD)+
|
||||
apply (simp add: word_ops_nth_size word_size)
|
||||
apply blast
|
||||
apply (insert word_plus_and_or_coroll2[where x="p + d" and w="mask n"])
|
||||
apply simp
|
||||
done
|
||||
|
||||
|
||||
lemma is_aligned_sub_helper:
|
||||
"\<lbrakk> is_aligned (p - d) n; d < 2 ^ n \<rbrakk>
|
||||
\<Longrightarrow> (p && mask n = d) \<and> (p && (~~ mask n) = p - d)"
|
||||
by (drule(1) is_aligned_add_helper, simp)
|
||||
|
||||
|
||||
lemmas n1_ge_32[simp] = word_n1_ge[where 'a=32, simplified]
|
||||
|
||||
|
||||
lemma mask_in_range:
|
||||
"is_aligned ptr bits \<Longrightarrow>
|
||||
(ptr' && (~~ mask bits) = ptr) = (ptr' \<in> {ptr .. ptr + 2 ^ bits - 1})"
|
||||
|
|
|
@ -1206,9 +1206,9 @@ definition
|
|||
|
||||
definition
|
||||
"page_inv_duplicates_valid iv \<equiv> case iv of
|
||||
ArchInvocation_A.PageMap cap ct_slot entries \<Rightarrow>
|
||||
ArchInvocation_A.PageMap asid cap ct_slot entries \<Rightarrow>
|
||||
page_inv_entries_safe entries
|
||||
| ArchInvocation_A.page_invocation.PageRemap entries \<Rightarrow>
|
||||
| ArchInvocation_A.page_invocation.PageRemap asid entries \<Rightarrow>
|
||||
page_inv_entries_safe entries
|
||||
| _ \<Rightarrow> \<top>"
|
||||
|
||||
|
@ -1484,6 +1484,8 @@ lemma set_cap_page_inv_entries_safe:
|
|||
Let_def split:if_splits option.splits)
|
||||
done
|
||||
|
||||
crunch valid_pdpt[wp]: pte_check_if_mapped, pde_check_if_mapped "valid_pdpt_objects"
|
||||
|
||||
lemma perform_page_valid_pdpt[wp]:
|
||||
"\<lbrace>valid_pdpt_objs and valid_page_inv pinv and page_inv_duplicates_valid pinv\<rbrace>
|
||||
perform_page_invocation pinv \<lbrace>\<lambda>rv. valid_pdpt_objs\<rbrace>"
|
||||
|
@ -1493,7 +1495,7 @@ lemma perform_page_valid_pdpt[wp]:
|
|||
split: sum.split arch_cap.split option.split,
|
||||
safe intro!: hoare_gen_asm hoare_gen_asm[unfolded K_def],
|
||||
simp_all add: mapM_x_Nil mapM_x_Cons mapM_x_map)
|
||||
apply (wp store_pte_valid_pdpt store_pde_valid_pdpt
|
||||
apply (wp store_pte_valid_pdpt store_pde_valid_pdpt get_master_pte_wp get_master_pde_wp
|
||||
store_pte_non_master_valid_pdpt store_pde_non_master_valid_pdpt
|
||||
mapM_x_wp'[OF store_invalid_pte_valid_pdpt
|
||||
[where pte=ARM_Structs_A.pte.InvalidPTE, simplified]]
|
||||
|
@ -1503,9 +1505,11 @@ lemma perform_page_valid_pdpt[wp]:
|
|||
hoare_vcg_imp_lift[OF set_cap_arch_obj_neg] hoare_vcg_all_lift
|
||||
| clarsimp simp: valid_page_inv_def cte_wp_at_weakenE[OF _ TrueI] obj_at_def
|
||||
pte_range_sz_def pde_range_sz_def swp_def valid_page_inv_def
|
||||
valid_slots_def page_inv_entries_safe_def
|
||||
split: ARM_Structs_A.pte.splits ARM_Structs_A.pde.splits)+
|
||||
done
|
||||
valid_slots_def page_inv_entries_safe_def pte_check_if_mapped_def
|
||||
pde_check_if_mapped_def
|
||||
split: ARM_Structs_A.pte.splits ARM_Structs_A.pde.splits
|
||||
| wp_once hoare_drop_imps)+
|
||||
done
|
||||
|
||||
definition
|
||||
"pti_duplicates_valid iv \<equiv>
|
||||
|
|
|
@ -392,14 +392,7 @@ lemma (in pspace_update_eq) pspace_no_overlap_update [simp]:
|
|||
"pspace_no_overlap ptr bits (f s) = pspace_no_overlap ptr bits s"
|
||||
by (simp add: pspace_no_overlap_def pspace)
|
||||
|
||||
|
||||
(* FIXME: Move *)
|
||||
lemma mask_twice:
|
||||
"(x && mask n) && mask m = x && mask (min m n)"
|
||||
apply (rule word_eqI)
|
||||
apply (simp add: word_size conj_ac)
|
||||
done
|
||||
|
||||
lemma multi_lessD:
|
||||
"\<lbrakk>(a::nat)*b < c;0<a;0<b\<rbrakk> \<Longrightarrow> a < c \<and> b < c"
|
||||
by (cases a, simp_all,cases b,simp_all)
|
||||
|
@ -936,70 +929,6 @@ lemma null_filter_caps_of_state_foldr:
|
|||
done
|
||||
|
||||
|
||||
lemma aligned_shift:
|
||||
"\<lbrakk>x < 2 ^ n; is_aligned (y :: 'a :: len word) n;n \<le> len_of TYPE('a)\<rbrakk>
|
||||
\<Longrightarrow> x + y >> n = y >> n"
|
||||
apply (subst word_plus_and_or_coroll)
|
||||
apply (rule word_eqI)
|
||||
apply (clarsimp simp: is_aligned_nth)
|
||||
apply (drule(1) nth_bounded)
|
||||
apply simp
|
||||
apply simp
|
||||
apply (rule word_eqI)
|
||||
apply (simp add: nth_shiftr)
|
||||
apply safe
|
||||
apply (drule(1) nth_bounded)
|
||||
apply simp+
|
||||
done
|
||||
|
||||
|
||||
lemma aligned_shift':
|
||||
"\<lbrakk>x < 2 ^ n; is_aligned (y :: 'a :: len word) n;n \<le> len_of TYPE('a)\<rbrakk>
|
||||
\<Longrightarrow> y + x >> n = y >> n"
|
||||
apply (subst word_plus_and_or_coroll)
|
||||
apply (rule word_eqI)
|
||||
apply (clarsimp simp: is_aligned_nth)
|
||||
apply (drule(1) nth_bounded)
|
||||
apply simp
|
||||
apply simp
|
||||
apply (rule word_eqI)
|
||||
apply (simp add: nth_shiftr)
|
||||
apply safe
|
||||
apply (drule(1) nth_bounded)
|
||||
apply simp+
|
||||
done
|
||||
|
||||
|
||||
lemma neg_mask_add_mask:
|
||||
"((x:: 'a :: len word) && ~~ mask n) + (2 ^ n - 1) = x || mask n"
|
||||
apply (simp add:mask_2pm1[symmetric])
|
||||
apply (rule word_eqI)
|
||||
apply (rule iffI)
|
||||
apply (clarsimp simp:word_size not_less)
|
||||
apply (cut_tac w = "((x && ~~ mask n) + mask n)" and
|
||||
m = n and n = "na - n" in nth_shiftr[symmetric])
|
||||
apply clarsimp
|
||||
apply (subst (asm) aligned_shift')
|
||||
apply (simp add:mask_lt_2pn nth_shiftr is_aligned_neg_mask word_size word_bits_def )+
|
||||
apply (case_tac "na<n")
|
||||
apply clarsimp
|
||||
apply (subst word_plus_and_or_coroll)
|
||||
apply (rule iffD1[OF is_aligned_mask])
|
||||
apply (simp add:is_aligned_neg_mask word_or_nth word_size not_less)+
|
||||
apply (cut_tac w = "((x && ~~ mask n) + mask n)" and
|
||||
m = n and n = "na - n" in nth_shiftr[symmetric])
|
||||
apply clarsimp
|
||||
apply (subst (asm) aligned_shift')
|
||||
apply (simp add:mask_lt_2pn is_aligned_neg_mask word_bits_def nth_shiftr neg_mask_bang)+
|
||||
done
|
||||
|
||||
(* FIXME: move *)
|
||||
lemma subtract_mask:
|
||||
"p - (p && mask n) = (p && ~~ mask n)"
|
||||
"p - (p && ~~ mask n) = (p && mask n)"
|
||||
by (simp add: field_simps word_plus_and_or_coroll2)+
|
||||
|
||||
|
||||
lemma retype_addrs_fold:
|
||||
" map (\<lambda>p. ptr_add ptr' (p * 2 ^ obj_bits_api ty us)) [0..< n ]
|
||||
= retype_addrs ptr' ty n us"
|
||||
|
@ -1290,6 +1219,15 @@ crunch cte_wp_at[wp]: init_arch_objects "\<lambda>s. P (cte_wp_at P' p s)"
|
|||
crunch typ_at[wp]: init_arch_objects "\<lambda>s. P (typ_at T p s)"
|
||||
(ignore: clearMemory wp: crunch_wps)
|
||||
|
||||
lemma mdb_cte_at_store_pde[wp]:
|
||||
"\<lbrace>\<lambda>s. mdb_cte_at (swp (cte_wp_at (op \<noteq> cap.NullCap)) s) (cdt s)\<rbrace>
|
||||
store_pde y pde
|
||||
\<lbrace>\<lambda>r s. mdb_cte_at (swp (cte_wp_at (op \<noteq> cap.NullCap)) s) (cdt s)\<rbrace>"
|
||||
apply (clarsimp simp:mdb_cte_at_def)
|
||||
apply (simp only: imp_conv_disj)
|
||||
apply (wp hoare_vcg_disj_lift hoare_vcg_all_lift)
|
||||
done
|
||||
|
||||
(*
|
||||
FIXME: This didn't used to assume the slot wasn't in kernel_mapping_slots.
|
||||
It may be necessary to put another rule in the wp set that says
|
||||
|
|
|
@ -100,6 +100,37 @@ proof -
|
|||
done
|
||||
qed
|
||||
|
||||
lemma storeWord_valid_irq_states:
|
||||
"\<lbrace>\<lambda>m. valid_irq_states (s\<lparr>machine_state := m\<rparr>)\<rbrace> storeWord x y
|
||||
\<lbrace>\<lambda>a b. valid_irq_states (s\<lparr>machine_state := b\<rparr>)\<rbrace>"
|
||||
apply(simp add: valid_irq_states_def | wp | simp add: no_irq_storeWord)+
|
||||
done
|
||||
|
||||
lemma dmo_storeWord_valid_irq_states[wp]:
|
||||
"\<lbrace>valid_irq_states\<rbrace> do_machine_op (storeWord x y) \<lbrace>\<lambda>_. valid_irq_states\<rbrace>"
|
||||
apply(simp add: do_machine_op_def | wp | wpc)+
|
||||
apply clarsimp
|
||||
apply(erule use_valid[OF _ storeWord_valid_irq_states])
|
||||
by simp
|
||||
|
||||
lemma dmo_mapM_storeWord_0_invs[wp]:
|
||||
"valid invs (do_machine_op (mapM (\<lambda>p. storeWord p 0) S)) (\<lambda>_. invs)"
|
||||
apply (simp add: dom_mapM ef_storeWord)
|
||||
apply (rule mapM_UNIV_wp)
|
||||
apply (simp add: do_machine_op_def split_def)
|
||||
apply wp
|
||||
apply (clarsimp simp: invs_def valid_state_def cur_tcb_def
|
||||
valid_machine_state_def)
|
||||
apply (rule conjI)
|
||||
apply(erule use_valid[OF _ storeWord_valid_irq_states], simp)
|
||||
apply (clarsimp simp: disj_commute[of "in_user_frame p s", standard])
|
||||
apply (drule_tac x=p in spec, simp)
|
||||
apply (erule use_valid)
|
||||
apply (simp add: storeWord_def word_rsplit_0)
|
||||
apply wp
|
||||
apply simp
|
||||
done
|
||||
|
||||
lemma dmo_kheap_arch_state[wp]:
|
||||
"\<lbrace>\<lambda>s. P (kheap s) (arch_state s)\<rbrace>
|
||||
do_machine_op a
|
||||
|
|
|
@ -744,6 +744,13 @@ lemma gts_wf[wp]: "\<lbrace>tcb_at t and invs\<rbrace> get_thread_state t \<lbra
|
|||
apply (clarsimp simp: valid_obj_def valid_tcb_def)
|
||||
done
|
||||
|
||||
lemma idle_thread_idle[wp]:
|
||||
"\<lbrace>\<lambda>s. valid_idle s \<and> t = idle_thread s\<rbrace> get_thread_state t \<lbrace>\<lambda>r s. idle r\<rbrace>"
|
||||
apply (clarsimp simp: valid_def get_thread_state_def thread_get_def bind_def return_def gets_the_def gets_def get_def assert_opt_def get_tcb_def
|
||||
fail_def valid_idle_def st_tcb_at_def obj_at_def
|
||||
split: option.splits Structures_A.kernel_object.splits)
|
||||
done
|
||||
|
||||
lemma set_thread_state_valid_objs[wp]:
|
||||
"\<lbrace>valid_objs and valid_tcb_state st and
|
||||
(\<lambda>s. (\<forall>a b. st = Structures_A.BlockedOnReceive a b \<longrightarrow>
|
||||
|
|
|
@ -936,6 +936,31 @@ lemma mdb_Null_None:
|
|||
apply (erule(1) mdb_cte_at_Null_None)
|
||||
done
|
||||
|
||||
lemma not_waiting_reply_slot_no_descendants:
|
||||
"\<lbrakk> st_tcb_at (Not \<circ> awaiting_reply) t s;
|
||||
valid_reply_caps s; valid_objs s; valid_mdb s \<rbrakk>
|
||||
\<Longrightarrow> descendants_of (t, tcb_cnode_index 2) (cdt s) = {}"
|
||||
apply (rule ccontr, erule nonemptyE)
|
||||
apply (clarsimp simp: valid_mdb_def reply_mdb_def reply_masters_mdb_def)
|
||||
apply (frule_tac ref="tcb_cnode_index 2" in tcb_at_cte_at[OF st_tcb_at_tcb_at])
|
||||
apply (simp add: domI)
|
||||
apply (clarsimp simp: cte_wp_at_caps_of_state)
|
||||
apply (frule(1) tcb_cap_valid_caps_of_stateD)
|
||||
apply (clarsimp simp: tcb_cap_valid_def st_tcb_at_tcb_at)
|
||||
apply (clarsimp simp: st_tcb_def2)
|
||||
apply (erule disjE)
|
||||
apply (clarsimp simp: cte_wp_at_caps_of_state is_cap_simps)
|
||||
apply (elim allE, drule(1) mp, clarsimp)
|
||||
apply (drule(1) bspec)
|
||||
apply (drule has_reply_cap_cte_wpD[OF caps_of_state_cteD])
|
||||
apply (erule notE[rotated], strengthen reply_cap_doesnt_exist_strg)
|
||||
apply (simp add: st_tcb_def2)
|
||||
apply clarsimp
|
||||
apply (frule mdb_Null_descendants[OF caps_of_state_cteD])
|
||||
apply (simp add: valid_mdb_def reply_mdb_def reply_masters_mdb_def)
|
||||
apply simp
|
||||
done
|
||||
|
||||
lemma more_revokables[simp]:
|
||||
"pspace_distinct (is_original_cap_update f s) = pspace_distinct s"
|
||||
"pspace_aligned (is_original_cap_update f s) = pspace_aligned s"
|
||||
|
|
|
@ -1510,17 +1510,20 @@ definition
|
|||
(* FIXME cache-sprint: we didn't need any new invariants for PageFlush - is this bad? *)
|
||||
definition
|
||||
"valid_page_inv pi \<equiv> case pi of
|
||||
ArchInvocation_A.PageMap cap ptr m \<Rightarrow>
|
||||
ArchInvocation_A.PageMap asid cap ptr m \<Rightarrow>
|
||||
cte_wp_at (is_arch_update cap and (op = None \<circ> vs_cap_ref)) ptr
|
||||
and cte_wp_at is_pg_cap ptr
|
||||
and (\<lambda>s. same_refs m cap s)
|
||||
and valid_slots m
|
||||
and valid_cap cap
|
||||
and K (is_pg_cap cap \<and> empty_refs m)
|
||||
and K (is_pg_cap cap \<and> empty_refs m \<and> asid \<le> mask asid_bits \<and> asid \<noteq> 0)
|
||||
and (\<lambda>s. \<exists>slot. cte_wp_at (parent_for_refs m) slot s)
|
||||
| ArchInvocation_A.PageRemap m \<Rightarrow>
|
||||
valid_slots m and K (empty_refs m)
|
||||
and (\<lambda>s. \<exists>pd. pd_at_asid asid pd s)
|
||||
| ArchInvocation_A.PageRemap asid m \<Rightarrow>
|
||||
valid_slots m and K (empty_refs m \<and> asid \<le> mask asid_bits \<and> asid \<noteq> 0)
|
||||
and (\<lambda>s. \<exists>slot. cte_wp_at (parent_for_refs m) slot s)
|
||||
and (\<lambda>s. \<exists>slot. cte_wp_at (\<lambda>cap. same_refs m cap s) slot s)
|
||||
and (\<lambda>s. \<exists>pd. pd_at_asid asid pd s)
|
||||
| ArchInvocation_A.PageUnmap cap ptr \<Rightarrow>
|
||||
\<lambda>s. \<exists>r R sz m. cap = ARM_Structs_A.PageCap r R sz m \<and>
|
||||
option_case True (valid_unmap sz) m \<and>
|
||||
|
@ -2777,6 +2780,35 @@ lemma no_cap_to_obj_with_diff_ref_map:
|
|||
lemmas store_pte_cte_wp_at1[wp]
|
||||
= hoare_cte_wp_caps_of_state_lift [OF store_pte_caps_of_state]
|
||||
|
||||
lemma mdb_cte_at_store_pte[wp]:
|
||||
"\<lbrace>\<lambda>s. mdb_cte_at (swp (cte_wp_at (op \<noteq> cap.NullCap)) s) (cdt s)\<rbrace>
|
||||
store_pte y pte
|
||||
\<lbrace>\<lambda>r s. mdb_cte_at (swp (cte_wp_at (op \<noteq> cap.NullCap)) s) (cdt s)\<rbrace>"
|
||||
apply (clarsimp simp:mdb_cte_at_def)
|
||||
apply (simp only: imp_conv_disj)
|
||||
apply (wp hoare_vcg_disj_lift hoare_vcg_all_lift)
|
||||
apply (simp add:store_pte_def set_pt_def)
|
||||
apply wp
|
||||
apply (rule hoare_drop_imp)
|
||||
apply (wp|simp)+
|
||||
done
|
||||
|
||||
lemma valid_idle_store_pte[wp]:
|
||||
"\<lbrace>valid_idle\<rbrace> store_pte y pte \<lbrace>\<lambda>rv. valid_idle\<rbrace>"
|
||||
apply (simp add:store_pte_def)
|
||||
apply wp
|
||||
apply (rule hoare_vcg_precond_imp[where Q="valid_idle"])
|
||||
apply (simp add:set_pt_def)
|
||||
apply wp
|
||||
apply (simp add:get_object_def)
|
||||
apply wp
|
||||
apply (clarsimp simp:obj_at_def
|
||||
split:Structures_A.kernel_object.splits arch_kernel_obj.splits)
|
||||
apply (fastforce simp:is_tcb_def)
|
||||
apply (assumption)
|
||||
apply (wp|simp)+
|
||||
done
|
||||
|
||||
lemma mapM_swp_store_pte_invs[wp]:
|
||||
"\<lbrace>invs and (\<lambda>s. (\<exists>p\<in>set slots. (\<exists>\<rhd> (p && ~~ mask pt_bits)) s) \<longrightarrow>
|
||||
valid_pte pte s) and
|
||||
|
@ -4780,6 +4812,11 @@ lemma unmap_page_section_unmapped:
|
|||
|
||||
crunch global_refs: store_pde "\<lambda>s. P (global_refs s)"
|
||||
|
||||
crunch invs[wp]: pte_check_if_mapped, pde_check_if_mapped "invs"
|
||||
|
||||
crunch vs_lookup[wp]: pte_check_if_mapped, pde_check_if_mapped "\<lambda>s. P (vs_lookup s)"
|
||||
|
||||
crunch valid_pte[wp]: pte_check_if_mapped "\<lambda>s. P (valid_pte p s)"
|
||||
|
||||
lemma perform_page_invs [wp]:
|
||||
"\<lbrace>invs and valid_page_inv pi\<rbrace> perform_page_invocation pi \<lbrace>\<lambda>_. invs\<rbrace>"
|
||||
|
@ -4788,26 +4825,26 @@ lemma perform_page_invs [wp]:
|
|||
-- "PageMap"
|
||||
apply clarsimp
|
||||
apply (rule hoare_pre)
|
||||
apply (wp dmo_ccr_invs arch_update_cap_invs_map
|
||||
hoare_vcg_const_Ball_lift
|
||||
hoare_vcg_const_imp_lift hoare_vcg_all_lift set_cap_typ_at
|
||||
mapM_swp_store_pde_invs_unmap
|
||||
store_pde_invs_unmap' mapM_swp_store_pte_invs
|
||||
hoare_vcg_ex_lift hoare_vcg_ball_lift set_cap_arch_obj
|
||||
set_cap_vs_lookup
|
||||
|wpc|simp del: fun_upd_apply|subst cte_wp_at_caps_of_state)+
|
||||
apply (wp get_master_pte_wp get_master_pde_wp mapM_swp_store_pde_invs_unmap store_pde_invs_unmap' hoare_vcg_const_imp_lift hoare_vcg_all_lift set_cap_arch_obj arch_update_cap_invs_map
|
||||
| wpc
|
||||
| simp add: pte_check_if_mapped_def pde_check_if_mapped_def del: fun_upd_apply
|
||||
| subst cte_wp_at_caps_of_state)+
|
||||
apply (wp_once hoare_drop_imp)
|
||||
apply (wp arch_update_cap_invs_map)
|
||||
apply (rule hoare_vcg_conj_lift)
|
||||
apply (rule hoare_lift_Pf[where f=vs_lookup, OF _ set_cap_vs_lookup])
|
||||
apply (rule_tac f="valid_pte xa" in hoare_lift_Pf[OF _ set_cap_valid_pte_stronger])
|
||||
apply wp
|
||||
apply (rule hoare_lift_Pf2[where f=vs_lookup, OF _ set_cap_vs_lookup])
|
||||
apply (wp dmo_ccr_invs arch_update_cap_invs_map
|
||||
apply ((wp dmo_ccr_invs arch_update_cap_invs_map
|
||||
hoare_vcg_const_Ball_lift
|
||||
hoare_vcg_const_imp_lift hoare_vcg_all_lift set_cap_typ_at
|
||||
hoare_vcg_ex_lift hoare_vcg_ball_lift set_cap_arch_obj
|
||||
set_cap_vs_lookup
|
||||
| wpc | simp add: same_refs_def del: fun_upd_apply
|
||||
| subst cte_wp_at_caps_of_state)+
|
||||
set_cap_vs_lookup
|
||||
| wpc | simp add: same_refs_def del: fun_upd_apply split del: split_if
|
||||
| subst cte_wp_at_caps_of_state)+)
|
||||
apply (wp_once hoare_drop_imp)
|
||||
apply (wp arch_update_cap_invs_map hoare_vcg_ex_lift set_cap_arch_obj)
|
||||
apply (clarsimp simp: valid_page_inv_def cte_wp_at_caps_of_state neq_Nil_conv
|
||||
valid_slots_def empty_refs_def parent_for_refs_def
|
||||
simp del: fun_upd_apply del: exE
|
||||
|
@ -4848,9 +4885,9 @@ lemma perform_page_invs [wp]:
|
|||
apply (rule conjI)
|
||||
apply (clarsimp simp: pde_at_def obj_at_def
|
||||
caps_of_state_cteD'[where P=\<top>, simplified])
|
||||
apply (drule_tac cap=capb and ptr="(aa,ba)"
|
||||
apply (drule_tac cap=capc and ptr="(aa,ba)"
|
||||
in valid_global_refsD[OF invs_valid_global_refs])
|
||||
apply assumption+
|
||||
apply assumption+
|
||||
apply (clarsimp simp: cap_range_def)
|
||||
apply (clarsimp)
|
||||
apply (rule conjI)
|
||||
|
@ -4860,41 +4897,37 @@ lemma perform_page_invs [wp]:
|
|||
apply (erule ballEI)
|
||||
apply (clarsimp simp: pde_at_def obj_at_def
|
||||
caps_of_state_cteD'[where P=\<top>, simplified])
|
||||
apply (drule_tac cap=capb and ptr="(aa,ba)"
|
||||
apply (drule_tac cap=capc and ptr="(aa,ba)"
|
||||
in valid_global_refsD[OF invs_valid_global_refs])
|
||||
apply assumption+
|
||||
apply (drule_tac x=x in imageI[where f="\<lambda>x. x && ~~ mask pd_bits"])
|
||||
apply (drule (1) subsetD)
|
||||
apply (clarsimp simp: cap_range_def)
|
||||
-- "PageRemap"
|
||||
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"
|
||||
apply (rule hoare_pre)
|
||||
apply (wp dmo_ccr_invs hoare_vcg_const_imp_lift hoare_vcg_all_lift
|
||||
mapM_swp_store_pte_invs arch_update_cap_invs_map get_cap_wp
|
||||
mapM_x_swp_store_pde_invs_unmap mapM_swp_store_pde_invs_unmap
|
||||
hoare_vcg_ex_lift
|
||||
| wpc
|
||||
| simp
|
||||
| (rule hoare_vcg_conj_lift, rule_tac slots=ba in store_pde_invs_unmap'))+
|
||||
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
|
||||
| (rule hoare_vcg_conj_lift, rule_tac slots=ba in store_pde_invs_unmap'))+
|
||||
apply (clarsimp simp: valid_page_inv_def cte_wp_at_caps_of_state
|
||||
valid_slots_def empty_refs_def neq_Nil_conv
|
||||
split: sum.splits)
|
||||
apply (clarsimp simp: parent_for_refs_def same_refs_def is_cap_simps cap_asid_def split:option.splits)
|
||||
apply (rule conjI, fastforce)
|
||||
apply (rule conjI, fastforce)
|
||||
apply (rule conjI)
|
||||
apply clarsimp
|
||||
apply (rule_tac x=ac in exI, rule_tac x=bc in exI, rule_tac x=capa in exI)
|
||||
apply clarsimp
|
||||
apply (erule (2) ref_is_unique[OF _ _ reachable_page_table_not_global])
|
||||
apply (simp_all add: invs_def valid_state_def valid_arch_state_def
|
||||
valid_arch_caps_def valid_pspace_def valid_objs_caps)[9]
|
||||
apply (simp_all add: invs_def valid_state_def valid_arch_state_def
|
||||
valid_arch_caps_def valid_pspace_def valid_objs_caps)[9]
|
||||
apply fastforce
|
||||
apply( frule valid_global_refsD2)
|
||||
apply (clarsimp simp: cap_range_def parent_for_refs_def)+
|
||||
apply (rule conjI)
|
||||
apply (rule conjI, rule impI)
|
||||
apply (rule exI, rule exI, rule exI)
|
||||
apply (erule conjI)
|
||||
apply clarsimp
|
||||
apply (rule conjI)
|
||||
apply (rule conjI, rule impI)
|
||||
apply (rule_tac x=ac in exI, rule_tac x=bc in exI, rule_tac x=capa in exI)
|
||||
apply (clarsimp simp: same_refs_def pde_ref_def pde_ref_pages_def
|
||||
valid_pde_def invs_def valid_state_def valid_pspace_def)
|
||||
|
@ -4910,14 +4943,12 @@ lemma perform_page_invs [wp]:
|
|||
ARM_Structs_A.arch_kernel_obj.splits option.splits
|
||||
arch_cap.splits)+)[2]
|
||||
apply (clarsimp simp: pde_at_def obj_at_def a_type_def)
|
||||
apply (case_tac ko, (simp_all split: split_if_asm))
|
||||
apply (case_tac arch_kernel_obj, simp_all)
|
||||
apply (rule conjI)
|
||||
apply clarsimp
|
||||
apply (drule_tac ptr="(ab,bb)" in
|
||||
valid_global_refsD[OF invs_valid_global_refs caps_of_state_cteD])
|
||||
apply simp+
|
||||
apply force
|
||||
apply force
|
||||
apply (erule ballEI)
|
||||
apply clarsimp
|
||||
apply (drule_tac ptr="(ab,bb)" in
|
||||
|
|
|
@ -727,11 +727,6 @@ lemma pt_slot_eq:
|
|||
apply clarsimp
|
||||
done
|
||||
|
||||
lemma ekheap_kheap_dom:
|
||||
"\<lbrakk>ekheap s x = Some etcb; valid_etcbs s\<rbrakk>
|
||||
\<Longrightarrow> \<exists>tcb. kheap s x = Some (TCB tcb)"
|
||||
by (fastforce simp: valid_etcbs_def st_tcb_at_def obj_at_def is_etcb_at_def)
|
||||
|
||||
-- "set_other_obj_corres unfortunately doesn't work here"
|
||||
lemma set_pd_corres:
|
||||
"pde_relation_aligned (p>>2) pde pde' \<Longrightarrow>
|
||||
|
|
|
@ -1123,7 +1123,7 @@ lemma inv_arch_corres:
|
|||
"archinv_relation ai ai' \<Longrightarrow>
|
||||
corres (intr \<oplus> op=)
|
||||
(einvs and ct_active and valid_arch_inv ai)
|
||||
(invs' and ct_active' and valid_arch_inv' ai')
|
||||
(invs' and ct_active' and valid_arch_inv' ai' and (\<lambda>s. vs_valid_duplicates' (ksPSpace s)))
|
||||
(arch_perform_invocation ai) (ArchRetypeDecls_H.performInvocation ai')"
|
||||
apply (clarsimp simp: arch_perform_invocation_def
|
||||
ArchRetype_H.performInvocation_def
|
||||
|
|
|
@ -1083,18 +1083,6 @@ section "Lemmas"
|
|||
lemmas objBits_simps = objBits_def objBitsKO_def word_size_def
|
||||
lemmas objBitsKO_simps = objBitsKO_def word_size_def objBits_simps
|
||||
|
||||
(* FIXME: move *)
|
||||
lemma and_neg_mask_plus_mask_mono: "(p && ~~ mask n) + mask n \<ge> (p::word32)"
|
||||
apply (subst add_commute)
|
||||
apply (rule word_le_minus_cancel[where x = "p && ~~ mask n"])
|
||||
apply (clarsimp simp: subtract_mask)
|
||||
using word_and_le1[where a = "mask n" and y = p]
|
||||
apply (clarsimp simp: mask_def word_le_less_eq)
|
||||
apply (subst add_commute)
|
||||
apply (rule is_aligned_no_overflow'[folded mask_2pm1])
|
||||
apply (clarsimp simp: is_aligned_neg_mask)
|
||||
done
|
||||
|
||||
lemma valid_duplicates'_D:
|
||||
"\<lbrakk>vs_valid_duplicates' m; m (p::word32) = Some ko;is_aligned p' 2;
|
||||
p && ~~ mask (vs_ptr_align ko) = p' && ~~ mask (vs_ptr_align ko)\<rbrakk>
|
||||
|
|
|
@ -2605,14 +2605,6 @@ lemma setupCallerCap_tcb_at'[wp]:
|
|||
crunch ct'[wp]: setupCallerCap "\<lambda>s. P (ksCurThread s)"
|
||||
(wp: crunch_wps)
|
||||
|
||||
lemma corres_symmetric_bool_cases:
|
||||
"\<lbrakk> P = P'; \<lbrakk> P; P' \<rbrakk> \<Longrightarrow> corres r Q Q' f g;
|
||||
\<lbrakk> \<not> P; \<not> P' \<rbrakk> \<Longrightarrow> corres r R R' f g \<rbrakk>
|
||||
\<Longrightarrow> corres r (\<lambda>s. (P \<longrightarrow> Q s) \<and> (\<not> P \<longrightarrow> R s))
|
||||
(\<lambda>s. (P' \<longrightarrow> Q' s) \<and> (\<not> P' \<longrightarrow> R' s))
|
||||
f g"
|
||||
by (cases P, simp_all)
|
||||
|
||||
lemma cteInsert_sch_act_wf[wp]:
|
||||
"\<lbrace>\<lambda>s. sch_act_wf (ksSchedulerAction s) s\<rbrace>
|
||||
cteInsert newCap srcSlot destSlot
|
||||
|
|
|
@ -1062,14 +1062,6 @@ lemma set_other_obj_corres:
|
|||
apply (clarsimp simp: a_type_def other_obj_relation_def etcb_relation_def is_other_obj_relation_type t exst_same_def split: Structures_A.kernel_object.splits Structures_H.kernel_object.splits ARM_Structs_A.arch_kernel_obj.splits)
|
||||
done
|
||||
|
||||
lemma ekheap_tcb_at:
|
||||
"\<lbrakk>ekheap s x = Some y; valid_etcbs s\<rbrakk> \<Longrightarrow> tcb_at x s"
|
||||
by (fastforce simp: valid_etcbs_def is_etcb_at_def st_tcb_at_def obj_at_def is_tcb_def)
|
||||
|
||||
lemma kheap_ekheap_dom:
|
||||
"\<lbrakk>kheap s x = Some (TCB tcb); valid_etcbs s\<rbrakk> \<Longrightarrow> \<exists>etcb. ekheap s x = Some etcb"
|
||||
by (force simp: valid_etcbs_def is_etcb_at_def st_tcb_at_def obj_at_def)
|
||||
|
||||
lemma set_ep_corres:
|
||||
"ep_relation e e' \<Longrightarrow>
|
||||
corres dc (ep_at ptr) (ep_at' ptr)
|
||||
|
|
|
@ -2137,7 +2137,7 @@ lemma performPageInvocation_no_orphans [wp]:
|
|||
apply (simp add: performPageInvocation_def
|
||||
cong: page_invocation.case_cong)
|
||||
apply (rule hoare_pre)
|
||||
apply (wp mapM_x_wp' mapM_wp' static_imp_wp | wpc | clarsimp)+
|
||||
apply (wp mapM_x_wp' mapM_wp' static_imp_wp | wpc | clarsimp simp: pdeCheckIfMapped_def pteCheckIfMapped_def)+
|
||||
done
|
||||
|
||||
lemma performASIDControlInvocation_no_orphans [wp]:
|
||||
|
|
|
@ -2331,6 +2331,18 @@ lemma invokeCNode_valid_duplicates'[wp]:
|
|||
apply (simp add:invs_valid_objs' invs_pspace_aligned')
|
||||
done
|
||||
|
||||
lemma getObject_pte_sp:
|
||||
"\<lbrace>P\<rbrace> getObject r \<lbrace>\<lambda>t::pte. P and ko_at' t r\<rbrace>"
|
||||
apply (wp getObject_ko_at)
|
||||
apply (auto simp: objBits_simps archObjSize_def)
|
||||
done
|
||||
|
||||
lemma getObject_pde_sp:
|
||||
"\<lbrace>P\<rbrace> getObject r \<lbrace>\<lambda>t::pde. P and ko_at' t r\<rbrace>"
|
||||
apply (wp getObject_ko_at)
|
||||
apply (auto simp: objBits_simps archObjSize_def)
|
||||
done
|
||||
|
||||
lemma performPageInvocation_valid_duplicates'[wp]:
|
||||
"\<lbrace>invs' and valid_arch_inv' (invocation.InvokePage page_invocation)
|
||||
and (\<lambda>s. vs_valid_duplicates' (ksPSpace s))\<rbrace>
|
||||
|
@ -2338,46 +2350,61 @@ lemma performPageInvocation_valid_duplicates'[wp]:
|
|||
\<lbrace>\<lambda>y a. vs_valid_duplicates' (ksPSpace a)\<rbrace>"
|
||||
apply (rule hoare_name_pre_state)
|
||||
apply (case_tac page_invocation)
|
||||
apply (simp_all add:performPageInvocation_def)
|
||||
-- "PageFlush"
|
||||
apply (simp_all add:performPageInvocation_def pteCheckIfMapped_def pdeCheckIfMapped_def)
|
||||
apply (wp|simp|wpc)+
|
||||
apply (case_tac sum)
|
||||
-- "PageRemap"
|
||||
apply (case_tac sum)
|
||||
apply (case_tac a)
|
||||
apply (case_tac aa)
|
||||
apply (clarsimp simp:valid_arch_inv'_def
|
||||
valid_page_inv'_def valid_slots'_def
|
||||
valid_slots_duplicated'_def mapM_x_singleton)
|
||||
apply (wp PageTableDuplicates.storePTE_no_duplicates')
|
||||
valid_slots_duplicated'_def mapM_singleton)
|
||||
apply (wp PageTableDuplicates.storePTE_no_duplicates' getPTE_wp | simp add: if_cancel)+
|
||||
apply (simp add:vs_entry_align_def)
|
||||
apply (wp|wpc|simp)+
|
||||
apply (subst mapM_discarded)
|
||||
apply simp
|
||||
apply (rule hoare_seq_ext[OF _ getObject_pte_sp])
|
||||
apply (wp|simp)+
|
||||
apply (clarsimp simp:valid_arch_inv'_def
|
||||
valid_page_inv'_def valid_slots'_def
|
||||
valid_slots_duplicated'_def)
|
||||
apply (rule hoare_pre)
|
||||
apply (rule_tac sz = 6 and ptr = "p && ~~ mask ptBits" and word = p
|
||||
in mapM_x_storePTE_update_helper)
|
||||
apply (simp add:invs_pspace_aligned' pageBits_def ptBits_def)
|
||||
apply (simp add:invs_pspace_aligned' pageBits_def ptBits_def)
|
||||
apply (subst mapM_discarded)
|
||||
apply (clarsimp simp:valid_arch_inv'_def
|
||||
valid_page_inv'_def valid_slots'_def
|
||||
valid_slots_duplicated'_def mapM_x_singleton)
|
||||
apply (wp PageTableDuplicates.storePTE_no_duplicates')
|
||||
apply (rule hoare_seq_ext[OF _ getObject_pte_sp])
|
||||
apply (wp PageTableDuplicates.storePTE_no_duplicates' | simp)+
|
||||
apply (simp add:vs_entry_align_def)
|
||||
apply (subst mapM_discarded)+
|
||||
apply (case_tac b)
|
||||
apply (case_tac a)
|
||||
apply (clarsimp simp:valid_arch_inv'_def
|
||||
valid_page_inv'_def valid_slots'_def
|
||||
valid_slots_duplicated'_def mapM_x_singleton)
|
||||
apply (wp PageTableDuplicates.storePDE_no_duplicates')
|
||||
apply (wp|wpc|simp add:vs_entry_align_def)+
|
||||
apply (rule hoare_seq_ext[OF _ getObject_pde_sp])
|
||||
apply (wp PageTableDuplicates.storePDE_no_duplicates' | simp add: when_def)+
|
||||
apply (simp add: vs_entry_align_def)+
|
||||
apply (rule hoare_seq_ext[OF _ getObject_pde_sp])
|
||||
apply (wp|wpc|simp add:vs_entry_align_def)+
|
||||
apply (clarsimp simp:valid_arch_inv'_def vs_entry_align_def
|
||||
valid_page_inv'_def valid_slots'_def
|
||||
valid_slots_duplicated'_def mapM_x_singleton)
|
||||
apply (wp PageTableDuplicates.storePDE_no_duplicates' | wpc | simp)+
|
||||
apply (wp|wpc|simp add:vs_entry_align_def)+
|
||||
apply ((wp PageTableDuplicates.storePDE_no_duplicates' | wpc | simp)+)[1]
|
||||
apply (simp add: vs_entry_align_def)+
|
||||
apply (rule hoare_seq_ext[OF _ getObject_pde_sp])
|
||||
apply (wp|wpc|simp add:vs_entry_align_def)+
|
||||
apply (clarsimp simp:valid_arch_inv'_def
|
||||
valid_page_inv'_def valid_slots'_def
|
||||
valid_slots_duplicated'_def mapM_x_singleton)
|
||||
apply (wp PageTableDuplicates.storePDE_no_duplicates')
|
||||
apply (wp|wpc|simp add:vs_entry_align_def)+
|
||||
apply (simp add: vs_entry_align_def)+
|
||||
apply (rule hoare_seq_ext[OF _ getObject_pde_sp])
|
||||
apply (wp|wpc|simp add:vs_entry_align_def)+
|
||||
apply (clarsimp simp:valid_arch_inv'_def
|
||||
valid_page_inv'_def valid_slots'_def
|
||||
valid_slots_duplicated'_def mapM_x_singleton)
|
||||
|
@ -2387,53 +2414,55 @@ lemma performPageInvocation_valid_duplicates'[wp]:
|
|||
apply clarsimp
|
||||
apply (simp add:invs_pspace_aligned' ptBits_def
|
||||
pdBits_def field_simps pageBits_def)+
|
||||
-- "PageMap"
|
||||
apply (clarsimp simp: pteCheckIfMapped_def pdeCheckIfMapped_def)
|
||||
apply (clarsimp simp:valid_pde_slots'_def valid_page_inv'_def
|
||||
valid_slots_duplicated'_def valid_arch_inv'_def)
|
||||
valid_slots_duplicated'_def valid_arch_inv'_def )
|
||||
apply (case_tac sum)
|
||||
apply (case_tac a)
|
||||
apply (case_tac aa)
|
||||
apply clarsimp
|
||||
apply (clarsimp simp: pteCheckIfMapped_def)
|
||||
apply (wp mapM_x_mapM_valid |wpc
|
||||
| simp)+
|
||||
apply (clarsimp simp:valid_slots_duplicated'_def mapM_x_singleton)+
|
||||
apply (rule PageTableDuplicates.storePTE_no_duplicates')
|
||||
apply wp
|
||||
apply (simp add:vs_entry_align_def)+
|
||||
apply clarsimp
|
||||
apply (wp mapM_x_mapM_valid)
|
||||
apply (rule_tac sz = 6 and ptr = "p && ~~ mask ptBits" and word = p in
|
||||
mapM_x_storePTE_update_helper)
|
||||
apply wp
|
||||
apply (simp add:ptBits_def pageBits_def)+
|
||||
apply (simp add:invs_pspace_aligned')
|
||||
apply simp
|
||||
apply (clarsimp simp:mapM_singleton)
|
||||
apply (wp PageTableDuplicates.storePTE_no_duplicates')
|
||||
apply (clarsimp simp:valid_slots_duplicated'_def mapM_x_singleton)+
|
||||
apply (rule PageTableDuplicates.storePTE_no_duplicates', rule getPTE_wp)
|
||||
apply (wp hoare_vcg_all_lift hoare_drop_imps)
|
||||
apply (simp add:vs_entry_align_def)+
|
||||
apply clarsimp
|
||||
apply (clarsimp simp: pteCheckIfMapped_def)
|
||||
apply (wp mapM_x_mapM_valid | simp)+
|
||||
apply (rule_tac sz = 6 and ptr = "p && ~~ mask ptBits" and word = p in
|
||||
mapM_x_storePTE_update_helper)
|
||||
apply (wp getPTE_wp hoare_vcg_all_lift hoare_drop_imps)
|
||||
apply (simp add:ptBits_def pageBits_def)+
|
||||
apply (simp add:invs_pspace_aligned')
|
||||
apply simp
|
||||
apply (clarsimp simp:mapM_singleton pteCheckIfMapped_def)
|
||||
apply (wp PageTableDuplicates.storePTE_no_duplicates' getPTE_wp hoare_drop_imps | simp)+
|
||||
apply (simp add:vs_entry_align_def)+
|
||||
apply (clarsimp simp: pdeCheckIfMapped_def)
|
||||
apply (case_tac a)
|
||||
apply (clarsimp simp:valid_arch_inv'_def
|
||||
valid_page_inv'_def valid_slots'_def
|
||||
valid_slots_duplicated'_def mapM_singleton)
|
||||
apply (wp PageTableDuplicates.storePDE_no_duplicates')
|
||||
apply (wp PageTableDuplicates.storePDE_no_duplicates' getPDE_wp hoare_drop_imps | simp)+
|
||||
apply (simp add:vs_entry_align_def)+
|
||||
apply (clarsimp simp:valid_arch_inv'_def
|
||||
valid_page_inv'_def valid_slots'_def
|
||||
valid_slots_duplicated'_def mapM_singleton)
|
||||
apply (wp PageTableDuplicates.storePDE_no_duplicates')
|
||||
apply (wp PageTableDuplicates.storePDE_no_duplicates' getPDE_wp hoare_drop_imps | simp)+
|
||||
apply (simp add:vs_entry_align_def)+
|
||||
apply (clarsimp simp:valid_arch_inv'_def
|
||||
valid_page_inv'_def valid_slots'_def
|
||||
valid_slots_duplicated'_def mapM_singleton)
|
||||
apply (wp PageTableDuplicates.storePDE_no_duplicates')
|
||||
apply (wp PageTableDuplicates.storePDE_no_duplicates' getPDE_wp hoare_drop_imps | simp)+
|
||||
apply (simp add:vs_entry_align_def)+
|
||||
apply (clarsimp simp:valid_arch_inv'_def
|
||||
valid_page_inv'_def valid_slots'_def
|
||||
valid_slots_duplicated'_def mapM_x_singleton)
|
||||
apply (wp mapM_x_mapM_valid)
|
||||
apply (rule_tac sz = 6 and ptr = "p && ~~ mask pdBits" and word = p
|
||||
in mapM_x_storePDE_update_helper)
|
||||
apply wp
|
||||
apply (wp mapM_x_mapM_valid | simp)+
|
||||
apply (rule_tac sz = 6 and ptr = "p && ~~ mask pdBits" and word = p
|
||||
in mapM_x_storePDE_update_helper)
|
||||
apply wp
|
||||
apply (simp add:pageBits_def pdBits_def ptBits_def)+
|
||||
apply (simp add:invs_pspace_aligned')+
|
||||
apply clarsimp
|
||||
|
|
|
@ -800,10 +800,6 @@ lemma return_wp_exs_valid [wp]: "\<lbrace> P x \<rbrace> return x \<exists>\<lbr
|
|||
lemma get_exs_valid [wp]: "\<lbrace>op = s\<rbrace> get \<exists>\<lbrace>\<lambda>r. op = s\<rbrace>"
|
||||
by (simp add: get_def exs_valid_def)
|
||||
|
||||
lemma tcb_at_is_etcb_at:
|
||||
"\<lbrakk>tcb_at t s; valid_etcbs s\<rbrakk> \<Longrightarrow> is_etcb_at t s"
|
||||
by (simp add: valid_etcbs_def tcb_at_st_tcb_at)
|
||||
|
||||
lemma switch_thread_corres:
|
||||
"corres dc (valid_arch_state and valid_objs and valid_asid_map
|
||||
and valid_arch_objs and pspace_aligned and pspace_distinct
|
||||
|
|
|
@ -497,7 +497,7 @@ lemma pinv_corres:
|
|||
and simple_sched_action
|
||||
and ct_active
|
||||
and (\<lambda>s. (\<exists>w w2 b. i = Invocations_A.InvokeEndpoint w w2 b) \<longrightarrow> st_tcb_at simple (cur_thread s) s))
|
||||
(invs' and sch_act_simple and valid_invocation' i' and ct_active')
|
||||
(invs' and sch_act_simple and valid_invocation' i' and ct_active' and (\<lambda>s. vs_valid_duplicates' (ksPSpace s)))
|
||||
(perform_invocation block call i) (performInvocation block call i')"
|
||||
apply (simp add: performInvocation_def)
|
||||
apply (case_tac i)
|
||||
|
@ -1308,6 +1308,35 @@ done
|
|||
lemmas set_thread_state_active_valid_sched =
|
||||
set_thread_state_runnable_valid_sched[simplified runnable_eq_active]
|
||||
|
||||
lemma setTCB_valid_duplicates'[wp]:
|
||||
"\<lbrace>\<lambda>s. vs_valid_duplicates' (ksPSpace s)\<rbrace>
|
||||
setObject a (tcb::tcb) \<lbrace>\<lambda>rv s. vs_valid_duplicates' (ksPSpace s)\<rbrace>"
|
||||
apply (clarsimp simp: setObject_def split_def valid_def in_monad
|
||||
projectKOs pspace_aligned'_def ps_clear_upd'
|
||||
objBits_def[symmetric] lookupAround2_char1
|
||||
split: split_if_asm)
|
||||
apply (frule pspace_storable_class.updateObject_type[where v = tcb,simplified])
|
||||
apply (clarsimp simp:updateObject_default_def assert_def bind_def
|
||||
alignCheck_def in_monad when_def alignError_def magnitudeCheck_def
|
||||
assert_opt_def return_def fail_def typeError_def
|
||||
split:if_splits option.splits Structures_H.kernel_object.splits)
|
||||
apply (erule valid_duplicates'_non_pd_pt_I[rotated 3],simp+)+
|
||||
done
|
||||
|
||||
crunch valid_duplicates'[wp]: threadSet "\<lambda>s. vs_valid_duplicates' (ksPSpace s)"
|
||||
(ignore: getObject setObject wp: setObject_ksInterrupt updateObject_default_inv)
|
||||
|
||||
lemma tcbSchedEnqueue_valid_duplicates'[wp]:
|
||||
"\<lbrace>\<lambda>s. vs_valid_duplicates' (ksPSpace s)\<rbrace>
|
||||
tcbSchedEnqueue a \<lbrace>\<lambda>rv s. vs_valid_duplicates' (ksPSpace s)\<rbrace>"
|
||||
by (simp add:tcbSchedEnqueue_def unless_def setQueue_def | wp | wpc)+
|
||||
|
||||
crunch valid_duplicates'[wp]: rescheduleRequired "\<lambda>s. vs_valid_duplicates' (ksPSpace s)"
|
||||
(ignore: getObject setObject wp: setObject_ksInterrupt updateObject_default_inv)
|
||||
|
||||
crunch valid_duplicates'[wp]: setThreadState "\<lambda>s. vs_valid_duplicates' (ksPSpace s)"
|
||||
(ignore: getObject setObject)
|
||||
|
||||
(*FIXME: move to NonDetMonadVCG.valid_validE_R *)
|
||||
lemma hinv_corres:
|
||||
"c \<longrightarrow> b \<Longrightarrow>
|
||||
|
@ -1371,7 +1400,8 @@ lemma hinv_corres:
|
|||
apply (rule_tac Q="\<lambda>rv. invs' and valid_invocation' rve'
|
||||
and (\<lambda>s. thread = ksCurThread s)
|
||||
and st_tcb_at' active' thread
|
||||
and (\<lambda>s. ksSchedulerAction s = ResumeCurrentThread)"
|
||||
and (\<lambda>s. ksSchedulerAction s = ResumeCurrentThread)
|
||||
and (\<lambda>s. vs_valid_duplicates' (ksPSpace s))"
|
||||
in hoare_post_imp)
|
||||
apply (clarsimp simp: ct_in_state'_def)
|
||||
apply (frule(1) ct_not_ksQ)
|
||||
|
|
|
@ -1701,12 +1701,6 @@ lemma threadGet_obj_at':
|
|||
"\<lbrace>obj_at' (\<lambda>t. P (f t) t) t\<rbrace> threadGet f t \<lbrace>\<lambda>rv. obj_at' (P rv) t\<rbrace>"
|
||||
by (simp add: threadGet_def o_def | wp getObject_obj_at_tcb)+
|
||||
|
||||
lemma tcb_at_ekheap_dom:
|
||||
"\<lbrakk>tcb_at x s; valid_etcbs s\<rbrakk> \<Longrightarrow> (\<exists>etcb. ekheap s x = Some etcb)"
|
||||
apply (clarsimp simp: obj_at_def is_tcb)
|
||||
apply (erule (1) kheap_ekheap_dom)
|
||||
done
|
||||
|
||||
lemma fun_if_triv[simp]:
|
||||
"(\<lambda>x. if x = y then f y else f x) = f"
|
||||
by (force intro: ext)
|
||||
|
|
|
@ -1724,12 +1724,13 @@ lemma perform_page_directory_corres:
|
|||
|
||||
definition
|
||||
"page_invocation_map pi pi' \<equiv> case pi of
|
||||
ArchInvocation_A.PageMap c ptr m \<Rightarrow>
|
||||
\<exists>c' m'. pi' = PageMap c' (cte_map ptr) m' \<and>
|
||||
ArchInvocation_A.PageMap a c ptr m \<Rightarrow>
|
||||
\<exists>c' m'. pi' = PageMap a c' (cte_map ptr) m' \<and>
|
||||
cap_relation c c' \<and>
|
||||
mapping_map m m'
|
||||
| ArchInvocation_A.PageRemap m \<Rightarrow>
|
||||
\<exists>m'. pi' = PageRemap m' \<and> mapping_map m m'
|
||||
|
||||
| ArchInvocation_A.PageRemap a m \<Rightarrow>
|
||||
\<exists>m'. pi' = PageRemap a m' \<and> mapping_map m m'
|
||||
| ArchInvocation_A.PageUnmap c ptr \<Rightarrow>
|
||||
\<exists>c'. pi' = PageUnmap c' (cte_map ptr) \<and>
|
||||
acap_relation c c'
|
||||
|
@ -1894,10 +1895,10 @@ lemma valid_slots_duplicated_updateCap[wp]:
|
|||
|
||||
definition
|
||||
"valid_page_inv' pi \<equiv> case pi of
|
||||
PageMap cap ptr m \<Rightarrow>
|
||||
PageMap asid cap ptr m \<Rightarrow>
|
||||
cte_wp_at' (is_arch_update' cap) ptr and valid_slots' m and valid_cap' cap
|
||||
and K (valid_pde_slots' m) and (valid_slots_duplicated' m)
|
||||
| PageRemap m \<Rightarrow> valid_slots' m and K (valid_pde_slots' m) and (valid_slots_duplicated' m)
|
||||
| PageRemap asid m \<Rightarrow> valid_slots' m and K (valid_pde_slots' m) and (valid_slots_duplicated' m)
|
||||
| PageUnmap cap ptr \<Rightarrow>
|
||||
\<lambda>s. \<exists>r R sz m. cap = PageCap r R sz m \<and>
|
||||
cte_wp_at' (is_arch_update' (ArchObjectCap cap)) ptr s \<and>
|
||||
|
@ -1962,11 +1963,131 @@ lemma updateCap_valid_slots'[wp]:
|
|||
apply (wp hoare_vcg_ball_lift)
|
||||
done
|
||||
|
||||
lemma pte_check_if_mapped_corres:
|
||||
"corres (op =) (pte_at slot) ((\<lambda>s. vs_valid_duplicates' (ksPSpace s)) and pspace_aligned' and pspace_distinct') (pte_check_if_mapped slot) (pteCheckIfMapped slot)"
|
||||
apply (simp add: pte_check_if_mapped_def pteCheckIfMapped_def)
|
||||
apply (rule corres_guard_imp)
|
||||
apply (rule corres_split[OF _ get_master_pte_corres', simplified])
|
||||
apply (rule corres_return[where P="pte_at slot" and
|
||||
P'="pspace_aligned' and pspace_distinct'", THEN iffD2])
|
||||
apply (clarsimp simp: pte_relation'_def split: )
|
||||
apply (case_tac pt, simp_all)[1]
|
||||
apply wp
|
||||
apply (simp)
|
||||
apply simp
|
||||
done
|
||||
|
||||
lemma pde_check_if_mapped_corres:
|
||||
"corres (op =) (pde_at slot) ((\<lambda>s. vs_valid_duplicates' (ksPSpace s)) and pspace_aligned' and pspace_distinct') (pde_check_if_mapped slot) (pdeCheckIfMapped slot)"
|
||||
apply (simp add: pde_check_if_mapped_def pdeCheckIfMapped_def)
|
||||
apply (rule corres_guard_imp)
|
||||
apply (rule corres_split[OF _ get_master_pde_corres', simplified])
|
||||
apply (rule corres_return[where P="pde_at slot" and
|
||||
P'="pspace_aligned' and pspace_distinct'", THEN iffD2])
|
||||
apply (clarsimp simp: pte_relation'_def split: )
|
||||
apply (case_tac pd, simp_all)[1]
|
||||
apply wp
|
||||
apply (clarsimp simp: pte_relation_aligned_def split: split_if_asm)
|
||||
apply simp
|
||||
done
|
||||
|
||||
crunch unique_table_refs[wp]: do_machine_op, store_pte "\<lambda>s. (unique_table_refs (caps_of_state s))"
|
||||
crunch valid_asid_map[wp]: store_pte "valid_asid_map"
|
||||
|
||||
lemma store_pte_valid_global_objs[wp]:
|
||||
"\<lbrace>valid_global_objs and valid_arch_state and K (aligned_pte pte)\<rbrace> store_pte slot pte \<lbrace>\<lambda>_. valid_global_objs\<rbrace>"
|
||||
apply (simp add: store_pte_def)
|
||||
apply (wp)
|
||||
apply clarsimp
|
||||
apply (frule valid_global_ptsD, simp)
|
||||
apply clarsimp
|
||||
apply (subgoal_tac "pt=pta")
|
||||
apply (auto simp: obj_at_def)
|
||||
done
|
||||
|
||||
lemma set_cap_pd_at_asid [wp]:
|
||||
"\<lbrace>pd_at_asid asid pd\<rbrace> set_cap t st \<lbrace>\<lambda>rv. pd_at_asid asid pd\<rbrace>"
|
||||
apply (simp add: pd_at_asid_def)
|
||||
apply wp
|
||||
done
|
||||
|
||||
lemma set_cap_valid_slots_inv[wp]:
|
||||
"\<lbrace>valid_slots m\<rbrace> set_cap t st \<lbrace>\<lambda>rv. valid_slots m\<rbrace>"
|
||||
by (cases m, (clarsimp simp: valid_slots_def, wp hoare_vcg_ball_lift set_cap_vs_lookup set_cap_typ_ats)+)
|
||||
|
||||
lemma set_cap_same_refs_inv[wp]:
|
||||
"\<lbrace>\<lambda>s. same_refs m cap s\<rbrace> set_cap t st \<lbrace>\<lambda>rv s. same_refs m cap s\<rbrace>"
|
||||
by (cases m, (clarsimp simp: same_refs_def, wp)+)
|
||||
|
||||
definition
|
||||
"valid_page_map_inv asid cap ptr m \<equiv> (\<lambda>s. caps_of_state s ptr = Some cap) and same_refs m cap and
|
||||
valid_slots m and
|
||||
valid_cap cap and
|
||||
K (is_pg_cap cap \<and> empty_refs m \<and> asid \<le> mask asid_bits \<and> asid \<noteq> 0) and
|
||||
(\<lambda>s. \<exists>sl. cte_wp_at (parent_for_refs m) sl s) and
|
||||
(\<lambda>s. \<exists>pd. pd_at_asid asid pd s)"
|
||||
|
||||
lemma set_cap_valid_page_map_inv:
|
||||
"\<lbrace>valid_page_inv (ArchInvocation_A.page_invocation.PageMap asid cap slot m)\<rbrace> set_cap cap slot \<lbrace>\<lambda>rv. valid_page_map_inv asid cap slot m\<rbrace>"
|
||||
apply (simp add: valid_page_inv_def valid_page_map_inv_def)
|
||||
apply (wp set_cap_cte_wp_at_cases hoare_vcg_ex_lift| simp)+
|
||||
apply (simp_all)
|
||||
apply clarsimp
|
||||
apply (rule conjI, fastforce simp: cte_wp_at_def)
|
||||
apply (rule_tac x=a in exI, rule_tac x=b in exI)
|
||||
apply (subgoal_tac "(a,b) \<noteq> slot")
|
||||
apply clarsimp
|
||||
apply (clarsimp simp: cte_wp_at_def parent_for_refs_def)
|
||||
apply (auto simp: is_pt_cap_def is_pg_cap_def is_pd_cap_def split: sum.splits)
|
||||
done
|
||||
|
||||
lemma duplicate_address_set_simp:
|
||||
"\<lbrakk>koTypeOf m \<noteq> ArchT PDET; koTypeOf m \<noteq> ArchT PTET \<rbrakk>
|
||||
\<Longrightarrow> p && ~~ mask (vs_ptr_align m) = p"
|
||||
by (auto simp:vs_ptr_align_def koTypeOf_def
|
||||
split:kernel_object.splits arch_kernel_object.splits)+
|
||||
|
||||
lemma valid_duplicates'_non_pd_pt_I:
|
||||
"\<lbrakk>koTypeOf ko \<noteq> ArchT PDET; koTypeOf ko \<noteq> ArchT PTET;
|
||||
vs_valid_duplicates' (ksPSpace s) ; ksPSpace s p = Some ko; koTypeOf ko = koTypeOf m\<rbrakk>
|
||||
\<Longrightarrow> vs_valid_duplicates' (ksPSpace s(p \<mapsto> m))"
|
||||
apply (subst vs_valid_duplicates'_def)
|
||||
apply (intro allI impI)
|
||||
apply (clarsimp split:if_splits simp:duplicate_address_set_simp option.splits)
|
||||
apply (intro conjI impI allI)
|
||||
apply (frule_tac p = x and p' = p in valid_duplicates'_D)
|
||||
apply assumption
|
||||
apply simp
|
||||
apply (simp add:duplicate_address_set_simp)+
|
||||
apply (drule_tac m = "ksPSpace s"
|
||||
and p = x in valid_duplicates'_D)
|
||||
apply simp+
|
||||
done
|
||||
|
||||
lemma setCTE_valid_duplicates'[wp]:
|
||||
"\<lbrace>\<lambda>s. vs_valid_duplicates' (ksPSpace s)\<rbrace>
|
||||
setCTE p cte \<lbrace>\<lambda>rv s. vs_valid_duplicates' (ksPSpace s)\<rbrace>"
|
||||
apply (simp add:setCTE_def)
|
||||
apply (clarsimp simp: setObject_def split_def valid_def in_monad
|
||||
projectKOs pspace_aligned'_def ps_clear_upd'
|
||||
objBits_def[symmetric] lookupAround2_char1
|
||||
split: split_if_asm)
|
||||
apply (frule pspace_storable_class.updateObject_type[where v = cte,simplified])
|
||||
apply (clarsimp simp:ObjectInstances_H.updateObject_cte assert_def bind_def
|
||||
alignCheck_def in_monad when_def alignError_def magnitudeCheck_def
|
||||
assert_opt_def return_def fail_def typeError_def
|
||||
split:if_splits option.splits Structures_H.kernel_object.splits)
|
||||
apply (erule valid_duplicates'_non_pd_pt_I[rotated 3],simp+)+
|
||||
done
|
||||
crunch valid_duplicates'[wp]: updateCap "\<lambda>s. vs_valid_duplicates' (ksPSpace s)"
|
||||
(wp: crunch_wps
|
||||
simp: crunch_simps unless_def
|
||||
ignore: getObject setObject)
|
||||
|
||||
lemma perform_page_corres:
|
||||
assumes "page_invocation_map pi pi'"
|
||||
shows "corres dc (invs and valid_etcbs and valid_page_inv pi)
|
||||
(valid_objs' and pspace_aligned' and pspace_distinct' and no_0_obj'
|
||||
and cur_tcb' and valid_arch_state' and valid_page_inv' pi')
|
||||
(invs' and valid_page_inv' pi' and (\<lambda>s. vs_valid_duplicates' (ksPSpace s)))
|
||||
(perform_page_invocation pi) (performPageInvocation pi')"
|
||||
proof -
|
||||
have pull_out_P:
|
||||
|
@ -1975,170 +2096,302 @@ proof -
|
|||
show ?thesis
|
||||
using assms
|
||||
apply (cases pi)
|
||||
apply (clarsimp simp: perform_page_invocation_def performPageInvocation_def
|
||||
page_invocation_map_def)
|
||||
apply (rule corres_guard_imp)
|
||||
apply (rule_tac R="\<lambda>_. pspace_aligned and valid_arch_objs and valid_slots sum
|
||||
and K (empty_refs sum) and valid_etcbs
|
||||
and valid_arch_state and same_refs sum cap"
|
||||
and R'="\<lambda>_. valid_slots' m' and pspace_aligned' and valid_slots_duplicated' m'
|
||||
and pspace_distinct'" in corres_split)
|
||||
prefer 2
|
||||
apply (erule updateCap_same_master)
|
||||
apply (case_tac sum, case_tac aa)
|
||||
apply (clarsimp simp: mapping_map_def valid_slots'_def valid_slots_def neq_Nil_conv)
|
||||
apply (rule corres_name_pre)
|
||||
apply (clarsimp simp:mapM_Cons bind_assoc split del:if_splits)
|
||||
apply (rule corres_guard_imp)
|
||||
apply (rule corres_split[OF _ store_pte_corres'])
|
||||
apply (rule corres_split[where r' = dc,OF corres_machine_op[OF corres_Id]])
|
||||
apply (simp add: last_byte_pte_def objBits_simps archObjSize_def)
|
||||
apply (clarsimp simp: perform_page_invocation_def performPageInvocation_def
|
||||
page_invocation_map_def)
|
||||
apply (rule corres_guard_imp)
|
||||
apply (rule_tac R="\<lambda>_. invs and (valid_page_map_inv word cap (a,b) sum) and valid_etcbs and (\<lambda>s. caps_of_state s (a,b) = Some cap)"
|
||||
and R'="\<lambda>_. invs' and valid_slots' m' and pspace_aligned' and valid_slots_duplicated' m'
|
||||
and pspace_distinct' and (\<lambda>s. vs_valid_duplicates' (ksPSpace s))" in corres_split)
|
||||
prefer 2
|
||||
apply (erule updateCap_same_master)
|
||||
apply (case_tac sum, case_tac aa)
|
||||
apply (clarsimp simp: mapping_map_def valid_slots'_def valid_slots_def valid_page_inv_def neq_Nil_conv valid_page_map_inv_def)
|
||||
apply (rule corres_name_pre)
|
||||
apply (clarsimp simp:mapM_Cons bind_assoc split del:if_splits)
|
||||
apply (rule corres_guard_imp)
|
||||
apply (rule corres_split[OF _ pte_check_if_mapped_corres])
|
||||
apply (rule corres_split[OF _ store_pte_corres'])
|
||||
apply (rule corres_split[where r' = dc, OF _ corres_store_pte_with_invalid_tail])
|
||||
apply (rule corres_split[where r'=dc, OF _ corres_machine_op[OF corres_Id]])
|
||||
apply (clarsimp simp add: when_def)
|
||||
apply (rule invalidate_tlb_by_asid_corres_ex)
|
||||
apply (simp add: last_byte_pte_def objBits_simps archObjSize_def)
|
||||
apply simp
|
||||
apply (rule no_fail_cleanCacheRange_PoU)
|
||||
apply (wp hoare_vcg_ex_lift)
|
||||
apply (clarsimp simp:pte_relation_aligned_def)
|
||||
apply (clarsimp dest!:valid_slots_duplicated_pteD') (*
|
||||
apply (wp mapM_UNIV_wp | simp add: swp_def del: fun_upd_apply)+ *)
|
||||
apply (rule_tac Q="\<lambda>_. K (word \<le> mask asid_bits \<and> word \<noteq> 0) and invs and (\<lambda>s. \<exists>pd. pd_at_asid word pd s)" in hoare_strengthen_post)
|
||||
prefer 2
|
||||
apply auto[1]
|
||||
apply (wp mapM_swp_store_pte_invs[where pte="ARM_Structs_A.pte.InvalidPTE", simplified] hoare_vcg_ex_lift)
|
||||
apply (wp mapM_UNIV_wp | simp add: swp_def del: fun_upd_apply)+
|
||||
apply (clarsimp simp:pte_relation_aligned_def)
|
||||
apply (clarsimp dest!:valid_slots_duplicated_pteD')
|
||||
apply (clarsimp simp del: fun_upd_apply simp add: cte_wp_at_caps_of_state)
|
||||
apply (wp_trace add: hoare_vcg_const_Ball_lift store_pte_typ_at store_pte_cte_wp_at hoare_vcg_ex_lift)
|
||||
apply (wp | simp add: pteCheckIfMapped_def)+
|
||||
apply (clarsimp simp add: cte_wp_at_caps_of_state valid_slots_def parent_for_refs_def empty_refs_def invs_psp_aligned simp del: fun_upd_apply)
|
||||
apply (rule conjI)
|
||||
apply (rule_tac x=aa in exI, rule_tac x=ba in exI)
|
||||
apply (clarsimp simp: is_pt_cap_def cap_asid_def image_def neq_Nil_conv Collect_disj_eq split: Structures_A.cap.splits arch_cap.splits option.splits)
|
||||
apply (rule conjI)
|
||||
apply clarsimp
|
||||
apply (drule same_refs_lD)
|
||||
apply (rule_tac x=a in exI, rule_tac x=b in exI)
|
||||
apply clarify
|
||||
apply (drule_tac x=refa in spec)
|
||||
apply clarsimp
|
||||
apply (rule conjI[rotated], fastforce)
|
||||
apply auto[1]
|
||||
apply clarsimp
|
||||
apply (case_tac ba)
|
||||
apply (clarsimp simp: mapping_map_def valid_slots_def valid_slots'_def neq_Nil_conv valid_page_inv_def valid_page_map_inv_def)
|
||||
apply (rule corres_name_pre)
|
||||
apply (clarsimp simp:mapM_Cons bind_assoc split del:if_splits)
|
||||
apply (rule corres_guard_imp)
|
||||
apply (rule corres_split[OF _ pde_check_if_mapped_corres])
|
||||
apply (rule corres_split[OF _ store_pde_corres'])
|
||||
apply (rule corres_split[where r'=dc, OF _ corres_store_pde_with_invalid_tail])
|
||||
apply (rule corres_split[where r'=dc,OF _ corres_machine_op[OF corres_Id]])
|
||||
apply (clarsimp simp: when_def)
|
||||
apply (rule invalidate_tlb_by_asid_corres_ex)
|
||||
apply (simp add: last_byte_pde_def objBits_simps archObjSize_def)
|
||||
apply simp
|
||||
apply (rule no_fail_cleanCacheRange_PoU)
|
||||
apply (wp hoare_vcg_ex_lift)
|
||||
apply (clarsimp simp: pde_relation_aligned_def)
|
||||
apply (clarsimp dest!:valid_slots_duplicated_pdeD' )
|
||||
apply (rule_tac Q="\<lambda>_. K (word \<le> mask asid_bits \<and> word \<noteq> 0) and invs and (\<lambda>s. \<exists>pd. pd_at_asid word pd s)" in hoare_strengthen_post)
|
||||
prefer 2
|
||||
apply auto[1]
|
||||
apply (wp mapM_swp_store_pde_invs_unmap[where pde="ARM_Structs_A.pde.InvalidPDE", simplified] hoare_vcg_ex_lift)
|
||||
apply (wp mapM_UNIV_wp store_pde_pd_at_asid | clarsimp simp add: swp_def del: fun_upd_apply)+
|
||||
apply (clarsimp simp: pde_relation_aligned_def)
|
||||
apply (clarsimp dest!:valid_slots_duplicated_pdeD')
|
||||
apply (clarsimp simp add: cte_wp_at_caps_of_state simp del: fun_upd_apply)
|
||||
apply (wp_trace hoare_vcg_const_Ball_lift store_pde_typ_at hoare_vcg_ex_lift store_pde_pd_at_asid)
|
||||
apply (rule hoare_vcg_conj_lift)
|
||||
apply (rule_tac slots="y # ys" in store_pde_invs_unmap')
|
||||
apply (wp hoare_vcg_const_Ball_lift store_pde_pd_at_asid hoare_vcg_ex_lift)
|
||||
apply (wp | simp add: pdeCheckIfMapped_def)+
|
||||
apply (clarsimp simp add: cte_wp_at_caps_of_state valid_slots_def parent_for_refs_def empty_refs_def invs_psp_aligned simp del: fun_upd_apply)
|
||||
apply (rule conjI, rule_tac x=ref in exI, clarsimp)
|
||||
apply (rule conjI)
|
||||
apply (rule_tac x=aa in exI, rule_tac x=ba in exI)
|
||||
apply auto[1]
|
||||
apply (rule conjI)
|
||||
apply (rule_tac x=a in exI, rule_tac x=b in exI, auto simp: same_refs_def)[1]
|
||||
apply (rule conjI)
|
||||
apply (clarsimp simp: pde_at_def obj_at_def
|
||||
caps_of_state_cteD'[where P=\<top>, simplified])
|
||||
apply (drule_tac cap=capa and ptr="(aa,ba)"
|
||||
in valid_global_refsD[OF invs_valid_global_refs])
|
||||
apply assumption+
|
||||
apply (clarsimp simp: cap_range_def)
|
||||
apply (rule conjI)
|
||||
apply (clarsimp simp: pde_at_def obj_at_def a_type_def)
|
||||
apply (clarsimp split: Structures_A.kernel_object.split_asm split_if_asm ARM_Structs_A.arch_kernel_obj.splits)
|
||||
apply (rule conjI[rotated], fastforce)
|
||||
apply (erule ballEI)
|
||||
apply (clarsimp simp: pde_at_def obj_at_def
|
||||
caps_of_state_cteD'[where P=\<top>, simplified])
|
||||
apply (drule_tac cap=capa and ptr="(aa,ba)"
|
||||
in valid_global_refsD[OF invs_valid_global_refs])
|
||||
apply assumption+
|
||||
apply (drule_tac x=x in imageI[where f="\<lambda>x. x && ~~ mask pd_bits"])
|
||||
apply (drule (1) subsetD)
|
||||
apply (clarsimp simp: cap_range_def)
|
||||
apply clarsimp
|
||||
apply (wp_trace arch_update_cap_invs_map set_cap_valid_page_map_inv)
|
||||
apply (wp arch_update_updateCap_invs)
|
||||
apply (clarsimp simp: invs_valid_objs invs_psp_aligned invs_distinct valid_page_inv_def cte_wp_at_caps_of_state is_arch_update_def is_cap_simps)
|
||||
apply (simp add: cap_master_cap_def split: cap.splits arch_cap.splits)
|
||||
apply (auto simp: cte_wp_at_ctes_of valid_page_inv'_def)[1]
|
||||
-- "PageRemap"
|
||||
apply (clarsimp simp: perform_page_invocation_def performPageInvocation_def
|
||||
page_invocation_map_def)
|
||||
apply (case_tac sum)
|
||||
apply simp
|
||||
apply (case_tac a, simp)
|
||||
apply (clarsimp simp: mapping_map_def)
|
||||
apply (rule corres_name_pre)
|
||||
apply (clarsimp simp:mapM_Cons mapM_x_mapM bind_assoc valid_slots_def valid_page_inv_def
|
||||
neq_Nil_conv split del:if_splits )
|
||||
apply (rule corres_guard_imp)
|
||||
apply (rule corres_split[OF _ pte_check_if_mapped_corres])
|
||||
apply (rule corres_split[OF _ store_pte_corres'])
|
||||
apply (rule corres_split[OF _ corres_store_pte_with_invalid_tail])
|
||||
apply (rule corres_split[where r' = dc,OF _ corres_machine_op[OF corres_Id]])
|
||||
apply (clarsimp simp: when_def)
|
||||
apply (rule invalidate_tlb_by_asid_corres_ex)
|
||||
apply (simp add: last_byte_pte_def objBits_simps archObjSize_def)
|
||||
apply simp
|
||||
apply (rule no_fail_cleanCacheRange_PoU)
|
||||
apply (wp hoare_vcg_ex_lift)
|
||||
apply (clarsimp simp:valid_page_inv'_def)
|
||||
apply (clarsimp dest!:valid_slots_duplicated_pteD')
|
||||
apply (rule_tac Q="\<lambda>_. K (word \<le> mask asid_bits \<and> word \<noteq> 0) and invs and (\<lambda>s. \<exists>pd. pd_at_asid word pd s)" in hoare_strengthen_post)
|
||||
prefer 2
|
||||
apply auto[1]
|
||||
apply (wp mapM_swp_store_pte_invs[where pte="ARM_Structs_A.pte.InvalidPTE", simplified] hoare_vcg_ex_lift)
|
||||
apply (wp mapM_UNIV_wp | simp add: swp_def del: fun_upd_apply)+
|
||||
apply (clarsimp simp:pte_relation_aligned_def valid_page_inv'_def)
|
||||
apply (clarsimp dest!:valid_slots_duplicated_pteD')
|
||||
apply (clarsimp simp del: fun_upd_apply simp add: cte_wp_at_caps_of_state)
|
||||
apply (wp_trace add: hoare_vcg_const_Ball_lift store_pte_typ_at store_pte_cte_wp_at hoare_vcg_ex_lift)
|
||||
apply (wp | simp add: pteCheckIfMapped_def)+
|
||||
apply (clarsimp simp add: cte_wp_at_caps_of_state valid_slots_def parent_for_refs_def empty_refs_def invs_psp_aligned is_cap_simps simp del: fun_upd_apply)
|
||||
apply (rule conjI)
|
||||
apply fastforce
|
||||
apply (rule conjI)
|
||||
apply clarsimp
|
||||
apply (rule_tac x=ac in exI, rule_tac x=bb in exI, rule_tac x=capa in exI)
|
||||
apply (drule same_refs_lD)
|
||||
apply clarsimp
|
||||
apply (erule (2) ref_is_unique[OF _ _ reachable_page_table_not_global])
|
||||
apply (simp_all add: invs_def valid_state_def valid_arch_state_def valid_arch_caps_def valid_pspace_def valid_objs_caps)[9]
|
||||
apply fastforce
|
||||
apply auto[1]
|
||||
apply simp
|
||||
apply (case_tac b)
|
||||
apply (rule corres_name_pre)
|
||||
apply (clarsimp simp: mapping_map_def valid_page_inv_def
|
||||
mapM_x_mapM mapM_Cons bind_assoc
|
||||
valid_slots_def neq_Nil_conv split del:if_splits)
|
||||
apply (rule corres_guard_imp)
|
||||
apply (rule corres_split[OF _ pde_check_if_mapped_corres])
|
||||
apply (rule corres_split[OF _ store_pde_corres'])
|
||||
apply (rule corres_split[where r'=dc, OF _ corres_store_pde_with_invalid_tail])
|
||||
apply (rule corres_split[where r'=dc,OF _ corres_machine_op[OF corres_Id]])
|
||||
apply (clarsimp simp: when_def)
|
||||
apply (rule invalidate_tlb_by_asid_corres_ex)
|
||||
apply (simp add: last_byte_pde_def objBits_simps archObjSize_def)
|
||||
apply simp
|
||||
apply (rule no_fail_cleanCacheRange_PoU)
|
||||
apply (rule corres_store_pte_with_invalid_tail)
|
||||
apply (clarsimp dest!:valid_slots_duplicated_pteD')
|
||||
apply wp
|
||||
apply (clarsimp simp:pte_relation_aligned_def)
|
||||
apply (clarsimp dest!:valid_slots_duplicated_pteD')
|
||||
apply (wp hoare_vcg_const_Ball_lift store_pte_typ_at)
|
||||
apply simp
|
||||
apply simp
|
||||
apply (case_tac ba)
|
||||
apply (clarsimp simp: mapping_map_def valid_slots_def valid_slots'_def neq_Nil_conv)
|
||||
apply (rule corres_name_pre)
|
||||
apply (clarsimp simp:mapM_Cons bind_assoc split del:if_splits)
|
||||
apply (rule corres_guard_imp)
|
||||
apply (rule corres_split[OF _ store_pde_corres'])
|
||||
apply (rule corres_split[where r'=dc,OF corres_machine_op[OF corres_Id]])
|
||||
apply (simp add: last_byte_pde_def objBits_simps archObjSize_def)
|
||||
apply simp
|
||||
apply (rule no_fail_cleanCacheRange_PoU)
|
||||
apply (rule corres_store_pde_with_invalid_tail)
|
||||
apply (clarsimp simp: pde_relation_aligned_def)
|
||||
apply (clarsimp dest!:valid_slots_duplicated_pdeD' )
|
||||
apply wp
|
||||
apply (clarsimp simp: pde_relation_aligned_def)
|
||||
apply (clarsimp dest!:valid_slots_duplicated_pdeD')
|
||||
apply (wp hoare_vcg_const_Ball_lift valid_pde_lift')
|
||||
apply simp
|
||||
apply simp
|
||||
apply (wp | simp add:empty_refs_def same_refs_def)+
|
||||
apply (clarsimp simp: valid_page_inv_def same_refs_def empty_refs_def)
|
||||
apply (clarsimp simp: cte_wp_at_caps_of_state is_arch_update_def is_cap_simps)
|
||||
apply (simp add: cap_master_cap_def split: cap.splits arch_cap.splits)
|
||||
apply auto[1]
|
||||
apply (clarsimp simp: cte_wp_at_ctes_of valid_page_inv'_def)
|
||||
apply (clarsimp simp: perform_page_invocation_def performPageInvocation_def
|
||||
apply (wp hoare_vcg_ex_lift)
|
||||
apply (clarsimp simp: pde_relation_aligned_def valid_page_inv'_def)
|
||||
apply (clarsimp dest!:valid_slots_duplicated_pdeD' )
|
||||
apply (rule_tac Q="\<lambda>_. K (word \<le> mask asid_bits \<and> word \<noteq> 0) and invs and (\<lambda>s. \<exists>pd. pd_at_asid word pd s)" in hoare_strengthen_post)
|
||||
prefer 2
|
||||
apply auto[1]
|
||||
apply (wp mapM_swp_store_pde_invs_unmap[where pde="ARM_Structs_A.pde.InvalidPDE", simplified] hoare_vcg_ex_lift)
|
||||
apply (wp mapM_UNIV_wp store_pde_pd_at_asid | clarsimp simp add: swp_def del: fun_upd_apply)+
|
||||
apply (clarsimp simp: pde_relation_aligned_def valid_page_inv'_def)
|
||||
apply (clarsimp dest!:valid_slots_duplicated_pdeD')
|
||||
apply (clarsimp simp add: cte_wp_at_caps_of_state simp del: fun_upd_apply)
|
||||
apply (wp_trace hoare_vcg_const_Ball_lift store_pde_typ_at hoare_vcg_ex_lift store_pde_pd_at_asid)
|
||||
apply (rule hoare_vcg_conj_lift)
|
||||
apply (rule_tac slots="y # ys" in store_pde_invs_unmap')
|
||||
apply (wp hoare_vcg_const_Ball_lift store_pde_pd_at_asid hoare_vcg_ex_lift)
|
||||
apply (wp | simp add: pdeCheckIfMapped_def)+
|
||||
apply (clarsimp simp add: cte_wp_at_caps_of_state valid_slots_def parent_for_refs_def empty_refs_def invs_psp_aligned simp del: fun_upd_apply)
|
||||
apply (rule conjI, rule_tac x=ref in exI, clarsimp)
|
||||
apply (frule valid_global_refsD2)
|
||||
apply (clarsimp simp: cap_range_def)+
|
||||
apply (rule conjI)
|
||||
apply (rule_tac x=ac in exI, rule_tac x=bb in exI, rule_tac x=cap in exI)
|
||||
apply (erule conjI)
|
||||
apply clarsimp
|
||||
apply (rule conjI)
|
||||
apply (rule_tac x=ad in exI, rule_tac x=bc in exI, rule_tac x=capa in exI)
|
||||
apply (clarsimp simp: same_refs_def pde_ref_def pde_ref_pages_def valid_pde_def invs_def valid_state_def valid_pspace_def)
|
||||
apply (drule valid_objs_caps)
|
||||
apply (clarsimp simp: valid_caps_def)
|
||||
apply (drule spec, drule spec, drule_tac x=capa in spec, drule (1) mp)
|
||||
apply (case_tac aa, simp_all)
|
||||
apply ((clarsimp simp: valid_cap_def obj_at_def a_type_def is_ep_def
|
||||
is_aep_def is_cap_table_def is_tcb_def
|
||||
is_pg_cap_def
|
||||
split: cap.splits Structures_A.kernel_object.splits
|
||||
split_if_asm
|
||||
ARM_Structs_A.arch_kernel_obj.splits option.splits
|
||||
arch_cap.splits)+)[2]
|
||||
apply (clarsimp simp: pde_at_def obj_at_def a_type_def)
|
||||
apply (case_tac ko, (simp_all split: split_if_asm))
|
||||
apply (case_tac arch_kernel_obj, simp_all)
|
||||
apply (rule conjI)
|
||||
apply clarsimp
|
||||
apply (drule_tac ptr="(ac,bb)" in
|
||||
valid_global_refsD[OF invs_valid_global_refs caps_of_state_cteD])
|
||||
apply simp+
|
||||
apply force
|
||||
apply (rule conjI[rotated], fastforce)
|
||||
apply (erule ballEI)
|
||||
apply clarsimp
|
||||
apply (drule_tac ptr="(ac,bb)" in
|
||||
valid_global_refsD[OF invs_valid_global_refs caps_of_state_cteD])
|
||||
apply simp+
|
||||
apply force
|
||||
apply auto[1]
|
||||
apply (clarsimp simp: performPageInvocation_def perform_page_invocation_def
|
||||
page_invocation_map_def)
|
||||
apply (case_tac sum)
|
||||
apply simp
|
||||
apply (case_tac a, simp)
|
||||
apply (clarsimp simp: mapping_map_def)
|
||||
apply (rule corres_name_pre)
|
||||
apply (clarsimp simp:mapM_Cons mapM_x_mapM bind_assoc valid_slots_def valid_page_inv_def
|
||||
neq_Nil_conv split del:if_splits )
|
||||
apply (rule corres_assume_pre)
|
||||
apply (clarsimp simp: valid_page_inv_def valid_page_inv'_def isCap_simps is_page_cap_def cong: option.case_cong prod.case_cong)
|
||||
apply (case_tac m)
|
||||
apply simp
|
||||
apply (rule corres_guard_imp)
|
||||
apply (rule corres_split [where r'="acap_relation"])
|
||||
prefer 2
|
||||
apply simp
|
||||
apply (rule corres_rel_imp)
|
||||
apply (rule get_cap_corres_all_rights_P[where P=is_arch_cap], rule refl)
|
||||
apply (clarsimp simp: is_cap_simps)
|
||||
apply (rule_tac F="is_page_cap cap" in corres_gen_asm)
|
||||
apply (rule updateCap_same_master)
|
||||
apply (clarsimp simp: is_page_cap_def update_map_data_def)
|
||||
apply (wp get_cap_wp getSlotCap_wp)
|
||||
apply (clarsimp simp: cte_wp_at_caps_of_state is_arch_diminished_def)
|
||||
apply (drule (2) diminished_is_update')+
|
||||
apply (clarsimp simp: cap_rights_update_def acap_rights_update_def update_map_data_def is_cap_simps)
|
||||
apply auto[1]
|
||||
apply (auto simp: cte_wp_at_ctes_of)[1]
|
||||
apply clarsimp
|
||||
apply (rule corres_guard_imp)
|
||||
apply (rule corres_split[OF _ store_pte_corres'])
|
||||
apply (rule corres_split[where r' = dc,OF corres_machine_op[OF corres_Id]])
|
||||
apply (simp add: last_byte_pte_def objBits_simps archObjSize_def)
|
||||
apply simp
|
||||
apply (rule no_fail_cleanCacheRange_PoU)
|
||||
apply (rule corres_store_pte_with_invalid_tail)
|
||||
apply (clarsimp simp:valid_page_inv'_def)
|
||||
apply (clarsimp dest!:valid_slots_duplicated_pteD')
|
||||
apply wp
|
||||
apply (clarsimp simp: valid_page_inv'_def pte_relation_aligned_def)
|
||||
apply (clarsimp dest!:valid_slots_duplicated_pteD')
|
||||
apply (wp hoare_vcg_const_Ball_lift store_pte_typ_at)
|
||||
apply fastforce
|
||||
apply simp
|
||||
apply (case_tac b)
|
||||
apply (rule corres_name_pre)
|
||||
apply (clarsimp simp: mapping_map_def valid_page_inv_def
|
||||
mapM_x_mapM mapM_Cons bind_assoc
|
||||
valid_slots_def neq_Nil_conv split del:if_splits)
|
||||
apply (rule corres_guard_imp)
|
||||
apply (rule corres_split[OF _ store_pde_corres'])
|
||||
apply (rule corres_split[where r' = dc,OF corres_machine_op[OF corres_Id]])
|
||||
apply (simp add: last_byte_pde_def objBits_simps archObjSize_def)
|
||||
apply simp
|
||||
apply (rule no_fail_cleanCacheRange_PoU)
|
||||
apply (rule corres_store_pde_with_invalid_tail)
|
||||
apply (clarsimp simp:valid_page_inv'_def pde_relation_aligned_def)
|
||||
apply (clarsimp dest!:valid_slots_duplicated_pdeD')
|
||||
apply wp
|
||||
apply (clarsimp simp:valid_page_inv'_def pde_relation_aligned_def)
|
||||
apply (clarsimp dest!:valid_slots_duplicated_pdeD')
|
||||
apply (wp hoare_vcg_const_Ball_lift store_pte_typ_at)
|
||||
apply fastforce
|
||||
apply simp
|
||||
apply (clarsimp simp: performPageInvocation_def perform_page_invocation_def
|
||||
page_invocation_map_def)
|
||||
apply (rule corres_assume_pre)
|
||||
apply (clarsimp simp: valid_page_inv_def valid_page_inv'_def isCap_simps is_page_cap_def cong: option.case_cong prod.case_cong)
|
||||
apply (case_tac m)
|
||||
apply simp
|
||||
apply (rule corres_guard_imp)
|
||||
apply (rule corres_split [where r'="acap_relation"])
|
||||
prefer 2
|
||||
apply simp
|
||||
apply (rule corres_rel_imp)
|
||||
apply (rule get_cap_corres_all_rights_P[where P=is_arch_cap], rule refl)
|
||||
apply (clarsimp simp: is_cap_simps)
|
||||
apply (rule_tac F="is_page_cap cap" in corres_gen_asm)
|
||||
apply (rule updateCap_same_master)
|
||||
apply (clarsimp simp: is_page_cap_def update_map_data_def)
|
||||
apply (wp get_cap_wp getSlotCap_wp)
|
||||
apply (clarsimp simp: cte_wp_at_caps_of_state is_arch_diminished_def)
|
||||
apply (drule (2) diminished_is_update')+
|
||||
apply (clarsimp simp: cap_rights_update_def acap_rights_update_def update_map_data_def is_cap_simps)
|
||||
apply auto[1]
|
||||
apply (clarsimp simp: cte_wp_at_ctes_of)
|
||||
apply clarsimp
|
||||
apply (rule corres_guard_imp)
|
||||
apply (rule corres_split)
|
||||
prefer 2
|
||||
apply (rule unmap_page_corres)
|
||||
apply (rule corres_split [where r'=acap_relation])
|
||||
apply (rule corres_split)
|
||||
prefer 2
|
||||
apply simp
|
||||
apply (rule corres_rel_imp)
|
||||
apply (rule get_cap_corres_all_rights_P[where P=is_arch_cap], rule refl)
|
||||
apply (clarsimp simp: is_cap_simps)
|
||||
apply (rule_tac F="is_page_cap cap" in corres_gen_asm)
|
||||
apply (rule updateCap_same_master)
|
||||
apply (clarsimp simp: is_page_cap_def update_map_data_def)
|
||||
apply (wp get_cap_wp getSlotCap_wp)
|
||||
apply (simp add: cte_wp_at_caps_of_state)
|
||||
apply (strengthen pull_out_P)+
|
||||
apply wp
|
||||
apply (simp add: cte_wp_at_ctes_of)
|
||||
apply wp
|
||||
apply (clarsimp simp: valid_unmap_def cte_wp_at_caps_of_state)
|
||||
apply (clarsimp simp: is_arch_diminished_def is_cap_simps split: cap.splits arch_cap.splits)
|
||||
apply (drule (2) diminished_is_update')+
|
||||
apply (clarsimp simp: cap_rights_update_def is_page_cap_def cap_master_cap_simps update_map_data_def acap_rights_update_def)
|
||||
apply (clarsimp simp add: valid_cap_def mask_def)
|
||||
apply auto[1]
|
||||
apply (clarsimp simp: cte_wp_at_ctes_of)
|
||||
apply (clarsimp simp: performPageInvocation_def perform_page_invocation_def
|
||||
page_invocation_map_def)
|
||||
apply (rule corres_guard_imp)
|
||||
apply (rule corres_when, simp)
|
||||
apply (rule corres_split [OF _ set_vm_root_for_flush_corres])
|
||||
apply (rule corres_split [OF _ corres_machine_op])
|
||||
prefer 2
|
||||
apply (rule do_flush_corres)
|
||||
apply (rule corres_when, simp)
|
||||
apply (rule corres_split [OF _ gct_corres])
|
||||
apply simp
|
||||
apply (rule set_vm_root_corres)
|
||||
apply wp
|
||||
apply (simp add: cur_tcb_def [symmetric] cur_tcb'_def [symmetric])
|
||||
apply (wp hoare_drop_imps)
|
||||
apply (simp add: cur_tcb_def [symmetric] cur_tcb'_def [symmetric])
|
||||
apply (wp hoare_drop_imps)
|
||||
apply (auto simp: valid_page_inv_def)
|
||||
apply (rule unmap_page_corres)
|
||||
apply (rule corres_split [where r'=acap_relation])
|
||||
prefer 2
|
||||
apply simp
|
||||
apply (rule corres_rel_imp)
|
||||
apply (rule get_cap_corres_all_rights_P[where P=is_arch_cap], rule refl)
|
||||
apply (clarsimp simp: is_cap_simps)
|
||||
apply (rule_tac F="is_page_cap cap" in corres_gen_asm)
|
||||
apply (rule updateCap_same_master)
|
||||
apply (clarsimp simp: is_page_cap_def update_map_data_def)
|
||||
apply (wp get_cap_wp getSlotCap_wp)
|
||||
apply (simp add: cte_wp_at_caps_of_state)
|
||||
apply (strengthen pull_out_P)+
|
||||
apply wp
|
||||
apply (simp add: cte_wp_at_ctes_of)
|
||||
apply wp
|
||||
apply (clarsimp simp: valid_unmap_def cte_wp_at_caps_of_state)
|
||||
apply (clarsimp simp: is_arch_diminished_def is_cap_simps split: cap.splits arch_cap.splits)
|
||||
apply (drule (2) diminished_is_update')+
|
||||
apply (clarsimp simp: cap_rights_update_def is_page_cap_def cap_master_cap_simps update_map_data_def acap_rights_update_def)
|
||||
apply (clarsimp simp add: valid_cap_def mask_def)
|
||||
apply auto[1]
|
||||
apply (auto simp: cte_wp_at_ctes_of)[1]
|
||||
apply (clarsimp simp: performPageInvocation_def perform_page_invocation_def
|
||||
page_invocation_map_def)
|
||||
apply (rule corres_guard_imp)
|
||||
apply (rule corres_when, simp)
|
||||
apply (rule corres_split [OF _ set_vm_root_for_flush_corres])
|
||||
apply (rule corres_split [OF _ corres_machine_op])
|
||||
prefer 2
|
||||
apply (rule do_flush_corres)
|
||||
apply (rule corres_when, simp)
|
||||
apply (rule corres_split [OF _ gct_corres])
|
||||
apply simp
|
||||
apply (rule set_vm_root_corres)
|
||||
apply wp
|
||||
apply (simp add: cur_tcb_def [symmetric] cur_tcb'_def [symmetric])
|
||||
apply (wp hoare_drop_imps)
|
||||
apply (simp add: cur_tcb_def [symmetric] cur_tcb'_def [symmetric])
|
||||
apply (wp hoare_drop_imps)
|
||||
apply (auto simp: valid_page_inv_def)
|
||||
done
|
||||
qed
|
||||
|
||||
|
@ -3263,6 +3516,15 @@ lemma unmapPage_invs' [wp]:
|
|||
|
||||
crunch (no_irq) no_irq[wp]: doFlush
|
||||
|
||||
crunch invs'[wp]: pteCheckIfMapped, pdeCheckIfMapped "invs'"
|
||||
(ignore: getObject)
|
||||
|
||||
crunch valid_pte'[wp]: pteCheckIfMapped, pdeCheckIfMapped "valid_pte' pte"
|
||||
(ignore: getObject)
|
||||
|
||||
crunch valid_pde'[wp]: pteCheckIfMapped, pdeCheckIfMapped "valid_pde' pde"
|
||||
(ignore: getObject)
|
||||
|
||||
lemma perform_pt_invs [wp]:
|
||||
"\<lbrace>invs' and valid_page_inv' pt\<rbrace> performPageInvocation pt \<lbrace>\<lambda>_. invs'\<rbrace>"
|
||||
apply (simp add: performPageInvocation_def)
|
||||
|
|
|
@ -15,125 +15,14 @@ imports
|
|||
"../../spec/capDL/Types_D"
|
||||
begin
|
||||
|
||||
(**************************************
|
||||
* Start of lemmas to move elsewhere. *
|
||||
**************************************)
|
||||
|
||||
lemma inter_empty_not_both:
|
||||
"\<lbrakk>x \<in> A; A \<inter> B = {}\<rbrakk> \<Longrightarrow> x \<notin> B"
|
||||
by fastforce
|
||||
|
||||
lemma union_intersection:
|
||||
"A \<inter> (A \<union> B) = A"
|
||||
"B \<inter> (A \<union> B) = B"
|
||||
"(A \<union> B) \<inter> A = A"
|
||||
"(A \<union> B) \<inter> B = B"
|
||||
by fastforce+
|
||||
|
||||
lemma union_intersection1: "A \<inter> (A \<union> B) = A"
|
||||
by (rule inf_sup_absorb)
|
||||
lemma union_intersection2: "B \<inter> (A \<union> B) = B"
|
||||
by fastforce
|
||||
|
||||
(* This lemma is strictly weaker than restrict_map_disj. *)
|
||||
lemma restrict_map_disj':
|
||||
"S \<inter> T = {} \<Longrightarrow> h |` S \<bottom> h' |` T"
|
||||
by (auto simp: map_disj_def restrict_map_def dom_def)
|
||||
|
||||
lemma map_add_restrict_comm:
|
||||
"S \<inter> T = {} \<Longrightarrow> h |` S ++ h' |` T = h' |` T ++ h |` S"
|
||||
apply (drule restrict_map_disj')
|
||||
apply (erule map_add_com)
|
||||
done
|
||||
|
||||
lemma restrict_map_not_self_UNIV [simp]:
|
||||
"h |` (UNIV - dom h) = empty"
|
||||
by (rule ext, clarsimp simp: restrict_map_def)
|
||||
|
||||
(************************************
|
||||
* End of lemmas to move elsewhere. *
|
||||
************************************)
|
||||
|
||||
definition
|
||||
the_set :: "cdl_component set \<Rightarrow> nat set"
|
||||
where
|
||||
"the_set xs = {x. Slot x \<in> xs}"
|
||||
|
||||
lemma in_the_set [simp]:
|
||||
"(x \<in> the_set xs) = (Slot x \<in> xs)"
|
||||
by (clarsimp simp: the_set_def)
|
||||
|
||||
lemma the_set_union [simp]:
|
||||
"the_set (A \<union> B) = the_set A \<union> the_set B"
|
||||
by (fastforce simp: the_set_def)
|
||||
|
||||
lemma the_set_inter [simp]:
|
||||
"the_set (A \<inter> B) = the_set A \<inter> the_set B"
|
||||
by (fastforce simp: the_set_def)
|
||||
|
||||
lemma the_set_inter_empty:
|
||||
"A \<inter> B = {} \<Longrightarrow> the_set A \<inter> the_set B = {}"
|
||||
by (fastforce simp: the_set_def)
|
||||
|
||||
lemma the_set_empty[simp]: "the_set {} = {}"
|
||||
by (simp add: the_set_def)
|
||||
|
||||
lemma the_set_UNIV [simp]:
|
||||
"the_set UNIV = UNIV"
|
||||
by (clarsimp simp: the_set_def)
|
||||
|
||||
lemma the_set_slot [simp]:
|
||||
"the_set (Slot ` xs) = xs"
|
||||
by (fastforce simp: the_set_def)
|
||||
|
||||
lemma the_set_singleton_slot [simp]:
|
||||
"the_set {Slot x} = {x}"
|
||||
by (clarsimp simp: the_set_def)
|
||||
|
||||
lemma the_set_singleton_fields [simp]:
|
||||
"the_set {Fields} = {}"
|
||||
by (clarsimp simp: the_set_def)
|
||||
|
||||
(* The sep_entitys are the entities we are interested in sep_capdl,
|
||||
and it includes objects and caps
|
||||
*)
|
||||
(* The sep_entitie are the entities we are interested in our lifted heap.
|
||||
* The entities are either objects without capabilities or capabilities.
|
||||
*)
|
||||
datatype sep_entity = CDL_Object cdl_object | CDL_Cap "cdl_cap option"
|
||||
|
||||
(* "Cleans" slots to conform with the compontents. *)
|
||||
definition
|
||||
clean_slots :: "cdl_cap_map \<Rightarrow> cdl_components \<Rightarrow> cdl_cap_map"
|
||||
where
|
||||
"clean_slots slots cmp \<equiv> slots |` the_set cmp"
|
||||
|
||||
(* Sets the fields of an object to a "clean" state.
|
||||
Because a frame's size is part of it's type, we don't reset it. *)
|
||||
definition
|
||||
object_clean_fields :: "cdl_object \<Rightarrow> cdl_components \<Rightarrow> cdl_object"
|
||||
where
|
||||
"object_clean_fields obj cmp \<equiv> if Fields \<in> cmp then obj else case obj of
|
||||
Tcb x \<Rightarrow> Tcb (x\<lparr>cdl_tcb_fault_endpoint := undefined,
|
||||
cdl_tcb_intent := undefined,
|
||||
cdl_tcb_has_fault := undefined,
|
||||
cdl_tcb_domain := undefined\<rparr>)
|
||||
| CNode x \<Rightarrow> CNode (x\<lparr>cdl_cnode_size_bits := undefined \<rparr>)
|
||||
| _ \<Rightarrow> obj"
|
||||
|
||||
(* Sets the slots of an object to a "clean" state. *)
|
||||
definition
|
||||
object_clean_slots :: "cdl_object \<Rightarrow> cdl_components \<Rightarrow> cdl_object"
|
||||
where
|
||||
"object_clean_slots obj cmp \<equiv> update_slots (clean_slots (object_slots obj) cmp) obj"
|
||||
|
||||
(* Sets an object to a "clean" state. *)
|
||||
definition
|
||||
object_clean :: "cdl_object \<Rightarrow> cdl_components \<Rightarrow> cdl_object"
|
||||
where
|
||||
"object_clean obj cmps \<equiv> object_clean_slots (object_clean_fields obj cmps) cmps"
|
||||
|
||||
(* The state for separation logic is an option map
|
||||
from (ptr,component) to sep_entities
|
||||
* from (obj_id,component) to sep_entities
|
||||
*)
|
||||
|
||||
type_synonym sep_state_heap = "(cdl_object_id \<times> cdl_component) \<Rightarrow> sep_entity option"
|
||||
type_synonym sep_state_irq_map = "cdl_irq \<Rightarrow> cdl_object_id option"
|
||||
|
||||
|
@ -141,6 +30,9 @@ translations
|
|||
(type) "sep_state_heap" <=(type) "32 word \<times> cdl_component \<Rightarrow> sep_entity option"
|
||||
(type) "sep_state_irq_map" <=(type) "8 word \<Rightarrow> 32 word option"
|
||||
|
||||
|
||||
(* Our lifted state contains sep_entities and the IRQ table.
|
||||
*)
|
||||
datatype sep_state =
|
||||
SepState "(cdl_object_id \<times> cdl_component) \<Rightarrow> sep_entity option"
|
||||
"cdl_irq \<Rightarrow> cdl_object_id option"
|
||||
|
@ -152,21 +44,7 @@ where "sep_heap (SepState heap irqs) = heap"
|
|||
primrec sep_irq_node :: "sep_state \<Rightarrow> sep_state_irq_map"
|
||||
where "sep_irq_node (SepState heap irqs) = irqs"
|
||||
|
||||
definition
|
||||
sep_caps :: "sep_state_heap \<Rightarrow> 32 word \<Rightarrow> cdl_cap_map"
|
||||
where
|
||||
"sep_caps s ptr \<equiv> \<lambda>slot. case s (ptr, Slot slot) of
|
||||
Some (CDL_Cap x ) \<Rightarrow> x
|
||||
| _ \<Rightarrow> None"
|
||||
|
||||
definition
|
||||
sep_objects :: "sep_state_heap \<Rightarrow> cdl_heap"
|
||||
where
|
||||
"sep_objects s \<equiv> \<lambda>ptr. case s (ptr,Fields) of
|
||||
Some (CDL_Object obj) \<Rightarrow> Some (update_slots (sep_caps s ptr) obj)
|
||||
| _ \<Rightarrow> None"
|
||||
|
||||
(* Adding states adds the objects an their ownerships.
|
||||
(* Adding states adds the separation entity heap and the IRQ table.
|
||||
*)
|
||||
definition
|
||||
sep_state_add :: "sep_state \<Rightarrow> sep_state \<Rightarrow> sep_state"
|
||||
|
@ -176,11 +54,8 @@ where
|
|||
((sep_irq_node state_a) ++ sep_irq_node state_b)"
|
||||
|
||||
|
||||
(* Heaps are disjoint if for all of their objects:
|
||||
* the caps of their respective objects are disjoint,
|
||||
* their respective objects don't conflict,
|
||||
* they don't both own any of the same fields.
|
||||
*)
|
||||
(* State are disjoint the separation entity heaps and the IRQ tables are dijoint.
|
||||
*)
|
||||
definition
|
||||
sep_state_disj :: "sep_state \<Rightarrow> sep_state \<Rightarrow> bool"
|
||||
where
|
||||
|
@ -188,224 +63,10 @@ where
|
|||
(sep_heap state_a) \<bottom> (sep_heap state_b) \<and>
|
||||
(sep_irq_node state_a) \<bottom> (sep_irq_node state_b)"
|
||||
|
||||
lemma map_add_emptyE [elim!]:
|
||||
"\<lbrakk>a ++ b = empty; \<lbrakk>a = empty; b = empty\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
|
||||
by (metis map_add_None)
|
||||
|
||||
lemma clean_slots_Fields [simp]:
|
||||
"clean_slots slots {Fields} = empty"
|
||||
by (clarsimp simp: clean_slots_def the_set_def)
|
||||
|
||||
lemma clean_slots_empty [simp]:
|
||||
"clean_slots empty cmp = empty"
|
||||
by (clarsimp simp: clean_slots_def)
|
||||
|
||||
lemma clean_slots_singleton [simp]:
|
||||
"clean_slots [slot \<mapsto> cap] {Slot slot} = [slot \<mapsto> cap]"
|
||||
by (clarsimp simp: clean_slots_def)
|
||||
|
||||
lemma object_type_update_slots [simp]:
|
||||
"object_type (update_slots slots x) = object_type x"
|
||||
by (clarsimp simp: object_type_def update_slots_def split: cdl_object.splits)
|
||||
|
||||
lemma object_type_object_clean_slots [simp]:
|
||||
"object_type (object_clean_slots x cmp) = object_type x"
|
||||
by (clarsimp simp: object_clean_slots_def)
|
||||
|
||||
lemma object_type_object_clean_fields [simp]:
|
||||
"object_type (object_clean_fields x cmp) = object_type x"
|
||||
by (clarsimp simp: object_clean_fields_def object_type_def split: cdl_object.splits)
|
||||
|
||||
lemma object_type_object_clean [simp]:
|
||||
"object_type (object_clean x cmp) = object_type x"
|
||||
by (clarsimp simp: object_clean_def)
|
||||
|
||||
lemma object_slots_empty [simp]:
|
||||
"\<not> has_slots obj \<Longrightarrow> object_slots obj = empty"
|
||||
by (clarsimp simp: object_slots_def has_slots_def split: cdl_object.splits)
|
||||
|
||||
lemma object_slots_update_slots [simp]:
|
||||
"has_slots obj \<Longrightarrow> object_slots (update_slots slots obj) = slots"
|
||||
by (clarsimp simp: object_slots_def update_slots_def has_slots_def
|
||||
split: cdl_object.splits)
|
||||
|
||||
lemma object_slots_update_slots_empty [simp]:
|
||||
"\<not>has_slots obj \<Longrightarrow> object_slots (update_slots slots obj) = empty"
|
||||
by (clarsimp simp: object_slots_def update_slots_def has_slots_def
|
||||
split: cdl_object.splits)
|
||||
|
||||
lemma update_slots_no_slots [simp]:
|
||||
"\<not>has_slots obj \<Longrightarrow> update_slots slots obj = obj"
|
||||
by (clarsimp simp: update_slots_def has_slots_def split: cdl_object.splits)
|
||||
|
||||
lemma update_slots_update_slots [simp]:
|
||||
"update_slots slots (update_slots slots' obj) = update_slots slots obj"
|
||||
by (clarsimp simp: update_slots_def split: cdl_object.splits)
|
||||
|
||||
lemma update_slots_same_object:
|
||||
"a = b \<Longrightarrow> update_slots a obj = update_slots b obj"
|
||||
by (erule arg_cong)
|
||||
|
||||
lemma update_slots_eq_slots:
|
||||
"\<lbrakk>has_slots obj; update_slots slots obj = update_slots slots' obj'\<rbrakk> \<Longrightarrow> slots = slots'"
|
||||
by (clarsimp simp: update_slots_def has_slots_def cdl_tcb.splits cdl_cnode.splits
|
||||
cdl_asid_pool.splits cdl_page_table.splits cdl_page_directory.splits
|
||||
split: cdl_object.splits)
|
||||
|
||||
lemma has_slots_object_type_has_slots:
|
||||
"\<lbrakk>has_slots x; object_type x = object_type y\<rbrakk> \<Longrightarrow> has_slots y"
|
||||
by (clarsimp simp: object_type_def has_slots_def split: cdl_object.splits)
|
||||
|
||||
lemma object_type_has_slots_eq:
|
||||
"object_type y = object_type x \<Longrightarrow> has_slots x = has_slots y"
|
||||
by (clarsimp simp: object_type_def has_slots_def split: cdl_object.splits)
|
||||
|
||||
lemma object_slots_object_clean_fields [simp]:
|
||||
"object_slots (object_clean_fields obj cmp) = object_slots obj"
|
||||
by (clarsimp simp: object_slots_def object_clean_fields_def split: cdl_object.splits)
|
||||
|
||||
lemma object_slots_object_clean_slots [simp]:
|
||||
"object_slots (object_clean_slots obj cmp) = clean_slots (object_slots obj) cmp"
|
||||
by (clarsimp simp: object_clean_slots_def object_slots_def update_slots_def split: cdl_object.splits)
|
||||
|
||||
lemma object_slots_object_clean [simp]:
|
||||
"object_slots (object_clean obj cmp) = clean_slots (object_slots obj) cmp"
|
||||
by (clarsimp simp: object_clean_def)
|
||||
|
||||
|
||||
lemma update_slots_object_clean_slots [simp]:
|
||||
"update_slots slots (object_clean_slots obj cmp) = update_slots slots obj"
|
||||
by (clarsimp simp: object_clean_slots_def)
|
||||
|
||||
lemma object_clean_fields_idem [simp]:
|
||||
"object_clean_fields (object_clean_fields obj cmp) cmp = object_clean_fields obj cmp"
|
||||
by (clarsimp simp: object_clean_fields_def split: cdl_object.splits)
|
||||
|
||||
lemma object_clean_slots_idem [simp]:
|
||||
"object_clean_slots (object_clean_slots obj cmp) cmp = object_clean_slots obj cmp"
|
||||
apply (case_tac "has_slots obj")
|
||||
apply (clarsimp simp: object_clean_slots_def clean_slots_def)+
|
||||
done
|
||||
|
||||
lemma object_clean_fields_object_clean_slots [simp]:
|
||||
"object_clean_fields (object_clean_slots obj cmps) cmps = object_clean_slots (object_clean_fields obj cmps) cmps"
|
||||
by (clarsimp simp: object_clean_fields_def object_clean_slots_def
|
||||
clean_slots_def object_slots_def update_slots_def
|
||||
split: cdl_object.splits)
|
||||
|
||||
lemma object_clean_idem [simp]:
|
||||
"object_clean (object_clean obj cmp) cmp = object_clean obj cmp"
|
||||
by (clarsimp simp: object_clean_def)
|
||||
|
||||
lemma has_slots_object_clean_slots:
|
||||
"has_slots (object_clean_slots obj cmp) = has_slots obj"
|
||||
by (clarsimp simp: has_slots_def object_clean_slots_def update_slots_def split: cdl_object.splits)
|
||||
|
||||
lemma has_slots_object_clean_fields:
|
||||
"has_slots (object_clean_fields obj cmp) = has_slots obj"
|
||||
by (clarsimp simp: has_slots_def object_clean_fields_def split: cdl_object.splits)
|
||||
|
||||
lemma has_slots_object_clean:
|
||||
"has_slots (object_clean obj cmp) = has_slots obj"
|
||||
by (clarsimp simp: object_clean_def has_slots_object_clean_slots has_slots_object_clean_fields)
|
||||
|
||||
lemma object_slots_update_slots_object_clean_fields [simp]:
|
||||
"object_slots (update_slots slots (object_clean_fields obj cmp)) = object_slots (update_slots slots obj)"
|
||||
apply (case_tac "has_slots obj")
|
||||
apply (clarsimp simp: has_slots_object_clean_fields)+
|
||||
done
|
||||
|
||||
lemma object_clean_fields_update_slots [simp]:
|
||||
"object_clean_fields (update_slots slots obj) cmp = update_slots slots (object_clean_fields obj cmp)"
|
||||
by (clarsimp simp: object_clean_fields_def update_slots_def split: cdl_object.splits)
|
||||
|
||||
lemma object_clean_fields_twice [simp]:
|
||||
"(object_clean_fields (object_clean_fields obj cmp') cmp) = object_clean_fields obj (cmp \<inter> cmp')"
|
||||
by (clarsimp simp: object_clean_fields_def split: cdl_object.splits)
|
||||
|
||||
lemma object_clean_fields_slot [simp]:
|
||||
"object_clean_fields obj (Slot ` slots) = object_clean_fields obj {}"
|
||||
by (clarsimp simp: object_clean_fields_def)
|
||||
|
||||
lemma object_clean_fields_slot' [simp]:
|
||||
"object_clean_fields obj {Slot slot} = object_clean_fields obj {}"
|
||||
by (clarsimp simp: object_clean_fields_def)
|
||||
|
||||
lemma object_clean_slots_empty_slots:
|
||||
"object_clean_slots obj (Slot ` (UNIV - dom (object_slots obj))) = object_clean_slots obj {}"
|
||||
by (clarsimp simp: object_clean_slots_def clean_slots_def)
|
||||
|
||||
lemma object_clean_empty_slots [simp]:
|
||||
"object_clean obj (Slot ` (UNIV - dom (object_slots obj))) = object_clean obj {}"
|
||||
by (clarsimp simp: object_clean_def object_clean_slots_def clean_slots_def)
|
||||
|
||||
lemma update_slots_object_clean_fields:
|
||||
"\<lbrakk>Fields \<notin> cmps; Fields \<notin> cmps'; object_type obj = object_type obj'\<rbrakk>
|
||||
\<Longrightarrow> update_slots slots (object_clean_fields obj cmps) =
|
||||
update_slots slots (object_clean_fields obj' cmps')"
|
||||
by (fastforce simp: update_slots_def object_clean_fields_def object_type_def split: cdl_object.splits)
|
||||
|
||||
lemma object_clean_fields_no_slots:
|
||||
"\<lbrakk>Fields \<notin> cmps; Fields \<notin> cmps'; object_type obj = object_type obj'; \<not> has_slots obj\<rbrakk>
|
||||
\<Longrightarrow> object_clean_fields obj cmps = object_clean_fields obj' cmps'"
|
||||
by (fastforce simp: object_clean_fields_def object_type_def has_slots_def split: cdl_object.splits)
|
||||
|
||||
lemma update_slots_object_clean:
|
||||
"\<lbrakk>Fields \<notin> cmps; Fields \<notin> cmps'; object_type obj = object_type obj'\<rbrakk>
|
||||
\<Longrightarrow> update_slots slots (object_clean obj cmps) = update_slots slots (object_clean obj' cmps')"
|
||||
apply (clarsimp simp: object_clean_def object_clean_slots_def)
|
||||
apply (erule (2) update_slots_object_clean_fields)
|
||||
done
|
||||
|
||||
lemma clean_slots_slot_same:
|
||||
"\<lbrakk>object_slots obj slot = object_slots obj' slot\<rbrakk>
|
||||
\<Longrightarrow> clean_slots (object_slots obj) {Slot slot} = clean_slots (object_slots obj') {Slot slot}"
|
||||
by (fastforce simp: clean_slots_def restrict_map_def)
|
||||
|
||||
lemma object_clean_single_slot_equal:
|
||||
"\<lbrakk>object_slots obj slot = object_slots obj' slot; object_type obj = object_type obj'\<rbrakk>
|
||||
\<Longrightarrow> object_clean obj {Slot slot} = object_clean obj' {Slot slot}"
|
||||
apply (clarsimp simp: object_clean_def object_clean_slots_def)
|
||||
apply (drule clean_slots_slot_same, simp)
|
||||
apply (clarsimp simp: update_slots_def object_clean_fields_def object_type_def
|
||||
split: cdl_object.splits)
|
||||
done
|
||||
|
||||
lemma clean_slots_map_add_comm:
|
||||
"cmps_a \<inter> cmps_b = {}
|
||||
\<Longrightarrow> clean_slots slots_a cmps_a ++ clean_slots slots_b cmps_b =
|
||||
clean_slots slots_b cmps_b ++ clean_slots slots_a cmps_a"
|
||||
apply (clarsimp simp: clean_slots_def)
|
||||
apply (drule the_set_inter_empty)
|
||||
apply (erule map_add_restrict_comm)
|
||||
done
|
||||
|
||||
lemma object_clean_all:
|
||||
"object_type obj_a = object_type obj_b \<Longrightarrow> object_clean obj_b {} = object_clean obj_a {}"
|
||||
apply (clarsimp simp: object_clean_def object_clean_slots_def clean_slots_def)
|
||||
apply (rule_tac cmps'1="{}" and obj'1="obj_a" in trans [OF update_slots_object_clean_fields], fastforce+)
|
||||
done
|
||||
|
||||
lemma sep_state_add_comm:
|
||||
"sep_state_disj x y \<Longrightarrow> sep_state_add x y = sep_state_add y x"
|
||||
by (fastforce simp: sep_state_add_def sep_state_disj_def intro!:map_add_com)
|
||||
|
||||
lemma sep_state_add_disjL:
|
||||
"\<lbrakk>sep_state_disj y z; sep_state_disj x (sep_state_add y z)\<rbrakk>
|
||||
\<Longrightarrow> sep_state_disj x y"
|
||||
by (fastforce simp: sep_state_disj_def
|
||||
sep_state_add_def map_disj_def)
|
||||
|
||||
lemma sep_state_add_disjR:
|
||||
"\<lbrakk>sep_state_disj y z; sep_state_disj x (sep_state_add y z)\<rbrakk> \<Longrightarrow> sep_state_disj x z"
|
||||
by (fastforce simp: sep_state_disj_def sep_state_add_def map_disj_def )
|
||||
|
||||
lemma sep_state_add_disj:
|
||||
"\<lbrakk>sep_state_disj y z; sep_state_disj x y; sep_state_disj x z\<rbrakk> \<Longrightarrow> sep_state_disj x (sep_state_add y z)"
|
||||
by (fastforce simp: sep_state_disj_def sep_state_add_def map_disj_def )
|
||||
|
||||
|
||||
|
||||
(*********************************************)
|
||||
(* Definition of separation logic for capDL. *)
|
||||
(*********************************************)
|
||||
|
@ -434,11 +95,10 @@ instance
|
|||
apply (simp add: sep_disj_sep_state_def sep_state_disj_def zero_sep_state_def)
|
||||
(* x ## y \<Longrightarrow> y ## x *)
|
||||
apply (clarsimp simp: sep_disj_sep_state_def sep_state_disj_def Let_unfold
|
||||
map_disj_com Int_commute)
|
||||
map_disj_com Int_commute)
|
||||
(* x + 0 = x *)
|
||||
apply (simp add: plus_sep_state_def sep_state_add_def zero_sep_state_def)
|
||||
apply (case_tac x,simp)
|
||||
|
||||
(* x ## y \<Longrightarrow> x + y = y + x *)
|
||||
apply (clarsimp simp: plus_sep_state_def sep_disj_sep_state_def)
|
||||
apply (erule sep_state_add_comm)
|
||||
|
|
|
@ -25,25 +25,24 @@ begin
|
|||
***********************************)
|
||||
|
||||
lemma dom_obj_to_sep_state:
|
||||
"dom (obj_to_sep_state ptr cmp obj) = {ptr} \<times> cmp"
|
||||
"dom (obj_to_sep_state ptr obj cmp) = {ptr} \<times> cmp"
|
||||
apply (clarsimp simp:obj_to_sep_state_def split:if_splits)
|
||||
apply (auto simp:object_project_def split:if_splits cdl_component.splits)
|
||||
done
|
||||
|
||||
lemma obj_to_sep_state_restrict:
|
||||
"obj_to_sep_state obj_id M obj |` ({obj_id} \<times> m) = obj_to_sep_state obj_id (M \<inter> m) obj"
|
||||
"obj_to_sep_state obj_id obj M |` ({obj_id} \<times> m) = obj_to_sep_state obj_id obj (M \<inter> m)"
|
||||
by (rule ext,
|
||||
clarsimp simp:obj_to_sep_state_def restrict_map_def
|
||||
split:if_splits)
|
||||
|
||||
lemma obj_to_sep_state_add:
|
||||
"\<lbrakk>{obj_id} \<times> m \<inter> dom M = {};b\<in> m\<rbrakk> \<Longrightarrow>
|
||||
(obj_to_sep_state obj_id m obj ++ M) (obj_id, b)
|
||||
(obj_to_sep_state obj_id obj m ++ M) (obj_id, b)
|
||||
= Some (object_project b obj)"
|
||||
apply (subgoal_tac "(obj_id, b)\<in> {obj_id} \<times> m")
|
||||
apply (drule(1) inter_empty_not_both)
|
||||
apply (clarsimp simp:obj_to_sep_state_def map_add_def
|
||||
dom_def)
|
||||
apply (clarsimp simp:obj_to_sep_state_def map_add_def dom_def)
|
||||
apply fastforce
|
||||
done
|
||||
|
||||
|
@ -59,7 +58,7 @@ lemma set_object_wp:
|
|||
apply (clarsimp simp: set_object_def state_sep_projection_def)
|
||||
apply (clarsimp simp: sep_conj_def)
|
||||
apply (simp add:sep_map_o_def sep_map_general_def)
|
||||
apply (rule_tac x = "SepState (obj_to_sep_state obj_id UNIV obj) empty" in exI)
|
||||
apply (rule_tac x = "SepState (obj_to_sep_state obj_id obj UNIV) empty" in exI)
|
||||
apply (rule_tac x=y in exI)
|
||||
apply (clarsimp simp: plus_sep_state_def sep_disj_sep_state_def
|
||||
sep_state_disj_def sep_any_def map_disj_def
|
||||
|
@ -77,60 +76,29 @@ lemma set_object_wp:
|
|||
apply (clarsimp split:option.splits)
|
||||
done
|
||||
|
||||
lemma intent_reset_object_clean_fields_some [simp]:
|
||||
"intent_reset (object_clean_fields obj {Slot slot}) =
|
||||
object_clean_fields obj {Slot slot}"
|
||||
by (clarsimp simp: intent_reset_def object_clean_fields_def
|
||||
split: cdl_object.splits)
|
||||
|
||||
lemma intent_reset_object_clean_fields_empty [simp]:
|
||||
"intent_reset (object_clean_fields obj {}) =
|
||||
object_clean_fields obj {}"
|
||||
by (clarsimp simp: intent_reset_def object_clean_fields_def
|
||||
split: cdl_object.splits)
|
||||
|
||||
lemma intent_reset_object_clean_some [simp]:
|
||||
"intent_reset (object_clean obj {Slot slot}) =
|
||||
object_clean obj {Slot slot}"
|
||||
apply (clarsimp simp: object_clean_def)
|
||||
apply (case_tac obj,simp_all add:intent_reset_def)
|
||||
apply (simp add:object_clean_fields_def)
|
||||
done
|
||||
|
||||
lemma map_add_self:
|
||||
"m ++ m = m"
|
||||
by (auto simp add:map_add_def split:option.splits)
|
||||
|
||||
definition "object_fields \<equiv> \<lambda>obj. object_clean_slots obj {Fields}"
|
||||
|
||||
lemma object_eqI:
|
||||
"\<lbrakk>object_fields obj = object_fields obj'; object_slots obj = object_slots obj'\<rbrakk>
|
||||
"\<lbrakk>object_wipe_slots obj = object_wipe_slots obj'; object_slots obj = object_slots obj'\<rbrakk>
|
||||
\<Longrightarrow> obj = obj'"
|
||||
by (simp add: object_slots_def object_fields_def
|
||||
object_clean_slots_def update_slots_def
|
||||
by (simp add: object_slots_def object_wipe_slots_def update_slots_def
|
||||
cdl_tcb.splits cdl_cnode.splits
|
||||
split: cdl_object.splits)
|
||||
|
||||
lemma object_filds_update_slots[simp]:
|
||||
"object_fields (update_slots slots obj) = object_fields obj"
|
||||
"object_wipe_slots (update_slots slots obj) = object_wipe_slots obj"
|
||||
apply (simp add:update_slots_def split:cdl_object.splits)
|
||||
apply (simp add:object_fields_def object_clean_slots_def update_slots_def)
|
||||
apply (simp add:object_wipe_slots_def update_slots_def)
|
||||
done
|
||||
|
||||
lemma object_fields_clean_slots[simp]:
|
||||
"object_fields (object_clean_slots obj cmps) = object_fields obj"
|
||||
by (case_tac obj,simp_all add:object_clean_slots_def object_fields_def object_slots_def
|
||||
clean_slots_def update_slots_def)
|
||||
lemma object_wipe_slots_idem[simp]:
|
||||
"object_wipe_slots (object_wipe_slots obj) = object_wipe_slots obj"
|
||||
by (case_tac obj,simp_all add: object_wipe_slots_def object_slots_def
|
||||
update_slots_def)
|
||||
|
||||
lemma object_slots_update_slots[simp]:
|
||||
"\<lbrakk>has_slots obj'; object_type obj = object_type obj'\<rbrakk>
|
||||
\<Longrightarrow> object_slots (update_slots slots obj) slot = slots slot"
|
||||
by (clarsimp simp:has_slots_def object_type_def split:cdl_object.splits)
|
||||
|
||||
|
||||
lemmas object_clean_fields_simps = object_clean_fields_def[split_simps cdl_object.split]
|
||||
lemmas object_clean_slots_simps = object_clean_slots_def[split_simps cdl_object.split]
|
||||
|
||||
lemma disjoint_union_diff:
|
||||
"a \<inter> b = {} \<Longrightarrow> a \<union> b - a = b"
|
||||
by auto
|
||||
|
@ -143,15 +111,15 @@ lemma intent_reset_update_slots_single:
|
|||
apply (auto simp:object_slots_def)
|
||||
done
|
||||
|
||||
lemma object_lift_update_slots_single:
|
||||
"object_lift (update_slots (object_slots obj(slot \<mapsto> cap)) obj)
|
||||
= update_slots (object_slots (object_lift obj)(slot \<mapsto> reset_cap_asid cap)) (object_lift obj)"
|
||||
by (auto simp: object_lift_def intent_reset_def asid_reset_def
|
||||
lemma object_clean_update_slots_single:
|
||||
"object_clean (update_slots (object_slots obj(slot \<mapsto> cap)) obj)
|
||||
= update_slots (object_slots (object_clean obj)(slot \<mapsto> reset_cap_asid cap)) (object_clean obj)"
|
||||
by (auto simp: object_clean_def intent_reset_def asid_reset_def
|
||||
update_slots_def object_slots_def fun_eq_iff
|
||||
split: cdl_object.splits)
|
||||
|
||||
lemma ptr_in_obj_to_sep_dom:
|
||||
"(ptr,Slot y) \<in> dom (obj_to_sep_state ptr {Slot y} obj)"
|
||||
"(ptr,Slot y) \<in> dom (obj_to_sep_state ptr obj {Slot y})"
|
||||
by (simp add:dom_obj_to_sep_state)
|
||||
|
||||
lemma sep_any_exist:
|
||||
|
@ -170,7 +138,7 @@ lemma sep_map_c_conj:
|
|||
sep_map_general_def sep_state_disj_def map_disj_def)
|
||||
apply (subst obj_to_sep_state_add)
|
||||
apply (simp add:dom_obj_to_sep_state singleton_iff)+
|
||||
apply (simp add:object_project_def object_slots_object_lift)
|
||||
apply (simp add:object_project_def object_slots_object_clean)
|
||||
apply (erule arg_cong[where f = R,THEN iffD1,rotated])
|
||||
apply (case_tac y)
|
||||
apply simp
|
||||
|
@ -194,14 +162,14 @@ lemma sep_map_c_conj:
|
|||
apply (clarsimp simp:object_slots_def)
|
||||
apply (rule ext)
|
||||
apply (clarsimp simp:object_project_def object_slots_def
|
||||
object_lift_def asid_reset_def update_slots_def
|
||||
object_clean_def asid_reset_def update_slots_def
|
||||
intent_reset_def split:if_splits option.splits)
|
||||
done
|
||||
|
||||
lemma sep_map_f_conj:
|
||||
"(ptr \<mapsto>f obj \<and>* R) s =
|
||||
(let slot = (ptr,Fields)
|
||||
in (sep_heap s) slot = Some (CDL_Object (object_clean_slots (object_lift obj) {})) \<and>
|
||||
in (sep_heap s) slot = Some (CDL_Object (object_wipe_slots (object_clean obj))) \<and>
|
||||
R (SepState ((sep_heap s)(slot:= None)) (sep_irq_node s)))"
|
||||
apply (clarsimp simp:sep_conj_def Let_def plus_sep_state_def sep_state_add_def)
|
||||
apply (rule iffI)
|
||||
|
@ -217,7 +185,7 @@ lemma sep_map_f_conj:
|
|||
apply (clarsimp simp:dom_def map_add_def
|
||||
obj_to_sep_state_def split:if_splits option.splits)
|
||||
apply (clarsimp simp:sep_map_c_def)
|
||||
apply (rule_tac x = "SepState [(ptr,Fields)\<mapsto> (CDL_Object (object_clean_slots (object_lift obj) {}))]
|
||||
apply (rule_tac x = "SepState [(ptr,Fields)\<mapsto> (CDL_Object (object_wipe_slots (object_clean obj)))]
|
||||
empty" in exI)
|
||||
apply (rule_tac x = "SepState ((sep_heap s)((ptr,Fields):= None))
|
||||
(sep_irq_node s)" in exI)
|
||||
|
@ -234,7 +202,7 @@ lemma sep_map_f_conj:
|
|||
|
||||
lemma object_project_slot:
|
||||
"object_project (Slot slot) obj =
|
||||
CDL_Cap (object_slots (object_lift obj) slot)"
|
||||
CDL_Cap (object_slots (object_clean obj) slot)"
|
||||
by (clarsimp simp:object_project_def)
|
||||
|
||||
lemma intent_reset_has_slots:
|
||||
|
@ -249,40 +217,39 @@ lemma asid_reset_has_slots:
|
|||
has_slots_def update_slots_def
|
||||
split:cdl_object.splits)
|
||||
|
||||
lemma object_lift_has_slots:
|
||||
"object_lift obj = object_lift obj' \<Longrightarrow> has_slots obj = has_slots obj'"
|
||||
by (clarsimp simp:object_lift_def2
|
||||
lemma object_clean_has_slots:
|
||||
"object_clean obj = object_clean obj' \<Longrightarrow> has_slots obj = has_slots obj'"
|
||||
by (clarsimp simp:object_clean_def2
|
||||
dest!: asid_reset_has_slots intent_reset_has_slots)
|
||||
|
||||
lemma set_object_slot_wp_helper:
|
||||
"\<lbrace>\<lambda>s. <(obj_id, slot) \<mapsto>c - \<and>* R> s \<and> cdl_objects s obj_id = Some obj
|
||||
\<and> object_lift obj = object_lift obj'\<rbrace>
|
||||
\<and> object_clean obj = object_clean obj'\<rbrace>
|
||||
set_object obj_id (update_slots (object_slots obj' (slot \<mapsto> cap)) obj')
|
||||
\<lbrace>\<lambda>rv. <(obj_id, slot) \<mapsto>c cap \<and>* R>\<rbrace>"
|
||||
apply (clarsimp simp: set_object_def sep_any_def)
|
||||
apply (clarsimp simp: sep_map_c_conj sep_conj_exists Let_def)
|
||||
apply (clarsimp simp: lift_def state_sep_projection_def
|
||||
sep_conj_def object_project_slot object_slots_object_lift)
|
||||
sep_conj_def object_project_slot object_slots_object_clean)
|
||||
apply (case_tac "has_slots obj")
|
||||
apply (frule object_lift_has_slots)
|
||||
apply (frule object_clean_has_slots)
|
||||
apply (simp add: object_slots_update_slots[where obj' = obj'])
|
||||
apply (erule arg_cong[where f = R,THEN iffD1,rotated])
|
||||
apply clarsimp
|
||||
apply (rule ext)
|
||||
apply (clarsimp split:if_splits option.splits)
|
||||
apply (clarsimp simp:object_project_def
|
||||
object_clean_slots_def clean_slots_def
|
||||
object_slots_object_lift object_lift_update_slots_single
|
||||
split:cdl_component.splits if_splits)
|
||||
apply (frule object_lift_has_slots)
|
||||
apply (clarsimp simp:has_slots_def
|
||||
object_lift_update_slots_single split:option.splits)
|
||||
apply (clarsimp split: if_splits option.splits)
|
||||
apply (clarsimp simp: object_project_def object_wipe_slots_def
|
||||
object_slots_object_clean object_clean_update_slots_single
|
||||
split: cdl_component.splits if_splits)
|
||||
apply (frule object_clean_has_slots)
|
||||
apply (clarsimp simp: has_slots_def object_clean_update_slots_single
|
||||
split: option.splits)
|
||||
done
|
||||
|
||||
lemma set_object_slot_wp:
|
||||
"\<lbrace>\<lambda>s. <(obj_id, slot) \<mapsto>c - \<and>* R> s \<and>
|
||||
cdl_objects s obj_id = Some obj \<and>
|
||||
(\<exists>obj'. object_lift obj = object_lift obj' \<and>
|
||||
(\<exists>obj'. object_clean obj = object_clean obj' \<and>
|
||||
nobj = (update_slots (object_slots obj' (slot \<mapsto> cap)) obj'))\<rbrace>
|
||||
set_object obj_id nobj
|
||||
\<lbrace>\<lambda>rv. <(obj_id, slot) \<mapsto>c cap \<and>* R>\<rbrace>"
|
||||
|
@ -293,17 +260,15 @@ lemma set_object_slot_wp:
|
|||
apply auto
|
||||
done
|
||||
|
||||
lemma object_fields_object_lift:
|
||||
"object_fields (object_lift obj) = object_lift (object_fields obj)"
|
||||
apply (case_tac obj, simp_all add: object_fields_def object_lift_def
|
||||
intent_reset_def asid_reset_def object_clean_slots_def
|
||||
update_slots_def object_slots_def)
|
||||
lemma object_wipe_slots_object_clean:
|
||||
"object_wipe_slots (object_clean obj) = object_clean (object_wipe_slots obj)"
|
||||
apply (case_tac obj, simp_all add: object_wipe_slots_def object_clean_def
|
||||
intent_reset_def asid_reset_def update_slots_def object_slots_def)
|
||||
done
|
||||
|
||||
lemma object_fields_intent_reset:
|
||||
"object_fields (intent_reset obj) = intent_reset (object_fields obj)"
|
||||
by (case_tac obj,simp_all add:object_fields_def object_clean_slots_def update_slots_def
|
||||
intent_reset_def)
|
||||
lemma object_wipe_slots_intent_reset:
|
||||
"object_wipe_slots (intent_reset obj) = intent_reset (object_wipe_slots obj)"
|
||||
by (case_tac obj,simp_all add:object_wipe_slots_def update_slots_def intent_reset_def)
|
||||
|
||||
abbreviation
|
||||
"tcb_at' P ptr s \<equiv> object_at (\<lambda>obj. \<exists>tcb. obj = Tcb tcb \<and> P tcb) ptr s"
|
||||
|
@ -311,55 +276,54 @@ abbreviation
|
|||
lemma sep_map_f_tcb_at:
|
||||
"\<lbrakk><ptr \<mapsto>f tcb \<and>* R> s ; is_tcb tcb\<rbrakk>
|
||||
\<Longrightarrow> tcb_at'
|
||||
(\<lambda>t. object_lift (object_fields (Tcb t)) = object_lift (object_fields tcb))
|
||||
(\<lambda>t. object_clean (object_wipe_slots (Tcb t)) = object_clean (object_wipe_slots tcb))
|
||||
ptr s"
|
||||
apply (simp add:state_sep_projection_def)
|
||||
apply (clarsimp simp:sep_map_f_conj
|
||||
object_project_def is_tcb_def
|
||||
split:option.splits cdl_object.splits)
|
||||
apply (case_tac z,
|
||||
simp_all add:object_clean_slots_def
|
||||
object_lift_def asid_reset_def object_fields_def
|
||||
update_slots_def intent_reset_def)
|
||||
apply (simp add:object_at_def opt_object_def object_slots_def
|
||||
object_fields_def object_clean_slots_def update_slots_def)
|
||||
simp_all add: object_wipe_slots_def object_clean_def asid_reset_def
|
||||
update_slots_def intent_reset_def)
|
||||
apply (simp add:object_at_def opt_object_def object_slots_def object_wipe_slots_def
|
||||
update_slots_def)
|
||||
apply (case_tac cdl_tcb_ext)
|
||||
apply clarsimp
|
||||
apply (case_tac cdl_tcb_exta)
|
||||
apply clarsimp
|
||||
done
|
||||
|
||||
lemma object_slots_update_slots_empty [simp]:
|
||||
"object_slots (update_slots empty object) = empty"
|
||||
by (case_tac "has_slots object", simp_all)
|
||||
|
||||
lemma set_object_cdl_field_wp:
|
||||
assumes slot: "object_slots obj_n = object_slots obj"
|
||||
and type: "object_type obj_n = object_type obj"
|
||||
"object_type obj_np = object_type obj"
|
||||
"object_type obj_p = object_type obj"
|
||||
and fields: "object_fields (object_lift obj_np) = object_fields (object_lift obj_n)"
|
||||
and fields: "object_wipe_slots (object_clean obj_np) = object_wipe_slots (object_clean obj_n)"
|
||||
shows
|
||||
"\<lbrace>\<lambda>s. <obj_id \<mapsto>f obj_p \<and>* R> s \<and> object_at (op = obj) obj_id s\<rbrace>
|
||||
set_object obj_id obj_n
|
||||
\<lbrace>\<lambda>rv. <obj_id \<mapsto>f obj_np \<and>* R>\<rbrace>"
|
||||
apply (clarsimp simp: set_object_def)
|
||||
apply (clarsimp simp: state_sep_projection_def
|
||||
sep_map_f_conj)
|
||||
apply (clarsimp simp: state_sep_projection_def sep_map_f_conj)
|
||||
apply (rule conjI)
|
||||
apply (simp add:object_project_def)
|
||||
apply (rule object_eqI)
|
||||
apply (cut_tac fields)
|
||||
apply (simp add:object_fields_object_lift)
|
||||
apply (simp add:object_wipe_slots_object_clean)
|
||||
apply (rule ext)
|
||||
apply (simp add:object_type_simps type)
|
||||
apply (cut_tac slot)
|
||||
apply (clarsimp split:if_splits simp:clean_slots_def )
|
||||
apply (simp add:object_wipe_slots_def type)
|
||||
apply (erule arg_cong[where f = R, THEN iffD1,rotated])
|
||||
apply (clarsimp simp:opt_object_def object_at_def)
|
||||
apply (rule ext)
|
||||
apply (clarsimp split:if_splits option.splits)
|
||||
apply (cut_tac slot)
|
||||
apply (rule ccontr)
|
||||
apply (clarsimp simp:object_project_def
|
||||
object_slots_object_lift
|
||||
split:cdl_component.splits)
|
||||
apply (clarsimp simp: object_project_def object_slots_object_clean
|
||||
split: cdl_component.splits)
|
||||
done
|
||||
|
||||
lemma set_cap_wp:
|
||||
|
@ -378,7 +342,7 @@ lemma set_cap_wp:
|
|||
apply (rule conjI, clarsimp)
|
||||
apply (rename_tac intent)
|
||||
apply (rule_tac x = "Tcb (tcb\<lparr>cdl_tcb_intent := intent\<rparr>)" in exI)
|
||||
apply (clarsimp simp: object_lift_def intent_reset_def asid_reset_def
|
||||
apply (clarsimp simp: object_clean_def intent_reset_def asid_reset_def
|
||||
update_slots_def object_slots_def)
|
||||
apply clarsimp
|
||||
apply (rule_tac x = "Tcb tcb" in exI)
|
||||
|
@ -391,8 +355,8 @@ lemma set_cap_wp:
|
|||
*)
|
||||
lemma set_cdl_tcb_field_wp:
|
||||
assumes fields_cong:
|
||||
"\<And>tcb tcb'. object_lift (object_fields (Tcb tcb)) = object_lift (object_fields (Tcb tcb'))
|
||||
\<Longrightarrow> object_lift (object_fields (Tcb (f tcb))) = object_lift (object_fields (Tcb (f tcb')))"
|
||||
"\<And>tcb tcb'. object_clean (object_wipe_slots (Tcb tcb)) = object_clean (object_wipe_slots (Tcb tcb'))
|
||||
\<Longrightarrow> object_clean (object_wipe_slots (Tcb (f tcb))) = object_clean (object_wipe_slots (Tcb (f tcb')))"
|
||||
assumes slots_simp:
|
||||
"\<And>tcb. object_slots (Tcb (f tcb)) = object_slots (Tcb tcb)"
|
||||
shows
|
||||
|
@ -406,10 +370,10 @@ lemma set_cdl_tcb_field_wp:
|
|||
apply (clarsimp simp:object_at_def)
|
||||
apply (wp|wpc|simp)+
|
||||
apply (rule_tac P =
|
||||
"object_lift (object_fields (Tcb cdl_tcb_ext)) = object_lift (object_fields (Tcb tcb))"
|
||||
"object_clean (object_wipe_slots (Tcb cdl_tcb_ext)) = object_clean (object_wipe_slots (Tcb tcb))"
|
||||
in hoare_gen_asm)
|
||||
apply (rule_tac obj_p = "Tcb tcb" in set_object_cdl_field_wp[OF slots_simp ])
|
||||
apply (simp add:object_type_simps object_fields_object_lift)+
|
||||
apply (simp add:object_type_simps object_wipe_slots_object_clean)+
|
||||
apply (drule_tac tcb = cdl_tcb_ext in fields_cong)
|
||||
apply simp
|
||||
apply wp
|
||||
|
@ -423,10 +387,9 @@ lemma set_cdl_tcb_intent_wp:
|
|||
\<lbrace>\<lambda>rv s. <obj_id \<mapsto>f Tcb (tcb\<lparr>cdl_tcb_intent := x\<rparr>) \<and>* R> s\<rbrace>"
|
||||
apply (rule hoare_weaken_pre)
|
||||
apply (wp set_cdl_tcb_field_wp)
|
||||
apply (clarsimp simp:object_fields_def object_lift_def
|
||||
update_slots_def asid_reset_def object_slots_def
|
||||
intent_reset_def object_clean_slots_def)
|
||||
apply (clarsimp simp:object_slots_def)
|
||||
apply (clarsimp simp: object_wipe_slots_def object_clean_def update_slots_def
|
||||
asid_reset_def object_slots_def intent_reset_def)
|
||||
apply (clarsimp simp: object_slots_def)
|
||||
apply simp
|
||||
done
|
||||
|
||||
|
@ -436,9 +399,8 @@ lemma set_cdl_tcb_fault_endpoint_wp:
|
|||
\<lbrace>\<lambda>rv s. <obj_id \<mapsto>f Tcb (tcb\<lparr>cdl_tcb_fault_endpoint := x\<rparr>) \<and>* R> s\<rbrace>"
|
||||
apply (rule hoare_weaken_pre)
|
||||
apply (wp set_cdl_tcb_field_wp)
|
||||
apply (clarsimp simp:object_fields_def asid_reset_def
|
||||
update_slots_def object_lift_def object_slots_def
|
||||
intent_reset_def object_clean_slots_def)
|
||||
apply (clarsimp simp: object_wipe_slots_def asid_reset_def update_slots_def
|
||||
object_clean_def object_slots_def intent_reset_def)
|
||||
apply (case_tac tcb,case_tac tcb',simp)
|
||||
apply (clarsimp simp:object_slots_def)
|
||||
apply simp
|
||||
|
@ -446,7 +408,7 @@ lemma set_cdl_tcb_fault_endpoint_wp:
|
|||
|
||||
lemma sep_map_o_conj:
|
||||
"(ptr \<mapsto>o obj \<and>* R) s = (( (sep_heap s) |` ({ptr} \<times> UNIV)
|
||||
= obj_to_sep_state ptr UNIV obj )
|
||||
= obj_to_sep_state ptr obj UNIV)
|
||||
\<and> R (SepState (sep_heap s |` (UNIV - {ptr}\<times>UNIV))
|
||||
(sep_irq_node s)))"
|
||||
apply (clarsimp simp:sep_conj_def Let_def plus_sep_state_def sep_state_add_def
|
||||
|
@ -467,7 +429,7 @@ lemma sep_map_o_conj:
|
|||
apply (subst restrict_map_univ_disj_eq)
|
||||
apply (fastforce simp:map_disj_def)
|
||||
apply simp
|
||||
apply (rule_tac x = "SepState (obj_to_sep_state ptr UNIV obj) empty" in exI)
|
||||
apply (rule_tac x = "SepState (obj_to_sep_state ptr obj UNIV) empty" in exI)
|
||||
apply (rule_tac x = "SepState (sep_heap s |` (UNIV - {ptr}\<times>UNIV))
|
||||
(sep_irq_node s) " in exI)
|
||||
apply (clarsimp simp:sep_map_c_def sep_disj_sep_state_def
|
||||
|
@ -511,13 +473,11 @@ lemma create_one_wp:
|
|||
apply (clarsimp simp: state_sep_projection_def sep_map_o_conj)
|
||||
apply (intro conjI ext)
|
||||
apply (drule_tac x = xa in fun_cong)
|
||||
apply (clarsimp simp: object_at_def Let_unfold heap_lift_def
|
||||
split: option.splits)
|
||||
apply (clarsimp simp:object_project_def
|
||||
restrict_map_def obj_to_sep_state_def)
|
||||
apply (clarsimp simp: obj_to_sep_state_def)
|
||||
apply (erule arg_cong[where f = R,THEN iffD1,rotated])
|
||||
apply clarsimp
|
||||
apply (rule ext)
|
||||
apply (clarsimp simp:restrict_map_def split:if_splits)
|
||||
done
|
||||
|
||||
end
|
||||
|
|
|
@ -71,6 +71,10 @@ lemma dom_map_comp' [simp]:
|
|||
"dom (f \<circ>\<^sub>M m) = dom m"
|
||||
by (clarsimp simp: dom_def split: option.splits)
|
||||
|
||||
lemma restrict_map_comp' [simp]:
|
||||
"f \<circ>\<^sub>M g |` S = (f \<circ>\<^sub>M g) |` S"
|
||||
by (rule ext, clarsimp simp: map_comp'_def restrict_map_def)
|
||||
|
||||
|
||||
(**************************
|
||||
* Rules about intervals. *
|
||||
|
@ -85,6 +89,59 @@ lemma nat_list_not_UNIV [simp]:
|
|||
***********************************)
|
||||
|
||||
|
||||
(*************************************
|
||||
* Rules that should go in MapExtra. *
|
||||
*************************************)
|
||||
|
||||
lemma map_add_self:
|
||||
"m ++ m = m"
|
||||
by (auto simp add:map_add_def split:option.splits)
|
||||
|
||||
(* This lemma is strictly weaker than restrict_map_disj. *)
|
||||
lemma restrict_map_disj':
|
||||
"S \<inter> T = {} \<Longrightarrow> h |` S \<bottom> h' |` T"
|
||||
by (auto simp: map_disj_def restrict_map_def dom_def)
|
||||
|
||||
lemma map_add_restrict_comm:
|
||||
"S \<inter> T = {} \<Longrightarrow> h |` S ++ h' |` T = h' |` T ++ h |` S"
|
||||
apply (drule restrict_map_disj')
|
||||
apply (erule map_add_com)
|
||||
done
|
||||
|
||||
lemma restrict_map_not_self_UNIV [simp]:
|
||||
"h |` (UNIV - dom h) = empty"
|
||||
by (rule ext, clarsimp simp: restrict_map_def)
|
||||
|
||||
lemma map_add_emptyE [elim!]:
|
||||
"\<lbrakk>a ++ b = empty; \<lbrakk>a = empty; b = empty\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
|
||||
by (metis map_add_None)
|
||||
|
||||
lemma restrict_map_sub_singleton [simp]:
|
||||
"s \<noteq> t \<Longrightarrow> h `- {t} |` {s} = h |` {s}"
|
||||
by (rule ext)(clarsimp simp: restrict_map_def sub_restrict_map_def)
|
||||
|
||||
lemma restrict_map_sub_add': "h `- S ++ h |` S = h"
|
||||
by (fastforce simp: sub_restrict_map_def map_add_def
|
||||
split: option.splits)
|
||||
|
||||
lemma map_add_restrict_singleton:
|
||||
"s \<noteq> t \<Longrightarrow> (m |` {s} ++ m' |` {t}) |` {s} = m |` {s}"
|
||||
apply (rule ext)
|
||||
apply (clarsimp simp: map_add_def restrict_map_def split: option.splits)
|
||||
done
|
||||
|
||||
lemma map_add_restrict_singleton':
|
||||
"s \<noteq> t \<Longrightarrow> (m |` {s} ++ m' |` {t}) |` {t} = m' |` {t}"
|
||||
apply (rule ext)
|
||||
apply (clarsimp simp: map_add_def restrict_map_def split: option.splits)
|
||||
done
|
||||
|
||||
(********************************************
|
||||
* End of rules that should go in MapExtra. *
|
||||
********************************************)
|
||||
|
||||
|
||||
|
||||
(**********************************************************
|
||||
* CapDL specific lemmas, that could be moved into capDL. *
|
||||
* These have nothing to do with the separation logic. *
|
||||
|
@ -927,6 +984,118 @@ lemma safe_for_derive_not_non:
|
|||
|
||||
|
||||
|
||||
(* Given an object, this returns what the default state is for the object.
|
||||
* This should be the same as what is created by a retype operation.
|
||||
*)
|
||||
definition
|
||||
object_default_state :: "cdl_object \<Rightarrow> cdl_object"
|
||||
where
|
||||
"object_default_state object \<equiv>
|
||||
the $ default_object (object_type object)
|
||||
(object_size_bits object)
|
||||
(object_domain object)"
|
||||
|
||||
lemma object_default_state_def2:
|
||||
"object_default_state obj = (
|
||||
case obj of
|
||||
Untyped \<Rightarrow> Untyped
|
||||
| Endpoint \<Rightarrow> Endpoint
|
||||
| AsyncEndpoint \<Rightarrow> AsyncEndpoint
|
||||
| Tcb tcb \<Rightarrow> Tcb (default_tcb (cdl_tcb_domain tcb))
|
||||
| CNode cnode \<Rightarrow> CNode (empty_cnode (cdl_cnode_size_bits cnode))
|
||||
| AsidPool ap \<Rightarrow> AsidPool \<lparr>cdl_asid_pool_caps = empty_cap_map asid_low_bits\<rparr>
|
||||
| PageTable pt \<Rightarrow> PageTable \<lparr> cdl_page_table_caps = empty_cap_map 8 \<rparr>
|
||||
| PageDirectory pd \<Rightarrow> PageDirectory \<lparr> cdl_page_directory_caps = empty_cap_map 12 \<rparr>
|
||||
| Frame frame \<Rightarrow> Frame \<lparr> cdl_frame_size_bits = (cdl_frame_size_bits frame) \<rparr>)"
|
||||
by (clarsimp simp: object_default_state_def object_type_def default_object_def
|
||||
object_size_bits_def object_domain_def
|
||||
split: cdl_object.splits)
|
||||
|
||||
(******************************************************
|
||||
* More rules, that should be put in the right place. *
|
||||
******************************************************)
|
||||
|
||||
lemma object_type_update_slots [simp]:
|
||||
"object_type (update_slots slots x) = object_type x"
|
||||
by (clarsimp simp: object_type_def update_slots_def split: cdl_object.splits)
|
||||
|
||||
lemma object_slots_empty [simp]:
|
||||
"\<not> has_slots obj \<Longrightarrow> object_slots obj = empty"
|
||||
by (clarsimp simp: object_slots_def has_slots_def split: cdl_object.splits)
|
||||
|
||||
lemma object_slots_update_slots [simp]:
|
||||
"has_slots obj \<Longrightarrow> object_slots (update_slots slots obj) = slots"
|
||||
by (clarsimp simp: object_slots_def update_slots_def has_slots_def
|
||||
split: cdl_object.splits)
|
||||
|
||||
lemma object_slots_update_slots_empty [simp]:
|
||||
"\<not>has_slots obj \<Longrightarrow> object_slots (update_slots slots obj) = empty"
|
||||
by (clarsimp simp: object_slots_def update_slots_def has_slots_def
|
||||
split: cdl_object.splits)
|
||||
|
||||
lemma update_slots_no_slots [simp]:
|
||||
"\<not>has_slots obj \<Longrightarrow> update_slots slots obj = obj"
|
||||
by (clarsimp simp: update_slots_def has_slots_def split: cdl_object.splits)
|
||||
|
||||
lemma update_slots_update_slots [simp]:
|
||||
"update_slots slots (update_slots slots' obj) = update_slots slots obj"
|
||||
by (clarsimp simp: update_slots_def split: cdl_object.splits)
|
||||
|
||||
lemma update_slots_same_object:
|
||||
"a = b \<Longrightarrow> update_slots a obj = update_slots b obj"
|
||||
by (erule arg_cong)
|
||||
|
||||
lemma update_slots_eq_slots:
|
||||
"\<lbrakk>has_slots obj; update_slots slots obj = update_slots slots' obj'\<rbrakk> \<Longrightarrow> slots = slots'"
|
||||
by (clarsimp simp: update_slots_def has_slots_def cdl_tcb.splits cdl_cnode.splits
|
||||
cdl_asid_pool.splits cdl_page_table.splits cdl_page_directory.splits
|
||||
split: cdl_object.splits)
|
||||
|
||||
lemma has_slots_object_type_has_slots:
|
||||
"\<lbrakk>has_slots x; object_type x = object_type y\<rbrakk> \<Longrightarrow> has_slots y"
|
||||
by (clarsimp simp: object_type_def has_slots_def split: cdl_object.splits)
|
||||
|
||||
lemma object_type_has_slots_eq:
|
||||
"object_type y = object_type x \<Longrightarrow> has_slots x = has_slots y"
|
||||
by (clarsimp simp: object_type_def has_slots_def split: cdl_object.splits)
|
||||
|
||||
|
||||
lemma object_type_object_default_state [simp]:
|
||||
"object_type (object_default_state obj) = object_type obj"
|
||||
by (clarsimp simp: object_default_state_def2 object_type_def split: cdl_object.splits)
|
||||
|
||||
lemma is_cnode_object_default_state [simp]:
|
||||
"is_cnode (object_default_state obj) = is_cnode obj"
|
||||
by (clarsimp simp: object_default_state_def2 is_cnode_def split: cdl_object.splits)
|
||||
|
||||
(* FIXME, is this a bad idea to add to the simpset? *)
|
||||
lemma range_Slot [simp]:
|
||||
"(range Slot) = (UNIV - {Fields})"
|
||||
apply (clarsimp simp: image_def)
|
||||
apply rule
|
||||
apply clarsimp+
|
||||
apply (case_tac x)
|
||||
apply auto
|
||||
done
|
||||
|
||||
lemma update_slots_same [simp]:
|
||||
"object_slots obj = cap_map \<Longrightarrow> update_slots cap_map obj = obj"
|
||||
by (clarsimp simp: update_slots_def object_slots_def split: cdl_object.splits)
|
||||
|
||||
lemma dom_sub_restrict [simp]:
|
||||
"dom (m `- A) = dom m \<inter> -A"
|
||||
by (auto simp: sub_restrict_map_def dom_def split: split_if_asm)
|
||||
|
||||
lemma Slot_slot_union:
|
||||
"insert (Slot slot) (Slot ` (UNIV - {slot})) = UNIV - {Fields}"
|
||||
apply rule
|
||||
apply clarsimp+
|
||||
apply (case_tac x)
|
||||
apply auto
|
||||
done
|
||||
|
||||
lemma inter_empty_not_both:
|
||||
"\<lbrakk>x \<in> A; A \<inter> B = {}\<rbrakk> \<Longrightarrow> x \<notin> B"
|
||||
by fastforce
|
||||
|
||||
end
|
||||
|
|
|
@ -34,167 +34,6 @@ lemma word_bits_eq [simp]:
|
|||
|
||||
hide_const (open) Types_D.word_bits
|
||||
|
||||
lemma restrict_map_comp' [simp]:
|
||||
"f \<circ>\<^sub>M g |` S = (f \<circ>\<^sub>M g) |` S"
|
||||
by (rule ext, clarsimp simp: map_comp'_def restrict_map_def)
|
||||
|
||||
(********************
|
||||
* Move to elsewhere *
|
||||
*********************)
|
||||
|
||||
lemma sep_list_conj_impl:
|
||||
"\<lbrakk> list_all2 (\<lambda>x y. \<forall>s. x s \<longrightarrow> y s) xs ys; (\<And>* xs) s \<rbrakk> \<Longrightarrow> (\<And>* ys) s"
|
||||
apply (induct arbitrary: s rule: list_all2_induct)
|
||||
apply simp
|
||||
apply simp
|
||||
apply (erule sep_conj_impl, simp_all)
|
||||
done
|
||||
|
||||
lemma sep_list_conj_exists:
|
||||
"(\<exists>x. (\<And>* map (\<lambda>y s. P x y s) ys) s) \<Longrightarrow> ((\<And>* map (\<lambda>y s. \<exists>x. P x y s) ys) s)"
|
||||
apply clarsimp
|
||||
apply (erule sep_list_conj_impl[rotated])
|
||||
apply (rule list_all2I)
|
||||
apply simp_all
|
||||
apply (fastforce simp add: zip_map1 zip_map2 set_zip_same)
|
||||
done
|
||||
|
||||
(* Given an object, this returns what the default state is for the object.
|
||||
* This should be the same as what is created by a retype operation.
|
||||
*)
|
||||
definition
|
||||
object_default_state :: "cdl_object \<Rightarrow> cdl_object"
|
||||
where
|
||||
"object_default_state object \<equiv>
|
||||
the $ default_object (object_type object)
|
||||
(object_size_bits object)
|
||||
(object_domain object)"
|
||||
|
||||
lemma object_default_state_def2:
|
||||
"object_default_state obj = (
|
||||
case obj of
|
||||
Untyped \<Rightarrow> Untyped
|
||||
| Endpoint \<Rightarrow> Endpoint
|
||||
| AsyncEndpoint \<Rightarrow> AsyncEndpoint
|
||||
| Tcb tcb \<Rightarrow> Tcb (default_tcb (cdl_tcb_domain tcb))
|
||||
| CNode cnode \<Rightarrow> CNode (empty_cnode (cdl_cnode_size_bits cnode))
|
||||
| AsidPool ap \<Rightarrow> AsidPool \<lparr>cdl_asid_pool_caps = empty_cap_map asid_low_bits\<rparr>
|
||||
| PageTable pt \<Rightarrow> PageTable \<lparr> cdl_page_table_caps = empty_cap_map 8 \<rparr>
|
||||
| PageDirectory pd \<Rightarrow> PageDirectory \<lparr> cdl_page_directory_caps = empty_cap_map 12 \<rparr>
|
||||
| Frame frame \<Rightarrow> Frame \<lparr> cdl_frame_size_bits = (cdl_frame_size_bits frame) \<rparr>)"
|
||||
by (clarsimp simp: object_default_state_def object_type_def default_object_def
|
||||
object_size_bits_def object_domain_def
|
||||
split: cdl_object.splits)
|
||||
|
||||
lemma object_type_object_default_state [simp]:
|
||||
"object_type (object_default_state obj) = object_type obj"
|
||||
by (clarsimp simp: object_default_state_def2 object_type_def split: cdl_object.splits)
|
||||
|
||||
lemma is_cnode_object_default_state [simp]:
|
||||
"is_cnode (object_default_state obj) = is_cnode obj"
|
||||
by (clarsimp simp: object_default_state_def2 is_cnode_def split: cdl_object.splits)
|
||||
|
||||
(* FIXME, is this a bad idea to add to the simpset? *)
|
||||
lemma range_Slot [simp]:
|
||||
"(range Slot) = (UNIV - {Fields})"
|
||||
apply (clarsimp simp: image_def)
|
||||
apply rule
|
||||
apply clarsimp+
|
||||
apply (case_tac x)
|
||||
apply auto
|
||||
done
|
||||
|
||||
lemma update_slots_same [simp]:
|
||||
"object_slots obj = cap_map \<Longrightarrow> update_slots cap_map obj = obj"
|
||||
by (clarsimp simp: update_slots_def object_slots_def split: cdl_object.splits)
|
||||
|
||||
lemma dom_sub_restrict [simp]:
|
||||
"dom (m `- A) = dom m \<inter> -A"
|
||||
by (auto simp: sub_restrict_map_def dom_def split: split_if_asm)
|
||||
|
||||
lemma Slot_slot_union:
|
||||
"insert (Slot slot) (Slot ` (UNIV - {slot})) = UNIV - {Fields}"
|
||||
apply rule
|
||||
apply clarsimp+
|
||||
apply (case_tac x)
|
||||
apply auto
|
||||
done
|
||||
|
||||
lemma restrict_map_sub_singleton [simp]:
|
||||
"s \<noteq> t \<Longrightarrow> h `- {t} |` {s} = h |` {s}"
|
||||
by (rule ext)(clarsimp simp: restrict_map_def sub_restrict_map_def)
|
||||
|
||||
lemma restrict_map_sub_add': "h `- S ++ h |` S = h"
|
||||
by (fastforce simp: sub_restrict_map_def map_add_def
|
||||
split: option.splits)
|
||||
|
||||
lemma map_add_restrict_singleton:
|
||||
"s \<noteq> t \<Longrightarrow> (m |` {s} ++ m' |` {t}) |` {s} = m |` {s}"
|
||||
apply (rule ext)
|
||||
apply (clarsimp simp: map_add_def restrict_map_def split: option.splits)
|
||||
done
|
||||
|
||||
lemma map_add_restrict_singleton':
|
||||
"s \<noteq> t \<Longrightarrow> (m |` {s} ++ m' |` {t}) |` {t} = m' |` {t}"
|
||||
apply (rule ext)
|
||||
apply (clarsimp simp: map_add_def restrict_map_def split: option.splits)
|
||||
done
|
||||
|
||||
lemma object_clean_fields_fields [simp]:
|
||||
"Fields \<in> cmps \<Longrightarrow> (object_clean_fields obj cmps) = obj"
|
||||
by (clarsimp simp: object_clean_fields_def)
|
||||
|
||||
lemma object_clean_fields_not_none_eq:
|
||||
"\<lbrakk>Fields \<notin> cmps; Fields \<notin> cmps'\<rbrakk>
|
||||
\<Longrightarrow> object_clean_fields obj cmps = object_clean_fields obj cmps'"
|
||||
by (clarsimp simp: object_clean_fields_def)
|
||||
|
||||
lemma clean_slots_map_addL [simp]:
|
||||
"cmps \<inter> cmps' = {}
|
||||
\<Longrightarrow> clean_slots (clean_slots slots cmps ++ clean_slots slots' cmps') cmps = clean_slots slots cmps"
|
||||
by (force simp: clean_slots_def map_add_def restrict_map_def)
|
||||
|
||||
lemma clean_slots_map_addR [simp]:
|
||||
"cmps \<inter> cmps' = {}
|
||||
\<Longrightarrow> clean_slots (clean_slots slots cmps ++ clean_slots slots' cmps') cmps' = clean_slots slots' cmps'"
|
||||
by (force simp: clean_slots_def map_add_def restrict_map_def
|
||||
split: option.splits)
|
||||
|
||||
lemma empty_fun_upd [simp]:
|
||||
"(\<lambda>x. {})(obj_id := {}) = (\<lambda>x. {})"
|
||||
by clarsimp
|
||||
|
||||
lemma clean_slots_UNIV [simp]:
|
||||
"clean_slots slots UNIV = slots"
|
||||
by (clarsimp simp: clean_slots_def)
|
||||
|
||||
lemma object_clean_slots_UNIV [simp]:
|
||||
"object_clean_slots obj UNIV = obj"
|
||||
by (clarsimp simp: object_clean_slots_def)
|
||||
|
||||
lemma object_clean_UNIV [simp]:
|
||||
"object_clean obj UNIV = obj"
|
||||
by (clarsimp simp: object_clean_def)
|
||||
|
||||
lemma object_clean_fields_object_clean:
|
||||
"object_clean_fields obj {Slot slot} = object_clean_fields obj' {Slot slot}
|
||||
\<Longrightarrow> object_clean obj {Slot slot} = object_clean obj' {Slot slot}"
|
||||
by (clarsimp simp: object_clean_def)
|
||||
|
||||
lemma object_clean_slots_object_clean:
|
||||
"object_clean_slots obj {Slot slot} = object_clean_slots obj' {Slot slot}
|
||||
\<Longrightarrow> object_clean obj {Slot slot} = object_clean obj' {Slot slot}"
|
||||
apply (subgoal_tac "object_type obj = object_type obj'")
|
||||
apply (clarsimp simp: object_clean_def object_clean_fields_def)
|
||||
apply (fastforce simp: object_clean_slots_def object_type_def object_slots_def
|
||||
update_slots_def cdl_tcb.splits cdl_cnode.splits
|
||||
split: cdl_object.splits)+
|
||||
done
|
||||
|
||||
(************************
|
||||
* End of things to move *
|
||||
*************************)
|
||||
|
||||
|
||||
|
||||
definition
|
||||
|
@ -238,27 +77,29 @@ where
|
|||
(* intent_reset is easier to bubble through object_slots,
|
||||
so is put on the outside. *)
|
||||
definition
|
||||
object_lift :: "cdl_object \<Rightarrow> cdl_object"
|
||||
object_clean :: "cdl_object \<Rightarrow> cdl_object"
|
||||
where
|
||||
"object_lift \<equiv> intent_reset \<circ> asid_reset"
|
||||
"object_clean \<equiv> intent_reset \<circ> asid_reset"
|
||||
|
||||
lemma object_lift_def2:
|
||||
"object_lift = asid_reset \<circ> intent_reset"
|
||||
lemma object_clean_def2:
|
||||
"object_clean = asid_reset \<circ> intent_reset"
|
||||
apply (rule ext)
|
||||
apply (clarsimp simp: object_lift_def intent_reset_def asid_reset_def
|
||||
apply (clarsimp simp: object_clean_def intent_reset_def asid_reset_def
|
||||
map_comp'_def object_slots_def update_slots_def
|
||||
split: cdl_object.splits)
|
||||
done
|
||||
|
||||
|
||||
(* Clears the capability slots of an object. *)
|
||||
definition
|
||||
heap_lift :: "cdl_heap \<Rightarrow> cdl_heap"
|
||||
object_wipe_slots :: "cdl_object \<Rightarrow> cdl_object"
|
||||
where
|
||||
"heap_lift heap \<equiv> \<lambda>obj_id. option_map object_lift (heap obj_id)"
|
||||
"object_wipe_slots obj \<equiv> update_slots empty obj"
|
||||
|
||||
definition object_project :: "cdl_component \<Rightarrow> cdl_object \<Rightarrow> sep_entity"
|
||||
where "object_project component obj \<equiv> case component of
|
||||
Fields \<Rightarrow> CDL_Object (object_clean_slots (object_lift obj) {})
|
||||
| Slot slot \<Rightarrow> CDL_Cap (object_slots (object_lift obj) slot)"
|
||||
Fields \<Rightarrow> CDL_Object (object_wipe_slots (object_clean obj))
|
||||
| Slot slot \<Rightarrow> CDL_Cap (object_slots (object_clean obj) slot)"
|
||||
|
||||
definition
|
||||
state_sep_projection :: "cdl_state \<Rightarrow> sep_state"
|
||||
|
@ -294,11 +135,11 @@ where
|
|||
else None"
|
||||
|
||||
definition
|
||||
obj_to_sep_state :: "cdl_object_id \<Rightarrow> cdl_components \<Rightarrow> cdl_object \<Rightarrow> sep_state_heap"
|
||||
obj_to_sep_state :: "cdl_object_id \<Rightarrow> cdl_object \<Rightarrow> cdl_components \<Rightarrow> sep_state_heap"
|
||||
where
|
||||
"obj_to_sep_state p cmp obj \<equiv> \<lambda>(ptr,component).
|
||||
if p = ptr \<and> component \<in> cmp
|
||||
then Some (object_project component obj)
|
||||
"obj_to_sep_state obj_id object components \<equiv> \<lambda>(obj_id',component).
|
||||
if obj_id = obj_id' \<and> component \<in> components
|
||||
then Some (object_project component object)
|
||||
else None"
|
||||
|
||||
(*********************
|
||||
|
@ -310,45 +151,53 @@ where
|
|||
definition
|
||||
sep_map_general :: "cdl_object_id \<Rightarrow> cdl_object \<Rightarrow> cdl_components \<Rightarrow> sep_pred"
|
||||
where
|
||||
"sep_map_general p obj cmps \<equiv> \<lambda>s.
|
||||
sep_heap s = (obj_to_sep_state p cmps obj) \<and>
|
||||
"sep_map_general obj_id object components \<equiv> \<lambda>s.
|
||||
sep_heap s = (obj_to_sep_state obj_id object components) \<and>
|
||||
sep_irq_node s = empty"
|
||||
|
||||
(* There is an object there. *)
|
||||
definition
|
||||
sep_map_o :: "cdl_object_id \<Rightarrow> cdl_object \<Rightarrow> sep_pred" ("_ \<mapsto>o _" [76,71] 76)
|
||||
where
|
||||
"p \<mapsto>o obj \<equiv> sep_map_general p obj UNIV"
|
||||
"obj_id \<mapsto>o object \<equiv> sep_map_general obj_id object UNIV"
|
||||
|
||||
(* The fields are there (and there are no caps). *)
|
||||
definition
|
||||
sep_map_f :: "cdl_object_id \<Rightarrow> cdl_object \<Rightarrow> sep_pred" ("_ \<mapsto>f _" [76,71] 76)
|
||||
where
|
||||
"p \<mapsto>f obj \<equiv> sep_map_general p obj {Fields}"
|
||||
"obj_id \<mapsto>f object \<equiv> sep_map_general obj_id object {Fields}"
|
||||
|
||||
(* There is a clean object there that has the same caps in the same slots. *)
|
||||
definition
|
||||
sep_map_S :: "cdl_object_id \<Rightarrow> cdl_object \<Rightarrow> sep_pred" ("_ \<mapsto>S _" [76,71] 76)
|
||||
where
|
||||
"p \<mapsto>S obj \<equiv> sep_map_general p obj (UNIV - {Fields})"
|
||||
"obj_id \<mapsto>S object \<equiv> sep_map_general obj_id object (UNIV - {Fields})"
|
||||
|
||||
(* There is a clean object there that has the same caps in the same slots, restricted to the slots "slots" *)
|
||||
definition
|
||||
sep_map_S' :: "(cdl_object_id \<times> cdl_cnode_index set) \<Rightarrow> cdl_object \<Rightarrow> sep_pred" ("_ \<mapsto>S' _" [76,71] 76)
|
||||
where
|
||||
"p \<mapsto>S' obj \<equiv> let (obj_id, slots) = p in sep_map_general obj_id obj (Slot ` slots)"
|
||||
"p \<mapsto>S' object \<equiv> let (obj_id, slots) = p in sep_map_general obj_id object (Slot ` slots)"
|
||||
|
||||
lemma sep_map_S'_def2:
|
||||
"(obj_id, slots) \<mapsto>S' object \<equiv> sep_map_general obj_id object (Slot ` slots)"
|
||||
by (simp add: sep_map_S'_def)
|
||||
|
||||
(* Consumes the ownership of the empty slots of an object. *)
|
||||
definition
|
||||
sep_map_E :: "cdl_object_id \<Rightarrow> cdl_object \<Rightarrow> sep_pred" ("_ \<mapsto>E _" [76,71] 76)
|
||||
where
|
||||
"p \<mapsto>E obj \<equiv> (p, UNIV- dom (object_slots obj)) \<mapsto>S' obj"
|
||||
"obj_id \<mapsto>E object \<equiv> (obj_id, UNIV- dom (object_slots object)) \<mapsto>S' object"
|
||||
|
||||
(* There is a clean object there that has the same cap in the same slot. *)
|
||||
definition
|
||||
sep_map_s :: "cdl_cap_ref \<Rightarrow> cdl_object \<Rightarrow> sep_pred" ("_ \<mapsto>s _" [76,71] 76)
|
||||
where
|
||||
"p \<mapsto>s obj \<equiv> let (obj_id, slot) = p in sep_map_general obj_id obj {Slot slot}"
|
||||
"p \<mapsto>s object \<equiv> let (obj_id, slot) = p in sep_map_general obj_id object {Slot slot}"
|
||||
|
||||
lemma sep_map_s_def2:
|
||||
"(obj_id, slot) \<mapsto>s object \<equiv> sep_map_general obj_id object {Slot slot}"
|
||||
by (simp add: sep_map_s_def)
|
||||
|
||||
(* There is that cap there. *)
|
||||
definition
|
||||
|
@ -464,9 +313,9 @@ lemma offset_slot:
|
|||
(* sep_map predcates are injective. *)
|
||||
|
||||
lemma sep_map_general_inj:
|
||||
"cmps \<noteq> {} \<Longrightarrow> inj (\<lambda>obj_id. sep_map_general obj_id obj cmps)"
|
||||
"cmps \<noteq> {} \<Longrightarrow> inj (\<lambda>obj_id. sep_map_general obj_id object cmps)"
|
||||
apply (clarsimp simp: inj_on_def fun_eq_iff sep_map_general_def)
|
||||
apply (erule_tac x="SepState (obj_to_sep_state x cmps obj) empty" in allE)
|
||||
apply (erule_tac x="SepState (obj_to_sep_state x object cmps) empty" in allE)
|
||||
apply (fastforce simp: obj_to_sep_state_def
|
||||
split: split_if_asm)
|
||||
done
|
||||
|
@ -482,7 +331,7 @@ lemma sep_map_f_inj:
|
|||
lemma sep_map_s_inj:
|
||||
"inj (\<lambda>obj_id. obj_id \<mapsto>s obj)"
|
||||
apply (clarsimp simp: inj_on_def fun_eq_iff sep_map_s_def sep_map_general_def)
|
||||
apply (erule_tac x="SepState (obj_to_sep_state a {Slot b} obj) empty" in allE)
|
||||
apply (erule_tac x="SepState (obj_to_sep_state a obj {Slot b}) empty" in allE)
|
||||
apply (fastforce simp: obj_to_sep_state_def
|
||||
split: split_if_asm)
|
||||
done
|
||||
|
@ -490,9 +339,10 @@ lemma sep_map_s_inj:
|
|||
lemma sep_map_c_inj:
|
||||
"inj (\<lambda>(obj_id,slot). (obj_id,slot) \<mapsto>c cap)"
|
||||
apply (clarsimp simp: inj_on_def fun_eq_iff sep_map_c_def sep_map_general_def)
|
||||
apply (erule_tac x="SepState (obj_to_sep_state a {Slot b}
|
||||
apply (erule_tac x="SepState (obj_to_sep_state a
|
||||
(CNode \<lparr> cdl_cnode_caps = [b \<mapsto> cap],
|
||||
cdl_cnode_size_bits = 0 \<rparr>))
|
||||
cdl_cnode_size_bits = 0 \<rparr>)
|
||||
{Slot b})
|
||||
empty" in allE)
|
||||
apply (auto simp: obj_to_sep_state_def object_project_def object_slots_def
|
||||
split: split_if_asm)
|
||||
|
@ -511,9 +361,9 @@ lemma object_type_asid_reset [simp]:
|
|||
"object_type (asid_reset obj) = object_type obj"
|
||||
by (clarsimp simp: asid_reset_def)
|
||||
|
||||
lemma object_type_object_lift [simp]:
|
||||
"object_type (object_lift obj) = object_type obj"
|
||||
by (clarsimp simp: object_lift_def)
|
||||
lemma object_type_object_clean [simp]:
|
||||
"object_type (object_clean obj) = object_type obj"
|
||||
by (clarsimp simp: object_clean_def)
|
||||
|
||||
lemma dom_map_restrict:
|
||||
"dom (f |` a) \<subseteq> dom f"
|
||||
|
@ -540,6 +390,10 @@ lemma sep_map_general_disjoint:
|
|||
* Object_empty predicate decomposition *
|
||||
****************************************)
|
||||
|
||||
lemma object_type_object_wipe_slots[simp]:
|
||||
"object_type (object_wipe_slots obj) = object_type obj"
|
||||
by (clarsimp simp: object_wipe_slots_def)
|
||||
|
||||
lemma object_slots_intent_reset [simp]:
|
||||
"object_slots (intent_reset obj) = object_slots obj"
|
||||
by (clarsimp simp: intent_reset_def object_slots_def split: cdl_object.splits)
|
||||
|
@ -548,46 +402,20 @@ lemma update_slots_intent_reset [simp]:
|
|||
"intent_reset (update_slots slots obj) = update_slots slots (intent_reset obj)"
|
||||
by (clarsimp simp: intent_reset_def update_slots_def split: cdl_object.splits)
|
||||
|
||||
lemma object_clean_fields_intent_reset [simp]:
|
||||
"intent_reset (object_clean_fields obj cmps) =
|
||||
object_clean_fields (intent_reset obj) cmps"
|
||||
by (clarsimp simp: intent_reset_def object_clean_fields_def split: cdl_object.splits)
|
||||
lemma object_wipe_slots_intent_reset [simp]:
|
||||
"intent_reset (object_wipe_slots obj)
|
||||
= object_wipe_slots (intent_reset obj)"
|
||||
by (clarsimp simp: object_wipe_slots_def)
|
||||
|
||||
lemma object_clean_fields_asid_reset [simp]:
|
||||
"object_clean_fields (asid_reset obj) cmps = asid_reset (object_clean_fields obj cmps)"
|
||||
by (clarsimp simp: asid_reset_def)
|
||||
|
||||
lemma object_clean_slots_intent_reset [simp]:
|
||||
"intent_reset (object_clean_slots obj cmps)
|
||||
= object_clean_slots (intent_reset obj) cmps"
|
||||
by (clarsimp simp: object_clean_slots_def)
|
||||
|
||||
lemma object_clean_fields_object_lift [simp]:
|
||||
"object_clean_fields (object_lift obj) cmps = object_lift (object_clean_fields obj cmps)"
|
||||
by (clarsimp simp: object_lift_def2)
|
||||
|
||||
lemma object_clean_slots_asid_reset [simp]:
|
||||
"object_clean_slots (asid_reset obj) cmps = asid_reset (object_clean_slots obj cmps)"
|
||||
apply (clarsimp simp: object_clean_slots_def asid_reset_def clean_slots_def)
|
||||
lemma object_wipe_slots_asid_reset [simp]:
|
||||
"object_wipe_slots (asid_reset obj) = asid_reset (object_wipe_slots obj)"
|
||||
apply (clarsimp simp: object_wipe_slots_def asid_reset_def)
|
||||
apply (case_tac "has_slots obj", simp_all)
|
||||
done
|
||||
|
||||
lemma object_clean_intent_reset [simp]:
|
||||
"intent_reset (object_clean obj cmps) =
|
||||
object_clean (intent_reset obj) cmps"
|
||||
by (clarsimp simp: object_clean_def)
|
||||
|
||||
lemma object_clean_asid_reset [simp]:
|
||||
"object_clean (asid_reset obj) cmps = asid_reset (object_clean obj cmps)"
|
||||
by (clarsimp simp: object_clean_def)
|
||||
|
||||
lemma object_clean_object_lift [simp]:
|
||||
"object_clean (object_lift obj) cmps = object_lift (object_clean obj cmps)"
|
||||
by (clarsimp simp: object_lift_def2)
|
||||
|
||||
lemma object_slots_object_lift:
|
||||
"object_slots (object_lift obj) = reset_cap_asid \<circ>\<^sub>M object_slots obj"
|
||||
apply (clarsimp simp: object_lift_def)
|
||||
lemma object_slots_object_clean:
|
||||
"object_slots (object_clean obj) = reset_cap_asid \<circ>\<^sub>M object_slots obj"
|
||||
apply (clarsimp simp: object_clean_def)
|
||||
apply (clarsimp simp: asid_reset_def object_slots_def update_slots_def
|
||||
split: cdl_object.splits)
|
||||
done
|
||||
|
@ -601,29 +429,13 @@ lemma restrict_map_remerge':
|
|||
done
|
||||
|
||||
|
||||
lemma "((reset_cap_asid \<circ>\<^sub>M object_slots obj |` the_set cmps) ++
|
||||
(reset_cap_asid \<circ>\<^sub>M object_slots obj |` the_set cmps'))
|
||||
=
|
||||
(reset_cap_asid \<circ>\<^sub>M
|
||||
object_slots
|
||||
(update_slots (object_slots obj |` (the_set cmps \<union> the_set cmps'))
|
||||
obj))"
|
||||
apply (case_tac "has_slots obj", simp_all)
|
||||
apply (rule restrict_map_remerge')
|
||||
done
|
||||
|
||||
lemma object_clean_fields_union [simp]:
|
||||
"Fields \<notin> cmps'
|
||||
\<Longrightarrow> object_clean_fields obj (cmps \<union> cmps') = object_clean_fields obj cmps"
|
||||
by (clarsimp simp: object_clean_fields_def)
|
||||
|
||||
lemma product_union:
|
||||
"{p} \<times> (cmps \<union> cmps') = ({p} \<times> cmps) \<union> ({p} \<times> cmps')"
|
||||
by (rule Sigma_Un_distrib2)
|
||||
|
||||
lemma obj_to_sep_state_add:
|
||||
"obj_to_sep_state p cmps obj ++ obj_to_sep_state p cmps' obj =
|
||||
obj_to_sep_state p (cmps \<union> cmps') obj"
|
||||
"obj_to_sep_state p obj cmps ++ obj_to_sep_state p obj cmps' =
|
||||
obj_to_sep_state p obj (cmps \<union> cmps')"
|
||||
apply (rule ext)
|
||||
apply (clarsimp simp:obj_to_sep_state_def split_def
|
||||
map_add_def split:if_splits option.splits)
|
||||
|
@ -640,8 +452,8 @@ lemma sep_map_decomp:
|
|||
apply (clarsimp simp: sep_map_general_def plus_sep_state_def
|
||||
sep_state_add_def obj_to_sep_state_add)
|
||||
apply (clarsimp simp:sep_map_general_def)
|
||||
apply (rule_tac x = "SepState (obj_to_sep_state p cmps obj) empty" in exI)
|
||||
apply (rule_tac x = "SepState (obj_to_sep_state p cmps' obj) empty" in exI)
|
||||
apply (rule_tac x = "SepState (obj_to_sep_state p obj cmps) empty" in exI)
|
||||
apply (rule_tac x = "SepState (obj_to_sep_state p obj cmps') empty" in exI)
|
||||
apply (clarsimp simp: plus_sep_state_def sep_state_add_def obj_to_sep_state_add)
|
||||
apply (intro conjI disjointI)
|
||||
apply (clarsimp simp:dom_def obj_to_sep_state_def
|
||||
|
@ -659,11 +471,19 @@ lemma sep_map_o_decomp:
|
|||
done
|
||||
|
||||
lemma sep_map_c_def2:
|
||||
"((obj_id,slot) \<mapsto>c cap) s \<Longrightarrow> \<exists>obj. ((obj_id,slot) \<mapsto>s obj) s \<and> ((object_slots obj) slot = Some cap)"
|
||||
"((obj_id,slot) \<mapsto>c cap) s \<Longrightarrow>
|
||||
\<exists>obj. ((obj_id,slot) \<mapsto>s obj) s \<and> ((object_slots obj) slot = Some cap)"
|
||||
apply (clarsimp simp: sep_map_c_def sep_map_s_def sep_map_S_def sep_map_general_def)
|
||||
apply (rule_tac x="update_slots [slot \<mapsto> cap] obj" in exI)
|
||||
apply (case_tac "has_slots obj")
|
||||
apply (clarsimp simp: object_clean_def object_clean_slots_def)+
|
||||
apply (clarsimp simp: object_wipe_slots_def)+
|
||||
done
|
||||
|
||||
lemma intent_reset_has_slots[simp]:
|
||||
"has_slots (intent_reset obj) = has_slots obj"
|
||||
apply (simp add:intent_reset_def
|
||||
split:cdl_object.splits)
|
||||
apply (simp add:has_slots_def)
|
||||
done
|
||||
|
||||
lemma sep_map_c_sep_map_s:
|
||||
|
@ -676,38 +496,27 @@ lemma sep_map_c_sep_map_s:
|
|||
apply (case_tac obj)
|
||||
apply (clarsimp simp:obj_to_sep_state_def
|
||||
object_project_def update_slots_def
|
||||
object_lift_def intent_reset_def
|
||||
object_clean_def intent_reset_def
|
||||
asid_reset_def object_slots_def)+
|
||||
apply (rule ext)
|
||||
apply (clarsimp simp: update_slots_def object_slots_def
|
||||
split: cdl_object.splits)
|
||||
done
|
||||
|
||||
lemma intent_reset_has_slots[simp]:
|
||||
"has_slots (intent_reset obj) = has_slots obj"
|
||||
apply (simp add:intent_reset_def
|
||||
split:cdl_object.splits)
|
||||
apply (simp add:has_slots_def)
|
||||
(* Fastforce doesn't seem to be able to prove this is if the assumptions are in the other order. *)
|
||||
lemma sep_map_s_sep_map_c [rotated]:
|
||||
"\<lbrakk>object_slots obj slot = Some cap; ((obj_id, slot) \<mapsto>c cap) s\<rbrakk> \<Longrightarrow>
|
||||
((obj_id, slot) \<mapsto>s obj) s"
|
||||
apply (fastforce simp: sep_map_c_def sep_map_s_def sep_map_S_def sep_map_general_def
|
||||
obj_to_sep_state_def object_project_def update_slots_def
|
||||
object_clean_def intent_reset_def asid_reset_def object_slots_def
|
||||
split: cdl_object.splits)
|
||||
done
|
||||
|
||||
lemma sep_map_s_sep_map_c:
|
||||
"\<lbrakk>((obj_id, slot) \<mapsto>c cap \<and>* obj_id \<mapsto>f obj') s;has_slots obj'\<rbrakk> \<Longrightarrow>
|
||||
\<exists>obj. ((obj_id, slot) \<mapsto>s obj \<and>* obj_id \<mapsto>f obj') s \<and>
|
||||
object_slots obj slot = Some cap \<and>
|
||||
object_type obj = object_type obj'"
|
||||
apply (rule_tac x="update_slots [slot \<mapsto> cap] obj'" in exI)
|
||||
apply (clarsimp simp: sep_map_c_def sep_map_f_def sep_conj_exists)
|
||||
apply (frule sep_map_general_disjoint)
|
||||
apply (clarsimp simp: sep_map_c_def sep_map_s_def sep_map_S_def sep_map_general_def)
|
||||
apply (clarsimp simp: sep_conj_def)
|
||||
apply (clarsimp simp: sep_disj_sep_state_def sep_state_disj_def)
|
||||
apply (rule_tac x = x in exI)
|
||||
apply (rule_tac x = y in exI)
|
||||
apply clarsimp
|
||||
apply (rule ext)
|
||||
apply (clarsimp simp:obj_to_sep_state_def
|
||||
object_project_def asid_reset_def object_slots_object_lift)
|
||||
done
|
||||
lemma sep_map_s_sep_map_c_eq:
|
||||
"\<lbrakk>object_slots obj slot = Some cap\<rbrakk> \<Longrightarrow>
|
||||
(obj_id, slot) \<mapsto>s obj = (obj_id, slot) \<mapsto>c cap"
|
||||
by (fastforce simp: sep_map_c_sep_map_s sep_map_s_sep_map_c)
|
||||
|
||||
|
||||
lemma sep_map_s_object_slots_equal:
|
||||
|
@ -716,25 +525,14 @@ lemma sep_map_s_object_slots_equal:
|
|||
apply (clarsimp simp: sep_map_s_def sep_map_general_def split: sep_state.splits)
|
||||
apply (intro iffI ext conjI |
|
||||
clarsimp simp: obj_to_sep_state_def object_project_def
|
||||
object_slots_object_lift)+
|
||||
done
|
||||
|
||||
lemma sep_map_s_sep_map_c':
|
||||
"\<lbrakk>((obj_id, slot) \<mapsto>c cap \<and>* obj_id \<mapsto>f obj) s;
|
||||
object_slots obj slot = Some cap\<rbrakk>
|
||||
\<Longrightarrow> ((obj_id, slot) \<mapsto>s obj \<and>* obj_id \<mapsto>f obj) s"
|
||||
apply (drule sep_map_s_sep_map_c)
|
||||
apply (rule object_slots_has_slots, clarsimp)
|
||||
apply clarsimp
|
||||
apply (cut_tac obj=obja and obj'=obj and obj_id=obj_id and slot=slot in
|
||||
sep_map_s_object_slots_equal, simp+)
|
||||
object_slots_object_clean)+
|
||||
done
|
||||
|
||||
lemma sep_map_S_def2:
|
||||
"(obj_id \<mapsto>S obj) = ((obj_id,UNIV) \<mapsto>S' obj)"
|
||||
by (clarsimp simp: sep_map_S_def sep_map_S'_def)
|
||||
|
||||
lemma sep_map_S'_def2:
|
||||
lemma sep_map_S'_def_singleton:
|
||||
"(obj_id, {slot}) \<mapsto>S' obj = (obj_id, slot) \<mapsto>s obj"
|
||||
by (rule ext)(clarsimp simp: sep_map_S'_def sep_map_s_def)
|
||||
|
||||
|
|
|
@ -70,10 +70,12 @@ datatype asid_pool_invocation =
|
|||
|
||||
datatype page_invocation
|
||||
= PageMap
|
||||
asid
|
||||
cap
|
||||
cslot_ptr
|
||||
"pte \<times> (obj_ref list) + pde \<times> (obj_ref list)"
|
||||
| PageRemap
|
||||
| PageRemap
|
||||
asid
|
||||
"pte \<times> (obj_ref list) + pde \<times> (obj_ref list)"
|
||||
| PageUnmap
|
||||
arch_cap
|
||||
|
@ -89,20 +91,30 @@ datatype page_invocation
|
|||
primrec
|
||||
page_map_cap :: "page_invocation \<Rightarrow> cap"
|
||||
where
|
||||
"page_map_cap (PageMap c p x) = c"
|
||||
"page_map_cap (PageMap a c p x) = c"
|
||||
|
||||
primrec
|
||||
page_map_asid :: "page_invocation \<Rightarrow> asid"
|
||||
where
|
||||
"page_map_asid (PageMap a c p x) = a"
|
||||
primrec
|
||||
page_map_ct_slot :: "page_invocation \<Rightarrow> cslot_ptr"
|
||||
where
|
||||
"page_map_ct_slot (PageMap c p x) = p"
|
||||
"page_map_ct_slot (PageMap a c p x) = p"
|
||||
primrec
|
||||
page_map_entries :: "page_invocation \<Rightarrow> pte \<times> (obj_ref list) + pde \<times> (obj_ref list)"
|
||||
where
|
||||
"page_map_entries (PageMap c p x) = x"
|
||||
"page_map_entries (PageMap a c p x) = x"
|
||||
|
||||
primrec
|
||||
page_remap_entries :: "page_invocation \<Rightarrow> pte \<times> (obj_ref list) + pde \<times> (obj_ref list)"
|
||||
where
|
||||
"page_remap_entries (PageRemap x) = x"
|
||||
"page_remap_entries (PageRemap a x) = x"
|
||||
|
||||
primrec
|
||||
page_remap_asid :: "page_invocation \<Rightarrow> asid"
|
||||
where
|
||||
"page_remap_asid (PageRemap a x) = a"
|
||||
|
||||
primrec
|
||||
page_unmap_cap :: "page_invocation \<Rightarrow> arch_cap"
|
||||
|
|
|
@ -104,6 +104,22 @@ where
|
|||
od
|
||||
| PageDirectoryNothing \<Rightarrow> return ()"
|
||||
|
||||
definition
|
||||
pte_check_if_mapped :: "32 word \<Rightarrow> (bool, 'z::state_ext) s_monad"
|
||||
where
|
||||
"pte_check_if_mapped slot \<equiv> do
|
||||
pt \<leftarrow> get_master_pte slot;
|
||||
return (pt \<noteq> InvalidPTE)
|
||||
od"
|
||||
|
||||
definition
|
||||
pde_check_if_mapped :: "32 word \<Rightarrow> (bool, 'z::state_ext) s_monad"
|
||||
where
|
||||
"pde_check_if_mapped slot \<equiv> do
|
||||
pd \<leftarrow> get_master_pde slot;
|
||||
return (pd \<noteq> InvalidPDE)
|
||||
od"
|
||||
|
||||
text {* The Page capability confers the authority to map, unmap and flush the
|
||||
memory page. The remap system call is a convenience operation that ensures the
|
||||
page is mapped in the same location as this cap was previously used to map it
|
||||
|
@ -111,33 +127,41 @@ in. *}
|
|||
definition
|
||||
perform_page_invocation :: "page_invocation \<Rightarrow> (unit,'z::state_ext) s_monad" where
|
||||
"perform_page_invocation iv \<equiv> case iv of
|
||||
PageMap cap ct_slot entries \<Rightarrow> do
|
||||
PageMap asid cap ct_slot entries \<Rightarrow> do
|
||||
set_cap cap ct_slot;
|
||||
case entries of
|
||||
Inl (pte, slots) \<Rightarrow> do
|
||||
flush \<leftarrow> pte_check_if_mapped (hd slots);
|
||||
store_pte (hd slots) pte;
|
||||
mapM (swp store_pte InvalidPTE) (tl slots);
|
||||
do_machine_op $ cleanCacheRange_PoU (hd slots) (last_byte_pte (last slots))
|
||||
(addrFromPPtr (hd slots))
|
||||
(addrFromPPtr (hd slots));
|
||||
if flush then (invalidate_tlb_by_asid asid) else return ()
|
||||
od
|
||||
| Inr (pde, slots) \<Rightarrow> do
|
||||
flush \<leftarrow> pde_check_if_mapped (hd slots);
|
||||
store_pde (hd slots) pde;
|
||||
mapM (swp store_pde InvalidPDE) (tl slots);
|
||||
do_machine_op $ cleanCacheRange_PoU (hd slots) (last_byte_pde (last slots))
|
||||
(addrFromPPtr (hd slots))
|
||||
(addrFromPPtr (hd slots));
|
||||
if flush then (invalidate_tlb_by_asid asid) else return ()
|
||||
od
|
||||
od
|
||||
| PageRemap (Inl (pte, slots)) \<Rightarrow> do
|
||||
| PageRemap asid (Inl (pte, slots)) \<Rightarrow> do
|
||||
flush \<leftarrow> pte_check_if_mapped (hd slots);
|
||||
store_pte (hd slots) pte;
|
||||
mapM_x (swp store_pte InvalidPTE) (tl slots);
|
||||
do_machine_op $ cleanCacheRange_PoU (hd slots) (last_byte_pte (last slots))
|
||||
(addrFromPPtr (hd slots))
|
||||
(addrFromPPtr (hd slots));
|
||||
if flush then (invalidate_tlb_by_asid asid) else return ()
|
||||
od
|
||||
| PageRemap (Inr (pde, slots)) \<Rightarrow> do
|
||||
| PageRemap asid (Inr (pde, slots)) \<Rightarrow> do
|
||||
flush \<leftarrow> pde_check_if_mapped (hd slots);
|
||||
store_pde (hd slots) pde;
|
||||
mapM_x (swp store_pde InvalidPDE) (tl slots);
|
||||
do_machine_op $ cleanCacheRange_PoU (hd slots) (last_byte_pde (last slots))
|
||||
(addrFromPPtr (hd slots))
|
||||
(addrFromPPtr (hd slots));
|
||||
if flush then (invalidate_tlb_by_asid asid) else return ()
|
||||
od
|
||||
| PageUnmap cap ct_slot \<Rightarrow>
|
||||
case cap of
|
||||
|
|
|
@ -176,7 +176,7 @@ where
|
|||
vaddr pgsz vm_rights
|
||||
(attribs_from_word attr) pd;
|
||||
ensure_safe_mapping entries;
|
||||
returnOk $ InvokePage $ PageMap
|
||||
returnOk $ InvokePage $ PageMap asid
|
||||
(ArchObjectCap $ PageCap p R pgsz (Some (asid, vaddr)))
|
||||
cte entries
|
||||
odE
|
||||
|
@ -202,7 +202,7 @@ where
|
|||
vaddr pgsz vm_rights
|
||||
(attribs_from_word attr) pd;
|
||||
ensure_safe_mapping entries;
|
||||
returnOk $ InvokePage $ PageRemap entries
|
||||
returnOk $ InvokePage $ PageRemap asid' entries
|
||||
odE
|
||||
else throwError TruncatedMessage
|
||||
else if invocation_type label = ARMPageUnmap
|
||||
|
|
|
@ -40,7 +40,7 @@ where
|
|||
target_slot \<leftarrow> lookup_slot_for_cnode_op cspace_cap index (unat depth);
|
||||
ensure_empty target_slot;
|
||||
|
||||
returnOk $ MakePool (set_avaiable_range untyped_cap {}) untyped_cap_ref
|
||||
returnOk $ MakePool (set_available_range untyped_cap {}) untyped_cap_ref
|
||||
(cap_objects untyped_cap) target_slot (base, 0)
|
||||
odE \<sqinter> throw"
|
||||
|
||||
|
|
|
@ -165,7 +165,6 @@ where
|
|||
"invoke_cnode params \<equiv> case params of
|
||||
(* Insert a new cap. *)
|
||||
InsertCall cap src_slot dest_slot \<Rightarrow>
|
||||
(* FIXME: probably want to be more precise here *)
|
||||
liftE $
|
||||
insert_cap_sibling cap src_slot dest_slot
|
||||
\<sqinter>
|
||||
|
|
|
@ -58,18 +58,18 @@ where
|
|||
|
||||
|
||||
|
||||
primrec avaiable_range :: "cdl_cap \<Rightarrow> cdl_object_id set"
|
||||
where "avaiable_range (UntypedCap r avaiable) = avaiable"
|
||||
primrec available_range :: "cdl_cap \<Rightarrow> cdl_object_id set"
|
||||
where "available_range (UntypedCap r available) = available"
|
||||
|
||||
primrec set_avaiable_range :: "cdl_cap \<Rightarrow> cdl_object_id set \<Rightarrow> cdl_cap"
|
||||
where "set_avaiable_range (UntypedCap r avaiable) nrange = UntypedCap r nrange"
|
||||
primrec set_available_range :: "cdl_cap \<Rightarrow> cdl_object_id set \<Rightarrow> cdl_cap"
|
||||
where "set_available_range (UntypedCap r available) nrange = UntypedCap r nrange"
|
||||
|
||||
definition
|
||||
set_untyped_cap_as_full :: "cdl_cap \<Rightarrow> cdl_cap \<Rightarrow> cdl_cap_ref \<Rightarrow> unit k_monad"
|
||||
where "set_untyped_cap_as_full src_cap new_cap src_slot \<equiv>
|
||||
if (is_untyped_cap src_cap \<and> is_untyped_cap new_cap
|
||||
\<and> cap_objects src_cap = cap_objects new_cap) then
|
||||
(set_cap src_slot (set_avaiable_range src_cap {}))
|
||||
(set_cap src_slot (set_available_range src_cap {}))
|
||||
else return ()"
|
||||
|
||||
(* Insert a new cap into an object. The cap will be a sibling. *)
|
||||
|
|
|
@ -113,12 +113,6 @@ where
|
|||
|
||||
(* --- caps --- *)
|
||||
|
||||
(* FIXME: It is unclear whether this definition is precise enough.
|
||||
We might need to make the type "cdl_cap option lookup" to
|
||||
distinguish between caps that are not present (NullCap) and
|
||||
caps that can't exist at the specificied location.
|
||||
At the moment these two outcomes are identified. *)
|
||||
|
||||
(* The slots of an object, returns an empty map for non-existing objects
|
||||
or objects that do not have caps *)
|
||||
definition
|
||||
|
|
|
@ -75,11 +75,8 @@ where
|
|||
(* FrameCaps, PageTableCaps and PageDirectoryCaps can either be
|
||||
* "real" cap or "fake" cap. Real caps are installed in CNodes,
|
||||
* and fake caps represent a page table mapping.
|
||||
*
|
||||
* The cdl_cap_ref for fake caps point to the real cap it was installed from.
|
||||
* TODO: This is not currently implemented.
|
||||
*)
|
||||
datatype cdl_frame_cap_type = Real | Fake (* cdl_cap_ref *)
|
||||
datatype cdl_frame_cap_type = Real | Fake
|
||||
|
||||
(*
|
||||
* Kernel capabilities.
|
||||
|
|
|
@ -88,7 +88,6 @@ where
|
|||
|
||||
|
||||
(* Insert a cap for a new object in the given location. *)
|
||||
(* FIXME: we do not do update_untyped_as_full here for refinedment ? *)
|
||||
definition
|
||||
create_cap :: "cdl_object_type \<Rightarrow> nat \<Rightarrow> cdl_cap_ref \<Rightarrow> (cdl_cap_ref \<times> cdl_object_id set) \<Rightarrow> unit k_monad"
|
||||
where
|
||||
|
@ -106,7 +105,7 @@ definition
|
|||
where "update_available_range orange newids cap_ref cap \<equiv>
|
||||
do
|
||||
new_range \<leftarrow> select {x. x \<subseteq> orange - set newids};
|
||||
set_cap cap_ref $ set_avaiable_range cap new_range
|
||||
set_cap cap_ref $ set_available_range cap new_range
|
||||
od"
|
||||
|
||||
|
||||
|
@ -143,7 +142,7 @@ where
|
|||
(* If the untyped hasn't already been typed into something we reuse our names. *)
|
||||
unless has_kids $ modify (detype (cap_objects untyped_cap));
|
||||
|
||||
new_range \<leftarrow> return $ if has_kids then (avaiable_range untyped_cap) else (cap_objects untyped_cap);
|
||||
new_range \<leftarrow> return $ if has_kids then (available_range untyped_cap) else (cap_objects untyped_cap);
|
||||
new_obj_refs \<leftarrow> generate_object_ids num_objects new_type new_range;
|
||||
|
||||
update_available_range new_range (map (\<lambda>s. pick s) new_obj_refs) untyped_ref untyped_cap;
|
||||
|
|
|
@ -198,14 +198,14 @@ where
|
|||
|
||||
datatype page_invocation =
|
||||
PageFlush flush_type vptr vptr paddr machine_word asid
|
||||
| PageRemap "(pte * machine_word list) + (pde * machine_word list)"
|
||||
| PageMap capability machine_word "(pte * machine_word list) + (pde * machine_word list)"
|
||||
| PageRemap asid "(pte * machine_word list) + (pde * machine_word list)"
|
||||
| PageMap asid capability machine_word "(pte * machine_word list) + (pde * machine_word list)"
|
||||
| PageUnmap arch_capability machine_word
|
||||
|
||||
primrec
|
||||
pageMapCap :: "page_invocation \<Rightarrow> capability"
|
||||
where
|
||||
"pageMapCap (PageMap v0 v1 v2) = v0"
|
||||
"pageMapCap (PageMap v0 v1 v2 v3) = v1"
|
||||
|
||||
primrec
|
||||
pageUnmapCapSlot :: "page_invocation \<Rightarrow> machine_word"
|
||||
|
@ -245,12 +245,12 @@ where
|
|||
primrec
|
||||
pageRemapEntries :: "page_invocation \<Rightarrow> (pte * machine_word list) + (pde * machine_word list)"
|
||||
where
|
||||
"pageRemapEntries (PageRemap v0) = v0"
|
||||
"pageRemapEntries (PageRemap v0 v1) = v1"
|
||||
|
||||
primrec
|
||||
pageMapCTSlot :: "page_invocation \<Rightarrow> machine_word"
|
||||
pageMapASID :: "page_invocation \<Rightarrow> asid"
|
||||
where
|
||||
"pageMapCTSlot (PageMap v0 v1 v2) = v1"
|
||||
"pageMapASID (PageMap v0 v1 v2 v3) = v0"
|
||||
|
||||
primrec
|
||||
pageFlushEnd :: "page_invocation \<Rightarrow> vptr"
|
||||
|
@ -260,12 +260,22 @@ where
|
|||
primrec
|
||||
pageMapEntries :: "page_invocation \<Rightarrow> (pte * machine_word list) + (pde * machine_word list)"
|
||||
where
|
||||
"pageMapEntries (PageMap v0 v1 v2) = v2"
|
||||
"pageMapEntries (PageMap v0 v1 v2 v3) = v3"
|
||||
|
||||
primrec
|
||||
pageRemapASID :: "page_invocation \<Rightarrow> asid"
|
||||
where
|
||||
"pageRemapASID (PageRemap v0 v1) = v0"
|
||||
|
||||
primrec
|
||||
pageMapCTSlot :: "page_invocation \<Rightarrow> machine_word"
|
||||
where
|
||||
"pageMapCTSlot (PageMap v0 v1 v2 v3) = v2"
|
||||
|
||||
primrec
|
||||
pageMapCap_update :: "(capability \<Rightarrow> capability) \<Rightarrow> page_invocation \<Rightarrow> page_invocation"
|
||||
where
|
||||
"pageMapCap_update f (PageMap v0 v1 v2) = PageMap (f v0) v1 v2"
|
||||
"pageMapCap_update f (PageMap v0 v1 v2 v3) = PageMap v0 (f v1) v2 v3"
|
||||
|
||||
primrec
|
||||
pageUnmapCapSlot_update :: "(machine_word \<Rightarrow> machine_word) \<Rightarrow> page_invocation \<Rightarrow> page_invocation"
|
||||
|
@ -305,12 +315,12 @@ where
|
|||
primrec
|
||||
pageRemapEntries_update :: "(((pte * machine_word list) + (pde * machine_word list)) \<Rightarrow> ((pte * machine_word list) + (pde * machine_word list))) \<Rightarrow> page_invocation \<Rightarrow> page_invocation"
|
||||
where
|
||||
"pageRemapEntries_update f (PageRemap v0) = PageRemap (f v0)"
|
||||
"pageRemapEntries_update f (PageRemap v0 v1) = PageRemap v0 (f v1)"
|
||||
|
||||
primrec
|
||||
pageMapCTSlot_update :: "(machine_word \<Rightarrow> machine_word) \<Rightarrow> page_invocation \<Rightarrow> page_invocation"
|
||||
pageMapASID_update :: "(asid \<Rightarrow> asid) \<Rightarrow> page_invocation \<Rightarrow> page_invocation"
|
||||
where
|
||||
"pageMapCTSlot_update f (PageMap v0 v1 v2) = PageMap v0 (f v1) v2"
|
||||
"pageMapASID_update f (PageMap v0 v1 v2 v3) = PageMap (f v0) v1 v2 v3"
|
||||
|
||||
primrec
|
||||
pageFlushEnd_update :: "(vptr \<Rightarrow> vptr) \<Rightarrow> page_invocation \<Rightarrow> page_invocation"
|
||||
|
@ -320,7 +330,17 @@ where
|
|||
primrec
|
||||
pageMapEntries_update :: "(((pte * machine_word list) + (pde * machine_word list)) \<Rightarrow> ((pte * machine_word list) + (pde * machine_word list))) \<Rightarrow> page_invocation \<Rightarrow> page_invocation"
|
||||
where
|
||||
"pageMapEntries_update f (PageMap v0 v1 v2) = PageMap v0 v1 (f v2)"
|
||||
"pageMapEntries_update f (PageMap v0 v1 v2 v3) = PageMap v0 v1 v2 (f v3)"
|
||||
|
||||
primrec
|
||||
pageRemapASID_update :: "(asid \<Rightarrow> asid) \<Rightarrow> page_invocation \<Rightarrow> page_invocation"
|
||||
where
|
||||
"pageRemapASID_update f (PageRemap v0 v1) = PageRemap (f v0) v1"
|
||||
|
||||
primrec
|
||||
pageMapCTSlot_update :: "(machine_word \<Rightarrow> machine_word) \<Rightarrow> page_invocation \<Rightarrow> page_invocation"
|
||||
where
|
||||
"pageMapCTSlot_update f (PageMap v0 v1 v2 v3) = PageMap v0 v1 (f v2) v3"
|
||||
|
||||
abbreviation (input)
|
||||
PageFlush_trans :: "(flush_type) \<Rightarrow> (vptr) \<Rightarrow> (vptr) \<Rightarrow> (paddr) \<Rightarrow> (machine_word) \<Rightarrow> (asid) \<Rightarrow> page_invocation" ("PageFlush'_ \<lparr> pageFlushType= _, pageFlushStart= _, pageFlushEnd= _, pageFlushPStart= _, pageFlushPD= _, pageFlushASID= _ \<rparr>")
|
||||
|
@ -328,14 +348,14 @@ where
|
|||
"PageFlush_ \<lparr> pageFlushType= v0, pageFlushStart= v1, pageFlushEnd= v2, pageFlushPStart= v3, pageFlushPD= v4, pageFlushASID= v5 \<rparr> == PageFlush v0 v1 v2 v3 v4 v5"
|
||||
|
||||
abbreviation (input)
|
||||
PageRemap_trans :: "((pte * machine_word list) + (pde * machine_word list)) \<Rightarrow> page_invocation" ("PageRemap'_ \<lparr> pageRemapEntries= _ \<rparr>")
|
||||
PageRemap_trans :: "(asid) \<Rightarrow> ((pte * machine_word list) + (pde * machine_word list)) \<Rightarrow> page_invocation" ("PageRemap'_ \<lparr> pageRemapASID= _, pageRemapEntries= _ \<rparr>")
|
||||
where
|
||||
"PageRemap_ \<lparr> pageRemapEntries= v0 \<rparr> == PageRemap v0"
|
||||
"PageRemap_ \<lparr> pageRemapASID= v0, pageRemapEntries= v1 \<rparr> == PageRemap v0 v1"
|
||||
|
||||
abbreviation (input)
|
||||
PageMap_trans :: "(capability) \<Rightarrow> (machine_word) \<Rightarrow> ((pte * machine_word list) + (pde * machine_word list)) \<Rightarrow> page_invocation" ("PageMap'_ \<lparr> pageMapCap= _, pageMapCTSlot= _, pageMapEntries= _ \<rparr>")
|
||||
PageMap_trans :: "(asid) \<Rightarrow> (capability) \<Rightarrow> (machine_word) \<Rightarrow> ((pte * machine_word list) + (pde * machine_word list)) \<Rightarrow> page_invocation" ("PageMap'_ \<lparr> pageMapASID= _, pageMapCap= _, pageMapCTSlot= _, pageMapEntries= _ \<rparr>")
|
||||
where
|
||||
"PageMap_ \<lparr> pageMapCap= v0, pageMapCTSlot= v1, pageMapEntries= v2 \<rparr> == PageMap v0 v1 v2"
|
||||
"PageMap_ \<lparr> pageMapASID= v0, pageMapCap= v1, pageMapCTSlot= v2, pageMapEntries= v3 \<rparr> == PageMap v0 v1 v2 v3"
|
||||
|
||||
abbreviation (input)
|
||||
PageUnmap_trans :: "(arch_capability) \<Rightarrow> (machine_word) \<Rightarrow> page_invocation" ("PageUnmap'_ \<lparr> pageUnmapCap= _, pageUnmapCapSlot= _ \<rparr>")
|
||||
|
@ -353,14 +373,14 @@ definition
|
|||
isPageRemap :: "page_invocation \<Rightarrow> bool"
|
||||
where
|
||||
"isPageRemap v \<equiv> case v of
|
||||
PageRemap v0 \<Rightarrow> True
|
||||
PageRemap v0 v1 \<Rightarrow> True
|
||||
| _ \<Rightarrow> False"
|
||||
|
||||
definition
|
||||
isPageMap :: "page_invocation \<Rightarrow> bool"
|
||||
where
|
||||
"isPageMap v \<equiv> case v of
|
||||
PageMap v0 v1 v2 \<Rightarrow> True
|
||||
PageMap v0 v1 v2 v3 \<Rightarrow> True
|
||||
| _ \<Rightarrow> False"
|
||||
|
||||
definition
|
||||
|
|
|
@ -224,6 +224,12 @@ performPageDirectoryInvocation :: "page_directory_invocation \<Rightarrow> unit
|
|||
consts
|
||||
performPageTableInvocation :: "page_table_invocation \<Rightarrow> unit kernel"
|
||||
|
||||
consts
|
||||
pteCheckIfMapped :: "machine_word \<Rightarrow> bool kernel"
|
||||
|
||||
consts
|
||||
pdeCheckIfMapped :: "machine_word \<Rightarrow> bool kernel"
|
||||
|
||||
consts
|
||||
performPageInvocation :: "page_invocation \<Rightarrow> unit kernel"
|
||||
|
||||
|
|
|
@ -1070,6 +1070,7 @@ defs decodeARMMMUInvocation_def:
|
|||
(attribsFromWord attr) pd;
|
||||
ensureSafeMapping entries;
|
||||
returnOk $ InvokePage $ PageMap_ \<lparr>
|
||||
pageMapASID= asid,
|
||||
pageMapCap= ArchObjectCap $
|
||||
cap \<lparr> capVPMappedAddress:= Just (asid, VPtr vaddr) \<rparr>,
|
||||
pageMapCTSlot= cte,
|
||||
|
@ -1094,6 +1095,7 @@ defs decodeARMMMUInvocation_def:
|
|||
vaddr (capVPSize cap) vmRights (attribsFromWord attr) pd;
|
||||
ensureSafeMapping entries;
|
||||
returnOk $ InvokePage $ PageRemap_ \<lparr>
|
||||
pageRemapASID= asidCheck,
|
||||
pageRemapEntries= entries \<rparr>
|
||||
odE)
|
||||
| (ARMPageRemap, _, _) \<Rightarrow> throw TruncatedMessage
|
||||
|
@ -1249,40 +1251,60 @@ defs performPageTableInvocation_def:
|
|||
od)
|
||||
)"
|
||||
|
||||
defs pteCheckIfMapped_def:
|
||||
"pteCheckIfMapped slot\<equiv> (do
|
||||
pt \<leftarrow> getObject slot;
|
||||
return $ pt \<noteq> InvalidPTE
|
||||
od)"
|
||||
|
||||
defs pdeCheckIfMapped_def:
|
||||
"pdeCheckIfMapped slot\<equiv> (do
|
||||
pd \<leftarrow> getObject slot;
|
||||
return $ pd \<noteq> InvalidPDE
|
||||
od)"
|
||||
|
||||
defs performPageInvocation_def:
|
||||
"performPageInvocation x0\<equiv> (case x0 of
|
||||
(PageMap cap ctSlot entries) \<Rightarrow> (do
|
||||
(PageMap asid cap ctSlot entries) \<Rightarrow> (do
|
||||
updateCap ctSlot cap;
|
||||
(case entries of
|
||||
Inl (pte, slots) \<Rightarrow> (do
|
||||
tlbFlush \<leftarrow> pteCheckIfMapped (head slots);
|
||||
mapM (flip storePTE pte) slots;
|
||||
doMachineOp $
|
||||
cleanCacheRange_PoU (VPtr $ fromPPtr $ head slots)
|
||||
(VPtr $ (fromPPtr (last slots)) + (bit (objBits (undefined::pte)) - 1))
|
||||
(addrFromPPtr (head slots))
|
||||
(addrFromPPtr (head slots));
|
||||
when tlbFlush $ invalidateTLBByASID asid
|
||||
od)
|
||||
| Inr (pde, slots) \<Rightarrow> (do
|
||||
tlbFlush \<leftarrow> pdeCheckIfMapped (head slots);
|
||||
mapM (flip storePDE pde) slots;
|
||||
doMachineOp $
|
||||
cleanCacheRange_PoU (VPtr $ fromPPtr $ head slots)
|
||||
(VPtr $ (fromPPtr (last slots)) + (bit (objBits (undefined::pde)) - 1))
|
||||
(addrFromPPtr (head slots))
|
||||
(addrFromPPtr (head slots));
|
||||
when tlbFlush $ invalidateTLBByASID asid
|
||||
od)
|
||||
)
|
||||
od)
|
||||
| (PageRemap (Inl (pte, slots))) \<Rightarrow> (do
|
||||
mapM_x (flip storePTE pte) slots;
|
||||
| (PageRemap asid (Inl (pte, slots))) \<Rightarrow> (do
|
||||
tlbFlush \<leftarrow> pteCheckIfMapped (head slots);
|
||||
mapM (flip storePTE pte) slots;
|
||||
doMachineOp $
|
||||
cleanCacheRange_PoU (VPtr $ fromPPtr $ head slots)
|
||||
(VPtr $ (fromPPtr (last slots)) + (bit (objBits (undefined::pte)) - 1))
|
||||
(addrFromPPtr (head slots))
|
||||
(addrFromPPtr (head slots));
|
||||
when tlbFlush $ invalidateTLBByASID asid
|
||||
od)
|
||||
| (PageRemap (Inr (pde, slots))) \<Rightarrow> (do
|
||||
mapM_x (flip storePDE pde) slots;
|
||||
| (PageRemap asid (Inr (pde, slots))) \<Rightarrow> (do
|
||||
tlbFlush \<leftarrow> pdeCheckIfMapped (head slots);
|
||||
mapM (flip storePDE pde) slots;
|
||||
doMachineOp $
|
||||
cleanCacheRange_PoU (VPtr $ fromPPtr $ head slots)
|
||||
(VPtr $ (fromPPtr (last slots)) + (bit (objBits (undefined::pde)) - 1))
|
||||
(addrFromPPtr (head slots))
|
||||
(addrFromPPtr (head slots));
|
||||
when tlbFlush $ invalidateTLBByASID asid
|
||||
od)
|
||||
| (PageUnmap cap ctSlot) \<Rightarrow> (do
|
||||
(case capVPMappedAddress cap of
|
||||
|
|
|
@ -482,6 +482,12 @@ defs isAligned_def:
|
|||
"isAligned x n \<equiv> x && mask n = 0"
|
||||
|
||||
|
||||
consts
|
||||
newKSDomSchedule :: "(domain \<times> machine_word) list"
|
||||
newKSDomScheduleIdx :: nat
|
||||
newKSCurDomain :: domain
|
||||
newKSDomainTime :: machine_word
|
||||
|
||||
definition
|
||||
newKernelState :: "machine_word \<Rightarrow> kernel_state"
|
||||
where
|
||||
|
@ -489,10 +495,10 @@ where
|
|||
ksPSpace= newPSpace,
|
||||
gsUserPages= (\<lambda>x. None),
|
||||
gsCNodes= (\<lambda>x. None),
|
||||
ksDomScheduleIdx = 0,
|
||||
ksDomSchedule = [(0, 15), (2, 42), (1, 73)],
|
||||
ksCurDomain = 0,
|
||||
ksDomainTime = 15,
|
||||
ksDomScheduleIdx = newKSDomScheduleIdx,
|
||||
ksDomSchedule = newKSDomSchedule,
|
||||
ksCurDomain = newKSCurDomain,
|
||||
ksDomainTime = newKSDomainTime,
|
||||
ksReadyQueues= const [],
|
||||
ksCurThread= error [],
|
||||
ksIdleThread= error [],
|
||||
|
|
|
@ -20,6 +20,12 @@ begin
|
|||
|
||||
#INCLUDE_HASKELL SEL4/Kernel/Init.lhs bodies_only NOT funArray newKernelState distinct rangesBy InitData doKernelOp runInit
|
||||
|
||||
consts
|
||||
newKSDomSchedule :: "(domain \<times> machine_word) list"
|
||||
newKSDomScheduleIdx :: nat
|
||||
newKSCurDomain :: domain
|
||||
newKSDomainTime :: machine_word
|
||||
|
||||
definition
|
||||
newKernelState :: "machine_word \<Rightarrow> kernel_state"
|
||||
where
|
||||
|
@ -27,10 +33,10 @@ where
|
|||
ksPSpace= newPSpace,
|
||||
gsUserPages= (\<lambda>x. None),
|
||||
gsCNodes= (\<lambda>x. None),
|
||||
ksDomScheduleIdx = 0,
|
||||
ksDomSchedule = [(0, 15), (2, 42), (1, 73)],
|
||||
ksCurDomain = 0,
|
||||
ksDomainTime = 15,
|
||||
ksDomScheduleIdx = newKSDomScheduleIdx,
|
||||
ksDomSchedule = newKSDomSchedule,
|
||||
ksCurDomain = newKSCurDomain,
|
||||
ksDomainTime = newKSDomainTime,
|
||||
ksReadyQueues= const [],
|
||||
ksCurThread= error [],
|
||||
ksIdleThread= error [],
|
||||
|
|
|
@ -1,10 +1,4 @@
|
|||
Built from hg repo at ../haskell on xinmac-204.keg.ertos.in.nicta.com.au by xgao
|
||||
Built from git repo at ../../../seL4/haskell on squires by jbeeren
|
||||
|
||||
Generated from changeset:
|
||||
|
||||
changeset: 3a9f1ef9bd4b
|
||||
branch: default
|
||||
user: Xin Gao <Xin.Gao@nicta.com.au>
|
||||
date: Fri, 23 May 2014 20:46:23 +1000
|
||||
description: fix (start,end) checking when decoding armPageDirectory invocation
|
||||
|
||||
|
|
|
@ -214,7 +214,7 @@ lemma object_slots_cnode_half:
|
|||
\<Longrightarrow> object_slots (cnode_half spec obj_id obj) slot =
|
||||
object_slots obj slot"
|
||||
apply (case_tac "has_slots obj")
|
||||
apply (clarsimp simp: cnode_half_def clean_slots_def restrict_map_def)
|
||||
apply (clarsimp simp: cnode_half_def restrict_map_def)
|
||||
apply (clarsimp simp: cnode_half_def)
|
||||
done
|
||||
|
||||
|
@ -228,7 +228,7 @@ lemma cnode_slot_half_initialised_not_original_slot:
|
|||
apply (clarsimp simp: sep_map_s_def sep_map_general_def)
|
||||
apply (rule ext)
|
||||
apply (clarsimp simp: obj_to_sep_state_def object_project_def
|
||||
object_slots_object_lift
|
||||
object_slots_object_clean
|
||||
split: option.splits)
|
||||
apply (cut_tac obj = "cnode_half spec obj_id spec_object" and
|
||||
obj' = spec_object and slot=slot and t=t in object_slot_spec2s)
|
||||
|
@ -237,7 +237,7 @@ lemma cnode_slot_half_initialised_not_original_slot:
|
|||
apply (clarsimp simp: sep_map_s_def sep_map_general_def)
|
||||
apply (rule ext)
|
||||
apply (clarsimp simp: obj_to_sep_state_def object_project_def
|
||||
object_slots_object_lift
|
||||
object_slots_object_clean
|
||||
split: option.splits)
|
||||
apply (cut_tac obj = "cnode_half spec obj_id spec_object" and
|
||||
obj' = spec_object and slot=slot and t=t in object_slot_spec2s)
|
||||
|
@ -249,27 +249,26 @@ lemma slots_empty_cnode1:
|
|||
"slot < 2 ^ sz
|
||||
\<Longrightarrow> object_slots (CNode (empty_cnode sz)) slot = Some NullCap"
|
||||
by (fastforce simp: object_slots_def empty_cnode_def empty_cap_map_def
|
||||
clean_slots_def restrict_map_def cdl_cnode.splits)
|
||||
restrict_map_def cdl_cnode.splits)
|
||||
|
||||
lemma slots_empty_cnode2:
|
||||
"\<not> slot < 2 ^ sz
|
||||
\<Longrightarrow> object_slots (CNode (empty_cnode sz)) slot = None"
|
||||
by (fastforce simp: object_slots_def empty_cnode_def empty_cap_map_def
|
||||
clean_slots_def restrict_map_def cdl_cnode.splits)
|
||||
restrict_map_def cdl_cnode.splits)
|
||||
|
||||
lemma slots_spec2s_cnode_half1:
|
||||
"\<lbrakk>slot < 2 ^ sz; original_cap_at (obj_id, slot) spec; (cdl_cnode_caps cnode slot) \<noteq> None\<rbrakk>
|
||||
\<Longrightarrow> object_slots (spec2s t (cnode_half spec obj_id (CNode cnode))) slot
|
||||
= Some NullCap"
|
||||
by (fastforce simp: object_slots_def cnode_half_def spec2s_def update_slots_def
|
||||
clean_slots_def)
|
||||
by (fastforce simp: object_slots_def cnode_half_def spec2s_def update_slots_def)
|
||||
|
||||
lemma slots_spec2s_cnode_half2:
|
||||
"\<lbrakk>\<not> slot < 2 ^ sz; original_cap_at (obj_id, slot) spec; (cdl_cnode_caps cnode slot) = None\<rbrakk>
|
||||
\<Longrightarrow> object_slots (spec2s t (cnode_half spec obj_id (CNode cnode))) slot
|
||||
= None"
|
||||
by (fastforce simp: object_slots_def cnode_half_def spec2s_def update_slots_def
|
||||
clean_slots_def restrict_map_def)
|
||||
restrict_map_def)
|
||||
|
||||
lemma object_slots_spec2s_cnode_half_object_default_state:
|
||||
"\<lbrakk>well_formed spec; original_cap_at (obj_id, slot) spec;
|
||||
|
@ -305,13 +304,13 @@ lemma cnode_slot_half_initialised_original_slot:
|
|||
apply (rule ext, rule iffI)
|
||||
apply (clarsimp simp: sep_map_s_def sep_map_general_def)
|
||||
apply (rule ext, clarsimp simp:obj_to_sep_state_def
|
||||
object_project_def object_slots_object_lift)
|
||||
object_project_def object_slots_object_clean)
|
||||
apply (subst object_slots_spec2s_cnode_half_object_default_state)
|
||||
apply simp+
|
||||
apply (clarsimp simp: object_at_def)+
|
||||
apply (clarsimp simp: sep_map_s_def sep_map_general_def)
|
||||
apply (rule ext)
|
||||
apply (clarsimp simp:obj_to_sep_state_def object_project_def object_slots_object_lift)
|
||||
apply (clarsimp simp:obj_to_sep_state_def object_project_def object_slots_object_clean)
|
||||
apply (subst object_slots_spec2s_cnode_half_object_default_state, simp+)
|
||||
apply (clarsimp split: option.splits)
|
||||
done
|
||||
|
@ -604,22 +603,15 @@ lemma mint_post:
|
|||
apply (clarsimp simp: object_at_def object_type_is_object cap_has_object_cap_type')
|
||||
apply (frule_tac obj_id=dest_id in empty_cnode_object_size_bits, clarsimp)
|
||||
apply (cut_tac slot=slot in offset_slot, assumption, simp, simp)
|
||||
apply (sep_select_asm 3)
|
||||
apply (sep_select_asm 9)
|
||||
apply (sep_drule sep_map_s_sep_map_c [where slot=slot], (clarsimp simp: sep_conj_exists)+)
|
||||
apply (cut_tac slot=slot and obj_id=dest_id and obj=obj and obj'="spec2s t spec_obj"
|
||||
in sep_map_s_object_slots_equal)
|
||||
apply (clarsimp simp: opt_cap_def slots_of_def opt_object_def)
|
||||
apply (subst object_slots_spec2s [rotated],assumption+,clarsimp+)
|
||||
apply (subgoal_tac "Some (object_type spec_cap_object) = cap_type spec_cap", clarsimp)
|
||||
apply (frule (2) well_formed_well_formed_cap, clarsimp simp: cap_has_object_def)
|
||||
apply (frule (2) well_formed_vm_cap_has_asid)
|
||||
apply (frule (1) well_formed_is_fake_vm_cap,
|
||||
(assumption|simp add: object_type_is_object)+)
|
||||
apply (clarsimp simp: cap_rights_inter_default_cap_rights)
|
||||
apply (subst update_cap_rights_and_data,(assumption|clarsimp)+)
|
||||
apply (subst (asm) offset_slot', assumption)
|
||||
apply (subst (asm) offset_slot', assumption)
|
||||
apply (subst sep_map_s_sep_map_c_eq [where cap="update_cap_object client_object_id spec_cap"])
|
||||
apply (rule object_slots_spec2s, (clarsimp simp: opt_cap_def slots_of_def opt_object_def)+)
|
||||
apply (frule (2) well_formed_well_formed_cap, clarsimp simp: cap_has_object_def)
|
||||
apply (frule (2) well_formed_vm_cap_has_asid)
|
||||
apply (frule (1) well_formed_is_fake_vm_cap,
|
||||
(assumption|simp add: object_type_is_object)+)
|
||||
apply (clarsimp simp: cap_rights_inter_default_cap_rights)
|
||||
apply (subst (asm) update_cap_rights_and_data,(assumption|clarsimp)+)
|
||||
apply (subst (asm) offset_slot', assumption)+
|
||||
apply sep_solve
|
||||
done
|
||||
|
||||
|
@ -670,26 +662,16 @@ lemma mutate_post:
|
|||
apply (clarsimp simp: object_at_def object_type_is_object cap_has_object_cap_type')
|
||||
apply (frule_tac obj_id=dest_id in empty_cnode_object_size_bits, clarsimp)
|
||||
apply (cut_tac slot=slot in offset_slot, assumption, simp, simp)
|
||||
apply (sep_select_asm 3)
|
||||
apply (sep_select_asm 9)
|
||||
apply (sep_drule sep_map_s_sep_map_c [where slot=slot])
|
||||
apply (clarsimp simp: sep_conj_exists)+
|
||||
apply (subst (asm) offset_slot', assumption)
|
||||
apply (subst (asm) offset_slot', assumption)
|
||||
apply sep_cancel+
|
||||
apply (cut_tac slot=slot and obj_id=dest_id and obj=obj and obj'="spec2s t spec_obj"
|
||||
in sep_map_s_object_slots_equal)
|
||||
apply (clarsimp simp: opt_cap_def slots_of_def opt_object_def)
|
||||
apply (subst object_slots_spec2s [rotated], assumption+, clarsimp+)
|
||||
apply (subgoal_tac "Some (object_type spec_cap_object) = cap_type spec_cap", clarsimp)
|
||||
apply (frule (2) well_formed_well_formed_cap, clarsimp simp: cap_has_object_def)
|
||||
apply (frule (2) well_formed_vm_cap_has_asid)
|
||||
apply (frule (1) well_formed_is_fake_vm_cap,
|
||||
(assumption|simp add: object_type_is_object)+)
|
||||
apply (subst update_cap_data, (assumption|clarsimp)+)
|
||||
apply (subst well_formed_orig_caps, assumption+)
|
||||
apply (simp add: slots_of_def opt_object_def split: option.splits)
|
||||
apply simp+
|
||||
apply (subst sep_map_s_sep_map_c_eq [where cap="update_cap_object client_object_id spec_cap"])
|
||||
apply (rule object_slots_spec2s, (clarsimp simp: opt_cap_def slots_of_def opt_object_def)+)
|
||||
apply (frule (2) well_formed_well_formed_cap, clarsimp simp: cap_has_object_def)
|
||||
apply (frule (2) well_formed_vm_cap_has_asid)
|
||||
apply (frule (1) well_formed_is_fake_vm_cap,
|
||||
(assumption|simp add: object_type_is_object)+)
|
||||
apply (subst update_cap_data [symmetric], simp+)
|
||||
apply (erule well_formed_orig_caps, (simp add: slots_of_def opt_object_def)+)
|
||||
apply (subst (asm) offset_slot', assumption)+
|
||||
apply sep_solve
|
||||
done
|
||||
|
||||
lemma move_post:
|
||||
|
@ -736,22 +718,10 @@ lemma move_post:
|
|||
apply (clarsimp simp: object_at_def object_type_is_object cap_has_object_cap_type')
|
||||
apply (frule_tac obj_id=dest_id in empty_cnode_object_size_bits, clarsimp)
|
||||
apply (cut_tac slot=slot in offset_slot, assumption, simp, simp)
|
||||
apply (sep_select_asm 3)
|
||||
apply (sep_select_asm 9)
|
||||
apply (sep_drule sep_map_s_sep_map_c [where slot=slot])
|
||||
apply (clarsimp simp: sep_conj_exists)+
|
||||
apply (subst (asm) offset_slot', assumption)
|
||||
apply (subst (asm) offset_slot', assumption)
|
||||
apply sep_cancel+
|
||||
apply (cut_tac slot=slot and obj_id=dest_id and obj=obj and obj'="spec2s t spec_obj"
|
||||
in sep_map_s_object_slots_equal)
|
||||
apply (clarsimp simp: opt_cap_def slots_of_def opt_object_def)
|
||||
apply (subst object_slots_spec2s [rotated], assumption+, clarsimp)
|
||||
apply (subgoal_tac "Some (object_type spec_cap_object) = cap_type spec_cap", clarsimp)
|
||||
apply (frule (2) well_formed_well_formed_cap, clarsimp simp: cap_has_object_def)
|
||||
apply (frule (2) well_formed_vm_cap_has_asid)
|
||||
apply (frule (1) well_formed_is_fake_vm_cap,
|
||||
(assumption|simp add: object_type_is_object)+)
|
||||
apply (subst sep_map_s_sep_map_c_eq [where cap="update_cap_object client_object_id spec_cap"])
|
||||
apply (rule object_slots_spec2s, (clarsimp simp: opt_cap_def slots_of_def opt_object_def)+)
|
||||
apply (subst (asm) offset_slot', assumption)+
|
||||
apply sep_solve
|
||||
done
|
||||
|
||||
lemma move_post_irq_handler:
|
||||
|
@ -792,17 +762,9 @@ lemma move_post_irq_handler:
|
|||
apply (clarsimp simp: object_at_def object_type_is_object cap_has_object_cap_type')
|
||||
apply (frule_tac obj_id=dest_id in empty_cnode_object_size_bits, clarsimp)
|
||||
apply (cut_tac slot=slot in offset_slot, assumption, simp, simp)
|
||||
apply (sep_select_asm 3)
|
||||
apply (sep_select_asm 9)
|
||||
apply (sep_drule sep_map_s_sep_map_c [where slot=slot], clarsimp+)
|
||||
apply (clarsimp simp: sep_conj_exists)
|
||||
apply (subst (asm) offset_slot', assumption)
|
||||
apply (subst (asm) offset_slot', assumption)
|
||||
apply (cut_tac slot=slot and obj_id=dest_id and obj=obj and obj'="spec2s t spec_obj"
|
||||
in sep_map_s_object_slots_equal)
|
||||
apply (clarsimp simp: opt_cap_def slots_of_def opt_object_def)
|
||||
apply simp
|
||||
apply simp
|
||||
apply (subst sep_map_s_sep_map_c_eq [where cap=spec_cap],
|
||||
(clarsimp simp: opt_cap_def slots_of_def opt_object_def)+)
|
||||
apply (subst (asm) offset_slot', assumption)+
|
||||
apply sep_solve
|
||||
done
|
||||
|
||||
|
|
Some files were not shown because too many files have changed in this diff Show More
Loading…
Reference in New Issue