Theory Security
section ‹Definition of the SIFUM-Security Property›
theory Security
imports Main Preliminaries
begin
context sifum_security begin
subsection ‹Evaluation of Concurrent Programs›
abbreviation eval_abv :: "('Com, 'Var, 'Val) LocalConf ⇒ (_, _, _) LocalConf ⇒ bool"
(infixl ‹↝› 70)
where
"x ↝ y ≡ (x, y) ∈ eval"
abbreviation conf_abv :: "'Com ⇒ 'Var Mds ⇒ ('Var, 'Val) Mem ⇒ (_,_,_) LocalConf"
(‹⟨_, _, _⟩› [0, 0, 0] 1000)
where
"⟨ c, mds, mem ⟩ ≡ ((c, mds), mem)"
inductive_set meval :: "(_,_,_) GlobalConf rel"
and meval_abv :: "_ ⇒ _ ⇒ bool" (infixl ‹→› 70)
where
"conf → conf' ≡ (conf, conf') ∈ meval" |
meval_intro [iff]: "⟦ (cms ! n, mem) ↝ (cm', mem'); n < length cms ⟧ ⟹
((cms, mem), (cms [n := cm'], mem')) ∈ meval"
inductive_cases meval_elim [elim!]: "((cms, mem), (cms', mem')) ∈ meval"
abbreviation meval_clos :: "_ ⇒ _ ⇒ bool" (infixl ‹→⇧*› 70)
where
"conf →⇧* conf' ≡ (conf, conf') ∈ meval⇧*"
fun lc_set_var :: "(_, _, _) LocalConf ⇒ 'Var ⇒ 'Val ⇒ (_, _, _) LocalConf"
where
"lc_set_var (c, mem) x v = (c, mem (x := v))"
fun meval_k :: "nat ⇒ ('Com, 'Var, 'Val) GlobalConf ⇒ (_, _, _) GlobalConf ⇒ bool"
where
"meval_k 0 c c' = (c = c')" |
"meval_k (Suc n) c c' = (∃ c''. meval_k n c c'' ∧ c'' → c')"
abbreviation meval_k_abv :: "nat ⇒ (_, _, _) GlobalConf ⇒ (_, _, _) GlobalConf ⇒ bool"
(‹_ →ı _› [100, 100] 80)
where
"gc →⇘k ⇙gc' ≡ meval_k k gc gc'"
subsection ‹Low-equivalence and Strong Low Bisimulations›
definition low_eq :: "('Var, 'Val) Mem ⇒ (_, _) Mem ⇒ bool" (infixl ‹=⇧l› 80)
where
"mem⇩1 =⇧l mem⇩2 ≡ (∀ x. dma x = Low ⟶ mem⇩1 x = mem⇩2 x)"
definition low_mds_eq :: "'Var Mds ⇒ ('Var, 'Val) Mem ⇒ (_, _) Mem ⇒ bool"
(‹_ =ı⇧l _› [100, 100] 80)
where
"(mem⇩1 =⇘mds⇙⇧l mem⇩2) ≡ (∀ x. dma x = Low ∧ x ∉ mds AsmNoRead ⟶ mem⇩1 x = mem⇩2 x)"
definition "mds⇩s" :: "'Var Mds" where
"mds⇩s x = {}"
lemma [simp]: "mem =⇧l mem' ⟹ mem =⇘mds⇙⇧l mem'"
by (simp add: low_mds_eq_def low_eq_def)
lemma [simp]: "(∀ mds. mem =⇘mds⇙⇧l mem') ⟹ mem =⇧l mem'"
by (auto simp: low_mds_eq_def low_eq_def)
definition closed_glob_consistent :: "(('Com, 'Var, 'Val) LocalConf) rel ⇒ bool"
where
"closed_glob_consistent ℛ =
(∀ c⇩1 mds mem⇩1 c⇩2 mem⇩2. (⟨ c⇩1, mds, mem⇩1 ⟩, ⟨ c⇩2, mds, mem⇩2 ⟩) ∈ ℛ ⟶
(∀ x. ((dma x = High ∧ x ∉ mds AsmNoWrite) ⟶
(∀ v⇩1 v⇩2. (⟨ c⇩1, mds, mem⇩1 (x := v⇩1) ⟩, ⟨ c⇩2, mds, mem⇩2 (x := v⇩2) ⟩) ∈ ℛ)) ∧
((dma x = Low ∧ x ∉ mds AsmNoWrite) ⟶
(∀ v. (⟨ c⇩1, mds, mem⇩1 (x := v) ⟩, ⟨ c⇩2, mds, mem⇩2 (x := v) ⟩) ∈ ℛ))))"
definition strong_low_bisim_mm :: "(('Com, 'Var, 'Val) LocalConf) rel ⇒ bool"
where
"strong_low_bisim_mm ℛ ≡
sym ℛ ∧
closed_glob_consistent ℛ ∧
(∀ c⇩1 mds mem⇩1 c⇩2 mem⇩2. (⟨ c⇩1, mds, mem⇩1 ⟩, ⟨ c⇩2, mds, mem⇩2 ⟩) ∈ ℛ ⟶
(mem⇩1 =⇘mds⇙⇧l mem⇩2) ∧
(∀ c⇩1' mds' mem⇩1'. ⟨ c⇩1, mds, mem⇩1 ⟩ ↝ ⟨ c⇩1', mds', mem⇩1' ⟩ ⟶
(∃ c⇩2' mem⇩2'. ⟨ c⇩2, mds, mem⇩2 ⟩ ↝ ⟨ c⇩2', mds', mem⇩2' ⟩ ∧
(⟨ c⇩1', mds', mem⇩1' ⟩, ⟨ c⇩2', mds', mem⇩2' ⟩) ∈ ℛ)))"
inductive_set mm_equiv :: "(('Com, 'Var, 'Val) LocalConf) rel"
and mm_equiv_abv :: "('Com, 'Var, 'Val) LocalConf ⇒
('Com, 'Var, 'Val) LocalConf ⇒ bool" (infix ‹≈› 60)
where
"mm_equiv_abv x y ≡ (x, y) ∈ mm_equiv" |
mm_equiv_intro [iff]: "⟦ strong_low_bisim_mm ℛ ; (lc⇩1, lc⇩2) ∈ ℛ ⟧ ⟹ (lc⇩1, lc⇩2) ∈ mm_equiv"
inductive_cases mm_equiv_elim [elim]: "⟨ c⇩1, mds, mem⇩1 ⟩ ≈ ⟨ c⇩2, mds, mem⇩2 ⟩"
definition low_indistinguishable :: "'Var Mds ⇒ 'Com ⇒ 'Com ⇒ bool"
(‹_ ∼ı _› [100, 100] 80)
where "c⇩1 ∼⇘mds⇙ c⇩2 = (∀ mem⇩1 mem⇩2. mem⇩1 =⇘mds⇙⇧l mem⇩2 ⟶
⟨ c⇩1, mds, mem⇩1 ⟩ ≈ ⟨ c⇩2, mds, mem⇩2 ⟩)"
subsection ‹SIFUM-Security›
definition com_sifum_secure :: "'Com ⇒ bool"
where "com_sifum_secure c = c ∼⇘mds⇩s⇙ c"
definition add_initial_modes :: "'Com list ⇒ ('Com × 'Var Mds) list"
where "add_initial_modes cmds = zip cmds (replicate (length cmds) mds⇩s)"
definition no_assumptions_on_termination :: "'Com list ⇒ bool"
where "no_assumptions_on_termination cmds =
(∀ mem mem' cms'.
(add_initial_modes cmds, mem) →⇧* (cms', mem') ∧
list_all (λ c. c = stop) (map fst cms') ⟶
(∀ mds' ∈ set (map snd cms'). mds' AsmNoRead = {} ∧ mds' AsmNoWrite = {}))"
definition prog_sifum_secure :: "'Com list ⇒ bool"
where "prog_sifum_secure cmds =
(no_assumptions_on_termination cmds ∧
(∀ mem⇩1 mem⇩2. mem⇩1 =⇧l mem⇩2 ⟶
(∀ k cms⇩1' mem⇩1'.
(add_initial_modes cmds, mem⇩1) →⇘k⇙ (cms⇩1', mem⇩1') ⟶
(∃ cms⇩2' mem⇩2'. (add_initial_modes cmds, mem⇩2) →⇘k⇙ (cms⇩2', mem⇩2') ∧
map snd cms⇩1' = map snd cms⇩2' ∧
length cms⇩2' = length cms⇩1' ∧
(∀ x. dma x = Low ∧ (∀ i < length cms⇩1'.
x ∉ snd (cms⇩1' ! i) AsmNoRead) ⟶ mem⇩1' x = mem⇩2' x)))))"
subsection ‹Sound Mode Use›
definition doesnt_read :: "'Com ⇒ 'Var ⇒ bool"
where
"doesnt_read c x = (∀ mds mem c' mds' mem'.
⟨ c, mds, mem ⟩ ↝ ⟨ c', mds', mem' ⟩ ⟶
((∀ v. ⟨ c, mds, mem (x := v) ⟩ ↝ ⟨ c', mds', mem' (x := v) ⟩) ∨
(∀ v. ⟨ c, mds, mem (x := v) ⟩ ↝ ⟨ c', mds', mem' ⟩)))"
definition doesnt_modify :: "'Com ⇒ 'Var ⇒ bool"
where
"doesnt_modify c x = (∀ mds mem c' mds' mem'. (⟨ c, mds, mem ⟩ ↝ ⟨ c', mds', mem' ⟩) ⟶
mem x = mem' x)"
inductive_set loc_reach :: "('Com, 'Var, 'Val) LocalConf ⇒ ('Com, 'Var, 'Val) LocalConf set"
for lc :: "(_, _, _) LocalConf"
where
refl : "⟨fst (fst lc), snd (fst lc), snd lc⟩ ∈ loc_reach lc" |
step : "⟦ ⟨c', mds', mem'⟩ ∈ loc_reach lc;
⟨c', mds', mem'⟩ ↝ ⟨c'', mds'', mem''⟩ ⟧ ⟹
⟨c'', mds'', mem''⟩ ∈ loc_reach lc" |
mem_diff : "⟦ ⟨ c', mds', mem' ⟩ ∈ loc_reach lc;
(∀ x ∈ mds' AsmNoWrite. mem' x = mem'' x) ⟧ ⟹
⟨ c', mds', mem'' ⟩ ∈ loc_reach lc"
definition locally_sound_mode_use :: "(_, _, _) LocalConf ⇒ bool"
where
"locally_sound_mode_use lc =
(∀ c' mds' mem'. ⟨ c', mds', mem' ⟩ ∈ loc_reach lc ⟶
(∀ x. (x ∈ mds' GuarNoRead ⟶ doesnt_read c' x) ∧
(x ∈ mds' GuarNoWrite ⟶ doesnt_modify c' x)))"
definition compatible_modes :: "('Var Mds) list ⇒ bool"
where
"compatible_modes mdss = (∀ (i :: nat) x. i < length mdss ⟶
(x ∈ (mdss ! i) AsmNoRead ⟶
(∀ j < length mdss. j ≠ i ⟶ x ∈ (mdss ! j) GuarNoRead)) ∧
(x ∈ (mdss ! i) AsmNoWrite ⟶
(∀ j < length mdss. j ≠ i ⟶ x ∈ (mdss ! j) GuarNoWrite)))"
definition reachable_mode_states :: "('Com, 'Var, 'Val) GlobalConf ⇒ (('Var Mds) list) set"
where "reachable_mode_states gc =
{mdss. (∃ cms' mem'. gc →⇧* (cms', mem') ∧ map snd cms' = mdss)}"
definition globally_sound_mode_use :: "('Com, 'Var, 'Val) GlobalConf ⇒ bool"
where "globally_sound_mode_use gc =
(∀ mdss. mdss ∈ reachable_mode_states gc ⟶ compatible_modes mdss)"
primrec sound_mode_use :: "(_, _, _) GlobalConf ⇒ bool"
where
"sound_mode_use (cms, mem) =
(list_all (λ cm. locally_sound_mode_use (cm, mem)) cms ∧
globally_sound_mode_use (cms, mem))"
lemma mm_equiv_sym:
assumes equivalent: "⟨c⇩1, mds⇩1, mem⇩1⟩ ≈ ⟨c⇩2, mds⇩2, mem⇩2⟩"
shows "⟨c⇩2, mds⇩2, mem⇩2⟩ ≈ ⟨c⇩1, mds⇩1, mem⇩1⟩"
proof -
from equivalent obtain ℛ
where ℛ_bisim: "strong_low_bisim_mm ℛ ∧ (⟨c⇩1, mds⇩1, mem⇩1⟩, ⟨c⇩2, mds⇩2, mem⇩2⟩) ∈ ℛ"
by (metis mm_equiv.simps)
hence "sym ℛ"
by (auto simp: strong_low_bisim_mm_def)
hence "(⟨c⇩2, mds⇩2, mem⇩2⟩, ⟨c⇩1, mds⇩1, mem⇩1⟩) ∈ ℛ"
by (metis ℛ_bisim symE)
thus ?thesis
by (metis ℛ_bisim mm_equiv.intros)
qed
lemma low_indistinguishable_sym: "lc ∼⇘mds⇙ lc' ⟹ lc' ∼⇘mds⇙ lc"
by (auto simp: mm_equiv_sym low_indistinguishable_def low_mds_eq_def)
lemma mm_equiv_glob_consistent: "closed_glob_consistent mm_equiv"
unfolding closed_glob_consistent_def
apply clarify
apply (erule mm_equiv_elim)
by (auto simp: strong_low_bisim_mm_def closed_glob_consistent_def)
lemma mm_equiv_strong_low_bisim: "strong_low_bisim_mm mm_equiv"
unfolding strong_low_bisim_mm_def
proof (auto)
show "closed_glob_consistent mm_equiv" by (rule mm_equiv_glob_consistent)
next
fix c⇩1 mds mem⇩1 c⇩2 mem⇩2 x
assume "⟨ c⇩1, mds, mem⇩1 ⟩ ≈ ⟨ c⇩2, mds, mem⇩2 ⟩"
then obtain ℛ where
"strong_low_bisim_mm ℛ ∧ (⟨ c⇩1, mds, mem⇩1 ⟩, ⟨ c⇩2, mds, mem⇩2 ⟩) ∈ ℛ"
by blast
thus "mem⇩1 =⇘mds⇙⇧l mem⇩2" by (auto simp: strong_low_bisim_mm_def)
next
fix c⇩1 :: 'Com
fix mds mem⇩1 c⇩2 mem⇩2 c⇩1' mds' mem⇩1'
let ?lc⇩1 = "⟨ c⇩1, mds, mem⇩1 ⟩" and
?lc⇩1' = "⟨ c⇩1', mds', mem⇩1' ⟩" and
?lc⇩2 = "⟨ c⇩2, mds, mem⇩2 ⟩"
assume "?lc⇩1 ≈ ?lc⇩2"
then obtain ℛ where "strong_low_bisim_mm ℛ ∧ (?lc⇩1, ?lc⇩2) ∈ ℛ"
by (rule mm_equiv_elim, blast)
moreover assume "?lc⇩1 ↝ ?lc⇩1'"
ultimately show "∃ c⇩2' mem⇩2'. ?lc⇩2 ↝ ⟨ c⇩2', mds', mem⇩2' ⟩ ∧ ?lc⇩1' ≈ ⟨ c⇩2', mds', mem⇩2' ⟩"
by (simp add: strong_low_bisim_mm_def mm_equiv_sym, blast)
next
show "sym mm_equiv"
by (auto simp: sym_def mm_equiv_sym)
qed
end
end