Library Rustre.Correctness

Require Import Rustre.Nelist.
Require Import List.
Import List.ListNotations.
Open Scope list_scope.

Require Import Rustre.Common.
Open Scope positive.

Require Import Rustre.Memory.
Require Import Rustre.Dataflow.

Require Import Rustre.Dataflow.IsVariable.Decide.
Require Import Rustre.Dataflow.IsDefined.Decide.

Require Import Rustre.Minimp.
Require Import Rustre.Translation.
Require Import Rustre.Correctness.Proper.
Require Import Rustre.Correctness.IsPresent.
Require Import Rustre.Correctness.MemoryCorres.

Technical lemmas


Lemma exp_eval_tovar:
   x v menv env memories,
    exp_eval menv env (tovar memories x) v
     (exp_eval menv env (State x) v PS.In x memories)
         (exp_eval menv env (Var x) v ¬PS.In x memories).
Proof.
  split; intro Heval;
  destruct In_dec with x memories as [Hxm|Hxm];
    pose proof Hxm as Hxmt;
    apply PS.mem_spec in Hxmt || apply mem_spec_false in Hxmt;
    unfold tovar in *;
    rewrite Hxmt in *;
    intuition.
Qed.

Lemma stmt_eval_translate_eqns_cons:
   prog mems menv env menv´ env´ eq eqs,
    stmt_eval prog menv env (translate_eqns mems (eq :: eqs)) (menv´, env´)
    
    ( menv´´ env´´,
        stmt_eval prog menv env (translate_eqns mems eqs) (menv´´, env´´)
         stmt_eval prog menv´´ env´´ (translate_eqn mems eq) (menv´, env´)).
Proof.
  split.
  - intro H.
    unfold translate_eqns in H.
    simpl in H.
    apply stmt_eval_fold_left_shift in H.
    destruct H as [menv´´ [env´´ [H1 H2]]].
     menv´´, env´´.
    split; [now apply H1|].
    inversion_clear H2.
    inversion H0.
    subst.
    exact H.
  - intro H.
    destruct H as [menv´´ [env´´ [H1 H2]]].
    unfold translate_eqns.
    simpl.
    apply stmt_eval_fold_left_shift.
     menv´´, env´´.
    split; [now apply H1|].
    eapply Icomp. apply H2.
    apply Iskip.
Qed.

Control absorption theorems


Lemma stmt_eval_Control_fwd:
   prog menv env mems c s menv´ env´,
    stmt_eval prog menv env (Control mems c s) (menv´, env´)
    → (Is_present_in mems menv env c
         stmt_eval prog menv env s (menv´, env´))
        (Is_absent_in mems menv env c
            menv´ = menv env´ = env).
Proof.
  Hint Constructors Is_present_in Is_absent_in.
  intros prog menv env mems c s menv´ env´ Hs.
  revert s Hs.
  induction c; [now intuition|].
  intros s Hs.
  simpl in Hs.
  destruct b;
    specialize (IHc _ Hs); clear Hs;
    destruct IHc as [[Hp Hs]|[Hp [Hmenv Henv]]];
    try inversion_clear Hs;
  (left; now intuition)
   || (right;
        repeat progress
               match goal with
               | H: stmt_eval _ _ _ Skip _ |- _inversion H; subst; clear H
               | Hp: Is_present_in _ _ _ _,
                     He: exp_eval _ _ _ _ |- Is_absent_in _ _ _ _
                 ⇒ apply IsAbs2 with (1:=Hp) (2:=He)
               | _intuition
               end).
Qed.

Lemma stmt_eval_Control:
   prog mems menv env ck stmt,
    (Is_absent_in mems menv env ck
     → stmt_eval prog menv env (Control mems ck stmt) (menv, env))
    
    ( menv´ env´,
       Is_present_in mems menv env ck
       → stmt_eval prog menv env stmt (menv´, env´)
       → stmt_eval prog menv env (Control mems ck stmt) (menv´, env´)).
Proof.
  intros prog mems menv env ck.
  induction ck; intro s; split.
  - inversion 1.
  - intros menv´ env´ Hp Hs; exact Hs.
  - inversion_clear 1 as [ ? ? ? Hp|? ? ? ? Hp Hexp Hneq];
    destruct b;
    try (now apply IHck with (1:=Hp));
    apply not_eq_sym in Hneq;
      (apply Bool.not_true_is_false in Hneq
       || apply Bool.not_false_is_true in Hneq);
      subst;
      apply IHck with (1:=Hp);
      (apply Iifte_false with (1:=Hexp)
       || apply Iifte_true with (1:=Hexp));
      constructor.
  - inversion_clear 1 as [|? ? ? Hp Hexp];
    intro Hs;
    destruct b;
    eapply IHck; eauto.
Qed.

If the clock is absent, then the controlled statement evaluates as a Skip.

Theorem stmt_eval_Control_absent:
   prog mems menv env ck stmt,
    Is_absent_in mems menv env ck
    → stmt_eval prog menv env (Control mems ck stmt) (menv, env).
Proof.
  apply stmt_eval_Control.
Qed.

If the clock is present, then the controlled statement evaluates as the underlying one.

Theorem stmt_eval_Control_present:
   prog mems menv env ck stmt menv´ env´,
    Is_present_in mems menv env ck
    → stmt_eval prog menv env stmt (menv´, env´)
    → stmt_eval prog menv env (Control mems ck stmt) (menv´, env´).
Proof.
  apply stmt_eval_Control.
Qed.

More technical lemmas


Lemma stmt_eval_translate_cexp_menv_inv:
   prog menv env mems x menv´ env´ ce,
    stmt_eval prog menv env (translate_cexp mems x ce) (menv´, env´)
    → menv´ = menv.
Proof.
  intros prog menv env mems x menv´ env´.
  induction ce;
  (apply IHce || inversion_clear 1); auto.
Qed.

Lemma stmt_eval_translate_cexp_env_add:
   prog menv env mems x menv´ env´ ce,
    stmt_eval prog menv env (translate_cexp mems x ce) (menv´, env´)
    → c, env´ = PM.add x c env.
Proof.
  intros prog menv env mems x menv´ env´.
  induction ce;
    (apply IHce || inversion_clear 1); auto.
   v; rewrite <- H1; intuition.
Qed.

Lemma not_Is_defined_in_eq_stmt_eval_menv_inv:
   prog x eq menv env mems menv´ env´,
    ¬Is_defined_in_eq x eq
    → stmt_eval prog menv env (translate_eqn mems eq) (menv´, env´)
    → mfind_mem x menv´ = mfind_mem x menv.
Proof.
  intros prog x eq menv env mems menv´ env´ Hneq Heval.
  destruct eq as [y ck ce | y ck f les | y ck v0 lae]; simpl in Heval.
  - apply stmt_eval_Control_fwd in Heval; destruct Heval as [Heval|Heval];
    destruct Heval as [Heval1 Heval2].
    apply stmt_eval_translate_cexp_menv_inv in Heval2.
    rewrite Heval2. intuition.
    destruct Heval2 as [Hmenv]; rewrite Hmenv; intuition.
  - simpl in Heval.
    apply stmt_eval_Control_fwd in Heval; destruct Heval as [Heval|Heval];
    destruct Heval as [Hipi Heval].
    inversion_clear Heval.
    rewrite <- H2.
    reflexivity.
    destruct Heval as [Hmenv Henv]; rewrite Hmenv; intuition.
  - apply not_Is_defined_in_eq_EqFby in Hneq.
    unfold translate_eqn in Heval.
    apply stmt_eval_Control_fwd in Heval; destruct Heval as [Heval|Heval];
    destruct Heval as [Hipi Heval].
    inversion_clear Heval.
    rewrite <- H0.
    unfold mfind_mem, madd_mem.
    simpl; rewrite PM.gso; [intuition | apply Hneq].
    destruct Heval as [Hmenv Henv]; rewrite Hmenv; intuition.
Qed.

Lemma not_Is_defined_in_eq_stmt_eval_mobj_inv:
   prog x eq menv env mems menv´ env´,
    ¬Is_defined_in_eq x eq
    → stmt_eval prog menv env (translate_eqn mems eq) (menv´, env´)
    → mfind_inst x menv´ = mfind_inst x menv.
Proof.
  intros prog x eq menv env mems menv´ env´ Hneq Heval.
  destruct eq as [y ck ce | y ck f les | y ck v0 lae].
  - simpl in Heval.
    apply stmt_eval_Control_fwd in Heval; destruct Heval as [Heval|Heval];
    destruct Heval as [Heval1 Heval2].
    apply stmt_eval_translate_cexp_menv_inv in Heval2.
    rewrite Heval2. intuition.
    destruct Heval2 as [Hmenv]; rewrite Hmenv; intuition.
  - simpl in Heval.
    apply stmt_eval_Control_fwd in Heval; destruct Heval as [Heval|Heval];
    destruct Heval as [Hipi Heval].
    + inversion_clear Heval.
      rewrite <- H2.
      destruct (ident_eq_dec x y) as [Hxy|Hxny].
      × rewrite Hxy in Hneq; exfalso; apply Hneq; constructor.
      × rewrite mfind_inst_gso; [reflexivity|assumption].
    + destruct Heval as [HR1 HR2]; rewrite HR1; reflexivity.
  - unfold translate_eqn in Heval.
    apply stmt_eval_Control_fwd in Heval; destruct Heval as [Heval|Heval];
    destruct Heval as [Hipi Heval].
    inversion_clear Heval.
    rewrite <- H0.
    unfold mfind_inst, madd_mem.
    reflexivity.
    destruct Heval as [Hmenv Henv]; rewrite Hmenv; intuition.
