mirror of
https://github.com/achlipala/frap.git
synced 2024-11-12 17:17:50 +00:00
SharedMemory: change StepParProceed
This commit is contained in:
parent
e9849b9e9c
commit
ec5a4a198e
1 changed files with 18 additions and 77 deletions
|
@ -68,8 +68,8 @@ Inductive step : heap * locks * cmd -> heap * locks * cmd -> Prop :=
|
||||||
| StepParRecur2 : forall h l c1 c2 h' l' c2',
|
| StepParRecur2 : forall h l c1 c2 h' l' c2',
|
||||||
step (h, l, c2) (h', l', c2')
|
step (h, l, c2) (h', l', c2')
|
||||||
-> step (h, l, Par c1 c2) (h', l', Par c1 c2')
|
-> step (h, l, Par c1 c2) (h', l', Par c1 c2')
|
||||||
| StepParProceed : forall h l,
|
| StepParProceed : forall h l r c,
|
||||||
step (h, l, Par (Return 0) (Return 0)) (h, l, Return 0)
|
step (h, l, Par (Return r) c) (h, l, c)
|
||||||
|
|
||||||
| StepLock : forall h l a,
|
| StepLock : forall h l a,
|
||||||
~a \in l
|
~a \in l
|
||||||
|
@ -155,9 +155,9 @@ Fixpoint runLocal (c : cmd) : cmd :=
|
||||||
| Write _ _ => c
|
| Write _ _ => c
|
||||||
| Fail => c
|
| Fail => c
|
||||||
| Par c1 c2 =>
|
| Par c1 c2 =>
|
||||||
match runLocal c1, runLocal c2 with
|
match runLocal c1 with
|
||||||
| Return 0, Return 0 => Return 0
|
| Return _ => runLocal c2
|
||||||
| c1', c2' => Par c1' c2'
|
| c1' => Par c1' (runLocal c2)
|
||||||
end
|
end
|
||||||
| Lock _ => c
|
| Lock _ => c
|
||||||
| Unlock _ => c
|
| Unlock _ => c
|
||||||
|
@ -244,29 +244,13 @@ Proof.
|
||||||
rewrite IHc; auto.
|
rewrite IHc; auto.
|
||||||
rewrite IHc; auto.
|
rewrite IHc; auto.
|
||||||
cases (runLocal c1); simplify; eauto.
|
cases (runLocal c1); simplify; eauto.
|
||||||
cases (runLocal c2); simplify; eauto.
|
rewrite IHc1; equality.
|
||||||
cases r; eauto.
|
rewrite IHc2; equality.
|
||||||
cases r0; eauto.
|
rewrite IHc2; equality.
|
||||||
cases r; simplify; eauto.
|
rewrite IHc2; equality.
|
||||||
rewrite IHc2; auto.
|
rewrite IHc1; equality.
|
||||||
rewrite IHc2; auto.
|
rewrite IHc2; equality.
|
||||||
cases r; simplify; eauto.
|
rewrite IHc2; equality.
|
||||||
cases r; simplify; eauto.
|
|
||||||
cases r; simplify; eauto.
|
|
||||||
cases r; simplify; eauto.
|
|
||||||
rewrite IHc2; eauto.
|
|
||||||
rewrite IHc2; eauto.
|
|
||||||
cases r; simplify; eauto.
|
|
||||||
cases r; simplify; eauto.
|
|
||||||
rewrite IHc2; eauto.
|
|
||||||
rewrite IHc1; eauto.
|
|
||||||
equality.
|
|
||||||
equality.
|
|
||||||
equality.
|
|
||||||
rewrite IHc1; eauto.
|
|
||||||
equality.
|
|
||||||
equality.
|
|
||||||
equality.
|
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
Lemma runLocal_left : forall c1 c2,
|
Lemma runLocal_left : forall c1 c2,
|
||||||
|
@ -363,45 +347,9 @@ Proof.
|
||||||
rewrite H0; equality.
|
rewrite H0; equality.
|
||||||
|
|
||||||
right.
|
right.
|
||||||
cases (runLocal c1).
|
cases (runLocal c1); eauto.
|
||||||
cases r.
|
|
||||||
cases (runLocal c2).
|
|
||||||
invert H0.
|
|
||||||
eexists; propositional.
|
eexists; propositional.
|
||||||
eauto.
|
eauto.
|
||||||
simplify.
|
|
||||||
rewrite H1; equality.
|
|
||||||
eexists; propositional.
|
|
||||||
eauto.
|
|
||||||
simplify.
|
|
||||||
rewrite H1; equality.
|
|
||||||
eexists; propositional.
|
|
||||||
eauto.
|
|
||||||
simplify.
|
|
||||||
rewrite H1; equality.
|
|
||||||
eexists; propositional.
|
|
||||||
eauto.
|
|
||||||
simplify.
|
|
||||||
rewrite H1; equality.
|
|
||||||
eexists; propositional.
|
|
||||||
eauto.
|
|
||||||
simplify.
|
|
||||||
rewrite H1; equality.
|
|
||||||
eexists; propositional.
|
|
||||||
eauto.
|
|
||||||
simplify.
|
|
||||||
rewrite H1; equality.
|
|
||||||
eexists; propositional.
|
|
||||||
eauto.
|
|
||||||
simplify.
|
|
||||||
rewrite H1; equality.
|
|
||||||
eexists; propositional.
|
|
||||||
eauto.
|
|
||||||
simplify.
|
|
||||||
rewrite H1; equality.
|
|
||||||
eexists; propositional.
|
|
||||||
eauto.
|
|
||||||
|
|
||||||
rewrite runLocal_left.
|
rewrite runLocal_left.
|
||||||
rewrite <- Heq.
|
rewrite <- Heq.
|
||||||
rewrite runLocal_idem.
|
rewrite runLocal_idem.
|
||||||
|
@ -500,18 +448,11 @@ Theorem notAboutToFail_runLocal : forall c,
|
||||||
Proof.
|
Proof.
|
||||||
induct c; simplify; auto.
|
induct c; simplify; auto.
|
||||||
cases (runLocal c); simplify; auto.
|
cases (runLocal c); simplify; auto.
|
||||||
|
cases (runLocal c1); simplify; auto; propositional;
|
||||||
Ltac bool := simplify; auto; propositional;
|
repeat match goal with
|
||||||
repeat match goal with
|
| [ H : _ |- _ ] => apply andb_true_iff in H; propositional
|
||||||
| [ H : _ |- _ ] => apply andb_true_iff in H; propositional
|
| [ H : _ = _ |- _ ] => rewrite H
|
||||||
| [ H : _ = _ |- _ ] => rewrite H
|
end; try equality.
|
||||||
end; try equality.
|
|
||||||
|
|
||||||
cases (runLocal c1); bool.
|
|
||||||
cases (runLocal c2); bool.
|
|
||||||
cases r; bool.
|
|
||||||
cases r; bool.
|
|
||||||
cases r; bool.
|
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
Theorem step_stepL : forall h l c ,
|
Theorem step_stepL : forall h l c ,
|
||||||
|
|
Loading…
Reference in a new issue