mirror of
https://github.com/achlipala/frap.git
synced 2025-03-01 20:52:13 +00:00
Revising OperationalSemantics
This commit is contained in:
parent
254e2aedc6
commit
d6e8cebdc9
2 changed files with 1018 additions and 1097 deletions
|
@ -18,38 +18,16 @@ Inductive arith : Set :=
|
||||||
| Minus (e1 e2 : arith)
|
| Minus (e1 e2 : arith)
|
||||||
| Times (e1 e2 : arith).
|
| Times (e1 e2 : arith).
|
||||||
|
|
||||||
Inductive cmd :=
|
|
||||||
| Skip
|
|
||||||
| Assign (x : var) (e : arith)
|
|
||||||
| Sequence (c1 c2 : cmd)
|
|
||||||
| If (e : arith) (then_ else_ : cmd)
|
|
||||||
| While (e : arith) (body : cmd).
|
|
||||||
(* Important differences: we added [If] and switched [Repeat] to general
|
|
||||||
* [While]. *)
|
|
||||||
|
|
||||||
(* Here are some notations for the language, which again we won't really
|
|
||||||
* explain. *)
|
|
||||||
Coercion Const : nat >-> arith.
|
Coercion Const : nat >-> arith.
|
||||||
Coercion Var : var >-> arith.
|
Coercion Var : var >-> arith.
|
||||||
Infix "+" := Plus : arith_scope.
|
Infix "+" := Plus : arith_scope.
|
||||||
Infix "-" := Minus : arith_scope.
|
Infix "-" := Minus : arith_scope.
|
||||||
Infix "*" := Times : arith_scope.
|
Infix "*" := Times : arith_scope.
|
||||||
Delimit Scope arith_scope with arith.
|
Delimit Scope arith_scope with arith.
|
||||||
Notation "x <- e" := (Assign x e%arith) (at level 75).
|
|
||||||
Infix ";;" := Sequence (at level 76). (* This one changed slightly, to avoid parsing clashes. *)
|
|
||||||
Notation "'when' e 'then' then_ 'else' else_ 'done'" := (If e%arith then_ else_) (at level 75, e at level 0).
|
|
||||||
Notation "'while' e 'loop' body 'done'" := (While e%arith body) (at level 75).
|
|
||||||
|
|
||||||
(* Here's an adaptation of our factorial example from Chapter 3. *)
|
Definition valuation := fmap var nat.
|
||||||
Example factorial :=
|
|
||||||
"output" <- 1;;
|
|
||||||
while "input" loop
|
|
||||||
"output" <- "output" * "input";;
|
|
||||||
"input" <- "input" - 1
|
|
||||||
done.
|
|
||||||
|
|
||||||
(* Recall our use of a recursive function to interpret expressions. *)
|
(* Recall our use of a recursive function to interpret expressions. *)
|
||||||
Definition valuation := fmap var nat.
|
|
||||||
Fixpoint interp (e : arith) (v : valuation) : nat :=
|
Fixpoint interp (e : arith) (v : valuation) : nat :=
|
||||||
match e with
|
match e with
|
||||||
| Const n => n
|
| Const n => n
|
||||||
|
@ -63,6 +41,31 @@ Fixpoint interp (e : arith) (v : valuation) : nat :=
|
||||||
| Times e1 e2 => interp e1 v * interp e2 v
|
| Times e1 e2 => interp e1 v * interp e2 v
|
||||||
end.
|
end.
|
||||||
|
|
||||||
|
Module Simple.
|
||||||
|
Inductive cmd :=
|
||||||
|
| Skip
|
||||||
|
| Assign (x : var) (e : arith)
|
||||||
|
| Sequence (c1 c2 : cmd)
|
||||||
|
| If (e : arith) (then_ else_ : cmd)
|
||||||
|
| While (e : arith) (body : cmd).
|
||||||
|
(* Important differences: we added [If] and switched [Repeat] to general
|
||||||
|
* [While]. *)
|
||||||
|
|
||||||
|
(* Here are some notations for the language, which again we won't really
|
||||||
|
* explain. *)
|
||||||
|
Notation "x <- e" := (Assign x e%arith) (at level 75).
|
||||||
|
Infix ";;" := Sequence (at level 76). (* This one changed slightly, to avoid parsing clashes. *)
|
||||||
|
Notation "'when' e 'then' then_ 'else' else_ 'done'" := (If e%arith then_ else_) (at level 75, e at level 0).
|
||||||
|
Notation "'while' e 'loop' body 'done'" := (While e%arith body) (at level 75).
|
||||||
|
|
||||||
|
(* Here's an adaptation of our factorial example from Chapter 3. *)
|
||||||
|
Example factorial :=
|
||||||
|
"output" <- 1;;
|
||||||
|
while "input" loop
|
||||||
|
"output" <- "output" * "input";;
|
||||||
|
"input" <- "input" - 1
|
||||||
|
done.
|
||||||
|
|
||||||
(* Our old trick of interpreters won't work for this new language, because of
|
(* Our old trick of interpreters won't work for this new language, because of
|
||||||
* the general "while" loops. No such interpreter could terminate for all
|
* the general "while" loops. No such interpreter could terminate for all
|
||||||
* programs. Instead, we will use inductive predicates to explain program
|
* programs. Instead, we will use inductive predicates to explain program
|
||||||
|
@ -460,7 +463,7 @@ Qed.
|
||||||
(* Bonus material: here's how to make these proofs much more automatic. We
|
(* Bonus material: here's how to make these proofs much more automatic. We
|
||||||
* won't explain the features being used here. *)
|
* won't explain the features being used here. *)
|
||||||
|
|
||||||
Hint Constructors trc step eval.
|
Hint Constructors trc step eval : core.
|
||||||
|
|
||||||
Lemma step_star_Seq_snazzy : forall v c1 c2 v' c1',
|
Lemma step_star_Seq_snazzy : forall v c1 c2 v' c1',
|
||||||
step^* (v, c1) (v', c1')
|
step^* (v, c1) (v', c1')
|
||||||
|
@ -470,7 +473,7 @@ Proof.
|
||||||
cases y; eauto.
|
cases y; eauto.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
Hint Resolve step_star_Seq_snazzy.
|
Hint Resolve step_star_Seq_snazzy : core.
|
||||||
|
|
||||||
Theorem big_small_snazzy : forall v c v', eval v c v'
|
Theorem big_small_snazzy : forall v c v', eval v c v'
|
||||||
-> step^* (v, c) (v', Skip).
|
-> step^* (v, c) (v', Skip).
|
||||||
|
@ -488,7 +491,7 @@ Proof.
|
||||||
end; eauto.
|
end; eauto.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
Hint Resolve small_big''_snazzy.
|
Hint Resolve small_big''_snazzy : core.
|
||||||
|
|
||||||
Lemma small_big'_snazzy : forall v c v' c', step^* (v, c) (v', c')
|
Lemma small_big'_snazzy : forall v c v' c', step^* (v, c) (v', c')
|
||||||
-> forall v'', eval v' c' v''
|
-> forall v'', eval v' c' v''
|
||||||
|
@ -498,7 +501,7 @@ Proof.
|
||||||
cases y; eauto.
|
cases y; eauto.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
Hint Resolve small_big'_snazzy.
|
Hint Resolve small_big'_snazzy : core.
|
||||||
|
|
||||||
Theorem small_big_snazzy : forall v c v', step^* (v, c) (v', Skip)
|
Theorem small_big_snazzy : forall v c v', step^* (v, c) (v', Skip)
|
||||||
-> eval v c v'.
|
-> eval v c v'.
|
||||||
|
@ -685,8 +688,8 @@ Proof.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
(* That manual proof was quite a pain. Here's a bonus automated proof. *)
|
(* That manual proof was quite a pain. Here's a bonus automated proof. *)
|
||||||
Hint Constructors isEven.
|
Hint Constructors isEven : core.
|
||||||
Hint Resolve isEven_minus2 isEven_plus.
|
Hint Resolve isEven_minus2 isEven_plus : core.
|
||||||
|
|
||||||
Lemma manually_proved_invariant'_snazzy : forall n,
|
Lemma manually_proved_invariant'_snazzy : forall n,
|
||||||
isEven n
|
isEven n
|
||||||
|
@ -863,7 +866,7 @@ Qed.
|
||||||
(* Bonus material: here's how to make these proofs much more automatic. We
|
(* Bonus material: here's how to make these proofs much more automatic. We
|
||||||
* won't explain the features being used here. *)
|
* won't explain the features being used here. *)
|
||||||
|
|
||||||
Hint Constructors plug step0 cstep.
|
Hint Constructors plug step0 cstep : core.
|
||||||
|
|
||||||
Theorem step_cstep_snazzy : forall v c v' c',
|
Theorem step_cstep_snazzy : forall v c v' c',
|
||||||
step (v, c) (v', c')
|
step (v, c) (v', c')
|
||||||
|
@ -874,7 +877,7 @@ Proof.
|
||||||
end; eauto.
|
end; eauto.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
Hint Resolve step_cstep_snazzy.
|
Hint Resolve step_cstep_snazzy : core.
|
||||||
|
|
||||||
Lemma step0_step_snazzy : forall v c v' c',
|
Lemma step0_step_snazzy : forall v c v' c',
|
||||||
step0 (v, c) (v', c')
|
step0 (v, c) (v', c')
|
||||||
|
@ -883,7 +886,7 @@ Proof.
|
||||||
invert 1; eauto.
|
invert 1; eauto.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
Hint Resolve step0_step_snazzy.
|
Hint Resolve step0_step_snazzy : core.
|
||||||
|
|
||||||
Lemma cstep_step'_snazzy : forall C c0 c,
|
Lemma cstep_step'_snazzy : forall C c0 c,
|
||||||
plug C c0 c
|
plug C c0 c
|
||||||
|
@ -896,7 +899,7 @@ Proof.
|
||||||
end; eauto.
|
end; eauto.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
Hint Resolve cstep_step'_snazzy.
|
Hint Resolve cstep_step'_snazzy : core.
|
||||||
|
|
||||||
Theorem cstep_step_snazzy : forall v c v' c',
|
Theorem cstep_step_snazzy : forall v c v' c',
|
||||||
cstep (v, c) (v', c')
|
cstep (v, c) (v', c')
|
||||||
|
@ -957,7 +960,7 @@ Proof.
|
||||||
apply cstep_step.
|
apply cstep_step.
|
||||||
eassumption.
|
eassumption.
|
||||||
Qed.
|
Qed.
|
||||||
|
End Simple.
|
||||||
|
|
||||||
(** * Example of how easy it is to add concurrency to a contextual semantics *)
|
(** * Example of how easy it is to add concurrency to a contextual semantics *)
|
||||||
|
|
||||||
|
@ -1041,7 +1044,7 @@ Module Concurrent.
|
||||||
|| ("b" <- "n";;
|
|| ("b" <- "n";;
|
||||||
"n" <- "b" + 1).
|
"n" <- "b" + 1).
|
||||||
|
|
||||||
Hint Constructors plug step0 cstep.
|
Hint Constructors plug step0 cstep : core.
|
||||||
|
|
||||||
(* The naive "expected" answer is attainable. *)
|
(* The naive "expected" answer is attainable. *)
|
||||||
Theorem correctAnswer : forall n, exists v, cstep^* ($0 $+ ("n", n), prog) (v, Skip)
|
Theorem correctAnswer : forall n, exists v, cstep^* ($0 $+ ("n", n), prog) (v, Skip)
|
||||||
|
@ -1146,7 +1149,7 @@ Module Concurrent.
|
||||||
step (v, c2) (v', c2')
|
step (v, c2) (v', c2')
|
||||||
-> step (v, Parallel c1 c2) (v', Parallel c1 c2').
|
-> step (v, Parallel c1 c2) (v', Parallel c1 c2').
|
||||||
|
|
||||||
Hint Constructors step.
|
Hint Constructors step : core.
|
||||||
|
|
||||||
(* Now, an automated proof of equivalence. Actually, it's *exactly* the same
|
(* Now, an automated proof of equivalence. Actually, it's *exactly* the same
|
||||||
* proof we used for the old feature set! *)
|
* proof we used for the old feature set! *)
|
||||||
|
@ -1160,7 +1163,7 @@ Module Concurrent.
|
||||||
end; eauto.
|
end; eauto.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
Hint Resolve step_cstep.
|
Hint Resolve step_cstep : core.
|
||||||
|
|
||||||
Lemma step0_step : forall v c v' c',
|
Lemma step0_step : forall v c v' c',
|
||||||
step0 (v, c) (v', c')
|
step0 (v, c) (v', c')
|
||||||
|
@ -1169,7 +1172,7 @@ Module Concurrent.
|
||||||
invert 1; eauto.
|
invert 1; eauto.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
Hint Resolve step0_step.
|
Hint Resolve step0_step : core.
|
||||||
|
|
||||||
Lemma cstep_step' : forall C c0 c,
|
Lemma cstep_step' : forall C c0 c,
|
||||||
plug C c0 c
|
plug C c0 c
|
||||||
|
@ -1182,7 +1185,7 @@ Module Concurrent.
|
||||||
end; eauto.
|
end; eauto.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
Hint Resolve cstep_step'.
|
Hint Resolve cstep_step' : core.
|
||||||
|
|
||||||
Theorem cstep_step : forall v c v' c',
|
Theorem cstep_step : forall v c v' c',
|
||||||
cstep (v, c) (v', c')
|
cstep (v, c) (v', c')
|
||||||
|
|
|
@ -360,7 +360,7 @@ Qed.
|
||||||
(* Automated proofs used here, if only to save time in class.
|
(* Automated proofs used here, if only to save time in class.
|
||||||
* See book code for more manual proofs, too. *)
|
* See book code for more manual proofs, too. *)
|
||||||
|
|
||||||
Hint Constructors trc step eval.
|
Hint Constructors trc step eval : core.
|
||||||
|
|
||||||
Theorem big_small : forall v c v', eval v c v'
|
Theorem big_small : forall v c v', eval v c v'
|
||||||
-> step^* (v, c) (v', Skip).
|
-> step^* (v, c) (v', Skip).
|
||||||
|
@ -430,8 +430,8 @@ Proof.
|
||||||
assumption.
|
assumption.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
Hint Constructors isEven.
|
Hint Constructors isEven : core.
|
||||||
Hint Resolve isEven_minus2 isEven_plus.
|
Hint Resolve isEven_minus2 isEven_plus : core.
|
||||||
|
|
||||||
Definition my_loop :=
|
Definition my_loop :=
|
||||||
while "n" loop
|
while "n" loop
|
||||||
|
@ -581,7 +581,7 @@ Inductive cstep : valuation * cmd -> valuation * cmd -> Prop :=
|
||||||
|
|
||||||
(* We can prove equivalence between the two formulations. *)
|
(* We can prove equivalence between the two formulations. *)
|
||||||
|
|
||||||
Hint Constructors plug step0 cstep.
|
Hint Constructors plug step0 cstep : core.
|
||||||
|
|
||||||
Theorem step_cstep : forall v c v' c',
|
Theorem step_cstep : forall v c v' c',
|
||||||
step (v, c) (v', c')
|
step (v, c) (v', c')
|
||||||
|
@ -614,176 +614,6 @@ Admitted.
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
(** * Example of how easy it is to add concurrency to a contextual semantics *)
|
|
||||||
|
|
||||||
Module Concurrent.
|
|
||||||
(* Let's add a construct for *parallel execution* of two commands. Such
|
|
||||||
* parallelism may be nested arbitrarily. *)
|
|
||||||
Inductive cmd :=
|
|
||||||
| Skip
|
|
||||||
| Assign (x : var) (e : arith)
|
|
||||||
| Sequence (c1 c2 : cmd)
|
|
||||||
| If (e : arith) (then_ else_ : cmd)
|
|
||||||
| While (e : arith) (body : cmd)
|
|
||||||
| Parallel (c1 c2 : cmd).
|
|
||||||
|
|
||||||
Notation "x <- e" := (Assign x e%arith) (at level 75).
|
|
||||||
Infix ";;" := Sequence (at level 76). (* This one changed slightly, to avoid parsing clashes. *)
|
|
||||||
Notation "'when' e 'then' then_ 'else' else_ 'done'" := (If e%arith then_ else_) (at level 75, e at level 0).
|
|
||||||
Notation "'while' e 'loop' body 'done'" := (While e%arith body) (at level 75).
|
|
||||||
Infix "||" := Parallel.
|
|
||||||
|
|
||||||
|
|
||||||
(* We need surprisingly few changes to the contextual semantics, to explain
|
|
||||||
* this new feature. First, we allow a hole to appear on *either side* of a
|
|
||||||
* [Parallel]. In other words, the "scheduler" may choose either "thread" to
|
|
||||||
* run next. *)
|
|
||||||
Inductive context :=
|
|
||||||
| Hole
|
|
||||||
| CSeq (C : context) (c : cmd)
|
|
||||||
| CPar1 (C : context) (c : cmd)
|
|
||||||
| CPar2 (c : cmd) (C : context).
|
|
||||||
|
|
||||||
(* We explain the meaning of plugging the new contexts in the obvious way. *)
|
|
||||||
Inductive plug : context -> cmd -> cmd -> Prop :=
|
|
||||||
| PlugHole : forall c, plug Hole c c
|
|
||||||
| PlugSeq : forall c C c' c2,
|
|
||||||
plug C c c'
|
|
||||||
-> plug (CSeq C c2) c (Sequence c' c2)
|
|
||||||
| PlugPar1 : forall c C c' c2,
|
|
||||||
plug C c c'
|
|
||||||
-> plug (CPar1 C c2) c (Parallel c' c2)
|
|
||||||
| PlugPar2 : forall c C c' c1,
|
|
||||||
plug C c c'
|
|
||||||
-> plug (CPar2 c1 C) c (Parallel c1 c').
|
|
||||||
|
|
||||||
(* The only new step rules are for "cleaning up" finished "threads," which
|
|
||||||
* have reached the point of being [Skip] commands. *)
|
|
||||||
Inductive step0 : valuation * cmd -> valuation * cmd -> Prop :=
|
|
||||||
| Step0Assign : forall v x e,
|
|
||||||
step0 (v, Assign x e) (v $+ (x, interp e v), Skip)
|
|
||||||
| Step0Seq : forall v c2,
|
|
||||||
step0 (v, Sequence Skip c2) (v, c2)
|
|
||||||
| Step0IfTrue : forall v e then_ else_,
|
|
||||||
interp e v <> 0
|
|
||||||
-> step0 (v, If e then_ else_) (v, then_)
|
|
||||||
| Step0IfFalse : forall v e then_ else_,
|
|
||||||
interp e v = 0
|
|
||||||
-> step0 (v, If e then_ else_) (v, else_)
|
|
||||||
| Step0WhileTrue : forall v e body,
|
|
||||||
interp e v <> 0
|
|
||||||
-> step0 (v, While e body) (v, Sequence body (While e body))
|
|
||||||
| Step0WhileFalse : forall v e body,
|
|
||||||
interp e v = 0
|
|
||||||
-> step0 (v, While e body) (v, Skip)
|
|
||||||
| Step0Par1 : forall v c,
|
|
||||||
step0 (v, Parallel Skip c) (v, c).
|
|
||||||
|
|
||||||
Inductive cstep : valuation * cmd -> valuation * cmd -> Prop :=
|
|
||||||
| CStep : forall C v c v' c' c1 c2,
|
|
||||||
plug C c c1
|
|
||||||
-> step0 (v, c) (v', c')
|
|
||||||
-> plug C c' c2
|
|
||||||
-> cstep (v, c1) (v', c2).
|
|
||||||
|
|
||||||
(** Example: stepping a specific program. *)
|
|
||||||
|
|
||||||
(* Here's the classic cautionary-tale program about remembering to lock your
|
|
||||||
* concurrent threads. *)
|
|
||||||
Definition prog :=
|
|
||||||
("a" <- "n";;
|
|
||||||
"n" <- "a" + 1)
|
|
||||||
|| ("b" <- "n";;
|
|
||||||
"n" <- "b" + 1).
|
|
||||||
|
|
||||||
Hint Constructors plug step0 cstep.
|
|
||||||
|
|
||||||
(* The naive "expected" answer is attainable. *)
|
|
||||||
Theorem correctAnswer : forall n, exists v, cstep^* ($0 $+ ("n", n), prog) (v, Skip)
|
|
||||||
/\ v $? "n" = Some (n + 2).
|
|
||||||
Proof.
|
|
||||||
eexists; propositional.
|
|
||||||
unfold prog.
|
|
||||||
|
|
||||||
econstructor.
|
|
||||||
eapply CStep with (C := CPar1 (CSeq Hole _) _); eauto.
|
|
||||||
|
|
||||||
econstructor.
|
|
||||||
eapply CStep with (C := CPar1 Hole _); eauto.
|
|
||||||
|
|
||||||
econstructor.
|
|
||||||
eapply CStep with (C := CPar1 Hole _); eauto.
|
|
||||||
|
|
||||||
econstructor.
|
|
||||||
eapply CStep with (C := Hole); eauto.
|
|
||||||
|
|
||||||
econstructor.
|
|
||||||
eapply CStep with (C := CSeq Hole _); eauto.
|
|
||||||
|
|
||||||
econstructor.
|
|
||||||
eapply CStep with (C := Hole); eauto.
|
|
||||||
|
|
||||||
econstructor.
|
|
||||||
eapply CStep with (C := Hole); eauto.
|
|
||||||
|
|
||||||
econstructor.
|
|
||||||
|
|
||||||
simplify.
|
|
||||||
f_equal.
|
|
||||||
ring.
|
|
||||||
Qed.
|
|
||||||
|
|
||||||
(* But so is the "wrong" answer! *)
|
|
||||||
Theorem wrongAnswer : forall n, exists v, cstep^* ($0 $+ ("n", n), prog) (v, Skip)
|
|
||||||
/\ v $? "n" = Some (n + 1).
|
|
||||||
Proof.
|
|
||||||
eexists; propositional.
|
|
||||||
unfold prog.
|
|
||||||
Admitted.
|
|
||||||
|
|
||||||
|
|
||||||
(** Proving equivalence with non-contextual semantics *)
|
|
||||||
|
|
||||||
(* To give us something interesting to prove, let's also define a
|
|
||||||
* non-contextual small-step semantics. Note how we have to do some more
|
|
||||||
* explicit threading of mutable state through recursive invocations. *)
|
|
||||||
Inductive step : valuation * cmd -> valuation * cmd -> Prop :=
|
|
||||||
| StepAssign : forall v x e,
|
|
||||||
step (v, Assign x e) (v $+ (x, interp e v), Skip)
|
|
||||||
| StepSeq1 : forall v c1 c2 v' c1',
|
|
||||||
step (v, c1) (v', c1')
|
|
||||||
-> step (v, Sequence c1 c2) (v', Sequence c1' c2)
|
|
||||||
| StepSeq2 : forall v c2,
|
|
||||||
step (v, Sequence Skip c2) (v, c2)
|
|
||||||
| StepIfTrue : forall v e then_ else_,
|
|
||||||
interp e v <> 0
|
|
||||||
-> step (v, If e then_ else_) (v, then_)
|
|
||||||
| StepIfFalse : forall v e then_ else_,
|
|
||||||
interp e v = 0
|
|
||||||
-> step (v, If e then_ else_) (v, else_)
|
|
||||||
| StepWhileTrue : forall v e body,
|
|
||||||
interp e v <> 0
|
|
||||||
-> step (v, While e body) (v, Sequence body (While e body))
|
|
||||||
| StepWhileFalse : forall v e body,
|
|
||||||
interp e v = 0
|
|
||||||
-> step (v, While e body) (v, Skip)
|
|
||||||
| StepParSkip1 : forall v c,
|
|
||||||
step (v, Parallel Skip c) (v, c)
|
|
||||||
| StepPar1 : forall v c1 c2 v' c1',
|
|
||||||
step (v, c1) (v', c1')
|
|
||||||
-> step (v, Parallel c1 c2) (v', Parallel c1' c2)
|
|
||||||
| StepPar2 : forall v c1 c2 v' c2',
|
|
||||||
step (v, c2) (v', c2')
|
|
||||||
-> step (v, Parallel c1 c2) (v', Parallel c1 c2').
|
|
||||||
|
|
||||||
Hint Constructors step.
|
|
||||||
|
|
||||||
(* Now, an automated proof of equivalence. Actually, it's *exactly* the same
|
|
||||||
* proof we used for the old feature set! For full dramatic effect, copy and
|
|
||||||
* paste here from above. *)
|
|
||||||
End Concurrent.
|
|
||||||
|
|
||||||
|
|
||||||
(** * Determinism *)
|
(** * Determinism *)
|
||||||
|
@ -837,3 +667,91 @@ Proof.
|
||||||
apply cstep_step.
|
apply cstep_step.
|
||||||
eassumption.
|
eassumption.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
(** * Example of how easy it is to add concurrency to a contextual semantics *)
|
||||||
|
|
||||||
|
(** At this point, we add concurrency to the code we already wrote above. *)
|
||||||
|
|
||||||
|
(*
|
||||||
|
(* Here's the classic cautionary-tale program about remembering to lock your
|
||||||
|
* concurrent threads. *)
|
||||||
|
Definition prog :=
|
||||||
|
("a" <- "n";;
|
||||||
|
"n" <- "a" + 1)
|
||||||
|
|| ("b" <- "n";;
|
||||||
|
"n" <- "b" + 1).
|
||||||
|
|
||||||
|
Hint Constructors plug step0 cstep : core.
|
||||||
|
|
||||||
|
(* The naive "expected" answer is attainable. *)
|
||||||
|
Theorem correctAnswer : forall n, exists v, cstep^* ($0 $+ ("n", n), prog) (v, Skip)
|
||||||
|
/\ v $? "n" = Some (n + 2).
|
||||||
|
Proof.
|
||||||
|
eexists; propositional.
|
||||||
|
unfold prog.
|
||||||
|
|
||||||
|
econstructor.
|
||||||
|
eapply CStep with (C := CPar1 (CSeq Hole _) _); eauto.
|
||||||
|
|
||||||
|
econstructor.
|
||||||
|
eapply CStep with (C := CPar1 Hole _); eauto.
|
||||||
|
|
||||||
|
econstructor.
|
||||||
|
eapply CStep with (C := CPar1 Hole _); eauto.
|
||||||
|
|
||||||
|
econstructor.
|
||||||
|
eapply CStep with (C := Hole); eauto.
|
||||||
|
|
||||||
|
econstructor.
|
||||||
|
eapply CStep with (C := CSeq Hole _); eauto.
|
||||||
|
|
||||||
|
econstructor.
|
||||||
|
eapply CStep with (C := Hole); eauto.
|
||||||
|
|
||||||
|
econstructor.
|
||||||
|
eapply CStep with (C := Hole); eauto.
|
||||||
|
|
||||||
|
econstructor.
|
||||||
|
|
||||||
|
simplify.
|
||||||
|
f_equal.
|
||||||
|
ring.
|
||||||
|
Qed.
|
||||||
|
|
||||||
|
(* But so is the "wrong" answer! *)
|
||||||
|
Theorem wrongAnswer : forall n, exists v, cstep^* ($0 $+ ("n", n), prog) (v, Skip)
|
||||||
|
/\ v $? "n" = Some (n + 1).
|
||||||
|
Proof.
|
||||||
|
eexists; propositional.
|
||||||
|
unfold prog.
|
||||||
|
Admitted.
|
||||||
|
*)
|
||||||
|
|
Loading…
Add table
Reference in a new issue