Qed.

Lemma not_Is_variable_in_eq_stmt_eval_env_inv:
   prog x eq menv env mems menv´ env´,
    ¬Is_variable_in_eq x eq
    → stmt_eval prog menv env (translate_eqn mems eq) (menv´, env´)
    → PM.find x env´ = PM.find x env.
Proof.
  intros prog x eq menv env mems menv´ env´ Hnd Heval.
  destruct eq as [y ck ce | y ck f les | y ck v0 ce]; simpl in Heval.
  - unfold translate_eqn in Heval.
    apply stmt_eval_Control_fwd in Heval;
      destruct Heval as [[Hipi Heval]|[Habs [Hmenv Henv]]];
      [|rewrite Henv; reflexivity].
    apply stmt_eval_translate_cexp_env_add in Heval.
    destruct Heval; rewrite H; rewrite PM.gso;
    [reflexivity|intro HR; rewrite HR in *; apply Hnd; constructor].
  - apply stmt_eval_Control_fwd in Heval; destruct Heval as [Heval|Heval];
    destruct Heval as [Heval1 Heval2].
    inversion_clear Heval2.
    rewrite <- H3.
    rewrite PM.gso; [reflexivity|].
    intro Hxy; apply Hnd; rewrite Hxy; constructor.
    destruct Heval2 as [Hmenv Henv]; rewrite Henv; intuition.
  - apply stmt_eval_Control_fwd in Heval; destruct Heval as [Heval|Heval];
    destruct Heval as [Heval1 Heval2].
    inversion Heval2; intuition.
    destruct Heval2 as [Hmenv Henv]; rewrite Henv; intuition.
Qed.

Lemma not_Is_defined_in_eq_stmt_eval_env_inv:
   prog x eq menv env mems menv´ env´,
    ¬Is_defined_in_eq x eq
    → stmt_eval prog menv env (translate_eqn mems eq) (menv´, env´)
    → PM.find x env´ = PM.find x env.
Proof.
  intros prog x eq menv env mems menv´ env´ Hidi Hstmt.
  apply not_Is_defined_in_eq_not_Is_variable_in_eq in Hidi.
  now apply not_Is_variable_in_eq_stmt_eval_env_inv with (1:=Hidi) (2:=Hstmt).
Qed.

Lemma stmt_eval_translate_eqns_menv_inv:
   prog menv env mems eqs menv´ env´,
    stmt_eval prog menv env (translate_eqns mems eqs) (menv´, env´)
    → ( x, ¬Is_defined_in_eqs x eqs
                  mfind_mem x menv´ = mfind_mem x menv).
Proof.
  induction eqs as [ |eq].
  + inversion_clear 1; reflexivity.
  + intros menv´ env´ Heval x Hnmem.
    apply not_Is_defined_in_cons in Hnmem.
    destruct Hnmem as [H0 H1].
    apply stmt_eval_translate_eqns_cons in Heval.
    destruct Heval as [menv´´ [env´´ [Heval0 Heval1]]].
    apply IHeqs with (x:=x) (2:=H1) in Heval0.
    apply not_Is_defined_in_eq_stmt_eval_menv_inv with (1:=H0) in Heval1.
    rewrite Heval1, Heval0.
    reflexivity.
Qed.

Lemma stmt_eval_translate_eqns_minst_inv:
   prog menv env mems eqs menv´ env´,
    stmt_eval prog menv env (translate_eqns mems eqs) (menv´, env´)
    → ( x, ¬Is_defined_in_eqs x eqs
                  mfind_inst x menv´ = mfind_inst x menv).
Proof.
  induction eqs as [ |eq].
  + inversion_clear 1; reflexivity.
  + intros menv´ env´ Heval x Hnmem.
    apply not_Is_defined_in_cons in Hnmem.
    destruct Hnmem as [H0 H1].
    apply stmt_eval_translate_eqns_cons in Heval.
    destruct Heval as [menv´´ [env´´ [Heval0 Heval1]]].
    apply IHeqs with (x:=x) (2:=H1) in Heval0.
    apply not_Is_defined_in_eq_stmt_eval_mobj_inv with (1:=H0) in Heval1.
    rewrite Heval1, Heval0.
    reflexivity.
Qed.

Lemma stmt_eval_translate_eqns_env_inv:
   prog menv env mems eqs menv´ env´,
    stmt_eval prog menv env (translate_eqns mems eqs) (menv´, env´)
    → ( x, ¬Is_variable_in_eqs x eqs
                  PM.find x env´ = PM.find x env).
Proof.
  induction eqs as [ |eq].
  + inversion_clear 1; reflexivity.
  + intros menv´ env´ Heval x Hndef.
    apply not_Is_variable_in_cons in Hndef.
    destruct Hndef as [H0 H1].
    apply stmt_eval_translate_eqns_cons in Heval.
    destruct Heval as [menv´´ [env´´ [Heval0 Heval1]]].
    apply IHeqs with (x:=x) (2:=H1) in Heval0.
    apply not_Is_variable_in_eq_stmt_eval_env_inv with (1:=H0) in Heval1.
    rewrite Heval1, Heval0.
    reflexivity.
Qed.

Local Ltac split_env_assumption :=
  match goal with
  | Henv: context Is_free_in_lexp [_], Hsem: sem_var_instant _ ?y _
    |- _apply Henv in Hsem; [destruct Hsem |solve [auto]]; clear Henv
  | Henv: context Is_free_in_clock [_], Hsem: sem_var_instant _ ?y _
    |- _apply Henv in Hsem; [destruct Hsem |solve [auto]]; clear Henv
  end.

Correctness of expression's translation

An imperative stack env and an imperative memory menv are faithful to a dataflow environment R over a set of identifiers Is_free if their values agree with the dataflow environment.

Definition equiv_env (Is_free: identProp) R memories env menv :=
   x c, Is_free x
              → sem_var_instant R x (present c)
              → (¬PS.In x memoriesPM.find x env = Some c)
               (PS.In x memoriesmfind_mem x menv = Some c).
Hint Unfold equiv_env.

We often need to weaken an equivalence to a subterm: for example we know that the environments are equivalent for all Is_free_caexp x (AnnExp e ck), we can deduce that the environements are equivalent for all Is_free_exp x e.
Lemma equiv_env_map (Is_free1 Is_free2: identProp) H memories env menv:
  ( x, Is_free2 xIs_free1 x) →
  equiv_env Is_free1 H memories env menv
  equiv_env Is_free2 H memories env menv.
Proof.
  intros Hinv Hequiv1 x **; auto.
Qed.

Ltac weaken_equiv_env :=
  match goal with
    | [H: equiv_env ?is_free1 ?R ?mems ?env ?menv
       |- equiv_env ?is_free2 ?R ?mems ?env ?menv] ⇒
      eapply equiv_env_map; [|exact H]; constructor(auto)
  end.

Correctness of clock's translation


Lemma get_exp_eval_tovar:
   x mems menv env v,
    (¬ PS.In x memsPM.find x env = Some v)
    → (PS.In x memsmfind_mem x menv = Some v)
    → exp_eval menv env (tovar mems x) v.
Proof.
  intros x mems menv env v Hvar Hmem.
  unfold tovar.
  destruct (In_dec x mems) as [Hin|Hnin].
  - specialize (Hmem Hin).
    apply PS.mem_spec in Hin; rewrite Hin.
    constructor; exact Hmem.
  - specialize (Hvar Hnin).
    apply mem_spec_false in Hnin; rewrite Hnin.
    constructor; exact Hvar.
Qed.

Theorem clock_correct_true:
   base R memories menv env ck,
    equiv_env (fun xIs_free_in_clock x ck) R memories env menv
    → sem_clock_instant base R ck true
    → Is_present_in memories menv env ck.
Proof.
  Hint Constructors Is_present_in.
  Hint Constructors sem_clock_instant.
  Hint Constructors Is_free_in_clock.
  Hint Constructors exp_eval.
  intros until env.
  induction ck as [|? ? x]; [ intuition | ].
  intro Henv.
  inversion_clear 1.
  constructor. apply IHck; auto.
  intros.
  split_env_assumption.
  apply exp_eval_tovar.
    destruct In_dec with x memories;
    intuition.
Qed.

Theorem clock_correct_false:
   R memories menv env ck,
    equiv_env (fun xIs_free_in_clock x ck) R memories env menv
    → sem_clock_instant true R ck false
    → Is_absent_in memories menv env ck.
