Library unify

Unification between our behaviors and CompCert's.
Require Import Smallstep Events.
Require Behaviors smallstep.
Set Implicit Arguments.
Unset Strict Implicit.

Section IO.

  Variables input output: Type.

Section Forward_with_stuck.

  Variables S C : semantics input output.
  Variable fs: forward_simulation S C.

  Hypothesis s_c_stuck:
  forall s i c (MS: fs i s c),
    Nostep S s ->
    (forall r, ~ final_state S s r) ->
    exists c2,
      Star C c E0 c2 /\ Nostep C c2 /\ forall r, ~ final_state C c2 r.

  Lemma forward_with_stuck_state_behaves : forall i s c
    (Hmatch: fs i s c)
    beh
    (Hbeh: Behaviors.state_behaves S s beh), Behaviors.state_behaves C c beh.
  Proof.
    intros.
    generalize (Behaviors.forward_simulation_state_behaves fs).
    intro H.
    destruct (H _ _ _ _ Hmatch Hbeh).
    destruct H0.
    destruct H1.
     subst; eauto.
     destruct H1.
   destruct H1.
   clear x H0 H2.
   subst.
   inversion Hbeh; subst.
   destruct (Smallstep.simulation_star fs H1 _ _ Hmatch).
   destruct H0.
   destruct H0.
   destruct (s_c_stuck H4 H2 H3).
   destruct H5.
   destruct H6.
    econstructor.
    eapply Smallstep.star_trans.
    eassumption.
    eassumption.
    rewrite E0_right. reflexivity.
    assumption.
    assumption.
Qed.

End Forward_with_stuck.

CompCert behaviors use bisimilarity for infinite execution traces. So we need to lift bisimilarity to CompCert behaviors.

Definition beh_bisim (b: Behaviors.program_behavior output) : Behaviors.program_behavior output -> Prop :=
  match b with
    | Behaviors.Reacts tr => fun b' =>
      match b' with
        | Behaviors.Reacts tr' => traceinf_sim tr tr'
        | _ => False
      end
    | _ => eq b
  end.

Lemma beh_bisim_refl: forall b, beh_bisim b b.
Proof.
  destruct b; simpl; eauto using traceinf_sim_refl.
Qed.

Lemma beh_bisim_sym: forall b1 b2, beh_bisim b1 b2 -> beh_bisim b2 b1.
Proof.
  destruct b1; destruct b2; simpl; try discriminate; eauto using traceinf_sim_sym; tauto.
Qed.

Lemma beh_bisim_trans: forall b1 b2, beh_bisim b1 b2 -> forall b3, beh_bisim b2 b3 -> beh_bisim b1 b3.
Proof.
  destruct b1; destruct b2; simpl; destruct b3; simpl; try discriminate; intros; eauto using traceinf_sim_trans, eq_trans.
Qed.

Lemma traceinf_sim_sim': forall T1 T2,
  traceinf_sim T1 T2 ->
  traceinf_sim' T1 T2.
Proof.
  cofix COINDHYP.
  inversion 1; subst.
  eapply traceinf_sim'_cons with (t := cons e nil) (T1 := T0) (T2 := T3).
  clear COINDHYP; discriminate.
  eapply COINDHYP; clear COINDHYP; assumption.
Qed.

Lemma traceinf_sim_decomp: forall t T1 T2,
  traceinf_sim (Eappinf t T1) T2 ->
  exists T'2, T2 = Eappinf t T'2 /\ traceinf_sim T1 T'2.
Proof.
  induction t; simpl.
   eauto.
  inversion 1; subst.
  destruct (IHt _ _ H3).
  destruct H0.
  subst; eauto.
Qed.

CoFixpoint traceinf_of_stream (s: Streams.Stream event): traceinf :=
  match s with
    | Streams.Cons a s' => Econsinf a (traceinf_of_stream s')
  end.

Lemma traceinf_of_stream_appinf : forall l s,
  traceinf_of_stream (stream.appinf l s) = Eappinf l (traceinf_of_stream s).
Proof.
  induction l; simpl.
   trivial.
  intros.
  rewrite (unroll_traceinf (traceinf_of_stream (Streams.Cons a (stream.appinf l s)))).
  simpl.
  f_equal.
  auto.
Qed.

CoFixpoint stream_of_traceinf (t: traceinf): Streams.Stream event :=
  match t with
    | Econsinf a t' => Streams.Cons a (stream_of_traceinf t')
  end.

Lemma stream_of_traceinf_appinf : forall l s,
  stream_of_traceinf (Eappinf l s) = stream.appinf l (stream_of_traceinf s).
