Library smallstep_exists_beh

Require Export smallstep.

Set Implicit Arguments.
Unset Strict Implicit.

CPP 2015 submission: Lemma 1. Existence of behaviors.

This proof is taken from CompCert, module Behaviors.

Section Smallstep.

Variable EVENT: Type.

Variable final: Type.

Section Smallstep'.

Variable CONFIG: Type.

Variable config_final : CONFIG -> final -> Prop.

Variable (STEP: CONFIG -> list EVENT -> CONFIG -> Prop).

Section Reacts_forever_reactive.

Variable reacts_forever_reactive
     : forall
         (s0 : CONFIG),
       (forall (s1 : CONFIG) (t1 : list EVENT),
        star STEP s0 t1 s1 ->
        exists s2 : CONFIG,
          exists t2 : list EVENT, star STEP s1 t2 s2 /\ t2 <> nil) ->
       exists T : Stream EVENT, config_reacts STEP s0 T
         .

Require Export classic.

Theorem exists_behavior_weak:
  forall s, exists beh, config_behaviors STEP config_final s beh.
Proof.
  intros s0.
  destruct (classic (forall s1 t1, star STEP s0 t1 s1 -> exists s2, exists t2, STEP s1 t2 s2)).
  destruct (classic (exists s1, exists t1, star STEP s0 t1 s1 /\
                       (forall s2 t2, star STEP s1 t2 s2 ->
                        exists s3, STEP s2 nil s3))).
  destruct H0 as [s1 [t1 [A B] ] ].
  exists (Diverges _ t1); econstructor; eauto.
  clear s0 H t1 A.
  revert s1 B.
  cofix COINDHYP.
   intros.
   edestruct B.
   eleft.
   econstructor.
   eassumption.
   eapply COINDHYP.
   intros.
   eapply B.
   eapply star_trans.
   eright.
   eleft.
   eassumption.
   eassumption.
  destruct (@reacts_forever_reactive s0) as [T FR].
  intros.
  generalize (not_ex_all_not _ _ H0 s1). intro A; clear H0.
  generalize (not_ex_all_not _ _ A t1). intro B; clear A.
  destruct (not_and_or _ _ B). contradiction.
  destruct (not_all_ex_not _ _ H0) as [s2 C]; clear H0.
  destruct (not_all_ex_not _ _ C) as [t2 D]; clear C.
  destruct (imply_to_and _ _ D) as [E F]; clear D.
  destruct (H s2 (t1 ++ t2)) as [s3 [t3 G] ]. eapply star_trans; eauto.
  exists s3; exists (t2 ++ t3); split.
  eapply star_trans. eassumption. eright. eleft. assumption.
  red; intros. destruct (app_eq_nil t2 t3 H0). subst. elim F. exists s3; auto.
  exists (Reacts _ (func_of_stream T)).
   simpl.
   eapply config_reacts_bisim.
   eassumption.
   apply stream_of_func_of_stream.
  destruct (not_all_ex_not _ _ H) as [s1 A]; clear H.
  destruct (not_all_ex_not _ _ A) as [t1 B]; clear A.
  destruct (imply_to_and _ _ B) as [C D]; clear B.
  destruct (classic (exists r, config_final s1 r)) as [ [r FINAL] | NOTFINAL].
  exists (Terminates t1 r); econstructor; eauto.
  exists (Stuck _ t1); econstructor; eauto.
Qed.

End Reacts_forever_reactive.

Require Export IndefiniteDescription.

Section Reacts.

Variable s0: CONFIG.

Hypothesis reacts:
  forall s1 t1, star STEP s0 t1 s1 ->
  exists s2, exists t2, star STEP s1 t2 s2 /\ t2 <> nil.

Lemma reacts':
  forall s1 t1, star STEP s0 t1 s1 ->
  { s2 : CONFIG & { t2 : list EVENT | star STEP s1 t2 s2 /\ t2 <> nil } }.
Proof.
  intros.
  destruct (constructive_indefinite_description _ (reacts H)) as [s2 A].
  destruct (constructive_indefinite_description _ A) as [t2 [B C]].
  exists s2; exists t2; auto.
Qed.

CoFixpoint build_traceinf' (s1: CONFIG) (t1: list EVENT) (ST: star STEP s0 t1 s1) : traceinf' EVENT :=
  match reacts' ST with
  | existT s2 (exist t2 (conj A B)) =>
      Econsinf' t2
                (build_traceinf' (star_trans ST A))
                B
  end.

Lemma reacts_forever_reactive_rec:
  forall s1 t1 (ST: star STEP s0 t1 s1),
  config_reacts STEP s1 (traceinf_of_traceinf' (build_traceinf' ST)).
Proof.
  cofix COINDHYP; intros.
  rewrite (unroll_traceinf' (build_traceinf' ST)). simpl.
  destruct (reacts' ST) as [s2 [t2 [A B]]].
  rewrite traceinf_traceinf'_app.
  destruct t2.
   congruence.
  econstructor. eexact A. auto. f_equal.
Qed.

Lemma reacts_forever_reactive:
  exists T, config_reacts STEP s0 T.
Proof.
  esplit.
  eapply reacts_forever_reactive_rec with (ST := star_refl _ _).
Qed.

End Reacts.

Theorem exists_behavior:
  forall s, exists beh, config_behaviors STEP config_final s beh.
Proof.
  eapply exists_behavior_weak.
  eapply reacts_forever_reactive.
Qed.

End Smallstep'.

End Smallstep.