Proof.
  Hint Constructors Is_absent_in sem_clock_instant Is_free_in_clock exp_eval.
  intros until env.
  induction ck as [|? ? x]; [ now inversion 2 | ].
  intro Henv.
  inversion_clear 1.
  constructor; apply IHck; now auto.
  eapply clock_correct_true in H0; auto.
  eapply IsAbs2; eauto.
  split_env_assumption.
  destruct In_dec with x memories as [Hin|Hin];
    match goal with
    | H:¬PS.In _ __, Hin:¬PS.In _ _ |- _specialize (H Hin)
    | H:PS.In _ __, Hin:PS.In _ _ |- _specialize (H Hin)
    end;
    apply PS.mem_spec in Hin || apply mem_spec_false in Hin;
    unfold tovar;
    rewrite Hin;
    intuition.
Qed.

Correctness of translate_lexp


Theorem lexp_correct:
   R memories menv env c e,
    sem_lexp_instant true R e (present c)
    → equiv_env (fun xIs_free_in_lexp x e) R memories env menv
    → exp_eval menv env (translate_lexp memories e) c.
Proof.
Hint Constructors exp_eval.
intros until e. revert c.
induction e as [c0|y|e IH y yb| op les IHle] using lexp_ind2; intro c;
  inversion 1; try (subst; injection H1); intros; subst; try apply IH; try apply econst; auto.
× split_env_assumption;
    unfold translate_lexp;
    destruct (PS.mem y memories) eqn:Hm;
    rewrite PS.mem_spec in Hm || rewrite mem_spec_false in Hm;
    auto.
× subst. simpl. apply eop with cs.
  + clear H2 H4 H.
    assert (Hlen : Nelist.length les = Nelist.length cs).
    { rewrite <- (Nelist.map_length present). eapply Nelist.Forall2_length; eassumption. }
    revert cs Hlen H3. induction les; intros cs Hlen Hrec.
    - { destruct cs as [c1 | c1 cs].
        + constructor. inversion_clear Hrec. inversion_clear IHle. apply H0; trivial.
          weaken_equiv_env. constructor(eauto).
        + destruct cs; simpl in Hlen; discriminate. }
    - { destruct cs as [c1 | c1 cs].
        × inversion Hrec.
        × simpl. constructor.
          + inversion_clear IHle.
            apply H.
            - now inversion_clear Hrec.
            - weaken_equiv_env. constructor(auto).
          + inversion_clear Hrec. apply IHles; omega || trivial.
            - now inversion IHle.
            - weaken_equiv_env. inversion H1. constructor(assumption).
            - simpl in Hlen. omega. }
  + destruct (apply_op op cs); now inversion H2.
Qed.

Theorem lexps_correct:
   R memories menv env cs es,
    let vs := Nelist.map present cs in
    Nelist.Forall2 (fun e vsem_lexp_instant true R e v) es vs
    → equiv_env (fun xNelist.Exists (Is_free_in_lexp x) es) R memories env menv
    → Nelist.Forall2 (exp_eval menv env) (Nelist.map (translate_lexp memories) es) cs.
Proof.
  Hint Constructors Forall2.
  intros until cs.
  induction cs; intros es Hvs Hsem Hequiv; subst Hvs.
  + inv Hsem. constructor. eapply lexp_correct; eauto.
    repeat intro. apply Hequiv; trivial. now constructor.
  + inv Hsem. simpl. constructor.
    - eapply lexp_correct; eauto.
      repeat intro. apply Hequiv; trivial. now constructor.
    - apply IHcs; trivial. repeat intro. apply Hequiv; trivial. now constructor(assumption).
Qed.

Correctness of translate_cexp


Theorem cexp_correct:
   R memories prog menv env c x e,
    sem_cexp_instant true R e (present c)
    → equiv_env (fun xIs_free_in_cexp x e) R memories env menv
    → stmt_eval prog menv env (translate_cexp memories x e)
                                                        (menv, PM.add x c env).
Proof.
  intros until x.
  induction e as [b et IHt ef IHf|e].
  -
    inversion_clear 1; intro Henv.
    + apply Iifte_true.
      split_env_assumption.
      apply get_exp_eval_tovar; now auto.
      apply IHt; now auto.
    + apply Iifte_false.
      split_env_assumption.
      apply get_exp_eval_tovar; now auto.
      apply IHf; now auto.
  -
    inversion_clear 1; intro Henv.
    unfold translate_cexp.
    econstructor.
    eapply lexp_correct; [eassumption|now auto].
    reflexivity.
Qed.

Lemma Is_memory_in_msem_var:
   G bk H M n x eqs c,
    Is_defined_in_eqs x eqs
    → ¬Is_variable_in_eqs x eqs
    → sem_var_instant (restr H n) x (present c)
    → List.Forall (msem_equation G bk H M) eqs
    → ( ms, mfind_mem x M = Some ms ms n = c).
Proof.
  induction eqs as [|eq eqs IH];
  inversion_clear 1 as [? ? Hidi|? ? Hidi];
  intros Hnvi Hsv Hmsem;
  apply Forall_cons2 in Hmsem;
  apply not_Is_variable_in_cons in Hnvi;
  destruct Hnvi as [Hnvi0 Hnvi1];
  destruct Hmsem as [Hmeq Hmeqs].
  - destruct eq; inversion Hidi; subst;
    try (exfalso; apply Hnvi0; now constructor).
    inversion_clear Hmeq as [| |? ? ? ? ? ? ls ? ? ? Hmfind Hms0 Hsemls HxS Hmls].
     ms.
    split; [apply Hmfind|].
    specialize Hmls with n; specialize (HxS n); simpl in HxS.
    destruct (ls n);
      destruct Hmls as [Hms Hsv´].
    rewrite Hsv´ in HxS.
    + assert (present c = absent) by sem_det; discriminate.
    + cut (present (ms n) = present c); [injection 1; auto|].
      rewrite Hsv´ in HxS.
      sem_det.
  - apply IH; assumption.
Qed.

Correctness of translate_eqns


Definition equiv_node G prog f :=
   n xss ys M menv inputs output,
    Memory_Corres G n f M menv
    → msem_node G f xss M ys
    → present_list xss n inputs
    → ys n = present output
    → menv´,
         stmt_step_eval prog menv f inputs menv´ output
      Memory_Corres G (S n) f M menv´.

Definition equiv_prog G prog :=
   f,
    equiv_node G prog f.

Section IsStepCorrect.

Variables (G: global)
          (HG: Welldef_global G)
          (bk: stream bool)
          (H: history)
          (M: memory)
          (mems: PS.t)
          (alleqs : list equation)
          (Hsems: msem_equations G bk H M alleqs)
          (prog: program)
          (Hnode: equiv_prog G prog)
          (inputs: nelist ident)
          (Hinput: input, Nelist.In input inputs¬ PS.In input mems)
          (n: nat)
          (Hbase: bk n = true)
          (menv: heap)
          (env: stack).

The proof is by induction on eqs. In the step case, eqs = eq :: beqs, and we may assume that beqs satisfies the three properties. We also know that eqs is a suffix of alleqs and thus that it has a semantics and is well scheduled. For each of the three cases for eq, there are two sub-cases for the value of its clock.

Lemma is_step_correct:
   (eqs: list equation),
        ( oeqs, alleqs = oeqs ++ eqs)
        → ( x, PS.In x mems
                      → (Is_defined_in_eqs x alleqs
                           ¬Is_variable_in_eqs x alleqs))

        
        → ( c input, Nelist.In input inputs
                            (sem_var_instant (restr H n) input (present c)
                              PM.find input env = Some c))

        → ( x, Is_variable_in_eqs x eqsPM.find x env = None)
        → ( input, Nelist.In input inputs¬ Is_defined_in_eqs input eqs)

        
        → Is_well_sch mems inputs eqs

        
        → List.Forall (Memory_Corres_eq G n M menv) alleqs

        
        → ( menv´ env´,
               stmt_eval prog menv env (translate_eqns mems eqs) (menv´, env´)
                ( x, Is_variable_in_eqs x eqs
                             → c : const, sem_var_instant (restr H n) x (present c)
                                                   PM.find x env´ = Some c)
               
                List.Forall (Memory_Corres_eq G (S n) M menv´) eqs).
