SELFOUR-421: merge with master, fix wholesystem proofs

This commit is contained in:
Xin,Gao 2016-09-22 19:23:19 +10:00 committed by Rafal Kolanski
commit 8f3a4dee31
52 changed files with 1028 additions and 1550 deletions

View File

@ -419,7 +419,7 @@ where
cspace \<mapsto> Types_D.CNodeCap (the_cnode_of n) (cnode_guard spec n) (cnode_guard_size spec n)
(cnode_size_bits spec n),
vspace \<mapsto> Types_D.PageDirectoryCap (the_pd_of n) Real None,
ipc_buffer_slot \<mapsto> Types_D.FrameCap (the_ipc_buffer n 0) RW 12 Real None],
ipc_buffer_slot \<mapsto> Types_D.FrameCap False (the_ipc_buffer n 0) RW 12 Real None],
cdl_tcb_fault_endpoint = 0,
cdl_tcb_intent = undefined,
cdl_tcb_has_fault = False,
@ -430,7 +430,7 @@ where
cspace \<mapsto> Types_D.CNodeCap (the_cnode_of n) (cnode_guard spec n) (cnode_guard_size spec n)
(cnode_size_bits spec n),
vspace \<mapsto> Types_D.PageDirectoryCap (the_pd_of n) Real None,
ipc_buffer_slot \<mapsto> Types_D.FrameCap (the_ipc_buffer n (i + 1)) RW 12 Real None],
ipc_buffer_slot \<mapsto> Types_D.FrameCap False (the_ipc_buffer n (i + 1)) RW 12 Real None],
cdl_tcb_fault_endpoint = 0,
cdl_tcb_intent = undefined,
cdl_tcb_has_fault = False,

View File

