rushby-noninterference/ArrayMachine.v

146 lines
4.7 KiB
Coq
Raw Permalink Normal View History

2019-06-12 18:29:21 +02:00
(** Instantiation of the intransitive interference with an "array machine" *)
Require Import NI.Rushby.
From stdpp Require Import list relations gmap sets fin_sets.
2018-02-14 12:55:20 +01:00
2019-06-12 18:29:21 +02:00
Module ArrayMachine.
2018-02-14 12:55:20 +01:00
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.
2019-06-12 18:29:21 +02:00
End ArrayMachine.
2018-02-14 12:55:20 +01:00
Import ArrayMachine.
2019-06-12 18:29:21 +02:00
Instance ArrayMealyMachine : Mealy state action out :=
{ initial := initial;
step := step;
output := output }.
2018-02-14 12:55:20 +01:00
Eval compute in (do_actions [Write 1 1] 2).
2019-06-12 18:29:21 +02:00
(** ===> 0 *)
2018-02-14 12:55:20 +01:00
Definition domain := action.
Definition nameA := nat.
Definition valA := nat.
2019-06-12 18:29:21 +02:00
Definition observeA (u : domain) : gset nameA :=
2018-02-14 12:55:20 +01:00
match u with
| Read i => {[ i ]}
| Write _ _ =>
end.
2019-06-12 18:29:21 +02:00
Definition alterA (u : domain) : gset valA :=
2018-02-14 12:55:20 +01:00
match u with
| Read _ =>
| Write i _ => {[ i ]}
end.
2019-06-12 18:29:21 +02:00
Instance domain_dec : EqDecision domain.
Proof. intros u v.
2018-02-14 12:55:20 +01:00
unfold Decision.
repeat (decide equality).
Defined.
2019-06-12 18:29:21 +02:00
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 :=
2018-02-14 12:55:20 +01:00
{ name := nameA; value := valA; contents s n := s n
2019-06-12 18:29:21 +02:00
; observe := observeA; alter := alterA }.
2018-02-14 12:55:20 +01:00
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.
2019-06-12 18:29:21 +02:00
Instance: Set_ nameA (gset nameA).
apply _.
Qed.
Instance policy_ss : Policy domain :=
2018-02-14 12:55:20 +01:00
{ policy := interferenceR
; dom := fun (a : action) => (a : domain) }.
Proof.
2019-06-12 18:29:21 +02:00
- 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.
2018-02-14 12:55:20 +01:00
- intro u. constructor.
Defined.
Instance rma_yay : RefMonAssumptions.
2019-06-12 18:29:21 +02:00
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.
simpl. destruct Hn as [Hn|Hn]; unfold step, preform in Hn; simpl in Hn;
2018-02-14 12:55:20 +01:00
congruence.
2019-06-12 18:29:21 +02:00
- 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.
2018-02-14 12:55:20 +01:00
congruence.
Defined.
Theorem yay : isecurity.
Proof.
apply rma_secure_intransitive.
intros u v n A1 A2. simpl. right. firstorder.
Qed.