Proof.
  induction eqs as [|eq];
    [ intros; menv, env;
      split; [unfold translate_eqns; constructor|];
      split; intros; [ match goal with
                       | H:Is_variable_in_eqs _ nil |- _inversion H
                       end | now constructor ]| ].
  intros Hall Hinmems Hin Henv Hin2 Hwsch Hmc.

  assert ( menv´ env´,
             stmt_eval prog menv env (translate_eqns mems eqs) (menv´, env´)
              ( x, Is_variable_in_eqs x eqs
                           → c, sem_var_instant (restr H n) x (present c)
                                         PM.find x env´ = Some c)
              List.Forall (Memory_Corres_eq G (S n) M menv´) eqs) as IHeqs´.
  { eapply IHeqs; trivial.
    - apply List_shift_away with (1:=Hall).
    - intros; apply Henv; constructor(assumption).
    - intros; eapply not_Is_defined_in_cons; eauto.
    - eapply Is_well_sch_cons; eauto.
  }

  clear IHeqs.
  destruct IHeqs´ as [menv´ [env´ [Hstmt [IHeqs0 IHeqs1]]]].

  destruct Hall as [oeqs Hall].
  assert (Hsems´ := Hsems); rewrite Hall in Hsems´.

  apply Forall_app in Hsems´.
  destruct Hsems´ as [H0 Hsems´]; clear H0.
  apply Forall_cons2 in Hsems´.
  destruct Hsems´ as [Hsem Hsems´].

  inversion Hsem as [ bk0 H0 M0 i ck xs ce Hvar Hce HR1 HR2 HR3
                    | bk0 H0 M0 y ck f Mo les ls xs Hmfind Hlaes Hvar Hmsem HR1 HR2 HR3
                    | bk0 H0 M0 ms y ck ls yS v0 le Hmfind Hms0 Hlae HyS Hvar HR1 HR2 HR3];
    subst bk0 H0 M0 eq;

    specialize (Hvar n).
  -
     menv´.
    specialize (Hce n); simpl in ×.
    assert (equiv_env (fun xIs_free_in_caexp x ck ce) (restr H n) mems env´ menv´)
      as Hce´.
    {
      Hint Constructors Is_free_in_eq.
      intros.
      split; intro Hmems.

      - assert (Hdecide_x: Is_variable_in_eqs x eqs Nelist.In x inputs)
          by (eapply Is_well_sch_free_variable; eauto).

        destruct Hdecide_x; try subst x.
        + apply IHeqs0; assumption.
        + erewrite stmt_eval_translate_eqns_env_inv; try eassumption.
          now apply Hin.
          apply not_Is_defined_in_not_Is_variable_in.
          intro Hnot_def. eapply Hin2; try eauto.
          constructor(assumption).

      - assert (¬ Is_defined_in_eqs x eqs)
          by (eapply Is_well_sch_free_variable_in_mems; eauto).

        specialize (Hinmems _ Hmems); destruct Hinmems.
        erewrite stmt_eval_translate_eqns_menv_inv; try eassumption.
        eapply Is_memory_in_msem_var in H1; try eassumption. do 2 destruct H1; subst c.
        assert (Is_defined_in_eqs x alleqs) by intuition.
        assert (¬ Is_variable_in_eqs x alleqs) by intuition.
        erewrite Is_memory_in_Memory_Corres_eqs; try eauto.
    }

    inversion Hce; subst ck0 ce0;
    match goal with
      | H: present _ = xs n |- _rewrite <- H in ×
      | H: absent = xs n |- _rewrite <- H in ×
    end.

    +
      rename c into v.
       (PM.add i v env´).
      split; [|split].
      × apply stmt_eval_translate_eqns_cons.
         menv´, env´.
        split; [exact Hstmt|].
        apply stmt_eval_Control_present.
        eapply clock_correct_true; now eauto.
        rewrite Hbase in ×.
        eapply cexp_correct; now eauto.
      × intros x0 Hivi c.
        inversion_clear Hivi as [? ? Hivi´|]; [inversion_clear Hivi´|].
        { rewrite PM.gss; split; intro HH.
          - assert (c = v) by
                (cut (present c = present v); [injection 1; auto|];
                 sem_det).
            congruence.
          - injection HH; intro Heq. subst. assumption. }
        { destruct (ident_eq_dec x0 i) as [Hxy|Hnxy].
          - rewrite Hxy in *; clear Hxy.
            rewrite PM.gss.
            split; intro Hsv´.
            × assert (v = c)
                by (cut (present v = present c); [injection 1; auto|]; sem_det).
              congruence.
            × injection Hsv´; congruence.
          - erewrite PM.gso; try eassumption.
            apply IHeqs0; assumption.
        }
      × rewrite Hall in Hmc.
        apply Forall_app in Hmc; destruct Hmc as [Hmc0 Hmc]; clear Hmc0.
        apply Forall_cons2 in Hmc; destruct Hmc as [Hmc Hmc0]; clear Hmc0.
        inversion_clear Hmc.
        repeat constructor; assumption.

    +
       env´.
      split; [|split].
      × apply stmt_eval_translate_eqns_cons.
         menv´, env´.
        split; [exact Hstmt|].
        rewrite Hbase in ×.
        apply stmt_eval_Control_absent; auto.
        eapply clock_correct_false; eauto.
      × intros x0 Hivi c.
        destruct (Is_variable_in_dec x0 eqs) as [Hvin|Hvin];
          [now apply IHeqs0 with (1:=Hvin)|].
        apply stmt_eval_translate_eqns_env_inv with (2:=Hvin) in Hstmt.
        rewrite Hstmt.
        inversion_clear Hivi as [? ? Hivi´|];
          [|unfold Is_variable_in_eqs in Hvin; contradiction].
        inversion_clear Hivi´.
        split; intro Hsv´.
        assert (present c = absent) by sem_det. discriminate.
        assert (PM.find i env = None).
        apply Henv; now repeat constructor.
        rewrite Hsv´ in *; discriminate.
      × rewrite Hall in Hmc.
        apply Forall_app in Hmc; destruct Hmc as [Hmc0 Hmc]; clear Hmc0.
        apply Forall_cons2 in Hmc; destruct Hmc as [Hmc Hmc0]; clear Hmc0.
        inversion_clear Hmc.
        repeat constructor; assumption.

  -

    
    assert (equiv_env (fun xIs_free_in_laexps x ck les) (restr H n) mems env´ menv´)
      as Hlae´. {
      intros.
      split; intro Hmems.

      - assert (Hdecide_x: Is_variable_in_eqs x eqs Nelist.In x inputs)
          by (eapply Is_well_sch_free_variable;
              eassumption || constructor (assumption)).

        destruct Hdecide_x; try subst x.
        + apply IHeqs0; assumption.
        + erewrite stmt_eval_translate_eqns_env_inv; try eassumption.
          apply Hin; now eauto.
          apply not_Is_defined_in_not_Is_variable_in.
          intro Hnot_def. eapply Hin2; eauto.
          econstructor(eassumption).

      - assert (¬ Is_defined_in_eqs x eqs)
          by (eapply Is_well_sch_free_variable_in_mems;
              eassumption || constructor (assumption)).
        specialize (Hinmems _ Hmems); destruct Hinmems.
        erewrite stmt_eval_translate_eqns_menv_inv; try eassumption.
        eapply Is_memory_in_msem_var in H1; try eassumption. do 2 destruct H1; subst c.
        assert (Is_defined_in_eqs x alleqs) by intuition.
        assert (¬ Is_variable_in_eqs x alleqs) by intuition.
        erewrite Is_memory_in_Memory_Corres_eqs; try eauto.
    }

    
    rewrite Hall in Hmc.
    apply Forall_app in Hmc.
    destruct Hmc as [Hmc0 Hmc]; clear Hmc0.
    apply Forall_cons2 in Hmc.
    destruct Hmc as [Hmceq Hmceqs].
    inversion_clear Hmceq as [|? ? ? ? ? ? Hmc0|].
    specialize (Hmc0 _ Hmfind).
    destruct Hmc0 as [omenv [Hfindo Hmc0]].
    assert (Hmsem´:=Hmsem).
    inversion_clear Hmsem´ as [? ? ? ? ? i o neqs Hck Hfind Hnsem].
    destruct Hnsem as [Hn [Hlsn [Hxsn [Hout Hnsem]]]].

    assert (¬Is_defined_in_eqs y eqs) as Hniii
        by (inversion_clear Hwsch; eauto).

    specialize (Hlaes n);
      specialize (Hxsn n);
      specialize (Hout n);
      simpl in ×.

    inversion_clear Hlaes.
    +
      rename cs into inValues.

      assert ( c : const, xs n = present c) as [outValue Hxsc].
      {
        apply not_absent_present.
        intro HH.
        apply Hout in HH.
        eapply not_absent_present_list; eauto.
      }
      rewrite Hxsc in ×.

      assert ( menv´ : heap,
                stmt_step_eval prog omenv f inValues menv´ outValue
              Memory_Corres G (S n) f Mo menv´) as Hclass
          by (eapply Hnode; eauto).
      destruct Hclass as [omenv´ [Hnstmt Hnmc]].

      simpl in ×.
       (madd_obj y omenv´ menv´), (PM.add y outValue env´).
      split;[|split].
      × apply stmt_eval_translate_eqns_cons.
         menv´, env´.
        split; [exact Hstmt|].
        apply stmt_eval_Control_present.
        eapply clock_correct_true; now eauto.

        assert (equiv_env (fun x : identNelist.Exists (Is_free_in_lexp x) les)
                          (restr H n) mems env´ menv´)
          by weaken_equiv_env.

        assert (Nelist.Forall2 (exp_eval menv´ env´)
               (Nelist.map (translate_lexp mems) les) inValues).
        {
          rewrite Hbase in ×.
          eapply lexps_correct; eauto.
          match goal with
            | H: _ = Nelist.map present inValues |- _rewrite <- H
            | H: Nelist.map present inValues = _ |- _rewrite H
          end. eauto.
        }
        erewrite <-stmt_eval_translate_eqns_minst_inv in Hfindo; eauto.
      × {
          intros x Hivi c.
          apply Is_variable_in_cons in Hivi.
          destruct Hivi as [Hivi | [notxy Hivi] ].
          -
            inversion_clear Hivi.
            rewrite PM.gss. split; intro HH.
            + assert (present c = present outValue) by sem_det.
              congruence.
            + injection HH. intro Heq; rewrite <-Heq.
              apply Hvar.
          -
             not_Is_variable x y.
             rewrite PM.gso; [|assumption].
             apply IHeqs0; assumption.
        }
      × apply Forall_cons.
        2:now apply Memory_Corres_eqs_add_obj with (1:=IHeqs1) (2:=Hniii).
        constructor.
        intros Mo´ Hmfind´.
        rewrite Hmfind in Hmfind´.
        injection Hmfind´; intro Heq; rewrite <-Heq in *; clear Heq Hmfind´.
         omenv´.
        split; [rewrite mfind_inst_gss; reflexivity|exact Hnmc].

    +
       menv´, env´.

      assert (Habs: absent_list ls n
                      List.Forall (rhs_absent_instant (bk0 n) (restr Hn n)) neqs)
        by (eapply subrate_property_eqns; eauto).

      assert (absent_list ls n)
        by (unfold absent_list; rewrite H0, Nelist.map_compose; auto).

      assert (xs n = absent) as Hout´
          by (apply Hout; auto).

      split; [|split].
      × apply stmt_eval_translate_eqns_cons.
         menv´, env´.
        split; [exact Hstmt|].
        rewrite Hbase in ×.
        apply stmt_eval_Control_absent.
        eapply clock_correct_false; eauto.
      × intros x Hivi c.
        destruct (Is_variable_in_dec x eqs) as [Hvin|Hvin];
          [now apply IHeqs0 with (1:=Hvin)|].
        apply stmt_eval_translate_eqns_env_inv with (2:=Hvin) in Hstmt.
        rewrite Hstmt.
        inversion_clear Hivi as [? ? Hivi´|];
          [|unfold Is_variable_in_eqs in Hvin; contradiction].
        inversion Hivi´ as [| ck´ e HR1 [HR2 HR3 HR4]];
        subst ck´ x e.
        split; intro Hsv´.
        { inversion_clear Hsv´ as [Hfind´].
          inversion_clear Hvar as [Hfind´´].
          rewrite Hfind´ in Hfind´´.
          injection Hfind´´; intro HR1; rewrite <-HR1 in *; clear HR1 Hfind´´.
          discriminate. }
        { assert (PM.find y env = None) as Hnone
              by (apply Henv; repeat constructor).
          rewrite Hnone in Hsv´.
          discriminate. }
      × apply Forall_cons; [|now apply IHeqs1].
        constructor.
        intros Mo´ Hmfind´.
        rewrite Hmfind in Hmfind´.
        injection Hmfind´; intro He; rewrite <-He in *; clear He Hmfind´.
         omenv.
        rewrite stmt_eval_translate_eqns_minst_inv with (1:=Hstmt) (2:=Hniii).
        split; [exact Hfindo|].
        eapply Memory_Corres_unchanged; eauto.

  -
    specialize (Hlae n).
    assert (equiv_env (fun xIs_free_in_laexp x ck le) (restr H n) mems env´ menv´)
      as Hlae´.
    {
      intros.
      split; intro Hmems.

      - assert (Hdecide_x: Is_variable_in_eqs x eqs Nelist.In x inputs)
          by (eapply Is_well_sch_free_variable;
              eassumption || constructor (assumption)).

        destruct Hdecide_x; try subst x.
        + apply IHeqs0; assumption.
        + erewrite stmt_eval_translate_eqns_env_inv; try eassumption.
          apply Hin; now eauto.
          apply not_Is_defined_in_not_Is_variable_in.
          intro. eapply Hin2; eauto. econstructor(eassumption).

      - assert (¬ Is_defined_in_eqs x eqs)
          by (eapply Is_well_sch_free_variable_in_mems;
              eassumption || constructor (assumption)).
        specialize (Hinmems _ Hmems); destruct Hinmems.
        erewrite stmt_eval_translate_eqns_menv_inv; try eassumption.
        eapply Is_memory_in_msem_var in H1; try eassumption. do 2 destruct H1; subst c.
        assert (Is_defined_in_eqs x alleqs) by intuition.
        assert (¬ Is_variable_in_eqs x alleqs) by intuition.
        erewrite Is_memory_in_Memory_Corres_eqs; try eauto.
    }

    inversion Hlae; subst ck0 ce;
      match goal with
        | H: present _ = ls n |- _rewrite <- H in ×
        | H: absent = ls n |- _rewrite <- H in ×
      end;
      destruct Hvar as [Hms Hvar].
    +
      rename c into v.
       (madd_mem y v menv´), env´.
      split; [|split].
      × apply stmt_eval_translate_eqns_cons.
         menv´, env´.
        split; [exact Hstmt|].
        apply stmt_eval_Control_present.
        eapply clock_correct_true; now eauto.
        econstructor.
        rewrite Hbase in ×.
        eapply lexp_correct; now eauto.
        reflexivity.
      × intros x Hivi c.
        inversion_clear Hivi as [? ? Hivi´|]; [now inversion_clear Hivi´|].
        apply IHeqs0.
        assumption.
      × rewrite <-Hms.
        apply Forall_cons.
        2:now apply Memory_Corres_eqs_add_mem with (1:=Hmfind) (2:=IHeqs1).
        constructor.
        intros ms´ Hmfind´.
        rewrite Hmfind in Hmfind´.
        injection Hmfind´; intro Heq; rewrite <-Heq in *; clear Hmfind´ Heq.
        now apply mfind_mem_gss.
    +
       menv´, env´.
      split; [|split].
      × apply stmt_eval_translate_eqns_cons.
         menv´, env´.
        split; [exact Hstmt|].
        rewrite Hbase in *;
        apply stmt_eval_Control_absent.
        eapply clock_correct_false; eauto.
      × intros x Hivi c.
        destruct (Is_variable_in_dec x eqs) as [Hvin|Hvin];
          [now apply IHeqs0 with (1:=Hvin)|].
        apply stmt_eval_translate_eqns_env_inv with (2:=Hvin) in Hstmt.
        rewrite Hstmt.
        inversion_clear Hivi as [? ? Hivi´|];
          [|unfold Is_variable_in_eqs in Hvin; contradiction].
        inversion_clear Hivi´.
      × {
          constructor.
          2:assumption.
          constructor.
          intros ms0´ Hmfind´.
          rewrite Hmfind in Hmfind´.
          injection Hmfind´; intro Heq; rewrite <-Heq in *; clear Heq Hmfind´.
          destruct (Is_defined_in_dec y eqs) as [Hxin|Hxin].
          - Hint Constructors Is_defined_in_eq.
            exfalso.
            inversion_clear Hwsch;
              match goal with
                | H: context[¬ Is_defined_in_eqs _ _] |- _
                  eapply H
              end; eauto.

          - eauto.
            rewrite Hall in Hmc.
            apply Forall_app in Hmc.
            destruct Hmc as [HZ Hmc]; clear HZ.
            apply Forall_cons2 in Hmc.
            destruct Hmc as [Hmc HZ]; clear HZ.
            inversion_clear Hmc as [| |? ? ? ? ? ? Hfindc].
            rewrite Hms.
            eapply stmt_eval_translate_eqns_menv_inv in Hstmt;
              try eassumption.
            rewrite Hstmt.
            eapply Hfindc; auto.
        }