Proof.
  induction l; simpl.
   trivial.
  intros.
  rewrite (Streams.unfold_Stream (stream_of_traceinf (Econsinf a (l *** s)))).
  simpl.
  f_equal.
  auto.
Qed.

Lemma traceinf_of_stream_of_traceinf: forall t u,
  traceinf_sim u (traceinf_of_stream (stream_of_traceinf t)) ->
  traceinf_sim t u.
Proof.
  cofix COINDHYP.
  destruct t.
  intros.
  rewrite (Streams.unfold_Stream (stream_of_traceinf (Econsinf e t))) in H.
  simpl in H.
  rewrite (unroll_traceinf (traceinf_of_stream (Streams.Cons e (stream_of_traceinf t)))) in H.
  simpl in H.
  inversion H; subst.
  econstructor.
  eapply COINDHYP.
  assumption.
Qed.

Lemma stream_of_traceinf_of_stream: forall s r,
  stream.bisim r (stream_of_traceinf (traceinf_of_stream s)) ->
  stream.bisim s r.
Proof.
  cofix COINDHYP.
  destruct s.
  intros.
  rewrite (unroll_traceinf (traceinf_of_stream (Streams.Cons e s))) in H.
  simpl in H.
  rewrite (Streams.unfold_Stream (stream_of_traceinf (Econsinf e (traceinf_of_stream s)))) in H.
  simpl in H.
  inversion H; subst.
  constructor.
  apply COINDHYP.
  assumption.
Qed.

Lemma traceinf_of_stream_bisim: forall s1 s2,
  stream.bisim s1 s2 ->
  forall t1, traceinf_sim t1 (traceinf_of_stream s1) ->
  forall t2, traceinf_sim t2 (traceinf_of_stream s2) ->
  traceinf_sim t1 t2.
Proof.
  cofix COINDHYP.
  inversion 1; subst.
  intros.
  rewrite (unroll_traceinf (traceinf_of_stream (Streams.Cons e l1))) in H0.
  rewrite (unroll_traceinf (traceinf_of_stream (Streams.Cons e l2))) in H1.
  simpl in *.
  inversion H0; subst.
  inversion H1; subst.
  constructor.
  eapply COINDHYP.
  eexact CIHbisim.
  assumption.
  assumption.
Qed.

Lemma stream_of_traceinf_bisim: forall t1 t2,
  traceinf_sim t1 t2 ->
  forall s1, stream.bisim s1 (stream_of_traceinf t1) ->
  forall s2, stream.bisim s2 (stream_of_traceinf t2) ->
  stream.bisim s1 s2.
Proof.
  cofix COINDHYP.
  inversion 1; subst.
  intros.
  rewrite (Streams.unfold_Stream (stream_of_traceinf (Econsinf e T1))) in H1.
  rewrite (Streams.unfold_Stream (stream_of_traceinf (Econsinf e T2))) in H2.
  simpl in *.
  inversion H1; subst.
  inversion H2; subst.
  constructor.
  eapply COINDHYP.
  eexact H0.
  assumption.
  assumption.
Qed.

Function beh_of_compcert_beh (beh: Behaviors.program_behavior output) : smallstep.behavior event output :=
  match beh with
    | Behaviors.Terminates tr res => smallstep.Terminates tr res
    | Behaviors.Goes_wrong tr => smallstep.Stuck _ tr
    | Behaviors.Diverges tr => smallstep.Diverges _ tr
    | Behaviors.Reacts trinf => smallstep.Reacts _ (stream.func_of_stream (stream_of_traceinf trinf))
  end.

Lemma beh_of_compcert_beh_bisim: forall beh1 beh2,
  beh_bisim beh1 beh2 ->
  beh_of_compcert_beh beh1 = beh_of_compcert_beh beh2.
Proof.
  destruct beh1; simpl; intros; subst; auto.
  destruct beh2; try contradiction; simpl.
  f_equal.
  apply stream.bisim_func_eq.
  eapply stream_of_traceinf_bisim.
  eassumption.
  apply stream.bisim_refl.
  apply stream.bisim_refl.
Qed.

Lemma bisim_beh_of_compcert_beh: forall beh1 beh2,
  beh_of_compcert_beh beh1 = beh_of_compcert_beh beh2 ->
  beh_bisim beh1 beh2.
Proof.
  destruct beh2; simpl; functional inversion 1; subst; eauto using beh_bisim_refl; simpl.
  generalize (stream.func_eq_bisim H0).
  intro.
  generalize (traceinf_of_stream_bisim H1 (Events.traceinf_sim_refl _) (Events.traceinf_sim_refl _)).
  intro.
  generalize (traceinf_of_stream_of_traceinf H2).
  intro.
  apply Events.traceinf_sim_sym in H2.
  generalize (traceinf_of_stream_of_traceinf H2).
  intro.
  apply Events.traceinf_sim_sym in H3.
  eauto using Events.traceinf_sim_trans.
