(* * Copyright 2014, NICTA * * This software may be distributed and modified according to the terms of * the GNU General Public License version 2. Note that NO WARRANTY is provided. * See "LICENSE_GPLv2.txt" for details. * * @TAG(NICTA_GPL) *) chapter "Restricted capabilities in the Separation Kernel Abstract Specification" theory Separation imports "ASepSpec.Syscall_SA" "AInvs.AInvs" "Lib.Bisim_UL" "Lib.LemmaBucket" begin text \ The seL4 kernel, when appropriately restricted, is a separation kernel. Any two processes in separate domains should behave the same as if they were processes running on two physically separated machines. They should not be aware of each other's existence and should not be able to communicate with each other except through well-defined channels. Importantly, it must be possible to show that there are no back channels through which one process can determine whether another process exists or what it is doing. In seL4 we achieve this by restricting the capabilities that a thread may possess. The restrictions are summarised in the predicate @{text separate_state} below (which indirectly depends on further predicates @{text separate_cnode_cap}, @{text separate_cap}, etc). a) A thread may only possess \emph{notification capabilities} (@{text NotificationCap}). b) Threads do not have caller capabilities. (A caller capability is a capability, placed in a special slot in the TCB, to allow replies. Since the @{text Reply} capability is disallowed so is the caller capability.) c) Pointers to other capability tables are disallowed meaning that the capability tree is flat. i.e. of depth 1 Initialising the kernel so that these restrictions hold is not covered in the bisimulation proof, but can be achieved using the capDL initialiser. Note that this proof does not preclude threads from communicating via shared memory if the threads have been set up accordingly, which again can be done via the capDL initialiser. The proof does show that the kernel API after reaching a state that satisifies @{text separate_state} is that of a static separation kernel, that is, it only provides system calls for sending and receiving on notification objects and otherwise exhibits no dynamic behaviour. Systems with such a setup satisfy the preconditions of our separate non-intereference proof, which shows that information travels only along these authorised channels. \ definition separate_cap :: "cap \ bool" where "separate_cap cap \ case cap of NotificationCap ptr badge rights \ rights \ {AllowRecv, AllowSend} | NullCap \ True | _ \ False" lemma separate_capE: "\ separate_cap cap; cap = NullCap \ R; \ptr badge rights. \ cap = NotificationCap ptr badge rights \ \ R \ \ R" unfolding separate_cap_def by (fastforce split: cap.splits) definition "separate_cnode_cap cs cap \ case cap of CNodeCap p bits guard \ (bits + length guard = word_bits) \ (\off. case_option True separate_cap (cs (p, off))) | NullCap \ True | _ \ False" definition "separate_tcb p cs \ case_option True (separate_cnode_cap cs) (cs (p, tcb_cnode_index 0)) \ cs (p, tcb_cnode_index 3) = Some NullCap" \ \ctable and caller cap\ lemma separate_cnode_cap_rab: "\ separate_cnode_cap cs cap; length cref = word_bits \ \ resolve_address_bits (cap, cref) = (case cap of CNodeCap p bits guard \ if guard \ cref then returnOk ((p, drop (length guard) cref), []) else (throwError (GuardMismatch (length cref) guard)) | _ \ throwError InvalidRoot)" unfolding separate_cnode_cap_def resolve_address_bits_def by (auto simp: word_bits_def split: cap.split_asm) definition "separate_state s \ \p. tcb_at p s \ separate_tcb p (caps_of_state s)" lemma separate_cnode_capE: "\ separate_cnode_cap cs cap; cap = NullCap \ R; \p bits guard. \ cap = CNodeCap p bits guard; bits + length guard = word_bits; (\off cap'. cs (p, off) = Some cap' \ separate_cap cap') \ \ R \ \ R" unfolding separate_cnode_cap_def by (auto split: cap.splits option.splits) lemma valid_sep_cap_not_cnode: "\ s \ cap; \off cap'. caps_of_state s (p, off) = Some cap' \ separate_cap cap'; cap = CNodeCap p bits guard; bits \ length cref - length guard \ \ \cap'. caps_of_state s (p, take bits (drop (length guard) cref)) = Some cap' \ \ is_cnode_cap cap'" apply (clarsimp simp: valid_cap_simps not_less in_monad) apply (drule_tac offset = "take bits (drop (length guard) cref)" in cap_table_at_cte_at) apply simp apply (fastforce simp: cte_wp_at_caps_of_state separate_cap_def is_cap_simps) done lemma bisim_gen_asm_r: assumes bs: "F \ bisim_underlying sr r P P' a b" shows "bisim_underlying sr r P (P' and K F) a b" using bs by (fastforce intro!: bisim_underlyingI elim: bisim_underlyingE1 bisim_underlyingE2) lemma bisim_separate_cap_cases: assumes nc: "cap = NullCap \ bisim R Pn Pn' m m'" and ac: "\ptr badge rights. \ cap = NotificationCap ptr badge rights \ \ bisim R (Pa ptr badge rights) (Pa' ptr badge rights) m m'" shows "bisim R (\s. (cap = NullCap \ Pn s) \ (\ptr badge rights. cap = NotificationCap ptr badge rights \ Pa ptr badge rights s)) ((\s. (cap = NullCap \ Pn' s) \ (\ptr badge rights. cap = NotificationCap ptr badge rights \ Pa' ptr badge rights s)) and K (separate_cap cap)) m m'" using assms apply - apply (rule bisim_gen_asm_r) apply (erule separate_capE, simp_all) done lemma caps_of_state_tcb: "\ get_tcb p s = Some tcb; option_map fst (tcb_cap_cases idx) = Some getF \ \ caps_of_state s (p, idx) = Some (getF tcb)" apply (drule get_tcb_SomeD) apply clarsimp apply (drule (1) cte_wp_at_tcbI [where t = "(p, idx)" and P = "(=) (getF tcb)", simplified]) apply simp apply (clarsimp simp: cte_wp_at_caps_of_state) done lemma caps_of_state_tcb_cap_cases: "\ get_tcb p s = Some tcb; idx \ dom tcb_cap_cases \ \ caps_of_state s (p, idx) = Some ((the (option_map fst (tcb_cap_cases idx))) tcb)" apply (clarsimp simp: dom_def) apply (erule caps_of_state_tcb) apply simp done lemma separate_state_get_tcbD: "\separate_state s; get_tcb p s = Some tcb \ \ separate_cnode_cap (caps_of_state s) (tcb_ctable tcb) \ tcb_caller tcb = NullCap" unfolding separate_state_def apply (drule spec [where x = p]) apply (simp add: tcb_at_def separate_tcb_def caps_of_state_tcb_cap_cases dom_tcb_cap_cases) done end