Qed.

End IsStepCorrect.

Correctness of translate


Section IsNodeCorrect.

Lemma equiv_prog_empty: equiv_prog [] (translate []).
Proof.
  intro Hwdef.
  intros n **.
  exfalso.
  repeat match goal with
      | H: msem_node [] _ _ _ _ |- _inversion H; clear H
      | H: find_node _ [] = Some _ |- _inversion H; clear H
         end.
Qed.

Lemma assoc_inputs:
   R inArgs inputs c arg,
    Nelist.In arg inArgsNelist.NoDup inArgs
    sem_var_instant R arg (present c) →
    Nelist.Forall2 (sem_var_instant R) inArgs (Nelist.map present inputs) →
    Assoc inArgs inputs arg c.
Proof.
intros × Hin Hnodup Hsem_var. revert inputs.
induction inArgs as [|inArg inArgs´]; intros inputs Hall; inv Hall.
+ symmetry in H0. rewrite Nelist.map_eq_nebase in H0. destruct H0 as [? [? ?]]. subst.
  inv Hin.
  assert (x = c) by (cut (present x = present c); [ injection 1; congruence | sem_det]); subst.
  constructor.
+ symmetry in H1. rewrite Nelist.map_eq_necons in H1. decompose [ex and] H1. clear H1. subst.
  inv Hnodup. inv Hin.
  - assert (x = c) by (cut (present x = present c); [ injection 1; congruence | sem_det]); subst.
    constructor.
  - constructor; auto.
    intro. subst. contradiction.
Qed.

