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

View File

@ -1,7 +1,8 @@
Require Import Rushby.
From stdpp Require Import list relations collections fin_collections.
(** Instantiation of the intransitive interference with an "array machine" *)
Require Import NI.Rushby.
From stdpp Require Import list relations gmap sets fin_sets.
Module ArrayMachine <: Mealy.
Module ArrayMachine.
Definition state := nat -> nat.
Inductive command :=
@ -23,36 +24,48 @@ Definition preform (s : state) (a : action) : state * out :=
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.
Instance ArrayMealyMachine : Mealy state action out :=
{ initial := initial;
step := step;
output := output }.
Eval compute in (do_actions [Write 1 1] 2).
(** ===> 0 *)
Definition domain := action.
Definition nameA := nat.
Definition valA := nat.
Definition observeA (u : domain) : FinSet nameA :=
Definition observeA (u : domain) : gset nameA :=
match u with
| Read i => {[ i ]}
| Write _ _ =>
end.
Definition alterA (u : domain) : FinSet valA :=
Definition alterA (u : domain) : gset valA :=
match u with
| Read _ =>
| Write i _ => {[ i ]}
end.
Instance domain_dec : forall (u v : domain), Decision (u = v).
Proof. intros.
Instance domain_dec : EqDecision domain.
Proof. intros u v.
unfold Decision.
repeat (decide equality).
Defined.
Instance domain_countable : Countable domain.
refine (inj_countable' (λ x, match x with
| Write i1 i2 => (inl (i1, i2))
| Read i => (inr i)
end) (λ x, match x with
| inl (i1, i2) => Write i1 i2
| inr i => Read i
end) _); by intros [].
Defined.
Instance arraymachine_ss : StructuredState domain :=
{ name := nameA; value := valA; contents s n := s n
; observe := observeA; alter := alterA }.
@ -64,59 +77,64 @@ 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: Set_ nameA (gset nameA).
apply _.
Qed.
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.
- intros v w.
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 HHH).
right. intro. inversion H; subst. apply n0. auto.
inversion H0. unfold alterA, observeA in *. destruct H1 as [HH HHH].
apply (not_elem_of_empty x HHH).
+ 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. eapply (not_elem_of_empty x); eassumption.
+ 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 x); 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.
Proof.
split; simpl; unfold RMA_partition.
- intros a s t Hst;
unfold contents in *; simpl in *.
unfold output.
unfold preform. destruct a as [i j | i]; simpl in *. reflexivity.
apply Hst. eapply elem_of_singleton. reflexivity.
- intros a s t n Hst Hn.
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 Hn as [Hn|Hn].
unfold step in Hn; simpl in Hn. unfold extendS in Hn; simpl in Hn.
replace (beq_nat n i) with false in *; auto.
congruence. symmetry. apply beq_nat_false_iff. omega.
unfold step in Hn; simpl in Hn. unfold extendS in Hn; simpl in Hn.
replace (beq_nat n i) with false in *; auto.
congruence. symmetry. apply beq_nat_false_iff. omega.
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;
simpl. destruct Hn as [Hn|Hn]; unfold step, preform in Hn; simpl in Hn;
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.
- intros a s n Hst.
unfold step, preform in Hst; simpl in Hst. destruct a; simpl in *.
unfold extendS in Hst. 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.
@ -125,4 +143,3 @@ Proof.
apply rma_secure_intransitive.
intros u v n A1 A2. simpl. right. firstorder.
Qed.

View File

@ -5,20 +5,7 @@
(** printing -> #→# *)
(** printing (policy a b) #a ⇝ b# *)
From stdpp 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 : EqDecision A), FinCollection A (FinSet A)}.
(** end hide **)
From stdpp Require Import list relations gmap sets fin_sets.
(** * Mealy machines *)
@ -73,6 +60,7 @@ Class Policy (domain : Type) := {
instances. *)
domain_dec :> EqDecision domain;
domain_countable :> Countable domain;
policy :> relation domain;
policy_dec :> RelDecision policy;
policy_refl :> Reflexive policy
@ -139,7 +127,7 @@ 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)
view_partition_is_equiv :> forall v, Equivalence (view_partition v);
}.
Notation "S ≈{ U } T" := (view_partition U S T)
@ -227,8 +215,10 @@ Class StructuredState (domain : Type) := {
value : Type;
contents : state -> name -> value;
value_dec :> EqDecision value;
observe : domain -> FinSet name;
alter : domain -> FinSet name
name_dec :> EqDecision name;
name_countable :> Countable name;
observe : domain -> gset name;
alter : domain -> gset name
}.
(** This induces the view partition relation as follows: two state [s]
@ -240,13 +230,15 @@ Definition RMA_partition `{@StructuredState domain} (u : domain) s t := (forall
Transparent RMA_partition.
Instance RMA `{@StructuredState domain} : ViewPartition domain := { view_partition := RMA_partition }.
Instance RMA `{!StructuredState domain}
`{!EqDecision domain, !Countable 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.
intro u. split; try apply _; 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 *) 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 *)
@ -255,16 +247,18 @@ 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)
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 *)
@ -339,7 +333,7 @@ End ACI.
Section Intransitive.
(** Auxiliary definitions *)
Fixpoint sources `{Policy domain} (ls : list action) (d : domain) : FinSet domain :=
Fixpoint sources `{Policy domain} (ls : list action) (d : domain) : gset domain :=
match ls with
| [] => {[ d ]}
| a::tl => let src := sources tl d in
@ -408,9 +402,13 @@ Proof.
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.
Definition view_partition_general
`{!ViewPartition domain, !EqDecision domain, !Countable domain}
(C : gset domain) s t
:= forall (u: domain), u C -> view_partition u s t.
Global Instance view_partition_general_equiv `{ViewPartition domain}:
Global Instance view_partition_general_equiv
`{ViewPartition domain, !EqDecision domain, !Countable domain}:
forall V, Equivalence (view_partition_general V).
Proof.
intro V. split.

View File

@ -1,2 +1,3 @@
-R . ""
-R . NI
Rushby.v
ArrayMachine.v