Qed.

Function compcert_beh_of_beh (beh: smallstep.behavior event output) : Behaviors.program_behavior output :=
  match beh with
    | smallstep.Terminates tr res => Behaviors.Terminates tr res
    | smallstep.Stuck tr => Behaviors.Goes_wrong _ tr
    | smallstep.Diverges tr => Behaviors.Diverges _ tr
    | smallstep.Reacts trinf => Behaviors.Reacts _ (traceinf_of_stream (stream.stream_of_func trinf))
  end.

Lemma beh_of_compcert_of_beh : forall beh,
  beh_of_compcert_beh (compcert_beh_of_beh beh) = beh.
Proof.
  destruct beh; trivial; simpl.
  f_equal.
  rewrite <- (stream.func_of_stream_of_func evinfseq) at -1.
  apply stream.bisim_func_eq.
  apply stream.bisim_sym.
  eapply stream_of_traceinf_of_stream.
  apply stream.bisim_refl.
Qed.

Lemma compcert_of_beh_of_compcert : forall beh,
  beh_bisim beh (compcert_beh_of_beh (beh_of_compcert_beh beh)).
Proof.
  destruct beh; simpl; try discriminate; eauto.
  intros.
  eapply traceinf_of_stream_of_traceinf.
  eapply traceinf_of_stream_bisim.
  2: eapply traceinf_sim_refl.
  2: eapply traceinf_sim_refl.
  apply stream.bisim_sym.
  apply stream.stream_of_func_of_stream.
Qed.

Lemma beh_of_compcert_beh_app : forall l beh,
    beh_of_compcert_beh (Behaviors.behavior_app l beh) =
    smallstep.append1_behavior l (beh_of_compcert_beh beh).
Proof.
  rewrite smallstep.append_behavior.
  destruct beh; simpl; try reflexivity.
  f_equal.
  apply stream.bisim_func_eq.
  rewrite stream_of_traceinf_appinf.
  apply stream.bisim_appinf.
  apply stream.stream_of_func_of_stream.
Qed.

Lemma compcert_beh_of_beh_app : forall l beh,
  beh_bisim (compcert_beh_of_beh (smallstep.append1_behavior l beh)) (Behaviors.behavior_app l (compcert_beh_of_beh beh)).
Proof.
  rewrite smallstep.append_behavior.
  destruct beh; simpl; try reflexivity.
  rewrite <- traceinf_of_stream_appinf.
  eapply traceinf_of_stream_bisim.
  2: eapply traceinf_sim_refl.
  2: eapply traceinf_sim_refl.
  apply stream.bisim_sym.
  apply stream.stream_of_func_of_stream.
Qed.

Require ccimproves.

Lemma improves_of_compcert_improves : forall beh1 beh2,
  Behaviors.behavior_improves beh1 beh2 ->
  ccimproves.behavior_improves (beh_of_compcert_beh beh2) (beh_of_compcert_beh beh1).
Proof.
  destruct 1; subst.
   left. reflexivity.
  destruct H.
  destruct H.
  subst.
  simpl in *.
  eright.
  esplit.
  split.
  reflexivity.
  destruct H0.
  subst.
  esplit.
  eapply beh_of_compcert_beh_app.
Qed.


Definition compcert_refinement (B1 B2: Behaviors.program_behavior output -> Prop) :=
  forall beh1, B1 beh1 ->
    exists beh2, B2 beh2 /\ Behaviors.behavior_improves beh2 beh1.

Lemma improves_of_compcert_improves_set: forall B1 B2,
  compcert_refinement B1 B2 ->
  forall B'1, (forall beh, B1 beh <-> B'1 (beh_of_compcert_beh beh)) ->
  forall B'2, (forall beh, B2 beh <-> B'2 (beh_of_compcert_beh beh)) ->
    ccimproves.compcert_refinement B'1 B'2.
Proof.
  unfold compcert_refinement.
  intros.
  intro.
  intros.
  rewrite <- beh_of_compcert_of_beh in H2.
  generalize (let (_, J) := H0 _ in J H2).
  intro.
  destruct (H _ H3).
  destruct H4.
  generalize (improves_of_compcert_improves H5).
  rewrite beh_of_compcert_of_beh.
  intros.
  generalize (let (K, _) := H1 _ in K H4).
  eauto.
Qed.