Lemma sem_var_assoc:
   R inArg inputs c input,
    Assoc inArg inputs input c
    Nelist.Forall2 (sem_var_instant R) inArg (Nelist.map present inputs) →
    sem_var_instant R input (present c).
Proof.
  intros ** Hassoc Hsem_vars.
  induction Hassoc; inversion_clear Hsem_vars; eauto.
Qed.

Theorem is_node_correct:
   (G: global),
    Welldef_global G
    equiv_prog G (translate G).
Proof.
  induction G as [|node G IH].
  - intro; apply equiv_prog_empty.
  - intros Hwd f n xs ys M menv inputs output Hmc Hmsem Hxs Hys.
    set (nodeName := n_name node).

    assert (Welldef_global G) as HwdG
        by (eapply Welldef_global_cons; eassumption).

    assert (Ordered_nodes (node::G)) as Hord
        by (apply Welldef_global_Ordered_nodes; assumption).

    destruct (ident_eqb nodeName f) eqn:Hfeq.
    +
      assert (nodeName = f)
        by (apply Pos.eqb_eq; assumption).
      subst f.

      set (prog := translate (node :: G)).

      inversion_clear Hwd as [|? ? Hwd´ eqs inArg outArg
                               HnodupIn Hwsch Hndef_in Hdef_out Hne Hfind Hnodup].
      clear Hwd´.
      inversion_clear Hmsem as [? ? ? ? ? ? ? ? Hck Heqs
                                [H [Hin [Hout [Hrabs Hall]]]]].
      subst eqs inArg outArg nodeName.
      simpl in Heqs; rewrite Hfeq in Heqs; simpl in Heqs.
      injection Heqs. intro Hnode. rewrite Hnode in ×. clear Heqs. simpl in ×.

      rename i into inArg; rename o into outArg; rename eqs0 into eqs.

      set (env := adds inArg inputs sempty).

      assert (msem_equations G bk H M eqs)
        by (eapply Forall_msem_equation_global_tl; try eassumption).

      assert ( (menv´ : heap) (env´ : stack),
                stmt_eval (translate G) menv env (translate_eqns (memories eqs) eqs) (menv´, env´)
                ( x : ident,
                   Is_variable_in_eqs x eqs
                    c : const,
                     sem_var_instant (restr H n) x (present c)
                     PM.find x env´ = Some c)
                Forall (Memory_Corres_eq G (S n) M menv´) eqs) as His_step_correct.
      {
        assert (Hlen : Nelist.length inArg = Nelist.length inputs).
        {
          transitivity (Nelist.length (xs n)).
          × eapply Nelist.Forall2_length; eauto.
          × rewrite Hxs, Nelist.map_length. auto.
        }
        eauto.
        eapply is_step_correct; eauto.
        - apply Hck; eauto.
        - []; auto.
        - intros y Hinm.
          assert (NoDup_defs eqs) as Hndds
              by (eapply Is_well_sch_NoDup_defs; eauto).
          split; [now apply Is_defined_in_memories
                 |now apply not_Is_variable_in_memories].
        - intros c input Hinput.
          specialize (Hin n).
          split; intro HH.
          + subst env.
            apply gsss; eauto.
            rewrite nelist2list_NoDup in HnodupIn. eapply assoc_inputs; eauto; [].
            match goal with | H : context[inputs] |- _move H at bottom end.
            unfold present_list in Hxs. rewrite <- Hxs. auto.
          + subst env.
            apply gsss in HH; trivial.
            eapply sem_var_assoc; eauto.
            match goal with | H : context[inputs] |- _move H at bottom end.
            unfold present_list in Hxs.
            rewrite <- Hxs. auto.
        - intros x Hivi.
          subst env.
          rewrite gsos, PM.gempty; auto.
          rewrite <- Nelist.nelist2list_In. intro.
          apply Is_variable_in_eqs_Is_defined_in_eqs in Hivi.
          apply Hndef_in, Exists_exists; eauto.

        - intros input Hinput Hisdef.
          rewrite <- Nelist.nelist2list_In in Hinput.
          apply Hndef_in. apply Exists_exists.
          eauto.

        - inversion_clear Hmc as [? ? ? ? ? ? Hf Hmeqs].

          simpl in Hf.
          rewrite ident_eqb_refl in Hf.
          injection Hf; intros Heq0 Heq1 Heq2;
          rewrite <-Heq0, <-Heq1, <-Heq2 in *;
          clear Heq0 Heq1 Heq2 Hf.

          eapply Memory_Corres_eqs_node_tl; try eassumption.
      }

      destruct His_step_correct as [menv´ [env´ [Hstmt [Hsemvar Hmem]]]].
       menv´.
      split.
      × {
          econstructor; eauto.
          - simpl. rewrite Pos.eqb_refl. reflexivity.
          - subst env. simpl.
            assert (inArg = n_input node)
              by (rewrite Hnode; auto).
            assert (eqs = n_eqs node)
              by (rewrite Hnode; auto).
            subst inArg; subst eqs.
            rewrite ps_from_list_gather_eqs_memories. eapply Hstmt.
          - assert (outArg = n_output node)
              by (rewrite Hnode; auto).
            subst outArg.
            specialize (Hout n); simpl in Hout; rewrite Hys in Hout.
            simpl. apply Hsemvar; try assumption.
        }
      × {
          econstructor.
          - simpl; rewrite Pos.eqb_refl; reflexivity.
          - eapply Memory_Corres_eqs_node_tl; eassumption.
        }

    +
      assert (nodeName f) as Hfneq
          by (eapply Pos.eqb_neq; try eassumption).

      rewrite Memory_Corres_node_tl in Hmc; try eassumption.
      apply msem_node_cons in Hmsem; try eassumption.
      eapply IH in HwdG.
      edestruct HwdG as [menv´ [Hstmt´ Hmc´]]; try eassumption.
      inversion_clear Hstmt´.
       menv´; split.
      × econstructor; try eassumption.
        simpl; subst nodeName; rewrite Hfeq.
        eassumption.
      × rewrite Memory_Corres_node_tl; try assumption.
Qed.

End IsNodeCorrect.

Correctness of the reset code


Lemma stmt_eval_translate_reset_eqn_shift:
   prog eqs iacc menv env menv´ env´,
    stmt_eval prog menv env
              (List.fold_left translate_reset_eqn eqs iacc)
              (menv´, env´)
    
     menv´´ env´´,
      stmt_eval prog menv env
                (List.fold_left translate_reset_eqn eqs Skip)
                (menv´´, env´´)
      
      stmt_eval prog menv´´ env´´ iacc (menv´, env´).
Proof.
  Hint Constructors stmt_eval.
  induction eqs as [|eq eqs IH].
  - split; [ now eauto | ].
    intro H; do 2 destruct H.
    destruct H as [H0 H1].
    inversion_clear H0; apply H1.
  - intros.
    split.
    + intro H0.
      apply IH in H0.
      destruct H0 as [menv´´ [env´´ [H0 H1]]].
      destruct eq; [now eauto| |];
      inversion_clear H1;
       menv1; env1;
      split; try (simpl; apply IH); eauto.
    + intros.
      destruct eq; [ now (apply IH; auto) | |];
      (apply IH;
       simpl in H;
       destruct H as [menv´´ [env´´ [H0 H1]]];
       apply IH in H0;
       destruct H0 as [menv0 [env0 [H2 H3]]];
        menv0; env0;
       split; [now auto|];
       inversion_clear H3;
       inversion H0; subst;
       econstructor; eauto).
Qed.

Lemma stmt_eval_translate_reset_eqns_cons:
   prog menv env (eq:equation) eqs P,
    ( menv´´ env´´,
        stmt_eval prog menv env
                  (translate_reset_eqns (eq :: eqs)) (menv´´, env´´)
         P menv´´ env´´)
    
    ( menv´ env´ menv´´ env´´,
        stmt_eval prog menv env (translate_reset_eqns eqs) (menv´, env´)
         stmt_eval prog menv´ env´
                     (translate_reset_eqn Skip eq) (menv´´, env´´)
         P menv´´ env´´).
Proof.
  split.
  - intro H.
    destruct H as [menv´´ [env´´ H]].
    unfold translate_reset_eqns in H.
    simpl in H.
    destruct H as [H HP].
    apply stmt_eval_translate_reset_eqn_shift in H.
    destruct H as [menv´ [env´ [H1 H2]]].
     menv´, env´, menv´´, env´´.
    now intuition.
  - intro H.
    destruct H as [menv´ [env´ [menv´´ [env´´ [H1 [H2 HP]]]]]].
    unfold translate_reset_eqns.
    simpl.
     menv´´, env´´.
    intuition.
    apply stmt_eval_translate_reset_eqn_shift.
     menv´, env´.
    intuition.
Qed.

Definition equiv_reset G prog f :=
   xs ys M,
    msem_node G f xs M ys
    → menv´,
         stmt_reset_eval prog f menv´
       Memory_Corres G 0 f M menv´.

