summaryrefslogtreecommitdiff
path: root/common
diff options
context:
space:
mode:
authorGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2009-08-16 12:53:19 +0000
committerGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2009-08-16 12:53:19 +0000
commit1fe28ba1ec3dd0657b121c4a911ee1cb046bab09 (patch)
tree7f3b22fade6b3d7b7871624aa0ccf4ef52a10e84 /common
parentf8d59bccd57fd53b55de5e9dd96fbc1af150951a (diff)
Distinguish two kinds of nonterminating behaviors: silent divergence
and reactive divergence. As a consequence: - Removed the Enilinf constructor from traceinf (values of traceinf type are always infinite traces). - Traces are now uniquely defined. - Adapted proofs big step -> small step for Clight and Cminor accordingly. - Strengthened results in driver/Complements accordingly. - Added common/Determinism to collect generic results about deterministic semantics. git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@1123 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e
Diffstat (limited to 'common')
-rw-r--r--common/Determinism.v508
-rw-r--r--common/Events.v55
-rw-r--r--common/Smallstep.v287
3 files changed, 790 insertions, 60 deletions
diff --git a/common/Determinism.v b/common/Determinism.v
new file mode 100644
index 0000000..430ee93
--- /dev/null
+++ b/common/Determinism.v
@@ -0,0 +1,508 @@
+(* *********************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Xavier Leroy, INRIA Paris-Rocquencourt *)
+(* *)
+(* Copyright Institut National de Recherche en Informatique et en *)
+(* Automatique. All rights reserved. This file is distributed *)
+(* under the terms of the INRIA Non-Commercial License Agreement. *)
+(* *)
+(* *********************************************************************)
+
+(** Characterization and properties of deterministic semantics *)
+
+Require Import Classical.
+Require Import Coqlib.
+Require Import AST.
+Require Import Integers.
+Require Import Values.
+Require Import Events.
+Require Import Globalenvs.
+Require Import Smallstep.
+
+(** This file uses classical logic (the axiom of excluded middle), as
+ well as the following extensionality property over infinite
+ sequences of events. All these axioms are sound in a set-theoretic
+ model of Coq's logic. *)
+
+Axiom traceinf_extensionality:
+ forall T T', traceinf_sim T T' -> T = T'.
+
+(** * Deterministic worlds *)
+
+(** One source of possible nondeterminism is that our semantics leave
+ unspecified the results of calls to external
+ functions. Different results to e.g. a "read" operation can of
+ course lead to different behaviors for the program.
+ We address this issue by modeling a notion of deterministic
+ external world that uniquely determines the results of external calls. *)
+
+(** An external world is a function that, given the name of an
+ external call and its arguments, returns either [None], meaning
+ that this external call gets stuck, or [Some(r,w)], meaning
+ that this external call succeeds, has result [r], and changes
+ the world to [w]. *)
+
+Inductive world: Type :=
+ World: (ident -> list eventval -> option (eventval * world)) -> world.
+
+Definition nextworld (w: world) (evname: ident) (evargs: list eventval) :
+ option (eventval * world) :=
+ match w with World f => f evname evargs end.
+
+(** A trace is possible in a given world if all events correspond
+ to non-stuck external calls according to the given world.
+ Two predicates are defined, for finite and infinite traces respectively:
+- [possible_trace w t w'], where [w] is the initial state of the
+ world, [t] the finite trace of interest, and [w'] the state of the
+ world after performing trace [t].
+- [possible_traceinf w T], where [w] is the initial state of the
+ world and [T] the infinite trace of interest.
+*)
+
+Inductive possible_trace: world -> trace -> world -> Prop :=
+ | possible_trace_nil: forall w,
+ possible_trace w E0 w
+ | possible_trace_cons: forall w0 evname evargs evres w1 t w2,
+ nextworld w0 evname evargs = Some (evres, w1) ->
+ possible_trace w1 t w2 ->
+ possible_trace w0 (mkevent evname evargs evres :: t) w2.
+
+Lemma possible_trace_app:
+ forall t2 w2 w0 t1 w1,
+ possible_trace w0 t1 w1 -> possible_trace w1 t2 w2 ->
+ possible_trace w0 (t1 ** t2) w2.
+Proof.
+ induction 1; simpl; intros.
+ auto.
+ econstructor; eauto.
+Qed.
+
+Lemma possible_trace_app_inv:
+ forall t2 w2 t1 w0,
+ possible_trace w0 (t1 ** t2) w2 ->
+ exists w1, possible_trace w0 t1 w1 /\ possible_trace w1 t2 w2.
+Proof.
+ induction t1; simpl; intros.
+ exists w0; split. constructor. auto.
+ inv H. exploit IHt1; eauto. intros [w1 [A B]].
+ exists w1; split. econstructor; eauto. auto.
+Qed.
+
+CoInductive possible_traceinf: world -> traceinf -> Prop :=
+ | possible_traceinf_cons: forall w0 evname evargs evres w1 T,
+ nextworld w0 evname evargs = Some (evres, w1) ->
+ possible_traceinf w1 T ->
+ possible_traceinf w0 (Econsinf (mkevent evname evargs evres) T).
+
+Lemma possible_traceinf_app:
+ forall t2 w0 t1 w1,
+ possible_trace w0 t1 w1 -> possible_traceinf w1 t2 ->
+ possible_traceinf w0 (t1 *** t2).
+Proof.
+ induction 1; simpl; intros.
+ auto.
+ econstructor; eauto.
+Qed.
+
+Lemma possible_traceinf_app_inv:
+ forall t2 t1 w0,
+ possible_traceinf w0 (t1 *** t2) ->
+ exists w1, possible_trace w0 t1 w1 /\ possible_traceinf w1 t2.
+Proof.
+ induction t1; simpl; intros.
+ exists w0; split. constructor. auto.
+ inv H. exploit IHt1; eauto. intros [w1 [A B]].
+ exists w1; split. econstructor; eauto. auto.
+Qed.
+
+Ltac possibleTraceInv :=
+ match goal with
+ | [H: possible_trace _ E0 _ |- _] =>
+ inversion H; clear H; subst; possibleTraceInv
+ | [H: possible_trace _ (_ ** _) _ |- _] =>
+ let P1 := fresh "P" in
+ let w := fresh "w" in
+ let P2 := fresh "P" in
+ elim (possible_trace_app_inv _ _ _ _ H); clear H;
+ intros w [P1 P2];
+ possibleTraceInv
+ | [H: possible_traceinf _ (_ *** _) |- _] =>
+ let P1 := fresh "P" in
+ let w := fresh "w" in
+ let P2 := fresh "P" in
+ elim (possible_traceinf_app_inv _ _ _ H); clear H;
+ intros w [P1 P2];
+ possibleTraceInv
+ | [H: exists w, possible_trace _ _ w |- _] =>
+ let P := fresh "P" in let w := fresh "w" in
+ destruct H as [w P]; possibleTraceInv
+ | _ => idtac
+ end.
+
+Definition possible_behavior (w: world) (b: program_behavior) : Prop :=
+ match b with
+ | Terminates t r => exists w', possible_trace w t w'
+ | Diverges t => exists w', possible_trace w t w'
+ | Reacts T => possible_traceinf w T
+ | Goes_wrong t => exists w', possible_trace w t w'
+ end.
+
+(** Determinism properties of [event_match]. *)
+
+Remark eventval_match_deterministic:
+ forall ev1 ev2 ty v1 v2,
+ eventval_match ev1 ty v1 -> eventval_match ev2 ty v2 ->
+ (ev1 = ev2 <-> v1 = v2).
+Proof.
+ intros. inv H; inv H0; intuition congruence.
+Qed.
+
+Remark eventval_list_match_deterministic:
+ forall ev1 ty v, eventval_list_match ev1 ty v ->
+ forall ev2, eventval_list_match ev2 ty v -> ev1 = ev2.
+Proof.
+ induction 1; intros.
+ inv H. auto.
+ inv H1. decEq.
+ rewrite (eventval_match_deterministic _ _ _ _ _ H H6). auto.
+ eauto.
+Qed.
+
+(** * Deterministic semantics *)
+
+Section DETERM_SEM.
+
+(** We assume given a transition semantics that is internally
+ deterministic: the only source of non-determinism is the return
+ value of external calls. *)
+
+Variable genv: Type.
+Variable state: Type.
+Variable step: genv -> state -> trace -> state -> Prop.
+Variable initial_state: state -> Prop.
+Variable final_state: state -> int -> Prop.
+
+Inductive internal_determinism: trace -> state -> trace -> state -> Prop :=
+ | int_determ_0: forall s,
+ internal_determinism E0 s E0 s
+ | int_determ_1: forall s s' id arg res res',
+ (res = res' -> s = s') ->
+ internal_determinism (mkevent id arg res :: nil) s
+ (mkevent id arg res' :: nil) s'.
+
+Hypothesis step_internal_deterministic:
+ forall ge s t1 s1 t2 s2,
+ step ge s t1 s1 -> step ge s t2 s2 -> internal_determinism t1 s1 t2 s2.
+
+Hypothesis initial_state_determ:
+ forall s1 s2, initial_state s1 -> initial_state s2 -> s1 = s2.
+
+Hypothesis final_state_determ:
+ forall st r1 r2, final_state st r1 -> final_state st r2 -> r1 = r2.
+
+Hypothesis final_state_nostep:
+ forall ge st r, final_state st r -> nostep step ge st.
+
+(** Consequently, the [step] relation is deterministic if restricted
+ to traces that are possible in a deterministic world. *)
+
+Lemma step_deterministic:
+ forall ge s0 t1 s1 t2 s2 w0 w1 w2,
+ step ge s0 t1 s1 -> step ge s0 t2 s2 ->
+ possible_trace w0 t1 w1 -> possible_trace w0 t2 w2 ->
+ s1 = s2 /\ t1 = t2 /\ w1 = w2.
+Proof.
+ intros. exploit step_internal_deterministic. eexact H. eexact H0. intro ID.
+ inv ID.
+ inv H1. inv H2. auto.
+ inv H2. inv H11. inv H1. inv H11.
+ rewrite H10 in H9. inv H9.
+ intuition.
+Qed.
+
+Ltac use_step_deterministic :=
+ match goal with
+ | [ S1: step _ _ ?t1 _, P1: possible_trace _ ?t1 _,
+ S2: step _ _ ?t2 _, P2: possible_trace _ ?t2 _ |- _ ] =>
+ destruct (step_deterministic _ _ _ _ _ _ _ _ _ S1 S2 P1 P2)
+ as [EQ1 [EQ2 EQ3]]; subst
+ end.
+
+(** Determinism for finite transition sequences. *)
+
+Lemma star_step_diamond:
+ forall ge s0 t1 s1, star step ge s0 t1 s1 ->
+ forall t2 s2 w0 w1 w2, star step ge s0 t2 s2 ->
+ possible_trace w0 t1 w1 -> possible_trace w0 t2 w2 ->
+ exists t,
+ (star step ge s1 t s2 /\ possible_trace w1 t w2 /\ t2 = t1 ** t)
+ \/ (star step ge s2 t s1 /\ possible_trace w2 t w1 /\ t1 = t2 ** t).
+Proof.
+ induction 1; intros.
+ inv H0. exists t2; auto.
+ inv H2. inv H4. exists (t1 ** t2); right.
+ split. econstructor; eauto. auto.
+ possibleTraceInv. use_step_deterministic.
+ exploit IHstar. eexact H6. eauto. eauto.
+ intros [t A]. exists t.
+ destruct A. left; intuition. traceEq. right; intuition. traceEq.
+Qed.
+
+Ltac use_star_step_diamond :=
+ match goal with
+ | [ S1: star step _ _ ?t1 _, P1: possible_trace _ ?t1 _,
+ S2: star step _ _ ?t2 _, P2: possible_trace _ ?t2 _ |- _ ] =>
+ let t := fresh "t" in let P := fresh "P" in let Q := fresh "Q" in let EQ := fresh "EQ" in
+ destruct (star_step_diamond _ _ _ _ S1 _ _ _ _ _ S2 P1 P2)
+ as [t [ [P [Q EQ]] | [P [Q EQ]] ]]; subst
+ end.
+
+Ltac use_nostep :=
+ match goal with
+ | [ S: step _ ?s _ _, NO: nostep step _ ?s |- _ ] => elim (NO _ _ S)
+ end.
+
+Lemma star_step_triangle:
+ forall ge s0 t1 s1 t2 s2 w0 w1 w2,
+ star step ge s0 t1 s1 ->
+ star step ge s0 t2 s2 ->
+ possible_trace w0 t1 w1 -> possible_trace w0 t2 w2 ->
+ nostep step ge s2 ->
+ exists t,
+ star step ge s1 t s2 /\ possible_trace w1 t w2 /\ t2 = t1 ** t.
+Proof.
+ intros. use_star_step_diamond; possibleTraceInv.
+ exists t; auto.
+ inv P. inv Q. exists E0. split. constructor. split. constructor. traceEq.
+ use_nostep.
+Qed.
+
+Ltac use_star_step_triangle :=
+ match goal with
+ | [ S1: star step _ _ ?t1 _, P1: possible_trace _ ?t1 _,
+ S2: star step _ _ ?t2 ?s2, P2: possible_trace _ ?t2 _,
+ NO: nostep step _ ?s2 |- _ ] =>
+ let t := fresh "t" in let P := fresh "P" in let Q := fresh "Q" in let EQ := fresh "EQ" in
+ destruct (star_step_triangle _ _ _ _ _ _ _ _ _ S1 S2 P1 P2 NO)
+ as [t [P [Q EQ]]]; subst
+ end.
+
+Lemma steps_deterministic:
+ forall ge s0 t1 s1 t2 s2 w0 w1 w2,
+ star step ge s0 t1 s1 -> star step ge s0 t2 s2 ->
+ nostep step ge s1 -> nostep step ge s2 ->
+ possible_trace w0 t1 w1 -> possible_trace w0 t2 w2 ->
+ t1 = t2 /\ s1 = s2.
+Proof.
+ intros. use_star_step_triangle. inv P.
+ inv Q. split; auto; traceEq. use_nostep.
+Qed.
+
+Lemma terminates_not_goes_wrong:
+ forall ge s t1 s1 r w w1 t2 s2 w2,
+ star step ge s t1 s1 -> final_state s1 r -> possible_trace w t1 w1 ->
+ star step ge s t2 s2 -> nostep step ge s2 -> possible_trace w t2 w2 ->
+ (forall r, ~final_state s2 r) -> False.
+Proof.
+ intros.
+ assert (t1 = t2 /\ s1 = s2).
+ eapply steps_deterministic; eauto.
+ destruct H6; subst. elim (H5 _ H0).
+Qed.
+
+(** Determinism for infinite transition sequences. *)
+
+Lemma star_final_not_forever_silent:
+ forall ge s t s', star step ge s t s' ->
+ forall w w', possible_trace w t w' -> nostep step ge s' ->
+ forever_silent step ge s -> False.
+Proof.
+ induction 1; intros.
+ inv H1. use_nostep.
+ inv H4. possibleTraceInv. assert (possible_trace w E0 w) by constructor.
+ use_step_deterministic. eauto.
+Qed.
+
+Lemma star2_final_not_forever_silent:
+ forall ge s t1 s1 w w1 t2 s2 w2,
+ star step ge s t1 s1 -> possible_trace w t1 w1 -> nostep step ge s1 ->
+ star step ge s t2 s2 -> possible_trace w t2 w2 -> forever_silent step ge s2 ->
+ False.
+Proof.
+ intros. use_star_step_triangle. possibleTraceInv.
+ eapply star_final_not_forever_silent. eexact P. eauto. auto. auto.
+Qed.
+
+Lemma star_final_not_forever_reactive:
+ forall ge s t s', star step ge s t s' ->
+ forall w w' T, possible_trace w t w' -> possible_traceinf w T -> nostep step ge s' ->
+ forever_reactive step ge s T -> False.
+Proof.
+ induction 1; intros.
+ inv H2. inv H3. congruence. use_nostep.
+ inv H5. possibleTraceInv. inv H6. congruence. possibleTraceInv.
+ use_step_deterministic.
+ eapply IHstar with (T := t4 *** T0). eauto.
+ eapply possible_traceinf_app; eauto. auto.
+ eapply star_forever_reactive; eauto.
+Qed.
+
+Lemma star_forever_silent_inv:
+ forall ge s t s', star step ge s t s' ->
+ forall w w', possible_trace w t w' ->
+ forever_silent step ge s ->
+ t = E0 /\ forever_silent step ge s'.
+Proof.
+ induction 1; intros.
+ auto.
+ subst. possibleTraceInv. inv H3. assert (possible_trace w E0 w) by constructor.
+ use_step_deterministic. eauto.
+Qed.
+
+Lemma forever_silent_reactive_exclusive:
+ forall ge s w T,
+ forever_silent step ge s -> forever_reactive step ge s T ->
+ possible_traceinf w T -> False.
+Proof.
+ intros. inv H0. possibleTraceInv. exploit star_forever_silent_inv; eauto.
+ intros [A B]. contradiction.
+Qed.
+
+Lemma forever_reactive_inv2:
+ forall ge s t1 s1, star step ge s t1 s1 ->
+ forall t2 s2 T1 T2 w w1 w2,
+ possible_trace w t1 w1 ->
+ star step ge s t2 s2 -> possible_trace w t2 w2 ->
+ t1 <> E0 -> t2 <> E0 ->
+ forever_reactive step ge s1 T1 -> possible_traceinf w1 T1 ->
+ forever_reactive step ge s2 T2 -> possible_traceinf w2 T2 ->
+ exists s', exists e, exists T1', exists T2', exists w',
+ forever_reactive step ge s' T1' /\ possible_traceinf w' T1' /\
+ forever_reactive step ge s' T2' /\ possible_traceinf w' T2' /\
+ t1 *** T1 = Econsinf e T1' /\
+ t2 *** T2 = Econsinf e T2'.
+Proof.
+ induction 1; intros.
+ congruence.
+ inv H3. congruence. possibleTraceInv.
+ assert (ID: internal_determinism t3 s5 t1 s2). eauto.
+ inv ID.
+ possibleTraceInv. eauto.
+ inv P. inv P1. inv H17. inv H19. rewrite H18 in H16; inv H16.
+ assert (s5 = s2) by auto. subst s5.
+ exists s2; exists (mkevent id arg res');
+ exists (t2 *** T1); exists (t4 *** T2); exists w0.
+ split. eapply star_forever_reactive; eauto.
+ split. eapply possible_traceinf_app; eauto.
+ split. eapply star_forever_reactive; eauto.
+ split. eapply possible_traceinf_app; eauto.
+ auto.
+Qed.
+
+Lemma forever_reactive_determ:
+ forall ge s T1 T2 w,
+ forever_reactive step ge s T1 -> possible_traceinf w T1 ->
+ forever_reactive step ge s T2 -> possible_traceinf w T2 ->
+ traceinf_sim T1 T2.
+Proof.
+ cofix COINDHYP; intros.
+ inv H. inv H1. possibleTraceInv.
+ destruct (forever_reactive_inv2 _ _ _ _ H _ _ _ _ _ _ _ P H3 P1 H6 H4
+ H7 P0 H5 P2)
+ as [s' [e [T1' [T2' [w' [A [B [C [D [E G]]]]]]]]]].
+ rewrite E; rewrite G. constructor.
+ eapply COINDHYP; eauto.
+Qed.
+
+Lemma star_forever_reactive_inv:
+ forall ge s t s', star step ge s t s' ->
+ forall w w' T, possible_trace w t w' -> forever_reactive step ge s T ->
+ possible_traceinf w T ->
+ exists T', forever_reactive step ge s' T' /\ possible_traceinf w' T' /\ T = t *** T'.
+Proof.
+ induction 1; intros.
+ possibleTraceInv. exists T; auto.
+ inv H3. possibleTraceInv. inv H5. congruence. possibleTraceInv.
+ use_step_deterministic.
+ exploit IHstar. eauto. eapply star_forever_reactive. 2: eauto. eauto.
+ eapply possible_traceinf_app; eauto.
+ intros [T' [A [B C]]]. exists T'; intuition. traceEq. congruence.
+Qed.
+
+Lemma forever_silent_reactive_exclusive2:
+ forall ge s t s' w w' T,
+ star step ge s t s' -> possible_trace w t w' -> forever_silent step ge s' ->
+ forever_reactive step ge s T -> possible_traceinf w T ->
+ False.
+Proof.
+ intros. exploit star_forever_reactive_inv; eauto.
+ intros [T' [A [B C]]]. subst T.
+ eapply forever_silent_reactive_exclusive; eauto.
+Qed.
+
+(** Determinism for program executions *)
+
+Ltac use_init_state :=
+ match goal with
+ | [ H1: (initial_state _), H2: (initial_state _) |- _ ] =>
+ generalize (initial_state_determ _ _ H1 H2); intro; subst; clear H2
+ | [ H1: (initial_state _), H2: (forall s, ~initial_state s) |- _ ] =>
+ elim (H2 _ H1)
+ | _ => idtac
+ end.
+
+Theorem program_behaves_deterministic:
+ forall ge w beh1 beh2,
+ program_behaves step initial_state final_state ge beh1 -> possible_behavior w beh1 ->
+ program_behaves step initial_state final_state ge beh2 -> possible_behavior w beh2 ->
+ beh1 = beh2.
+Proof.
+ intros until beh2; intros BEH1 POS1 BEH2 POS2.
+ inv BEH1; inv BEH2; simpl in POS1; simpl in POS2;
+ possibleTraceInv; use_init_state.
+(* terminates, terminates *)
+ assert (t = t0 /\ s' = s'0). eapply steps_deterministic; eauto.
+ destruct H2. f_equal; auto. subst. eauto.
+(* terminates, diverges *)
+ byContradiction. eapply star2_final_not_forever_silent with (s1 := s') (s2 := s'0); eauto.
+(* terminates, reacts *)
+ byContradiction. eapply star_final_not_forever_reactive; eauto.
+(* terminates, goes_wrong *)
+ byContradiction. eapply terminates_not_goes_wrong with (s1 := s') (s2 := s'0); eauto.
+(* diverges, terminates *)
+ byContradiction. eapply star2_final_not_forever_silent with (s2 := s') (s1 := s'0); eauto.
+(* diverges, diverges *)
+ f_equal. use_star_step_diamond.
+ exploit star_forever_silent_inv. eexact P1. eauto. eauto.
+ intros [A B]. subst; traceEq.
+ exploit star_forever_silent_inv. eexact P1. eauto. eauto.
+ intros [A B]. subst; traceEq.
+(* diverges, reacts *)
+ byContradiction. eapply forever_silent_reactive_exclusive2; eauto.
+(* diverges, goes wrong *)
+ byContradiction. eapply star2_final_not_forever_silent with (s1 := s'0) (s2 := s'); eauto.
+(* reacts, terminates *)
+ byContradiction. eapply star_final_not_forever_reactive; eauto.
+(* reacts, diverges *)
+ byContradiction. eapply forever_silent_reactive_exclusive2; eauto.
+(* reacts, reacts *)
+ f_equal. apply traceinf_extensionality.
+ eapply forever_reactive_determ; eauto.
+(* reacts, goes wrong *)
+ byContradiction. eapply star_final_not_forever_reactive; eauto.
+(* goes wrong, terminate *)
+ byContradiction. eapply terminates_not_goes_wrong with (s1 := s'0) (s2 := s'); eauto.
+(* goes wrong, diverges *)
+ byContradiction. eapply star2_final_not_forever_silent with (s1 := s') (s2 := s'0); eauto.
+(* goes wrong, reacts *)
+ byContradiction. eapply star_final_not_forever_reactive; eauto.
+(* goes wrong, goes wrong *)
+ assert (t = t0 /\ s' = s'0). eapply steps_deterministic; eauto.
+ destruct H3. congruence.
+(* goes initially wrong, goes initially wrong *)
+ reflexivity.
+Qed.
+
+End DETERM_SEM.
diff --git a/common/Events.v b/common/Events.v
index 011c454..0f5d9d2 100644
--- a/common/Events.v
+++ b/common/Events.v
@@ -43,8 +43,7 @@ Record event : Type := mkevent {
}.
(** The dynamic semantics for programs collect traces of events.
- Traces are of two kinds: finite (type [trace])
- or possibly infinite (type [traceinf]). *)
+ Traces are of two kinds: finite (type [trace]) or infinite (type [traceinf]). *)
Definition trace := list event.
@@ -57,7 +56,6 @@ Definition Eextcall
Definition Eapp (t1 t2: trace) : trace := t1 ++ t2.
CoInductive traceinf : Type :=
- | Enilinf: traceinf
| Econsinf: event -> traceinf -> traceinf.
Fixpoint Eappinf (t: trace) (T: traceinf) {struct t} : traceinf :=
@@ -81,6 +79,9 @@ Proof. intros. unfold E0, Eapp. rewrite <- app_nil_end. auto. Qed.
Lemma Eapp_assoc: forall t1 t2 t3, (t1 ** t2) ** t3 = t1 ** (t2 ** t3).
Proof. intros. unfold Eapp, trace. apply app_ass. Qed.
+Lemma Eapp_E0_inv: forall t1 t2, t1 ** t2 = E0 -> t1 = E0 /\ t2 = E0.
+Proof (@app_eq_nil event).
+
Lemma E0_left_inf: forall T, E0 *** T = T.
Proof. auto. Qed.
@@ -226,8 +227,6 @@ End EVENT_MATCH_LESSDEF.
(** Bisimilarity between infinite traces. *)
CoInductive traceinf_sim: traceinf -> traceinf -> Prop :=
- | traceinf_sim_nil:
- traceinf_sim Enilinf Enilinf
| traceinf_sim_cons: forall e T1 T2,
traceinf_sim T1 T2 ->
traceinf_sim (Econsinf e T1) (Econsinf e T2).
@@ -236,7 +235,7 @@ Lemma traceinf_sim_refl:
forall T, traceinf_sim T T.
Proof.
cofix COINDHYP; intros.
- destruct T. constructor. constructor. apply COINDHYP.
+ destruct T. constructor. apply COINDHYP.
Qed.
Lemma traceinf_sim_sym:
@@ -252,15 +251,16 @@ Proof.
cofix COINDHYP;intros. inv H; inv H0; constructor; eauto.
Qed.
-(** The "is prefix of" relation between infinite traces. *)
+(** The "is prefix of" relation between a finite and an infinite trace. *)
-CoInductive traceinf_prefix: traceinf -> traceinf -> Prop :=
+Inductive traceinf_prefix: trace -> traceinf -> Prop :=
| traceinf_prefix_nil: forall T,
- traceinf_prefix Enilinf T
- | traceinf_prefix_cons: forall e T1 T2,
- traceinf_prefix T1 T2 ->
- traceinf_prefix (Econsinf e T1) (Econsinf e T2).
+ traceinf_prefix nil T
+ | traceinf_prefix_cons: forall e t1 T2,
+ traceinf_prefix t1 T2 ->
+ traceinf_prefix (e :: t1) (Econsinf e T2).
+(*
Lemma traceinf_prefix_compat:
forall T1 T2 T1' T2',
traceinf_prefix T1 T2 -> traceinf_sim T1 T1' -> traceinf_sim T2 T2' ->
@@ -271,31 +271,16 @@ Proof.
Qed.
Transparent trace E0 Eapp Eappinf.
+*)
Lemma traceinf_prefix_app:
- forall T1 T2 t,
- traceinf_prefix T1 T2 ->
- traceinf_prefix (t *** T1) (t *** T2).
+ forall t1 T2 t,
+ traceinf_prefix t1 T2 ->
+ traceinf_prefix (t ** t1) (t *** T2).
Proof.
- induction t; simpl; intros. auto. constructor. auto.
-Qed.
-
-Lemma traceinf_sim_prefix:
- forall T1 T2,
- traceinf_sim T1 T2 -> traceinf_prefix T1 T2.
-Proof.
- cofix COINDHYP; intros.
- inv H. constructor. constructor. apply COINDHYP; auto.
-Qed.
-
-Lemma traceinf_prefix_2_sim:
- forall T1 T2,
- traceinf_prefix T1 T2 -> traceinf_prefix T2 T1 ->
- traceinf_sim T1 T2.
-Proof.
- cofix COINDHYP; intros.
- inv H.
- inv H0. constructor.
- inv H0. constructor. apply COINDHYP; auto.
+ induction t; simpl; intros. auto.
+ change ((a :: t) ** t1) with (a :: (t ** t1)).
+ change ((a :: t) *** T2) with (Econsinf a (t *** T2)).
+ constructor. auto.
Qed.
diff --git a/common/Smallstep.v b/common/Smallstep.v
index 9e9063b..deab49b 100644
--- a/common/Smallstep.v
+++ b/common/Smallstep.v
@@ -46,6 +46,11 @@ Variable state: Type.
Variable step: genv -> state -> trace -> state -> Prop.
+(** No transitions: stuck state *)
+
+Definition nostep (ge: genv) (s: state) : Prop :=
+ forall t s', ~(step ge s t s').
+
(** Zero, one or several transitions. Also known as Kleene closure,
or reflexive transitive closure. *)
@@ -280,28 +285,102 @@ Proof.
subst. econstructor; eauto.
Qed.
+(** Infinitely many silent transitions *)
+
+CoInductive forever_silent (ge: genv): state -> Prop :=
+ | forever_silent_intro: forall s1 s2,
+ step ge s1 E0 s2 -> forever_silent ge s2 ->
+ forever_silent ge s1.
+
+(** An alternate definition. *)
+
+CoInductive forever_silent_N (ge: genv) : A -> state -> Prop :=
+ | forever_silent_N_star: forall s1 s2 a1 a2,
+ star ge s1 E0 s2 ->
+ order a2 a1 ->
+ forever_silent_N ge a2 s2 ->
+ forever_silent_N ge a1 s1
+ | forever_silent_N_plus: forall s1 s2 a1 a2,
+ plus ge s1 E0 s2 ->
+ forever_silent_N ge a2 s2 ->
+ forever_silent_N ge a1 s1.
+
+Lemma forever_silent_N_inv:
+ forall ge a s,
+ forever_silent_N ge a s ->
+ exists s', exists a',
+ step ge s E0 s' /\ forever_silent_N ge a' s'.
+Proof.
+ intros ge a0. pattern a0. apply (well_founded_ind order_wf).
+ intros. inv H0.
+ (* star case *)
+ inv H1.
+ (* no transition *)
+ apply H with a2. auto. auto.
+ (* at least one transition *)
+ exploit Eapp_E0_inv; eauto. intros [P Q]. subst.
+ exists s0; exists x.
+ split. auto. eapply forever_silent_N_star; eauto.
+ (* plus case *)
+ inv H1. exploit Eapp_E0_inv; eauto. intros [P Q]. subst.
+ exists s0; exists a2.
+ split. auto. inv H3. auto.
+ eapply forever_silent_N_plus. econstructor; eauto. eauto.
+Qed.
+
+Lemma forever_silent_N_forever:
+ forall ge a s, forever_silent_N ge a s -> forever_silent ge s.
+Proof.
+ cofix COINDHYP; intros.
+ destruct (forever_silent_N_inv H) as [s' [a' [P Q]]].
+ apply forever_silent_intro with s'. auto.
+ apply COINDHYP with a'; auto.
+Qed.
+
+(** Infinitely many non-silent transitions *)
+
+CoInductive forever_reactive (ge: genv): state -> traceinf -> Prop :=
+ | forever_reactive_intro: forall s1 s2 t T,
+ star ge s1 t s2 -> t <> E0 -> forever_reactive ge s2 T ->
+ forever_reactive ge s1 (t *** T).
+
+Lemma star_forever_reactive:
+ forall ge s1 t s2 T,
+ star ge s1 t s2 -> forever_reactive ge s2 T ->
+ forever_reactive ge s1 (t *** T).
+Proof.
+ intros. inv H0. rewrite <- Eappinf_assoc. econstructor.
+ eapply star_trans; eauto.
+ red; intro. exploit Eapp_E0_inv; eauto. intros [P Q]. contradiction.
+ auto.
+Qed.
+
(** * Outcomes for program executions *)
-(** The three possible outcomes for the execution of a program:
+(** The four possible outcomes for the execution of a program:
- Termination, with a finite trace of observable events
and an integer value that stands for the process exit code
(the return value of the main function).
-- Divergence with an infinite trace of observable events.
- (The actual events generated by the execution can be a
- finite prefix of this trace, or the whole trace.)
+- Divergence with a finite trace of observable events.
+ (At some point, the program runs forever without doing any I/O.)
+- Reactive divergence with an infinite trace of observable events.
+ (The program performs infinitely many I/O operations separated
+ by finite amounts of internal computations.)
- Going wrong, with a finite trace of observable events
performed before the program gets stuck.
*)
Inductive program_behavior: Type :=
| Terminates: trace -> int -> program_behavior
- | Diverges: traceinf -> program_behavior
+ | Diverges: trace -> program_behavior
+ | Reacts: traceinf -> program_behavior
| Goes_wrong: trace -> program_behavior.
Definition not_wrong (beh: program_behavior) : Prop :=
match beh with
| Terminates _ _ => True
| Diverges _ => True
+ | Reacts _ => True
| Goes_wrong _ => False
end.
@@ -319,16 +398,23 @@ Inductive program_behaves (ge: genv): program_behavior -> Prop :=
star ge s t s' ->
final_state s' r ->
program_behaves ge (Terminates t r)
- | program_diverges: forall s T,
+ | program_diverges: forall s t s',
initial_state s ->
- forever ge s T ->
- program_behaves ge (Diverges T)
+ star ge s t s' -> forever_silent ge s' ->
+ program_behaves ge (Diverges t)
+ | program_reacts: forall s T,
+ initial_state s ->
+ forever_reactive ge s T ->
+ program_behaves ge (Reacts T)
| program_goes_wrong: forall s t s',
initial_state s ->
star ge s t s' ->
- (forall t1 s1, ~step ge s' t1 s1) ->
+ nostep ge s' ->
(forall r, ~final_state s' r) ->
- program_behaves ge (Goes_wrong t).
+ program_behaves ge (Goes_wrong t)
+ | program_goes_initially_wrong:
+ (forall s, ~initial_state s) ->
+ program_behaves ge (Goes_wrong E0).
End CLOSURES.
@@ -363,7 +449,6 @@ Variable ge2: genv2.
This matching relation must be compatible with initial states
and with final states. *)
-
Variable match_states: state1 -> state2 -> Prop.
Hypothesis match_initial_states:
@@ -414,23 +499,35 @@ Proof.
auto.
Qed.
-Lemma simulation_star_forever:
- forall st1 st2 T,
- forever step1 ge1 st1 T -> match_states st1 st2 ->
- forever step2 ge2 st2 T.
+Lemma simulation_star_forever_silent:
+ forall st1 st2,
+ forever_silent step1 ge1 st1 -> match_states st1 st2 ->
+ forever_silent step2 ge2 st2.
Proof.
- assert (forall st1 st2 T,
- forever step1 ge1 st1 T -> match_states st1 st2 ->
- forever_N step2 order ge2 st1 st2 T).
+ assert (forall st1 st2,
+ forever_silent step1 ge1 st1 -> match_states st1 st2 ->
+ forever_silent_N step2 order ge2 st1 st2).
cofix COINDHYP; intros.
inversion H; subst.
destruct (simulation H1 H0) as [st2' [A B]].
destruct A as [C | [C D]].
- apply forever_N_plus with t st2' s2 T0.
- auto. apply COINDHYP. assumption. assumption. auto.
- apply forever_N_star with t st2' s2 T0.
- auto. auto. apply COINDHYP. assumption. auto. auto.
- intros. eapply forever_N_forever; eauto.
+ apply forever_silent_N_plus with st2' s2.
+ auto. apply COINDHYP. assumption. assumption.
+ apply forever_silent_N_star with st2' s2.
+ auto. auto. apply COINDHYP. assumption. auto.
+ intros. eapply forever_silent_N_forever; eauto.
+Qed.
+
+Lemma simulation_star_forever_reactive:
+ forall st1 st2 T,
+ forever_reactive step1 ge1 st1 T -> match_states st1 st2 ->
+ forever_reactive step2 ge2 st2 T.
+Proof.
+ cofix COINDHYP; intros.
+ inv H.
+ destruct (simulation_star_star H1 H0) as [st2' [A B]].
+ econstructor. eexact A. auto.
+ eapply COINDHYP. eauto. auto.
Qed.
Lemma simulation_star_wf_preservation:
@@ -444,9 +541,13 @@ Proof.
destruct (simulation_star_star H2 B) as [s2' [C D]].
econstructor; eauto.
destruct (match_initial_states H1) as [s2 [A B]].
+ destruct (simulation_star_star H2 B) as [s2' [C D]].
econstructor; eauto.
- eapply simulation_star_forever; eauto.
- contradiction.
+ eapply simulation_star_forever_silent; eauto.
+ destruct (match_initial_states H1) as [s2 [A B]].
+ econstructor; eauto. eapply simulation_star_forever_reactive; eauto.
+ contradiction.
+ contradiction.
Qed.
End SIMULATION_STAR_WF.
@@ -579,4 +680,140 @@ End SIMULATION_OPT.
End SIMULATION.
+(** * Additional results about infinite reduction sequences *)
+
+(** We now show that any infinite sequence of reductions is either of
+ the "reactive" kind or of the "silent" kind (after a finite number
+ of non-silent transitions). The proof necessitates the axiom of
+ excluded middle. This result is used in [Csem] and [Cminor] to
+ relate the coinductive big-step semantics for divergence with the
+ small-step notions of divergence. *)
+
+Require Import Classical.
+Unset Implicit Arguments.
+
+Section INF_SEQ_DECOMP.
+
+Variable genv: Type.
+Variable state: Type.
+Variable step: genv -> state -> trace -> state -> Prop.
+
+Variable ge: genv.
+
+Inductive State: Type :=
+ ST: forall (s: state) (T: traceinf), forever step ge s T -> State.
+
+Definition state_of_State (S: State): state :=
+ match S with ST s T F => s end.
+Definition traceinf_of_State (S: State) : traceinf :=
+ match S with ST s T F => T end.
+
+Inductive Step: trace -> State -> State -> Prop :=
+ | Step_intro: forall s1 t T s2 S F,
+ Step t (ST s1 (t *** T) (@forever_intro genv state step ge s1 t s2 T S F))
+ (ST s2 T F).
+
+Inductive Steps: State -> State -> Prop :=
+ | Steps_refl: forall S, Steps S S
+ | Steps_left: forall t S1 S2 S3, Step t S1 S2 -> Steps S2 S3 -> Steps S1 S3.
+
+Remark Steps_trans:
+ forall S1 S2, Steps S1 S2 -> forall S3, Steps S2 S3 -> Steps S1 S3.
+Proof.
+ induction 1; intros. auto. econstructor; eauto.
+Qed.
+
+Let Reactive (S: State) : Prop :=
+ forall S1,
+ Steps S S1 ->
+ exists S2, exists S3, exists t, Steps S1 S2 /\ Step t S2 S3 /\ t <> E0.
+
+Let Silent (S: State) : Prop :=
+ forall S1 t S2, Steps S S1 -> Step t S1 S2 -> t = E0.
+
+Lemma Reactive_or_Silent:
+ forall S, Reactive S \/ (exists S', Steps S S' /\ Silent S').
+Proof.
+ intros. destruct (classic (exists S', Steps S S' /\ Silent S')).
+ auto.
+ left. red; intros.
+ generalize (not_ex_all_not _ _ H S1). intros.
+ destruct (not_and_or _ _ H1). contradiction.
+ unfold Silent in H2.
+ generalize (not_all_ex_not _ _ H2). intros [S2 A].
+ generalize (not_all_ex_not _ _ A). intros [t B].
+ generalize (not_all_ex_not _ _ B). intros [S3 C].
+ generalize (imply_to_and _ _ C). intros [D F].
+ generalize (imply_to_and _ _ F). intros [G J].
+ exists S2; exists S3; exists t. auto.
+Qed.
+
+Lemma Steps_star:
+ forall S1 S2, Steps S1 S2 ->
+ exists t, star step ge (state_of_State S1) t (state_of_State S2)
+ /\ traceinf_of_State S1 = t *** traceinf_of_State S2.
+Proof.
+ induction 1.
+ exists E0; split. apply star_refl. auto.
+ inv H. destruct IHSteps as [t' [A B]].
+ exists (t ** t'); split.
+ simpl; eapply star_left; eauto.
+ simpl in *. subst T. traceEq.
+Qed.
+
+Lemma Silent_forever_silent:
+ forall S,
+ Silent S -> forever_silent step ge (state_of_State S).
+Proof.
+ cofix COINDHYP; intro S. case S. intros until f. simpl. case f. intros.
+ assert (Step t (ST s1 (t *** T0) (forever_intro s1 t s0 f0))
+ (ST s2 T0 f0)).
+ constructor.
+ assert (t = E0).
+ red in H. eapply H; eauto. apply Steps_refl.
+ apply forever_silent_intro with (state_of_State (ST s2 T0 f0)).
+ rewrite <- H1. assumption.
+ apply COINDHYP.
+ red; intros. eapply H. eapply Steps_left; eauto. eauto.
+Qed.
+
+Lemma Reactive_forever_reactive:
+ forall S,
+ Reactive S -> forever_reactive step ge (state_of_State S) (traceinf_of_State S).
+Proof.
+ cofix COINDHYP; intros.
+ destruct (H S) as [S1 [S2 [t [A [B C]]]]]. apply Steps_refl.
+ destruct (Steps_star _ _ A) as [t' [P Q]].
+ inv B. simpl in *. rewrite Q. rewrite <- Eappinf_assoc.
+ apply forever_reactive_intro with s2.
+ eapply star_right; eauto.
+ red; intros. destruct (Eapp_E0_inv _ _ H0). contradiction.
+ change (forever_reactive step ge (state_of_State (ST s2 T F)) (traceinf_of_State (ST s2 T F))).
+ apply COINDHYP.
+ red; intros. apply H.
+ eapply Steps_trans. eauto.
+ eapply Steps_left. constructor. eauto.
+Qed.
+
+Theorem forever_silent_or_reactive:
+ forall s T,
+ forever step ge s T ->
+ forever_reactive step ge s T \/
+ exists t, exists s', exists T',
+ star step ge s t s' /\ forever_silent step ge s' /\ T = t *** T'.
+Proof.
+ intros.
+ destruct (Reactive_or_Silent (ST s T H)).
+ left.
+ change (forever_reactive step ge (state_of_State (ST s T H)) (traceinf_of_State (ST s T H))).
+ apply Reactive_forever_reactive. auto.
+ destruct H0 as [S' [A B]].
+ exploit Steps_star; eauto. intros [t [C D]]. simpl in *.
+ right. exists t; exists (state_of_State S'); exists (traceinf_of_State S').
+ split. auto.
+ split. apply Silent_forever_silent. auto.
+ auto.
+Qed.
+
+End INF_SEQ_DECOMP.