Library Common
Require Export List.
Require Export ListSet.
Require Export Permutation.
Require Export Setoid.
Require Export Basic.
Require Export ListSet.
Require Export Permutation.
Require Export Setoid.
Require Export Basic.
Record DecType : Type :=
{ t :> Type;
eq_dec : forall x y:t, {x = y} + {x <> y} }.
Section DecidableType.
Variable M : DecType.
Definition eqb (x y:M) := if (eq_dec x y) then true else false.
Notation "x '=?' y" := (eqb x y).
Lemma eqb_eq : forall x y, (x =? y) = true <-> x = y.
Lemma eqb_neq : forall x y, (x =? y) = false <-> x <> y.
Lemma eqb_refl : forall x, (x =? x) = true.
Lemma eqb_sym : forall x y, (x =? y) = (y =? x).
Lemma eqb_trans : forall x y z,
(x =? y) = true -> (y =? z) = true -> (x =? z) = true.
Lemma eqb_ntrans : forall x y z,
(x =? y) = true -> (y =? z) = false -> (x =? z) = false.
Lemma eqb_ntrans' : forall x y z,
(x =? y) = false -> (y =? z) = true -> (x =? z) = false.
Lemma DecType_eq : forall T (t:M) (A B:T), (if eq_dec t t then A else B) = A.
Lemma DecType_neq : forall T (t t':M) (A B:T), t <> t' ->
(if eq_dec t t' then A else B) = B.
End DecidableType.
Notation "x '=?' y" := (eqb _ x y).
Section UnitType.
Lemma unit_dec : forall (x y:unit), {x = y} + {x <> y}.
Definition Unit := {| t := unit; eq_dec := unit_dec |}.
End UnitType.
Section DecProd.
Variable A B:DecType.
Lemma DP_eq_dec : forall (p1 p2:(A*B)), {p1 = p2} + {p1 <> p2}.
Definition DecProd := {| t := A*B; eq_dec := DP_eq_dec |}.
End DecProd.
Labels
We require that there be exactly two labels (left and right), as this is the default practice in many choreography languages. This is also a decidable type.Section Labels.
Inductive label : Type :=
| left : label
| right : label
.
Lemma eq_label_dec : forall (l l' : label), { l = l' } + { l <> l' }.
Definition Label := {| t := label; eq_dec := eq_label_dec |}.
End Labels.
Local states
A local state is the state of a particular process. It maps the values in the process's local set of variables to values. The set of variables needs to be decidable.Definition Leq_state (s s': LState) : Prop := forall x, s x = s' x.
Notation "s [=] s'" := (Leq_state s s') (at level 50).
Lemma Leq_state_neq : forall x s s', s x <> s' x -> ~ s [=] s'.
Lemma Leq_state_refl : reflexive _ Leq_state.
Lemma Leq_state_sym : symmetric _ Leq_state.
Lemma Leq_state_trans : transitive _ Leq_state.
Add Parametric Relation : LState Leq_state
reflexivity proved by Leq_state_refl
symmetry proved by Leq_state_sym
transitivity proved by Leq_state_trans
as Leq_state_rel.
Definition Lupdate (s:LState) (x:Var) (v:Value) : LState :=
fun (y:Var) => if (x =? y) then v else (s y).
Lemma Lupdate_read : forall s x v, Lupdate s x v x = v.
Lemma Lupdate_read'' : forall s x y v, x <> y ->
Lupdate s x v y = s y.
Lemma Lupdate_update : forall s x v w,
Lupdate (Lupdate s x w) x v [=] Lupdate s x v.
Lemma Lupdate_independent' : forall s x y v v', x <> y ->
Lupdate (Lupdate s y v') x v [=] Lupdate (Lupdate s x v) y v'.
End LState.
Notation "s [=] s'" := (Leq_state _ _ s s') (at level 50).
Global states
A global state is the state of a choreography, or of a network as a whole: it maps each process name to its state. Note that all processes are required to use the same sets of variables and values.Equivalence of states
We now define equivalence up to a set of processes: we are often only interested in some of those.Definition eq_state (P : list Pid) (s: State) (s': State) : Prop :=
forall p, In p P -> (s p) [=] (s' p).
Notation "s [= P =] s'" := (eq_state P s s') (at level 50).
Lemma eq_state_neq : forall p P (s s':State),
~ s p [=] s' p -> In p P -> ~ s [= P =] s'.
Lemma eq_state_nil : forall s s', s [= nil =] s'.
Lemma eq_state_cons : forall p P s s',
s p [=] s' p -> s [= P =] s' -> s [= p::P =] s'.
Lemma eq_state_hd : forall p P s s', s [= p::P =] s' -> s p [=] s' p.
Lemma eq_state_tl : forall p P s s', s [= p::P =] s' -> s [= P =] s'.
Lemma eq_state_cons_iff : forall p P s s',
s [= p::P =] s' <-> (s p [=] s' p) /\ s [= P =] s'.
Lemma eq_state_app : forall P P' s s',
s [= P =] s' -> s [= P' =] s' -> s [= P ++ P' =] s'.
Lemma eq_state_split : forall P P' s s',
s [= P ++ P' =] s' -> s [= P =] s' /\ s [= P =] s'.
Lemma eq_state_refl : forall P, reflexive _ (eq_state P).
Lemma eq_state_sym : forall P, symmetric _ (eq_state P).
Lemma eq_state_trans : forall P, transitive _ (eq_state P).
Add Parametric Relation P : State (eq_state P)
reflexivity proved by (eq_state_refl P)
symmetry proved by (eq_state_sym P)
transitivity proved by (eq_state_trans P)
as eq_state_rel.
Definition update (s:State) (p:Pid) (x:Var) (v:Value) : State :=
fun (q:Pid) => if (p =? q) then (Lupdate (s p) x v) else (s q).
Notation "s [ p , x => v ]" := (update s p x v) (at level 10).
Lemma update_read : forall s p x v, s[p,x => v] p x = v.
Lemma update_read' : forall s p q x y v, p <> q -> s[p,x => v] q y = s q y.
Lemma update_read'' : forall s p q x y v, x <> y -> s[p,x => v] q y = s q y.
Lemma update_update : forall s p x v w P, s[p,x => w][p,x => v] [= P =] s[p,x => v].
Lemma update_not_in : forall s p x v P, ~In p P -> s [= P =] s[p,x => v].
Definition eq_state_ext (s s': State) : Prop :=
forall p, s p [=] s' p.
Notation "s [==] s'" := (eq_state_ext s s') (at level 50).
Lemma eq_state_ext_refl : forall s, s [==] s.
Lemma eq_state_ext_sym : forall s s', s [==] s' -> s' [==] s.
Lemma eq_state_ext_trans : forall s s' s'', s [==] s' -> s' [==] s'' -> s [==] s''.
Lemma eq_state_ext_congr : forall s s' p x v,
s [==] s' -> s[p,x => v] [==] s'[p,x => v].
Lemma update_update_ext : forall s p x v w, s[p,x => w][p,x => v] [==] s[p,x => v].
Lemma update_independent : forall s p q x y v w, p <> q ->
s[q,y => w][p,x => v] [==] s[p,x => v][q,y => w].
Lemma update_independent' : forall s p x y v w, x <> y ->
s[p,y => w][p,x => v] [==] s[p,x => v][p,y => w].
Lemma update_idempotent : forall s p x, s[p,x=>s p x] [==] s.
End GState.
Notation "s '[[' p ',' x '=>' v ']]'" := (update _ _ _ s p x v) (at level 100).
Notation "s [==] s'" := (eq_state_ext _ _ _ s s') (at level 50).
Useful tactics while we can't rewrite.
Ltac ESEr := apply eq_state_ext_refl.
Ltac ESEs := apply eq_state_ext_sym; auto.
Ltac ESEt x := apply eq_state_ext_trans with x; auto.
Ltac eESEt := eapply eq_state_ext_trans; eauto.
Ltac ESEc := apply eq_state_ext_congr.
Evaluation
Evaluation is parameterized on the types of expressions and values. Decidability of expressions makes choreography equality decidable.Record Eval {Expression Vars Input Output : DecType} :=
{ eval : Expression -> (Vars -> Input) -> Output;
eval_wd : forall f f', (forall x, f x = f' x) ->
forall e, eval e f = eval e f'}.
Both CC and SP use two types of expressions - one
for values, one for Boolean expressions. Their shared properties
that depend on the state are proven here.
Section EvalState.
Variable Pid Expression Vars Input Output : DecType.
Variable Ev:@Eval Expression Vars Input Output.
Expression evaluation on the state of a process
Consistency with state equivalence.
Lemma eval_eq : forall e s s' p, s [==] s' ->
eval_on_state e s p = eval_on_state e s' p.
Lemma eval_neq : forall e s p q x v, p <> q ->
eval_on_state e s p = eval_on_state e (s[[q,x => v]]) p.
End EvalState.
The labels in the LTS as in the reference articles.
Inductive TransitionLabel : Type :=
| TL_Com (p:Pid) (v:Value) (q:Pid) : TransitionLabel
| TL_Sel (p:Pid) (q:Pid) (l:Label) : TransitionLabel
| TL_Tau (p:Pid) : TransitionLabel
.
Lemma TransitionLabel_eq_dec : forall (x y:TransitionLabel), {x=y} + {x<>y}.
Definition TL_pn (t:TransitionLabel) : list Pid :=
match t with
| TL_Com p v q => p::q::nil
| TL_Sel p q l => p::q::nil
| TL_Tau p => p::nil
end.
The inductive definition of the semantics in Coq uses an LTS with
more expressive labels.
Inductive RichLabel : Type :=
| RL_Com (p:Pid) (v:Value) (q:Pid) (x:Var) : RichLabel
| RL_Sel (p:Pid) (q:Pid) (l:Label) : RichLabel
| RL_Cond (p:Pid) : RichLabel
| RL_Call (X:RecVar) (p:Pid) : RichLabel
.
Lemma RichLabel_eq_dec : forall (x y:RichLabel), {x=y} + {x<>y}.
Definition RL_pn (t:RichLabel) : list Pid :=
match t with
| RL_Com p v q _ => p::q::nil
| RL_Sel p q l => p::q::nil
| RL_Cond p => p::nil
| RL_Call _ p => p::nil
end.
Definition forget (t:RichLabel) : TransitionLabel :=
match t with
| RL_Com p v q _ => TL_Com p v q
| RL_Sel p q l => TL_Sel p q l
| RL_Cond p => TL_Tau p
| RL_Call _ p => TL_Tau p
end.
Useful for rewriting in proofs.
Lemma forget_Com : forall x p v q, forget (RL_Com p v q x) = TL_Com p v q.
Lemma forget_Sel : forall p q l, forget (RL_Sel p q l) = TL_Sel p q l.
Lemma forget_Cond : forall p, forget (RL_Cond p) = TL_Tau p.
Lemma forget_Call : forall X p, forget (RL_Call X p) = TL_Tau p.
Disjointness between a process/list of processes and a label - used
for a lot of properties of the semantics, these lemmas are useful for
reasoning about it.
Definition disjoint_p_rl (p:Pid) (t:RichLabel) : Prop :=
match t with
| RL_Com r _ s _ => p <> r /\ p <> s
| RL_Sel r s _ => p <> r /\ p <> s
| RL_Cond r => p <> r
| RL_Call _ r => p <> r
end.
Lemma disjoint_p_rl_eq : forall p t t', forget t = forget t' ->
disjoint_p_rl p t -> disjoint_p_rl p t'.
Fixpoint disjoint_ps_rl (ps:list Pid) (t:RichLabel) : Prop :=
match ps with
| nil => True
| p::ps' => disjoint_p_rl p t /\ disjoint_ps_rl ps' t
end.
Lemma disjoint_ps_rl_In p ps t :
In p ps -> disjoint_ps_rl ps t -> disjoint_p_rl p t.
Lemma disjoint_ps_Sel : forall p q l ps,
disjoint_ps_rl ps (RL_Sel p q l) -> disjoint (p::q::nil) ps.
Lemma disjoint_ps_Com : forall p v q x ps,
disjoint_ps_rl ps (RL_Com p v q x) -> disjoint (p::q::nil) ps.
Lemma disjoint_ps_Cond : forall p ps, disjoint_ps_rl ps (RL_Cond p) -> ~In p ps.
Lemma disjoint_ps_Call : forall p ps X, disjoint_ps_rl ps (RL_Call X p) -> ~In p ps.
Lemma disjoint_ps_char : forall ps tl,
(forall p, In p ps -> disjoint_p_rl p tl) -> disjoint_ps_rl ps tl.
Lemma disjoint_ps_remove : forall p ps tl, disjoint_ps_rl ps tl ->
disjoint_ps_rl (set_remove' (@eq_dec _) p ps) tl.
End Transitions.