Lemma equiv_reset_empty: f, equiv_reset [] (translate []) f.
Proof.
  intro Hwdef.
  intros n **.
  exfalso.
  repeat match goal with
      | H: msem_node [] _ _ _ _ |- _inversion H; clear H
      | H: find_node _ [] = Some _ |- _inversion H; clear H
         end.
Qed.

Section IsResetCorrect.

Variables (G: global)
          (HG: Welldef_global G)
          (H: history)
          (M: memory)
          (mems: PS.t)
          (inputs: Nelist.nelist ident).

Lemma is_reset_correct:
   bk eqs,
    msem_equations G bk H M eqs
    Is_well_sch mems inputs eqs
    ( f, equiv_reset G (translate G) f) →
   menv´ env´,
    stmt_eval (translate G) hempty sempty (translate_reset_eqns eqs)
              (menv´, env´)
     Forall (Memory_Corres_eq G 0 M menv´) eqs.
Proof.
  intros × Hmsem Hwsch Hreset.
  induction eqs as [|eq eqs IH]; [eauto|].
  destruct eq as [i ck e | i ck f e | i ck v e];
  inversion_clear Hmsem as [| ? ? Hsem Hmsem´ ];
  inversion_clear Hwsch;
  edestruct IH as [menv´ [env´ [Hstmt Hmc]]]; try eassumption.
  -
    Hint Constructors Forall.
    Hint Constructors Memory_Corres_eq.
    eauto.
  -
    unfold translate_reset_eqns; simpl.
    inversion_clear Hsem as [|? ? ? ? ? ? Mo ? xs´ ys´ Hmfind Hxs´ Hys´ HsemNode|].
    assert ( omenv, stmt_reset_eval (translate G) f omenv
                        Memory_Corres G 0 f Mo omenv) as [omenv [Hstmt_reset Hmcf]]
by (eapply Hreset; try eassumption).

     (madd_obj i omenv menv´), env´.
    split.
    + rewrite stmt_eval_translate_reset_eqn_shift.
       menv´, env´.
      split; try assumption.
      econstructor; [|constructor].
      econstructor; auto.
      assumption.
    + repeat constructor; [| apply Memory_Corres_eqs_add_obj; eauto].
      intros Hmfind´.
      rewrite Hmfind in Hmfind´; injection Hmfind´; intro Heq; subst .
       omenv.
      rewrite mfind_inst_gss.
      auto.
  -
     (madd_mem i v menv´), env´.
    split.
    + unfold translate_reset_eqns; simpl;
        rewrite stmt_eval_translate_reset_eqn_shift.
       menv´, env´.
      Hint Constructors stmt_eval.
      eauto.
    + inversion_clear Hsem as [| | ? ? ? ? ? ? ? ? ? ? Hmfind Hms Hlae Hls].
      rewrite <- Hms.
      constructor; [|apply Memory_Corres_eqs_add_mem; assumption].
      constructor. intros ms´ Hmfind´.
      rewrite Hmfind in Hmfind´.
      injection Hmfind´; intro HR; rewrite HR in *; clear HR Hmfind´.
      rewrite mfind_mem_gss.
      reflexivity.
Qed.

End IsResetCorrect.

Theorem is_node_reset_correct:
   (G: global) f,
         Welldef_global G
         equiv_reset G (translate G) f.
Proof.
  induction G as [|node G IH].
  - intros f Hwd.
    apply equiv_reset_empty.
  - intros f Hwdef xs ys M Hmsem.

    assert (Ordered_nodes (node :: G)) as HordG
      by (apply Welldef_global_Ordered_nodes; assumption).

    set (nodeName := n_name node).

    destruct (ident_eqb nodeName f) eqn:Heqb.
    + assert (nodeName = f) as Hfeq
        by (apply Pos.eqb_eq; assumption).

      inversion_clear Hmsem as [? ? ? ? ? inArg outArg eqs Hbk Hfind
                                [H [Hin [Hout [Hrhs Hmsem´]]]]].
      rename Hmsem´ into Hmsem.

      specialize (Hin 0)%nat; specialize (Hout 0)%nat;
        simpl in Hin, Hout.

      simpl in Hfind; subst nodeName; rewrite Heqb in Hfind;
      injection Hfind; clear Hfind; intro Hfind.

      assert (msem_equations G bk H M eqs).
      {
        inversion_clear Hwdef. subst eqs0; rewrite Hfind in *; simpl in ×.
        eapply Forall_msem_equation_global_tl; try eassumption.
      }

      assert (Is_well_sch (memories eqs) inArg eqs)
        by (inversion Hwdef; subst ni eqs0;
            rewrite Hfind in *; simpl in *; assumption).

      assert (Welldef_global G)
        by (inversion Hwdef; assumption).

      assert ( menv´ env´,
             stmt_eval (translate G) hempty sempty (translate_reset_eqns eqs)
                       (menv´, env´)
              Forall (Memory_Corres_eq G 0 M menv´) eqs)
        as [menv´ [env´ [Hstmt Hmc]]].
      {
        eapply is_reset_correct; try eassumption.
        intro; apply IH; assumption.
      }

       menv´.
      split.
      × {
          econstructor.
          - simpl. rewrite Heqb. reflexivity.
          - subst node; eassumption.
        }
      × {
          econstructor.
          - simpl; rewrite Heqb. subst node. reflexivity.
          - apply Memory_Corres_eqs_node_tl; try assumption.
            inversion Hwdef. subst eqs0. rewrite Hfind in ×. simpl. assumption.
        }

    + assert (nodeName f) as Hfneq
        by (apply Pos.eqb_neq; assumption).

      apply Welldef_global_cons in Hwdef.
      apply msem_node_cons in Hmsem; try assumption.
      edestruct IH as [menv´ [Hstmt Hmc]]; try eassumption.
       menv´; split.
      × inversion_clear Hstmt.
        econstructor; try eassumption.
        simpl. subst nodeName; rewrite Heqb. assumption.
      × apply Memory_Corres_node_tl; eassumption.
Qed.

Correctness, from the point of view of the event loop


Section EventLoop.

Variables (G : global)
          (main : ident)
          (css : stream (nelist const))
          (ys : stream value)
          (r : ident)
          (obj : ident)
          (Hwdef : Welldef_global G).

Let xss := fun nNelist.map present (css n).

Variable (Hsem: sem_node G main xss ys).

Open Scope nat_scope.
Fixpoint step (n: nat) P r main obj css menv env: Prop :=
  match n with
    | 0 ⇒ stmt_eval P hempty sempty (Reset_ap main obj) (menv, env)
    | S nlet vs := Nelist.map Const (css n) in
       menvN envN, step n P r main obj css menvN envN
                     stmt_eval P menvN envN (Step_ap r main obj vs) (menv, env)
  end.

The proof proceeds by induction on G, a list of dataflow nodes in a reversed order of their instantiation dependencies, which requires showing correctness for the translation of node f under the hypothesis that the translations of the nodes it instantiates are correct.

Lemma is_event_loop_correctG:
   M,
    let P := translate G in
    msem_node G main xss M ys
     n,
     menv env omenv,
      step (S n) P r main obj css menv env
    mfind_inst obj menv = Some omenv
    Memory_Corres G (S n) main M omenv
    ( co, ys n = present co PM.find r env = Some co).
Proof.
  intros × Hmsem n.
  assert ( menv0,
            stmt_reset_eval (translate G) main menv0
          Memory_Corres G 0 main M menv0) as [menv0 [Hstmtr Hmc0]]
by (eapply is_node_reset_correct; try eassumption).

  set (ci0 := css 0).

  assert (Hpres: present_list xss 0 ci0)
    by (subst xss; unfold present_list; eauto).

  assert ( co0, ys 0 = present co0)%nat as [co0 Hco0].
  {
    inversion_clear Hmsem as
        [ ? ? ? ? ? ? ? ? Hbk Hfind
            [H [Hsem_in [Hsem_out [Habs Hsem_eqns]]]]].
    apply not_absent_present;
      rewrite <- Habs;
      eapply not_absent_present_list; eauto.
  }

  induction n.
  -
    assert ( menv,
              stmt_step_eval (translate G) menv0 main ci0 menv co0
               Memory_Corres G 1 main M menv) as [menv1 [Hstmt1 Hmem1]]
by (eapply is_node_correct; eauto).

    do 3 eexists.
    split; [|split; [| split]]; try eauto.
    + do 2 eexists; split; simpl step; eauto.
      econstructor; eauto.
      × apply mfind_inst_gss.
      × apply exps_eval_const.
    + apply mfind_inst_gss.
    + rewrite Hco0, PM.gss. intuition congruence.

  -

    destruct IHn as [menvN [envN [omenvN [HstepN [HmfindN [HmcN HeqN]]]]]].

    set (ciSn := css (S n)).

    assert (HpresN: present_list xss (S n) ciSn)
      by (subst xss; unfold present_list; eauto).

    assert ( coSn, ys (S n) = present coSn) as [coSn Hys].
    {
      inversion_clear Hmsem as
        [ ? ? ? ? ? ? ? ? Hbk Hfind
            [H [Hsem_in [Hsem_out [Habs Hsem_eqns]]]]].
      apply not_absent_present; rewrite <- Habs;
      eapply not_absent_present_list; eauto.
    }

    assert ( omenvSn,
              stmt_step_eval (translate G) omenvN main ciSn omenvSn coSn
           Memory_Corres G (S (S n)) main M omenvSn) as [omenvSn [HstmtSn HmemSn]]
