Initial import from darcs
This commit is contained in:
commit
855ac3eff4
|
@ -0,0 +1,14 @@
|
|||
*.vo
|
||||
*.vio
|
||||
*.v.d
|
||||
*.glob
|
||||
*.cache
|
||||
*.aux
|
||||
\#*\#
|
||||
.\#*
|
||||
*~
|
||||
*.bak
|
||||
build-dep/
|
||||
Makefile.coq
|
||||
Makefile.coq.conf
|
||||
*.crashcoqide
|
|
@ -0,0 +1,128 @@
|
|||
Require Import Rushby.
|
||||
Require Import list relations collections fin_collections.
|
||||
|
||||
Module ArrayMachine <: Mealy.
|
||||
|
||||
Definition state := nat -> nat.
|
||||
Inductive command :=
|
||||
| Write : nat -> nat -> command
|
||||
| Read : nat -> command
|
||||
.
|
||||
Definition action := command.
|
||||
Definition out := nat.
|
||||
|
||||
Definition extendS (s : state) (n : nat) (m : nat) : state :=
|
||||
fun x => if beq_nat x n then m else s x.
|
||||
|
||||
Definition preform (s : state) (a : action) : state * out :=
|
||||
match a with
|
||||
| Write i j => (extendS s i j, 0)
|
||||
| Read i => (s, s i)
|
||||
end.
|
||||
|
||||
Definition step s a := fst (preform s a).
|
||||
Definition output s a := snd (preform s a).
|
||||
Definition initial (x : nat) := 0.
|
||||
End ArrayMachine.
|
||||
|
||||
|
||||
Import ArrayMachine.
|
||||
Module M := Rushby.Rushby ArrayMachine.
|
||||
Import M.
|
||||
|
||||
Eval compute in (do_actions [Write 1 1] 2).
|
||||
|
||||
Definition domain := action.
|
||||
Definition nameA := nat.
|
||||
Definition valA := nat.
|
||||
Definition observeA (u : domain) : FinSet nameA :=
|
||||
match u with
|
||||
| Read i => {[ i ]}
|
||||
| Write _ _ => ∅
|
||||
end.
|
||||
Definition alterA (u : domain) : FinSet valA :=
|
||||
match u with
|
||||
| Read _ => ∅
|
||||
| Write i _ => {[ i ]}
|
||||
end.
|
||||
|
||||
|
||||
Instance domain_dec : forall (u v : domain), Decision (u = v).
|
||||
Proof. intros.
|
||||
unfold Decision.
|
||||
repeat (decide equality).
|
||||
Defined.
|
||||
|
||||
Instance arraymachine_ss : StructuredState domain :=
|
||||
{ name := nameA; value := valA; contents s n := s n
|
||||
; observe := observeA; alter := alterA }.
|
||||
|
||||
Definition interference (u v : domain) :=
|
||||
(exists (n : nameA), n ∈ alterA u ∧ n ∈ observeA v).
|
||||
|
||||
Inductive interferenceR : relation domain :=
|
||||
| interference_refl : forall (u : domain), interferenceR u u
|
||||
| interference_step : forall (u v: domain), interference u v -> interferenceR u v.
|
||||
|
||||
Instance policy_ss : Policy domain :=
|
||||
{ policy := interferenceR
|
||||
; dom := fun (a : action) => (a : domain) }.
|
||||
Proof.
|
||||
intros. unfold Decision.
|
||||
destruct v as [i j | i]; destruct w as [m n | m].
|
||||
- destruct (decide (i = m)). destruct (decide (j = n)); subst; auto.
|
||||
left. constructor.
|
||||
right. intro I. inversion I; subst. apply n0. auto.
|
||||
inversion H. unfold observeA in *. destruct H0 as [HH HHH].
|
||||
apply (not_elem_of_empty x); assumption.
|
||||
right. intro. inversion H; subst. apply n0. auto.
|
||||
inversion H0. unfold alterA, observeA in *. destruct H1.
|
||||
apply (not_elem_of_empty x); assumption.
|
||||
- destruct (decide (i = m)); subst. left. right. unfold interference.
|
||||
simpl. exists m. split; apply elem_of_singleton; auto.
|
||||
right. intro. inversion H; subst. inversion H0. simpl in H1.
|
||||
destruct H1. apply elem_of_singleton in H1; apply elem_of_singleton in H2. subst. apply n;auto.
|
||||
- right. intro. inversion H;subst. inversion H0; subst.
|
||||
simpl in H1; inversion H1. apply (not_elem_of_empty x); assumption.
|
||||
- destruct (decide (i = m)); subst.
|
||||
left. constructor.
|
||||
right. intro. inversion H; subst. auto.
|
||||
inversion H0; subst. simpl in H1; inversion H1. eapply not_elem_of_empty. eassumption.
|
||||
- intro u. constructor.
|
||||
Defined.
|
||||
|
||||
Check RefMonAssumptions.
|
||||
|
||||
Instance rma_yay : RefMonAssumptions.
|
||||
Proof. split; simpl; unfold RMA_partition; intros;
|
||||
unfold contents in *; simpl in H8.
|
||||
unfold output.
|
||||
unfold preform. destruct a as [i j | i]; simpl in *. reflexivity.
|
||||
apply H8. apply elem_of_singleton. reflexivity.
|
||||
|
||||
unfold step, preform. destruct a as [i j | i]. simpl in *.
|
||||
destruct (decide (i = n)); subst; unfold extendS; simpl.
|
||||
replace (beq_nat n n) with true; auto. apply beq_nat_refl.
|
||||
destruct H9.
|
||||
unfold step in H9; simpl in H9. unfold extendS in H9; simpl in H9.
|
||||
replace (beq_nat n i) with false in *; auto.
|
||||
congruence. SearchAbout beq_nat false. symmetry. apply beq_nat_false_iff. omega.
|
||||
unfold step in H9; simpl in H9. unfold extendS in H9; simpl in H9.
|
||||
replace (beq_nat n i) with false in *; auto.
|
||||
congruence. symmetry. apply beq_nat_false_iff. omega.
|
||||
|
||||
simpl. destruct H9; unfold step, preform in H9; simpl in H9;
|
||||
congruence.
|
||||
|
||||
unfold step, preform in H8; simpl in H8. destruct a; simpl in *.
|
||||
unfold extendS in H8. destruct (decide (n = n0)); subst.
|
||||
apply elem_of_singleton; auto. replace (beq_nat n n0) with false in *; try (congruence). symmetry. apply beq_nat_false_iff. assumption.
|
||||
congruence.
|
||||
Defined.
|
||||
|
||||
Theorem yay : isecurity.
|
||||
Proof.
|
||||
apply rma_secure_intransitive.
|
||||
intros u v n A1 A2. simpl. right. firstorder.
|
||||
Qed.
|
||||
|
File diff suppressed because it is too large
Load Diff
|
@ -0,0 +1,80 @@
|
|||
Require Export Coq.Bool.Bool.
|
||||
Require Export Coq.Arith.Arith.
|
||||
Require Export Coq.Arith.EqNat.
|
||||
Require Export Coq.omega.Omega.
|
||||
Require Export Coq.Lists.List.
|
||||
Require Export Coq.ZArith.ZArith.
|
||||
Require Export Coq.Numbers.Natural.Peano.NPeano.
|
||||
Require Export Coq.Setoids.Setoid.
|
||||
Require Export Coq.Program.Equality. (**r Necessary for 'dependent induction'. *)
|
||||
|
||||
Require Export LibTactics.
|
||||
|
||||
|
||||
Ltac autoinjection :=
|
||||
repeat match goal with
|
||||
| h: ?f _ = ?f _ |- _ => injection h; intros; clear h; subst
|
||||
| h: ?f _ _ = ?f _ _ |- _ => injection h; intros; clear h; subst
|
||||
| h: ?f _ _ _ = ?f _ _ _ |- _ => injection h; intros; clear h; subst
|
||||
| h: ?f _ _ _ _ = ?f _ _ _ _ |- _ => injection h; intros; clear h; subst
|
||||
| h: ?f _ _ _ _ _ = ?f _ _ _ _ _ |- _ => injection h; intros; clear h; subst
|
||||
end.
|
||||
|
||||
Ltac go :=
|
||||
simpl in *;
|
||||
repeat match goal with
|
||||
| h: ?x = _ |- context[match ?x with _ => _ end] => rewrite h
|
||||
end;
|
||||
autoinjection;
|
||||
try (congruence);
|
||||
try omega;
|
||||
subst;
|
||||
eauto 4 with zarith datatypes;
|
||||
try (econstructor ; (solve[go])).
|
||||
|
||||
Tactic Notation "go" := try (go; fail).
|
||||
|
||||
Ltac go_with b :=
|
||||
simpl in *;
|
||||
repeat match goal with
|
||||
| h: ?x = _ |- context[match ?x with _ => _ end] => rewrite h
|
||||
end;
|
||||
autoinjection;
|
||||
try (congruence);
|
||||
try omega;
|
||||
subst;
|
||||
eauto 4 with zarith datatypes b;
|
||||
try (econstructor ; (solve[go])).
|
||||
|
||||
Ltac inv H := inversion H; try subst; clear H.
|
||||
|
||||
Tactic Notation "flatten" ident(H) :=
|
||||
repeat match goal with
|
||||
| H: context[match ?x with | left _ => _ | right _ => _ end] |- _ => destruct x
|
||||
| H: context[match ?x with | _ => _ end] |- _ => let E := fresh "Eq" in destruct x eqn:E
|
||||
end; autoinjection; try congruence.
|
||||
|
||||
Tactic Notation "flatten" :=
|
||||
repeat match goal with
|
||||
| |- context[match ?x with | left _ => _ | right _ => _ end] => destruct x
|
||||
| |- context[match ?x with | _ => _ end] => let E := fresh "Eq" in destruct x eqn:E
|
||||
end; autoinjection; try congruence.
|
||||
|
||||
(*Tactic Notation "induction" ident(x) := dependent induction x.*)
|
||||
|
||||
Definition admit {T: Type} : T. Admitted.
|
||||
|
||||
Tactic Notation "solve_by_inversion_step" tactic(t) :=
|
||||
match goal with
|
||||
| H : _ |- _ => solve [ inversion H; subst; t ]
|
||||
end
|
||||
|| fail "because the goal is not solvable by inversion.".
|
||||
|
||||
Tactic Notation "solve" "by" "inversion" "1" :=
|
||||
solve_by_inversion_step idtac.
|
||||
Tactic Notation "solve" "by" "inversion" "2" :=
|
||||
solve_by_inversion_step (solve by inversion 1).
|
||||
Tactic Notation "solve" "by" "inversion" "3" :=
|
||||
solve_by_inversion_step (solve by inversion 2).
|
||||
Tactic Notation "solve" "by" "inversion" :=
|
||||
solve by inversion 1.
|
|
@ -0,0 +1,99 @@
|
|||
Require Import base.
|
||||
Require Import Monoids.
|
||||
|
||||
Class Mealy (state action out : Type) := {
|
||||
initial : state;
|
||||
step : state -> action -> state;
|
||||
output : state -> action -> out
|
||||
}.
|
||||
|
||||
(** ** Auxiliary functions
|
||||
|
||||
We define a [run] of a machine [M], which is an extension of
|
||||
the [step] function to lists of actions. We use a shortcut
|
||||
[do_actions] for running the machine from the [initial] state. *)
|
||||
|
||||
|
||||
Fixpoint run `{Mealy state action out} (s : state) (ls : list action) : state :=
|
||||
match ls with
|
||||
| nil => s
|
||||
| a::tl => run (step s a) tl
|
||||
end.
|
||||
|
||||
Definition do_actions `{Mealy state action out} : list action -> state := run initial.
|
||||
|
||||
(** The [test] function runs the required list of actions and examines the output of the resulting state on a specified action. *)
|
||||
Definition test `{Mealy state action out} (ls : list action) : action -> out := output (do_actions ls).
|
||||
|
||||
Global Instance MealyMRet `{Monoid out op e} : MRet (λ A, (A * out)%type) :=
|
||||
fun A x => (x, e).
|
||||
Global Instance MealyMBind `{Monoid out op e} : MBind (λ A, (A * out)%type) :=
|
||||
fun A B (f : A → B * out) (x : A * out) =>
|
||||
match x with
|
||||
(x₁,x₂) => match f x₁ with
|
||||
(y₁,y₂) => (y₁, op x₂ y₂)
|
||||
end
|
||||
end.
|
||||
Global Instance MealyMonad `{Monoid out op e} `{Mealy state action out} :
|
||||
Monad (λ A, (A * out)%type).
|
||||
Proof. split; intros;
|
||||
unfold mret, mbind; unfold MealyMRet, MealyMBind.
|
||||
destruct m. rewrite right_id; auto.
|
||||
apply H. (* TODO: why is this not being picked up automatically? *)
|
||||
|
||||
destruct (f x). rewrite left_id; auto. apply H.
|
||||
|
||||
destruct m. destruct (f a). destruct (g b).
|
||||
rewrite associative. reflexivity. apply H.
|
||||
Qed.
|
||||
About run.
|
||||
|
||||
Class MSync (state action out : Type)
|
||||
:= {
|
||||
msync_mealy : Mealy state action out;
|
||||
delayed : action -> (state -> Prop);
|
||||
delayed_dec :> forall a s, Decision (delayed a s)
|
||||
}.
|
||||
|
||||
Global Instance MSync_as_Mealy `{MS: MSync state action out}
|
||||
: Mealy state action (option out) :=
|
||||
{ initial := @initial _ _ _ msync_mealy
|
||||
; step := fun s a => if (decide (delayed a s))
|
||||
then s
|
||||
else @step _ _ _ msync_mealy s a
|
||||
; output := fun s a => if (decide (delayed a s))
|
||||
then None
|
||||
else Some (@output _ _ _ msync_mealy s a)
|
||||
}.
|
||||
|
||||
|
||||
Class MApi (domain : Type) (API : Type) (state action out : Type) :=
|
||||
{ mapi_mealy : Mealy state action out
|
||||
; domain_dec :> forall (x y : domain), Decision (x = y)
|
||||
; dom_api : API -> domain
|
||||
; sem_api : API -> list action
|
||||
}.
|
||||
|
||||
Definition upd `{MApi domain API s a o}
|
||||
(f : domain -> list a) a b :=
|
||||
fun d => if (decide (a = d)) then b else f d.
|
||||
|
||||
Global Instance MApi_as_Mealy `{MA: MApi domain API state action out}
|
||||
: Mealy (state * (domain -> list action)) API (option out)
|
||||
:=
|
||||
{ initial := (@initial _ _ _ mapi_mealy, (λ _,[]))
|
||||
; step := fun s a => let (s', f) := s
|
||||
in let step' := @step _ _ _ mapi_mealy
|
||||
in match (f (dom_api a)) with
|
||||
| [] => (s', upd f (dom_api a) (sem_api a))
|
||||
| (x::xs) => (step' s' x, upd f (dom_api a) xs)
|
||||
end
|
||||
; output := fun s a => let (s', f) := s
|
||||
in let output' := @output _ _ _ mapi_mealy
|
||||
in match (f (dom_api a)) with
|
||||
| [] => None
|
||||
| (x::xs) => Some (output' s' x)
|
||||
end
|
||||
}.
|
||||
|
||||
|
|
@ -0,0 +1,79 @@
|
|||
Require Import Rushby.
|
||||
Require Import base.
|
||||
|
||||
Class MSync (state action out : Type)
|
||||
:= {
|
||||
msync_mealy : Mealy state action out;
|
||||
delayed : action -> (state -> Prop);
|
||||
delayed_dec :> forall a s, Decision (delayed a s)
|
||||
}.
|
||||
Check @step.
|
||||
Instance MSync_as_Mealy `{MSync state action out}
|
||||
: Mealy state action (option out).
|
||||
Proof.
|
||||
split. exact (@initial _ _ _ msync_mealy).
|
||||
refine (fun s a => if (decide (delayed a s))
|
||||
then @step _ _ _ msync_mealy s a
|
||||
else s).
|
||||
refine (fun s a => if (decide (delayed a s))
|
||||
then Some (@output _ _ _ msync_mealy s a)
|
||||
else None).
|
||||
Qed.
|
||||
|
||||
Section MealyM.
|
||||
|
||||
Variable state action out : Type.
|
||||
Variable MM : MSync state action out.
|
||||
Variable domain : Type.
|
||||
Check Policy.
|
||||
Variable P : Policy domain.
|
||||
Variable VP : ViewPartition domain.
|
||||
|
||||
|
||||
Definition sync_separation `{Policy action domain} := forall (s : state) (b : action) (a : action),
|
||||
¬ (delayed a s) ∧ (delayed a (step s a))
|
||||
→ policy (Rushby.dom b) (Rushby.dom a).
|
||||
|
||||
Definition sync_interf `{Policy action domain} `{MMs : MMsync state action out}
|
||||
|
||||
Class SyncPartition {domain state action out : Type} (MMs : MMsync state action out) {P : Policy domain} {VP : @ViewPartition state domain} := {
|
||||
sync_partitioned : forall s t a, view_partition (Rushby.dom a) s t -> (delayed a s <-> delayed a t)
|
||||
}.
|
||||
|
||||
|
||||
Section Unwinding.
|
||||
Context {state action out : Type}.
|
||||
Context {domain : Type}.
|
||||
Context {P : @Policy action domain}.
|
||||
Context {VP : @ViewPartition state domain}.
|
||||
Context {MMs : MMsync state action out}.
|
||||
Definition MM := underlyingM.
|
||||
Definition N := (@MMsync_is_MM domain state action out MMs).
|
||||
|
||||
Context {OC : @OutputConsistent _ _ _ MM domain P VP}.
|
||||
Context {SP : SyncPartition MMs}.
|
||||
Instance OC' : @OutputConsistent _ _ _ N domain P VP.
|
||||
Proof.
|
||||
intros a s t L.
|
||||
unfold output; simpl.
|
||||
destruct (decide (delayed a s)) as [D | ND];
|
||||
destruct (decide (delayed a t)) as [D' | ND']; auto;
|
||||
try (erewrite output_consistent by (apply L); reflexivity);
|
||||
(exfalso; (apply ND' || apply ND);
|
||||
(rewrite <- (sync_partitioned s t a L) ||
|
||||
rewrite -> (sync_partitioned s t a L)); auto).
|
||||
Qed.
|
||||
|
||||
Theorem unwinding_sync : locally_respects_policy (MM:=MM) -> step_consistent (MM:=MM) -> sync_separation -> security (MM:=N).
|
||||
Proof. intros.
|
||||
apply unwinding.
|
||||
|
||||
(* local step consistency *)
|
||||
intros a u s NP. unfold step; simpl.
|
||||
destruct (decide (delayed a s)); auto. reflexivity.
|
||||
(* TODO auto cannot solve reflexivity? *)
|
||||
|
||||
intros a s t u L. unfold step; simpl.
|
||||
destruct (decide (delayed a s)) as [D | ND];
|
||||
destruct (decide (delayed a t)) as [D' | ND']; auto.
|
||||
transitivity t; auto. apply H. apply
|
|
@ -0,0 +1,30 @@
|
|||
Require Import base.
|
||||
|
||||
Class Monad M {ret : MRet M} {bind : MBind M} :=
|
||||
{ ret_unit_1 : forall {A} (m : M A) , m ≫= mret = m
|
||||
; ret_unit_2 : forall {A B} (x : A) (f : A → M B),
|
||||
(mret x) ≫= f = f x
|
||||
; bind_assoc : forall {A B C} (m : M A) (f : A → M B) (g : B → M C),
|
||||
(m ≫= f) ≫= g = m ≫= (fun x => f x ≫= g)
|
||||
}.
|
||||
|
||||
Global Instance Join_of_Monad M `{Monad M} : (MJoin M) :=
|
||||
fun A (m : M (M A)) => m ≫= id.
|
||||
Global Instance FMap_of_Monad M `{Monad M} : FMap M :=
|
||||
fun A B f ma => ma ≫= (fun x => mret (f x)).
|
||||
|
||||
Class Monoid A (op: A → A → A) (e : A) :=
|
||||
{ op_assoc : Associative (=) op
|
||||
; e_left : LeftId (=) e op
|
||||
; e_right : RightId (=) e op
|
||||
}.
|
||||
|
||||
Definition kleisli {A B C} `{mon : Monad M} (f : A → M B) (g : B → M C) (x : A) : M C := y ← f x; g y.
|
||||
Notation "f >=> g" := (kleisli f g) (at level 60, right associativity) : C_scope.
|
||||
Notation "( f >=>)" := (kleisli f) (only parsing) : C_scope.
|
||||
Notation "(>=> g )" := (fun f => kleisli f g) (only parsing) : C_scope.
|
||||
Notation "(>=>)" := (fun f g => kleisli f g) (only parsing) : C_scope.
|
||||
Notation "g <=< f" := (kleisli f g) (at level 60, right associativity) : C_scope.
|
||||
Notation "( g <=<)" := (fun f => kleisli f g) (only parsing) : C_scope.
|
||||
Notation "(<=< f )" := (kleisli f) (only parsing) : C_scope.
|
||||
|
|
@ -0,0 +1,19 @@
|
|||
Require Import base.
|
||||
|
||||
Class Policy (action domain : Type) := {
|
||||
dom : action -> domain;
|
||||
domain_dec :> forall (x y : domain), Decision (x = y);
|
||||
policy :> relation domain;
|
||||
policy_dec :> (forall v w, Decision (policy v w));
|
||||
policy_refl :> Reflexive policy
|
||||
}.
|
||||
|
||||
Infix "⇝" := policy (at level 70).
|
||||
|
||||
Fixpoint purge `{Policy action domain} (ls : list action) (v : domain) : list action :=
|
||||
match ls with
|
||||
| [] => []
|
||||
| a::tl => if (decide (dom a ⇝ v)) then a::(purge tl v)
|
||||
else purge tl v
|
||||
end.
|
||||
|
|
@ -0,0 +1,520 @@
|
|||
(** printing -> #→# *)
|
||||
(** printing (policy a b) #a ⇝ b# *)
|
||||
|
||||
(** Formalisation of "Noninterference, Transitivity, and Channel-Control Security Policies" by J. Rushby
|
||||
www.csl.sri.com/papers/csl-92-2/
|
||||
*)
|
||||
|
||||
|
||||
(** We use Robbert Krebbers' prelude, see https://github.com/robbertkrebbers/ch2o/tree/master/prelude **)
|
||||
Require Import list relations collections fin_collections.
|
||||
Parameter FinSet : Type -> Type.
|
||||
(** begin hide **)
|
||||
Context `{forall A, ElemOf A (FinSet A)}.
|
||||
Context `{forall A, Empty (FinSet A)}.
|
||||
Context `{forall A, Singleton A (FinSet A)}.
|
||||
Context `{forall A, Union (FinSet A)}.
|
||||
Context `{forall A, Intersection (FinSet A)}.
|
||||
Context `{forall A, Difference (FinSet A)}.
|
||||
Context `{forall A, Elements A (FinSet A)}.
|
||||
Context `{forall A, Collection A (FinSet A)}. (* TODO: i wrote this line down so that there is a Collection -> SimpleCollection -> JoinSemiLattice instance for FinSet; how come this is not automatically picked up by the next assumption? *)
|
||||
Context `{forall A (H : (forall (a b :A), Decision (a = b))), FinCollection A (FinSet A)}.
|
||||
(** end hide **)
|
||||
|
||||
(** * Mealy machines *)
|
||||
|
||||
(** A Mealy machine is a state machine with labels and outputs on arrows *)
|
||||
|
||||
Class Mealy (state action out : Type) := {
|
||||
initial : state;
|
||||
step : state -> action -> state;
|
||||
output : state -> action -> out
|
||||
}.
|
||||
|
||||
(** ** Auxiliary functions
|
||||
|
||||
We define a [run] of a machine [M], which is an extension of
|
||||
the [step] function to lists of actions. We use a shortcut
|
||||
[do_actions] for running the machine from the [initial] state. *)
|
||||
|
||||
|
||||
Fixpoint run `{Mealy state action out} (s : state) (ls : list action) : state :=
|
||||
match ls with
|
||||
| nil => s
|
||||
| a::tl => run (step s a) tl
|
||||
end.
|
||||
|
||||
Definition do_actions `{Mealy state action out} : list action -> state := run initial.
|
||||
|
||||
(** The [test] function runs the required list of actions and examines the output of the resulting state on a specified action. *)
|
||||
Definition test `{Mealy state action out} (ls : list action) : action -> out := output (do_actions ls).
|
||||
|
||||
Section Rushby.
|
||||
|
||||
(** We assume for the rest of the formalisation that we have a Mealy
|
||||
machine [M]. Thus, we parameterize our main development module by a machine [M]. *)
|
||||
|
||||
|
||||
Context `{MM : Mealy state action out}.
|
||||
|
||||
|
||||
(** * Security policies *)
|
||||
Section Policy.
|
||||
|
||||
(** In order to define the notion of security, we assume a type [domain] of security domains. Those can be, e.g. clearance levels. For each action [a] we assign a domain [dom a] of that action. For instance, if we interpret domains as clearance levels, the domain of an action can signify the clearance level required to observe/preform the action.
|
||||
|
||||
A *security policy* is defined to be a reflexive decidable relation on the type of domains. For instance, the policy relation can state that the clearance level [u] is below the clearance level [v]. *)
|
||||
Class Policy (domain : Type) := {
|
||||
dom : action -> domain;
|
||||
|
||||
(** Sidenote: I had a lot of issues with Coq not finding relevant
|
||||
typeclass instances because the next field was declared as
|
||||
[domain_dec : TYPE] instead of [domain_dec :> TYPE]. So we actually
|
||||
do need implicit coercions to get automatic resolution of typeclass
|
||||
instances. *)
|
||||
|
||||
domain_dec :> forall (x y : domain), Decision (x = y);
|
||||
policy :> relation domain;
|
||||
policy_dec :> (forall v w, Decision (policy v w));
|
||||
policy_refl :> Reflexive policy
|
||||
}.
|
||||
|
||||
(** TODO: Make this notation span over different sections? *)
|
||||
Delimit Scope policy_scope with P.
|
||||
Open Scope policy_scope.
|
||||
Infix "⇝" := policy (at level 70) : policy_scope.
|
||||
|
||||
(** Quoting Rushby:
|
||||
|
||||
<<
|
||||
We wish to define security in terms of information flow, so the
|
||||
next step is to capture the idea of the "flow of information"
|
||||
formally. The key observation is that information can be said to
|
||||
flow from a domain [u] to a domain [v] exactly when the acions
|
||||
submitted by domain [u] cause the behavior of the system
|
||||
percieved by domain [v] to be different from that perceived when
|
||||
those actions are not present.
|
||||
>>
|
||||
|
||||
Hence, we define a function [purge] that removes from the given list
|
||||
all the actions that are not supposed to "influence" the domain [v].
|
||||
*)
|
||||
|
||||
Fixpoint purge `{Policy domain} (ls : list action) (v : domain) : list action :=
|
||||
match ls with
|
||||
| [] => []
|
||||
| a::tl => if (decide (dom a ⇝ v)) then a::(purge tl v)
|
||||
else purge tl v
|
||||
end.
|
||||
|
||||
|
||||
(** Then, the system is _secure_ w.r.t. a given policy if for any set
|
||||
of actions that can be performed, the system cannot tell the
|
||||
difference between preforming the given actions and only performing
|
||||
the actions that are supposed to influence the outcome *)
|
||||
|
||||
Definition security `{Policy domain} := forall (ls : list action) (a : action),
|
||||
test ls a = test (purge ls (dom a)) a.
|
||||
|
||||
|
||||
End Policy.
|
||||
|
||||
(** * View partitions
|
||||
|
||||
As we have seen, the non-interference notion of security is defined by
|
||||
quantifying over all the possible paths of the system. We wish to
|
||||
develop techniques for establishing the security property by putting
|
||||
conditions on individual state transformations.
|
||||
|
||||
As a first step, we define a notion of a "view partition", which is an
|
||||
equivalence relation on the type of states of the system, siginfying
|
||||
that two states are "identitcal" or indistinguishable for a given
|
||||
domain.
|
||||
|
||||
*)
|
||||
|
||||
Section view_partitions.
|
||||
|
||||
(** Formally, a view partition is an assignment of an equivalence
|
||||
relation [≈{u}] for every domain [u] *)
|
||||
|
||||
Class ViewPartition (domain : Type) := {
|
||||
view_partition :> domain -> relation state;
|
||||
view_partition_is_equiv :> forall v, Equivalence (view_partition v)
|
||||
}.
|
||||
|
||||
Notation "S ≈{ U } T" := (view_partition U S T)
|
||||
(at level 70, format "S ≈{ U } T") : policy_scope.
|
||||
Open Scope policy_scope.
|
||||
(** We say that a system is _output consistent_ if for every two
|
||||
states [s, t] that are indistinguishable w.r.t. the domain [dom a],
|
||||
the output of the system [output s a] is the same as [output t a] *)
|
||||
|
||||
Class OutputConsistent `{P : Policy domain} `(ViewPartition domain) :=
|
||||
output_consistent : (forall a s t, s ≈{dom a} t -> output s a = output t a).
|
||||
|
||||
(** Our first lemma states that if we have a view partitioned system
|
||||
that is output consistent and satisfies
|
||||
|
||||
[[(do_actions ls) ≈{u} (do_actions (purge ls u))]]
|
||||
|
||||
then the system is secure. *)
|
||||
|
||||
(* Lemma 1 *)
|
||||
Lemma output_consist_security `{P : Policy domain} `{VP : ViewPartition domain} {OC : @OutputConsistent domain P VP} : (forall ls u,
|
||||
(do_actions ls) ≈{u} (do_actions (purge ls u)))
|
||||
-> security.
|
||||
Proof.
|
||||
intros L ls a. apply output_consistent.
|
||||
apply (L _ (dom a)).
|
||||
Qed.
|
||||
|
||||
(** Note that the conditions of Lemma [output_consist_security] still
|
||||
require us to quantify over all possible paths of the system. We wish
|
||||
to replace this condition with the following two "local" conditions.
|
||||
|
||||
The first condition, stating that the view partitioned system _locally
|
||||
respects policy_: if the domain [dom a] is not suppose to interfere
|
||||
with the domain [u], then, from the point of view of [u], the states
|
||||
[s] and [step s a] are indistinguishable.
|
||||
*)
|
||||
|
||||
Definition locally_respects_policy `{Policy domain} `{ViewPartition domain} := forall a u s,
|
||||
¬(policy (dom a) u) -> s ≈{u} (step s a).
|
||||
|
||||
(** We say that the system is *step consistent*, if the view partition
|
||||
is closed under the [step] function. *)
|
||||
|
||||
Definition step_consistent `{Policy domain} `{ViewPartition domain} := forall a s t u,
|
||||
s ≈{u} t -> (step s a) ≈{u} (step t a).
|
||||
|
||||
(** Given those two conditions we can prove that the system is secure,
|
||||
by applying the previous lemma *)
|
||||
|
||||
(* Theorem 1 *)
|
||||
Theorem unwinding `{P: Policy domain} `{VP: ViewPartition domain} `{@OutputConsistent domain P VP} : locally_respects_policy -> step_consistent -> security.
|
||||
Proof.
|
||||
intros LRP SC. apply output_consist_security.
|
||||
assert (forall ls u s t, view_partition u s t -> view_partition u (run s ls) (run t (purge ls u))) as General. (* TODO: a simple generalize would not suffice, because we actually need the s ≈ t assumption *)
|
||||
induction ls; simpl; auto.
|
||||
intros u s t HI.
|
||||
destruct (decide (policy (dom a) u)).
|
||||
(* DOESNT WORK (Lexer) : destruct (decide ((dom a) ⇝ u)). *)
|
||||
(** Case [(dom a) ~> u] *)
|
||||
apply IHls. apply SC; assumption.
|
||||
(** Case [(dom a) ~/> u] *)
|
||||
apply IHls.
|
||||
transitivity s. symmetry. unfold locally_respects_policy in LRP. apply LRP; assumption. assumption.
|
||||
unfold do_actions. intros ls u. apply General. reflexivity.
|
||||
Qed.
|
||||
|
||||
End view_partitions.
|
||||
|
||||
(** * Access control interpretation of the view partition. *)
|
||||
|
||||
Section ACI.
|
||||
|
||||
(** In this section we consider a formalisation of the access control mechansisms.
|
||||
|
||||
We say that the machine has _structured state_ if we have a collection
|
||||
of [name]s and [value]s (the latter being decidable), and some sort of
|
||||
the memory assigned to each state, which is formally given by the
|
||||
[contents] function. For each domain [u] we assign sets [observe u]
|
||||
and [alter u] of names that are allowed to be observed and altered,
|
||||
respectively, by the domain [u]. *)
|
||||
|
||||
Class StructuredState (domain : Type) := {
|
||||
name : Type;
|
||||
value : Type;
|
||||
contents : state -> name -> value;
|
||||
value_dec :> forall (v1 v2 : value), Decision (v1 = v2);
|
||||
observe : domain -> FinSet name;
|
||||
alter : domain -> FinSet name
|
||||
}.
|
||||
|
||||
(** This induces the view partition relation as follows: two state [s]
|
||||
and [t] are indistinguishable for the given domain, if all the
|
||||
observable contents at [s] is the same as the observable content at
|
||||
[t] *)
|
||||
|
||||
Definition RMA_partition `{@StructuredState domain} (u : domain) s t := (forall n, (n ∈ observe u) -> contents s n = contents t n).
|
||||
|
||||
Transparent RMA_partition.
|
||||
|
||||
Instance RMA `{@StructuredState domain} : ViewPartition domain := { view_partition := RMA_partition }.
|
||||
(* begin hide *)
|
||||
intro u. split; unfold RMA_partition.
|
||||
(* Reflexivity *) unfold Reflexive. auto.
|
||||
(* Symmetry *) unfold Symmetric. intros x y Sy.
|
||||
symmetry. apply Sy. assumption.
|
||||
(* Transitivity *) unfold Transitive. intros x y z T1 T2 n.
|
||||
transitivity (contents y n); [apply T1 | apply T2]; assumption.
|
||||
Defined. (* We have to use 'Defined' here instead of 'Qed' so that we can unfold 'RMA' later on *)
|
||||
(* end hide *)
|
||||
|
||||
Hint Resolve RMA.
|
||||
|
||||
(** ** Reference monitor assumptions *)
|
||||
(* TODO: explain those assumptions *)
|
||||
Class RefMonAssumptions `{Policy domain} `{StructuredState domain} :=
|
||||
{ rma1 : forall (a : action) s t,
|
||||
view_partition (dom a) s t -> output s a = output t a
|
||||
; rma2 : forall a s t n,
|
||||
view_partition (dom a) s t ->
|
||||
((contents (step s a) n) ≠ (contents s n)
|
||||
∨ (contents (step t a) n) ≠ (contents t n))
|
||||
-> contents (step s a) n = contents (step t a) n
|
||||
; rma3 : forall a s n,
|
||||
contents (step s a) n ≠ contents s n -> n ∈ alter (dom a)
|
||||
}.
|
||||
|
||||
(** If the reference monitor assumptions are satisfied, then the system is output-consistent *)
|
||||
Global Instance OC `{RefMonAssumptions}: OutputConsistent RMA.
|
||||
exact rma1. Defined.
|
||||
|
||||
(** The reference mointor assumptions provide the security for the system, if, furthermore, two additional requirements are satisfied:
|
||||
|
||||
- [u ~> v] implies [observe u ⊆ observe v]; that is, if [u] is supposed to interfere with [v], then [v] can observe at least as much as [u]
|
||||
- if [n ∈ alter u] and [n ∈ observe v], then [u ~> v]; that is, if [u] is allowed to alter a location that is observable by [v], then [u] is allowed to interfere with [v]
|
||||
*)
|
||||
|
||||
(* Theorem 2 *)
|
||||
Theorem RMA_secutity `{RefMonAssumptions} :
|
||||
(forall u v, (policy u v) → observe u ⊆ observe v)
|
||||
-> (forall u v n, (n ∈ alter u) → (n ∈ observe v) → (policy u v))
|
||||
-> security.
|
||||
Proof.
|
||||
intros Cond1 Cond2. apply unwinding.
|
||||
(** We apply the unwinding theorem, so we have to verify that
|
||||
we locally respect policy and that we have step-consistency *)
|
||||
unfold locally_respects_policy.
|
||||
intros a u s.
|
||||
|
||||
(** In order to prove that the system locally respects the policy
|
||||
relation, we first state the contrapositive, which we can prove
|
||||
because the policy relation and type of values are decidable *)
|
||||
|
||||
assert ((exists n, (n ∈ observe u ∧ (contents s n ≠ contents (step s a) n))) -> policy (dom a) u) as CP.
|
||||
intros opH. destruct opH as [n [??]].
|
||||
eapply Cond2. eapply rma3. eauto. assumption.
|
||||
intros NPolicy.
|
||||
unfold view_partition, RMA, RMA_partition.
|
||||
(* TODO: why can't decide automatically pick the instance value_dec? *)
|
||||
intros. destruct (decide (contents s n = contents (step s a) n)) as [e|Ne]. assumption. exfalso. apply NPolicy. apply CP.
|
||||
exists n; split; assumption.
|
||||
|
||||
(** In order to prove step-consistency we wish to apply the second
|
||||
reference monitor assumption. For that we must distinguish three
|
||||
cases: (contents (step s a) n ≠ contents s n), (contents (step t a)
|
||||
n ≠ contents t n), or if both of the equalities hold *)
|
||||
intros a s t u A.
|
||||
unfold view_partition, RMA, RMA_partition in *.
|
||||
intros n nO.
|
||||
destruct (decide (contents (step s a) n = contents s n)) as [E1 | NE1].
|
||||
destruct (decide (contents (step t a) n = contents t n)) as [E2 | NE2].
|
||||
(* Both equalities hold *)
|
||||
rewrite E1, E2. apply A; assumption.
|
||||
(* NE2 *)
|
||||
(* We use the Second RM assmption to deal with this case *)
|
||||
apply rma2. (* for this we have to show that s ~_(dom a) t *)
|
||||
unfold view_partition, RMA, RMA_partition.
|
||||
intros m L. apply A.
|
||||
apply Cond1 with (u:=dom a).
|
||||
apply Cond2 with (n:=n); [eapply rma3 | ]; eassumption.
|
||||
assumption.
|
||||
right. auto.
|
||||
(* NE1 case is similar *)
|
||||
(* TODO: repetition *)
|
||||
apply rma2. (* for this we have to show that s ~_(dom a) t *)
|
||||
unfold view_partition, RMA, RMA_partition.
|
||||
intros m L. apply A.
|
||||
apply Cond1 with (u:=dom a).
|
||||
apply Cond2 with (n:=n); [eapply rma3| ]; eassumption.
|
||||
assumption.
|
||||
left. auto.
|
||||
Qed.
|
||||
|
||||
End ACI.
|
||||
|
||||
(** * Intransitive security policy *)
|
||||
Section Intransitive.
|
||||
|
||||
(** Auxiliary definitions *)
|
||||
Fixpoint sources `{Policy domain} (ls : list action) (d : domain) : FinSet domain :=
|
||||
match ls with
|
||||
| [] => {[ d ]}
|
||||
| a::tl => let src := sources tl d in
|
||||
if (decide (exists (v: domain), ((v ∈ src) ∧ (policy (dom a) v))))
|
||||
then src ∪ {[ dom a ]} else src
|
||||
end.
|
||||
|
||||
(** The two properties of the [sources] function: it is monotone and [d \in sources ls d] *)
|
||||
Lemma sources_mon `{Policy} : forall a ls d, sources ls d ⊆ sources (a::ls) d.
|
||||
Proof.
|
||||
intros. simpl.
|
||||
destruct (decide _); [apply union_subseteq_l |]; auto.
|
||||
Qed.
|
||||
|
||||
Hint Resolve sources_mon.
|
||||
|
||||
Lemma sources_monotone `{Policy} : forall ls js d, ls `sublist` js → sources ls d ⊆ sources js d.
|
||||
Proof.
|
||||
intros ls js d M.
|
||||
induction M. simpl. reflexivity.
|
||||
simpl. destruct (decide (∃ v : domain, v ∈ sources l1 d ∧ policy (dom x) v)); destruct (decide (∃ v : domain, v ∈ sources l2 d ∧ policy (dom x) v)).
|
||||
apply union_preserving_r. assumption.
|
||||
exfalso. apply n. destruct e as [v [e1 e2]]. exists v; split; try (apply (IHM v)); assumption.
|
||||
transitivity (sources l2 d). assumption. apply union_subseteq_l.
|
||||
assumption.
|
||||
transitivity (sources l2 d); auto.
|
||||
Qed.
|
||||
|
||||
Lemma sources_in `{Policy} : forall ls d, d ∈ sources ls d.
|
||||
Proof.
|
||||
intros. induction ls; simpl. apply elem_of_singleton_2; auto.
|
||||
destruct (decide (∃ v : domain, v ∈ sources ls d ∧ policy (dom a) v)); simpl.
|
||||
apply elem_of_union_l. assumption. assumption.
|
||||
Qed.
|
||||
|
||||
Hint Resolve sources_in.
|
||||
|
||||
Fixpoint ipurge `{Policy} (ls : list action) (d : domain) :=
|
||||
match ls with
|
||||
| [] => []
|
||||
| a::tl => if (decide ((dom a) ∈ sources ls d))
|
||||
then a::(ipurge tl d)
|
||||
else ipurge tl d
|
||||
end.
|
||||
|
||||
(** The non-interference notion of security for intransitive policies
|
||||
is very similar to the transitive case, bu uses the [ipurge] function
|
||||
instead of [purge] *)
|
||||
|
||||
Definition isecurity `{Policy} := forall ls a,
|
||||
test ls a = test (ipurge ls (dom a)) a.
|
||||
|
||||
(** We can prove lemmas similar to the transitive case *)
|
||||
|
||||
Lemma output_consist_isecurity `{P : Policy domain} `{VP : ViewPartition domain} {OC : @OutputConsistent domain P VP} : (forall ls u,
|
||||
view_partition u (do_actions ls) (do_actions (ipurge ls u)))
|
||||
-> isecurity.
|
||||
Proof.
|
||||
unfold isecurity. intros H ls a.
|
||||
unfold test.
|
||||
apply output_consistent.
|
||||
apply (H ls (dom a)).
|
||||
Qed.
|
||||
|
||||
Definition view_partition_general `{ViewPartition domain} (C : FinSet domain) s t := forall (u: domain), u ∈ C -> view_partition u s t.
|
||||
|
||||
Global Instance view_partition_general_equiv `{ViewPartition domain}:
|
||||
forall V, Equivalence (view_partition_general V).
|
||||
Proof.
|
||||
intro V. split.
|
||||
intros x v A. reflexivity.
|
||||
intros x y A1 v A2. symmetry. apply (A1 v); assumption.
|
||||
intros x y z A1 A2 v A3. transitivity y. apply (A1 v); assumption. apply (A2 v); assumption.
|
||||
Qed.
|
||||
|
||||
Definition weakly_step_consistent `{Policy domain} `{ViewPartition domain} :=
|
||||
forall s t u a, view_partition u s t -> view_partition (dom a) s t -> view_partition u (step s a) (step t a).
|
||||
|
||||
Ltac exists_inside v :=
|
||||
let H := fresh "Holds" in
|
||||
let nH := fresh "notHolds" in
|
||||
destruct (decide _) as [H | []];
|
||||
[ try reflexivity | exists v; try auto].
|
||||
|
||||
Local Hint Resolve sources_mon.
|
||||
Local Hint Extern 1 (_ ∈ sources (_::_) _) => eapply sources_mon.
|
||||
Local Hint Immediate elem_of_singleton.
|
||||
(* Local Hint Extern 1 (_ ∈ {[ _ ]}) => apply elem_of_singleton; trivial. *)
|
||||
Local Hint Resolve elem_of_union.
|
||||
|
||||
(* Lemma 3 *)
|
||||
Lemma weakly_step_consistent_general `{Policy domain} `{ViewPartition domain} (s t : state) (a : action) ls (u: domain) : weakly_step_consistent -> locally_respects_policy ->
|
||||
view_partition_general (sources (a::ls) u) s t
|
||||
-> view_partition_general (sources ls u) (step s a) (step t a).
|
||||
Proof.
|
||||
intros WSC LRP P v vIn.
|
||||
unfold view_partition_general in P.
|
||||
unfold locally_respects_policy in LRP.
|
||||
destruct (decide (policy (dom a) v)).
|
||||
(* Case [dom a ~> v] *)
|
||||
apply WSC; apply P. auto.
|
||||
(* we need to show that [dom a ∈ sources (a::ls) v] *)
|
||||
simpl. exists_inside v. apply elem_of_union. right. auto. apply elem_of_singleton; trivial.
|
||||
(* Case [dom a ~/> v] *)
|
||||
transitivity s. symmetry. apply LRP; assumption.
|
||||
transitivity t. apply P. auto.
|
||||
apply LRP; assumption.
|
||||
Qed.
|
||||
|
||||
(* Lemma 4 *)
|
||||
Lemma locally_respects_gen `{Policy domain} `{ViewPartition domain} (WSC: weakly_step_consistent) (LRP : locally_respects_policy) s a ls u :
|
||||
¬ ((dom a) ∈ sources (a::ls) u) ->
|
||||
view_partition_general (sources ls u) s (step s a).
|
||||
Proof.
|
||||
intros domN v vIn.
|
||||
apply LRP. intro.
|
||||
(* If [dom a ~> v], then [dom a ∈ sources (a::ls) u], because [v ∈ sources ls u] *) (* again, a clusterfuck *)
|
||||
apply domN. simpl.
|
||||
exists_inside v; auto. apply elem_of_union; right; apply elem_of_singleton; reflexivity.
|
||||
Qed.
|
||||
|
||||
|
||||
(* Lemma 5 *)
|
||||
Lemma unwinding_gen `{Policy domain} `{ViewPartition domain} (WSC: weakly_step_consistent) (LRP : locally_respects_policy) s t ls u :
|
||||
view_partition_general (sources ls u) s t
|
||||
-> view_partition u (run s ls) (run t (ipurge ls u)).
|
||||
Proof.
|
||||
generalize dependent s.
|
||||
generalize dependent t.
|
||||
induction ls; intros s t.
|
||||
simpl. intro A. apply (A u). apply elem_of_singleton; reflexivity.
|
||||
intro VPG. simpl. unfold sources. fold (sources (a::ls) u).
|
||||
|
||||
destruct (decide _).
|
||||
(** Case [dom a ∈ sources (a::ls) u] *)
|
||||
simpl. apply IHls. apply weakly_step_consistent_general; auto.
|
||||
(** Case [dom a ∉ sources (a::ls) u] *)
|
||||
apply IHls. symmetry. transitivity t.
|
||||
- intros v vIn. symmetry. apply VPG. apply sources_mon; exact vIn.
|
||||
- apply locally_respects_gen; try(assumption).
|
||||
Qed.
|
||||
|
||||
|
||||
Theorem unwinding_intrans `{P : Policy domain} `{VP : ViewPartition domain}
|
||||
`{OC : @OutputConsistent domain P VP} (WSC : weakly_step_consistent) (LRP : locally_respects_policy) : isecurity.
|
||||
Proof.
|
||||
apply output_consist_isecurity.
|
||||
intros.
|
||||
apply unwinding_gen; try assumption.
|
||||
reflexivity.
|
||||
Qed.
|
||||
|
||||
Theorem rma_secure_intransitive `{RefMonAssumptions} : (forall u v n, n ∈ alter u -> n ∈ observe v -> policy u v) -> isecurity.
|
||||
Proof. intro policyA.
|
||||
apply @unwinding_intrans with (VP:=RMA). exact rma1.
|
||||
intros s t u a A1 A2.
|
||||
unfold view_partition. unfold RMA_partition; simpl. intros n L.
|
||||
destruct (decide (contents (step s a) n = contents s n)) as [E1 | NE1].
|
||||
destruct (decide (contents (step t a) n = contents t n)) as [E2 | NE2].
|
||||
(* Case [contents (step s a) n = contents s n /\ contents (step t a) n = contents t n] *)
|
||||
rewrite E1, E2. apply A1; assumption.
|
||||
(* Case [contents (step t a) n ≠ contents t n] *)
|
||||
apply rma2; [ | right]; assumption.
|
||||
(* Case [contents (step s a) n ≠ contents s n] *)
|
||||
apply rma2; [ | left]; assumption.
|
||||
|
||||
intros a u s A.
|
||||
assert ((exists n, (n ∈ observe u ∧ (contents s n ≠ contents (step s a) n))) -> policy (dom a) u) as CP.
|
||||
intros opH. destruct opH as [n [??]].
|
||||
eapply policyA. eapply rma3. eauto. assumption.
|
||||
intros n L. destruct (decide (contents s n = contents (step s a) n)). assumption. exfalso. apply A.
|
||||
apply CP. exists n. auto.
|
||||
Qed.
|
||||
|
||||
End Intransitive.
|
||||
|
||||
End Rushby.
|
|
@ -0,0 +1,184 @@
|
|||
(* rewrite using the view partition relation
|
||||
TODO *)
|
||||
|
||||
Require Import base.
|
||||
Require Import Mealy.
|
||||
Require Import Policy.
|
||||
Require Import ViewPartition.
|
||||
|
||||
Class IsSecure {domain : Type}
|
||||
`(M : Mealy state action out)
|
||||
(P : Policy action domain) :=
|
||||
noninterference :
|
||||
(forall (ls : list action) (a : action), test ls a = test (purge ls (dom a)) a).
|
||||
|
||||
Section MealySec.
|
||||
Variable (state action out : Type).
|
||||
Variable (M : Mealy state action out).
|
||||
Variable domain : Type.
|
||||
Variable (P : Policy action domain).
|
||||
Variable (VP : ViewPartition M domain).
|
||||
Lemma output_consist_view_security {OC : OutputConsistent P VP} : (forall ls u,
|
||||
(do_actions ls) ≈{u} (do_actions (purge ls u)))
|
||||
-> IsSecure M P.
|
||||
Proof.
|
||||
intros L ls a. apply output_consistent.
|
||||
apply (L _ (dom a)).
|
||||
Qed.
|
||||
|
||||
Instance unwinding {OC : OutputConsistent P VP} {LRP : LocallyRespectsPolicy P VP} {SC: StepConsistent P VP} : IsSecure M P.
|
||||
Proof.
|
||||
apply output_consist_view_security.
|
||||
assert (forall ls v s t, view_partition v s t -> view_partition v (run s ls) (run t (purge ls v))) as General. (* TODO: a simple generalize would not suffice, because we actually need the s ≈ t assumption *)
|
||||
induction ls; simpl; auto.
|
||||
intros u s t HI.
|
||||
destruct (decide (policy (dom a) u)).
|
||||
(* DOESNT WORK (Lexer) : destruct (decide ((dom a) ⇝ u)). *)
|
||||
(** Case [(dom a) ~> u] *)
|
||||
eapply IHls. apply step_consistent. assumption.
|
||||
(** Case [(dom a) ~/> u] *)
|
||||
eapply IHls.
|
||||
transitivity s; [| assumption]. symmetry. apply locally_respects_policy. assumption.
|
||||
unfold do_actions. intros ls u. eapply General; reflexivity.
|
||||
Qed.
|
||||
|
||||
Instance unwinding_trans {OC : OutputConsistent P VP} {LRP : LocallyRespectsPolicy P VP} {TP: Transitive policy} {SC: WeakStepConsistent P VP} : IsSecure M P.
|
||||
Proof.
|
||||
apply output_consist_view_security.
|
||||
assert (forall ls u s t, (forall v, policy v u -> view_partition v s t) -> (forall v, policy v u -> view_partition v (run s ls) (run t (purge ls v)))) as General.
|
||||
induction ls; simpl; auto.
|
||||
intros u s t S v HI.
|
||||
destruct (decide (policy (dom a) v)).
|
||||
(** Case [(dom a) ~> u] *)
|
||||
apply IHls with (u:=u); try assumption.
|
||||
intros v' Hv'.
|
||||
apply weak_step_consistent; apply S. etransitivity; eassumption.
|
||||
assumption.
|
||||
(** Case [(dom a) ~/> u] *)
|
||||
apply IHls with (u:=v); try assumption.
|
||||
intros v' Hv'.
|
||||
transitivity s. symmetry.
|
||||
apply locally_respects_policy. intro.
|
||||
apply n. rewrite H. assumption.
|
||||
apply S. rewrite Hv'. assumption. reflexivity.
|
||||
unfold do_actions. intros ls u. eapply General; reflexivity.
|
||||
Qed.
|
||||
|
||||
End MealySec.
|
||||
|
||||
Module MSyncSec.
|
||||
Variable (state action out : Type).
|
||||
Variable (M : MSync state action out).
|
||||
Definition M' : Mealy state action out := msync_mealy.
|
||||
Variable domain : Type.
|
||||
Variable (P : Policy action domain).
|
||||
Variable (VP' : ViewPartition M' domain).
|
||||
|
||||
Instance VP: ViewPartition (MSync_as_Mealy (MS:=M)) domain
|
||||
:=
|
||||
{ view_partition := @view_partition _ _ _ _ _ VP'
|
||||
; view_partition_is_equiv := @view_partition_is_equiv _ _ _ _ _ VP'
|
||||
}.
|
||||
|
||||
Class SyncPartitioned :=
|
||||
{ sync_partitioned : forall s t a, view_partition (dom a) s t -> (delayed a s <-> delayed a t)
|
||||
}.
|
||||
|
||||
|
||||
Global Instance oc2 {SP: SyncPartitioned} (OC: OutputConsistent P VP') :
|
||||
OutputConsistent P VP.
|
||||
Proof.
|
||||
intros a s t L.
|
||||
unfold output; simpl.
|
||||
destruct (decide (delayed a s)) as [D | ND];
|
||||
destruct (decide (delayed a t)) as [D' | ND'];
|
||||
try solve
|
||||
[erewrite output_consistent by (apply L);
|
||||
reflexivity | reflexivity ].
|
||||
exfalso. apply ND'. rewrite <- (sync_partitioned s t a L). assumption.
|
||||
exfalso. apply ND. rewrite -> (sync_partitioned s t a L). assumption.
|
||||
Qed.
|
||||
|
||||
Theorem unwinding_sync_trans {SP: SyncPartitioned} {OC : OutputConsistent P VP'} {LRP : LocallyRespectsPolicy P VP'} {TP : Transitive policy} {SC: WeakStepConsistent P VP'} : IsSecure (MSync_as_Mealy (MS:=M)) P.
|
||||
Proof.
|
||||
intros.
|
||||
apply unwinding_trans with (VP:=VP). (* TODO: i want this to be resolved automatically *) apply oc2. assumption.
|
||||
|
||||
(* Locally respects policy *)
|
||||
intros a u s NP. unfold step; simpl.
|
||||
destruct (decide (delayed a s)); auto. reflexivity.
|
||||
|
||||
(* Transitive policy *) assumption. (* again, can this be picked up automatically? *)
|
||||
|
||||
(* Step-consitency *)
|
||||
intros a s t u Ldom L. unfold step; simpl.
|
||||
destruct (decide (delayed a s)) as [D | ND];
|
||||
destruct (decide (delayed a t)) as [D' | ND'];
|
||||
try assumption.
|
||||
(* only s is delayed by a *)
|
||||
exfalso. apply ND'. rewrite sync_partitioned; try eassumption.
|
||||
symmetry; assumption.
|
||||
(* similar case, but symmetric *)
|
||||
exfalso. apply ND. rewrite sync_partitioned; try eassumption.
|
||||
(* both s and t are delayed by a *)
|
||||
eapply weak_step_consistent; assumption.
|
||||
Qed.
|
||||
|
||||
End MSyncSec.
|
||||
|
||||
Module MApiSec.
|
||||
Variable (state action out : Type).
|
||||
Variable domain : Type.
|
||||
Variable API : Type.
|
||||
Variable (P : Policy API domain).
|
||||
Variable (M : MApi domain API state action out).
|
||||
Definition M' : Mealy state action out := mapi_mealy.
|
||||
Variable (VP' : ViewPartition M' domain).
|
||||
|
||||
Definition vpeq : domain -> relation (state * (domain -> list action)) := fun d s1 s2 =>
|
||||
let (s1', f) := s1
|
||||
in let (s2', g) := s2
|
||||
in @view_partition _ _ _ _ _ VP' d s1' s2' /\ f = g.
|
||||
|
||||
Lemma vpeq_eq (d : domain) : Equivalence (vpeq d).
|
||||
Proof.
|
||||
intros. split; unfold vpeq.
|
||||
intros [s ?]. split; reflexivity.
|
||||
|
||||
intros [s ?] [t ?] [? ?]. split; symmetry; auto.
|
||||
|
||||
intros [s ?] [t ?] [z ?] [? ?] [? ?]. split; etransitivity; eassumption.
|
||||
Qed.
|
||||
|
||||
Instance VP: ViewPartition (MApi_as_Mealy (MA:=M)) domain
|
||||
:=
|
||||
{ view_partition := vpeq
|
||||
; view_partition_is_equiv := vpeq_eq
|
||||
}.
|
||||
|
||||
Instance: IsSecure (MApi_as_Mealy (MA:=M)) P.
|
||||
Proof. intros. apply unwinding_trans with (VP:=VP).
|
||||
|
||||
(* output consistency *)
|
||||
intro. intros [s f] [t g].
|
||||
simpl.
|
||||
remember (f (dom_api a)) as L1;
|
||||
remember (g (dom_api a)) as L2.
|
||||
destruct L1;
|
||||
destruct L2; try (subst; reflexivity).
|
||||
|
||||
intros [? Z]. exfalso. rewrite Z in *.
|
||||
rewrite <- HeqL2 in *.
|
||||
inversion HeqL1.
|
||||
|
||||
intros [? Z]. exfalso. rewrite Z in *.
|
||||
rewrite <- HeqL2 in *.
|
||||
inversion HeqL1.
|
||||
|
||||
intros [? Z]. f_equal.
|
||||
rewrite Z in *.
|
||||
rewrite <- HeqL1 in *.
|
||||
inversion HeqL2.
|
||||
apply output_consistency.
|
||||
(* TODO replace dom in policy with the Dom class *)
|
||||
End MApiSec.
|
|
@ -0,0 +1,29 @@
|
|||
Require Import base.
|
||||
Require Import Policy.
|
||||
Require Import Mealy.
|
||||
(** Formally, a view partition is an assignment of an equivalence
|
||||
relation [≈{u}] for every domain [u] *)
|
||||
|
||||
Class ViewPartition `(M : Mealy state action out) (domain : Type) := {
|
||||
view_partition :> domain -> relation state;
|
||||
view_partition_is_equiv :> forall v, Equivalence (view_partition v)
|
||||
}.
|
||||
|
||||
Notation "S ≈{ U } T" := (view_partition U S T)
|
||||
(at level 70, format "S ≈{ U } T").
|
||||
|
||||
Class OutputConsistent {domain : Type} `{M : Mealy state action out}
|
||||
(P : Policy action domain) (VP: ViewPartition M domain) :=
|
||||
output_consistent : (forall a s t, s ≈{dom a} t -> output s a = output t a).
|
||||
|
||||
Class LocallyRespectsPolicy {domain : Type} `{M : Mealy state action out}
|
||||
(P : Policy action domain) (VP: ViewPartition M domain) :=
|
||||
locally_respects_policy : (forall a u s, ¬(policy (dom a) u) -> s ≈{u} (step s a)).
|
||||
|
||||
Class WeakStepConsistent {domain : Type} `{M : Mealy state action out}
|
||||
(P : Policy action domain) (VP: ViewPartition M domain) :=
|
||||
weak_step_consistent : (forall a s t u, s ≈{dom a} t -> s ≈{u} t -> (step s a) ≈{u} (step t a)).
|
||||
|
||||
Class StepConsistent {domain : Type} `{M : Mealy state action out}
|
||||
(P : Policy action domain) (VP: ViewPartition M domain) :=
|
||||
step_consistent : (forall a s t u, s ≈{u} t -> (step s a) ≈{u} (step t a)).
|
Loading…
Reference in New Issue