@ -252,15 +252,19 @@ ML {*
structure CrunchBCorresInstance : CrunchInstance =
struct
type extra = term;
val eq_extra = ae_conv;
val name = "bcorres";
val has_preconds = false;
fun mk_term _ body extra =
(Syntax.parse_term @{context} "bcorres_underlying") $ extra $ body $ body;
fun get_precond (Const (@{const_name "bcorres_underlying"}, _) $ _ $ _ $ _ ) = Var (("Dummy", 0), dummyT)
fun get_precond (Const (@{const_name "bcorres_underlying"}, _) $ _ $ _ $ _ ) = Term.dummy
| get_precond _ = error "get_precond: not an bcorres term";
fun put_precond _ ((v as Const (@{const_name "bcorres_underlying"}, _)) $ extra $ body $ body')
= v $ extra $ body $ body'
| put_precond _ _ = error "put_precond: not an bcorres term";
fun dest_term (Const (@{const_name "bcorres_underlying"}, _) $ extra $ body $ _)
= SOME (Term.dummy, body, extra)
| dest_term _ = NONE
val pre_thms = [];
val wpc_tactic = WeakestPreCases.wp_cases_tac @{thms wpc_processors};
fun parse_extra ctxt extra

View File

@ -27,6 +27,7 @@ fun add_crunch_instance name instance lthy =
structure CrunchValidInstance : CrunchInstance =
struct
type extra = term;
val eq_extra = ae_conv;
val name = "valid";
val has_preconds = true;
fun mk_term pre body post =
@ -36,6 +37,9 @@ struct
fun put_precond pre ((v as Const (@{const_name "valid"}, _)) $ _ $ body $ post)
= v $ pre $ body $ post
| put_precond _ _ = error "put_precond: not a hoare triple";
fun dest_term ((Const (@{const_name "valid"}, _)) $ pre $ body $ post)
= SOME (pre, body, betapply (post, Bound 0))
| dest_term _ = NONE
val pre_thms = @{thms "hoare_pre"};
val wpc_tactic = wp_cases_tactic_weak;
fun parse_extra ctxt extra
@ -51,6 +55,7 @@ structure CrunchValid : CRUNCH = Crunch(CrunchValidInstance);
structure CrunchNoFailInstance : CrunchInstance =
struct
type extra = unit;
val eq_extra = op =;
val name = "no_fail";
val has_preconds = true;
fun mk_term pre body _ =
@ -60,6 +65,9 @@ struct
fun put_precond pre ((v as Const (@{const_name "no_fail"}, _)) $ _ $ body)
= v $ pre $ body
| put_precond _ _ = error "put_precond: not a no_fail term";
fun dest_term ((Const (@{const_name "no_fail"}, _)) $ pre $ body)
= SOME (pre, body, ())
| dest_term _ = NONE
val pre_thms = @{thms "no_fail_pre"};
val wpc_tactic = wp_cases_tactic_weak;
fun parse_extra ctxt extra
@ -75,12 +83,16 @@ structure CrunchNoFail : CRUNCH = Crunch(CrunchNoFailInstance);
structure CrunchEmptyFailInstance : CrunchInstance =
struct
type extra = unit;
val eq_extra = op =;
val name = "empty_fail";
val has_preconds = false;
fun mk_term _ body _ =
(Syntax.parse_term @{context} "empty_fail") $ body;
fun get_precond _ = error "crunch empty_fail should not be calling get_precond";
fun put_precond _ _ = error "crunch empty_fail should not be calling put_precond";
fun dest_term (Const (@{const_name empty_fail}, _) $ b)
= SOME (Term.dummy, b, ())
| dest_term _ = NONE
val pre_thms = [];
val wpc_tactic = wp_cases_tactic_weak;
fun parse_extra ctxt extra
@ -96,6 +108,7 @@ structure CrunchEmptyFail : CRUNCH = Crunch(CrunchEmptyFailInstance);
structure CrunchValidEInstance : CrunchInstance =
struct
type extra = term * term;
fun eq_extra ((a, b), (c, d)) = (ae_conv (a, c) andalso ae_conv (b, d));
val name = "valid_E";
val has_preconds = true;
fun mk_term pre body extra =
@ -106,6 +119,9 @@ struct
fun put_precond pre ((v as Const (@{const_name "validE"}, _)) $ _ $ body $ post $ post')
= v $ pre $ body $ post $ post'
| put_precond _ _ = error "put_precond: not a validE term";
fun dest_term (Const (@{const_name "validE"}, _) $ pre $ body $ p1 $ p2)
= SOME (pre, body, (betapply (p1, Bound 0), betapply (p2, Bound 0)))
| dest_term _ = NONE
val pre_thms = @{thms "hoare_pre"};
val wpc_tactic = wp_cases_tactic_weak;
fun parse_extra ctxt extra
@ -162,7 +178,7 @@ val crunchP =
(((Scan.optional (P.$$$ "(" |-- P.name --| P.$$$ ")") "" -- P.name
-- Parse.opt_attribs --| P.$$$ ":") -- P.list1 P.name -- Scan.optional P.term ""
-- Scan.optional
(P.$$$ "(" |-- read_sections [wp_sect,wp_del_sect,ignore_sect,simp_sect,simp_del_sect,lift_sect,ignore_del_sect,unfold_sect] --| P.$$$ ")")
(P.$$$ "(" |-- read_sections [wp_sect,wp_del_sect,ignore_sect,simp_sect,simp_del_sect,rule_sect,rule_del_sect,ignore_del_sect] --| P.$$$ ")")
[]
)
>> (fn (((((crunch_instance, prp_name), att_srcs), consts), extra), wpigs) =>

View File

@ -14,6 +14,7 @@ keywords "crunch" "crunch_ignore" :: thy_decl
begin
named_theorems "crunch_def"
named_theorems "crunch_rules"
ML_file "crunch-cmd.ML"
ML_file "Crunch.ML"

View File

@ -95,12 +95,11 @@ lemma crunch_foo4_alt:
crunch gt3: crunch_foo4 "\<lambda>x. x > y"
(ignore: modify bind)
*)
crunch (no_fail) no_fail2: crunch_foo4
(unfold: "(crunch_foo4, crunch_foo4_alt)" ignore: modify bind)
(rule: crunch_foo4_alt ignore: modify bind)
crunch gt3: crunch_foo4 "\<lambda>x. x > y"
(unfold: "(crunch_foo4, crunch_foo4_alt)" ignore: modify bind)
(rule: crunch_foo4_alt ignore: modify bind)
crunch gt4: crunch_foo5 "\<lambda>x. x > y"
(ignore: modify bind)

View File

@ -8,6 +8,35 @@
* @TAG(NICTA_BSD)
*)
(*
What is crunch? Crunch is for proving easily proved properties over various
functions within a (shallow monadic) specification. Suppose we have a toplevel
specification function f and some simple notion of correctness, P \<turnstile> f, which
we want to prove. However f is defined in terms of subfunctions g etc.
To prove P \<turnstile> f, we will first show many lemmas P \<turnstile> g, which crunch lets us
do easily.
As a first step, crunch must discover "cruch rules" which structure the proof.
For instance, if f is a constant with a definition, crunch may discover the
definition "f == f_body" and derive the crunch rule "P \<turnstile> f_body \<Longrightarrow> P \<turnstile> f".
A crunch rule will be used if all its premises are either proofs of the
same notion of correctness (P \<turnstile> ?f) or can be solved trivially by wp/simp.
The user can supply crunch rules with the rule: section or crunch_rule
attribute, which will be tried both as rewrites and
as introduction rules. Crunch also has a number of builtin strategies
for finding definitional rules in the context.
Once a crunch rule for P \<turnstile> f is found, crunch will recursively apply to
all monadic constants in all the premises. (The exploration terminates
where crunch rules can be found without premises.) Crunch will obtain
proofs e.g. P \<turnstile> g for all f's dependencies, then finally try to prove
P \<turnstile> f using the crunch rule, wp (with the given dependencies) and simp.
Additional wp and simp rules can also be given with the wp: and simp:
sections.
*)
(* Tracing debug. *)
val should_debug = Unsynchronized.ref false
fun debug_trace s =
@ -27,22 +56,34 @@ fun handle_int exn func = if Exn.is_interrupt exn then reraise exn else func
val wp_sect = "wp";
val wp_del_sect = "wp_del";
val ignore_sect = "ignore";
val ignore_del_sect = "ignore_del";
val simp_sect = "simp";
val simp_del_sect = "simp_del";
val lift_sect = "lift";
val ignore_del_sect = "ignore_del";
val unfold_sect = "unfold";
val rule_sect = "rule";
val rule_del_sect = "rule_del";
fun read_const ctxt = Proof_Context.read_const {proper = true, strict = false} ctxt;
fun gen_term_eq (f $ x, f' $ x') = gen_term_eq (f, f') andalso gen_term_eq (x, x')
| gen_term_eq (Abs (_, _, t), Abs (_, _, t')) = gen_term_eq (t, t')
| gen_term_eq (Const (nm, _), Const (nm', _)) = (nm = nm')
| gen_term_eq (Free (nm, _), Free (nm', _)) = (nm = nm')
| gen_term_eq (Var (nm, _), Var (nm', _)) = (nm = nm')
| gen_term_eq (Bound i, Bound j) = (i = j)
| gen_term_eq _ = false
fun ae_conv (t, t') = gen_term_eq
(Envir.beta_eta_contract t, Envir.beta_eta_contract t')
signature CrunchInstance =
sig
type extra;
val name : string;
val has_preconds : bool;
val mk_term : term -> term -> extra -> term;
val dest_term : term -> (term * term * extra) option;
val get_precond : term -> term;
val put_precond : term -> term -> term;
val eq_extra : extra * extra -> bool;
val pre_thms : thm list;
val wpc_tactic : Proof.context -> tactic;
val parse_extra : Proof.context -> string -> term * extra;
@ -53,8 +94,9 @@ signature CRUNCH =
sig
type extra;
(* Crunch configuration: theory, naming scheme, lifting rules, wp rules *)
type crunch_cfg = {ctxt: local_theory, prp_name: string, nmspce: string option, lifts: thm list,
wps: (string * thm) list, igs: string list, simps: thm list, ig_dels: string list, unfolds: (string * thm) list};
type crunch_cfg = {ctxt: local_theory, prp_name: string, nmspce: string option,
wps: (string * thm) list, igs: string list, simps: thm list, ig_dels: string list,
rules: thm list};
(* Crunch takes a configuration, a precondition, any extra information needed, a debug stack,
a constant name, and a list of previously proven theorems, and returns a theorem for
@ -68,6 +110,8 @@ sig
-> string list -> local_theory -> local_theory;
val crunch_ignore_add_del : string list -> string list -> theory -> theory
val mism_term_trace : (term * extra) list Unsynchronized.ref
end
functor Crunch (Instance : CrunchInstance) =
@ -75,8 +119,9 @@ struct
type extra = Instance.extra;
type crunch_cfg = {ctxt: local_theory, prp_name: string, nmspce: string option, lifts: thm list,
wps: (string * thm) list, igs: string list, simps: thm list, ig_dels: string list, unfolds: (string * thm) list};
type crunch_cfg = {ctxt: local_theory, prp_name: string, nmspce: string option,
wps: (string * thm) list, igs: string list, simps: thm list, ig_dels: string list,
rules: thm list};
structure CrunchIgnore = Generic_Data
(struct
@ -95,7 +140,7 @@ fun crunch_ignore_del thms thy =
fun crunch_ignore_add_del adds dels thy =
thy |> crunch_ignore_add adds |> crunch_ignore_del dels
val def_sfx = "_def";
val def_sfx = "_def";
val induct_sfx = ".induct";
val simps_sfx = ".simps";
val param_name = "param_a";
@ -131,7 +176,7 @@ fun get_monad ctxt f xs = if is_Const f then
fun num_args (Type ("fun", [v,p])) n =
if is_product v p then SOME n else num_args p (n + 1)
| num_args _ _ = NONE
in case num_args T 0 of NONE => []
| SOME n => [list_comb (f, Library.take n (xs @ map Bound (1 upto n)))]
end
@ -148,8 +193,8 @@ val get_thm = Proof_Context.get_thm
val get_thms = Proof_Context.get_thms
fun maybe_thms thy name = get_thms thy name handle exn => handle_int exn []
fun thy_maybe_thms thy name = Global_Theory.get_thms thy name handle exn => handle_int exn []
fun maybe_thms thy name = get_thms thy name handle ERROR _ => []
fun thy_maybe_thms thy name = Global_Theory.get_thms thy name handle ERROR _ => []
fun add_thm thm atts name ctxt =
Local_Theory.notes [((Binding.name name, atts), [([thm], atts)])] ctxt |> #2
@ -165,7 +210,7 @@ fun get_thm_name (cfg : crunch_cfg) const_name
fun get_stored cfg n = get_thm (#ctxt cfg) (get_thm_name cfg n)
fun get_stored_bluh cfg n =
fun get_stored_bluh cfg n =
let val r = (maybe_thms (#ctxt cfg) (get_thm_name cfg n)) @ (maybe_thms (#ctxt cfg) (get_thm_name_bluh cfg n));
in (case r of [] => error ("") | _ => (hd r))
end
@ -190,7 +235,7 @@ fun deep_search_thms ctxt defn const nmspce =
val thy = Proof_Context.theory_of ctxt
val thys = thy :: Theory.ancestors_of thy;
val filt = filter (const_is_lhs const nmspce ctxt);
fun search [] = error("not found: const: " ^ PolyML.makestring const ^ " defn: " ^ PolyML.makestring defn)
| search (t::ts) = (case (filt (thy_maybe_thms t defn)) of
[] => search ts
@ -207,19 +252,17 @@ val unfold_get_params = @{thms Let_def return_bind returnOk_bindE
fun def_from_ctxt ctxt const =
let
val crunch_defs = Named_Theorems.get ctxt "Crunch.crunch_def"
val crunch_defs = Named_Theorems.get ctxt @{named_theorems crunch_def}
val abs_def = Local_Defs.meta_rewrite_rule ctxt #> Drule.abs_def;
fun do_filter thm =
let
val (Const (nm, _), _) = Logic.dest_equals (Thm.prop_of (abs_def thm));
in nm = const end
in
in
case crunch_defs |> filter do_filter of
[x] => [x]
| [] => []
| _ => raise Fail ("Multiple definitions declared for:" ^ const)
| _ => raise Fail ("Multiple definitions declared for:" ^ const)
end
fun unfold ctxt const triple nmspce =
@ -242,8 +285,8 @@ fun unfold ctxt const triple nmspce =
then error ("Unfold rule generated for " ^ const ^ " does not apply")
else (ms, unfold_rule) end
fun mk_apps t n m =
if n = 0
fun mk_apps t n m =
if n = 0
then t
else mk_apps t (n-1) (m+1) $ Bound m
@ -263,7 +306,7 @@ fun resolve_abbreviated ctxt abbrev =
val (_::_::_::nil) = Long_Name.explode originn;
in origin end handle exn => handle_int exn abbrev
fun map_consts f =
fun map_consts f =
let
fun map_aux (Const a) = f (Const a)
| map_aux (t $ u) = map_aux t $ map_aux u
@ -390,15 +433,91 @@ fun funkyfold _ [] _ = ([], [])
exception WrongType
fun make_goal const_term const pre extra ctxt =
let val nns = const_term |> fastype_of |> num_args |>
Name.invent Name.context param_name;
let val nns = const_term |> fastype_of |> num_args |>
Name.invent Name.context param_name;
val parse = Syntax.parse_term ctxt;
val check = Syntax.check_term ctxt;
val body = parse (String.concat (separate " " (const :: nns)));
in check (Instance.mk_term pre body extra) end;
fun seq_try_apply f x = Seq.map_filter (try f) (Seq.single x)
fun crunch_known_rule cfg const const_long_name thms goal_prop =
let
val thms_proof = Seq.filter (is_proof_of cfg const) (Seq.of_list thms)
val stored = seq_try_apply (get_stored cfg) const
val ctxt = #ctxt cfg
val cgoal_in = Goal.init (Thm.cterm_of ctxt goal_prop)
val empty_ref = Unsynchronized.ref []
val wps = Seq.filter (fn (s, t) => s = const_long_name andalso
(is_some o SINGLE (WeakestPre.apply_rules_tac_n false ctxt [t] empty_ref 1)) cgoal_in)
(Seq.of_list (#wps cfg))
val seq = Seq.append (Seq.map snd thms_proof) (Seq.append stored (Seq.map snd wps))
in Seq.pull seq |> Option.map fst end
val mism_term_trace = Unsynchronized.ref []
fun crunch_rule cfg const goal extra thms =
let
(* first option: produce a terminal rule via wp *)
val ctxt = #ctxt cfg
val empty_ref = Unsynchronized.ref []
fun wp rules = WeakestPre.apply_rules_tac_n false ctxt
(map snd (thms @ #wps cfg) @ rules) empty_ref
val vgoal_prop = goal |> var_precond |> HOLogic.mk_Trueprop
val goal_prop = goal |> HOLogic.mk_Trueprop
val ctxt' = Variable.auto_fixes goal_prop
(Variable.auto_fixes vgoal_prop ctxt)
val wp_seq = seq_try_apply (Goal.prove ctxt' [] [] goal_prop) (fn _ =>
TRY (wp [] 1))
|> Seq.map (singleton (Proof_Context.export ctxt' ctxt))
|> Seq.map (pair NONE)
(* second option: apply a supplied rule *)
val cgoal = vgoal_prop |> Thm.cterm_of ctxt
val base_rule = Thm.trivial cgoal
fun app_rew r t = Seq.single (Simplifier.rewrite_goals_rule ctxt [r] t)
|> Seq.filter (fn t' => not (Thm.eq_thm_prop (t, t')))
val supplied_seq = Seq.of_list (#rules cfg)
|> Seq.maps (fn r => Seq.append
(resolve_tac ctxt [r] 1 base_rule) (app_rew r base_rule))
|> Seq.map (pair NONE)
(* third option: builtins *)
val unfolds' = (map (pair "") (def_from_ctxt ctxt const))
val unf_seq = Seq.map (unfold_data ctxt' const cgoal (#nmspce cfg))
(Seq.single unfolds')
|> Seq.map (apfst SOME)
val seq = foldr1 (uncurry Seq.append) [wp_seq, supplied_seq, unf_seq]
fun fail_tac t _ _ = (writeln "discarding crunch rule, unsolved premise:";
Syntax.pretty_term ctxt t |> Pretty.writeln;
mism_term_trace := (t, extra) :: (! mism_term_trace);
Seq.empty)
val goal_extra = goal |> Instance.dest_term |> the |> #3
val finalise = ALLGOALS (SUBGOAL (fn (t, i)
=> if try (Logic.strip_assums_concl #> HOLogic.dest_Trueprop
#> Instance.dest_term #> the #> #3 #> curry Instance.eq_extra goal_extra)
t = SOME true
then all_tac
else DETERM (((wp [] ORELSE' simp_tac ctxt) THEN_ALL_NEW fail_tac t) i)))
val seq = Seq.maps (fn (ms, t) => Seq.map (pair ms) (finalise t)) seq
val (ms, thm) = case Seq.pull seq of SOME ((ms, thm), _) => (ms, thm)
| NONE => error ("could not find crunch rule for " ^ const)
val _ = Pretty.writeln (Pretty.block
[Pretty.str "crunch rule: ", Thm.pretty_thm ctxt thm])
val ms = case ms of SOME ms => ms
| NONE => Thm.prems_of thm |> maps (monads_of ctxt)
in (ms, thm) end
fun crunch cfg pre extra stack const' thms =
let
let
val ctxt = #ctxt cfg |> Context_Position.set_visible false;
val const = real_const_from_name const' (#nmspce cfg) ctxt;
val empty_ref = Unsynchronized.ref [] : thm list Unsynchronized.ref (* FIXME: avoid refs *)
@ -410,95 +529,79 @@ fun crunch cfg pre extra stack const' thms =
val goal = make_goal const_term const pre extra ctxt
handle exn => handle_int exn (raise WrongType);
val goal_prop = HOLogic.mk_Trueprop goal;
in
let val v = find_first (is_proof_of cfg const) thms
in (SOME (snd (the v)), []) end
handle exn => handle_int exn
(SOME (get_stored cfg const), [])
handle exn => handle_int exn
let val const_long_name = real_const_term |> dest_Const |> fst;
val cgoal_prop = Thm.cterm_of ctxt goal_prop;
val v = find_first (fn (s, t) => s = const_long_name andalso
(is_some o SINGLE (WeakestPre.apply_rules_tac_n false ctxt [t] empty_ref 1)) (Goal.init cgoal_prop)) (#wps cfg)
in (SOME (snd (the v)), []) end
handle exn => handle_int exn
let
fun wp rules = WeakestPre.apply_rules_tac_n false ctxt
(map snd (thms @ #wps cfg) @ rules) empty_ref
val const_long_name = real_const_term |> dest_Const |> fst;
in (* first check: has this constant already been done or supplied? *)
case crunch_known_rule cfg const const_long_name thms goal_prop
of SOME thm => (SOME thm, [])
| NONE => let (* not already known, find a crunch rule. *)
val (ms, rule) = crunch_rule cfg const goal extra thms
(* and now crunch *)
val ctxt' = Variable.auto_fixes goal ctxt;
val ms = ms
|> map (fn t => (real_const_from_name (fst (dest_Const (head_of t))) (#nmspce cfg) ctxt', t))
|> subtract (fn (a, b) => a = (fst b))
(subtract (op =) (#ig_dels cfg) (#igs cfg @ CrunchIgnore.get (Context.Proof ctxt')))
|> filter_out (fn (s, _) => s = const');
val stack' = const :: stack;
val _ = if (length stack' > 20) then
(writeln "Crunch call stack:";
map writeln (const::stack);
error("probably infinite loop")) else ();
val (goals, thms') = funkyfold (crunch cfg pre extra stack') (map fst ms) thms;
val goals' = map_filter I goals
val ctxt'' = ctxt' addsimps ((#simps cfg) @ goals')
|> Splitter.del_split split_if
val ctxt' = Variable.auto_fixes goal_prop ctxt
val (thm, thms')
= ( Goal.prove ctxt' [] [] goal_prop (fn prems =>
TRY (wp [] 1) THEN TRY (wp (#lifts cfg) 1))
|> singleton (Proof_Context.export ctxt' ctxt)
, [] )
handle exn => handle_int exn
let
val ctxt' = Variable.auto_fixes goal ctxt;
val cgoal = goal |> var_precond |> HOLogic.mk_Trueprop |> Thm.cterm_of ctxt';
val unfolds' = filter (fn (a, _) => a = const) (#unfolds cfg) @ (map (pair "") (def_from_ctxt ctxt const))
val (ms, unfold_rule) = unfold_data ctxt' const cgoal (#nmspce cfg) unfolds'
|>> map (fn t => (real_const_from_name (fst (dest_Const (head_of t))) (#nmspce cfg) ctxt', t))
|>> subtract (fn (a, b) => a = (fst b))
(subtract (op =) (#ig_dels cfg) (#igs cfg @ CrunchIgnore.get (Context.Proof ctxt')));
val stack' = const :: stack;
val _ = if (length stack' > 20) then
(writeln "Crunch call stack:";
map writeln (const::stack);
error("probably infinite loop")) else ();
val (goals, thms') = funkyfold (crunch cfg pre extra stack') (map fst ms) thms;
val goals' = map_filter I goals
val ctxt'' = ctxt' addsimps ((#simps cfg) @ goals')
|> Splitter.del_split split_if
fun collect_preconds pre =
let val preconds = map_filter (fn (x, SOME y) => SOME (x, y) | (_, NONE) => NONE) (map snd ms ~~ goals)
fun collect_preconds pre =
let val preconds = map_filter (fn (x, SOME y) => SOME (x, y) | (_, NONE) => NONE) (map snd ms ~~ goals)
|> map_filter (get_inst_precond ctxt'' pre extra);
val precond = combine_preconds ctxt'' (Instance.get_precond goal) preconds;
in Instance.put_precond precond goal |> HOLogic.mk_Trueprop end;
val goal_prop2 = if Instance.has_preconds then collect_preconds pre else goal_prop;
val precond = combine_preconds ctxt'' (Instance.get_precond goal) preconds;
in Instance.put_precond precond goal |> HOLogic.mk_Trueprop end;
val goal_prop2 = if Instance.has_preconds then collect_preconds pre else goal_prop;
val ctxt''' = ctxt'' |> Variable.auto_fixes goal_prop2
val _ = writeln ("attempting: " ^ Syntax.string_of_term ctxt''' goal_prop2);
in (Goal.prove_future ctxt''' [] [] goal_prop2
( (*DupSkip.goal_prove_wrapper *) (fn _ =>
resolve_tac ctxt''' [unfold_rule] 1
THEN maybe_cheat_tac ctxt'''
THEN ALLGOALS (fn n =>
simp_tac ctxt''' n
THEN
TRY (resolve_tac ctxt''' Instance.pre_thms n)
THEN
REPEAT_DETERM (
wp goals' n
ORELSE
CHANGED (clarsimp_tac ctxt''' n)
ORELSE
assume_tac ctxt''' n
ORELSE
Instance.wpc_tactic ctxt'''
ORELSE
safe_tac ctxt'''
ORELSE
CHANGED (simp_tac ctxt''' n)
)))) |> singleton (Proof_Context.export ctxt''' ctxt)
handle e =>
(writeln "Crunch call stack:";
map writeln (const::stack);
raise e)
, thms') end
val ctxt''' = ctxt'' |> Variable.auto_fixes goal_prop2
val _ = writeln ("attempting: " ^ Syntax.string_of_term ctxt''' goal_prop2);
fun wp rules = WeakestPre.apply_rules_tac_n false ctxt
(map snd (thms @ #wps cfg) @ rules) empty_ref
val thm = Goal.prove_future ctxt''' [] [] goal_prop2
( (*DupSkip.goal_prove_wrapper *) (fn _ =>
resolve_tac ctxt''' [rule] 1
THEN maybe_cheat_tac ctxt'''
THEN ALLGOALS (fn n =>
simp_tac ctxt''' n
THEN
TRY (resolve_tac ctxt''' Instance.pre_thms n)
THEN
REPEAT_DETERM (
wp goals' n
ORELSE
CHANGED (clarsimp_tac ctxt''' n)
ORELSE
assume_tac ctxt''' n
ORELSE
Instance.wpc_tactic ctxt'''
ORELSE
safe_tac ctxt'''
ORELSE
CHANGED (simp_tac ctxt''' n)
)))) |> singleton (Proof_Context.export ctxt''' ctxt)
handle e =>
(writeln "Crunch call stack:";
map writeln (const::stack);
raise e)
in (SOME thm, (get_thm_name cfg const, thm) :: thms') end
end
handle WrongType =>
let val _ = writeln ("The constant " ^ const ^ " has the wrong type and is being ignored")
in (NONE, []) end
end
end
(*Todo: Remember mapping from locales to theories*)
fun get_locale_origins full_const_names ctxt =
let
fun get_locale_origin abbrev =
fun get_locale_origin abbrev =
let
(*Check if the given const is an abbreviation*)
val (origin,_) = dest_Const (head_of (snd ((Consts.the_abbreviation o Proof_Context.consts_of) ctxt abbrev)));
@ -506,20 +609,12 @@ fun get_locale_origins full_const_names ctxt =
val [_,_,_] = Long_Name.explode origin;
(*Remember the theory for the abbreviation*)
val [qual,nm] = Long_Name.explode abbrev
val [qual,nm] = Long_Name.explode abbrev
in SOME qual end handle exn => handle_int exn NONE
in fold (curry (fn (abbrev,qual) => case (get_locale_origin abbrev) of
SOME q => SOME q
| NONE => NONE)) full_const_names NONE
end
fun parse_unfold_pair unfold =
let
val parse_string = Scan.repeat (Scan.unless ($$ "," || $$ ")") (Scan.one Symbol.not_eof))
val ((first, second), rem) = ($$ "(" |-- parse_string --| Scan.this_string ", " -- parse_string --| $$ ")") (Symbol.explode unfold)
handle exn => handle_int exn error "Could not parse unfold data, expected: (const, thm)"
in if rem = [] then (implode first, implode second)
else error "Could not parse unfold data, expected: (const, thm)" end
fun crunch_x atts extra prp_name wpigs consts ctxt =
let
@ -535,32 +630,28 @@ fun crunch_x atts extra prp_name wpigs consts ctxt =
val simp_dels = wpigs |> filter (fn (s,_) => s = simp_del_sect) |> map #2
|> maps (get_thms ctxt)
val igs = wpigs |> filter (fn (s,_) => s = ignore_sect)
val igs = wpigs |> filter (fn (s,_) => s = ignore_sect)
|> map (const_name o #2)
val lifts = wpigs |> filter (fn (s,_) => s = lift_sect) |> map #2
val rules = wpigs |> filter (fn (s,_) => s = rule_sect) |> map #2
|> maps (get_thms ctxt)
val rules = rules @ Named_Theorems.get ctxt @{named_theorems crunch_rules}
val ig_dels = wpigs |> filter (fn (s,_) => s = ignore_del_sect)
|> map (const_name o #2)
fun read_unfold (first, second) = (const_name first, get_thm ctxt second)
val unfolds = wpigs |> filter (fn (s,_) => s = unfold_sect) |> map #2
|> map (read_unfold o parse_unfold_pair)
fun mk_wp thm =
fun mk_wp thm =
let val ms = Thm.prop_of thm |> monads_of ctxt;
val m = if length ms = 1
then
val m = if length ms = 1
then
hd ms |> head_of |> dest_Const |> fst
else
else
dummy_monad_name;
in (m, thm) end;
val wps = maps (get_thms ctxt) wps' |> map mk_wp;
val full_const_names = map const_name consts;
val nmspce = get_locale_origins full_const_names ctxt;
val (pre', extra') = Instance.parse_extra ctxt extra
@ -575,8 +666,8 @@ fun crunch_x atts extra prp_name wpigs consts ctxt =
val ctxt'' = ctxt' delsimps simp_dels;
val (_, thms) = funkyfold (crunch {ctxt = ctxt'', prp_name = prp_name, nmspce = nmspce, lifts = lifts,
wps = wps, igs = igs, simps = simps, ig_dels = ig_dels, unfolds = unfolds} pre' extra' [])
val (_, thms) = funkyfold (crunch {ctxt = ctxt'', prp_name = prp_name, nmspce = nmspce,
wps = wps, igs = igs, simps = simps, ig_dels = ig_dels, rules = rules} pre' extra' [])
full_const_names [];
val atts' = map (Attrib.check_src ctxt) atts;

View File

@ -355,17 +355,24 @@ def print_test_line(test_name, color, status, real_time=None, cpu_time=None, mem
#
def rglob(base_dir, pattern):
matches = []
extras = []
for root, dirnames, filenames in os.walk(base_dir):
for filename in fnmatch.filter(filenames, pattern):
matches.append(os.path.join(root, filename))
return matches
for filename in fnmatch.filter(filenames, 'extra_tests'):
f = os.path.join(root, filename)
extras.extend([os.path.join(root, l.strip())
for l in open(f) if l.strip()])
matches.extend([f for e in extras for f in rglob(e, pattern)])
return sorted(set(matches))
#
# Run tests.
#
def main():
# Parse arguments
parser = argparse.ArgumentParser(description="Parallel Regression Framework")
parser = argparse.ArgumentParser(description="Parallel Regression Framework",
epilog="RUN_TESTS_DEFAULT can be used to overwrite the default set of tests")
parser.add_argument("-s", "--strict", action="store_true",
help="be strict when parsing test XML files")
parser.add_argument("-d", "--directory", action="store",
@ -420,10 +427,10 @@ def main():
# Calculate which tests should be run.
tests_to_run = []
if len(args.tests) == 0:
if len(args.tests) == 0 and not os.environ.get('RUN_TESTS_DEFAULT'):
tests_to_run = tests
else:
desired_names = set(args.tests)
desired_names = set(args.tests) or set(os.environ.get('RUN_TESTS_DEFAULT').split())
bad_names = desired_names - set([t.name for t in tests])
if len(bad_names) > 0:
parser.error("Unknown test names: %s" % (", ".join(sorted(bad_names))))

View File

@ -93,6 +93,7 @@ mv kernel_all.txt kernel_all_2.txt
[ -z ${AST+x} ] || mv ckernel_ast.txt ckernel_ast_2.txt
set +e
ERRORS=false
# Check for differences in
echo -en "${RED}"
if ! diff ckernel_names_1.txt ckernel_names_2.txt &> /dev/null
@ -103,8 +104,6 @@ then
echo -e "#################################\n"
echo "please address this issue appropriately"
echo "and run make_muge.sh afterwards."
echo -e "\n${NC}${YEL}kernel_all diff:${NC}"
diff -uw kernel_all_1.txt kernel_all_2.txt
echo -e "\n${YEL}Symbols diff:${NC}"
diff -uw ckernel_names_1.txt ckernel_names_2.txt
ERRORS=true
@ -125,7 +124,20 @@ fi
echo -en "${NC}"
if [ -z ${ERRORS+x} ]
if ! ${ERRORS}
then echo -e "${GRE}Everything seems fine${NC}"
else exit 1
else
echo -en "${RED}"
if ! diff ckernel_all_1.txt ckernel_all_2.txt &> /dev/null
then
echo -e "${RED}"
echo "#################################"
echo "# Something has changed #"
echo -e "#################################\n"
echo "please address this issue appropriately"
echo "and run make_muge.sh afterwards."
echo -e "\n${NC}${YEL}kernel_all diff:${NC}"
diff -uw kernel_all_1.txt kernel_all_2.txt
fi
exit 1
fi

View File

@ -33,7 +33,7 @@ crunch valid_queues'[wp]: handleFaultReply valid_queues'
crunch sch_act_wf: handleFaultReply "\<lambda>s. sch_act_wf (ksSchedulerAction s) s"
crunch valid_ipc_buffer_ptr' [wp]: copyMRs "valid_ipc_buffer_ptr' p"
(lift: hoare_valid_ipc_buffer_ptr_typ_at' wp: crunch_wps)
(rule: hoare_valid_ipc_buffer_ptr_typ_at' wp: crunch_wps)
lemma threadSet_obj_at'_nontcb:
"koType TYPE('a::pspace_storable) \<noteq> koType TYPE(Structures_H.tcb) \<Longrightarrow>
@ -2328,7 +2328,7 @@ lemma maskedAsFull_again:
apply (simp_all add:maskedAsFull_def isCap_simps split: split_if)+
done
lemma ccap_relation_lift:
lemma ccap_relation_lift:
"ccap_relation cap cap'
\<Longrightarrow> (cap_to_H (the (cap_lift cap'))) = cap"
apply (case_tac "cap_lift cap'")

View File

@ -518,7 +518,7 @@ lemma finalise_cap_makes_halted_proof[CNodeInv_AI_assms]:
crunch emptyable[wp, CNodeInv_AI_assms]: finalise_cap "emptyable sl"
(simp: crunch_simps lift: emptyable_lift
(simp: crunch_simps rule: emptyable_lift
wp: crunch_wps suspend_emptyable unbind_notification_invs unbind_maybe_notification_invs)

View File

@ -30,12 +30,16 @@ ML {*
structure CrunchNoIrqInstance : CrunchInstance =
struct
type extra = unit;
val eq_extra = op =;
val name = "no_irq";
val has_preconds = false;
fun mk_term _ body _ =
(Syntax.parse_term @{context} "no_irq") $ body;
fun get_precond _ = error "crunch no_irq should not be calling get_precond";
fun put_precond _ _ = error "crunch no_irq should not be calling put_precond";
fun dest_term (Const (@{const_name no_irq}, _) $ body)
= SOME (Term.dummy, body, ())
| dest_term _ = NONE;
val pre_thms = [];
val wpc_tactic = wp_cases_tactic_weak;
fun parse_extra ctxt extra

View File

@ -1076,20 +1076,13 @@ next
done
qed (auto simp: rec_del_dom rec_del_fails)
lemmas rec_del_preservation =
lemmas rec_del_preservation[crunch_rules] =
validE_valid [OF use_spec(2) [OF rec_del_preservation']]
end
lemma cap_swap_fd_typ_at:
"\<lbrace>\<lambda>s. P (typ_at T p s)\<rbrace> cap_swap_for_delete src dst \<lbrace>\<lambda>_ s. P (typ_at T p s)\<rbrace>"
apply(simp add: cap_swap_for_delete_def)
apply(wp cap_swap_typ_at)
apply(simp)
done
crunch typ_at: cap_swap_for_delete "\<lambda>s. P (typ_at T p s)"
lemma cap_swap_valid_cap:
"\<lbrace>valid_cap c\<rbrace> cap_swap_for_delete x y \<lbrace>\<lambda>_. valid_cap c\<rbrace>"
@ -1109,12 +1102,8 @@ lemma cap_swap_cte_at:
context CNodeInv_AI begin
lemma rec_del_typ_at:
"\<And>P T p call. \<lbrace>\<lambda>s::'state_ext state. P (typ_at T p s)\<rbrace> rec_del call \<lbrace>\<lambda>_ s. P (typ_at T p s)\<rbrace>"
by (wp rec_del_preservation cancel_all_ipc_typ_at cancel_all_signals_typ_at
cap_swap_fd_typ_at empty_slot_typ_at set_cap_typ_at
irq_state_independent_AI preemption_point_inv
| simp)+
crunch typ_at: rec_del "\<lambda>s::'state_ext state. P (typ_at T p s)"
(ignore: preemption_point wp: preemption_point_inv)
lemma rec_del_cte_at:
"\<And>c call. \<lbrace>cte_at c :: 'state_ext state \<Rightarrow> bool\<rbrace> rec_del call \<lbrace>\<lambda>_. cte_at c\<rbrace>"
@ -2470,10 +2459,10 @@ lemma suspend_emptyable[wp]:
crunch emptyable[wp]: do_machine_op "emptyable sl"
(lift: emptyable_lift)
(rule: emptyable_lift)
crunch emptyable[wp]: set_irq_state "emptyable sl"
(lift: emptyable_lift)
(rule: emptyable_lift)
declare get_irq_slot_real_cte [wp]

View File

@ -29,15 +29,17 @@ crunch_ignore (add:
liftE whenE unlessE throw_opt
assertE liftM liftME sequence_x
zipWithM_x mapM_x sequence mapM sequenceE_x
mapME_x catch select_f
sequenceE mapME mapME_x catch select_f
handleE' handleE handle_elseE forM forM_x
zipWithM)
zipWithM filterM forME_x)
crunch_ignore (add: cap_fault_on_failure lookup_error_on_failure)
crunch_ignore (add: cap_fault_on_failure lookup_error_on_failure
without_preemption const_on_failure ignore_failure
empty_on_failure unify_failure throw_on_false)
crunch_ignore (add:
storeWord storeWordVM loadWord setRegister getRegister getRestartPC
set_register get_register setNextPC maskInterrupt throw_on_false)
set_register get_register setNextPC maskInterrupt)
crunch_ignore (add:
cap_swap_ext cap_move_ext cap_insert_ext empty_slot_ext create_cap_ext tcb_sched_action

View File

@ -216,7 +216,6 @@ crunch typ_at[wp]: cancel_ipc, reply_cancel_ipc, unbind_maybe_notification
(wp: crunch_wps hoare_vcg_split_ifE select_wp
simp: crunch_simps unless_def)
lemma cancel_ipc_tcb [wp]:
"\<lbrace>tcb_at t\<rbrace> cancel_ipc t' \<lbrace>\<lambda>rv. tcb_at t\<rbrace>"
by (simp add: tcb_at_typ) wp

View File

@ -438,7 +438,7 @@ lemma valid_remove_rights_If[simp]:
declare const_on_failure_wp [wp]
crunch ex_cte_cap_wp_to [wp]: set_extra_badge "ex_cte_cap_wp_to P p"
(lift: ex_cte_cap_to_pres)
(rule: ex_cte_cap_to_pres)
lemma cap_insert_assume_null:
"\<lbrace>P\<rbrace> cap_insert cap src dest \<lbrace>Q\<rbrace> \<Longrightarrow>

View File

@ -24,16 +24,20 @@ crunch_ignore (add:
liftE whenE unlessE throw_opt
assertE liftM liftME sequence_x
zipWithM_x mapM_x sequence mapM sequenceE_x
mapME_x catch select_f
sequenceE mapME mapME_x catch select_f
handleE' handleE handle_elseE forM forM_x
zipWithM)
zipWithM filterM forME_x)
crunch_ignore (add:
withoutFailure throw catchFailure rethrowFailure
capFaultOnFailure lookupErrorOnFailure
nullCapOnFailure nothingOnFailure
withoutPreemption preemptionPoint
cap_fault_on_failure lookup_error_on_failure)
without_preemption withoutPreemption preemptionPoint
cap_fault_on_failure lookup_error_on_failure
const_on_failure ignore_failure ignoreFailure
empty_on_failure emptyOnFailure unifyFailure
unify_failure throw_on_false
)
context Arch begin (*FIXME: arch_split*)

View File

@ -5771,7 +5771,7 @@ proof (rule mdb_chunked_update_final [OF chunk, OF slot])
qed
crunch tcb_in_cur_domain'[wp]: updateCap "tcb_in_cur_domain' t"
(wp: crunch_wps simp: crunch_simps lift: tcb_in_cur_domain'_lift)
(wp: crunch_wps simp: crunch_simps rule: tcb_in_cur_domain'_lift)
lemma make_zombie_invs':
"\<lbrace>\<lambda>s. invs' s \<and> s \<turnstile>' cap \<and>
@ -9571,7 +9571,7 @@ lemma updateCap_noop_irq_handlers:
done
crunch ct_idle_or_in_cur_domain'[wp]: updateCap ct_idle_or_in_cur_domain'
(lift: ct_idle_or_in_cur_domain'_lift2)
(rule: ct_idle_or_in_cur_domain'_lift2)
lemma updateCap_noop_invs:
"\<lbrace>invs' and cte_wp_at' (\<lambda>cte. cteCap cte = cap) slot\<rbrace>

View File

@ -4199,7 +4199,7 @@ lemma setupReplyMaster_irq_handlers'[wp]:
crunch irq_states' [wp]: setupReplyMaster valid_irq_states'
crunch irqs_makes' [wp]: setupReplyMaster irqs_masked'
(lift: irqs_masked_lift)
(rule: irqs_masked_lift)
crunch pde_mappings' [wp]: setupReplyMaster valid_pde_mappings'
crunch pred_tcb_at' [wp]: setupReplyMaster "pred_tcb_at' proj P t"

View File

@ -2952,7 +2952,7 @@ crunch sch_act_simple[wp]: cteDeleteOne, unbindNotification sch_act_simple
(wp: crunch_wps ssa_sch_act_simple sts_sch_act_simple getObject_inv
loadObject_default_inv
simp: crunch_simps unless_def
lift: sch_act_simple_lift
rule: sch_act_simple_lift
ignore: getObject)
crunch valid_queues[wp]: setSchedulerAction "Invariants_H.valid_queues"
@ -3406,7 +3406,7 @@ crunch nosch[wp]: "Arch.finaliseCap" "\<lambda>s. P (ksSchedulerAction s)"
crunch sch_act_simple[wp]: finaliseCap sch_act_simple
(simp: crunch_simps
lift: sch_act_simple_lift
rule: sch_act_simple_lift
wp: getObject_inv loadObject_default_inv crunch_wps
ignore: getObject)

View File

@ -680,7 +680,7 @@ qed
declare constOnFailure_wp [wp]
lemma transferCapsToSlots_pres1:
lemma transferCapsToSlots_pres1[crunch_rules]:
assumes x: "\<And>cap src dest. \<lbrace>P\<rbrace> cteInsert cap src dest \<lbrace>\<lambda>rv. P\<rbrace>"
assumes eb: "\<And>b n. \<lbrace>P\<rbrace> setExtraBadge buffer b n \<lbrace>\<lambda>_. P\<rbrace>"
shows "\<lbrace>P\<rbrace> transferCapsToSlots ep buffer n caps slots mi \<lbrace>\<lambda>rv. P\<rbrace>"
@ -731,7 +731,7 @@ lemma lsfco_cte_wp_at_univ':
done
crunch ex_cte_cap_wp_to' [wp]: setExtraBadge "ex_cte_cap_wp_to' P p"
(lift: ex_cte_cap_to'_pres)
(rule: ex_cte_cap_to'_pres)
crunch valid_objs' [wp]: setExtraBadge valid_objs'
crunch aligned' [wp]: setExtraBadge pspace_aligned'
@ -907,7 +907,8 @@ lemma transferCapsToSlots_vp[wp]:
apply (fastforce simp: cte_wp_at_ctes_of dest: ctes_of_valid')
done
crunch sch_act [wp]: setExtraBadge "\<lambda>s. P (ksSchedulerAction s)"
crunch sch_act [wp]: setExtraBadge, doIPCTransfer "\<lambda>s. P (ksSchedulerAction s)"
(wp: crunch_wps mapME_wp' simp: zipWithM_x_mapM)
crunch pred_tcb_at' [wp]: setExtraBadge "\<lambda>s. pred_tcb_at' proj P p s"
crunch ksCurThread[wp]: setExtraBadge "\<lambda>s. P (ksCurThread s)"
crunch ksCurDomain[wp]: setExtraBadge "\<lambda>s. P (ksCurDomain s)"
@ -1141,7 +1142,6 @@ lemma transferCapsToSlots_vms[wp]:
crunch pspace_domain_valid[wp]: setExtraBadge, transferCapsToSlots
"pspace_domain_valid"
(lift: transferCapsToSlots_pres1)
crunch ct_not_inQ[wp]: setExtraBadge "ct_not_inQ"
@ -1153,15 +1153,11 @@ lemma tcts_ct_not_inQ[wp]:
crunch ct_idle_or_in_cur_domain'[wp]: setExtraBadge ct_idle_or_in_cur_domain'
crunch ct_idle_or_in_cur_domain'[wp]: transferCapsToSlots ct_idle_or_in_cur_domain'
(lift: transferCapsToSlots_pres1)
crunch ksCurDomain[wp]: transferCapsToSlots "\<lambda>s. P (ksCurDomain s)"
(lift: transferCapsToSlots_pres1)
crunch ksDomSchedule[wp]: setExtraBadge "\<lambda>s. P (ksDomSchedule s)"
crunch ksDomScheduleIdx[wp]: setExtraBadge "\<lambda>s. P (ksDomScheduleIdx s)"
crunch ksDomSchedule[wp]: transferCapsToSlots "\<lambda>s. P (ksDomSchedule s)"
(lift: transferCapsToSlots_pres1)
crunch ksDomScheduleIdx[wp]: transferCapsToSlots "\<lambda>s. P (ksDomScheduleIdx s)"
(lift: transferCapsToSlots_pres1)
lemma transferCapsToSlots_invs[wp]:
@ -2018,7 +2014,7 @@ crunch pde_mappings'[wp]: doIPCTransfer "valid_pde_mappings'"
(wp: crunch_wps simp: crunch_simps)
crunch irqs_masked'[wp]: doIPCTransfer "irqs_masked'"
(wp: crunch_wps simp: crunch_simps lift: irqs_masked_lift)
(wp: crunch_wps simp: crunch_simps rule: irqs_masked_lift)
lemma doIPCTransfer_invs[wp]:
"\<lbrace>invs' and tcb_at' s and tcb_at' r\<rbrace>
@ -2032,9 +2028,7 @@ lemma doIPCTransfer_invs[wp]:
done
crunch nosch[wp]: doIPCTransfer "\<lambda>s. P (ksSchedulerAction s)"
(ignore: getRestartPC setRegister transferCapsToSlots
wp: hoare_drop_imps hoare_vcg_split_case_option
transferCapsToSlots_pres1 mapM_wp'
(wp: hoare_drop_imps hoare_vcg_split_case_option mapM_wp'
simp: split_def zipWithM_x_mapM)
lemma handle_fault_reply_registers_corres:
@ -2643,7 +2637,7 @@ crunch ksDomSchedule[wp]: setupCallerCap, switchIfRequiredTo, doIPCTransfer, att
(wp: crunch_wps simp: zipWithM_x_mapM)
crunch tcbDomain_obj_at'[wp]: doIPCTransfer "obj_at' (\<lambda>tcb. P (tcbDomain tcb)) t"
(wp: crunch_wps constOnFailure_wp simp: crunch_simps lift: transferCapsToSlots_pres1)
(wp: crunch_wps constOnFailure_wp simp: crunch_simps)
lemma send_ipc_corres:
(* call is only true if called in handleSyscall SysCall, which
@ -3113,7 +3107,7 @@ crunch it'[wp]: sendSignal "\<lambda>s. P (ksIdleThread s)"
(wp: crunch_wps simp: crunch_simps)
crunch irqs_masked'[wp]: sendSignal "irqs_masked'"
(wp: crunch_wps getObject_inv loadObject_default_inv simp: crunch_simps unless_def lift: irqs_masked_lift ignore: getObject)
(wp: crunch_wps getObject_inv loadObject_default_inv simp: crunch_simps unless_def rule: irqs_masked_lift ignore: getObject)
lemma sts_running_valid_queues:
"runnable' st \<Longrightarrow> \<lbrace> Invariants_H.valid_queues \<rbrace> setThreadState st t \<lbrace>\<lambda>_. Invariants_H.valid_queues \<rbrace>"
@ -3357,6 +3351,8 @@ lemma ntfn_q_refs_no_TCBBound':
"(x, TCBBound) \<notin> ntfn_q_refs_of' ntfn"
by (simp add: ntfn_q_refs_of'_def split: ntfn.splits)
crunch nosch[wp]: setMRs "\<lambda>s. P (ksSchedulerAction s)"
lemma sai_invs'[wp]:
"\<lbrace>invs' and ex_nonz_cap_to' ntfnptr\<rbrace>
sendSignal ntfnptr badge \<lbrace>\<lambda>y. invs'\<rbrace>"
@ -3881,7 +3877,7 @@ lemma setupCallerCap_state_refs_of[wp]:
crunch sch_act_wf: setupCallerCap
"\<lambda>s. sch_act_wf (ksSchedulerAction s) s"
(wp: crunch_wps ssa_sch_act sts_sch_act lift: sch_act_wf_lift ignore:setObject)
(wp: crunch_wps ssa_sch_act sts_sch_act rule: sch_act_wf_lift ignore:setObject)
lemma setCTE_valid_queues[wp]:
"\<lbrace>Invariants_H.valid_queues\<rbrace> setCTE ptr val \<lbrace>\<lambda>rv. Invariants_H.valid_queues\<rbrace>"
@ -4061,7 +4057,7 @@ crunch pde_mappings' [wp]: setupCallerCap valid_pde_mappings'
(wp: crunch_wps)
crunch irqs_masked' [wp]: receiveIPC "irqs_masked'"
(wp: crunch_wps lift: irqs_masked_lift)
(wp: crunch_wps rule: irqs_masked_lift)
crunch ct_not_inQ[wp]: getThreadCallerSlot "ct_not_inQ"
crunch ct_not_inQ[wp]: getThreadReplySlot "ct_not_inQ"
@ -4076,15 +4072,13 @@ crunch ksQ[wp]: copyMRs "\<lambda>s. P (ksReadyQueues s)"
(wp: mapM_wp' hoare_drop_imps simp: crunch_simps)
crunch ksQ[wp]: doIPCTransfer "\<lambda>s. P (ksReadyQueues s)"
(ignore: getRestartPC setRegister transferCapsToSlots
wp: hoare_drop_imps hoare_vcg_split_case_option
transferCapsToSlots_pres1 mapM_wp'
(wp: hoare_drop_imps hoare_vcg_split_case_option
mapM_wp'
simp: split_def zipWithM_x_mapM)
crunch ct'[wp]: doIPCTransfer "\<lambda>s. P (ksCurThread s)"
(ignore: getRestartPC setRegister transferCapsToSlots
wp: hoare_drop_imps hoare_vcg_split_case_option
transferCapsToSlots_pres1 mapM_wp'
(wp: hoare_drop_imps hoare_vcg_split_case_option
mapM_wp'
simp: split_def zipWithM_x_mapM)
lemma asUser_ct_not_inQ[wp]:
@ -4099,7 +4093,7 @@ crunch ct_not_inQ[wp]: copyMRs "ct_not_inQ"
crunch ct_not_inQ[wp]: doIPCTransfer "ct_not_inQ"
(ignore: getRestartPC setRegister transferCapsToSlots
wp: hoare_drop_imps hoare_vcg_split_case_option
transferCapsToSlots_pres1 mapM_wp'
mapM_wp'
simp: split_def zipWithM_x_mapM)
lemma ntfn_q_refs_no_bound_refs': "rf : ntfn_q_refs_of' (ntfnObj ob) \<Longrightarrow> rf ~: ntfn_bound_refs' (ntfnBoundTCB ob')"

View File

@ -1467,7 +1467,7 @@ crunch st_tcb_at'[wp]: replyFromKernel "st_tcb_at' P t"
crunch cap_to'[wp]: replyFromKernel "ex_nonz_cap_to' p"
crunch it'[wp]: replyFromKernel "\<lambda>s. P (ksIdleThread s)"
crunch sch_act_simple[wp]: replyFromKernel sch_act_simple
(lift: sch_act_simple_lift)
(rule: sch_act_simple_lift)
lemma rfk_ksQ[wp]:
"\<lbrace>\<lambda>s. P (ksReadyQueues s p)\<rbrace> replyFromKernel t x1 \<lbrace>\<lambda>_ s. P (ksReadyQueues s p)\<rbrace>"
@ -1668,7 +1668,7 @@ lemma deleteCallerCap_nonz_cap:
crunch sch_act_sane[wp]: cteDeleteOne sch_act_sane
(wp: crunch_wps ss_sch_act_sane_weak loadObject_default_inv getObject_inv
simp: crunch_simps unless_def
lift: sch_act_sane_lift
rule: sch_act_sane_lift
ignore: getObject)
crunch sch_act_sane[wp]: deleteCallerCap sch_act_sane

View File

@ -337,7 +337,7 @@ lemma readreg_corres:
done
crunch sch_act_simple [wp]: asUser "sch_act_simple"
(lift: sch_act_simple_lift)
(rule: sch_act_simple_lift)
lemma writereg_corres:
"corres (intr \<oplus> op =) (einvs and tcb_at dest and ex_nonz_cap_to dest)
@ -1032,7 +1032,7 @@ lemma assertDerived_wp:
done
crunch sch_act_simple[wp]: setPriority sch_act_simple
(wp: ssa_sch_act_simple crunch_wps lift: sch_act_simple_lift simp: crunch_simps)
(wp: ssa_sch_act_simple crunch_wps rule: sch_act_simple_lift simp: crunch_simps)
(* For some reason, when this was embedded in a larger expression clarsimp wouldn't remove it. Adding it as a simp rule does *)
lemma inQ_tc_corres_helper:

View File

@ -5741,7 +5741,7 @@ crunch vq'[wp]: insertNewCap valid_queues'
(wp: crunch_wps)
crunch irqs_masked' [wp]: insertNewCap irqs_masked'
(wp: crunch_wps lift: irqs_masked_lift)
(wp: crunch_wps rule: irqs_masked_lift)
crunch valid_machine_state'[wp]: insertNewCap valid_machine_state'
(wp: crunch_wps)

View File

@ -107,6 +107,7 @@ lemmas (in Arch) [simp] =
instantiation ARM_H.pde :: pspace_storable
begin
interpretation Arch .

View File

@ -18,6 +18,7 @@ imports
begin
context Arch begin global_naming ARM_H
type_synonym asid = "word32"
definition

View File

@ -103,6 +103,8 @@ lemmas (in Arch) [simp] =
-- --------------------------------------
#INCLUDE_SETTINGS keep_constructor = asidpool
#INCLUDE_HASKELL_PREPARSE SEL4/Object/Structures/ARM.lhs
#INCLUDE_HASKELL_PREPARSE SEL4/Machine/Hardware/ARM.lhs

View File

@ -16,6 +16,8 @@ imports
begin
context Arch begin global_naming ARM_H
#INCLUDE_SETTINGS keep_constructor=asidpool
#INCLUDE_HASKELL SEL4/Object/Structures/ARM.lhs CONTEXT ARM_H decls_only
#INCLUDE_HASKELL SEL4/Object/Structures/ARM.lhs CONTEXT ARM_H instanceproofs
#INCLUDE_HASKELL SEL4/Object/Structures/ARM.lhs CONTEXT ARM_H bodies_only

View File

@ -3,777 +3,3 @@ Built from git repo at /Users/xgao/verification/l4v/spec/haskell by xgao
Generated from changeset:
75c159b untypeddevice-proof: fix drefine
Warning - uncomitted changes used:
M ../../.gitignore
A ../../.licenseignore
M ../../README.md
M ../../ROOTS
M ../../camkes/ROOT
M ../../camkes/adl-spec/document/root.tex
M ../../camkes/cdl-refine/Generator_CAMKES_CDL.thy
M ../../camkes/cdl-refine/Types_CAMKES_CDL.thy
M ../../camkes/glue-proofs/document/eventfrom-emit-underlying.c
M ../../camkes/glue-proofs/document/eventto-poll.c
M ../../camkes/glue-proofs/document/eventto-wait.c
M ../../camkes/glue-proofs/document/from-echo-int.c
M ../../camkes/glue-proofs/document/root.tex
M ../../camkes/glue-proofs/document/simple.camkes
M ../../camkes/glue-proofs/document/to-echo-int.c
M ../../camkes/glue-spec/document/dataport.camkes
M ../../camkes/glue-spec/document/event.camkes
M ../../camkes/glue-spec/document/filter.camkes
M ../../camkes/glue-spec/document/root.tex
M ../../camkes/glue-spec/document/simple.camkes
A ../../lib/ARM/WordSetup.thy
R ../../lib/wp/AdjustSchematic.thy -> ../../lib/AdjustSchematic.thy
M ../../lib/Apply_Trace.thy
M ../../lib/AutoLevity.thy
A ../../lib/AutoLevity_Base.thy
A ../../lib/AutoLevity_Hooks.thy
M ../../lib/AutoLevity_Run.thy
A ../../lib/AutoLevity_Test.thy
A ../../lib/AutoLevity_Theory_Report.thy
M ../../lib/BCorres_UL.thy
M ../../lib/Bisim_UL.thy
A ../../lib/BitFieldProofsLib.thy
M ../../lib/CTranslationNICTA.thy
M ../../lib/Conjuncts.thy
M ../../lib/Corres_UL.thy
M ../../lib/Crunch.thy
M ../../lib/Crunch_Test.thy
M ../../lib/Defs.thy
D ../../lib/DistinctProp.thy
D ../../lib/DistinctPropLemmaBucket.thy
A ../../lib/Distinct_Prop.thy
M ../../lib/Eisbach_Methods.thy
M ../../lib/EmptyFailLib.thy
D ../../lib/Enumeration.thy
M ../../lib/Etanercept.thy
M ../../lib/ExpandAll.thy
A ../../lib/Extend_Locale.thy
M ../../lib/ExtraCorres.thy
M ../../lib/GenericLib.thy
D ../../lib/HOLLemmaBucket.thy
M ../../lib/HaskellLemmaBucket.thy
M ../../lib/HaskellLib_H.thy
M ../../lib/Hoare_Sep_Tactics/Hoare_Sep_Tactics.thy
M ../../lib/Insulin.thy
M ../../lib/LemmaBucket.thy
M ../../lib/LemmaBucket_C.thy
M ../../lib/Lib.thy
D ../../lib/Methods.thy
M ../../lib/MonadEq.thy
R ../../lib/wp/NonDetMonad.thy -> ../../lib/Monad_WP/NonDetMonad.thy
R ../../lib/wp/NonDetMonadLemmas.thy -> ../../lib/Monad_WP/NonDetMonadLemmas.thy
R ../../lib/wp/NonDetMonadVCG.thy -> ../../lib/Monad_WP/NonDetMonadVCG.thy
R ../../lib/OptionMonad.thy -> ../../lib/Monad_WP/OptionMonad.thy
R ../../lib/OptionMonadND.thy -> ../../lib/Monad_WP/OptionMonadND.thy
R ../../lib/OptionMonadWP.thy -> ../../lib/Monad_WP/OptionMonadWP.thy
R ../../lib/Strengthen.thy -> ../../lib/Monad_WP/Strengthen.thy
R ../../lib/WhileLoopRules.thy -> ../../lib/Monad_WP/WhileLoopRules.thy
R ../../lib/WhileLoopRulesCompleteness.thy -> ../../lib/Monad_WP/WhileLoopRulesCompleteness.thy
R ../../lib/wp/Eisbach_WP.thy -> ../../lib/Monad_WP/wp/Eisbach_WP.thy
R ../../lib/wp/WP-method.ML -> ../../lib/Monad_WP/wp/WP-method.ML
R ../../lib/wp/WP.thy -> ../../lib/Monad_WP/wp/WP.thy
R ../../lib/wp/WPBang.thy -> ../../lib/Monad_WP/wp/WPBang.thy
R ../../lib/wp/WPC.thy -> ../../lib/Monad_WP/wp/WPC.thy
R ../../lib/wp/WPEx.thy -> ../../lib/Monad_WP/wp/WPEx.thy
R ../../lib/wp/WPI.thy -> ../../lib/Monad_WP/wp/WPI.thy
M ../../lib/MonadicRewrite.thy
D ../../lib/MoreDivides.thy
D ../../lib/NICTACompat.thy
M ../../lib/NICTATools.thy
M ../../lib/NonDetMonadLemmaBucket.thy
M ../../lib/ProvePart.thy
A ../../lib/Qualify.thy
A ../../lib/Requalify.thy
M ../../lib/Rule_By_Method.thy
M ../../lib/ShowTypes.thy
M ../../lib/SimplRewrite.thy
M ../../lib/Solves_Tac.thy
M ../../lib/StateMonad.thy
M ../../lib/SubMonadLib.thy
M ../../lib/TSubst.thy
M ../../lib/Trace_Attribs.thy
M ../../lib/TypHeapLib.thy
M ../../lib/WPTutorial.thy
D ../../lib/WordBitwiseSigned.thy
D ../../lib/WordSetup.thy
R ../../lib/Aligned.thy -> ../../lib/Word_Lib/Aligned.thy
A ../../lib/Word_Lib/Enumeration.thy
A ../../lib/Word_Lib/HOL_Lemmas.thy
A ../../lib/Word_Lib/Hex_Words.thy
A ../../lib/Word_Lib/More_Divides.thy
A ../../lib/Word_Lib/Norm_Words.thy
A ../../lib/Word_Lib/ROOT
R ../../lib/SignedWords.thy -> ../../lib/Word_Lib/Signed_Words.thy
A ../../lib/Word_Lib/WordBitwise_Signed.thy
R ../../lib/WordEnum.thy -> ../../lib/Word_Lib/Word_Enum.thy
R ../../lib/WordLemmaBucket.thy -> ../../lib/Word_Lib/Word_Lemmas.thy
A ../../lib/Word_Lib/Word_Lemmas_32.thy
A ../../lib/Word_Lib/Word_Lemmas_64.thy
R ../../lib/WordLib.thy -> ../../lib/Word_Lib/Word_Lib.thy
A ../../lib/Word_Lib/Word_Setup_32.thy
A ../../lib/Word_Lib/Word_Setup_64.thy
A ../../lib/Word_Lib/Word_Syntax.thy
A ../../lib/Word_Lib/document/root.tex
M ../../lib/clib/AutoCorresCRefine.thy
A ../../lib/clib/AutoCorresModifiesProofs.thy
M ../../lib/clib/CCorresLemmas.thy
A ../../lib/clib/CCorres_Rewrite.thy
M ../../lib/clib/Corres_C.thy
M ../../lib/crunch-cmd.ML
M ../../lib/defs.ML
M ../../lib/ml-helpers/TermPatternAntiquote.thy
M ../../lib/sep_algebra/Generic_Separation_Algebras.thy
M ../../lib/sep_algebra/Map_Extra.thy
M ../../lib/sep_algebra/MonadSep.thy
M ../../lib/sep_algebra/Sep_Attribs.thy
M ../../lib/sep_algebra/Sep_Cancel.thy
M ../../lib/sep_algebra/Sep_Cancel_Example.thy
M ../../lib/sep_algebra/Sep_Cancel_Set.thy
M ../../lib/sep_algebra/Sep_Eq.thy
M ../../lib/sep_algebra/Sep_Heap_Instance.thy
M ../../lib/sep_algebra/Sep_ImpI.thy
M ../../lib/sep_algebra/Sep_MP_Example.thy
M ../../lib/sep_algebra/Sep_Provers.thy
M ../../lib/sep_algebra/Sep_Provers_Example.thy
M ../../lib/sep_algebra/Sep_Rotate.thy
M ../../lib/sep_algebra/Sep_Rule_Ext.thy
M ../../lib/sep_algebra/Sep_Select.thy
M ../../lib/sep_algebra/Sep_Select_Example.thy
M ../../lib/sep_algebra/Sep_Solve.thy
M ../../lib/sep_algebra/Sep_Solve_Example.thy
M ../../lib/sep_algebra/Sep_Tactic_Helpers.thy
M ../../lib/sep_algebra/Separation_Algebra.thy
M ../../lib/sep_algebra/Separation_Algebra_Alt.thy
M ../../lib/sep_algebra/ex/Sep_Tactics_Test.thy
M ../../lib/sep_algebra/ex/Simple_Separation_Example.thy
M ../../lib/sep_algebra/ex/VM_Example.thy
M ../../lib/sep_algebra/ex/capDL/Abstract_Separation_D.thy
M ../../lib/sep_algebra/ex/capDL/Separation_D.thy
M ../../lib/sep_algebra/ex/capDL/Types_D.thy
M ../../lib/sep_algebra/sep_tactics.ML
M ../../lib/subgoal_focus/Subgoal_Methods.thy
M ../../lib/trace_attribs.ML
M ../../misc/git-commit-emails/commit-email.py
M ../../misc/regression/run_tests.py
M ../../misc/regression/tests.xml
A ../../misc/scripts/check_theory_imports
A ../../misc/scripts/test_munge.sh
M ../../misc/scripts/test_spec.sh
A ../../misc/scripts/tests.xml
M ../../misc/scripts/thydeps
M ../../misc/zsh/_isabelle
M ../../proof/ROOT
M ../../proof/access-control/ADT_AC.thy
M ../../proof/access-control/Access.thy
M ../../proof/access-control/Arch_AC.thy
M ../../proof/access-control/CNode_AC.thy
M ../../proof/access-control/DomainSepInv.thy
M ../../proof/access-control/Dpolicy.thy
M ../../proof/access-control/ExampleSystem.thy
M ../../proof/access-control/Finalise_AC.thy
M ../../proof/access-control/Interrupt_AC.thy
M ../../proof/access-control/Ipc_AC.thy
M ../../proof/access-control/Retype_AC.thy
M ../../proof/access-control/Syscall_AC.thy
M ../../proof/access-control/Tcb_AC.thy
M ../../proof/asmrefine/SEL4GraphRefine.thy
M ../../proof/bisim/Syscall_S.thy
M ../../proof/bisim/document/root.tex
M ../../proof/capDL-api/IRQ_DP.thy
M ../../proof/crefine/ADT_C.thy
UU ../../proof/crefine/Arch_C.thy
A ../../proof/crefine/AutoCorresTest.thy
A ../../proof/crefine/AutoCorres_C.thy
M ../../proof/crefine/BuildRefineCache_C.thy
M ../../proof/crefine/CLevityCatch.thy
M ../../proof/crefine/CSpace_C.thy
M ../../proof/crefine/CSpace_RAB_C.thy
M ../../proof/crefine/Cache.thy
M ../../proof/crefine/Delete_C.thy
M ../../proof/crefine/DetWP.thy
M ../../proof/crefine/Detype_C.thy
M ../../proof/crefine/Fastpath_C.thy
M ../../proof/crefine/Finalise_C.thy
M ../../proof/crefine/Interrupt_C.thy
M ../../proof/crefine/Invoke_C.thy
M ../../proof/crefine/IpcCancel_C.thy
UU ../../proof/crefine/Ipc_C.thy
M ../../proof/crefine/Machine_C.thy
UU ../../proof/crefine/Recycle_C.thy
UU ../../proof/crefine/Refine_C.thy
M ../../proof/crefine/Refine_nondet_C.thy
UU ../../proof/crefine/Retype_C.thy
UU ../../proof/crefine/SR_lemmas_C.thy
M ../../proof/crefine/Schedule_C.thy
M ../../proof/crefine/StateRelation_C.thy
M ../../proof/crefine/StoreWord_C.thy
M ../../proof/crefine/SyscallArgs_C.thy
M ../../proof/crefine/Syscall_C.thy
M ../../proof/crefine/TcbAcc_C.thy
UU ../../proof/crefine/Tcb_C.thy
M ../../proof/crefine/VSpace_C.thy
M ../../proof/crefine/Wellformed_C.thy
UU ../../proof/drefine/Arch_DR.thy
UU ../../proof/drefine/CNode_DR.thy
M ../../proof/drefine/Corres_D.thy
M ../../proof/drefine/Finalise_DR.thy
M ../../proof/drefine/Intent_DR.thy
M ../../proof/drefine/Interrupt_DR.thy
M ../../proof/drefine/Ipc_DR.thy
M ../../proof/drefine/KHeap_DR.thy
M ../../proof/drefine/MoreCorres.thy
M ../../proof/drefine/Refine_D.thy
M ../../proof/drefine/Schedule_DR.thy
M ../../proof/drefine/StateTranslationProofs_DR.thy
UU ../../proof/drefine/StateTranslation_D.thy
M ../../proof/drefine/Syscall_DR.thy
M ../../proof/drefine/Tcb_DR.thy
UU ../../proof/drefine/Untyped_DR.thy
M ../../proof/infoflow/ADT_IF.thy
UU ../../proof/infoflow/ADT_IF_Refine.thy
UU ../../proof/infoflow/ADT_IF_Refine_C.thy
M ../../proof/infoflow/Arch_IF.thy
M ../../proof/infoflow/CNode_IF.thy
M ../../proof/infoflow/Decode_IF.thy
M ../../proof/infoflow/Example_Valid_State.thy
M ../../proof/infoflow/Example_Valid_StateH.thy
M ../../proof/infoflow/FinalCaps.thy
M ../../proof/infoflow/Finalise_IF.thy
UU ../../proof/infoflow/IRQMasks_IF.thy
M ../../proof/infoflow/InfoFlow.thy
M ../../proof/infoflow/Interrupt_IF.thy
UU ../../proof/infoflow/Ipc_IF.thy
M ../../proof/infoflow/Noninterference.thy
M ../../proof/infoflow/Noninterference_Base_Enabledness_weak_asym.thy
M ../../proof/infoflow/Noninterference_Refinement.thy
M ../../proof/infoflow/PasUpdates.thy
M ../../proof/infoflow/Retype_IF.thy
M ../../proof/infoflow/Scheduler_IF.thy
M ../../proof/infoflow/Syscall_IF.thy
M ../../proof/infoflow/Tcb_IF.thy
M ../../proof/infoflow/UserOp_IF.thy
M ../../proof/invariant-abstract/ADT_AI.thy
MM ../../proof/invariant-abstract/AInvs.thy
AM ../../proof/invariant-abstract/AInvsPre.thy
A ../../proof/invariant-abstract/ARM/ArchADT_AI.thy
AM ../../proof/invariant-abstract/ARM/ArchAInvsPre.thy
R ../../proof/invariant-abstract/ArchAcc_AI.thy -> ../../proof/invariant-abstract/ARM/ArchAcc_AI.thy
C ../../proof/invariant-abstract/Arch_AI.thy -> ../../proof/invariant-abstract/ARM/ArchArch_AI.thy
A ../../proof/invariant-abstract/ARM/ArchBCorres2_AI.thy
A ../../proof/invariant-abstract/ARM/ArchBits_AI.thy
A ../../proof/invariant-abstract/ARM/ArchCNodeInv_AI.thy
A ../../proof/invariant-abstract/ARM/ArchCSpaceInvPre_AI.thy
A ../../proof/invariant-abstract/ARM/ArchCSpaceInv_AI.thy
A ../../proof/invariant-abstract/ARM/ArchCSpacePre_AI.thy
A ../../proof/invariant-abstract/ARM/ArchCSpace_AI.thy
A ../../proof/invariant-abstract/ARM/ArchCrunchSetup_AI.thy
AM ../../proof/invariant-abstract/ARM/ArchDetSchedAux_AI.thy
AM ../../proof/invariant-abstract/ARM/ArchDeterministic_AI.thy
A ../../proof/invariant-abstract/ARM/ArchDetype_AI.thy
A ../../proof/invariant-abstract/ARM/ArchEmptyFail_AI.thy
C ../../proof/invariant-abstract/Finalise_AI.thy -> ../../proof/invariant-abstract/ARM/ArchFinalise_AI.thy
A ../../proof/invariant-abstract/ARM/ArchInterruptAcc_AI.thy
A ../../proof/invariant-abstract/ARM/ArchInterrupt_AI.thy
A ../../proof/invariant-abstract/ARM/ArchInvariants_AI.thy
A ../../proof/invariant-abstract/ARM/ArchIpcCancel_AI.thy
A ../../proof/invariant-abstract/ARM/ArchIpc_AI.thy
A ../../proof/invariant-abstract/ARM/ArchKHeap_AI.thy
AM ../../proof/invariant-abstract/ARM/ArchKernelInit_AI.thy
A ../../proof/invariant-abstract/ARM/ArchLevityCatch_AI.thy
A ../../proof/invariant-abstract/ARM/ArchRetype_AI.thy
A ../../proof/invariant-abstract/ARM/ArchSchedule_AI.thy
A ../../proof/invariant-abstract/ARM/ArchSyscall_AI.thy
A ../../proof/invariant-abstract/ARM/ArchTcbAcc_AI.thy
A ../../proof/invariant-abstract/ARM/ArchTcb_AI.thy
A ../../proof/invariant-abstract/ARM/ArchUntyped_AI.thy
RM ../../proof/invariant-abstract/PDPTEntries_AI.thy -> ../../proof/invariant-abstract/ARM/ArchVSpaceEntries_AI.thy
C ../../proof/invariant-abstract/VSpace_AI.thy -> ../../proof/invariant-abstract/ARM/ArchVSpace_AI.thy
R ../../proof/invariant-abstract/Machine_AI.thy -> ../../proof/invariant-abstract/ARM/Machine_AI.thy
M ../../proof/invariant-abstract/Arch_AI.thy
M ../../proof/invariant-abstract/BCorres2_AI.thy
M ../../proof/invariant-abstract/BCorres_AI.thy
M ../../proof/invariant-abstract/Bits_AI.thy
M ../../proof/invariant-abstract/CNodeInv_AI.thy
A ../../proof/invariant-abstract/CSpaceInvPre_AI.thy
M ../../proof/invariant-abstract/CSpaceInv_AI.thy
A ../../proof/invariant-abstract/CSpacePre_AI.thy
M ../../proof/invariant-abstract/CSpace_AI.thy
MM ../../proof/invariant-abstract/DetSchedAux_AI.thy
M ../../proof/invariant-abstract/DetSchedInvs_AI.thy
MM ../../proof/invariant-abstract/DetSchedSchedule_AI.thy
MM ../../proof/invariant-abstract/Deterministic_AI.thy
M ../../proof/invariant-abstract/Detype_AI.thy
M ../../proof/invariant-abstract/EmptyFail_AI.thy
M ../../proof/invariant-abstract/Finalise_AI.thy
M ../../proof/invariant-abstract/Include_AI.thy
M ../../proof/invariant-abstract/InterruptAcc_AI.thy
M ../../proof/invariant-abstract/Interrupt_AI.thy
A ../../proof/invariant-abstract/InvariantsPre_AI.thy
M ../../proof/invariant-abstract/Invariants_AI.thy
M ../../proof/invariant-abstract/IpcCancel_AI.thy
M ../../proof/invariant-abstract/Ipc_AI.thy
A ../../proof/invariant-abstract/KHeapPre_AI.thy
M ../../proof/invariant-abstract/KHeap_AI.thy
M ../../proof/invariant-abstract/KernelInitSepProofs_AI.thy
MM ../../proof/invariant-abstract/KernelInit_AI.thy
M ../../proof/invariant-abstract/LevityCatch_AI.thy
M ../../proof/invariant-abstract/Retype_AI.thy
M ../../proof/invariant-abstract/Schedule_AI.thy
M ../../proof/invariant-abstract/Syscall_AI.thy
M ../../proof/invariant-abstract/TcbAcc_AI.thy
M ../../proof/invariant-abstract/Tcb_AI.thy
M ../../proof/invariant-abstract/Untyped_AI.thy
A ../../proof/invariant-abstract/VSpaceEntries_AI.thy
A ../../proof/invariant-abstract/VSpacePre_AI.thy
M ../../proof/invariant-abstract/VSpace_AI.thy
M ../../proof/refine/ADT_H.thy
M ../../proof/refine/ArchAcc_R.thy
M ../../proof/refine/Arch_R.thy
M ../../proof/refine/Bits_R.thy
M ../../proof/refine/CNodeInv_R.thy
M ../../proof/refine/CSpace1_R.thy
M ../../proof/refine/CSpace_I.thy
M ../../proof/refine/CSpace_R.thy
M ../../proof/refine/Corres.thy
UU ../../proof/refine/Detype_R.thy
M ../../proof/refine/EmptyFail.thy
M ../../proof/refine/EmptyFail_H.thy
M ../../proof/refine/Finalise_R.thy
M ../../proof/refine/InterruptAcc_R.thy
M ../../proof/refine/Interrupt_R.thy
UU ../../proof/refine/Invariants_H.thy
M ../../proof/refine/Invocations_R.thy
M ../../proof/refine/IpcCancel_R.thy
M ../../proof/refine/Ipc_R.thy
M ../../proof/refine/KHeap_R.thy
M ../../proof/refine/LevityCatch.thy
M ../../proof/refine/Machine_R.thy
M ../../proof/refine/Orphanage.thy
UU ../../proof/refine/PageTableDuplicates.thy
M ../../proof/refine/RAB_FN.thy
M ../../proof/refine/Refine.thy
UU ../../proof/refine/Retype_R.thy
M ../../proof/refine/Schedule_R.thy
M ../../proof/refine/StateRelation.thy
M ../../proof/refine/SubMonad_R.thy
M ../../proof/refine/Syscall_R.thy
UU ../../proof/refine/TcbAcc_R.thy
M ../../proof/refine/Tcb_R.thy
UU ../../proof/refine/Untyped_R.thy
M ../../proof/refine/VSpace_R.thy
M ../../proof/sep-capDL/Lookups_D.thy
M ../../proof/sep-capDL/Separation_SD.thy
M ../ROOT
M ../abstract/ARM/ArchCSpace_A.thy
M ../abstract/ARM/ArchDecode_A.thy
M ../abstract/ARM/ArchInvocation_A.thy
M ../abstract/ARM/ArchRetype_A.thy
M ../abstract/ARM/ArchVMRights_A.thy
M ../abstract/ARM/ArchVSpaceAcc_A.thy
M ../abstract/ARM/ArchVSpace_A.thy
M ../abstract/ARM/Arch_A.thy
M ../abstract/ARM/Arch_Structs_A.thy
M ../abstract/ARM/Init_A.thy
M ../abstract/ARM/Machine_A.thy
M ../abstract/CSpace_A.thy
M ../abstract/Decode_A.thy
M ../abstract/Deterministic_A.thy
M ../abstract/Interrupt_A.thy
M ../abstract/InvocationLabels_A.thy
M ../abstract/Invocations_A.thy
M ../abstract/Ipc_A.thy
M ../abstract/KHeap_A.thy
M ../abstract/KernelInit_A.thy
M ../abstract/MiscMachine_A.thy
M ../abstract/README.md
M ../abstract/Retype_A.thy
M ../abstract/Schedule_A.thy
MM ../abstract/Structures_A.thy
M ../abstract/Syscall_A.thy
M ../abstract/TcbAcc_A.thy
M ../abstract/Tcb_A.thy
A ../abstract/VMRights_A.thy
M ../abstract/document/root.tex
M ../capDL/CSpace_D.thy
M ../capDL/Intents_D.thy
M ../capDL/Interrupt_D.thy
M ../capDL/Monads_D.thy
M ../capDL/PageTableUnmap_D.thy
M ../capDL/Types_D.thy
M ../cspec/KernelState_C.thy
M ../cspec/Kernel_C.thy
M ../cspec/SetSorryModifiesProofs.thy
M ../cspec/Substitute.thy
M ../design/API_H.thy
M ../design/ARM/ArchInterruptDecls_H.thy
M ../design/ARM/ArchInterrupt_H.thy
M ../design/ARM/ArchInvocationLabels_H.thy
M ../design/ARM/ArchLabelFuns_H.thy
M ../design/ARM/ArchObjInsts_H.thy
UU ../design/ARM/ArchRetypeDecls_H.thy
UU ../design/ARM/ArchRetype_H.thy
M ../design/ARM/ArchStateData_H.thy
MM ../design/ARM/ArchStructures_H.thy
M ../design/ARM/ArchTCB_H.thy
M ../design/ARM/ArchThreadDecls_H.thy
M ../design/ARM/ArchThread_H.thy
M ../design/ARM/ArchTypes_H.thy
M ../design/ARM/ArchVSpaceDecls_H.thy
UU ../design/ARM/ArchVSpace_H.thy
M ../design/ARM/Arch_Structs_B.thy
M ../design/ARM/Hardware_H.thy
UU ../design/ARM/Intermediate_H.thy
M ../design/ARM/RegisterSet_H.thy
M ../design/ARM/State_H.thy
UU ../design/CNode_H.thy
M ../design/CSpaceDecls_H.thy
M ../design/CSpace_H.thy
M ../design/Config_H.thy
M ../design/Delete_H.thy
M ../design/EndpointDecls_H.thy
M ../design/Endpoint_H.thy
M ../design/Event_H.thy
M ../design/FaultHandlerDecls_H.thy
M ../design/FaultHandler_H.thy
M ../design/FaultMonad_H.thy
M ../design/Fault_H.thy
M ../design/InterruptDecls_H.thy
M ../design/Interrupt_H.thy
M ../design/InvocationLabels_H.thy
M ../design/Invocations_H.thy
UU ../design/KI_Decls_H.thy
M ../design/KernelInitMonad_H.thy
M ../design/KernelInit_H.thy
M ../design/KernelStateData_H.thy
M ../design/Kernel_H.thy
M ../design/NotificationDecls_H.thy
M ../design/Notification_H.thy
MM ../design/ObjectInstances_H.thy
M ../design/Object_H.thy
M ../design/PSpaceFuns_H.thy
M ../design/PSpaceStorable_H.thy
M ../design/PSpaceStruct_H.thy
UU ../design/RetypeDecls_H.thy
UU ../design/Retype_H.thy
UU ../design/Structures_H.thy
M ../design/Syscall_H.thy
M ../design/TCBDecls_H.thy
M ../design/TCB_H.thy
M ../design/ThreadDecls_H.thy
M ../design/Thread_H.thy
UU ../design/Types_H.thy
M ../design/Untyped_H.thy
M ../design/VSpace_H.thy
MM ../design/m-skel/ARM/MachineTypes.thy
MM ../design/skel/ARM/ArchInterruptDecls_H.thy
M ../design/skel/ARM/ArchInterrupt_H.thy
M ../design/skel/ARM/ArchInvocationLabels_H.thy
M ../design/skel/ARM/ArchLabelFuns_H.thy
M ../design/skel/ARM/ArchObjInsts_H.thy
M ../design/skel/ARM/ArchRetypeDecls_H.thy
M ../design/skel/ARM/ArchRetype_H.thy
M ../design/skel/ARM/ArchStateData_H.thy
M ../design/skel/ARM/ArchStructures_H.thy
M ../design/skel/ARM/ArchTCB_H.thy
M ../design/skel/ARM/ArchThreadDecls_H.thy
M ../design/skel/ARM/ArchThread_H.thy
M ../design/skel/ARM/ArchTypes_H.thy
M ../design/skel/ARM/ArchVSpaceDecls_H.thy
M ../design/skel/ARM/ArchVSpace_H.thy
M ../design/skel/ARM/Arch_Structs_B.thy
M ../design/skel/ARM/Hardware_H.thy
M ../design/skel/ARM/RegisterSet_H.thy
M ../design/skel/ARM/State_H.thy
M ../design/skel/Delete_H.thy
M ../design/skel/Event_H.thy
M ../design/skel/FaultHandler_H.thy
M ../design/skel/FaultMonad_H.thy
M ../design/skel/Interrupt_H.thy
M ../design/skel/InvocationLabels_H.thy
M ../design/skel/Invocations_H.thy
M ../design/skel/KernelInit_H.thy
M ../design/skel/KernelStateData_H.thy
M ../design/skel/Kernel_H.thy
M ../design/skel/Notification_H.thy
MM ../design/skel/ObjectInstances_H.thy
M ../design/skel/PSpaceFuns_H.thy
M ../design/skel/PSpaceStorable_H.thy
M ../design/skel/Retype_H.thy
M ../design/skel/Structures_H.thy
M ../design/skel/TCBDecls_H.thy
M ../design/skel/TCB_H.thy
M ../design/skel/Thread_H.thy
MM ../design/skel/Types_H.thy
M ../design/skel/VSpace_H.thy
UU ../design/version
A .gitignore
A Makefile
A README.md
A SEL4.cabal
A Setup.hs
A check-newlines.sh
A configure
A doc/Makefile
A doc/README
A doc/disy.sty
A doc/figures/clientkernel.mp
A doc/figures/clientserver.mp
A doc/figures/derive.mp
A doc/figures/ipctransfer.mp
A doc/figures/modules.dot
A doc/figures/truncation.mp
A doc/haskell.tex
A doc/lambdaTeX.sty
A doc/mkfunctions.pl
A doc/mkmodulesgraph.pl
A doc/overview.tex
A doc/packages
A doc/pubs.bib
A doc/refman.tex
A include/gic.h
A include/mct.h
A include/mptimer.h
R ../../proof/crefine/CREGRESSION_S.ML -> include/sel4model.h
A mkhsboot.pl
A src/Data/BinaryTree.hs
A src/Data/Helpers.hs
A src/SEL4.lhs
A src/SEL4/API.lhs
A src/SEL4/API/Failures.lhs
A src/SEL4/API/Faults.lhs
AM src/SEL4/API/Invocation.lhs
A src/SEL4/API/Invocation/ARM.lhs
A src/SEL4/API/InvocationLabels.lhs
A src/SEL4/API/InvocationLabels/ARM.lhs
A src/SEL4/API/Syscall.lhs
AM src/SEL4/API/Types.lhs
AM src/SEL4/API/Types/ARM.lhs
A src/SEL4/API/Types/Universal.lhs
A src/SEL4/Config.lhs
A src/SEL4/Kernel.lhs
AM src/SEL4/Kernel/BootInfo.lhs
A src/SEL4/Kernel/CSpace.lhs
A src/SEL4/Kernel/FaultHandler.lhs
AM src/SEL4/Kernel/Init.lhs
A src/SEL4/Kernel/Thread.lhs
A src/SEL4/Kernel/Thread/ARM.lhs
A src/SEL4/Kernel/VSpace.lhs
AM src/SEL4/Kernel/VSpace/ARM.lhs
A src/SEL4/Machine.lhs
A src/SEL4/Machine/Hardware.lhs
A src/SEL4/Machine/Hardware/ARM.lhs
A src/SEL4/Machine/Hardware/ARM/Callbacks.hs
A src/SEL4/Machine/Hardware/ARM/Exynos4210.hs
A src/SEL4/Machine/Hardware/ARM/KZM.hs
A src/SEL4/Machine/Hardware/ARM/Sabre.hs
A src/SEL4/Machine/Hardware/GICInterface.hsc
A src/SEL4/Machine/Hardware/MCTInterface.hsc
A src/SEL4/Machine/Hardware/MPTimerInterface.hsc
A src/SEL4/Machine/RegisterSet.lhs
A src/SEL4/Machine/RegisterSet/ARM.lhs
A src/SEL4/Machine/Target.lhs
A src/SEL4/Model.lhs
A src/SEL4/Model/Failures.lhs
A src/SEL4/Model/PSpace.lhs
A src/SEL4/Model/Preemption.lhs
A src/SEL4/Model/StateData.lhs
A src/SEL4/Model/StateData/ARM.lhs
A src/SEL4/Model/Syscall.lhs
AM src/SEL4/Object.lhs
AM src/SEL4/Object/CNode.lhs
A src/SEL4/Object/Endpoint.lhs
AM src/SEL4/Object/Instances.lhs
A src/SEL4/Object/Instances/ARM.lhs
A src/SEL4/Object/Interrupt.lhs
A src/SEL4/Object/Interrupt/ARM.lhs
A src/SEL4/Object/Notification.lhs
AM src/SEL4/Object/ObjectType.lhs
AM src/SEL4/Object/ObjectType/ARM.lhs
AM src/SEL4/Object/Structures.lhs
AM src/SEL4/Object/Structures/ARM.lhs
A src/SEL4/Object/TCB.lhs
A src/SEL4/Object/TCB/ARM.lhs
AM src/SEL4/Object/Untyped.lhs
A src/Simulation/FFIBindings.hs
M ../machine/ARM/MachineOps.thy
MM ../machine/ARM/MachineTypes.thy
M ../machine/ARM/Platform.thy
A ../machine/MachineExports.thy
AM ../machine/MachineMonad.thy
A ../machine/Setup_Locale.thy
M ../take-grant/Confine_S.thy
M ../take-grant/System_S.thy
M ../tests.xml
M ../../sys-init/CreateObjects_SI.thy
M ../../sys-init/ExampleSpecIRQ_SI.thy
M ../../sys-init/ExampleSpec_SI.thy
M ../../sys-init/InitCSpace_SI.thy
M ../../sys-init/InitVSpace_SI.thy
M ../../sys-init/Proof_SI.thy
M ../../sys-init/RootTask_SI.thy
M ../../sys-init/SysInit_SI.thy
M ../../sys-init/WellFormed_SI.thy
M ../../tools/asmrefine/CommonOpsLemmas.thy
M ../../tools/asmrefine/GhostAssertions.thy
M ../../tools/asmrefine/GlobalsSwap.thy
M ../../tools/asmrefine/GraphRefine.thy
M ../../tools/asmrefine/ProveGraphRefine.thy
M ../../tools/asmrefine/SimplExport.thy
M ../../tools/asmrefine/TailrecPre.thy
M ../../tools/autocorres/AbstractArrays.thy
M ../../tools/autocorres/AutoCorres.thy
D ../../tools/autocorres/AutoCorresAttributes.thy
M ../../tools/autocorres/AutoCorresSimpset.thy
M ../../tools/autocorres/CCorresE.thy
M ../../tools/autocorres/CorresXF.thy
M ../../tools/autocorres/ExceptionRewrite.thy
M ../../tools/autocorres/HeapLift.thy
M ../../tools/autocorres/L1Defs.thy
M ../../tools/autocorres/L1Peephole.thy
M ../../tools/autocorres/L2Defs.thy
M ../../tools/autocorres/L2Opt.thy
M ../../tools/autocorres/L2Peephole.thy
M ../../tools/autocorres/L4VerifiedLinks.thy
M ../../tools/autocorres/LocalVarExtract.thy
M ../../tools/autocorres/MonadMono.thy
M ../../tools/autocorres/NonDetMonadEx.thy
M ../../tools/autocorres/Polish.thy
M ../../tools/autocorres/README.md
M ../../tools/autocorres/ROOT
M ../../tools/autocorres/SimplConv.thy
M ../../tools/autocorres/TestSEL4.thy
M ../../tools/autocorres/TypHeapSimple.thy
M ../../tools/autocorres/TypeStrengthen.thy
M ../../tools/autocorres/WordAbstract.thy
D ../../tools/autocorres/attributes.ML
M ../../tools/autocorres/autocorres.ML
M ../../tools/autocorres/autocorres_data.ML
M ../../tools/autocorres/autocorres_trace.ML
M ../../tools/autocorres/autocorres_util.ML
M ../../tools/autocorres/doc/quickstart/document/root.bib
M ../../tools/autocorres/doc/quickstart/minmax.c
M ../../tools/autocorres/doc/quickstart/mult_by_add.c
M ../../tools/autocorres/doc/quickstart/swap.c
M ../../tools/autocorres/exception_rewrite.ML
M ../../tools/autocorres/experiments/alloc-proof/Alloc_Simp.thy
M ../../tools/autocorres/function_info.ML
M ../../tools/autocorres/heap_lift.ML
M ../../tools/autocorres/heap_lift_base.ML
M ../../tools/autocorres/l2_opt.ML
M ../../tools/autocorres/local_var_extract.ML
M ../../tools/autocorres/monad_convert.ML
M ../../tools/autocorres/program_info.ML
M ../../tools/autocorres/simpl_conv.ML
D ../../tools/autocorres/statistics.ML
M ../../tools/autocorres/tests/examples/AC_Rename.thy
M ../../tools/autocorres/tests/examples/Factorial.thy
A ../../tools/autocorres/tests/examples/FunctionInfoDemo.thy
A ../../tools/autocorres/tests/examples/Incremental.thy
A ../../tools/autocorres/tests/examples/Suzuki.thy
M ../../tools/autocorres/tests/examples/TraceDemo.thy
A ../../tools/autocorres/tests/examples/function_info.c
M ../../tools/autocorres/tests/examples/rename.c
M ../../tools/autocorres/tests/examples/simple.c
A ../../tools/autocorres/tests/examples/suzuki.c
M ../../tools/autocorres/tests/examples/trace_demo.c
M ../../tools/autocorres/tests/examples/type_strengthen.c
M ../../tools/autocorres/tests/examples/type_strengthen_tricks.thy
A ../../tools/autocorres/tests/failing/dirty_frees.c
A ../../tools/autocorres/tests/failing/dirty_frees.thy
A ../../tools/autocorres/tests/failing/jira_ver_591.c
A ../../tools/autocorres/tests/failing/jira_ver_591.thy
M ../../tools/autocorres/tests/parse-tests/heap_lift_array.c
A ../../tools/autocorres/tests/parse-tests/l2_opt_invariant.c
A ../../tools/autocorres/tests/proof-tests/Test_Spec_Translation.thy
M ../../tools/autocorres/tests/proof-tests/WordAbsFnCall.thy
R ../../tools/autocorres/tests/parse-tests/badnames.c -> ../../tools/autocorres/tests/proof-tests/badnames.c
A ../../tools/autocorres/tests/proof-tests/badnames.thy
M ../../tools/autocorres/tests/proof-tests/prototyped_functions.c
M ../../tools/autocorres/tests/proof-tests/prototyped_functions.thy
M ../../tools/autocorres/tests/proof-tests/skip_heap_abs.thy
M ../../tools/autocorres/tests/proof-tests/struct.thy
M ../../tools/autocorres/tests/proof-tests/struct2.thy
M ../../tools/autocorres/tests/proof-tests/struct3.thy
A ../../tools/autocorres/tests/proof-tests/test_spec_translation.c
M ../../tools/autocorres/tests/proof-tests/word_abs_fn_call.c
M ../../tools/autocorres/tools/release.py
M ../../tools/autocorres/tools/release_files/ChangeLog
M ../../tools/autocorres/tools/release_files/LIB_FILES
M ../../tools/autocorres/tools/release_files/README
M ../../tools/autocorres/tools/stats/schorr_waite.c
M ../../tools/autocorres/type_strengthen.ML
M ../../tools/autocorres/utils.ML
M ../../tools/autocorres/word_abstract.ML
A ../../tools/c-parser/Absyn-Serial.ML
M ../../tools/c-parser/CProof.thy
M ../../tools/c-parser/CTranslation.thy
A ../../tools/c-parser/Listsort.ML
M ../../tools/c-parser/MANIFEST
M ../../tools/c-parser/MString.ML
M ../../tools/c-parser/PackedTypes.thy
M ../../tools/c-parser/README
M ../../tools/c-parser/Simpl/ROOT
M ../../tools/c-parser/calculate_state.ML
M ../../tools/c-parser/mkrelease
M ../../tools/c-parser/program_analysis.ML
A ../../tools/c-parser/standalone-parser/GetOpt.sig
A ../../tools/c-parser/standalone-parser/GetOpt.sml
M ../../tools/c-parser/standalone-parser/c-parser.mlb
M ../../tools/c-parser/standalone-parser/main.sml
A ../../tools/c-parser/standalone-parser/make_munge.sh
M ../../tools/c-parser/termstypes.ML
M ../../tools/c-parser/testfiles/errors/addrlocal.thy
A ../../tools/c-parser/testfiles/extern_builtin.c
A ../../tools/c-parser/testfiles/extern_builtin.thy
M ../../tools/c-parser/testfiles/factorial.thy
M ../../tools/c-parser/testfiles/includes/test_locality.thy
M ../../tools/c-parser/testfiles/isa2014.c
M ../../tools/c-parser/testfiles/isa2014.thy
M ../../tools/c-parser/testfiles/jiraver310.c
M ../../tools/c-parser/testfiles/jiraver310.thy
M ../../tools/c-parser/testfiles/jiraver337.c
M ../../tools/c-parser/testfiles/jiraver337.thy
M ../../tools/c-parser/testfiles/jiraver422.c
M ../../tools/c-parser/testfiles/jiraver422.thy
M ../../tools/c-parser/testfiles/jiraver426.c
M ../../tools/c-parser/testfiles/jiraver426.thy
M ../../tools/c-parser/testfiles/jiraver429.c
M ../../tools/c-parser/testfiles/jiraver429.thy
M ../../tools/c-parser/testfiles/jiraver432.c
M ../../tools/c-parser/testfiles/jiraver432.thy
M ../../tools/c-parser/testfiles/jiraver434.c
M ../../tools/c-parser/testfiles/jiraver434.thy
M ../../tools/c-parser/testfiles/jiraver439.c
M ../../tools/c-parser/testfiles/jiraver439.thy
M ../../tools/c-parser/testfiles/jiraver440.c
M ../../tools/c-parser/testfiles/jiraver440.thy
M ../../tools/c-parser/testfiles/jiraver443.c
M ../../tools/c-parser/testfiles/jiraver443.thy
M ../../tools/c-parser/testfiles/jiraver443a.c
M ../../tools/c-parser/testfiles/jiraver443a.thy
M ../../tools/c-parser/testfiles/jiraver456.c
M ../../tools/c-parser/testfiles/jiraver456.thy
M ../../tools/c-parser/testfiles/jiraver464.c
M ../../tools/c-parser/testfiles/jiraver464.thy
M ../../tools/c-parser/testfiles/jiraver473.c
M ../../tools/c-parser/testfiles/jiraver473.thy
A ../../tools/c-parser/testfiles/jiraver550.c
A ../../tools/c-parser/testfiles/jiraver550.thy
M ../../tools/c-parser/testfiles/longlong.thy
M ../../tools/c-parser/testfiles/many_local_vars.thy
M ../../tools/c-parser/testfiles/ptr_modifies.thy
M ../../tools/c-parser/umm_heap/ArrayAssertion.thy
M ../../tools/c-parser/umm_heap/CTypes.thy
M ../../tools/c-parser/umm_heap/CTypesBase.thy
M ../../tools/c-parser/umm_heap/CTypesDefs.thy
M ../../tools/c-parser/umm_heap/TypHeap.thy
M ../../tools/c-parser/umm_heap/Vanilla32.thy
M ../../tools/haskell-translator/lhs_pars.py
M ../../tools/haskell-translator/make_spec.sh
M ../../tools/haskell-translator/pars_skl.py
M ../../tools/tests.xml
?? ../haskell2/

View File

@ -74,3 +74,4 @@ There are three ARM-specific object types: virtual pages, page tables, and page
> isFrameType SectionObject = True
> isFrameType SuperSectionObject = True
> isFrameType _ = False

View File

@ -80,7 +80,10 @@ The second argument specifies the size of the object, for the types for which th
The value of this argument is the base 2 logarithm of the actual required size. The unit depends on the object type: one byte for untyped memory objects, the architecture's minimum page size for data frames, and one capability slot for CNodes.
> let userObjSize = fromIntegral userObjSizeW
> rangeCheck userObjSize 0 $ finiteBitSize nullPointer - 3 -- max untyped size is 2^30
max untyped size is $2^30$
> rangeCheck userObjSize 0 $ finiteBitSize nullPointer - 3
The kernel does not allow creation of CNodes containing only one entry; this is done to avoid non-terminating loops in capability lookup. Note that it is possible for a single entry CNode to translate bits using its guard; this is not allowed, however, to avoid having to check for it during capability lookups.

View File

@ -174,7 +174,7 @@ lemma create_irq_cap_sep:
lemma create_irq_caps_sep_helper:
"\<lbrace>\<guillemotleft>((\<And>* cptr \<in> set (take (card (used_irqs spec)) free_cptrs).
((si_cnode_id, unat cptr) \<mapsto>c NullCap)) \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* si_irq_nodes spec \<and>* R) and
K (well_formed spec \<and>
list_all (\<lambda>n. n < 2 ^ si_cnode_size) free_cptrs \<and>
@ -184,7 +184,7 @@ lemma create_irq_caps_sep_helper:
\<lbrace>\<lambda>rv s. \<exists>(t'::32 word \<Rightarrow> 32 word option).
\<guillemotleft>(irqs_empty spec t' (used_irqs spec) \<and>*
si_irq_caps_at (fst rv) spec (used_irqs spec) \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R)
and K ((map_of (zip (used_irq_list spec) free_cptrs), drop (card (used_irqs spec)) free_cptrs) = rv \<and>
inj_on t' (used_irq_nodes spec) \<and>
@ -205,7 +205,7 @@ lemma create_irq_caps_sep_helper:
|` used_irq_nodes spec) irq \<and>*
si_irq_cap_at (map_of (zip (used_irq_list spec) free_cptrs)) spec irq" and
I1 = "si_objects" and
R1 = "si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
R1 = "si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
R" in hoare_chain [OF mapM_x_set_sep])
apply (metis distinct_zipI2)
apply (clarsimp split:prod.splits)
@ -280,7 +280,7 @@ lemma object_empty_cong:
lemma si_cap_at_cong:
"t obj_id = t' obj_id
\<Longrightarrow> si_cap_at t si_caps spec obj_id = si_cap_at t' si_caps spec obj_id"
\<Longrightarrow> si_cap_at t si_caps spec dev obj_id = si_cap_at t' si_caps spec dev obj_id"
by (clarsimp simp: si_cap_at_def)
lemma irq_empty_map_add:
@ -301,7 +301,7 @@ lemma object_empty_map_add:
lemma si_caps_at_map_add:
"\<lbrakk>dom t = obj_ids; map_disj t t'\<rbrakk>
\<Longrightarrow> si_caps_at t si_caps spec obj_ids = si_caps_at (t++t') si_caps spec obj_ids"
\<Longrightarrow> si_caps_at t si_caps spec dev obj_ids = si_caps_at (t++t') si_caps spec dev obj_ids"
apply (clarsimp simp: si_caps_at_def)
apply (rule sep.setprod.cong, simp)
apply (subst si_cap_at_cong [where t'="t++t'" and t=t], simp_all)
@ -375,7 +375,7 @@ lemma si_objects_extra_caps'_split:
lemma create_irq_caps_sep:
"\<lbrace>\<lambda>s. \<exists>t_real.
\<guillemotleft>(objects_empty spec t_real {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t_real orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t_real orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>*
si_objects_extra_caps' {obj_id. real_object_at obj_id spec} free_cptrs_orig untyped_cptrs \<and>*
si_irq_nodes spec \<and>* R) and
@ -391,7 +391,7 @@ lemma create_irq_caps_sep:
\<lbrace>\<lambda>rv s. \<exists>(t::32 word \<Rightarrow> 32 word option).
\<guillemotleft>(objects_empty spec t {obj_id. real_object_at obj_id spec} \<and>*
irqs_empty spec t (used_irqs spec) \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_irq_caps_at (fst rv) spec (used_irqs spec) \<and>*
si_objects \<and>*
si_objects_extra_caps' (dom (cdl_objects spec)) free_cptrs_orig untyped_cptrs \<and>*

View File

@ -27,7 +27,7 @@ lemma seL4_Untyped_Retype_cdt_inc':
\<lbrace>\<lambda>s.
\<guillemotleft> si_tcb_id \<mapsto>f Tcb tcb \<and>*
si_cnode_id \<mapsto>f CNode (empty_cnode si_cnode_size) \<and>*
(si_cnode_id, unat untyped_cptr) \<mapsto>c UntypedCap obj_range free_range \<and>*
(si_cnode_id, unat untyped_cptr) \<mapsto>c UntypedCap dev obj_range free_range \<and>*
(si_tcb_id, tcb_cspace_slot) \<mapsto>c si_cnode_cap \<and>*
(si_cnode_id, unat seL4_CapInitThreadCNode) \<mapsto>c si_cnode_cap \<and>*
sep_true \<guillemotright> s \<and>
@ -46,7 +46,7 @@ lemma seL4_Untyped_Retype_has_children_wp:
\<lbrace>\<lambda>s.
\<guillemotleft>si_tcb_id \<mapsto>f Tcb tcb \<and>*
si_cnode_id \<mapsto>f CNode (empty_cnode si_cnode_size) \<and>*
(si_cnode_id, unat untyped_cptr) \<mapsto>c UntypedCap obj_range free_range \<and>*
(si_cnode_id, unat untyped_cptr) \<mapsto>c UntypedCap dev obj_range free_range \<and>*
(si_tcb_id, tcb_cspace_slot) \<mapsto>c si_cnode_cap \<and>*
(si_cnode_id, unat seL4_CapInitThreadCNode) \<mapsto>c si_cnode_cap \<and>*
sep_true \<guillemotright> s \<and>
@ -69,7 +69,7 @@ lemma seL4_Untyped_Retype_list_all_has_children_index_wp:
\<lbrace>\<lambda>s.
\<guillemotleft>si_tcb_id \<mapsto>f Tcb tcb \<and>*
si_cnode_id \<mapsto>f CNode (empty_cnode si_cnode_size) \<and>*
(si_cnode_id, unat untyped_cptr) \<mapsto>c UntypedCap obj_range free_range \<and>*
(si_cnode_id, unat untyped_cptr) \<mapsto>c UntypedCap dev obj_range free_range \<and>*
(si_tcb_id, tcb_cspace_slot) \<mapsto>c si_cnode_cap \<and>*
(si_cnode_id, unat seL4_CapInitThreadCNode) \<mapsto>c si_cnode_cap \<and>*
sep_true\<guillemotright> s \<and>
@ -94,7 +94,7 @@ lemma seL4_Untyped_Retype_sep_cdt_inc:
\<guillemotleft>si_tcb_id \<mapsto>f (Tcb tcb)
\<and>* (si_tcb_id, tcb_pending_op_slot) \<mapsto>c RunningCap
\<and>* (cap_object si_cnode_cap \<mapsto>f CNode (empty_cnode si_cnode_size))
\<and>* (si_cnode_id, unat untyped_cptr) \<mapsto>c UntypedCap obj_range free_range
\<and>* (si_cnode_id, unat untyped_cptr) \<mapsto>c UntypedCap dev obj_range free_range
\<and>* (si_cnode_id, unat ncptr ) \<mapsto>c NullCap
\<and>* (\<And>* ptr\<in>tot_free_range. ptr \<mapsto>o Untyped)
\<and>* (si_tcb_id, tcb_cspace_slot) \<mapsto>c si_cnode_cap
@ -108,11 +108,11 @@ lemma seL4_Untyped_Retype_sep_cdt_inc:
\<lbrace>\<lambda>r s. (\<not> r \<longrightarrow> (\<exists>oid free_range'. (\<guillemotleft>
(si_tcb_id, tcb_pending_op_slot) \<mapsto>c RunningCap
\<and>* si_tcb_id \<mapsto>f (Tcb tcb)
\<and>* (si_cnode_id, unat ncptr) \<mapsto>c (default_cap nt {oid} (unat ts))
\<and>* (si_cnode_id, unat ncptr) \<mapsto>c (default_cap nt {oid} (unat ts) dev)
\<and>* oid \<mapsto>o obj
\<and>* (cap_object si_cnode_cap \<mapsto>f CNode (empty_cnode si_cnode_size))
\<and>* (\<And>* ptr\<in>tot_free_range - {oid}. ptr \<mapsto>o Untyped)
\<and>* (si_cnode_id, unat untyped_cptr) \<mapsto>c UntypedCap obj_range free_range'
\<and>* (si_cnode_id, unat untyped_cptr) \<mapsto>c UntypedCap dev obj_range free_range'
\<and>* (si_tcb_id, tcb_cspace_slot) \<mapsto>c si_cnode_cap
\<and>* (cap_object si_cnode_cap, unat seL4_CapInitThreadCNode) \<mapsto>c si_cnode_cap
\<and>* R \<guillemotright> s ) \<and> free_range' \<subseteq> free_range - {oid} \<and> oid \<in> free_range)
@ -121,7 +121,7 @@ lemma seL4_Untyped_Retype_sep_cdt_inc:
(si_tcb_id, tcb_pending_op_slot) \<mapsto>c RunningCap
\<and>* si_tcb_id \<mapsto>f (Tcb tcb)
\<and>* (cap_object si_cnode_cap \<mapsto>f CNode (empty_cnode si_cnode_size))
\<and>* (si_cnode_id,unat untyped_cptr) \<mapsto>c UntypedCap obj_range free_range
\<and>* (si_cnode_id,unat untyped_cptr) \<mapsto>c UntypedCap dev obj_range free_range
\<and>* (si_cnode_id, unat ncptr) \<mapsto>c NullCap
\<and>* (\<And>* ptr\<in>tot_free_range. ptr \<mapsto>o Untyped)
\<and>* (si_tcb_id, tcb_cspace_slot) \<mapsto>c si_cnode_cap
@ -151,6 +151,7 @@ lemma seL4_Untyped_Retype_sep_cdt_inc:
apply clarsimp
apply sep_solve
apply clarsimp
apply sep_solve
apply clarsimp
done
@ -175,7 +176,7 @@ lemma retype_untyped_wp:
sz < 2 ^ word_bits;
type \<noteq> UntypedType\<rbrakk>
\<Longrightarrow>
\<lbrace>\<lambda>s. \<guillemotleft>(si_cnode_id, unat untyped_cptr) \<mapsto>c UntypedCap cover_ids available_ids \<and>*
\<lbrace>\<lambda>s. \<guillemotleft>(si_cnode_id, unat untyped_cptr) \<mapsto>c UntypedCap dev cover_ids available_ids \<and>*
(\<And>* obj_id \<in> all_available_ids. (obj_id \<mapsto>o Untyped)) \<and>*
(si_cnode_id, unat free_cptr) \<mapsto>c NullCap \<and>*
si_tcb_id \<mapsto>f root_tcb \<and>*
@ -188,10 +189,10 @@ lemma retype_untyped_wp:
retype_untyped free_cptr untyped_cptr type sz
\<lbrace>\<lambda>rv s. (\<not>rv \<longrightarrow> (\<exists>new_id available_ids'.
new_id \<in> available_ids \<and> available_ids' \<subseteq> available_ids - {new_id} \<and>
\<guillemotleft>(si_cnode_id, unat untyped_cptr) \<mapsto>c UntypedCap cover_ids available_ids' \<and>*
\<guillemotleft>(si_cnode_id, unat untyped_cptr) \<mapsto>c UntypedCap dev cover_ids available_ids' \<and>*
(\<And>* obj_id \<in> all_available_ids - {new_id}. (obj_id \<mapsto>o Untyped)) \<and>*
new_id \<mapsto>o new_object \<and>*
(si_cnode_id, unat free_cptr) \<mapsto>c default_cap type {new_id} sz \<and>*
(si_cnode_id, unat free_cptr) \<mapsto>c default_cap type {new_id} sz dev \<and>*
si_tcb_id \<mapsto>f root_tcb \<and>*
(si_tcb_id, tcb_pending_op_slot) \<mapsto>c RunningCap \<and>*
(si_tcb_id, tcb_cspace_slot) \<mapsto>c si_cspace_cap \<and>*
@ -200,7 +201,7 @@ lemma retype_untyped_wp:
has_children (si_cnode_id,unat untyped_cptr) (kernel_state s) \<and>
list_all (\<lambda>index. has_children (si_cnode_id, untyped_slots ! index) (kernel_state s)) indices) \<and>
(rv \<longrightarrow>
\<guillemotleft>(si_cnode_id, unat untyped_cptr) \<mapsto>c UntypedCap cover_ids available_ids \<and>*
\<guillemotleft>(si_cnode_id, unat untyped_cptr) \<mapsto>c UntypedCap dev cover_ids available_ids \<and>*
(\<And>* obj_id \<in> all_available_ids. (obj_id \<mapsto>o Untyped)) \<and>*
(si_cnode_id, unat free_cptr) \<mapsto>c NullCap \<and>*
si_tcb_id \<mapsto>f root_tcb \<and>*
@ -236,14 +237,14 @@ lemma retype_untyped_wp:
done
lemma retype_untyped_wp_success:
"\<lbrakk>default_object type sz minBound = Some new_object;
"\<lbrakk>default_object type sz minBound= Some new_object;
available_ids \<subseteq> all_available_ids;
free_cptr < 2 ^ si_cnode_size;
untyped_cptr < 2 ^ si_cnode_size;
sz < 2 ^ word_bits;
type \<noteq> UntypedType\<rbrakk>
\<Longrightarrow>
\<lbrace>\<lambda>s. \<guillemotleft>(si_cnode_id, unat untyped_cptr) \<mapsto>c UntypedCap cover_ids available_ids \<and>*
\<lbrace>\<lambda>s. \<guillemotleft>(si_cnode_id, unat untyped_cptr) \<mapsto>c UntypedCap dev cover_ids available_ids \<and>*
(\<And>* obj_id \<in> all_available_ids. (obj_id \<mapsto>o Untyped)) \<and>*
(si_cnode_id, unat free_cptr) \<mapsto>c NullCap \<and>*
si_tcb_id \<mapsto>f root_tcb \<and>*
@ -256,10 +257,10 @@ lemma retype_untyped_wp_success:
retype_untyped free_cptr untyped_cptr type sz
\<lbrace>\<lambda>rv s. \<not>rv \<longrightarrow> (\<exists>new_id available_ids'.
new_id \<in> available_ids \<and> available_ids' \<subseteq> available_ids - {new_id} \<and>
\<guillemotleft>(si_cnode_id, unat untyped_cptr) \<mapsto>c UntypedCap cover_ids available_ids' \<and>*
\<guillemotleft>(si_cnode_id, unat untyped_cptr) \<mapsto>c UntypedCap dev cover_ids available_ids' \<and>*
(\<And>* obj_id \<in> all_available_ids - {new_id}. (obj_id \<mapsto>o Untyped)) \<and>*
new_id \<mapsto>o new_object \<and>*
(si_cnode_id, unat free_cptr) \<mapsto>c default_cap type {new_id} sz \<and>*
(si_cnode_id, unat free_cptr) \<mapsto>c default_cap type {new_id} sz dev \<and>*
si_tcb_id \<mapsto>f root_tcb \<and>*
(si_tcb_id, tcb_pending_op_slot) \<mapsto>c RunningCap \<and>*
(si_tcb_id, tcb_cspace_slot) \<mapsto>c si_cspace_cap \<and>*
@ -279,7 +280,7 @@ lemma retype_untyped_wp_fail:
sz < 2 ^ word_bits;
type \<noteq> UntypedType\<rbrakk>
\<Longrightarrow>
\<lbrace>\<lambda>s. \<guillemotleft>(si_cnode_id, unat untyped_cptr) \<mapsto>c UntypedCap cover_ids available_ids \<and>*
\<lbrace>\<lambda>s. \<guillemotleft>(si_cnode_id, unat untyped_cptr) \<mapsto>c UntypedCap dev cover_ids available_ids \<and>*
(\<And>* obj_id \<in> all_available_ids. (obj_id \<mapsto>o Untyped)) \<and>*
(si_cnode_id, unat free_cptr) \<mapsto>c NullCap \<and>*
si_tcb_id \<mapsto>f root_tcb \<and>*
@ -291,7 +292,7 @@ lemma retype_untyped_wp_fail:
list_all (\<lambda>index. has_children (si_cnode_id, untyped_slots ! index) (kernel_state s)) indices\<rbrace>
retype_untyped free_cptr untyped_cptr type sz
\<lbrace>\<lambda>rv s. rv \<longrightarrow>
\<guillemotleft>(si_cnode_id, unat untyped_cptr) \<mapsto>c UntypedCap cover_ids available_ids \<and>*
\<guillemotleft>(si_cnode_id, unat untyped_cptr) \<mapsto>c UntypedCap dev cover_ids available_ids \<and>*
(\<And>* obj_id \<in> all_available_ids. (obj_id \<mapsto>o Untyped)) \<and>*
(si_cnode_id, unat free_cptr) \<mapsto>c NullCap \<and>*
si_tcb_id \<mapsto>f root_tcb \<and>*
@ -321,7 +322,7 @@ lemma retype_untyped_bij_success:
dom t = used_spec_ids \<and>
obj_id \<notin> used_spec_ids \<and>
available_ids \<subseteq> all_available_ids \<and>
\<guillemotleft>(si_cnode_id, untyped_slot) \<mapsto>c UntypedCap cover_ids available_ids \<and>*
\<guillemotleft>(si_cnode_id, untyped_slot) \<mapsto>c UntypedCap dev cover_ids available_ids \<and>*
(\<And>* obj_id \<in> all_available_ids. (obj_id \<mapsto>o Untyped)) \<and>*
(si_cnode_id, free_slot) \<mapsto>c NullCap \<and>*
si_tcb_id \<mapsto>f root_tcb \<and>*
@ -342,11 +343,11 @@ lemma retype_untyped_bij_success:
obj_id \<notin> used_spec_ids \<and>
new_id \<in> available_ids \<and>
available_ids' \<subseteq> available_ids - {new_id} \<and>
\<guillemotleft>(si_cnode_id, untyped_slot) \<mapsto>c UntypedCap cover_ids available_ids' \<and>*
\<guillemotleft>(si_cnode_id, untyped_slot) \<mapsto>c UntypedCap dev cover_ids available_ids' \<and>*
(\<And>* obj_id \<in> all_available_ids - {new_id}. (obj_id \<mapsto>o Untyped)) \<and>*
object_empty spec (t(obj_id \<mapsto> new_id)) obj_id \<and>*
si_cap_at (t(obj_id \<mapsto> new_id)) (si_caps(obj_id \<mapsto> free_cptr))
spec obj_id \<and>*
spec dev obj_id \<and>*
si_tcb_id \<mapsto>f root_tcb \<and>*
(si_tcb_id, tcb_pending_op_slot) \<mapsto>c RunningCap \<and>*
(si_tcb_id, tcb_cspace_slot) \<mapsto>c si_cspace_cap \<and>*
@ -390,15 +391,15 @@ lemma retype_untyped_bij_success:
done
lemma si_cap_at_update:
"\<lbrakk>(si_cap_at t si_caps spec obj_id) s; obj_id \<noteq> obj_id'\<rbrakk>
\<Longrightarrow> (si_cap_at t (si_caps(obj_id' \<mapsto> cap_ptr)) spec obj_id) s"
"\<lbrakk>(si_cap_at t si_caps spec dev obj_id) s; obj_id \<noteq> obj_id'\<rbrakk>
\<Longrightarrow> (si_cap_at t (si_caps(obj_id' \<mapsto> cap_ptr)) spec dev obj_id) s"
by (clarsimp simp: si_cap_at_def)
lemma map_si_cap_at_update_old:
"\<lbrakk>distinct obj_ids; obj_id_index < length obj_ids;
obj_id = obj_ids ! obj_id_index;
(\<And>* map (si_cap_at t si_caps spec) (take obj_id_index obj_ids)) s\<rbrakk>
\<Longrightarrow> (\<And>* map (si_cap_at t (si_caps(obj_id \<mapsto> cap_ptr)) spec)
(\<And>* map (si_cap_at t si_caps spec dev) (take obj_id_index obj_ids)) s\<rbrakk>
\<Longrightarrow> (\<And>* map (si_cap_at t (si_caps(obj_id \<mapsto> cap_ptr)) spec dev)
(take obj_id_index obj_ids)) s"
apply (erule sep_list_conj_map_impl [rotated])
apply (erule si_cap_at_update)
@ -407,10 +408,10 @@ lemma map_si_cap_at_update_old:
done
lemma map_si_cap_at_update': (* Need better tactics. *)
"\<lbrakk>(\<And>* map (si_cap_at t si_caps spec) (take obj_id_index obj_ids) \<and>* R) s;
"\<lbrakk>(\<And>* map (si_cap_at t si_caps spec dev) (take obj_id_index obj_ids) \<and>* R) s;
distinct obj_ids; obj_id_index < length obj_ids;
obj_id = obj_ids ! obj_id_index\<rbrakk>
\<Longrightarrow> (\<And>* map (si_cap_at t (si_caps(obj_id \<mapsto> cap_ptr)) spec)
\<Longrightarrow> (\<And>* map (si_cap_at t (si_caps(obj_id \<mapsto> cap_ptr)) spec dev)
(take obj_id_index obj_ids) \<and>* R) s"
by (drule sep_conj_impl, erule map_si_cap_at_update_old, assumption+)
@ -457,7 +458,7 @@ lemma all_available_ids_updates:
done
lemma untyped_cap_eq:
"is_untyped_cap cap \<Longrightarrow> UntypedCap (cap_objects cap) (cap_free_ids cap) = cap"
"is_untyped_cap cap \<Longrightarrow> UntypedCap (is_device_cap cap) (cap_objects cap) (cap_free_ids cap) = cap"
by (clarsimp simp: cap_type_def cap_free_ids_def split: cdl_cap.splits)
lemma retype_untyped_loop_inv_pre:
@ -470,7 +471,7 @@ lemma retype_untyped_loop_inv_pre:
\<And>* map (\<lambda>slot. (si_cnode_id, slot) \<mapsto>c NullCap) (drop obj_id_index free_slots) \<and>*
\<And>* map (object_empty spec t) (take obj_id_index obj_ids) \<and>*
(\<And>* obj_id\<in>all_available_ids. obj_id \<mapsto>o Untyped) \<and>*
\<And>* map (si_cap_at t si_caps spec) (take obj_id_index obj_ids) \<and>*
\<And>* map (si_cap_at t si_caps spec dev) (take obj_id_index obj_ids) \<and>*
si_tcb_id \<mapsto>f root_tcb \<and>*
(si_tcb_id, tcb_pending_op_slot) \<mapsto>c RunningCap \<and>*
(si_tcb_id, tcb_cspace_slot) \<mapsto>c si_cspace_cap \<and>*
@ -482,7 +483,7 @@ lemma retype_untyped_loop_inv_pre:
[index\<leftarrow>[0..<length untyped_caps] . \<not> is_full_untyped_cap (untyped_caps ! index)]
\<Longrightarrow>
\<guillemotleft>(si_cnode_id, untyped_slots ! untyped_index) \<mapsto>c
UntypedCap (cap_objects (untyped_caps ! untyped_index))
UntypedCap (is_device_cap (untyped_caps ! untyped_index)) (cap_objects (untyped_caps ! untyped_index))
(cap_free_ids (untyped_caps ! untyped_index)) \<and>*
(\<And>* obj_id\<in>all_available_ids. obj_id \<mapsto>o Untyped) \<and>*
(si_cnode_id, free_slots ! obj_id_index) \<mapsto>c NullCap \<and>*
@ -497,7 +498,7 @@ lemma retype_untyped_loop_inv_pre:
(drop (Suc untyped_index) (zip untyped_slots untyped_caps)) \<and>*
\<And>* map (\<lambda>slot. (si_cnode_id, slot) \<mapsto>c NullCap) (drop (Suc obj_id_index) free_slots) \<and>*
\<And>* map (object_empty spec t) (take obj_id_index obj_ids) \<and>*
\<And>* map (si_cap_at t si_caps spec) (take obj_id_index obj_ids) \<and>*
\<And>* map (si_cap_at t si_caps spec dev) (take obj_id_index obj_ids) \<and>*
R\<guillemotright> s \<and>
(\<not> has_children (si_cnode_id, untyped_slots ! untyped_index) (kernel_state s) \<longrightarrow>
cap_objects (untyped_caps ! untyped_index) =
@ -533,7 +534,7 @@ lemma retype_untyped_loop_inv_post:
object_empty spec t obj_id \<and>*
si_cap_at t
(si_caps(obj_id \<mapsto> of_nat (free_slots ! obj_id_index)))
spec obj_id \<and>*
spec dev obj_id \<and>*
si_tcb_id \<mapsto>f root_tcb \<and>*
(si_tcb_id, tcb_pending_op_slot) \<mapsto>c RunningCap \<and>*
(si_tcb_id, tcb_cspace_slot) \<mapsto>c si_cspace_cap \<and>*
@ -546,7 +547,7 @@ lemma retype_untyped_loop_inv_post:
\<And>* map (\<lambda>slot. (si_cnode_id, slot) \<mapsto>c NullCap)
(drop (Suc obj_id_index) free_slots) \<and>*
\<And>* map (object_empty spec t) (take obj_id_index obj_ids) \<and>*
\<And>* map (si_cap_at t si_caps spec) (take obj_id_index obj_ids) \<and>*
\<And>* map (si_cap_at t si_caps spec dev) (take obj_id_index obj_ids) \<and>*
R\<guillemotright> s
\<Longrightarrow>
\<guillemotleft>\<And>* map (\<lambda>(slot, cap). (si_cnode_id, slot) \<mapsto>c cap)
@ -558,7 +559,7 @@ lemma retype_untyped_loop_inv_post:
(\<And>* obj_id\<in>all_available_ids - {new_id}. obj_id \<mapsto>o Untyped) \<and>*
\<And>* map (si_cap_at t
(si_caps(obj_id \<mapsto> of_nat (free_slots ! obj_id_index)))
spec) (take (Suc obj_id_index) obj_ids) \<and>*
spec dev) (take (Suc obj_id_index) obj_ids) \<and>*
si_tcb_id \<mapsto>f root_tcb \<and>*
(si_tcb_id, tcb_pending_op_slot) \<mapsto>c RunningCap \<and>*
(si_tcb_id, tcb_cspace_slot) \<mapsto>c si_cspace_cap \<and>*
@ -611,8 +612,8 @@ lemma map_object_empty_update:
lemma map_si_cap_at_update:
"\<lbrakk>i < length obj_ids; distinct obj_ids\<rbrakk>
\<Longrightarrow> \<And>* map (si_cap_at (t(obj_ids ! i \<mapsto> obj_id')) si_caps spec) (take i obj_ids)
= \<And>* map (si_cap_at t si_caps spec) (take i obj_ids)"
\<Longrightarrow> \<And>* map (si_cap_at (t(obj_ids ! i \<mapsto> obj_id')) si_caps spec dev) (take i obj_ids)
= \<And>* map (si_cap_at t si_caps spec dev) (take i obj_ids)"
apply (erule (1) sep_map_conj_f_update)
apply (clarsimp simp: si_cap_at_def)
done
@ -631,6 +632,17 @@ lemma ran_insert_new:
"\<lbrakk>a \<notin> dom m; b \<notin> ran m\<rbrakk> \<Longrightarrow> ran (m(a \<mapsto> b)) = insert b (ran m)"
by auto
lemma remove_free_ids_is_device[simp]:
"is_untyped_cap a \<Longrightarrow> is_device_cap (remove_free_ids a b) = is_device_cap a"
by (simp add:remove_free_ids_def split:cdl_cap.splits)
lemma list_all_conj:
"(list_all P xs \<and> list_all Q xs) = list_all (P and Q) xs"
by (auto simp:pred_list_def)
lemmas list_all_conjI = list_all_conj[THEN iffD1,unfolded pred_conj_def,OF conjI]
lemma retype_untyped_loop_inv_success:
"\<lbrakk>well_formed spec;
distinct obj_ids;
@ -646,13 +658,13 @@ lemma retype_untyped_loop_inv_success:
type = object_type object;
object_size = of_nat (object_at_pointer_size_bits spec obj_id)\<rbrakk>
\<Longrightarrow>
\<lbrace>\<lambda>s. \<exists>untyped_caps t all_available_ids.
\<lbrace>\<lambda>s. \<exists>untyped_caps t all_available_ids.
\<guillemotleft>\<And>* map (\<lambda>(slot, cap). (si_cnode_id, slot) \<mapsto>c cap)
(zip untyped_slots untyped_caps) \<and>*
\<And>* map (\<lambda>slot. (si_cnode_id, slot) \<mapsto>c NullCap) (drop obj_id_index free_slots) \<and>*
\<And>* map (object_empty spec t) (take obj_id_index obj_ids) \<and>*
(\<And>* obj_id\<in>all_available_ids. obj_id \<mapsto>o Untyped) \<and>*
\<And>* map (si_cap_at t si_caps spec) (take obj_id_index obj_ids) \<and>*
\<And>* map (si_cap_at t si_caps spec dev) (take obj_id_index obj_ids) \<and>*
si_tcb_id \<mapsto>f root_tcb \<and>*
(si_tcb_id, tcb_pending_op_slot) \<mapsto>c RunningCap \<and>*
@ -661,6 +673,7 @@ lemma retype_untyped_loop_inv_success:
(si_cnode_id, unat seL4_CapInitThreadCNode) \<mapsto>c si_cnode_cap \<and>* R\<guillemotright> s \<and>
length untyped_slots = length untyped_caps \<and>
list_all is_untyped_cap untyped_caps \<and>
list_all (\<lambda>c. is_device_cap c = dev) untyped_caps \<and>
list_all well_formed_untyped_cap untyped_caps \<and>
list_all (\<lambda>n. n < 2 ^ si_cnode_size) untyped_slots \<and>
list_all (\<lambda>n. n < 2 ^ si_cnode_size) free_slots \<and>
@ -680,7 +693,7 @@ lemma retype_untyped_loop_inv_success:
\<And>* map (object_empty spec t)
(take (Suc obj_id_index) obj_ids) \<and>*
(\<And>* obj_id\<in>all_available_ids. obj_id \<mapsto>o Untyped) \<and>*
\<And>* map (si_cap_at t (si_caps(obj_id \<mapsto> free_cptr)) spec)
\<And>* map (si_cap_at t (si_caps(obj_id \<mapsto> free_cptr)) spec dev)
(take (Suc obj_id_index) obj_ids) \<and>*
si_tcb_id \<mapsto>f root_tcb \<and>*
(si_tcb_id, tcb_pending_op_slot) \<mapsto>c RunningCap \<and>*
@ -689,6 +702,7 @@ lemma retype_untyped_loop_inv_success:
(si_cnode_id, unat seL4_CapInitThreadCNode) \<mapsto>c si_cnode_cap \<and>* R\<guillemotright> s \<and>
length untyped_slots = length untyped_caps \<and>
list_all is_untyped_cap untyped_caps \<and>
list_all (\<lambda>c. is_device_cap c = dev) untyped_caps \<and>
list_all well_formed_untyped_cap untyped_caps \<and>
distinct_sets (map cap_free_ids untyped_caps) \<and>
(\<Union>x\<in>set untyped_caps. cap_free_ids x) \<subseteq> all_available_ids \<and>
@ -722,7 +736,7 @@ lemma retype_untyped_loop_inv_success:
(drop (Suc obj_id_index) free_slots) \<and>*
\<And>* map (object_empty spec t)
(take obj_id_index obj_ids) \<and>*
\<And>* map (si_cap_at t si_caps spec)
\<And>* map (si_cap_at t si_caps spec dev)
(take obj_id_index obj_ids) \<and>* R" and
indices = "[index\<leftarrow>[0..<length untyped_slots].
\<not> is_full_untyped_cap (untyped_caps ! index)]" and
@ -751,7 +765,7 @@ lemma retype_untyped_loop_inv_success:
apply (rule conjI)
apply (metis UN_subset_iff nth_mem)
apply (rule retype_untyped_loop_inv_pre, simp+)
apply clarsimp
apply (clarsimp)
apply (rule_tac x="untyped_caps[untyped_index :=
remove_free_ids (untyped_caps ! untyped_index)
((cap_free_ids (untyped_caps ! untyped_index)) - available_ids')]" in exI)
@ -764,12 +778,16 @@ lemma retype_untyped_loop_inv_success:
apply (subst remove_free_ids_simps)
apply (rule list_all_nth, assumption, assumption)
apply fast
apply (drule list_all_nth [where P = "\<lambda>c. is_device_cap c = dev"], simp)
apply simp
apply (rule conjI)
apply (subst length_list_update, rule refl)
apply (subst (asm) take_insert_nth, assumption)
apply (subst conj_assoc[symmetric])
apply (subst list_all_conj)
apply (rule conjI)
apply (erule (1) list_all_update, simp)
apply (rule list_all_update,simp)
apply (rule list_all_conjI,simp+)
apply (rule conjI)
apply (erule (1) list_all_update)
apply (erule well_formed_untyped_cap_remove_free_ids)
@ -784,7 +802,6 @@ lemma retype_untyped_loop_inv_success:
apply (subst ran_insert_new, simp, force)
apply clarsimp
apply (rule conjI)
apply clarsimp
apply (rule take_insert_nth)
apply clarsimp
apply (rule conjI)
@ -819,7 +836,7 @@ lemma retype_untyped_bij_fail:
\<lbrace>\<lambda>s.
bij_betw_map t used_spec_ids used_ids \<and> obj_id \<notin> used_spec_ids \<and>
available_ids \<subseteq> all_available_ids \<and>
\<guillemotleft>(si_cnode_id, unat untyped_cptr) \<mapsto>c UntypedCap cover_ids available_ids \<and>*
\<guillemotleft>(si_cnode_id, unat untyped_cptr) \<mapsto>c UntypedCap dev cover_ids available_ids \<and>*
(\<And>* obj_id \<in> all_available_ids. (obj_id \<mapsto>o Untyped)) \<and>*
(si_cnode_id, unat free_cptr) \<mapsto>c NullCap \<and>*
si_tcb_id \<mapsto>f root_tcb \<and>*
@ -835,7 +852,7 @@ lemma retype_untyped_bij_fail:
\<lbrace>\<lambda>rv s. rv \<longrightarrow>
(bij_betw_map t used_spec_ids used_ids \<and> obj_id \<notin> used_spec_ids \<and>
available_ids \<subseteq> all_available_ids \<and>
\<guillemotleft>(si_cnode_id, unat untyped_cptr) \<mapsto>c UntypedCap cover_ids available_ids \<and>*
\<guillemotleft>(si_cnode_id, unat untyped_cptr) \<mapsto>c UntypedCap dev cover_ids available_ids \<and>*
(\<And>* obj_id \<in> all_available_ids. (obj_id \<mapsto>o Untyped)) \<and>*
(si_cnode_id, unat free_cptr) \<mapsto>c NullCap \<and>*
si_tcb_id \<mapsto>f root_tcb \<and>*
@ -900,7 +917,7 @@ lemma retype_untyped_loop_inv_fail:
\<And>* map (\<lambda>slot. (si_cnode_id, slot) \<mapsto>c NullCap) (drop obj_id_index free_slots) \<and>*
\<And>* map (object_empty spec t) (take obj_id_index obj_ids) \<and>*
(\<And>* obj_id \<in> all_available_ids. obj_id \<mapsto>o Untyped) \<and>*
\<And>* map (si_cap_at t si_caps spec) (take obj_id_index obj_ids) \<and>*
\<And>* map (si_cap_at t si_caps spec dev) (take obj_id_index obj_ids) \<and>*
si_tcb_id \<mapsto>f root_tcb \<and>*
(si_tcb_id, tcb_pending_op_slot) \<mapsto>c RunningCap \<and>*
(si_tcb_id, tcb_cspace_slot) \<mapsto>c si_cspace_cap \<and>*
@ -908,6 +925,7 @@ lemma retype_untyped_loop_inv_fail:
(si_cnode_id, unat seL4_CapInitThreadCNode) \<mapsto>c si_cnode_cap \<and>* R\<guillemotright> s \<and>
length untyped_slots = length untyped_caps \<and>
list_all is_untyped_cap untyped_caps \<and>
list_all (\<lambda>c. is_device_cap c = dev) untyped_caps \<and>
list_all well_formed_untyped_cap untyped_caps \<and>
list_all (\<lambda>n. n < 2 ^ si_cnode_size) untyped_slots \<and>
list_all (\<lambda>n. n < 2 ^ si_cnode_size) free_slots \<and>
@ -927,7 +945,7 @@ lemma retype_untyped_loop_inv_fail:
\<And>* map (object_empty spec t)
(take obj_id_index obj_ids) \<and>*
(\<And>* obj_id \<in> all_available_ids. obj_id \<mapsto>o Untyped) \<and>*
\<And>* map (si_cap_at t si_caps spec)
\<And>* map (si_cap_at t si_caps spec dev)
(take obj_id_index obj_ids) \<and>*
si_tcb_id \<mapsto>f root_tcb \<and>*
(si_tcb_id, tcb_pending_op_slot) \<mapsto>c RunningCap \<and>*
@ -936,6 +954,7 @@ lemma retype_untyped_loop_inv_fail:
(si_cnode_id, unat seL4_CapInitThreadCNode) \<mapsto>c si_cnode_cap \<and>* R\<guillemotright> s \<and>
length untyped_slots = length untyped_caps \<and>
list_all is_untyped_cap untyped_caps \<and>
list_all (\<lambda>c. is_device_cap c = dev) untyped_caps \<and>
list_all well_formed_untyped_cap untyped_caps \<and>
distinct_sets (map cap_free_ids untyped_caps) \<and>
(\<Union>x\<in>set untyped_caps. cap_free_ids x) \<subseteq> all_available_ids \<and>
@ -966,7 +985,7 @@ lemma retype_untyped_loop_inv_fail:
(drop (Suc obj_id_index) free_slots) \<and>*
\<And>* map (object_empty spec t)
(take obj_id_index obj_ids) \<and>*
\<And>* map (si_cap_at t si_caps spec)
\<And>* map (si_cap_at t si_caps spec dev)
(take obj_id_index obj_ids) \<and>* R" and
untyped_slots1 = untyped_slots and
untyped_caps1 = untyped_caps and
@ -1019,7 +1038,7 @@ lemma retype_untyped_loop_inv_fail:
apply (clarsimp simp: untyped_cap_eq sep_conj_assoc)
apply (subst (asm) unat_of_nat32)
apply (metis (full_types) si_cnode_size_less_than_word_size unat_less_word_bits unat_power_lower32)
apply sep_solve
apply sep_cancel+
apply (subst (asm) list_all_imp_filter2, simp)
done
@ -1044,7 +1063,7 @@ lemma retype_untyped_loop_inv_helper:
\<And>* map (\<lambda>slot. (si_cnode_id, slot) \<mapsto>c NullCap) (drop obj_id_index free_slots) \<and>*
\<And>* map (object_empty spec t) (take obj_id_index obj_ids) \<and>*
(\<And>* obj_id \<in> all_available_ids. obj_id \<mapsto>o Untyped) \<and>*
\<And>* map (si_cap_at t si_caps spec) (take obj_id_index obj_ids) \<and>*
\<And>* map (si_cap_at t si_caps spec dev) (take obj_id_index obj_ids) \<and>*
si_tcb_id \<mapsto>f root_tcb \<and>*
(si_tcb_id, tcb_pending_op_slot) \<mapsto>c RunningCap \<and>*
(si_tcb_id, tcb_cspace_slot) \<mapsto>c si_cspace_cap \<and>*
@ -1052,6 +1071,7 @@ lemma retype_untyped_loop_inv_helper:
(si_cnode_id, unat seL4_CapInitThreadCNode) \<mapsto>c si_cnode_cap \<and>* R\<guillemotright> s \<and>
length untyped_slots = length untyped_caps \<and>
list_all is_untyped_cap untyped_caps \<and>
list_all (\<lambda>c. is_device_cap c = dev) untyped_caps \<and>
list_all well_formed_untyped_cap untyped_caps \<and>
list_all (\<lambda>n. n < 2 ^ si_cnode_size) untyped_slots \<and>
list_all (\<lambda>n. n < 2 ^ si_cnode_size) free_slots \<and>
@ -1074,7 +1094,7 @@ lemma retype_untyped_loop_inv_helper:
(take (if rv then obj_id_index else Suc obj_id_index) obj_ids) \<and>*
(\<And>* obj_id \<in> all_available_ids. obj_id \<mapsto>o Untyped) \<and>*
\<And>* map (si_cap_at t (if rv then si_caps
else si_caps(obj_id \<mapsto> free_cptr)) spec)
else si_caps(obj_id \<mapsto> free_cptr)) spec dev)
(take (if rv then obj_id_index else Suc obj_id_index) obj_ids) \<and>*
si_tcb_id \<mapsto>f root_tcb \<and>*
(si_tcb_id, tcb_pending_op_slot) \<mapsto>c RunningCap \<and>*
@ -1083,6 +1103,7 @@ lemma retype_untyped_loop_inv_helper:
(si_cnode_id, unat seL4_CapInitThreadCNode) \<mapsto>c si_cnode_cap \<and>* R\<guillemotright> s \<and>
length untyped_slots = length untyped_caps \<and>
list_all is_untyped_cap untyped_caps \<and>
list_all (\<lambda>c. is_device_cap c = dev) untyped_caps \<and>
list_all well_formed_untyped_cap untyped_caps \<and>
distinct_sets (map cap_free_ids untyped_caps) \<and>
(\<Union>x\<in>set untyped_caps. cap_free_ids x) \<subseteq> all_available_ids \<and>
@ -1101,19 +1122,10 @@ lemma retype_untyped_loop_inv_helper:
apply (case_tac r, simp_all)
done
lemma nth_mem_sub:
"\<lbrakk>set xs \<subseteq> dom f; n < length xs\<rbrakk> \<Longrightarrow> f (xs ! n) = Some (the (f (xs ! n)))"
by (metis Some_the nth_mem set_rev_mp)
lemma retype_untyped_loop_inv:
"\<lbrace>\<lambda>s. \<exists>untyped_caps t all_available_ids.
\<guillemotleft>\<And>* map (\<lambda>(slot, cap). (si_cnode_id, slot) \<mapsto>c cap)
@ -1121,7 +1133,7 @@ lemma retype_untyped_loop_inv:
\<And>* map (\<lambda>slot. (si_cnode_id, slot) \<mapsto>c NullCap) (drop obj_id_index free_slots) \<and>*
\<And>* map (object_empty spec t) (take obj_id_index obj_ids) \<and>*
(\<And>* obj_id \<in> all_available_ids. obj_id \<mapsto>o Untyped) \<and>*
\<And>* map (si_cap_at t si_caps spec) (take obj_id_index obj_ids) \<and>*
\<And>* map (si_cap_at t si_caps spec dev) (take obj_id_index obj_ids) \<and>*
si_tcb_id \<mapsto>f root_tcb \<and>*
(si_tcb_id, tcb_pending_op_slot) \<mapsto>c RunningCap \<and>*
(si_tcb_id, tcb_cspace_slot) \<mapsto>c si_cspace_cap \<and>*
@ -1143,6 +1155,7 @@ lemma retype_untyped_loop_inv:
untyped_cptr = untyped_cptrs ! untyped_index \<and>
type = object_type (the (cdl_objects spec obj_id)) \<and>
list_all is_untyped_cap untyped_caps \<and>
list_all (\<lambda>c. is_device_cap c = dev) untyped_caps \<and>
list_all well_formed_untyped_cap untyped_caps \<and>
list_all (\<lambda>n. n < 2 ^ si_cnode_size) untyped_slots \<and>
list_all (\<lambda>n. n < 2 ^ si_cnode_size) free_slots \<and>
@ -1165,7 +1178,7 @@ lemma retype_untyped_loop_inv:
(take (if rv then obj_id_index else Suc obj_id_index) obj_ids) \<and>*
(\<And>* obj_id \<in> all_available_ids. obj_id \<mapsto>o Untyped) \<and>*
\<And>* map (si_cap_at t (if rv then si_caps
else si_caps(obj_id \<mapsto> free_cptr)) spec)
else si_caps(obj_id \<mapsto> free_cptr)) spec dev)
(take (if rv then obj_id_index else Suc obj_id_index) obj_ids) \<and>*
si_tcb_id \<mapsto>f root_tcb \<and>*
(si_tcb_id, tcb_pending_op_slot) \<mapsto>c RunningCap \<and>*
@ -1178,6 +1191,7 @@ lemma retype_untyped_loop_inv:
length obj_ids \<le> length free_slots \<and>
length untyped_slots = length untyped_caps \<and>
list_all is_untyped_cap untyped_caps \<and>
list_all (\<lambda>c. is_device_cap c = dev) untyped_caps \<and>
list_all well_formed_untyped_cap untyped_caps \<and>
distinct_sets (map cap_free_ids untyped_caps) \<and>
(\<Union>x\<in>set untyped_caps. cap_free_ids x) \<subseteq> all_available_ids \<and>
@ -1247,6 +1261,7 @@ lemma retype_untypeds_wp_helper:
(si_cnode_id, unat seL4_CapInitThreadCNode) \<mapsto>c si_cnode_cap \<and>* R\<guillemotright> s \<and>
length untyped_slots = length untyped_caps \<and>
list_all is_full_untyped_cap untyped_caps \<and>
list_all (\<lambda>c. is_device_cap c = dev) untyped_caps \<and>
list_all well_formed_untyped_cap untyped_caps \<and>
distinct_sets (map cap_free_ids untyped_caps)\<rbrace>
@ -1257,7 +1272,7 @@ lemma retype_untypeds_wp_helper:
\<And>* map (\<lambda>slot. (si_cnode_id, slot) \<mapsto>c NullCap) (drop (length [obj\<leftarrow>obj_ids. real_object_at obj spec]) free_slots) \<and>*
\<And>* map (object_empty spec t) [obj\<leftarrow>obj_ids. real_object_at obj spec] \<and>*
(\<And>* obj_id \<in> all_available_ids. obj_id \<mapsto>o Untyped) \<and>*
\<And>* map (si_cap_at t si_caps spec) [obj\<leftarrow>obj_ids. real_object_at obj spec] \<and>*
\<And>* map (si_cap_at t si_caps spec dev) [obj\<leftarrow>obj_ids. real_object_at obj spec] \<and>*
si_tcb_id \<mapsto>f root_tcb \<and>*
(si_tcb_id, tcb_pending_op_slot) \<mapsto>c RunningCap \<and>*
(si_tcb_id, tcb_cspace_slot) \<mapsto>c si_cspace_cap \<and>*
@ -1277,7 +1292,7 @@ lemma retype_untypeds_wp_helper:
\<And>* map (\<lambda>slot. (si_cnode_id, slot) \<mapsto>c NullCap) (drop obj_id_index free_slots) \<and>*
\<And>* map (object_empty spec t) (take obj_id_index [obj\<leftarrow>obj_ids. real_object_at obj spec]) \<and>*
(\<And>* obj_id \<in> all_available_ids. obj_id \<mapsto>o Untyped) \<and>*
\<And>* map (si_cap_at t si_caps spec) (take obj_id_index [obj\<leftarrow>obj_ids. real_object_at obj spec]) \<and>*
\<And>* map (si_cap_at t si_caps spec dev) (take obj_id_index [obj\<leftarrow>obj_ids. real_object_at obj spec]) \<and>*
si_tcb_id \<mapsto>f root_tcb \<and>*
(si_tcb_id, tcb_pending_op_slot) \<mapsto>c RunningCap \<and>*
(si_tcb_id, tcb_cspace_slot) \<mapsto>c si_cspace_cap \<and>*
@ -1288,6 +1303,7 @@ lemma retype_untypeds_wp_helper:
untyped_index \<le> length untyped_cptrs \<and>
length untyped_slots = length untyped_caps \<and>
list_all is_untyped_cap untyped_caps \<and>
list_all (\<lambda>c. is_device_cap c = dev) untyped_caps \<and>
list_all well_formed_untyped_cap untyped_caps \<and>
distinct_sets (map cap_free_ids untyped_caps) \<and>
(\<Union>x\<in>set untyped_caps. cap_free_ids x) \<subseteq> all_available_ids \<and>
@ -1309,7 +1325,7 @@ lemma retype_untypeds_wp_helper:
free_slots=free_slots and free_cptrs=free_cptrs and
untyped_slots=untyped_slots and
untyped_cptrs=untyped_cptrs and
si_caps=si_caps and
si_caps=si_caps and dev = dev and
obj_id=obj_id and R=R
in retype_untyped_loop_inv)
apply (erule pre_post_ex)+
@ -1365,6 +1381,7 @@ lemma create_objects_sep:
distinct untyped_cptrs \<and>
length untyped_cptrs = length untyped_caps \<and>
list_all is_full_untyped_cap untyped_caps \<and>
list_all (\<lambda>c. is_device_cap c = dev) untyped_caps \<and>
list_all well_formed_untyped_cap untyped_caps \<and>
distinct_sets (map cap_free_ids untyped_caps) \<and>
card {obj_id. real_object_at obj_id spec} \<le> length free_cptrs \<and>
@ -1375,7 +1392,7 @@ lemma create_objects_sep:
\<lbrace>\<lambda>rv s. \<exists>t.
\<guillemotleft>(objects_empty spec t {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t (fst rv) spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t (fst rv) spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>*
si_objects_extra_caps' {obj_id. real_object_at obj_id spec} free_cptrs untyped_cptrs \<and>*
R) and
@ -1390,7 +1407,7 @@ lemma create_objects_sep:
apply (rule hoare_chain)
apply (wp retype_untypeds_wp_helper
[where R="(si_cnode_id, unat seL4_CapIRQControl) \<mapsto>c IrqControlCap \<and>* si_asid \<and>* R"
and untyped_slots = "map unat untyped_cptrs"
and untyped_slots = "map unat untyped_cptrs" and dev = dev
and free_slots = "map unat free_cptrs"],
(simp|clarsimp)+)
apply (fastforce simp: list_all_iff unat_less_2_si_cnode_size')
@ -1401,7 +1418,7 @@ lemma create_objects_sep:
apply (subst comp_apply)+
apply (subst sep_list_conj_sep_map_set_conj, simp add: distinct_zipI1)
apply (subst sep_list_conj_sep_map_set_conj, simp add: distinct_zipI1)
apply sep_solve
apply sep_cancel+
apply (clarsimp simp: length_real_object_at_card)
apply (rule_tac x=t in exI)
apply (clarsimp simp: sep_conj_exists sep_conj_assoc

View File

@ -37,18 +37,18 @@ lemma sep_map_c_unat [simp]:
by fastforce
lemma wfdc_obj_not_untyped:
"well_formed_cap (default_cap t oid sz) \<Longrightarrow> t \<noteq> UntypedType"
"well_formed_cap (default_cap t oid sz dev) \<Longrightarrow> t \<noteq> UntypedType"
by (clarsimp simp:default_cap_def well_formed_cap_def)
lemma derived_cap_default:
"derived_cap (default_cap ty oid sz)
= (default_cap ty oid sz)"
"derived_cap (default_cap ty oid sz dev)
= (default_cap ty oid sz dev)"
by (case_tac ty,
simp_all add:derived_cap_def default_cap_def)
lemma cnode_or_tcb_at_wfc:
"\<lbrakk>cnode_or_tcb_at obj_id spec; cdl_objects spec obj_id = Some obj; sz \<le> 32\<rbrakk>
\<Longrightarrow> well_formed_cap (default_cap (object_type obj) oid sz)"
\<Longrightarrow> well_formed_cap (default_cap (object_type obj) oid sz dev)"
apply (elim disjE)
apply (simp add: object_at_def is_tcb_def default_cap_def well_formed_cap_def
is_cnode_def object_type_def guard_bits_def
@ -58,7 +58,7 @@ lemma cnode_or_tcb_at_wfc:
lemma seL4_CNode_Copy_sep_helper:
"\<lbrakk>(free_cptr::32 word) < 2 ^ si_cnode_size;
(cap_ptr::32 word) < 2 ^ si_cnode_size;
well_formed_cap (default_cap (object_type obj) {kobj_id} (object_size_bits obj))\<rbrakk>
well_formed_cap (default_cap (object_type obj) {kobj_id} (object_size_bits obj) dev)\<rbrakk>
\<Longrightarrow>
\<lbrace>\<lambda>s. \<guillemotleft>si_tcb_id \<mapsto>f root_tcb \<and>*
(si_tcb_id, tcb_cspace_slot) \<mapsto>c si_cspace_cap \<and>*
@ -66,7 +66,7 @@ lemma seL4_CNode_Copy_sep_helper:
si_cnode_id \<mapsto>f CNode (empty_cnode si_cnode_size) \<and>*
(si_cnode_id, unat seL4_CapInitThreadCNode) \<mapsto>c si_cnode_cap \<and>*
(si_cnode_id, unat cap_ptr) \<mapsto>c default_cap (object_type obj) {kobj_id}
(object_size_bits obj) \<and>*
(object_size_bits obj) dev \<and>*
(si_cnode_id, unat free_cptr) \<mapsto>c NullCap \<and>* R\<guillemotright> s\<rbrace>
seL4_CNode_Copy seL4_CapInitThreadCNode free_cptr 32
seL4_CapInitThreadCNode cap_ptr 32 UNIV
@ -76,16 +76,16 @@ lemma seL4_CNode_Copy_sep_helper:
si_cnode_id \<mapsto>f CNode (empty_cnode si_cnode_size) \<and>*
(si_cnode_id, unat seL4_CapInitThreadCNode) \<mapsto>c si_cnode_cap \<and>*
(si_cnode_id, unat cap_ptr) \<mapsto>c default_cap (object_type obj) {kobj_id}
(object_size_bits obj) \<and>*
(object_size_bits obj) dev \<and>*
(si_cnode_id, unat free_cptr) \<mapsto>c default_cap (object_type obj) {kobj_id}
(object_size_bits obj) \<and>* R\<guillemotright>\<rbrace>"
(object_size_bits obj) dev \<and>* R\<guillemotright>\<rbrace>"
apply (rule hoare_chain)
apply (rule_tac src_index=cap_ptr and
cnode_cap=si_cspace_cap and
cnode_cap'=si_cnode_cap and
root_size=si_cnode_size and
src_cap="default_cap (object_type obj) {kobj_id}
(object_size_bits obj)" and
(object_size_bits obj) dev" and
R=R in
seL4_CNode_Copy_sep, (simp add: wfdc_obj_not_untyped offset_slot offset_slot'|clarsimp)+)
apply (rule conjI)
@ -103,12 +103,12 @@ lemma duplicate_cap_sep_helper:
set obj_ids = dom (cdl_objects spec)\<rbrakk>
\<Longrightarrow>
\<lbrace>\<guillemotleft>(si_cnode_id, unat free_cptr) \<mapsto>c NullCap \<and>*
si_cap_at t orig_caps spec obj_id \<and>* si_objects \<and>* R\<guillemotright>\<rbrace>
si_cap_at t orig_caps spec dev obj_id \<and>* si_objects \<and>* R\<guillemotright>\<rbrace>
duplicate_cap spec orig_caps (obj_id, free_cptr)
\<lbrace>\<lambda>_ s.
\<guillemotleft>si_cap_at t (map_of (zip [obj\<leftarrow>obj_ids. cnode_or_tcb_at obj spec] free_cptrs))
spec obj_id \<and>*
si_cap_at t orig_caps spec obj_id \<and>* si_objects \<and>* R\<guillemotright> s\<rbrace>"
spec dev obj_id \<and>*
si_cap_at t orig_caps spec dev obj_id \<and>* si_objects \<and>* R\<guillemotright> s\<rbrace>"
apply (rule hoare_assume_pre)
apply (clarsimp simp: duplicate_cap_def si_cap_at_def sep_conj_exists)
apply (rule_tac x=free_cptr in hoare_exI)
@ -121,7 +121,7 @@ lemma duplicate_cap_sep_helper:
apply (clarsimp simp: si_objects_def Ball_set_list_all[symmetric])
apply (wp hoare_drop_imps)
apply (rule hoare_chain)
apply (rule_tac free_cptr=free_cptr and cap_ptr=cap_ptr and
apply (rule_tac free_cptr=free_cptr and cap_ptr=cap_ptr and dev = dev and
R="(si_cnode_id, unat seL4_CapIRQControl) \<mapsto>c IrqControlCap \<and>* si_asid \<and>* R" in
seL4_CNode_Copy_sep_helper)
apply (rule unat_less_2_si_cnode_size, simp)
@ -136,7 +136,7 @@ lemma duplicate_cap_sep_helper:
lemma duplicate_cap_sep:
"\<lbrace>\<guillemotleft>(si_cnode_id, unat free_cptr) \<mapsto>c NullCap \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>* si_objects \<and>* R\<guillemotright> and K (
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>* si_objects \<and>* R\<guillemotright> and K (
well_formed spec \<and> distinct obj_ids \<and>
list_all (\<lambda>n. n < 2 ^ si_cnode_size) (map unat free_cptrs) \<and>
(obj_id, free_cptr) \<in> set (zip [obj\<leftarrow>obj_ids. cnode_or_tcb_at obj spec] free_cptrs) \<and>
@ -144,18 +144,18 @@ lemma duplicate_cap_sep:
duplicate_cap spec orig_caps (obj_id, free_cptr)
\<lbrace>\<lambda>_.
\<guillemotleft>si_cap_at t (map_of (zip [obj\<leftarrow>obj_ids. cnode_or_tcb_at obj spec] free_cptrs))
spec obj_id \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>* si_objects \<and>* R\<guillemotright>\<rbrace>"
spec dev obj_id \<and>*
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>* si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (rule hoare_gen_asm)
apply clarsimp
apply (frule well_formed_finite [where obj_id=obj_id])
apply (clarsimp simp: si_caps_at_def)
apply (rule hoare_chain [where
P="\<guillemotleft>((si_cnode_id, unat free_cptr) \<mapsto>c NullCap \<and>* si_objects) \<and>*
(\<And>* obj_id \<in> {obj_id. real_object_at obj_id spec}. si_cap_at t orig_caps spec obj_id) \<and>* R\<guillemotright>" and
(\<And>* obj_id \<in> {obj_id. real_object_at obj_id spec}. si_cap_at t orig_caps spec dev obj_id) \<and>* R\<guillemotright>" and
Q="\<lambda>rv.\<guillemotleft>(si_cap_at t (map_of (zip [obj\<leftarrow>obj_ids. cnode_or_tcb_at obj spec]
free_cptrs)) spec obj_id \<and>* si_objects) \<and>*
(\<And>* obj_id \<in> {obj_id. real_object_at obj_id spec}. si_cap_at t orig_caps spec obj_id) \<and>* R\<guillemotright>"])
free_cptrs)) spec dev obj_id \<and>* si_objects) \<and>*
(\<And>* obj_id \<in> {obj_id. real_object_at obj_id spec}. si_cap_at t orig_caps spec dev obj_id) \<and>* R\<guillemotright>"])
apply (rule sep_set_conj_map_singleton_wp [where x=obj_id])
apply simp
apply (drule in_set_zip1)
@ -170,7 +170,7 @@ lemma duplicate_cap_sep:
done
lemma duplicate_caps_sep_helper:
"\<lbrace>\<guillemotleft>si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
"\<lbrace>\<guillemotleft>si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
(\<And>* (obj_id, free_cptr) \<in> set (zip [obj\<leftarrow>obj_ids. cnode_or_tcb_at obj spec] free_cptrs).
(si_cnode_id, unat free_cptr) \<mapsto>c NullCap) \<and>*
si_objects \<and>* R\<guillemotright> and K(
@ -180,8 +180,8 @@ lemma duplicate_caps_sep_helper:
length [obj\<leftarrow>obj_ids . cnode_or_tcb_at obj spec] \<le> length free_cptrs)\<rbrace>
duplicate_caps spec orig_caps obj_ids free_cptrs
\<lbrace>\<lambda>dup_caps.
\<guillemotleft>si_caps_at t dup_caps spec {obj_id. cnode_or_tcb_at obj_id spec} \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>* si_objects \<and>* R\<guillemotright>\<rbrace>"
\<guillemotleft>si_caps_at t dup_caps spec dev {obj_id. cnode_or_tcb_at obj_id spec} \<and>*
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>* si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (rule hoare_gen_asm)
apply (clarsimp simp: duplicate_caps_def si_caps_at_def)
apply (wp)
@ -192,8 +192,8 @@ lemma duplicate_caps_sep_helper:
P="\<lambda>(obj_id,free_cptr). (si_cnode_id, unat free_cptr) \<mapsto>c NullCap" and
Q="\<lambda>(obj_id,free_cptr). (si_cap_at t (map_of
(zip [obj\<leftarrow>obj_ids. cnode_or_tcb_at obj spec] free_cptrs))
spec obj_id)" and
I="si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>* si_objects" and
spec dev obj_id)" and
I="si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>* si_objects" and
R=R])
apply (rule distinct_zipI1, simp)
apply (clarsimp simp: sep_conj_assoc)
@ -218,7 +218,7 @@ lemma distinct_length_filter':
by (metis distinct_length_filter set_conj_Int_simp inf_commute)
lemma duplicate_caps_sep_no_rv:
"\<lbrace>\<guillemotleft>si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
"\<lbrace>\<guillemotleft>si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>*
si_objects_extra_caps' (set obj_ids) free_cptrs untyped_cptrs \<and>*
R\<guillemotright> and K(well_formed spec \<and> distinct obj_ids \<and> distinct free_cptrs \<and>
@ -228,8 +228,8 @@ lemma duplicate_caps_sep_no_rv:
free_cptrs' = drop (length obj_ids) free_cptrs)\<rbrace>
duplicate_caps spec orig_caps obj_ids free_cptrs'
\<lbrace>\<lambda>dup_caps s.
\<guillemotleft>si_caps_at t dup_caps spec {obj_id. cnode_or_tcb_at obj_id spec} \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
\<guillemotleft>si_caps_at t dup_caps spec dev {obj_id. cnode_or_tcb_at obj_id spec} \<and>*
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>*
si_objects_extra_caps (set obj_ids) free_cptrs untyped_cptrs spec \<and>*
R\<guillemotright> s\<rbrace>"
@ -275,7 +275,7 @@ lemma duplicate_caps_sep_no_rv:
done
lemma duplicate_caps_rv:
"\<lbrace>\<guillemotleft>si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
"\<lbrace>\<guillemotleft>si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>*
si_objects_extra_caps' (set obj_ids) free_cptrs untyped_cptrs \<and>* R\<guillemotright>\<rbrace>
duplicate_caps spec orig_caps obj_ids free_cptrs'
@ -285,7 +285,7 @@ lemma duplicate_caps_rv:
done
lemma duplicate_caps_sep:
"\<lbrace>\<guillemotleft>(si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
"\<lbrace>\<guillemotleft>(si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>*
si_objects_extra_caps' (dom (cdl_objects spec)) free_cptrs_orig untyped_cptrs \<and>* R) and
K (well_formed spec \<and>
@ -297,8 +297,8 @@ lemma duplicate_caps_sep:
length obj_ids + card {obj_id. cnode_or_tcb_at obj_id spec} \<le> length free_cptrs_orig)\<guillemotright>\<rbrace>
duplicate_caps spec orig_caps obj_ids free_cptrs
\<lbrace>\<lambda>dup_caps.
\<guillemotleft>(si_caps_at t dup_caps spec {obj_id. cnode_or_tcb_at obj_id spec} \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
\<guillemotleft>(si_caps_at t dup_caps spec dev {obj_id. cnode_or_tcb_at obj_id spec} \<and>*
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>*
si_objects_extra_caps (dom (cdl_objects spec)) free_cptrs_orig untyped_cptrs spec \<and>* R) and
K (dup_caps = map_of (zip [obj\<leftarrow>obj_ids. cnode_or_tcb_at obj spec] free_cptrs))\<guillemotright> \<rbrace>"

View File

@ -65,7 +65,7 @@ definition
tcb_vspace_slot \<mapsto> PageDirectoryCap pd_a_id Real None,
tcb_replycap_slot \<mapsto> NullCap,
tcb_caller_slot \<mapsto> NullCap,
tcb_ipcbuffer_slot \<mapsto> FrameCap frame_a1_id {AllowRead, AllowWrite} small_frame_size Real None,
tcb_ipcbuffer_slot \<mapsto> FrameCap False frame_a1_id {AllowRead, AllowWrite} small_frame_size Real None,
tcb_pending_op_slot \<mapsto> NullCap,
tcb_boundntfn_slot \<mapsto> NullCap],
cdl_tcb_fault_endpoint = 0,
@ -80,7 +80,7 @@ definition
tcb_vspace_slot \<mapsto> PageDirectoryCap pd_b_id Real None,
tcb_replycap_slot \<mapsto> NullCap,
tcb_caller_slot \<mapsto> NullCap,
tcb_ipcbuffer_slot \<mapsto> FrameCap frame_b_id {AllowRead, AllowWrite} small_section_size Real None,
tcb_ipcbuffer_slot \<mapsto> FrameCap False frame_b_id {AllowRead, AllowWrite} small_section_size Real None,
tcb_pending_op_slot \<mapsto> NullCap,
tcb_boundntfn_slot \<mapsto> NullCap],
cdl_tcb_fault_endpoint = 0,
@ -110,9 +110,9 @@ definition
2 \<mapsto> CNodeCap cnode_a1_id guard guard_size cnode_a1_size,
3 \<mapsto> PageDirectoryCap pd_a_id Real None,
4 \<mapsto> PageTableCap pt_a_id Real None,
8 \<mapsto> FrameCap frame_a1_id {AllowRead, AllowWrite} small_frame_size Real None,
8 \<mapsto> FrameCap False frame_a1_id {AllowRead, AllowWrite} small_frame_size Real None,
10 \<mapsto> NotificationCap ntfn_id 0 {Read},
11 \<mapsto> FrameCap frame_a2_id {AllowRead, AllowWrite} small_frame_size Real None,
11 \<mapsto> FrameCap False frame_a2_id {AllowRead, AllowWrite} small_frame_size Real None,
12 \<mapsto> IrqHandlerCap 4]"
definition
@ -122,7 +122,7 @@ definition
2 \<mapsto> CNodeCap cnode_b_id guard guard_size cnode_b_size,
4 \<mapsto> EndpointCap ep_id 0 {Read},
7 \<mapsto> PageDirectoryCap pd_b_id Real None,
8 \<mapsto> FrameCap frame_b_id {AllowRead, AllowWrite} small_section_size Real None,
8 \<mapsto> FrameCap False frame_b_id {AllowRead, AllowWrite} small_section_size Real None,
254 \<mapsto> IrqHandlerCap 254]"
definition
@ -137,12 +137,12 @@ definition
definition
"pd_b \<equiv> \<lparr>cdl_page_directory_caps = new_cap_map pd_size
[2 \<mapsto> FrameCap frame_b_id {AllowRead, AllowWrite} small_section_size Fake None]\<rparr>"
[2 \<mapsto> FrameCap False frame_b_id {AllowRead, AllowWrite} small_section_size Fake None]\<rparr>"
definition
"pt_a \<equiv> \<lparr>cdl_page_table_caps = new_cap_map pt_size
[0 \<mapsto> FrameCap frame_a1_id {AllowRead, AllowWrite} small_frame_size Fake None,
255 \<mapsto> FrameCap frame_a2_id {AllowRead, AllowWrite} small_frame_size Fake None]\<rparr>"
[0 \<mapsto> FrameCap False frame_a1_id {AllowRead, AllowWrite} small_frame_size Fake None,
255 \<mapsto> FrameCap False frame_a2_id {AllowRead, AllowWrite} small_frame_size Fake None]\<rparr>"
definition
"empty_frame \<equiv> \<lparr>cdl_frame_size_bits = small_frame_size\<rparr>"
@ -265,7 +265,7 @@ lemma is_fake_pt_cap_simps:
by (clarsimp simp: is_fake_pt_cap_def)+
lemma frame_cap_not_cnode:
"\<not>is_cnode_cap (FrameCap a b c d e)"
"\<not>is_cnode_cap (FrameCap dev a b c d e)"
by (clarsimp simp: cap_type_def)
lemma empty_cap_map_NullCap [simp]:
@ -892,8 +892,8 @@ lemma well_formed_example:
apply (clarsimp simp: cnode_at_example_spec)
by (auto simp: example_spec_def object_size_bits_def object_default_state_def2
pd_size_def word_bits_def empty_cnode_def is_cnode_def
object_slots_def empty_cap_map_def tcb_slot_defs
default_tcb_def obj_defs
object_slots_def empty_cap_map_def tcb_slot_defs slots_of_def
default_tcb_def obj_defs cap_at_def opt_cap_def opt_object_def
small_frame_size_def small_section_size_def pt_size_def
new_cnode_def new_cap_map_def empty_irq_node_def
new_irq_node_def

View File

@ -49,7 +49,7 @@ lemma well_formed_empty:
cdl_asid_table = undefined,
cdl_current_domain = undefined
\<rparr>"
by (clarsimp simp: well_formed_def well_formed_orig_caps_unique_def
by (clarsimp simp: well_formed_def well_formed_orig_caps_unique_def cap_at_def
well_formed_irqhandler_caps_unique_def well_formed_irqhandler_caps_def
well_formed_irq_table_def down_ucast_inj is_down
well_formed_fake_pt_caps_unique_def irq_nodes_def object_at_def
@ -68,7 +68,7 @@ definition
tcb_vspace_slot \<mapsto> PageDirectoryCap pd_id Real None,
tcb_replycap_slot \<mapsto> NullCap,
tcb_caller_slot \<mapsto> NullCap,
tcb_ipcbuffer_slot \<mapsto> FrameCap frame_id {AllowRead, AllowWrite} small_frame_size Real None,
tcb_ipcbuffer_slot \<mapsto> FrameCap False frame_id {AllowRead, AllowWrite} small_frame_size Real None,
tcb_pending_op_slot \<mapsto> NullCap,
tcb_boundntfn_slot \<mapsto> NullCap],
cdl_tcb_fault_endpoint = 0,
@ -103,7 +103,7 @@ definition
[0 \<mapsto> TcbCap tcb_id,
1 \<mapsto> CNodeCap cnode_id 0 0 2,
2 \<mapsto> PageDirectoryCap pd_id Real None,
3 \<mapsto> FrameCap frame_id {AllowRead, AllowWrite} small_frame_size Real None],
3 \<mapsto> FrameCap False frame_id {AllowRead, AllowWrite} small_frame_size Real None],
cdl_cnode_size_bits = 2\<rparr>),
pd_id \<mapsto> PageDirectory \<lparr>cdl_page_directory_caps = empty_cap_map pd_size\<rparr>,
frame_id \<mapsto> Frame \<lparr>cdl_frame_size_bits = small_frame_size\<rparr>],
@ -337,27 +337,34 @@ lemma well_formed_example:
"well_formed example_spec"
apply (clarsimp simp: well_formed_def)
apply (intro conjI)
apply (rule well_formed_orig_caps_unique_example)
apply (rule well_formed_irqhandler_caps_unique_example)
apply (rule well_formed_fake_pt_caps_unique_example)
apply (rule well_formed_irqhandler_caps_example)
apply (clarsimp split: option.splits, rename_tac obj)
apply (clarsimp simp: well_formed_cap_to_object_example
apply (rule well_formed_orig_caps_unique_example)
apply (rule well_formed_irqhandler_caps_unique_example)
apply (rule well_formed_fake_pt_caps_unique_example)
apply (rule well_formed_irqhandler_caps_example)
apply (clarsimp split: option.splits, rename_tac obj)
apply (clarsimp simp: well_formed_cap_to_object_example
well_formed_orig_caps_unique_example)
apply (rule conjI)
apply (erule well_formed_tcb_example)
apply (rule conjI)
apply (fact well_formed_vspace_example)
apply (rule conjI)
apply (fact well_formed_irq_node_example)
by (fastforce simp: example_spec_def object_size_bits_def object_default_state_def2
apply (rule conjI)
apply (erule well_formed_tcb_example)
apply (rule conjI)
apply (fact well_formed_vspace_example)
apply (rule conjI)
apply (fact well_formed_irq_node_example)
apply (fastforce simp: example_spec_def object_size_bits_def object_default_state_def2
pd_size_def word_bits_def empty_cnode_def is_cnode_def
object_slots_def empty_cap_map_def tcb_slot_defs
default_tcb_def example_tcb_def
small_frame_size_def object_at_def
irq_nodes_def range_example_irq_node
split: split_if_asm)
apply (clarsimp simp: example_spec_def object_size_bits_def object_default_state_def2
pd_size_def word_bits_def empty_cnode_def is_cnode_def
object_slots_def empty_cap_map_def tcb_slot_defs
default_tcb_def example_tcb_def cap_at_def opt_cap_def slots_of_def
small_frame_size_def object_at_def opt_object_def
irq_nodes_def range_example_irq_node
split: split_if_asm)
done
end
end

File diff suppressed because it is too large Load Diff

View File

@ -18,7 +18,7 @@ begin
lemma seL4_IRQHandler_SetEndpoint_irq_initialised_helper_sep:
"\<lbrace>\<guillemotleft>irq_empty spec t irq \<and>*
si_cap_at t orig_caps spec ntfn_id \<and>*
si_cap_at t orig_caps spec dev ntfn_id \<and>*
si_irq_cap_at irq_caps spec irq \<and>*
si_objects \<and>* R\<guillemotright> and
K(well_formed spec \<and>
@ -40,7 +40,7 @@ lemma seL4_IRQHandler_SetEndpoint_irq_initialised_helper_sep:
seL4_IRQHandler_SetEndpoint irq_handler_cptr endpoint_cptr
\<lbrace>\<lambda>_.
\<guillemotleft>irq_initialised spec t irq \<and>*
si_cap_at t orig_caps spec ntfn_id \<and>*
si_cap_at t orig_caps spec dev ntfn_id \<and>*
si_irq_cap_at irq_caps spec irq \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (rule hoare_gen_asm, clarsimp)
@ -98,7 +98,7 @@ lemma seL4_IRQHandler_SetEndpoint_irq_initialised_helper_sep:
lemma seL4_IRQHandler_SetEndpoint_irq_initialised_sep:
"\<lbrace>\<guillemotleft>irq_empty spec t irq \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_irq_caps_at irq_caps spec (bound_irqs spec) \<and>*
si_objects \<and>* R\<guillemotright> and
K(well_formed spec \<and>
@ -108,7 +108,7 @@ lemma seL4_IRQHandler_SetEndpoint_irq_initialised_sep:
seL4_IRQHandler_SetEndpoint irq_handler_cptr endpoint_cptr
\<lbrace>\<lambda>_.
\<guillemotleft>irq_initialised spec t irq \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_irq_caps_at irq_caps spec (bound_irqs spec) \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (rule hoare_gen_asm)
@ -121,10 +121,10 @@ lemma seL4_IRQHandler_SetEndpoint_irq_initialised_sep:
apply (frule slots_of_cdl_objects, clarsimp)
apply (rule hoare_chain [OF sep_set_conj_map_singleton_wp
[where P = "irq_empty spec t irq \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_objects"
and Q = "irq_initialised spec t irq \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_objects"
and I = "si_irq_cap_at irq_caps spec"
and x = irq
@ -137,7 +137,7 @@ lemma seL4_IRQHandler_SetEndpoint_irq_initialised_sep:
and Q = "irq_initialised spec t irq \<and>*
si_irq_cap_at irq_caps spec irq \<and>*
si_objects"
and I = "si_cap_at t orig_caps spec"
and I = "si_cap_at t orig_caps spec dev"
and x = "cap_object (the (opt_cap (cdl_irq_node spec irq, 0) spec))"
and xs = "{obj_id. real_object_at obj_id spec}"]], simp+)
apply (wp sep_wp: seL4_IRQHandler_SetEndpoint_irq_initialised_helper_sep [where t=t and spec=spec and irq=irq
@ -156,14 +156,14 @@ lemma seL4_IRQHandler_SetEndpoint_irq_initialised_sep:
lemma init_irq_sep:
"\<lbrace>\<guillemotleft>irq_empty spec t irq \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_irq_caps_at irq_caps spec (bound_irqs spec) \<and>*
si_objects \<and>* R\<guillemotright> and
K(well_formed spec \<and>
irq \<in> bound_irqs spec)\<rbrace>
init_irq spec orig_caps irq_caps irq
\<lbrace>\<lambda>_. \<guillemotleft>irq_initialised spec t irq \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_irq_caps_at irq_caps spec (bound_irqs spec) \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (rule hoare_gen_asm, clarsimp)
@ -180,13 +180,13 @@ lemma init_irq_sep:
lemma init_irqs_bound_irqs_sep:
"\<lbrace>\<guillemotleft>irqs_empty spec t (bound_irqs spec) \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_irq_caps_at irq_caps spec (bound_irqs spec) \<and>*
si_objects \<and>* R\<guillemotright> and
K(well_formed spec)\<rbrace>
init_irqs spec orig_caps irq_caps
\<lbrace>\<lambda>_.\<guillemotleft>irqs_initialised spec t (bound_irqs spec) \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_irq_caps_at irq_caps spec (bound_irqs spec) \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (rule hoare_gen_asm)
@ -196,7 +196,7 @@ lemma init_irqs_bound_irqs_sep:
apply (rule mapM_x_set_sep' [where
P="irq_empty spec t" and
Q="irq_initialised spec t" and
I="si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
I="si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_irq_caps_at irq_caps spec (bound_irqs spec) \<and>*
si_objects" and
xs="sorted_list_of_set (bound_irqs spec)" and
@ -258,13 +258,13 @@ lemma irq_empty_initialised:
*)
lemma init_irqs_sep:
"\<lbrace>\<guillemotleft>irqs_empty spec t (used_irqs spec) \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_irq_caps_at irq_caps spec (used_irqs spec) \<and>*
si_objects \<and>* R\<guillemotright> and
K(well_formed spec)\<rbrace>
init_irqs spec orig_caps irq_caps
\<lbrace>\<lambda>_.\<guillemotleft>irqs_initialised spec t (used_irqs spec) \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_irq_caps_at irq_caps spec (used_irqs spec) \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (rule hoare_gen_asm)
@ -272,9 +272,9 @@ lemma init_irqs_sep:
apply (frule well_formed_bound_irqs_are_used_irqs)
apply (frule sep_set_conj_subset_wp
[where P = "sep_map_set_conj (irq_empty spec t) (used_irqs spec) \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>* si_objects"
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>* si_objects"
and Q = "sep_map_set_conj (irq_initialised spec t) (used_irqs spec) \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>* si_objects"
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>* si_objects"
and f = "init_irqs spec orig_caps irq_caps"], simp+)
apply (subst sep.setprod.subset_diff, assumption, simp)+
apply (sep_wp init_irqs_bound_irqs_sep [where t=t])

View File

@ -189,7 +189,7 @@ lemma tcb_half_decomp [simplified]:
lemma default_cap_size_0:
"type \<noteq> CNodeType
\<Longrightarrow> default_cap type obj_id sz = default_cap type obj_id 0"
\<Longrightarrow> default_cap type obj_id sz dev = default_cap type obj_id 0 dev"
by (clarsimp simp: default_cap_def split: cdl_object_type.splits)
lemma tcb_configure_pre:
@ -223,16 +223,16 @@ lemma tcb_configure_pre:
vspace_slot = offset vspace_index si_cnode_size;
buffer_frame_slot = offset buffer_frame_index si_cnode_size;
tcb_cap = default_cap TcbType {k_obj_id} 0;
k_cspace_cap = default_cap CNodeType {cspace_kobj_id} cnode_size;
k_vspace_cap = default_cap PageDirectoryType {vspace_kobj_id} 0;
k_buffer_frame_cap = default_cap buffer_frame_type {buffer_frame_kobj_id} 0;
tcb_cap = default_cap TcbType {k_obj_id} 0 False;
k_cspace_cap = default_cap CNodeType {cspace_kobj_id} cnode_size False;
k_vspace_cap = default_cap PageDirectoryType {vspace_kobj_id} 0 False;
k_buffer_frame_cap = default_cap buffer_frame_type {buffer_frame_kobj_id} 0 False;
\<guillemotleft>object_empty spec t obj_id \<and>*
si_cap_at t orig_caps spec obj_id \<and>*
si_cap_at t orig_caps spec cspace_id \<and>*
si_cap_at t orig_caps spec vspace_id \<and>*
si_cap_at t orig_caps spec buffer_frame_id \<and>*
si_cap_at t orig_caps spec False obj_id \<and>*
si_cap_at t orig_caps spec False cspace_id \<and>*
si_cap_at t orig_caps spec False vspace_id \<and>*
si_cap_at t orig_caps spec False buffer_frame_id \<and>*
si_objects \<and>* R\<guillemotright> s\<rbrakk>
\<Longrightarrow>
\<guillemotleft>si_tcb_id \<mapsto>f root_tcb \<and>*
@ -291,7 +291,7 @@ lemma tcb_configure_pre:
apply (clarsimp simp: object_type_is_object)
apply (subst (asm) (2) default_cap_size_0 [where type=TcbType], simp)
apply (subst (asm) (2) default_cap_size_0 [where type=PageDirectoryType], simp)
apply (cut_tac type="FrameType sz" and sz="(object_size_bits obja)" and
apply (cut_tac type="FrameType sz" and sz="(object_size_bits obja)" and dev=False and
obj_id="{buffer_frame_kobj_id}" in default_cap_size_0, simp+)
apply sep_solve
done
@ -310,8 +310,9 @@ lemma well_formed_cnode_object_size_bits_eq2:
lemma default_cap_update_cap_object_non_cnode:
"\<lbrakk>cap_type cap = Some type; is_default_cap cap; cnode_cap_size cap \<le> 32;
type \<noteq> UntypedType; type \<noteq> AsidPoolType; type \<noteq> CNodeType; type \<noteq> IRQNodeType\<rbrakk>
\<Longrightarrow> default_cap type {obj_id} sz =
type \<noteq> UntypedType; type \<noteq> AsidPoolType; type \<noteq> CNodeType; type \<noteq> IRQNodeType;
dev = (is_device_cap cap)\<rbrakk>
\<Longrightarrow> default_cap type {obj_id} sz dev =
update_cap_object obj_id cap"
apply (frule (5) default_cap_update_cap_object [where obj_id=obj_id])
apply (subst default_cap_size_0, simp+)
@ -327,6 +328,15 @@ lemma sep_map_f_eq_tcb_fault_endpoint:
apply (clarsimp simp: update_tcb_fault_endpoint_def default_tcb_def)
done
lemma cnode_not_device[simp]:
"is_cnode_cap spec_cspace_cap \<Longrightarrow> \<not> is_device_cap spec_cspace_cap"
by (auto simp: is_device_cap_def split:cdl_cap.splits)
lemma well_formed_nondevice_cap_in_tcb:
"\<lbrakk>well_formed spec;cdl_objects spec obj_id = Some (Tcb spec_tcb)\<rbrakk> \<Longrightarrow>
\<not> cap_at (\<lambda>c. is_device_cap c) (obj_id, slot) spec"
by (simp add:well_formed_def)
lemma tcb_configure_post:
"\<lbrakk>well_formed spec; tcb_at obj_id spec;
cdl_objects spec obj_id = Some (Tcb spec_tcb);
@ -345,9 +355,9 @@ lemma tcb_configure_post:
cap_data spec_cspace_cap = cspace_cap_data;
cap_data spec_vspace_cap = vspace_cap_data;
cspace_cap = default_cap CNodeType {cspace_kobj_id} cnode_size;
vspace_cap = default_cap PageDirectoryType {vspace_kobj_id} 0;
buffer_frame_cap = default_cap buffer_frame_type {buffer_frame_kobj_id} 0;
cspace_cap = default_cap CNodeType {cspace_kobj_id} cnode_size False;
vspace_cap = default_cap PageDirectoryType {vspace_kobj_id} 0 False;
buffer_frame_cap = default_cap buffer_frame_type {buffer_frame_kobj_id} 0 False;
orig_caps obj_id = Some tcb_index;
orig_caps cspace_id = Some cspace_index;
@ -373,11 +383,11 @@ lemma tcb_configure_post:
si_cnode_id \<mapsto>f CNode (empty_cnode si_cnode_size) \<and>*
(si_cnode_id, unat seL4_CapInitThreadCNode) \<mapsto>c si_cnode_cap \<and>*
(si_cnode_id, offset tcb_index si_cnode_size) \<mapsto>c
default_cap TcbType {k_obj_id} 0 \<and>*
default_cap TcbType {k_obj_id} 0 False \<and>*
(si_cnode_id, offset cspace_index si_cnode_size) \<mapsto>c
default_cap CNodeType {cspace_kobj_id} cnode_size \<and>*
default_cap CNodeType {cspace_kobj_id} cnode_size False \<and>*
(si_cnode_id, offset vspace_index si_cnode_size) \<mapsto>c
default_cap PageDirectoryType {vspace_kobj_id} 0 \<and>*
default_cap PageDirectoryType {vspace_kobj_id} 0 False \<and>*
(si_cnode_id, offset buffer_frame_index si_cnode_size) \<mapsto>c buffer_frame_cap \<and>*
(si_cnode_id, unat seL4_CapIRQControl) \<mapsto>c IrqControlCap \<and>*
si_asid \<and>*
@ -385,8 +395,8 @@ lemma tcb_configure_post:
k_obj_id \<mapsto>f Tcb (update_tcb_fault_endpoint (cdl_tcb_fault_endpoint spec_tcb)
(default_tcb minBound)) \<and>*
(k_obj_id, tcb_cspace_slot) \<mapsto>c update_cap_data_det cspace_cap_data
(default_cap CNodeType {cspace_kobj_id} cnode_size) \<and>*
(k_obj_id, tcb_vspace_slot) \<mapsto>c default_cap PageDirectoryType {vspace_kobj_id} 0 \<and>*
(default_cap CNodeType {cspace_kobj_id} cnode_size False) \<and>*
(k_obj_id, tcb_vspace_slot) \<mapsto>c default_cap PageDirectoryType {vspace_kobj_id} 0 False \<and>*
(k_obj_id, tcb_ipcbuffer_slot) \<mapsto>c buffer_frame_cap \<and>*
(k_obj_id, tcb_replycap_slot) \<mapsto>c NullCap \<and>*
(k_obj_id, tcb_caller_slot) \<mapsto>c NullCap \<and>*
@ -394,10 +404,10 @@ lemma tcb_configure_post:
(k_obj_id, tcb_boundntfn_slot) \<mapsto>c NullCap \<and>*
k_obj_id \<mapsto>E Tcb (default_tcb minBound) \<and>* R\<guillemotright> s\<rbrakk>
\<Longrightarrow> \<guillemotleft>tcb_half_initialised spec t obj_id \<and>*
si_cap_at t orig_caps spec obj_id \<and>*
si_cap_at t orig_caps spec cspace_id \<and>*
si_cap_at t orig_caps spec vspace_id \<and>*
si_cap_at t orig_caps spec buffer_frame_id \<and>*
si_cap_at t orig_caps spec False obj_id \<and>*
si_cap_at t orig_caps spec False cspace_id \<and>*
si_cap_at t orig_caps spec False vspace_id \<and>*
si_cap_at t orig_caps spec False buffer_frame_id \<and>*
si_objects \<and>* R\<guillemotright> s"
apply (frule (1) well_formed_tcb_cspace_cap)
apply (frule (1) well_formed_tcb_vspace_cap)
@ -412,6 +422,7 @@ lemma tcb_configure_post:
apply (frule (1) well_formed_cap_object [where slot=tcb_vspace_slot], clarsimp)
apply (frule (1) well_formed_cap_object [where slot=tcb_ipcbuffer_slot],
clarsimp simp: cap_type_def)
apply (frule (1) well_formed_nondevice_cap_in_tcb[where slot = tcb_ipcbuffer_slot])
apply clarsimp
apply (frule (1) well_formed_types_match [where slot=tcb_cspace_slot], fastforce+)
apply (frule (1) well_formed_types_match [where slot=tcb_vspace_slot], fastforce+)
@ -444,6 +455,7 @@ lemma tcb_configure_post:
apply (subst cap_transform_update_cap_object
[where obj_id="cap_object spec_vspace_cap"],
(assumption|simp)+)
apply (clarsimp simp:cap_at_def opt_cap_def slots_of_def opt_object_def)
apply (subst cap_transform_update_cap_object
[where obj_id="cap_object spec_buffer_frame_cap"],
(assumption|simp)+)
@ -453,7 +465,7 @@ lemma tcb_configure_post:
assumption, assumption, simp+)
apply (subst default_cap_update_cap_object_non_cnode,
assumption, assumption, simp+)
apply (subst default_cap_update_cap_object_pd [THEN sym],
apply (subst default_cap_update_cap_object_pd [where dev =False,THEN sym],
assumption, assumption, simp+)
apply (cut_tac type = "FrameType sz"
and obj_id = buffer_frame_kobj_id
@ -464,7 +476,7 @@ lemma tcb_configure_post:
apply (clarsimp simp: sep_conj_assoc)
apply (clarsimp simp: object_type_simps)
apply (subst default_cap_size_0 [where type=TcbType], simp)
apply (cut_tac type=PageDirectoryType and sz="(object_size_bits obj)" and
apply (cut_tac type=PageDirectoryType and sz="(object_size_bits obj)" and dev = False and
obj_id="{vspace_kobj_id}" in default_cap_size_0, simp+)
apply (cut_tac type="FrameType sz" and sz="(object_size_bits obja)" and
obj_id="{buffer_frame_kobj_id}" in default_cap_size_0, simp+)
@ -654,20 +666,20 @@ lemma seL4_TCB_Configure_object_initialised_sep_helper:
fault_ep = cdl_tcb_fault_endpoint tcb\<rbrakk>
\<Longrightarrow>
\<lbrace>\<guillemotleft>object_empty spec t obj_id \<and>*
si_cap_at t orig_caps spec obj_id \<and>*
si_cap_at t orig_caps spec cspace_id \<and>*
si_cap_at t orig_caps spec vspace_id \<and>*
si_cap_at t orig_caps spec buffer_frame_id \<and>*
si_cap_at t orig_caps spec False obj_id \<and>*
si_cap_at t orig_caps spec False cspace_id \<and>*
si_cap_at t orig_caps spec False vspace_id \<and>*
si_cap_at t orig_caps spec False buffer_frame_id \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>
seL4_TCB_Configure tcb_index fault_ep priority
cspace_index cspace_root_data
vspace_index vspace_root_data
buffer_addr buffer_frame_index
\<lbrace>\<lambda>_. \<guillemotleft>tcb_half_initialised spec t obj_id \<and>*
si_cap_at t orig_caps spec obj_id \<and>*
si_cap_at t orig_caps spec cspace_id \<and>*
si_cap_at t orig_caps spec vspace_id \<and>*
si_cap_at t orig_caps spec buffer_frame_id \<and>*
si_cap_at t orig_caps spec False obj_id \<and>*
si_cap_at t orig_caps spec False cspace_id \<and>*
si_cap_at t orig_caps spec False vspace_id \<and>*
si_cap_at t orig_caps spec False buffer_frame_id \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (frule (1) well_formed_tcb_vspace_cap, elim exE conjE)
apply (frule (1) well_formed_tcb_ipcbuffer_cap, clarsimp)
@ -683,10 +695,10 @@ lemma seL4_TCB_Configure_object_initialised_sep_helper:
apply (cut_tac tcb="default_tcb minBound" and
cnode_cap = si_cspace_cap and
cnode_cap' = si_cnode_cap and
tcb_cap = "default_cap TcbType {k_obj_id} 0" and
cspace_cap = "default_cap CNodeType {cspace_kobj_id} (object_size_bits spec_cnode)" and
vspace_cap = "default_cap PageDirectoryType {vspace_kobj_id} 0" and
buffer_frame_cap = "default_cap (FrameType sz) {buffer_frame_kobj_id} 0" and
tcb_cap = "default_cap TcbType {k_obj_id} 0 False" and
cspace_cap = "default_cap CNodeType {cspace_kobj_id} (object_size_bits spec_cnode) False" and
vspace_cap = "default_cap PageDirectoryType {vspace_kobj_id} 0 False" and
buffer_frame_cap = "default_cap (FrameType sz) {buffer_frame_kobj_id} 0 False" and
cspace_root = cspace_index and
vspace_root = vspace_index and
buffer_frame_root = buffer_frame_index and
@ -730,20 +742,20 @@ lemma seL4_TCB_Configure_object_initialised_sep:
orig_caps buffer_frame_id = Some buffer_frame_index \<and>
\<guillemotleft>object_empty spec t obj_id \<and>*
si_cap_at t orig_caps spec obj_id \<and>*
si_cap_at t orig_caps spec cspace_id \<and>*
si_cap_at t orig_caps spec vspace_id \<and>*
si_cap_at t orig_caps spec buffer_frame_id \<and>*
si_cap_at t orig_caps spec False obj_id \<and>*
si_cap_at t orig_caps spec False cspace_id \<and>*
si_cap_at t orig_caps spec False vspace_id \<and>*
si_cap_at t orig_caps spec False buffer_frame_id \<and>*
si_objects \<and>* R\<guillemotright> s\<rbrace>
seL4_TCB_Configure tcb_index fault_ep priority
cspace_index cspace_root_data
vspace_index vspace_root_data
buffer_addr buffer_frame_index
\<lbrace>\<lambda>_. \<guillemotleft>tcb_half_initialised spec t obj_id \<and>*
si_cap_at t orig_caps spec obj_id \<and>*
si_cap_at t orig_caps spec cspace_id \<and>*
si_cap_at t orig_caps spec vspace_id \<and>*
si_cap_at t orig_caps spec buffer_frame_id \<and>*
si_cap_at t orig_caps spec False obj_id \<and>*
si_cap_at t orig_caps spec False cspace_id \<and>*
si_cap_at t orig_caps spec False vspace_id \<and>*
si_cap_at t orig_caps spec False buffer_frame_id \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (rule hoare_assume_pre)
apply (elim conjE)
@ -769,17 +781,17 @@ lemma init_tcb_sep':
cap_type tcb_ipcbuffer_cap = Some buffer_frame_type;
cap_object tcb_ipcbuffer_cap = buffer_frame_id\<rbrakk> \<Longrightarrow>
\<lbrace>\<guillemotleft>object_empty spec t obj_id \<and>*
si_cap_at t orig_caps spec obj_id \<and>*
si_cap_at t orig_caps spec cspace_id \<and>*
si_cap_at t orig_caps spec vspace_id \<and>*
si_cap_at t orig_caps spec buffer_frame_id \<and>*
si_cap_at t orig_caps spec False obj_id \<and>*
si_cap_at t orig_caps spec False cspace_id \<and>*
si_cap_at t orig_caps spec False vspace_id \<and>*
si_cap_at t orig_caps spec False buffer_frame_id \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>
init_tcb spec orig_caps obj_id
\<lbrace>\<lambda>_.\<guillemotleft>tcb_half_initialised spec t obj_id \<and>*
si_cap_at t orig_caps spec obj_id \<and>*
si_cap_at t orig_caps spec cspace_id \<and>*
si_cap_at t orig_caps spec vspace_id \<and>*
si_cap_at t orig_caps spec buffer_frame_id \<and>*
si_cap_at t orig_caps spec False obj_id \<and>*
si_cap_at t orig_caps spec False cspace_id \<and>*
si_cap_at t orig_caps spec False vspace_id \<and>*
si_cap_at t orig_caps spec False buffer_frame_id \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (clarsimp)
apply (subgoal_tac "\<exists>tcb. cdl_objects spec obj_id = Some (Tcb tcb)", clarsimp)
@ -802,11 +814,11 @@ lemma init_tcb_sep:
"\<lbrakk>well_formed spec; obj_id \<in> set tcbs; distinct tcbs;
set tcbs = {obj_id. tcb_at obj_id spec}\<rbrakk> \<Longrightarrow>
\<lbrace>\<guillemotleft>object_empty spec t obj_id \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>
init_tcb spec orig_caps obj_id
\<lbrace>\<lambda>_.\<guillemotleft>tcb_half_initialised spec t obj_id \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (frule well_formed_tcb_cspace_cap, fastforce)
apply (frule well_formed_tcb_vspace_cap, fastforce)
@ -856,11 +868,11 @@ lemma init_tcbs_sep_helper:
"\<lbrakk>well_formed spec; distinct tcbs;
set tcbs = {obj_id \<in> dom (cdl_objects spec). tcb_at obj_id spec}\<rbrakk> \<Longrightarrow>
\<lbrace>\<guillemotleft>objects_empty spec t {obj_id. tcb_at obj_id spec} \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>
mapM_x (init_tcb spec orig_caps) tcbs
\<lbrace>\<lambda>_.\<guillemotleft>tcbs_half_initialised spec t {obj_id. tcb_at obj_id spec} \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (clarsimp simp: objects_empty_def tcbs_half_initialised_def)
apply (rule hoare_name_pre_state)
@ -869,7 +881,7 @@ lemma init_tcbs_sep_helper:
mapM_x_set_sep [where
P="\<lambda>obj_id. object_empty spec t obj_id" and
Q="\<lambda>obj_id. tcb_half_initialised spec t obj_id" and
I="si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
I="si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects" and
xs="tcbs",
simplified sep_conj_assoc], simp+)
@ -877,18 +889,18 @@ lemma init_tcbs_sep_helper:
done
lemma is_tcb_default_cap:
"is_tcb obj \<Longrightarrow> default_cap (object_type obj) {obj_id} sz = TcbCap obj_id"
"is_tcb obj \<Longrightarrow> default_cap (object_type obj) {obj_id} sz dev = TcbCap obj_id"
by (clarsimp simp: default_cap_def is_tcb_obj_type)
lemma configure_tcb_sep:
"\<lbrace>\<guillemotleft>tcb_half_initialised spec t obj_id \<and>*
si_cap_at t orig_caps spec obj_id \<and>*
si_cap_at t orig_caps spec False obj_id \<and>*
si_objects \<and>* R\<guillemotright> and
K(well_formed spec \<and> obj_id \<in> set tcbs \<and> distinct tcbs \<and>
set tcbs = {obj_id. tcb_at obj_id spec})\<rbrace>
configure_tcb spec orig_caps obj_id
\<lbrace>\<lambda>_.\<guillemotleft>tcb_half_initialised spec t obj_id \<and>*
si_cap_at t orig_caps spec obj_id \<and>*
si_cap_at t orig_caps spec False obj_id \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (rule hoare_gen_asm)
apply (clarsimp simp: configure_tcb_def object_initialised_def tcb_half_initialised_def object_initialised_general_def
@ -910,20 +922,20 @@ lemma configure_tcb_sep:
lemma configure_tcbs_sep:
"\<lbrace>\<guillemotleft>tcbs_half_initialised spec t {obj_id. tcb_at obj_id spec} \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright> and
K(well_formed spec \<and> distinct tcbs \<and>
set tcbs = {obj_id \<in> dom (cdl_objects spec). tcb_at obj_id spec})\<rbrace>
mapM_x (configure_tcb spec orig_caps) tcbs
\<lbrace>\<lambda>_.\<guillemotleft>tcbs_half_initialised spec t {obj_id. tcb_at obj_id spec} \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (rule hoare_gen_asm)
apply (clarsimp simp: objects_empty_def tcbs_half_initialised_def)
apply (rule mapM_x_set_sep' [where
P="\<lambda>obj_id. tcb_half_initialised spec t obj_id" and
Q="\<lambda>obj_id. tcb_half_initialised spec t obj_id" and
I="si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
I="si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects" and
xs="tcbs" and
X="{obj_id. tcb_at obj_id spec}" and
@ -934,7 +946,7 @@ lemma configure_tcbs_sep:
and xs = "{obj_id. real_object_at obj_id spec}"
and P = "tcb_half_initialised spec t obj_id \<and>* si_objects"
and Q = "tcb_half_initialised spec t obj_id \<and>* si_objects"
and I = "si_cap_at t orig_caps spec"
and I = "si_cap_at t orig_caps spec False"
and R=R
in sep_set_conj_map_singleton_wp [simplified], simp_all add: object_at_real_object_at)
apply (wp sep_wp: configure_tcb_sep [where t=t and tcbs=tcbs], (simp|sep_solve)+)
@ -942,12 +954,12 @@ lemma configure_tcbs_sep:
lemma init_tcbs_sep:
"\<lbrace>\<guillemotleft>objects_empty spec t {obj_id. tcb_at obj_id spec} \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright> and
K(well_formed spec \<and> set obj_ids = dom (cdl_objects spec) \<and> distinct obj_ids)\<rbrace>
init_tcbs spec orig_caps obj_ids
\<lbrace>\<lambda>_.\<guillemotleft>tcbs_half_initialised spec t {obj_id. tcb_at obj_id spec} \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (rule hoare_gen_asm)
apply (clarsimp simp: init_tcbs_def)

View File

@ -48,7 +48,7 @@ lemma cap_has_object_simps [simp]:
"cap_has_object (PendingSyncSendCap obj_id badge a b c)"
"cap_has_object (PendingSyncRecvCap obj_id d)"
"cap_has_object (PendingNtfnRecvCap obj_id)"
"cap_has_object (FrameCap obj_id rights sz type maddr)"
"cap_has_object (FrameCap dev obj_id rights sz type maddr)"
"cap_has_object (PageTableCap obj_id cdl_frame_cap_type maddr)"
"cap_has_object (PageDirectoryCap obj_id cdl_frame_cap_type asid)"
"cap_has_object (AsidPoolCap obj_id as)"
@ -169,12 +169,12 @@ lemma valid_vm_rights_rw:
lemma seL4_Section_Map_object_initialised_sep:
"\<lbrace>\<guillemotleft>object_slot_empty spec t spec_pd_ptr (unat (vaddr >> 20)) \<and>*
(si_cnode_id , offset sel4_section si_cnode_size) \<mapsto>c (FrameCap section_ptr {AllowRead, AllowWrite} n Real None) \<and>*
(si_cnode_id , offset sel4_section si_cnode_size) \<mapsto>c (FrameCap dev section_ptr {AllowRead, AllowWrite} n Real None) \<and>*
(si_cnode_id , offset sel4_pd si_cnode_size) \<mapsto>c (PageDirectoryCap pd_ptr Real None) \<and>*
si_objects \<and>* R\<guillemotright> and
K(pd_at spec_pd_ptr spec \<and>
opt_cap (spec_pd_ptr,unat (vaddr >> 20)) spec
= Some (FrameCap spec_section_ptr (validate_vm_rights rights) n Fake None) \<and>
= Some (FrameCap False spec_section_ptr (validate_vm_rights rights) n Fake None) \<and>
sel4_section < 2 ^ si_cnode_size \<and>
sel4_pd < 2 ^ si_cnode_size \<and>
@ -185,7 +185,7 @@ lemma seL4_Section_Map_object_initialised_sep:
t spec_section_ptr = Some section_ptr)\<rbrace>
seL4_Page_Map sel4_section sel4_pd vaddr rights vmattribs
\<lbrace>\<lambda>rv. \<guillemotleft>object_slot_initialised spec t spec_pd_ptr (unat (vaddr >> 20)) \<and>*
(si_cnode_id , offset sel4_section si_cnode_size) \<mapsto>c (FrameCap section_ptr {AllowRead, AllowWrite} n Real None) \<and>*
(si_cnode_id , offset sel4_section si_cnode_size) \<mapsto>c (FrameCap dev section_ptr {AllowRead, AllowWrite} n Real None) \<and>*
(si_cnode_id , offset sel4_pd si_cnode_size) \<mapsto>c (PageDirectoryCap pd_ptr Real None) \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (rule hoare_gen_asm)
@ -227,7 +227,7 @@ lemma well_formed_cap_obj_match_frame:
opt_cap cap_ref spec = Some cap;
cap_has_object cap;
cap_object cap = ptr\<rbrakk>
\<Longrightarrow> \<exists>is_real. cap = FrameCap ptr (validate_vm_rights (cap_rights cap)) (cdl_frame_size_bits frame) is_real None"
\<Longrightarrow> \<exists>is_real dev. cap = FrameCap dev ptr (validate_vm_rights (cap_rights cap)) (cdl_frame_size_bits frame) is_real None"
apply (case_tac cap_ref, clarsimp)
apply (frule (1) well_formed_well_formed_cap', simp)
apply (frule (3) well_formed_types_match)
@ -251,8 +251,8 @@ lemma well_formed_cap_obj_match_pt:
done
lemma sep_caps_at_split: "a \<in> A \<Longrightarrow>
si_caps_at t orig_caps spec A = (
si_cap_at t orig_caps spec a \<and>* si_caps_at t orig_caps spec (A - {a}))"
si_caps_at t orig_caps spec dev A = (
si_cap_at t orig_caps spec dev a \<and>* si_caps_at t orig_caps spec dev (A - {a}))"
apply (simp add:si_caps_at_def)
apply (subst sep.setprod.union_disjoint [where A = "{a}", simplified, symmetric])
apply simp
@ -263,17 +263,17 @@ lemma map_page_in_pd_wp:
"\<lbrakk>well_formed spec; pd_at spec_pd_ptr spec\<rbrakk>
\<Longrightarrow>
\<lbrace>\<guillemotleft>object_slot_empty spec t spec_pd_ptr (unat (shiftr vaddr 20)) \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright> and
K ((pt_at spec_pt_section_ptr spec \<longrightarrow>
opt_cap (spec_pd_ptr, unat (shiftr vaddr 20)) spec = Some (PageTableCap spec_pt_section_ptr Fake None)) \<and>
(frame_at spec_pt_section_ptr spec \<longrightarrow>
(\<exists>n. (n = 20 \<or> n = 24) \<and>
opt_cap (spec_pd_ptr, unat (shiftr vaddr 20)) spec =
Some (FrameCap spec_pt_section_ptr (validate_vm_rights rights) n Fake None))))\<rbrace>
Some (FrameCap False spec_pt_section_ptr (validate_vm_rights rights) n Fake None))))\<rbrace>
map_page spec orig_caps spec_pt_section_ptr spec_pd_ptr rights vaddr
\<lbrace>\<lambda>_. \<guillemotleft>object_slot_initialised spec t spec_pd_ptr (unat (shiftr vaddr 20)) \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (rule hoare_gen_asm)
apply (clarsimp simp: map_page_def dest!:domE)
@ -351,7 +351,7 @@ lemma well_formed_frame_in_pt:
lemma well_formed_frame_in_pd:
"\<lbrakk>well_formed spec; opt_cap (pd, pt_slot) spec = Some frame_cap; pd_at pd spec; is_frame_cap frame_cap \<rbrakk> \<Longrightarrow>
(\<exists>sz. cap_type frame_cap = Some (FrameType sz) \<and> (sz = 20 \<or> sz = 24)) \<and> is_fake_vm_cap frame_cap"
(\<exists>sz. cap_type frame_cap = Some (FrameType sz) \<and> (sz = 20 \<or> sz = 24)) \<and> is_fake_vm_cap frame_cap \<and> \<not> is_device_cap frame_cap"
apply (clarsimp simp: well_formed_def object_at_def)
apply (drule_tac x = pd in spec)
apply (clarsimp simp: well_formed_vspace_def opt_cap_def slots_of_def opt_object_def
@ -360,18 +360,23 @@ lemma well_formed_frame_in_pd:
apply (drule_tac x = frame_cap in spec)
apply (clarsimp simp: is_fake_pt_cap_def cap_type_def
split: cdl_cap.splits)
apply (clarsimp simp:cap_at_def opt_cap_def slots_of_def opt_object_def
simp del:split_paired_All)
apply (drule_tac x = pd in spec)
apply (drule_tac x = pt_slot in spec)
apply fastforce
done
lemma map_page_directory_slot_wp:
"\<lbrace>\<guillemotleft>object_slot_empty spec t spec_pd_ptr slot \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright> and K(
well_formed spec \<and>
slot < 0x1000 \<and>
pd_at spec_pd_ptr spec)\<rbrace>
map_page_directory_slot spec orig_caps spec_pd_ptr slot
\<lbrace>\<lambda>_. \<guillemotleft>object_slot_initialised spec t spec_pd_ptr slot \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (rule hoare_gen_asm)
apply (clarsimp simp: map_page_directory_slot_def)
@ -410,6 +415,7 @@ lemma map_page_directory_slot_wp:
apply simp+
apply (simp add:object_at_def)
apply (case_tac a,simp_all)
apply clarsimp
apply (clarsimp simp:is_fake_vm_cap_def split:cdl_cap.split_asm cdl_frame_cap_type.split_asm)
apply (frule object_slot_empty_initialised_NullCap [where obj_id=spec_pd_ptr and slot=slot and t=t])
apply (clarsimp simp: object_at_def object_type_is_object)
@ -442,13 +448,13 @@ lemma well_formed_pt_slot_limited:
lemma map_map_page_directory_slot_wp':
"\<lbrace>\<guillemotleft>object_slots_empty spec t obj_id \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright> and K(
well_formed spec \<and> pd_at obj_id spec)\<rbrace>
mapM_x (map_page_directory_slot spec orig_caps obj_id)
(slots_of_list spec obj_id)
\<lbrace>\<lambda>_. \<guillemotleft>object_slots_initialised spec t obj_id \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (rule hoare_gen_asm)
apply (subst object_slots_empty_decomp, simp+)
@ -456,7 +462,7 @@ lemma map_map_page_directory_slot_wp':
apply (subst object_empty_slots_empty_initialised, simp)
apply (simp add:sep_conj_assoc)
apply (rule mapM_x_set_sep' [where I = "object_empty_slots_initialised spec t obj_id \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_objects"
and xs = "slots_of_list spec obj_id", unfolded sep_conj_assoc])
apply clarsimp
@ -515,14 +521,14 @@ lemma map_object_empty_initialised_frame:
lemma map_map_page_directory_slot_wp:
"\<lbrace>\<guillemotleft>object_empty spec t obj_id \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright> and K(
well_formed spec \<and> pd_at obj_id spec)\<rbrace>
mapM_x (map_page_directory_slot spec orig_caps obj_id)
(slots_of_list spec obj_id)
\<lbrace>\<lambda>_.
\<guillemotleft>object_initialised spec t obj_id \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (rule hoare_gen_asm)
apply clarsimp
@ -541,7 +547,7 @@ lemma set_asid_rewrite:
orig_caps obj_id = Some pd_offset;
pd_offset < 2 ^ si_cnode_size;
t obj_id = Some kobj_id\<rbrakk> \<Longrightarrow>
((si_cnode_id, unat pd_offset) \<mapsto>c default_cap (object_type pd) {kobj_id} (object_size_bits pd) \<and>*
((si_cnode_id, unat pd_offset) \<mapsto>c default_cap (object_type pd) {kobj_id} (object_size_bits pd) dev \<and>*
si_objects \<and>* R)
=
((si_tcb_id, tcb_pending_op_slot) \<mapsto>c RunningCap \<and>*
@ -562,12 +568,12 @@ lemma set_asid_rewrite:
done
lemma set_asid_wp:
"\<lbrace>\<guillemotleft>si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
"\<lbrace>\<guillemotleft>si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright> and K(
well_formed spec \<and>
pd_at obj_id spec)\<rbrace>
set_asid spec orig_caps obj_id
\<lbrace>\<lambda>rv. \<guillemotleft>si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
\<lbrace>\<lambda>rv. \<guillemotleft>si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (rule hoare_gen_asm, clarsimp)
apply (frule (1) object_at_real_object_at)
@ -598,16 +604,16 @@ lemma set_asid_wp:
lemma map_map_page_directory_wp':
"\<lbrace>\<guillemotleft>(\<And>* pd_id \<in> {pd_id. pd_at pd_id spec}. object_empty spec t pd_id) \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright> and
K(well_formed spec \<and> set obj_ids = dom (cdl_objects spec) \<and> distinct obj_ids)\<rbrace>
mapM_x (map_page_directory spec orig_caps)
[obj\<leftarrow>obj_ids. pd_at obj spec]
\<lbrace>\<lambda>rv. \<guillemotleft>(\<And>* pd_id \<in> {pd_id. pd_at pd_id spec}. object_initialised spec t pd_id) \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (rule hoare_gen_asm)
apply (rule mapM_x_set_sep' [where I = "si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>* si_objects",
apply (rule mapM_x_set_sep' [where I = "si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>* si_objects",
simplified sep_conj_assoc])
apply simp
apply (fastforce simp: object_at_def)
@ -618,7 +624,7 @@ lemma map_map_page_directory_wp':
lemma map_map_page_directory_wp:
"\<lbrace>\<guillemotleft>(\<And>* pt_id \<in> {pt_id. pt_at pt_id spec}. object_empty spec t pt_id) \<and>*
(\<And>* pd_id \<in> {pd_id. pd_at pd_id spec}. object_empty spec t pd_id) \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright> and K(
well_formed spec \<and> set obj_ids = dom (cdl_objects spec) \<and> distinct obj_ids)\<rbrace>
mapM_x (map_page_directory spec orig_caps)
@ -626,7 +632,7 @@ lemma map_map_page_directory_wp:
\<lbrace>\<lambda>_.
\<guillemotleft>(\<And>* pt_id \<in> {pt_id. pt_at pt_id spec}. object_empty spec t pt_id) \<and>*
(\<And>* pd_id \<in> {pd_id. pd_at pd_id spec}. object_initialised spec t pd_id) \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec dev {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (rule hoare_gen_asm)
apply (wp sep_wp: map_map_page_directory_wp' [where t=t], simp)
@ -774,7 +780,7 @@ lemma cdl_lookup_pd_slot_compute:
done
lemma well_formed_frame_valid:
"\<lbrakk>well_formed spec; opt_cap cap_ref spec = Some (FrameCap ptr seta sz real option)\<rbrakk>
"\<lbrakk>well_formed spec; opt_cap cap_ref spec = Some (FrameCap dev ptr seta sz real option)\<rbrakk>
\<Longrightarrow> validate_vm_rights seta = seta"
apply (case_tac cap_ref, clarsimp)
apply (frule (1) well_formed_well_formed_cap', simp)
@ -786,13 +792,18 @@ lemma empty_cap_map_NullCap:
"pt_slot < 2 ^n \<Longrightarrow> empty_cap_map n pt_slot = Some NullCap"
by (simp add:empty_cap_map_def)
(* FIXME: current cdl_init does not consider device caps *)
lemma well_formed_opt_cap_nondevice:
"\<lbrakk>well_formed spec; opt_cap slot spec = Some cap\<rbrakk> \<Longrightarrow> \<not> is_device_cap cap"
by (simp add:well_formed_def cap_at_def del:split_paired_All)
(***********************
* Mapping page tables *
***********************)
lemma map_page_in_pt_sep:
"\<lbrace>\<guillemotleft>object_slot_empty spec t (cap_object pt_cap) pt_slot \<and>*
object_slot_initialised spec t obj_id pd_slot \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright> and K(
well_formed spec \<and>
pd_at obj_id spec \<and>
@ -808,7 +819,7 @@ lemma map_page_in_pt_sep:
(of_nat pt_slot << small_frame_size))
\<lbrace>\<lambda>_. \<guillemotleft>object_slot_initialised spec t (cap_object pt_cap) pt_slot \<and>*
object_slot_initialised spec t obj_id pd_slot \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>*
R\<guillemotright>\<rbrace>"
@ -846,10 +857,11 @@ lemma map_page_in_pt_sep:
apply (clarsimp simp: cap_type_def is_fake_vm_cap_def cap_object_simps
cap_transform_def is_fake_pt_cap_def
split: cdl_cap.split_asm cdl_frame_cap_type.split_asm)
apply (frule(1) well_formed_opt_cap_nondevice)
apply (rule hoare_pre)
apply (rule hoare_strengthen_post)
apply (rule_tac pd_ptr = kobj_id in seL4_Page_Map_wp[where cnode_cap = si_cspace_cap
and root_size = si_cnode_size
and root_size = si_cnode_size and dev = False
and rights = "{AllowRead, AllowWrite}"])
apply (simp add:word_bits_def guard_equal_si_cspace_cap)+
apply (simp add:si_objects_def sep_conj_assoc pt_slot_compute)
@ -883,6 +895,7 @@ lemma map_page_in_pt_sep:
apply (sep_drule sep_map_c_sep_map_s)
apply (simp add: object_default_state_def object_type_def
default_object_def object_slots_def empty_cap_map_NullCap)
apply (simp add:fun_upd_def)
apply sep_solve
apply clarsimp
apply (clarsimp simp: is_pt_def object_at_def
@ -895,7 +908,7 @@ lemma map_page_in_pt_sep:
lemma map_page_table_slot_wp:
"\<lbrace>\<guillemotleft>object_slot_empty spec t (cap_object page_cap) pt_slot \<and>*
object_slot_initialised spec t obj_id pd_slot \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright> and K(
well_formed spec \<and>
pd_at obj_id spec \<and>
@ -910,7 +923,7 @@ lemma map_page_table_slot_wp:
pt_slot
\<lbrace>\<lambda>_. \<guillemotleft>object_slot_initialised spec t (cap_object page_cap) pt_slot \<and>*
object_slot_initialised spec t obj_id pd_slot \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (clarsimp simp: map_page_table_slot_def)
apply (wp map_page_in_pt_sep)
@ -929,7 +942,7 @@ lemma map_page_table_slots_wp'':
opt_cap (obj_id, pd_slot) spec = Some page_cap\<rbrakk> \<Longrightarrow>
\<lbrace>\<guillemotleft>object_slots_empty spec t (cap_ref_object (obj_id, pd_slot) spec) \<and>*
object_slot_initialised spec t obj_id pd_slot \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>
mapM_x (map_page_table_slot spec orig_caps obj_id
(cap_object page_cap)
@ -937,7 +950,7 @@ lemma map_page_table_slots_wp'':
(slots_of_list spec (cap_object page_cap))
\<lbrace>\<lambda>_. \<guillemotleft>object_slots_initialised spec t (cap_ref_object (obj_id, pd_slot) spec) \<and>*
object_slot_initialised spec t obj_id pd_slot \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (frule well_formed_distinct_slots_of_list [where obj_id="cap_ref_object (obj_id, pd_slot) spec"])
apply (frule well_formed_finite [where obj_id="cap_ref_object (obj_id, pd_slot) spec"])
@ -947,7 +960,7 @@ lemma map_page_table_slots_wp'':
apply (clarsimp simp: sep_conj_assoc cap_ref_object_def cap_at_def)
apply (rule mapM_x_set_sep' [where I = "object_empty_slots_initialised spec t (cap_object page_cap) \<and>*
object_slot_initialised spec t obj_id pd_slot \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects", unfolded sep_conj_assoc], simp, clarsimp)
apply (frule is_fake_pt_cap_is_pt_cap)
apply (frule (1) well_formed_cap_object, clarsimp)
@ -967,7 +980,7 @@ lemma map_page_table_slots_wp'':
lemma map_page_table_slots_wp':
"\<lbrace>\<guillemotleft>object_empty spec t (cap_ref_object (obj_id, pd_slot) spec) \<and>*
object_slot_initialised spec t obj_id pd_slot \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>*
R\<guillemotright> and K(
well_formed spec \<and>
@ -979,7 +992,7 @@ lemma map_page_table_slots_wp':
(slots_of_list spec (cap_object page_cap))
\<lbrace>\<lambda>_. \<guillemotleft>object_initialised spec t (cap_ref_object (obj_id, pd_slot) spec) \<and>*
object_slot_initialised spec t obj_id pd_slot \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>*
R\<guillemotright>\<rbrace>"
apply (rule hoare_gen_asm)
@ -1000,7 +1013,7 @@ lemma map_page_table_slots_wp':
lemma map_page_table_slots_wp:
"\<lbrace>\<guillemotleft>object_empty spec t (cap_ref_object (obj_id, pd_slot) spec) \<and>*
object_slot_initialised spec t obj_id pd_slot \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright> and K(
well_formed spec \<and>
pd_at obj_id spec \<and>
@ -1009,7 +1022,7 @@ lemma map_page_table_slots_wp:
map_page_table_slots spec orig_caps obj_id pd_slot
\<lbrace>\<lambda>_. \<guillemotleft>object_initialised spec t (cap_ref_object (obj_id, pd_slot) spec) \<and>*
object_slot_initialised spec t obj_id pd_slot \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (rule hoare_gen_asm)
apply (clarsimp simp: map_page_table_slots_def)
@ -1031,7 +1044,7 @@ lemma map_page_directory_page_tables_wp:
then (object_empty spec t (cap_ref_object (obj_id, slot) spec))
else \<box>) \<and>*
object_initialised spec t obj_id \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright> and
K(well_formed spec \<and> pd_at obj_id spec)\<rbrace>
map_page_directory_page_tables spec orig_caps obj_id
@ -1040,13 +1053,13 @@ lemma map_page_directory_page_tables_wp:
then (object_initialised spec t (cap_ref_object (obj_id, slot) spec))
else \<box>) \<and>*
object_initialised spec t obj_id \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (rule hoare_gen_asm, clarsimp)
apply (frule well_formed_distinct_slots_of_list [where obj_id=obj_id])
apply (clarsimp simp: map_page_directory_page_tables_def)
apply (rule mapM_x_set_sep'[where I="object_initialised spec t obj_id \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects",
unfolded sep_conj_assoc,simplified], simp+)
apply clarsimp
@ -1071,7 +1084,7 @@ lemma map_map_page_directory_page_tables_wp'':
obj_id = cap_object cap)}.
object_empty spec t obj_id) \<and>*
(\<And>* obj_id \<in> {obj_id. pd_at obj_id spec}. object_initialised spec t obj_id) \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright> and
K (well_formed spec \<and> set obj_ids = dom (cdl_objects spec) \<and> distinct obj_ids)\<rbrace>
mapM_x (map_page_directory_page_tables spec orig_caps)
@ -1081,7 +1094,7 @@ lemma map_map_page_directory_page_tables_wp'':
obj_id = cap_object cap)}.
object_initialised spec t obj_id) \<and>*
(\<And>* obj_id \<in> {obj_id. pd_at obj_id spec}. object_initialised spec t obj_id) \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (rule hoare_gen_asm, clarsimp)
apply (subst fake_pt_cap_rewrite, assumption)+
@ -1090,7 +1103,7 @@ lemma map_map_page_directory_page_tables_wp'':
apply (rule mapM_x_set_sep'[where
I = "(\<And>* obj_id \<in> {obj_id. pd_at obj_id spec}.
object_initialised spec t obj_id) \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>* si_objects",
si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>* si_objects",
unfolded sep_conj_assoc,simplified])
apply simp
apply simp
@ -1106,14 +1119,14 @@ lemma map_map_page_directory_page_tables_wp'':
lemma map_map_page_directory_page_tables_wp:
"\<lbrace>\<guillemotleft>(\<And>* obj_id \<in> {obj_id. pt_at obj_id spec}. object_empty spec t obj_id) \<and>*
(\<And>* obj_id \<in> {obj_id. pd_at obj_id spec}. object_initialised spec t obj_id) \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright> and
K(well_formed spec \<and> set obj_ids = dom (cdl_objects spec) \<and> distinct obj_ids)\<rbrace>
mapM_x (map_page_directory_page_tables spec orig_caps)
[obj\<leftarrow>obj_ids. pd_at obj spec]
\<lbrace>\<lambda>_. \<guillemotleft>(\<And>* obj_id \<in> {obj_id. pt_at obj_id spec}. object_initialised spec t obj_id) \<and>*
(\<And>* obj_id \<in> {obj_id. pd_at obj_id spec}. object_initialised spec t obj_id) \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (rule hoare_gen_asm, clarsimp)
apply (subst sep_map_set_conj_restrict [where xs="{obj_id. pt_at obj_id spec}" and
@ -1126,12 +1139,12 @@ lemma map_map_page_directory_page_tables_wp:
lemma init_vspace_sep:
"\<lbrace>\<guillemotleft>objects_empty spec t {obj_id. table_at obj_id spec} \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright> and
K(well_formed spec \<and> set obj_ids = dom (cdl_objects spec) \<and> distinct obj_ids)\<rbrace>
init_vspace spec orig_caps obj_ids
\<lbrace>\<lambda>_. \<guillemotleft>objects_initialised spec t {obj_id. table_at obj_id spec} \<and>*
si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (rule hoare_gen_asm, clarsimp)
apply (clarsimp simp: objects_empty_def objects_initialised_def)
@ -1144,10 +1157,10 @@ lemma init_vspace_sep:
done
lemma init_pd_asids_sep:
"\<lbrace>\<guillemotleft>si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
"\<lbrace>\<guillemotleft>si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright> and K(well_formed spec)\<rbrace>
init_pd_asids spec orig_caps obj_ids
\<lbrace>\<lambda>_. \<guillemotleft>si_caps_at t orig_caps spec {obj_id. real_object_at obj_id spec} \<and>*
\<lbrace>\<lambda>_. \<guillemotleft>si_caps_at t orig_caps spec False {obj_id. real_object_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (rule hoare_gen_asm)
apply (clarsimp simp: init_pd_asids_def)

View File

@ -458,7 +458,7 @@ lemma is_default_cap_cap_transform [simp]:
lemma default_cap_cap_transform:
"\<lbrakk>is_default_cap cap; well_formed_cap cap; t (cap_object cap) = Some obj_id;
cap_type cap = Some type; type \<noteq> IRQNodeType\<rbrakk>
\<Longrightarrow> default_cap type {obj_id} (cnode_cap_size cap) = cap_transform t cap"
\<Longrightarrow> default_cap type {obj_id} (cnode_cap_size cap) (is_device_cap cap) = cap_transform t cap"
by (clarsimp simp: is_default_cap_def default_cap_def cap_transform_def cap_type_def
well_formed_cap_def cap_has_object_def
update_cap_object_def split: cdl_cap.splits)+
@ -470,20 +470,29 @@ lemma cap_transform_update_cap_object:
cap_object_def cap_has_object_def
split: cdl_cap.splits)
lemma is_default_cap_def2:
"is_default_cap cap =
((\<exists>type. cap_type cap = Some type \<and> cap = default_cap type (cap_objects cap) (cnode_cap_size cap) (is_device_cap cap)) \<or>
is_irqhandler_cap cap)"
apply (clarsimp simp:is_default_cap_def)
apply (case_tac cap)
apply (auto simp:default_cap_def cap_type_def)
done
lemma default_cap_update_cap_object:
"\<lbrakk>is_default_cap cap; cap_type cap = Some type; cnode_cap_size cap \<le> 32;
type \<noteq> UntypedType; type \<noteq> AsidPoolType; type \<noteq> IRQNodeType\<rbrakk>
\<Longrightarrow> default_cap type {obj_id} (cnode_cap_size cap) = update_cap_object obj_id cap"
\<Longrightarrow> default_cap type {obj_id} (cnode_cap_size cap) (is_device_cap cap) = update_cap_object obj_id cap"
apply (subst default_cap_cap_transform, simp_all)
apply (frule (1) default_cap_well_formed_cap2 [where obj_ids="cap_objects cap"
and sz = "(cnode_cap_size cap)"], simp+)
apply (fastforce simp: is_default_cap_def)
and sz = "(cnode_cap_size cap)" and dev = "is_device_cap cap"], simp+)
apply (fastforce simp: is_default_cap_def2)
apply (subst cap_transform_update_cap_object, simp_all)
done
lemma default_cap_update_cap_object_pd:
"\<lbrakk>is_pd_cap cap; \<not> vm_cap_has_asid cap; \<not> is_fake_vm_cap cap\<rbrakk>
\<Longrightarrow> default_cap PageDirectoryType {obj_id} (cnode_cap_size cap) = update_cap_object obj_id cap"
\<Longrightarrow> default_cap PageDirectoryType {obj_id} (cnode_cap_size cap) dev = update_cap_object obj_id cap"
by (clarsimp simp: default_cap_def update_cap_object_def cap_type_def
vm_cap_has_asid_def is_fake_vm_cap_def not_Some_eq_tuple
split: cdl_cap.splits cdl_frame_cap_type.splits)

View File

@ -201,6 +201,7 @@ lemma small_one:
distinct_sets (map cap_free_ids untyped_caps);
list_all is_full_untyped_cap untyped_caps;
list_all well_formed_untyped_cap untyped_caps;
list_all (\<lambda>c. \<not> is_device_cap c) untyped_caps;
bi_untypes bootinfo = (ustart, uend);
bi_free_slots bootinfo = (fstart, fend);
unat ustart < 2 ^ si_cnode_size;
@ -221,7 +222,7 @@ lemma small_one:
\<guillemotleft>objects_initialised spec t {obj_id. real_object_at obj_id spec} \<and>*
irqs_initialised spec t (used_irqs spec) \<and>*
(\<And>* cptr\<in>set (take (card (dom (cdl_objects spec))) free_cptrs). (si_cnode_id, unat cptr) \<mapsto>c NullCap) \<and>*
si_caps_at t dup_caps spec {obj_id. cnode_or_tcb_at obj_id spec} \<and>*
si_caps_at t dup_caps spec False {obj_id. cnode_or_tcb_at obj_id spec} \<and>*
si_objects \<and>*
si_objects_extra_caps (dom (cdl_objects spec))
(free_cptrs :: 32 word list)
@ -240,16 +241,16 @@ lemma small_one:
apply (rule hoare_ex_wp, rename_tac t, rule_tac t=t in init_tcbs_sep [sep_wandise])
apply (rule hoare_ex_wp, rename_tac t, rule_tac t=t in init_vspace_sep [sep_wandise])
apply (rule hoare_ex_wp, rename_tac t, rule_tac t=t in init_pd_asids_sep [sep_wandise])
apply (rule hoare_ex_wp, rename_tac t, rule_tac t=t in init_irqs_sep [sep_wandise])
apply (rule hoare_ex_wp, rename_tac t, rule_tac t=t and
apply (rule hoare_ex_wp, rename_tac t, rule_tac t=t and dev=False in init_irqs_sep [sep_wandise])
apply (rule hoare_ex_wp, rename_tac t, rule_tac t=t and dev=False and
untyped_cptrs = "[ustart .e. uend - 1]" and
free_cptrs_orig = "[fstart .e. fend - 1]" in duplicate_caps_sep [sep_wandise])
apply (rule create_irq_caps_sep [sep_wandise,
apply (rule create_irq_caps_sep [where dev = False,sep_wandise,
where free_cptrs_orig = "[fstart .e. fend - 1]"
and untyped_cptrs = "[ustart .e. uend - 1]"
and orig_caps = "map_of (zip [obj\<leftarrow>obj_ids. real_object_at obj spec] [fstart .e. fend - 1])"
and spec = spec])
apply (wp sep_wp: create_objects_sep [where untyped_caps = untyped_caps])
apply (wp sep_wp: create_objects_sep [where untyped_caps = untyped_caps and dev = False])
apply (wp sep_wp: parse_bootinfo_sep [where fstart = fstart
and fend = fend
and ustart = ustart
@ -342,4 +343,3 @@ lemma sys_init_paper:
done
end

View File

@ -292,12 +292,12 @@ lemma offset_seL4_CapIRQControl [simp]:
* This predicate can be used for spec objects, or the duplicated cnode caps.
*)
definition si_cap_at :: "(cdl_object_id \<Rightarrow> cdl_object_id option) \<Rightarrow>
(cdl_object_id \<Rightarrow> cdl_cptr option) \<Rightarrow> cdl_state \<Rightarrow>
(cdl_object_id \<Rightarrow> cdl_cptr option) \<Rightarrow> cdl_state \<Rightarrow> bool \<Rightarrow>
cdl_object_id \<Rightarrow> sep_pred"
where
"si_cap_at t si_caps spec obj_id \<equiv> \<lambda>s.
"si_cap_at t si_caps spec dev obj_id \<equiv> \<lambda>s.
\<exists>cap_ptr slot obj kobj_id.
((si_cnode_id, slot) \<mapsto>c default_cap (object_type obj) {kobj_id} (object_size_bits obj)) s \<and>
((si_cnode_id, slot) \<mapsto>c default_cap (object_type obj) {kobj_id} (object_size_bits obj) dev) s \<and>
si_caps obj_id = Some cap_ptr \<and>
unat cap_ptr = slot \<and>
cap_ptr < 2 ^ si_cnode_size \<and>
@ -315,11 +315,11 @@ where
cap_ptr < 2 ^ si_cnode_size"
definition si_caps_at :: "(cdl_object_id \<Rightarrow> cdl_object_id option) \<Rightarrow>
(cdl_object_id \<Rightarrow> cdl_cptr option) \<Rightarrow> cdl_state \<Rightarrow>
(cdl_object_id \<Rightarrow> cdl_cptr option) \<Rightarrow> cdl_state \<Rightarrow> bool \<Rightarrow>
cdl_object_id set \<Rightarrow> sep_pred"
where
"si_caps_at t si_caps spec obj_ids \<equiv>
\<And>* obj_id \<in> obj_ids. (si_cap_at t si_caps spec) obj_id"
"si_caps_at t si_caps spec dev obj_ids \<equiv>
\<And>* obj_id \<in> obj_ids. (si_cap_at t si_caps spec) dev obj_id"
definition si_irq_caps_at :: "(cdl_irq \<Rightarrow> cdl_cptr option) \<Rightarrow> cdl_state \<Rightarrow>
cdl_irq set \<Rightarrow> sep_pred"
@ -329,34 +329,34 @@ where
definition
si_obj_cap_at' :: " (cdl_object_id \<Rightarrow> cdl_object_id option) \<Rightarrow>
(cdl_object_id \<Rightarrow> cdl_cptr option) \<Rightarrow> cdl_state \<Rightarrow>
(cdl_object_id \<Rightarrow> cdl_cptr option) \<Rightarrow> cdl_state \<Rightarrow> bool \<Rightarrow>
cdl_object_id \<Rightarrow> cdl_cnode_index \<Rightarrow> sep_pred"
where
"si_obj_cap_at' t si_caps spec obj_id slot \<equiv> \<lambda>s. \<exists> spec_cap.
si_cap_at t si_caps spec (cap_object spec_cap) s \<and>
"si_obj_cap_at' t si_caps spec dev obj_id slot \<equiv> \<lambda>s. \<exists> spec_cap.
si_cap_at t si_caps spec dev (cap_object spec_cap) s \<and>
opt_cap (obj_id, slot) spec = Some spec_cap"
definition si_obj_cap_at where
"si_obj_cap_at t si_caps spec obj_id slot \<equiv>
"si_obj_cap_at t si_caps spec dev obj_id slot \<equiv>
if original_cap_at (obj_id, slot) spec \<and> cap_at cap_has_object (obj_id, slot) spec
then si_obj_cap_at' t si_caps spec obj_id slot
then si_obj_cap_at' t si_caps spec dev obj_id slot
else \<box>"
definition
si_obj_caps_at :: "(cdl_object_id \<Rightarrow> cdl_object_id option) \<Rightarrow>
(cdl_object_id \<Rightarrow> cdl_cptr option) \<Rightarrow> cdl_state \<Rightarrow>
(cdl_object_id \<Rightarrow> cdl_cptr option) \<Rightarrow> cdl_state \<Rightarrow> bool \<Rightarrow>
cdl_object_id \<Rightarrow> sep_pred"
where
"si_obj_caps_at t si_caps spec obj_id \<equiv>
\<And>* slot \<in> dom (slots_of obj_id spec). si_obj_cap_at t si_caps spec obj_id slot"
"si_obj_caps_at t si_caps spec dev obj_id \<equiv>
\<And>* slot \<in> dom (slots_of obj_id spec). si_obj_cap_at t si_caps spec dev obj_id slot"
definition
si_objs_caps_at :: "(cdl_object_id \<Rightarrow> cdl_object_id option) \<Rightarrow>
(cdl_object_id \<Rightarrow> cdl_cptr option) \<Rightarrow> cdl_state \<Rightarrow>
(cdl_object_id \<Rightarrow> cdl_cptr option) \<Rightarrow> cdl_state \<Rightarrow> bool \<Rightarrow>
cdl_object_id set \<Rightarrow> sep_pred"
where
"si_objs_caps_at t si_caps spec obj_ids \<equiv>
\<And>* obj_id \<in> obj_ids. (si_obj_caps_at t si_caps spec) obj_id"
"si_objs_caps_at t si_caps spec dev obj_ids \<equiv>
\<And>* obj_id \<in> obj_ids. (si_obj_caps_at t si_caps spec) dev obj_id"
definition
si_spec_irq_cap_at' :: "(cdl_irq \<Rightarrow> cdl_cptr option) \<Rightarrow> cdl_state \<Rightarrow>
@ -471,7 +471,7 @@ where
lemma si_cap_at_less_si_cnode_size:
"\<lbrakk>\<guillemotleft>si_cap_at t opt_sel4_cap spec obj_id \<and>* R\<guillemotright> s;
"\<lbrakk>\<guillemotleft>si_cap_at t opt_sel4_cap spec dev obj_id \<and>* R\<guillemotright> s;
Some cap_ptr = opt_sel4_cap obj_id\<rbrakk>
\<Longrightarrow> cap_ptr < 2 ^ si_cnode_size"
by (clarsimp simp: si_cap_at_def sep_conj_exists)
@ -483,7 +483,7 @@ lemma si_irq_cap_at_less_si_cnode_size:
by (clarsimp simp: si_irq_cap_at_def sep_conj_exists)
lemma si_cap_at_has_k_obj_id:
"\<lbrakk>\<guillemotleft>si_cap_at t opt_sel4_cap spec obj_id \<and>* R\<guillemotright> s\<rbrakk>
"\<lbrakk>\<guillemotleft>si_cap_at t opt_sel4_cap spec dev obj_id \<and>* R\<guillemotright> s\<rbrakk>
\<Longrightarrow> \<exists>cap_object_id. t obj_id = Some cap_object_id"
by (clarsimp simp: si_cap_at_def sep_conj_exists)
@ -493,16 +493,16 @@ lemma si_cap_at_has_k_obj_id:
lemma valid_si_caps_at_si_cap_at:
"\<lbrakk>finite obj_ids; obj_id \<in> obj_ids;
(\<And>R. \<lbrace>\<guillemotleft>si_cap_at t orig_caps spec obj_id \<and>* P \<and>* R\<guillemotright>\<rbrace>
(\<And>R. \<lbrace>\<guillemotleft>si_cap_at t orig_caps spec dev obj_id \<and>* P \<and>* R\<guillemotright>\<rbrace>
f
\<lbrace>\<lambda>_.\<guillemotleft>si_cap_at t orig_caps spec obj_id \<and>* Q \<and>* R\<guillemotright>\<rbrace>)\<rbrakk>
\<lbrace>\<lambda>_.\<guillemotleft>si_cap_at t orig_caps spec dev obj_id \<and>* Q \<and>* R\<guillemotright>\<rbrace>)\<rbrakk>
\<Longrightarrow>
\<lbrace>\<guillemotleft>si_caps_at t orig_caps spec obj_ids \<and>* P \<and>* R\<guillemotright>\<rbrace>
\<lbrace>\<guillemotleft>si_caps_at t orig_caps spec dev obj_ids \<and>* P \<and>* R\<guillemotright>\<rbrace>
f
\<lbrace>\<lambda>_.\<guillemotleft>si_caps_at t orig_caps spec obj_ids \<and>* Q \<and>* R\<guillemotright>\<rbrace>"
\<lbrace>\<lambda>_.\<guillemotleft>si_caps_at t orig_caps spec dev obj_ids \<and>* Q \<and>* R\<guillemotright>\<rbrace>"
apply (clarsimp simp: si_caps_at_def)
apply (drule sep_set_conj_map_singleton_wp [where f=f and
I="si_cap_at t orig_caps spec" and P=P and Q=Q and R=R, rotated])
I="si_cap_at t orig_caps spec dev" and P=P and Q=Q and R=R, rotated])
apply (clarsimp simp: sep_conj_ac)+
done
@ -526,6 +526,7 @@ where
distinct_sets (map cap_free_ids untyped_caps) \<and>
list_all is_full_untyped_cap untyped_caps \<and>
list_all well_formed_untyped_cap untyped_caps \<and>
list_all (\<lambda>c. \<not> is_device_cap c) untyped_caps \<and>
bi_untypes bootinfo = (ustart, uend) \<and>
bi_free_slots bootinfo = (fstart, fend) \<and>
ustart < 2 ^ si_cnode_size \<and>
@ -547,7 +548,7 @@ where
(\<And>* (cptr, untyped_cap) \<in> set (zip untyped_cptrs untyped_caps).
(si_cnode_id, unat cptr) \<mapsto>c untyped_cap) \<and>*
(\<And>* obj_id \<in> all_available_ids. obj_id \<mapsto>o Untyped) \<and>*
(\<And>* obj_id \<in> {obj_id. cnode_or_tcb_at obj_id spec}. (si_cap_at t dup_caps spec obj_id)) \<and>*
(\<And>* obj_id \<in> {obj_id. cnode_or_tcb_at obj_id spec}. (si_cap_at t dup_caps spec False obj_id)) \<and>*
si_objects) s"
(********************************************************
@ -668,8 +669,8 @@ lemma si_caps_at_conversion:
"\<lbrakk>well_formed spec;
real_ids = {obj_id. real_object_at obj_id spec};
cnode_ids = {obj_id. cnode_at obj_id spec}\<rbrakk>
\<Longrightarrow> si_objs_caps_at t si_caps spec cnode_ids =
si_caps_at t si_caps spec real_ids"
\<Longrightarrow> si_objs_caps_at t si_caps spec dev cnode_ids =
si_caps_at t si_caps spec dev real_ids"
apply (clarsimp simp: si_objs_caps_at_def si_obj_caps_at_def [abs_def]
si_obj_cap_at_def [abs_def] si_caps_at_def)
apply (subst sep.setprod.Sigma, clarsimp+)

View File

@ -69,12 +69,12 @@ lemma cap_transform_MasterReplyCap:
lemma start_thread_sep:
"\<lbrace>\<guillemotleft>tcb_half_initialised spec t obj_id \<and>*
si_cap_at t dup_caps spec obj_id \<and>*
si_cap_at t dup_caps spec False obj_id \<and>*
si_objects \<and>* R\<guillemotright> and
K(well_formed spec \<and> obj_id \<in> {obj_id. is_waiting_thread_at obj_id spec})\<rbrace>
start_thread spec dup_caps obj_id
\<lbrace>\<lambda>_.\<guillemotleft>object_initialised spec t obj_id \<and>*
si_cap_at t dup_caps spec obj_id \<and>*
si_cap_at t dup_caps spec False obj_id \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (rule hoare_gen_asm)
apply (clarsimp simp: start_thread_def object_initialised_def tcb_half_initialised_def object_initialised_general_def
@ -126,12 +126,12 @@ lemma tcb_half_initialised_object_initialised':
lemma start_threads_sep:
"\<lbrace>\<guillemotleft>tcbs_half_initialised spec t {obj_id. tcb_at obj_id spec} \<and>*
si_caps_at t dup_caps spec {obj_id. cnode_or_tcb_at obj_id spec} \<and>*
si_caps_at t dup_caps spec False {obj_id. cnode_or_tcb_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright> and
K(well_formed spec \<and> set obj_ids = dom (cdl_objects spec) \<and> distinct obj_ids)\<rbrace>
start_threads spec dup_caps obj_ids
\<lbrace>\<lambda>_.\<guillemotleft>objects_initialised spec t {obj_id. tcb_at obj_id spec} \<and>*
si_caps_at t dup_caps spec {obj_id. cnode_or_tcb_at obj_id spec} \<and>*
si_caps_at t dup_caps spec False {obj_id. cnode_or_tcb_at obj_id spec} \<and>*
si_objects \<and>* R\<guillemotright>\<rbrace>"
apply (rule hoare_gen_asm)
apply (clarsimp simp: start_threads_def tcbs_half_initialised_def objects_initialised_def)
@ -150,7 +150,7 @@ lemma start_threads_sep:
apply (rule mapM_x_set_sep' [where
P="\<lambda>obj_id. tcb_half_initialised spec t obj_id" and
Q="\<lambda>obj_id. object_initialised spec t obj_id" and
I="si_caps_at t dup_caps spec {obj_id. cnode_or_tcb_at obj_id spec} \<and>*
I="si_caps_at t dup_caps spec False {obj_id. cnode_or_tcb_at obj_id spec} \<and>*
si_objects" and
xs="[obj_id \<leftarrow> obj_ids. is_waiting_thread_at obj_id spec]" and
X="{obj_id. object_at is_waiting_thread obj_id spec}" and
@ -165,7 +165,7 @@ lemma start_threads_sep:
and xs = "{obj_id. cnode_or_tcb_at obj_id spec}"
and P = "tcb_half_initialised spec t obj_id \<and>* si_objects"
and Q = "object_initialised spec t obj_id \<and>* si_objects"
and I = "si_cap_at t dup_caps spec"
and I = "si_cap_at t dup_caps spec False"
and R=R
in sep_set_conj_map_singleton_wp [simplified], simp_all add: object_at_real_object_at)

View File

@ -115,13 +115,13 @@ where
| NotificationCap _ b r \<Rightarrow> (b < 2 ^ badge_bits) \<and> (r \<subseteq> {AllowRead, AllowWrite})
| CNodeCap _ g gs sz \<Rightarrow> (gs < guard_bits) \<and> (g < 2 ^ gs) \<and> (sz + gs \<le> 32)
| TcbCap _ \<Rightarrow> True
| FrameCap _ r _ _ ad \<Rightarrow> r \<in> {vm_read_write, vm_read_only} \<and> ad = None
| FrameCap _ _ r _ _ ad \<Rightarrow> r \<in> {vm_read_write, vm_read_only} \<and> ad = None
| PageTableCap _ _ ad \<Rightarrow> ad = None
| PageDirectoryCap _ _ ad \<Rightarrow> ad = None
| IrqHandlerCap _ \<Rightarrow> True
(* LIMITATION: The following should probably eventually be true. *)
| IrqControlCap \<Rightarrow> False
| UntypedCap _ _ \<Rightarrow> False
| UntypedCap _ _ _ \<Rightarrow> False
| AsidControlCap \<Rightarrow> False
| AsidPoolCap _ _ \<Rightarrow> False
| _ \<Rightarrow> False)"
@ -131,7 +131,7 @@ definition vm_cap_has_asid :: "cdl_cap \<Rightarrow> bool"
where
"vm_cap_has_asid cap \<equiv>
(case cap of
FrameCap _ _ _ _ ad \<Rightarrow> ad \<noteq> None
FrameCap _ _ _ _ _ ad \<Rightarrow> ad \<noteq> None
| PageTableCap _ _ ad \<Rightarrow> ad \<noteq> None
| PageDirectoryCap _ _ ad \<Rightarrow> ad \<noteq> None
| _ \<Rightarrow> False)"
@ -140,7 +140,7 @@ definition is_real_vm_cap :: "cdl_cap \<Rightarrow> bool"
where
"is_real_vm_cap cap \<equiv>
(case cap of
FrameCap _ _ _ Real _ \<Rightarrow> True
FrameCap _ _ _ _ Real _ \<Rightarrow> True
| PageTableCap _ Real _ \<Rightarrow> True
| PageDirectoryCap _ Real _ \<Rightarrow> True
| _ \<Rightarrow> False)"
@ -149,7 +149,7 @@ definition is_fake_vm_cap :: "cdl_cap \<Rightarrow> bool"
where
"is_fake_vm_cap cap \<equiv>
(case cap of
FrameCap _ _ _ Fake _ \<Rightarrow> True
FrameCap _ _ _ _ Fake _ \<Rightarrow> True
| PageTableCap _ Fake _ \<Rightarrow> True
| PageDirectoryCap _ Fake _ \<Rightarrow> True
| _ \<Rightarrow> False)"
@ -166,7 +166,7 @@ definition
well_formed_orig_cap :: "cdl_cap \<Rightarrow> bool"
where
"well_formed_orig_cap cap \<equiv>
(cap_has_type cap \<longrightarrow> cap_rights (default_cap (the (cap_type cap)) undefined undefined)
(cap_has_type cap \<longrightarrow> cap_rights (default_cap (the (cap_type cap)) undefined undefined undefined)
= cap_rights cap) \<and>
(ep_related_cap cap \<longrightarrow> cap_badge cap = 0)"
@ -364,7 +364,8 @@ where
object_size_bits (object_default_state obj) = object_size_bits obj \<and>
dom (object_slots (object_default_state obj)) = dom (object_slots obj) \<and>
(cnode_at obj_id spec \<longrightarrow> 0 < object_size_bits obj)
| None \<Rightarrow> True)"
| None \<Rightarrow> True)
\<and> (\<forall>slot. \<not> cap_at (\<lambda>c. is_device_cap c = True) slot spec)"
lemma dom_cap_map [simp]:
"dom (\<lambda>n. if n \<le> N then Some a else None) = {0::nat .. N}"
@ -398,7 +399,7 @@ lemma well_formed_cap_update_cap_object [simp]:
lemma cap_rights_inter_default_cap_rights:
"\<lbrakk>well_formed_cap cap; cap_type cap = Some type\<rbrakk>
\<Longrightarrow> cap_rights (default_cap type ids sz) \<inter> cap_rights cap
\<Longrightarrow> cap_rights (default_cap type ids sz dev) \<inter> cap_rights cap
= cap_rights cap"
by (fastforce simp: well_formed_cap_def default_cap_def cap_type_def cap_rights_def
validate_vm_rights_def vm_read_write_def
@ -777,7 +778,7 @@ lemma default_ntfn_cap[simp]:
lemma default_cap_well_formed_cap:
"\<lbrakk>well_formed_cap cap; cap_type cap = Some type; cnode_cap_size cap = sz\<rbrakk>
\<Longrightarrow> well_formed_cap (default_cap type obj_ids sz)"
\<Longrightarrow> well_formed_cap (default_cap type obj_ids sz dev)"
by (auto simp: well_formed_cap_def default_cap_def cap_type_def
word_gt_a_gt_0 vm_read_write_def cnode_cap_size_def
split: cdl_cap.splits)
@ -785,7 +786,7 @@ lemma default_cap_well_formed_cap:
lemma default_cap_well_formed_cap2:
"\<lbrakk>is_default_cap cap; cap_type cap = Some type; sz \<le> 32;
\<not> is_untyped_cap cap; \<not> is_asidpool_cap cap\<rbrakk>
\<Longrightarrow> well_formed_cap (default_cap type obj_ids sz)"
\<Longrightarrow> well_formed_cap (default_cap type obj_ids sz dev )"
apply (clarsimp simp: is_default_cap_def)
apply (clarsimp simp: default_cap_def well_formed_cap_def
word_gt_a_gt_0 badge_bits_def guard_bits_def
@ -826,8 +827,8 @@ lemma well_formed_orig_ep_cap_is_default:
done
lemma cap_rights_default_cap_eq:
"cap_rights (default_cap type obj_ids sz) =
cap_rights (default_cap type obj_ids' sz')"
"cap_rights (default_cap type obj_ids sz dev) =
cap_rights (default_cap type obj_ids' sz' dev')"
apply (clarsimp simp: cap_rights_def default_cap_def)
apply (case_tac type, simp_all)
done
@ -835,7 +836,7 @@ lemma cap_rights_default_cap_eq:
lemma well_formed_orig_caps:
"\<lbrakk>well_formed spec; original_cap_at (obj_id, slot) spec;
slots_of obj_id spec slot = Some cap; cap \<noteq> NullCap; cap_type cap = Some type\<rbrakk>
\<Longrightarrow> cap_rights (default_cap type obj_ids sz) = cap_rights cap"
\<Longrightarrow> cap_rights (default_cap type obj_ids sz dev) = cap_rights cap"
apply (frule well_formed_well_formed_orig_cap, simp add: opt_cap_def, assumption+)
apply (clarsimp simp: well_formed_orig_cap_def)
apply (subst (asm) cap_rights_default_cap_eq, fast)
@ -1582,7 +1583,7 @@ lemma update_cap_rights_and_data:
\<Longrightarrow> update_cap_data_det
(cap_data spec_cap)
(update_cap_rights (cap_rights spec_cap)
(default_cap type {client_object_id} (cnode_cap_size spec_cap))) =
(default_cap type {client_object_id} (cnode_cap_size spec_cap) (is_device_cap spec_cap))) =
update_cap_object client_object_id spec_cap"
apply (case_tac "\<not>is_cnode_cap spec_cap")
apply (case_tac spec_cap, simp_all add: cap_type_def,
@ -1644,15 +1645,16 @@ lemma update_cap_data:
cap_type spec_cap = Some type; cap_data spec_cap = data;
well_formed_cap spec_cap; \<not> is_untyped_cap spec_cap;
\<not> vm_cap_has_asid spec_cap; \<not> is_fake_vm_cap spec_cap; \<not> is_irqhandler_cap spec_cap;
cap_rights (default_cap type {obj_id} sz) = cap_rights spec_cap\<rbrakk>
\<Longrightarrow> update_cap_data_det data (default_cap type {client_object_id} (cnode_cap_size spec_cap)) =
cap_rights (default_cap type {obj_id} sz (is_device_cap spec_cap)) = cap_rights spec_cap;
dev = is_device_cap spec_cap\<rbrakk>
\<Longrightarrow> update_cap_data_det data (default_cap type {client_object_id} (cnode_cap_size spec_cap) dev) =
update_cap_object client_object_id spec_cap"
apply (frule (6) update_cap_rights_and_data)
apply clarsimp
apply (subgoal_tac "update_cap_rights
apply (subgoal_tac "\<And>dev. update_cap_rights
(cap_rights spec_cap)
(default_cap type {client_object_id} (cnode_cap_size spec_cap))
= default_cap type {client_object_id} (cnode_cap_size spec_cap)")
(default_cap type {client_object_id} (cnode_cap_size spec_cap) dev)
= default_cap type {client_object_id} (cnode_cap_size spec_cap) dev")
apply clarsimp
apply (subst well_formed_update_cap_rights_idem)
apply (erule (1) default_cap_well_formed_cap, simp)

View File

@ -1,20 +1,11 @@
(*
* Copyright 2016, NICTA
*
* This software may be distributed and modified according to the terms of
* the BSD 2-Clause license. Note that NO WARRANTY is provided.
* See "LICENSE_BSD2.txt" for details.
*
* @TAG(NICTA_BSD)
*)
theory badnames
imports "../../AutoCorres"
begin
(*
* Test handling of C idents that are unusual or at risk of conflicting with other names.
*)
theory badnames imports "../../AutoCorres" begin
declare [[allow_underscore_idents = true]]
declare [[allow_underscore_idents]]
install_C_file "badnames.c"
autocorres "badnames.c"
end

View File

@ -107,7 +107,6 @@ ML_file "recursive_records/recursive_record_pp.ML"
ML_file "recursive_records/recursive_record_package.ML"
ML_file "expression_typing.ML"
ML_file "UMM_Proofs.ML"
ML_file "Listsort.ML"
ML_file "program_analysis.ML"
ML_file "heapstatetype.ML"
ML_file "MemoryModelExtras-sig.ML"

View File

@ -1,58 +0,0 @@
(* Listsort *)
signature LISTSORT =
sig
val sort : ('a * 'a -> order) -> 'a list -> 'a list
val sorted : ('a * 'a -> order) -> 'a list -> bool
end;
(*
[sort ordr xs] sorts the list xs in nondecreasing order, using the
given ordering. Uses Richard O'Keefe's smooth applicative merge
sort.
[sorted ordr xs] checks that the list xs is sorted in nondecreasing
order, in the given ordering.
*)
structure Listsort : LISTSORT =
struct
(* Listsort *)
(** Smooth Applicative Merge Sort, Richard O'Keefe 1982 **)
(** From L.C. Paulson: ML for the Working Programmer, CUP 1991 **)
(** Optimized for Moscow ML **)
fun sort ordr [] = []
| sort ordr (xs as [_]) = xs
| sort ordr (xs as [x1, x2]) =
(case ordr(x1, x2) of
GREATER => [x2, x1]
| _ => xs)
| sort ordr xs =
let fun merge [] ys = ys
| merge xs [] = xs
| merge (x::xs) (y::ys) =
if ordr(x, y) <> GREATER then x :: merge xs (y::ys)
else y :: merge (x::xs) ys
fun mergepairs l1 [] k = [l1]
| mergepairs l1 (ls as (l2::lr)) k =
if k mod 2 = 1 then l1::ls
else mergepairs (merge l1 l2) lr (k div 2)
fun nextrun run [] = (run, [])
| nextrun run (xs as (x::xr)) =
if ordr(x, List.hd run) = LESS then (run, xs)
else nextrun (x::run) xr
fun sorting [] ls r = List.hd(mergepairs [] ls 0)
| sorting (x::xs) ls r =
let val (revrun, tail) = nextrun [x] xs
in sorting tail (mergepairs (List.rev revrun) ls (r+1)) (r+1) end
in sorting xs [] 0 end;
fun sorted ordr [] = true
| sorted ordr (y1 :: yr) =
let fun h x0 [] = true
| h x0 (x1::xr) = ordr(x0, x1) <> GREATER andalso h x1 xr
in h y1 yr end;
end;

View File

@ -377,7 +377,7 @@ local
defined_functions = defined_functions, allow_underscore_idents = aui,
recursive_functions = recursive_functions, caller_info = caller_info,
modify_locs = modify_locs, typedefs = typedefs, fnspecs = fnspecs,
read_globals = read_globals, globinits = globinits, owners = owners,
read_globals = read_globals, globinits = globinits, owners = owners,
munge_info_fname = munge_info_fname}
(* fields in reverse order to above *)
fun from' munge_info_fname owners globinits read_globals fnspecs typedefs modify_locs caller_info
@ -2060,7 +2060,7 @@ fun export_mungedb (c as CSE cse) =
fun foldthis (_, vinfos) acc = List.foldl tab_foldthis acc vinfos
fun recurse tab = Symtab.fold foldthis tab []
val data0 = recurse (get_vars c)
val data = Listsort.sort String.compare data0
val data = Library.sort String.compare data0
in
List.app (fn s => TextIO.output(outstrm, s)) data;
TextIO.closeOut outstrm

View File

@ -33,7 +33,7 @@ in
../expression_typing.ML
../FunctionalRecordUpdate.ML
../complit.ML
../Listsort.ML
library.ML
../program_analysis.ML
../syntax_transforms.ML
ann

View File

@ -0,0 +1,67 @@
(*
* Extracted from Isabelle sources (src/Pure/library.ML)
*
* @TAG(OTHER_BSD)
*)
signature LIBRARY =
sig
val sort : ('a * 'a -> order) -> 'a list -> 'a list
end
structure Library : LIBRARY =
struct
(*stable mergesort -- preserves order of equal elements*)
fun mergesort unique ord =
let
fun merge (xs as x :: xs') (ys as y :: ys') =
(case ord (x, y) of
LESS => x :: merge xs' ys
| EQUAL =>
if unique then merge xs ys'
else x :: merge xs' ys
| GREATER => y :: merge xs ys')
| merge [] ys = ys
| merge xs [] = xs;
fun merge_all [xs] = xs
| merge_all xss = merge_all (merge_pairs xss)
and merge_pairs (xs :: ys :: xss) = merge xs ys :: merge_pairs xss
| merge_pairs xss = xss;
fun runs (x :: y :: xs) =
(case ord (x, y) of
LESS => ascending y [x] xs
| EQUAL =>
if unique then runs (x :: xs)
else ascending y [x] xs
| GREATER => descending y [x] xs)
| runs xs = [xs]
and ascending x xs (zs as y :: ys) =
(case ord (x, y) of
LESS => ascending y (x :: xs) ys
| EQUAL =>
if unique then ascending x xs ys
else ascending y (x :: xs) ys
| GREATER => rev (x :: xs) :: runs zs)
| ascending x xs [] = [rev (x :: xs)]
and descending x xs (zs as y :: ys) =
(case ord (x, y) of
GREATER => descending y (x :: xs) ys
| EQUAL =>
if unique then descending x xs ys
else (x :: xs) :: runs zs
| LESS => (x :: xs) :: runs zs)
| descending x xs [] = [x :: xs];
in merge_all o runs end;
fun sort ord = mergesort false ord;
end

View File

@ -61,6 +61,18 @@ def parse(call):
return ['%s\n' % line for line in lines]
def settings_line(l):
"""Adjusts some global settings."""
bits = l.split (',')
for bit in bits:
bit = bit.strip ()
(kind, setting) = bit.split ('=')
kind = kind.strip ()
if kind == 'keep_constructor':
[cons] = setting.split ()
keep_conss[cons] = 1
else:
assert not "setting kind understood", bit
def set_global(_call):
global call
@ -675,7 +687,7 @@ def typename_transform(line, header, d):
return d
dontwrap = {'asidpool': 1}
keep_conss = {}
def simple_newtype_transform(line, header, d):
@ -705,7 +717,7 @@ def simple_newtype_transform(line, header, d):
arities.append((str(bits[0]), len(bits[1:])))
if list((dict(arities)).values()) == [1] and header not in dontwrap:
if list((dict(arities)).values()) == [1] and header not in keep_conss:
return type_wrapper_type(header, last_lhs, last_rhs, d)
d.body = [('datatype %s =' % header, [(line, []) for line in lines])]
@ -777,7 +789,7 @@ def named_newtype_transform(line, header, d):
arities = [(name, len(map)) for (name, map) in constructors]
if list((dict(arities)).values()) == [1]:
if list((dict(arities)).values()) == [1] and header not in keep_conss:
[(cons, map)] = constructors
[(name, type)] = map
return type_wrapper_type(header, cons, type, d, decons=(name, type))

View File

@ -92,6 +92,10 @@ for line in instructions:
pass
else:
output_f.writelines(parsed)
elif line.startswith ("#INCLUDE_SETTINGS"):
(_, settings) = line.split (None, 1)
settings = settings.strip ()
lhs_pars.settings_line (settings)
else:
output_f.write(line)