by (eapply is_node_correct; eauto).
     (madd_obj obj omenvSn menvN).
    do 2 eexists.

    split; [|split; [| split]]; eauto.
    + do 2 eexists; split; simpl step; try eauto.
      econstructor; eauto.
      apply exps_eval_const.
    + apply mfind_inst_gss.
    + rewrite Hys, PM.gss. intuition congruence.
Qed.

Theorem is_event_loop_correct:
sem_node G main xss ys
n, menv env,
    step (S n) (translate G) r main obj css menv env
     ( co, ys n = present co PM.find r env = Some co).
Proof.
  intros until n.

  assert ( M, msem_node G main xss M ys) as [M Hmsem]
    by (eapply sem_msem_node; eauto).

  assert ( menv env omenv,
            step (S n) (translate G) r main obj css menv env
          mfind_inst obj menv = Some omenv
          Memory_Corres G (S n) main M omenv
          ( co, ys n = present co PM.find r env = Some co))
    as [menv [env [omenv [Hstep [_ [_ Hys]]]]]]
by (eapply is_event_loop_correctG; eauto).

  do 2 eexists; eauto.
Qed.

End EventLoop.

Correctness of optimized code


Require Import Minimp.FuseIfte.

Lemma not_Can_write_in_translate_cexp:
   x mems ce i,
    x i¬ Can_write_in i (translate_cexp mems x ce).
Proof.
  induction ce.
  - intros j Hxni Hcw.
    specialize (IHce1 _ Hxni).
    specialize (IHce2 _ Hxni).
    inversion_clear Hcw; intuition.
  - intros j Hxni Hcw.
    inversion Hcw; intuition.
Qed.

Lemma Is_free_in_tovar:
   mems i j,
    Is_free_in_exp j (tovar mems i) i = j.
Proof.
  unfold tovar.
  intros mems i j.
  destruct (PS.mem i mems); split; intro HH;
  (inversion_clear HH; reflexivity || subst; now constructor).
Qed.

Lemma Is_fusible_translate_cexp:
   mems x ce,
    ( i, Is_free_in_cexp i cex i)
    → Is_fusible (translate_cexp mems x ce).
Proof.
  intros mems x ce Hfree.
  induction ce.
  - simpl; constructor;
    [apply IHce1; now auto|apply IHce2; now auto|].
    intros j Hfree´; split;
    (apply not_Can_write_in_translate_cexp;
      apply Is_free_in_tovar in Hfree´;
      subst; apply Hfree; constructor).
  - now constructor.
Qed.

Lemma Is_fusible_Control_caexp:
   mems ck f ce,
    ( i, Is_free_in_caexp i ck ce¬Can_write_in i (f ce))
    → Is_fusible (f ce)
    → Is_fusible (Control mems ck (f ce)).
Proof.
  induction ck as [|ck IH i b]; [now intuition|].
  intros f ce Hxni Hfce.
  simpl.
  destruct b.
  - apply IH with (f:=fun ceIfte (tovar mems i) (f ce) Skip).
    + intros j Hfree Hcw.
      apply Hxni with (i0:=j); [inversion_clear Hfree; now auto|].
      inversion_clear Hcw as [| | |? ? ? ? Hskip| | |];
        [assumption|inversion Hskip].
    + repeat constructor; [assumption| |now inversion 1].
      apply Hxni.
      match goal with
      | H:Is_free_in_exp _ (tovar mems _) |- _rename H into Hfree
      end.
      unfold tovar in Hfree.
      destruct (PS.mem i mems); inversion Hfree; subst; now auto.
  - apply IH with (f:=fun ceIfte (tovar mems i) Skip (f ce)).
    + intros j Hfree Hcw.
      apply Hxni with (i0:=j); [inversion_clear Hfree; now auto|].
      inversion_clear Hcw as [| |? ? ? ? Hskip| | | |];
        [inversion Hskip|assumption].
    + repeat constructor; [assumption|now inversion 1|].
      apply Hxni.
      match goal with
      | H:Is_free_in_exp _ (tovar mems _) |- _rename H into Hfree
      end.
      unfold tovar in Hfree.
      destruct (PS.mem i mems); inversion Hfree; subst; now auto.
Qed.

Lemma Is_fusible_Control_laexp:
   mems ck s,
    ( i, Is_free_in_clock i ck¬Can_write_in i s)
    → Is_fusible s
    → Is_fusible (Control mems ck s).
Proof.
  induction ck as [|ck IH i b]; [now intuition|].
  intros s Hxni Hfce.
  simpl.
  destruct b; apply IH.
  - intros j Hfree Hcw.
    apply Hxni with (i0:=j); [inversion_clear Hfree; now auto|].
    inversion_clear Hcw as [| | |? ? ? ? Hskip| | |];
      [assumption|inversion Hskip].
  - repeat constructor; [assumption| |now inversion 1].
    apply Hxni.
    match goal with
    | H:Is_free_in_exp _ (tovar mems _) |- _rename H into Hfree
    end.
    unfold tovar in Hfree.
    destruct (PS.mem i mems); inversion Hfree; subst; now auto.
  - intros j Hfree Hcw.
    apply Hxni with (i0:=j); [inversion_clear Hfree; now auto|].
    inversion_clear Hcw as [| |? ? ? ? Hskip| | | | ];
      [inversion Hskip|assumption].
  - repeat constructor; [assumption|now inversion 1|].
    apply Hxni.
    match goal with
    | H:Is_free_in_exp _ (tovar mems _) |- _rename H into Hfree
      end.
    unfold tovar in Hfree.
    destruct (PS.mem i mems); inversion Hfree; subst; now auto.
Qed.

Require Import Rustre.Dataflow.Clocking.
Require Import Rustre.Dataflow.Clocking.Properties.

Lemma translate_eqns_Is_fusible:
   C mems inputs eqs,
    Well_clocked_env C
    → Forall (Well_clocked_eq C) eqs
    → Is_well_sch mems inputs eqs
    → ( x, PS.In x mems¬Is_variable_in_eqs x eqs)
    → ( input, Nelist.In input inputs¬ Is_defined_in_eqs input eqs)
    → Is_fusible (translate_eqns mems eqs).
Proof.
  intros C mems inputs eqs Hwk Hwks Hwsch Hnvi Hnin.
  induction eqs as [|eq eqs IH]; [now constructor|].
  inversion Hwks as [|eq´ eqs´ Hwkeq Hwks´]; subst.
  specialize (IH Hwks´ (Is_well_sch_cons _ _ _ _ Hwsch)).
  unfold translate_eqns.
  simpl; apply Is_fusible_fold_left_shift.
  split.
  - apply IH.
    + intros x Hin; apply Hnvi in Hin.
      apply not_Is_variable_in_cons in Hin.
      now intuition.
    + intros x Hin. apply Hnin in Hin.
      apply not_Is_defined_in_cons in Hin.
      now intuition.
  - clear IH.
    repeat constructor.
    destruct eq as [x ck e|x ck f e|x ck v0 e]; simpl.
    + assert (¬PS.In x mems) as Hnxm
          by (intro Hin; apply Hnvi with (1:=Hin); repeat constructor).
      inversion_clear Hwsch as [|? ? Hwsch´ HH Hndef].
      assert ( i, Is_free_in_caexp i ck ex i) as Hfni.
      { intros i Hfree.
        assert (Hfree´: Is_free_in_eq i (EqDef x ck e)) by auto.
        eapply HH in Hfree´.
        destruct Hfree´ as [Hm Hnm].
        assert (¬ Nelist.In x inputs) as Hninp
            by (intro Hin; eapply Hnin; eauto; constructor(auto)).

        assert (¬PS.In x mems) as Hnxm´ by intuition.
        intro Hxi; rewrite Hxi in *; clear Hxi.
        specialize (Hnm Hnxm´).
        eapply Hndef; intuition.
        eapply Is_variable_in_eqs_Is_defined_in_eqs. auto. }
      apply Is_fusible_Control_caexp.
      intros i Hfree.
      apply (not_Can_write_in_translate_cexp).
      apply Hfni with (1:=Hfree).
      apply (Is_fusible_translate_cexp).
      intros i Hfree; apply Hfni; intuition.
    + assert (¬Is_free_in_clock x ck) as Hnfree
          by (apply Well_clocked_EqApp_not_Is_free_in_clock
              with (1:=Hwk) (2:=Hwkeq));
      apply Is_fusible_Control_laexp;
      [intros i Hfree Hcw; inversion Hcw; subst; contradiction|intuition].
    + assert (¬Is_free_in_clock x ck) as Hnfree
          by (apply Well_clocked_EqFby_not_Is_free_in_clock
              with (1:=Hwk) (2:=Hwkeq));
      apply Is_fusible_Control_laexp;
      [intros i Hfree Hcw; inversion Hcw; subst; contradiction|intuition].
Qed.