Lemma beh_of_compcert_beh_app_recip: forall beh l beh',
  beh_of_compcert_beh beh = smallstep.append1_behavior l beh' ->
  exists beh'', beh = Behaviors.behavior_app l beh'' /\ beh_of_compcert_beh beh'' = beh'.
Proof.
  intros until beh'.
  rewrite smallstep.append_behavior.
  destruct beh; destruct beh'; try discriminate.
  simpl.
  injection 1; intros; subst.
  eexists (Behaviors.Terminates _ _).
  simpl.
  split.
  reflexivity.
  reflexivity.
  simpl.
  injection 1; intros; subst.
  eexists (Behaviors.Diverges _ _).
  simpl.
  split.
  reflexivity.
  reflexivity.
  revert t evinfseq.
  rewrite <- smallstep.append_behavior.
  induction l; simpl.
   injection 1; intros; subst.
   eexists (Behaviors.Reacts _ _).
   split.
   simpl.
   reflexivity.
   reflexivity.
  rewrite smallstep.append_behavior.
  simpl.
  injection 1; intros; subst.
  destruct t.
  rewrite (Streams.unfold_Stream (stream_of_traceinf (Econsinf e t))) in H0.
  simpl in H0.
  generalize (stream.func_eq_bisim H0).
  inversion 1; subst.
  assert (beh_of_compcert_beh (Behaviors.Reacts output t) = smallstep.append1_behavior l (smallstep.Reacts output evinfseq)).
   rewrite smallstep.append_behavior.
   simpl.
   f_equal.
   apply stream.bisim_func_eq.
   eapply stream.bisim_trans.
   eassumption.
   apply stream.bisim_sym.
   apply stream.stream_of_func_of_stream.
  destruct (IHl _ _ H2).
  destruct H3.
  destruct x; try discriminate.
  exists (Behaviors.Reacts _ t0).
   split.
   simpl.
   simpl in H3.
   congruence.
   assumption.
  simpl.
  injection 1; intros; subst.
  eexists (Behaviors.Goes_wrong _ _).
  simpl.
  split.
  reflexivity.
  reflexivity.
Qed.

