mirror of
https://github.com/achlipala/frap.git
synced 2024-11-10 00:07:51 +00:00
692 lines
16 KiB
Coq
692 lines
16 KiB
Coq
(** Formal Reasoning About Programs <http://adam.chlipala.net/frap/>
|
|
* Chapter 13: Operational Semantics for Shared-Memory Concurrency
|
|
* Author: Adam Chlipala
|
|
* License: https://creativecommons.org/licenses/by-nc-nd/4.0/ *)
|
|
|
|
Require Import Frap.
|
|
|
|
Set Implicit Arguments.
|
|
Set Asymmetric Patterns.
|
|
|
|
(** * Shared notations and definitions; main material starts afterward. *)
|
|
|
|
Notation "m $! k" := (match m $? k with Some n => n | None => O end) (at level 30).
|
|
Definition heap := fmap nat nat.
|
|
Definition assertion := heap -> Prop.
|
|
|
|
Hint Extern 1 (_ <= _) => linear_arithmetic.
|
|
Hint Extern 1 (@eq nat _ _) => linear_arithmetic.
|
|
|
|
Ltac simp := repeat (simplify; subst; propositional;
|
|
try match goal with
|
|
| [ H : ex _ |- _ ] => invert H
|
|
end); try linear_arithmetic.
|
|
|
|
|
|
(** * An object language with shared-memory concurrency *)
|
|
|
|
(* Let's simplify the encoding by only working with commands that generate
|
|
* [nat]. *)
|
|
Inductive loop_outcome :=
|
|
| Done (a : nat)
|
|
| Again (a : nat).
|
|
|
|
Inductive cmd :=
|
|
| Return (r : nat)
|
|
| Bind (c1 : cmd) (c2 : nat -> cmd)
|
|
| Read (a : nat)
|
|
| Write (a v : nat)
|
|
| Fail
|
|
|
|
(* Now here's the new part: parallel composition of commands. *)
|
|
| Par (c1 c2 : cmd)
|
|
|
|
(* Let's also add locking commands, where locks are named by [nat]s. *)
|
|
| Lock (a : nat)
|
|
| Unlock (a : nat).
|
|
|
|
Notation "x <- c1 ; c2" := (Bind c1 (fun x => c2)) (right associativity, at level 80).
|
|
Infix "||" := Par.
|
|
|
|
Definition locks := set nat.
|
|
|
|
Inductive step : heap * locks * cmd -> heap * locks * cmd -> Prop :=
|
|
| StepBindRecur : forall c1 c1' c2 h h' l l',
|
|
step (h, l, c1) (h', l', c1')
|
|
-> step (h, l, Bind c1 c2) (h', l', Bind c1' c2)
|
|
| StepBindProceed : forall v c2 h l,
|
|
step (h, l, Bind (Return v) c2) (h, l, c2 v)
|
|
|
|
| StepRead : forall h l a,
|
|
step (h, l, Read a) (h, l, Return (h $! a))
|
|
| StepWrite : forall h l a v,
|
|
step (h, l, Write a v) (h $+ (a, v), l, Return 0)
|
|
|
|
| StepParRecur1 : forall h l c1 c2 h' l' c1',
|
|
step (h, l, c1) (h', l', c1')
|
|
-> step (h, l, Par c1 c2) (h', l', Par c1' c2)
|
|
| StepParRecur2 : forall h l c1 c2 h' l' c2',
|
|
step (h, l, c2) (h', l', c2')
|
|
-> step (h, l, Par c1 c2) (h', l', Par c1 c2')
|
|
| StepParProceed : forall h l r c,
|
|
step (h, l, Par (Return r) c) (h, l, c)
|
|
|
|
| StepLock : forall h l a,
|
|
~a \in l
|
|
-> step (h, l, Lock a) (h, l \cup {a}, Return 0)
|
|
| StepUnlock : forall h l a,
|
|
a \in l
|
|
-> step (h, l, Unlock a) (h, l \setminus {a}, Return 0).
|
|
|
|
Definition trsys_of (h : heap) (l : locks) (c : cmd) := {|
|
|
Initial := {(h, l, c)};
|
|
Step := step
|
|
|}.
|
|
|
|
|
|
Example two_increments_thread :=
|
|
_ <- Lock 0;
|
|
n <- Read 0;
|
|
_ <- Write 0 (n + 1);
|
|
Unlock 0.
|
|
|
|
Example two_increments :=
|
|
_ <- (two_increments_thread || two_increments_thread);
|
|
n <- Read 0;
|
|
if n ==n 2 then
|
|
Return 0
|
|
else
|
|
Fail.
|
|
|
|
Fixpoint notAboutToFail (c : cmd) : bool :=
|
|
match c with
|
|
| Fail => false
|
|
| Bind c1 _ => notAboutToFail c1
|
|
| Par c1 c2 => notAboutToFail c1 && notAboutToFail c2
|
|
| _ => true
|
|
end.
|
|
|
|
Theorem two_increments_ok :
|
|
invariantFor (trsys_of $0 {} two_increments)
|
|
(fun p => let '(_, _, c) := p in
|
|
notAboutToFail c = true).
|
|
Proof.
|
|
Admitted.
|
|
(* unfold two_increments, two_increments_thread.
|
|
simplify.
|
|
eapply invariant_weaken.
|
|
apply multiStepClosure_ok; simplify.
|
|
|
|
model_check_step.
|
|
model_check_step.
|
|
model_check_step.
|
|
model_check_step.
|
|
model_check_step.
|
|
model_check_step.
|
|
model_check_step.
|
|
model_check_step.
|
|
model_check_step.
|
|
model_check_step.
|
|
model_check_step.
|
|
model_check_step.
|
|
model_check_step.
|
|
model_check_step.
|
|
model_check_step.
|
|
model_check_step.
|
|
model_check_step.
|
|
model_check_step.
|
|
model_check_done.
|
|
|
|
simplify.
|
|
propositional; subst; equality.
|
|
Qed.*)
|
|
|
|
|
|
(** * Optimization #1: always run all purely local actions first. *)
|
|
|
|
Fixpoint runLocal (c : cmd) : cmd :=
|
|
match c with
|
|
| Return _ => c
|
|
| Bind c1 c2 =>
|
|
match runLocal c1 with
|
|
| Return v => runLocal (c2 v)
|
|
| c1' => Bind c1' c2
|
|
end
|
|
| Read _ => c
|
|
| Write _ _ => c
|
|
| Fail => c
|
|
| Par c1 c2 =>
|
|
match runLocal c1 with
|
|
| Return _ => runLocal c2
|
|
| c1' => Par c1' (runLocal c2)
|
|
end
|
|
| Lock _ => c
|
|
| Unlock _ => c
|
|
end.
|
|
|
|
Inductive stepL : heap * locks * cmd -> heap * locks * cmd -> Prop :=
|
|
| StepL : forall h l c h' l' c',
|
|
step (h, l, c) (h', l', c')
|
|
-> stepL (h, l, c) (h', l', runLocal c').
|
|
|
|
Definition trsys_ofL (h : heap) (l : locks) (c : cmd) := {|
|
|
Initial := {(h, l, runLocal c)};
|
|
Step := stepL
|
|
|}.
|
|
|
|
Hint Constructors step stepL.
|
|
|
|
Lemma run_Return : forall h l r h' l' c,
|
|
step^* (h, l, Return r) (h', l', c)
|
|
-> h' = h /\ l' = l /\ c = Return r.
|
|
Proof.
|
|
induct 1; eauto.
|
|
invert H.
|
|
Qed.
|
|
|
|
Lemma run_Bind : forall h l c1 c2 h' l' c',
|
|
step^* (h, l, Bind c1 c2) (h', l', c')
|
|
-> (exists c1', step^* (h, l, c1) (h', l', c1')
|
|
/\ c' = Bind c1' c2)
|
|
\/ (exists h'' l'' r, step^* (h, l, c1) (h'', l'', Return r)
|
|
/\ step^* (h'', l'', c2 r) (h', l', c')).
|
|
Proof.
|
|
induct 1; eauto.
|
|
invert H; eauto 10.
|
|
|
|
Ltac inst H :=
|
|
repeat match type of H with
|
|
| _ = _ -> _ => specialize (H eq_refl)
|
|
| forall x : ?T, _ =>
|
|
let y := fresh in evar (y : T); let y' := eval unfold y in y in
|
|
specialize (H y'); clear y
|
|
end.
|
|
|
|
inst IHtrc.
|
|
first_order; eauto 10.
|
|
Qed.
|
|
|
|
Lemma StepBindRecur_star : forall c1 c1' c2 h h' l l',
|
|
step^* (h, l, c1) (h', l', c1')
|
|
-> step^* (h, l, Bind c1 c2) (h', l', Bind c1' c2).
|
|
Proof.
|
|
induct 1; eauto.
|
|
cases y.
|
|
cases p.
|
|
eauto.
|
|
Qed.
|
|
|
|
Lemma StepParRecur1_star : forall h l c1 c2 h' l' c1',
|
|
step^* (h, l, c1) (h', l', c1')
|
|
-> step^* (h, l, Par c1 c2) (h', l', Par c1' c2).
|
|
Proof.
|
|
induct 1; eauto.
|
|
cases y.
|
|
cases p.
|
|
eauto.
|
|
Qed.
|
|
|
|
Lemma StepParRecur2_star : forall h l c1 c2 h' l' c2',
|
|
step^* (h, l, c2) (h', l', c2')
|
|
-> step^* (h, l, Par c1 c2) (h', l', Par c1 c2').
|
|
Proof.
|
|
induct 1; eauto.
|
|
cases y.
|
|
cases p.
|
|
eauto.
|
|
Qed.
|
|
|
|
Hint Resolve StepBindRecur_star StepParRecur1_star StepParRecur2_star.
|
|
|
|
Lemma runLocal_idem : forall c, runLocal (runLocal c) = runLocal c.
|
|
Proof.
|
|
induct c; simplify; eauto.
|
|
cases (runLocal c); simplify; eauto.
|
|
rewrite IHc; auto.
|
|
rewrite IHc; auto.
|
|
cases (runLocal c1); simplify; eauto.
|
|
rewrite IHc1; equality.
|
|
rewrite IHc2; equality.
|
|
rewrite IHc2; equality.
|
|
rewrite IHc2; equality.
|
|
rewrite IHc1; equality.
|
|
rewrite IHc2; equality.
|
|
rewrite IHc2; equality.
|
|
Qed.
|
|
|
|
Lemma runLocal_left : forall c1 c2,
|
|
(forall r, runLocal c1 <> Return r)
|
|
-> runLocal (c1 || c2) = runLocal c1 || runLocal c2.
|
|
Proof.
|
|
simplify.
|
|
cases (runLocal c1); eauto.
|
|
unfold not in *.
|
|
exfalso; eauto.
|
|
Qed.
|
|
|
|
Lemma step_runLocal : forall h l c h' l' c',
|
|
step (h, l, c) (h', l', c')
|
|
-> (runLocal c = runLocal c' /\ h = h' /\ l = l')
|
|
\/ exists c'', step (h, l, runLocal c) (h', l', c'')
|
|
/\ runLocal c'' = runLocal c'.
|
|
Proof.
|
|
induct 1; simplify; first_order; eauto.
|
|
|
|
rewrite H0; equality.
|
|
|
|
cases (runLocal c1).
|
|
invert H0.
|
|
rewrite <- H1; eauto.
|
|
rewrite <- H1; eauto.
|
|
rewrite <- H1; eauto.
|
|
rewrite <- H1; eauto.
|
|
rewrite <- H1; eauto.
|
|
rewrite <- H1; eauto.
|
|
rewrite <- H1; eauto.
|
|
|
|
rewrite H0; equality.
|
|
|
|
cases (runLocal c1).
|
|
invert H0.
|
|
rewrite <- H1.
|
|
right.
|
|
eexists.
|
|
propositional.
|
|
eauto.
|
|
simplify.
|
|
rewrite runLocal_idem.
|
|
equality.
|
|
rewrite <- H1.
|
|
right.
|
|
eexists.
|
|
propositional.
|
|
eauto.
|
|
simplify.
|
|
rewrite runLocal_idem.
|
|
equality.
|
|
rewrite <- H1.
|
|
right.
|
|
eexists.
|
|
propositional.
|
|
eauto.
|
|
simplify.
|
|
rewrite runLocal_idem.
|
|
equality.
|
|
rewrite <- H1.
|
|
right.
|
|
eexists.
|
|
propositional.
|
|
eauto.
|
|
simplify.
|
|
rewrite runLocal_idem.
|
|
equality.
|
|
rewrite <- H1.
|
|
right.
|
|
eexists.
|
|
propositional.
|
|
eauto.
|
|
simplify.
|
|
rewrite runLocal_idem.
|
|
equality.
|
|
rewrite <- H1.
|
|
right.
|
|
eexists.
|
|
propositional.
|
|
eauto.
|
|
simplify.
|
|
rewrite runLocal_idem.
|
|
equality.
|
|
rewrite <- H1.
|
|
right.
|
|
eexists.
|
|
propositional.
|
|
eauto.
|
|
simplify.
|
|
rewrite runLocal_idem.
|
|
equality.
|
|
|
|
rewrite H0; equality.
|
|
|
|
right.
|
|
cases (runLocal c1); eauto.
|
|
eexists; propositional.
|
|
eauto.
|
|
rewrite runLocal_left.
|
|
rewrite <- Heq.
|
|
rewrite runLocal_idem.
|
|
equality.
|
|
rewrite <- Heq.
|
|
rewrite runLocal_idem.
|
|
rewrite Heq.
|
|
equality.
|
|
|
|
eexists; propositional.
|
|
eauto.
|
|
rewrite runLocal_left.
|
|
rewrite <- Heq.
|
|
rewrite runLocal_idem.
|
|
equality.
|
|
rewrite <- Heq.
|
|
rewrite runLocal_idem.
|
|
rewrite Heq.
|
|
equality.
|
|
|
|
eexists; propositional.
|
|
eauto.
|
|
rewrite runLocal_left.
|
|
rewrite <- Heq.
|
|
rewrite runLocal_idem.
|
|
equality.
|
|
rewrite <- Heq.
|
|
rewrite runLocal_idem.
|
|
rewrite Heq.
|
|
equality.
|
|
|
|
eexists; propositional.
|
|
eauto.
|
|
rewrite runLocal_left.
|
|
rewrite <- Heq.
|
|
rewrite runLocal_idem.
|
|
equality.
|
|
rewrite <- Heq.
|
|
rewrite runLocal_idem.
|
|
rewrite Heq.
|
|
equality.
|
|
|
|
eexists; propositional.
|
|
eauto.
|
|
rewrite runLocal_left.
|
|
rewrite <- Heq.
|
|
rewrite runLocal_idem.
|
|
equality.
|
|
rewrite <- Heq.
|
|
rewrite runLocal_idem.
|
|
rewrite Heq.
|
|
equality.
|
|
|
|
eexists; propositional.
|
|
eauto.
|
|
rewrite runLocal_left.
|
|
rewrite <- Heq.
|
|
rewrite runLocal_idem.
|
|
equality.
|
|
rewrite <- Heq.
|
|
rewrite runLocal_idem.
|
|
rewrite Heq.
|
|
equality.
|
|
|
|
eexists; propositional.
|
|
eauto.
|
|
rewrite runLocal_left.
|
|
rewrite <- Heq.
|
|
rewrite runLocal_idem.
|
|
equality.
|
|
rewrite <- Heq.
|
|
rewrite runLocal_idem.
|
|
rewrite Heq.
|
|
equality.
|
|
Qed.
|
|
|
|
Lemma step_stepL' : forall h l c h' l' c',
|
|
step^* (h, l, c) (h', l', c')
|
|
-> stepL^* (h, l, runLocal c) (h', l', runLocal c').
|
|
Proof.
|
|
induct 1; simplify; eauto.
|
|
cases y.
|
|
cases p.
|
|
inst IHtrc.
|
|
apply step_runLocal in H; first_order; subst.
|
|
rewrite H; eauto.
|
|
econstructor.
|
|
econstructor.
|
|
eauto.
|
|
equality.
|
|
Qed.
|
|
|
|
Theorem notAboutToFail_runLocal : forall c,
|
|
notAboutToFail (runLocal c) = true
|
|
-> notAboutToFail c = true.
|
|
Proof.
|
|
induct c; simplify; auto.
|
|
cases (runLocal c); simplify; auto.
|
|
cases (runLocal c1); simplify; auto; propositional;
|
|
repeat match goal with
|
|
| [ H : _ |- _ ] => apply andb_true_iff in H; propositional
|
|
| [ H : _ = _ |- _ ] => rewrite H
|
|
end; try equality.
|
|
Qed.
|
|
|
|
Theorem step_stepL : forall h l c ,
|
|
invariantFor (trsys_ofL h l c) (fun p => let '(_, _, c) := p in
|
|
notAboutToFail c = true)
|
|
-> invariantFor (trsys_of h l c) (fun p =>
|
|
let '(_, _, c) := p in
|
|
notAboutToFail c = true).
|
|
Proof.
|
|
unfold invariantFor; simplify.
|
|
propositional; subst.
|
|
|
|
cases s'.
|
|
cases p.
|
|
apply step_stepL' in H1.
|
|
apply H in H1; eauto using notAboutToFail_runLocal.
|
|
Qed.
|
|
|
|
Theorem two_increments_ok_again :
|
|
invariantFor (trsys_of $0 {} two_increments)
|
|
(fun p => let '(_, _, c) := p in
|
|
notAboutToFail c = true).
|
|
Proof.
|
|
Admitted.
|
|
(* apply step_stepL.
|
|
unfold two_increments, two_increments_thread.
|
|
simplify.
|
|
eapply invariant_weaken.
|
|
apply multiStepClosure_ok; simplify.
|
|
|
|
model_check_step.
|
|
model_check_step.
|
|
model_check_step.
|
|
model_check_step.
|
|
model_check_step.
|
|
model_check_step.
|
|
model_check_step.
|
|
model_check_step.
|
|
model_check_step.
|
|
model_check_done.
|
|
|
|
simplify.
|
|
propositional; subst; equality.
|
|
Qed.*)
|
|
|
|
|
|
(** * Optimization #2: partial-order reduction *)
|
|
|
|
Example independent_threads :=
|
|
(a <- Read 0;
|
|
_ <- Write 1 (a + 1);
|
|
a <- Read 1;
|
|
if a ==n 1 then
|
|
Return 0
|
|
else
|
|
Fail)
|
|
|| (b <- Read 2;
|
|
Write 2 (b + 1)).
|
|
|
|
Theorem independent_threads_ok :
|
|
invariantFor (trsys_of $0 {} independent_threads)
|
|
(fun p => let '(_, _, c) := p in
|
|
notAboutToFail c = true).
|
|
Proof.
|
|
Admitted.
|
|
(* apply step_stepL.
|
|
unfold independent_threads.
|
|
simplify.
|
|
eapply invariant_weaken.
|
|
apply multiStepClosure_ok; simplify.
|
|
|
|
model_check_step.
|
|
model_check_step.
|
|
model_check_step.
|
|
model_check_step.
|
|
model_check_step.
|
|
model_check_done.
|
|
|
|
simplify.
|
|
propositional; subst; equality.
|
|
Qed.*)
|
|
|
|
Inductive firstThread : cmd -> cmd -> cmd -> Prop :=
|
|
| FtPar : forall c1 c2 c11 c12,
|
|
firstThread c1 c11 c12
|
|
-> firstThread (Par c1 c2) c11 (Par c12 c2)
|
|
| FtDone : forall c,
|
|
match c with
|
|
| Par _ _ => False
|
|
| _ => True
|
|
end
|
|
-> firstThread c c (Return 0).
|
|
|
|
Inductive nextAction : cmd -> cmd -> Prop :=
|
|
| NaRead : forall a,
|
|
nextAction (Read a) (Read a)
|
|
| NaWrite : forall a v,
|
|
nextAction (Write a v) (Write a v)
|
|
| NaLock : forall l,
|
|
nextAction (Lock l) (Lock l)
|
|
| NaUnlock : forall l,
|
|
nextAction (Unlock l) (Unlock l)
|
|
| NaBind : forall c1 c2 c,
|
|
nextAction c1 c
|
|
-> nextAction (Bind c1 c2) c.
|
|
|
|
Inductive commutes : cmd -> cmd -> Prop :=
|
|
| ComReadRead : forall a1 a2,
|
|
commutes (Read a1) (Read a2)
|
|
| ComReadWrite : forall a1 a2 v,
|
|
a1 <> a2
|
|
-> commutes (Read a1) (Write a2 v)
|
|
| ComReadLock : forall a1 a2,
|
|
commutes (Read a1) (Lock a2)
|
|
| ComReadUnlock : forall a1 a2,
|
|
commutes (Read a1) (Unlock a2)
|
|
|
|
| ComWriteRead : forall a1 v a2,
|
|
a1 <> a2
|
|
-> commutes (Write a1 v) (Read a2)
|
|
| ComWriteWrite : forall a1 a2 v1 v2,
|
|
a1 <> a2
|
|
-> commutes (Write a1 v1) (Write a2 v2)
|
|
| ComWriteLock : forall a1 v a2,
|
|
commutes (Write a1 v) (Lock a2)
|
|
| ComWriteUnlock : forall a1 v a2,
|
|
commutes (Write a1 v) (Unlock a2)
|
|
|
|
| ComLockRead : forall a1 a2,
|
|
commutes (Lock a1) (Read a2)
|
|
| ComLockWrite : forall a1 a2 v,
|
|
commutes (Lock a1) (Write a2 v)
|
|
| ComLockLock : forall a1 a2,
|
|
a1 <> a2
|
|
-> commutes (Lock a1) (Lock a2)
|
|
| ComLockUnlock : forall a1 a2,
|
|
a1 <> a2
|
|
-> commutes (Lock a1) (Unlock a2)
|
|
|
|
| ComUnlockRead : forall a1 a2,
|
|
commutes (Unlock a1) (Read a2)
|
|
| ComUnlockWrite : forall a1 a2 v,
|
|
commutes (Unlock a1) (Write a2 v)
|
|
| ComUnlockLock : forall a1 a2,
|
|
a1 <> a2
|
|
-> commutes (Unlock a1) (Lock a2)
|
|
| ComUnlockUnlock : forall a1 a2,
|
|
a1 <> a2
|
|
-> commutes (Unlock a1) (Unlock a2)
|
|
|
|
| CommFail : forall c,
|
|
commutes c Fail
|
|
| CommReturn : forall c r,
|
|
commutes c (Return r)
|
|
| CommBind : forall c c1 c2,
|
|
commutes c c1
|
|
-> (forall r, commutes c (c2 r))
|
|
-> commutes c (Bind c1 c2)
|
|
| CommPar : forall c c1 c2,
|
|
commutes c c1
|
|
-> commutes c c2
|
|
-> commutes c (Par c1 c2).
|
|
|
|
Lemma commutes_sound1 : forall h l c2 h' l' c2',
|
|
step (h, l, c2) (h', l', c2')
|
|
-> forall c1 h'' l'' c1', step (h', l', c1) (h'', l'', c1')
|
|
-> commutes c1 c2
|
|
-> exists h1 l1, step (h, l, c1) (h1, l1, c1')
|
|
/\ step (h1, l1, c2) (h'', l'', c2').
|
|
Proof.
|
|
induct 1; simplify; eauto.
|
|
|
|
invert H1.
|
|
apply IHstep in H0; first_order.
|
|
eauto.
|
|
|
|
invert H0; invert H; eauto.
|
|
do 2 eexists; propositional.
|
|
eauto.
|
|
replace (h' $! a) with (h' $+ (a1, v) $! a) by (simplify; equality).
|
|
eauto.
|
|
|
|
invert H0; invert H; eauto.
|
|
simplify.
|
|
eauto.
|
|
do 2 eexists; propositional.
|
|
eauto.
|
|
replace (h $+ (a, v) $+ (a1, v1)) with (h $+ (a1, v1) $+ (a, v)) by maps_equal.
|
|
eauto.
|
|
|
|
invert H1.
|
|
eapply IHstep in H5; eauto.
|
|
first_order; eauto.
|
|
|
|
invert H1.
|
|
eapply IHstep in H6; eauto.
|
|
first_order; eauto.
|
|
|
|
invert H1; invert H0; eauto.
|
|
do 2 eexists; propositional.
|
|
constructor.
|
|
sets.
|
|
replace ((l \cup {a}) \cup {a1}) with ((l \cup {a1}) \cup {a}) by sets.
|
|
constructor.
|
|
sets.
|
|
do 2 eexists; propositional.
|
|
constructor.
|
|
sets; propositional.
|
|
replace (l \cup {a} \setminus {a1}) with ((l \setminus {a1}) \cup {a}) by sets.
|
|
constructor.
|
|
sets.
|
|
|
|
invert H1; invert H0; eauto.
|
|
do 2 eexists; propositional.
|
|
constructor.
|
|
sets.
|
|
replace ((l \setminus {a}) \cup {a1}) with ((l \cup {a1}) \setminus {a}) by sets.
|
|
constructor.
|
|
sets.
|
|
do 2 eexists; propositional.
|
|
constructor.
|
|
sets; propositional.
|
|
replace ((l \setminus {a}) \setminus {a1}) with ((l \setminus {a1}) \setminus {a}) by sets.
|
|
constructor.
|
|
sets.
|
|
Qed.
|
|
|
|
Hint Constructors commutes.
|
|
|
|
Lemma commutes_sound2 : forall h l c2 h' l' c2',
|
|
step (h, l, c2) (h', l', c2')
|
|
-> forall c1, commutes c1 c2
|
|
-> commutes c1 c2'.
|
|
Proof.
|
|
induct 1; invert 1; simplify; eauto.
|
|
Qed.
|