This commit is contained in:
2019-06-12 18:29:21 +02:00
parent 8601f7c673
commit 4dd4e4f383
9 changed files with 144 additions and 128 deletions

99
experimental/Mealy.v Normal file
View File

@@ -0,0 +1,99 @@
From stdpp 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
}.

74
experimental/MealySync.v Normal file
View File

@@ -0,0 +1,74 @@
Require Import Rushby.
From stdpp 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 action domain.
Variable VP : @ViewPartition action 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).
End MealyM.
Section Unwinding.
Context {state action out : Type}.
Context {domain : Type}.
Context {P : @Policy action domain}.
Context {VP : @ViewPartition state domain}.
Context {MMs : MSync state action out}.
Definition MM := msync_mealy.
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

30
experimental/Monoids.v Normal file
View File

@@ -0,0 +1,30 @@
From stdpp 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.

19
experimental/Policy.v Normal file
View File

@@ -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.

184
experimental/Security.v Normal file
View File

@@ -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.

View File

@@ -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)).