Lemma compcert_improves_of_improves_set: forall B'1 B'2,
  ccimproves.compcert_refinement B'1 B'2 ->
  forall B1 B2,
    (forall beh, B1 beh <-> B'1 (beh_of_compcert_beh beh)) ->
    (forall beh, B2 beh <-> B'2 (beh_of_compcert_beh beh)) ->
    compcert_refinement B1 B2.
  Proof.
    unfold ccimproves.compcert_refinement.
    intros.
    intro.
    intros.
    generalize (let (J, _) := H0 _ in J H2).
    intro.
    destruct (H _ H3).
    destruct H4.
    destruct H5.
     subst.
     generalize (let (_, K) := H1 _ in K H4).
     intros.
     esplit.
     split.
     eassumption.
     left.
     reflexivity.
    destruct H5.
    destruct H5.
    destruct H6.
    subst.
    destruct (beh_of_compcert_beh_app_recip H6).
    destruct H5.
    subst.
    generalize (let (_, K) := H1 (Behaviors.Goes_wrong _ _) in K H4).
    intro.
    esplit.
    split.
    eassumption.
    right.
    esplit.
    split.
    reflexivity.
    esplit.
    reflexivity.
Qed.

Section Forever_bisim.

Variables genv state: Type.
Variable step: genv -> state -> trace -> state -> Prop.
Variable ge: genv.

Lemma forever_bisim: forall s T,
  forever_reactive step ge s T ->
  forall T', traceinf_sim T T' ->
    forever_reactive step ge s T'.
Proof.
  cofix COINDHYP.
  inversion 1; subst.
  intros.
  destruct (traceinf_sim_decomp H3).
  destruct H4.
  subst.
  econstructor.
  eassumption.
  assumption.
  eapply COINDHYP.
  eassumption.
  eassumption.
Qed.

End Forever_bisim.

Section State_behaves.

Variable S: semantics input output.

Lemma state_behaves_bisim: forall s beh,
  Behaviors.state_behaves S s beh ->
  forall beh', beh_bisim beh beh' ->
    Behaviors.state_behaves S s beh'.
Proof.
  inversion 1; subst; simpl; intros; subst; trivial.
  destruct beh'; try contradiction.
  econstructor.
  eapply forever_bisim.
  eassumption.
  assumption.
Qed.

Lemma program_behaves_bisim: forall inp beh,
  Behaviors.program_behaves S inp beh ->
  forall beh', beh_bisim beh beh' ->
    Behaviors.program_behaves S inp beh'.
Proof.
  inversion 1; subst.
   intros; eleft; eauto using state_behaves_bisim.
  simpl; intros; subst; eauto.
Qed.

Lemma star_star:
  smallstep.star (Step S) = Star S.
Proof.
  apply FunctionalExtensionality.functional_extensionality.
  intros.
  apply FunctionalExtensionality.functional_extensionality.
  intros.
  apply Ensembles.Extensionality_Ensembles.
  split; unfold Ensembles.Included, Ensembles.In; revert x x0.
   induction 1 using smallstep.star_ind3_left.
    eleft.
   eright.
   eassumption.
   eassumption.
   reflexivity.
  induction 1.
   eleft.
  subst.
  eapply smallstep.star_trans.
  eright.
  eleft.
  eassumption.
  assumption.
Qed.

Lemma silent_silent:
  smallstep.config_silently_diverges (Step S) =
  Forever_silent S.
Proof.
  apply Ensembles.Extensionality_Ensembles.
  split; unfold Ensembles.Included, Ensembles.In.
  cofix COINDHYP.
  intros.
  inversion H; subst.
  econstructor.
  eassumption.
  eapply COINDHYP.
  assumption.
  cofix COINDHYP.
  intros.
  inversion H; subst.
  econstructor.
  eassumption.
  eapply COINDHYP.
  assumption.
Qed.

Lemma reacts_forever_reactive : forall s t,
  smallstep.config_reacts (Step S) s t ->
  forall u, traceinf_sim u (traceinf_of_stream t) ->
  Forever_reactive S s u.
Proof.
  cofix COINDHYP.
  inversion 1; subst.
  intros.
  rewrite (traceinf_of_stream_appinf) in H0.
  apply traceinf_sim_sym in H0.
  destruct (traceinf_sim_decomp H0).
  destruct H1.
  subst.
  rewrite star_star in Hstep.
  econstructor.
  eassumption.
  discriminate.
  eapply COINDHYP.
  eassumption.
  apply traceinf_sim_sym.
  assumption.
Qed.

Lemma forever_reactive_reacts: forall s u,
  Forever_reactive S s u ->
  forall t, stream.bisim t (stream_of_traceinf u) ->
  smallstep.config_reacts (Step S) s t.
Proof.
  cofix COINDHYP.
  inversion 1; subst.
  destruct t.
   destruct H1. trivial.
  intros.
  rewrite (stream_of_traceinf_appinf) in H3.
  apply stream.bisim_sym in H3.
  destruct (stream.bisim_appinf_strong H3).
  subst.
  generalize (stream.bisim_sym (stream.appinf_bisim H3)).
  intro.
  rewrite <- star_star in H0.
  econstructor.
  eassumption.
  eapply COINDHYP.
  eassumption.
  eassumption.
  reflexivity.
Qed.

Lemma config_behaviors_state_behaves: forall s beh,
  Behaviors.state_behaves S s beh ->
  smallstep.config_behaviors (Step S) (final_state S) s (beh_of_compcert_beh beh).
Proof.
  inversion 1; subst; simpl.
   rewrite <- star_star in H0. econstructor. eassumption. assumption.
   rewrite <- star_star in H0. econstructor. eassumption. rewrite silent_silent. assumption.
   eapply forever_reactive_reacts.
   2: eapply stream.bisim_sym; eapply stream.stream_of_func_of_stream.
   assumption.
   rewrite <- star_star in H0. econstructor. eassumption. assumption. assumption.
Qed.

Lemma state_behaves_config_behaviors: forall s beh,
  smallstep.config_behaviors (Step S) (final_state S) s beh ->
  Behaviors.state_behaves S s (compcert_beh_of_beh beh).
Proof.
  destruct beh; simpl.
   inversion 1; subst. rewrite star_star in Hterm. econstructor. eassumption. assumption.
   inversion 1; subst. rewrite star_star in Hstar. econstructor. eassumption. assumption. assumption.
   inversion 1; subst. rewrite star_star in Hstar. econstructor. eassumption. rewrite <- silent_silent. assumption.
   intros. econstructor. eapply reacts_forever_reactive. eassumption. apply traceinf_sim_refl.
Qed.

Lemma config_behaviors_eq_state_behaves: forall s beh,
  Behaviors.state_behaves S s beh <->
  smallstep.config_behaviors (Step S) (final_state S) s (beh_of_compcert_beh beh).
Proof.
  split.
   eauto using config_behaviors_state_behaves.
   intros.
   eapply state_behaves_bisim.
   eapply state_behaves_config_behaviors.
   eassumption.
   apply beh_bisim_sym.
   apply compcert_of_beh_of_compcert.
Qed.

End State_behaves.

End IO.