Library smallstep_guided

Require Export smallstep.

Set Implicit Arguments.
Unset Strict Implicit.

Section Smallstep.

Variable EVENT: Type.

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

Extracting an infinite "guided" execution from a diverging/reacting execution.
A guided execution records all the configurations during the execution.

Definition guided_execution := (CONFIG * Stream (list EVENT * CONFIG))%type.

It actually allows defining a deterministic transition relation restricting the transition relation to the given infinite execution.

Inductive guided_step : guided_execution -> list EVENT -> guided_execution -> Prop :=
| guided_step_intro s tr s' q :
  guided_step (s, Cons (tr, s') q) tr (s', q)
.

Lemma guided_step_determ: forall conf tr1 conf1 (Hconf1: guided_step conf tr1 conf1) tr2 conf2 (Hconf2: guided_step conf tr2 conf2), (tr1,conf1) = (tr2,conf2).
Proof.
  inversion 1; subst; inversion 1; subst; trivial.
Qed.

CoInductive guided_execution_correct : guided_execution -> Prop :=
| guided_execution_correct_intro s tr s'
  (Hstep: STEP s tr s')
  q (HCOIND: guided_execution_correct (s', q)):
  guided_execution_correct (s, Cons (tr, s') q).

Lemma guided_execution_correct_inv: forall s tr s' q
  (Hcorrect: guided_execution_correct (s, (Cons (tr, s') q))),
  STEP s tr s' /\ guided_execution_correct (s', q).
Proof.
  inversion 1; subst; auto.
Qed.

Lemma guided_step_inv: forall s q, guided_execution_correct (s,q) -> forall tr s' q',
  guided_step (s,q) tr (s',q') ->
  STEP s tr s'.
Proof.
  inversion 1; subst; inversion 1; subst; assumption.
Qed.

Lemma guided_step_inv_correct: forall s q, guided_execution_correct (s,q) -> forall tr s' q',
  guided_step (s,q) tr (s',q') ->
  guided_execution_correct (s',q').
Proof.
  inversion 1; subst; inversion 1; subst; assumption.
Qed.

Guided executions are deterministic


Lemma guided_stepn_inv: forall n s q tr s' q',
  stepn_left (guided_step) n (s,q) tr (s',q') ->
  guided_execution_correct (s,q) ->
  stepn_left STEP n s tr s'.
Proof.
  induction n; inversion 1; subst.
   intros; constructor.
  inversion Hstep; subst.
  inversion 1; subst.
  econstructor.
  eassumption.
  eapply IHn.
  eassumption.
  assumption.
  reflexivity.
Qed.

Lemma guided_stepn_inv_correct: forall n s q tr s' q',
  stepn_left guided_step n (s,q) tr (s',q') ->
  guided_execution_correct (s,q) ->
  guided_execution_correct (s',q').
Proof.
  induction n; inversion 1; subst.
   auto.
  inversion Hstep; subst.
  inversion 1; subst.
  eauto.
Qed.

Lemma guided_step_silently_diverges: forall s q,
  config_silently_diverges (guided_step) (s, q) ->
  guided_execution_correct (s, q) ->
  config_silently_diverges STEP s.
Proof.
  cofix COINDHYP.
  inversion 1; subst.
  intros.
  destruct conf'.
  econstructor.
  eapply guided_step_inv.
  eassumption.
  eassumption.
  eapply COINDHYP.
  eassumption.
  eapply guided_step_inv_correct.
  eassumption.
  eassumption.
Qed.

Lemma guided_step_diverges: forall s q tr
  (Hdiv: config_diverges (guided_step) (s,q) tr),
  guided_execution_correct (s,q) ->
  config_diverges STEP s tr.
Proof.
  inversion 1; subst.
  clear Hdiv.
  destruct (star_stepn_left Hstar).
  destruct conf'.
  econstructor.
  eapply stepn_left_star.
  eapply guided_stepn_inv.
  eassumption.
  assumption.
  eapply guided_step_silently_diverges.
  eassumption.
  eapply guided_stepn_inv_correct.
  eassumption.
  assumption.
Qed.

Lemma guided_step_reacts: forall s q itr,
  config_reacts (guided_step) (s,q) itr ->
  guided_execution_correct (s,q) ->
  config_reacts STEP s itr.
Proof.
  cofix COINDHYP.
  inversion 1; subst.
  destruct (star_stepn_left Hstep).
  destruct conf'.
  intros.
  econstructor.
  eapply stepn_left_star.
  eapply guided_stepn_inv.
  eassumption.
  assumption.
  eapply COINDHYP.
  eassumption.
  eapply guided_stepn_inv_correct.
  eassumption.
  assumption.
  trivial.
Qed.

Lemma guided_step_exists : forall s, exists tr, exists s',
  guided_step s tr s'.
Proof.
  destruct s.
  destruct s.
  destruct p.
  esplit.
  esplit.
  econstructor.
Qed.

Lemma guided_stepn_exists : forall n s,
  exists tr, exists s',
    stepn_left (guided_step) n s tr s'.
Proof.
  induction n; intros.
   esplit.
   esplit.
   eleft.
  destruct (guided_step_exists s).
  destruct H.
  destruct (IHn x0).
  destruct H0.
  eauto using stepn_left_S.
Qed.

Lemma guided_behavior_inv : forall ret s b,
  config_behaviors (guided_step) (fun _ _ => False) s b ->
  (exists l, b = Diverges _ l) \/ (exists l, b = Reacts ret l).
Proof.
  destruct b; eauto; inversion 1; exfalso; subst.
   assumption.
  destruct (guided_step_exists cstuck).
  destruct H0.
  eauto.
Qed.

Building a guided execution from a diverging/reacting execution

Require Export IndefiniteDescription.

Lemma config_silently_diverges_ex: forall conf (Hsildiv: config_silently_diverges STEP conf), {conf' | STEP conf nil conf' /\ config_silently_diverges STEP conf'}.
Proof.
  intros.
  apply constructive_indefinite_description.
  inversion Hsildiv; subst.
  eauto.
Qed.

CoFixpoint config_silently_diverges_guided_execution_aux (conf: CONFIG) (Hsildiv: config_silently_diverges STEP conf): Stream (list EVENT * CONFIG) :=
  match config_silently_diverges_ex Hsildiv with
    | exist conf' (conj Hstep HCOIND) => Cons (nil, conf') (config_silently_diverges_guided_execution_aux HCOIND)
  end.

Lemma config_silently_diverges_guided_execution_aux_correct: forall conf (Hsildiv: config_silently_diverges STEP conf), guided_execution_correct (conf, config_silently_diverges_guided_execution_aux Hsildiv).
Proof.
  cofix COINDHYP.
  intros.
  rewrite (unfold_Stream (config_silently_diverges_guided_execution_aux Hsildiv)).
  unfold config_silently_diverges_guided_execution_aux.
  fold config_silently_diverges_guided_execution_aux.
  destruct (config_silently_diverges_ex Hsildiv).
  destruct a.
  econstructor.
  assumption.
  eapply COINDHYP.
Qed.

Lemma config_silently_diverges_guided_execution_aux_silently_diverges: forall conf (Hsildiv: config_silently_diverges STEP conf), config_silently_diverges (guided_step) (conf, config_silently_diverges_guided_execution_aux Hsildiv).
Proof.
  cofix COINDHYP.
  intros.
  rewrite (unfold_Stream (config_silently_diverges_guided_execution_aux Hsildiv)).
  unfold config_silently_diverges_guided_execution_aux.
  fold config_silently_diverges_guided_execution_aux.
  destruct (config_silently_diverges_ex Hsildiv).
  destruct a.
  econstructor.
  econstructor.
  eapply COINDHYP.
Qed.

Lemma config_diverges_guided_execution: forall conf tr (Hdiv: config_diverges STEP conf tr), exists st, guided_execution_correct (conf, st) /\ config_diverges (guided_step) (conf, st) tr.
Proof.
  inversion 1; subst; clear Hdiv.
  induction Hstar using (star_ind3_left).
   esplit.
    split.
    eapply config_silently_diverges_guided_execution_aux_correct with (Hsildiv := Hsilently_diverges).
    econstructor. eleft. eapply config_silently_diverges_guided_execution_aux_silently_diverges.
   destruct (IHHstar Hsilently_diverges).
   destruct H.
   inversion H0; subst.
   esplit.
   split.
   econstructor.
   eassumption.
   eassumption.
   econstructor.
   eapply star_trans.
   eright.
   eleft.
   econstructor.
   eassumption.
   assumption.
Qed.

Lemma config_reacts_N_min_ex: forall conf n st
(Hreacts: config_reacts_N_min STEP conf n st), {
  conf' : _ & {
    n' |
      STEP conf nil conf' /\
      config_reacts_N_min STEP conf' n' st /\
      n = S n'
  }
} + {
  tr : _ & {
    conf' : _ & {
      st' : _ & {
        n' |
          STEP conf tr conf' /\
          config_reacts_N_min STEP conf' n' st' /\
          tr <> nil /\
          st = appinf tr st' /\
          n = O
      }
    }
  }
}.
Proof.
  destruct n; intros.
   assert (
     exists tr, exists conf', exists st', exists n', STEP conf tr conf' /\ config_reacts_N_min STEP conf' n' st' /\ tr <> nil /\ st = appinf tr st' /\ O = O
   ).
    inversion Hreacts; subst; eauto 9.
   destruct (constructive_indefinite_description _ H); clear H.
   destruct (constructive_indefinite_description _ e); clear e.
   destruct (constructive_indefinite_description _ e0); clear e0.
   destruct (constructive_indefinite_description _ e); clear e.
   eauto 6.
  assert (
    exists conf', STEP conf nil conf' /\ config_reacts_N_min STEP conf' n st /\ S n = S n
  ).
   inversion Hreacts; subst.
   eauto.
  destruct (constructive_indefinite_description _ H); clear H.
  eauto.
Qed.

CoFixpoint config_reacts_N_min_guided_execution_aux (n: nat) (conf: CONFIG) (st: Stream EVENT) (Hreacts: config_reacts_N_min STEP conf n st): Stream (list EVENT * CONFIG) :=
  match config_reacts_N_min_ex Hreacts with
    | inl (existT conf' (exist n' (conj _ (conj CIHreacts _)))) =>
      Cons (nil, conf') (config_reacts_N_min_guided_execution_aux CIHreacts)
    | inr (existT tr (existT conf' (existT st' (exist n' (conj _ (conj CIHreacts _)))))) =>
      Cons (tr, conf') (config_reacts_N_min_guided_execution_aux CIHreacts)
  end.

Lemma config_reacts_N_min_guided_execution_aux_correct: forall n conf st (Hreacts: config_reacts_N_min STEP conf n st),
  guided_execution_correct (conf, config_reacts_N_min_guided_execution_aux Hreacts).
Proof.
  cofix COINDHYP.
  intros.
  rewrite (unfold_Stream (config_reacts_N_min_guided_execution_aux Hreacts)).
  unfold config_reacts_N_min_guided_execution_aux.
  fold config_reacts_N_min_guided_execution_aux.
  destruct (config_reacts_N_min_ex Hreacts).
  destruct s.
  destruct s.
  destruct a.
  destruct a.
  subst.
  econstructor.
  assumption.
  eapply COINDHYP.
  destruct s.
  destruct s.
  destruct s.
  destruct s.
  destruct a.
  destruct a.
  econstructor.
  assumption.
  eapply COINDHYP.
Qed.

Lemma config_reacts_N_min_guided_execution_aux_reacts_N_min: forall n conf st (Hreacts: config_reacts_N_min STEP conf n st),
  config_reacts_N_min (guided_step) (conf, config_reacts_N_min_guided_execution_aux Hreacts) n st.
Proof.
  cofix COINDHYP.
  intros.
  rewrite (unfold_Stream (config_reacts_N_min_guided_execution_aux Hreacts)).
  unfold config_reacts_N_min_guided_execution_aux.
  fold config_reacts_N_min_guided_execution_aux.
  destruct (config_reacts_N_min_ex Hreacts).
  destruct s.
  destruct s.
  destruct a.
  destruct a.
  subst.
  econstructor.
  econstructor.
  eapply COINDHYP.
  destruct s.
  destruct s.
  destruct s.
  destruct s.
  destruct a.
  destruct a.
  destruct a.
  destruct H0.
  subst.
  econstructor.
  eassumption.
  econstructor.
  eapply COINDHYP.
  reflexivity.
Qed.

Corollary config_reacts_guided_step : forall conf st (Hreacts: config_reacts STEP conf st), exists t, guided_execution_correct (conf, t) /\ config_reacts (guided_step) (conf, t) st.
Proof.
  intros.
  destruct (config_reacts_config_reacts_N_min Hreacts).
  esplit.
  split.
  eapply config_reacts_N_min_guided_execution_aux_correct with (Hreacts := H).
  eapply config_reacts_N_star_lt_config_reacts.
  eapply config_reacts_N_star_lt'_lt.
  eapply config_reacts_N_star_config_reacts_N_star_lt'.
  eapply config_reacts_N_config_reacts_N_star.
  eapply config_reacts_N_min_config_reacts_N.
  eapply config_reacts_N_min_guided_execution_aux_reacts_N_min.
Qed.

Simulation between a guided execution of some transition relation and another transition relation

Section Simult_guided.

Variables (CONFIG': Type) (STEP': CONFIG' -> list EVENT -> CONFIG' -> Prop)

 (I: guided_execution -> CONFIG' -> Prop) (Hsimult: simult (guided_step) STEP' I).

Lemma simult_guided_behaviors: forall (ret: Type) s b,
  config_behaviors (final := ret) (guided_step) (fun _ _ => False) s b ->
  forall s', I s s' ->
    forall config_final',
      config_behaviors STEP' config_final' s' b.
Proof.
  intros.
  destruct (guided_behavior_inv H).
   destruct H1.
   subst; simpl in *.
   eapply config_diverges_simult.
   2: eassumption.
   eapply guided_step_determ.
   eassumption.
   assumption.
  destruct H1.
  subst; simpl in *.
  eapply config_reacts_simult.
  2: eassumption.
  eapply guided_step_determ.
  eassumption.
  assumption.
Qed.

End Simult_guided.

End Smallstep.