Library Rustre.Minimp.FuseIfte


Require Import Coq.FSets.FMapPositive.
Require Import PArith.
Require Import Rustre.Common.
Require Import Rustre.Minimp.Syntax.
Require Import Rustre.Minimp.Semantics.
Require Import Rustre.Minimp.Equiv.
Require Import Rustre.Dataflow.Syntax.
Require Import Rustre.Dataflow.Memories.
Require Import Rustre.Dataflow.IsFree.
Require Import Rustre.Dataflow.IsDefined.
Require Import Rustre.Dataflow.IsVariable.
Require Import Rustre.Dataflow.WellFormed.
Require Import Rustre.Memory.

Ltac inv H := inversion H; subst; clear H.

Inductive Is_free_in_exp : identexpProp :=
| FreeVar: i,
    Is_free_in_exp i (Var i)
| FreeState: i,
    Is_free_in_exp i (State i)
| FreeOp: i op es,
    Nelist.Exists (Is_free_in_exp i) es
    Is_free_in_exp i (Op op es).

Inductive Can_write_in : identstmtProp :=
| CWIAssign: x e,
    Can_write_in x (Assign x e)
| CWIAssignSt: x e,
    Can_write_in x (AssignSt x e)
| CWIIfteTrue: x e s1 s2,
    Can_write_in x s1
    Can_write_in x (Ifte e s1 s2)
| CWIIfteFalse: x e s1 s2,
    Can_write_in x s2
    Can_write_in x (Ifte e s1 s2)
| CWIStep_ap: x cls obj e,
    Can_write_in x (Step_ap x cls obj e)
| CWIComp1: x s1 s2,
    Can_write_in x s1
    Can_write_in x (Comp s1 s2)
| CWIComp2: x s1 s2,
    Can_write_in x s2
    Can_write_in x (Comp s1 s2).

Lemma cannot_write_in_Ifte:
   x e s1 s2,
    ¬ Can_write_in x (Ifte e s1 s2)
    
    ¬ Can_write_in x s1 ¬ Can_write_in x s2.
Proof.
  Hint Constructors Can_write_in.
  intros; split; intro; try (intro HH; inversion_clear HH); intuition.
Qed.

Lemma cannot_write_in_Comp:
   x s1 s2,
    ¬ Can_write_in x (Comp s1 s2)
    
    ¬ Can_write_in x s1 ¬ Can_write_in x s2.
Proof.
  Hint Constructors Can_write_in.
  intros; split; intro; try (intro HH; inversion_clear HH); intuition.
Qed.

Ltac cannot_write :=
    repeat progress
           match goal with
           | |- x, Is_free_in_exp x __intros
           | Hfw: ( x, Is_free_in_exp x ?e_),
                  Hfree: Is_free_in_exp ?x ?e |- _
             ⇒ pose proof (Hfw x Hfree); clear Hfw
           | |- _ _split
           | |- ¬Can_write_in _ _intro
           | H: Can_write_in _ (Comp _ _) |- _inversion_clear H
           | _now intuition
           end.

Lemma not_free_aux : i op es,
  ¬Is_free_in_exp i (Op op es) → Nelist.Forall (fun e¬Is_free_in_exp i e) es.
Proof.
intros i op es Hfree. induction es.
+ constructor. intro Habs. apply Hfree. now do 2 constructor.
+ constructor.
  - intro Habs. apply Hfree. now do 2 constructor.
  - apply IHes. intro Habs. inversion_clear Habs.
    apply Hfree. constructor. now constructor 3.
Qed.

Ltac not_free :=
  lazymatch goal with
    | H : ¬Is_free_in_exp ?x (Var ?i) |- _let HH := fresh in
        assert (HH : i x) by (intro; subst; apply H; constructor);
        clear H; rename HH into H
    | H : ¬Is_free_in_exp ?x (State ?i) |- _let HH := fresh in
        assert (HH : i x) by (intro; subst; apply H; constructor);
        clear H; rename HH into H
    | H : ¬Is_free_in_exp ?x (Op ?op ?es) |- _apply not_free_aux in H
  end.

If we add irrelevent values to env, evaluation does not change.
Lemma exp_eval_extend_env : mem env x e v,
  ¬Is_free_in_exp x eexp_eval mem env e vexp_eval mem (PM.add x env) e v.
Proof.
intros mem env x e.
induction e using exp_ind2; intros v Hfree Heval.
+ inv Heval. constructor. not_free. now rewrite PM.gso.
+ inv Heval. now constructor.
+ inv Heval. constructor.
+ inv Heval. constructor 4 with cs; trivial.
  not_free.
  assert (Hrec : Nelist.Forall (fun e v,
          exp_eval mem env e vexp_eval mem (PM.add x env) e v) es).
  { rewrite Nelist.Forall_forall in IHes, Hfree |- ×. intros. apply IHes; trivial. now apply Hfree. }
  clear IHes Hfree H3. revert cs H1.
  induction es; intros cs Hes; inv Hrec; inv Hes; constructor; auto.
Qed.

If we add irrelevent values to mem, evaluation does not change.
Lemma exp_eval_extend_mem : mem env x e v,
  ¬Is_free_in_exp x eexp_eval mem env e vexp_eval (madd_mem x mem) env e v.
Proof.
intros mem env x e.
induction e using exp_ind2; intros v Hfree Heval.
+ inversion_clear Heval. now constructor.
+ inversion_clear Heval. constructor. not_free. now rewrite mfind_mem_gso.
+ inversion_clear Heval. constructor.
+ inversion_clear Heval. constructor 4 with cs; trivial.
  not_free.
  assert (Hrec : Nelist.Forall (fun e v,
          exp_eval mem env e vexp_eval (madd_mem x mem) env e v) es).
  { rewrite Nelist.Forall_forall in IHes, Hfree |- ×. intros. apply IHes; trivial. now apply Hfree. }
  clear IHes Hfree H0. revert cs H.
  induction es; intros cs Hes; inv Hrec; inv Hes; constructor; auto.
Qed.

If we add objcets to mem, evaluation does not change.
Lemma exp_eval_extend_mem_by_obj : mem env f obj e v,
  exp_eval mem env e vexp_eval (madd_obj f obj mem) env e v.
Proof.
intros mem env f e.
induction e using exp_ind2; intros v Heval.
+ inversion_clear Heval. now constructor.
+ inversion_clear Heval. constructor. now rewrite mfind_mem_add_inst.
+ inversion_clear Heval. constructor.
+ inversion_clear Heval. constructor 4 with cs; trivial.
  assert (Hrec : Nelist.Forall (fun e v,
          exp_eval mem env e vexp_eval (madd_obj f mem) env e v) es).
  { rewrite Nelist.Forall_forall in IHes |- ×. intros. now apply IHes. }
  clear IHes H0. revert cs H.
  induction es; intros cs Hes; inv Hrec; inv Hes; constructor; auto.
Qed.

Lemma cannot_write_exp_eval:
   prog s menv env menv´ env´ e v,
    ( x, Is_free_in_exp x e¬ Can_write_in x s)
    → exp_eval menv env e v
    → stmt_eval prog menv env s (menv´, env´)
    → exp_eval menv´ env´ e v.
Proof.
  Hint Constructors Is_free_in_exp Can_write_in exp_eval.
  induction s; intros menv env menv´ env´ v Hfree Hexp Hstmt.
  - inv Hstmt.
    apply exp_eval_extend_env; trivial.
    intro Habs. apply (Hfree i); auto.
  - inv Hstmt.
    apply exp_eval_extend_mem; trivial.
    intro Habs. apply (Hfree i); auto.
  - inv Hstmt; solve [eapply IHs1; eassumption || cannot_write | eapply IHs2; eassumption || cannot_write].
  - inv Hstmt.
    apply exp_eval_extend_env, exp_eval_extend_mem_by_obj; trivial.
    intro Habs. apply (Hfree i); auto.
  - inv Hstmt.
    now apply exp_eval_extend_mem_by_obj.
  - inv Hstmt.
    match goal with
    | Hs1: stmt_eval _ _ _ s1 _,
      Hs2: stmt_eval _ _ _ s2 _ |- _
      ⇒ apply IHs1 with (2:=Hexp) in Hs1;
         [apply IHs2 with (2:=Hs1) (3:=Hs2)|];
         now cannot_write
    end.
  - now inv Hstmt.
Qed.

Lemma lift_Ifte:
   e s1 s2 t1 t2,
    ( x, Is_free_in_exp x e
               → (¬Can_write_in x s1 ¬Can_write_in x s2))
    → stmt_eval_eq (Comp (Ifte e s1 s2) (Ifte e t1 t2))
                    (Ifte e (Comp s1 t1) (Comp s2 t2)).
Proof.
  intros prog e s1 s2 t1 t2 menv env menv´ env´ Hfw.
  split; intro Hstmt.
  - inversion_clear Hstmt as [| | | |? ? ? ? ? env´´ menv´´ ? ? Hs Ht| | | ].
    inversion_clear Hs as [| | | | |? ? ? ? ? ? ? ? Hexp Hs1
                                   |? ? ? ? ? ? ? ? Hexp Hs2| ];
    inversion_clear Ht as [| | | | |? ? ? ? ? ? ? ? Hexp´ Ht1
                                   |? ? ? ? ? ? ? ? Hexp´ Ht2| ].
    + constructor; [now apply Hexp|].
      econstructor; [now apply Hs1|now apply Ht1].
    + apply cannot_write_exp_eval with (2:=Hexp) in Hs1; [|now cannot_write].
      apply exp_eval_det with (1:=Hexp´) in Hs1; discriminate.
    + apply cannot_write_exp_eval with (2:=Hexp) in Hs2; [|now cannot_write].
      apply exp_eval_det with (1:=Hexp´) in Hs2; discriminate.
    + constructor 7; [now apply Hexp|].
      econstructor; [now apply Hs2|now apply Ht2].
  - inversion_clear Hstmt as [| | | | |? ? ? ? ? ? ? ? Hexp Hs|
                              ? ? ? ? ? ? ? ? Hexp Hs|].
    + inversion_clear Hs
        as [| | | |? ? ? ? ? env´´ menv´´ ? ? Hs1 Ht1| | | ].
      apply Icomp with (menv1:=menv´´) (env1:=env´´).
      apply Iifte_true with (1:=Hexp) (2:=Hs1).
      apply cannot_write_exp_eval with (2:=Hexp) in Hs1; [|now cannot_write].
      apply Iifte_true with (1:=Hs1) (2:=Ht1).
    + inversion_clear Hs
        as [| | | |? ? ? ? ? env´´ menv´´ ? ? Hs2 Ht2| | | ].
      apply Icomp with (menv1:=menv´´) (env1:=env´´).
      apply Iifte_false with (1:=Hexp) (2:=Hs2).
      apply cannot_write_exp_eval with (2:=Hexp) in Hs2; [|now cannot_write].
      apply Iifte_false with (1:=Hs2) (2:=Ht2).
Qed.

Inductive Is_fusible : stmtProp :=
| IFWAssign: x e,
    Is_fusible (Assign x e)
| IFWAssignSt: x e,
    Is_fusible (AssignSt x e)
| IFWIfte: e s1 s2,
    Is_fusible s1
    Is_fusible s2
    ( x, Is_free_in_exp x e¬Can_write_in x s1 ¬Can_write_in x s2) →
    Is_fusible (Ifte e s1 s2)
| IFWStep_ap: x cls obj e,
    Is_fusible (Step_ap x cls obj e)
| IFWComp: s1 s2,
    Is_fusible s1
    Is_fusible s2
    Is_fusible (Comp s1 s2)
| IFWSkip:
    Is_fusible Skip.

Lemma lift_Is_fusible:
   e s1 s2 t1 t2,
    Is_fusible (Comp (Ifte e s1 s2) (Ifte e t1 t2))
    
    Is_fusible (Ifte e (Comp s1 t1) (Comp s2 t2)).
Proof.
  Hint Constructors Is_fusible.
  intros e s1 s2 t1 t2.
  split; intro Hifw.
  - inversion_clear Hifw as [| | | |? ? Hs Ht|].
    constructor; inversion_clear Hs; inversion_clear Ht; now cannot_write.
  - inversion_clear Hifw as [| |? ? ? Hfree1 Hfree2 Hcw| | | ].
    inversion_clear Hfree1; inversion_clear Hfree2.
    repeat constructor; cannot_write.
Qed.

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

Lemma Is_fusible_fold_left_shift:
   A f (xs : list A) iacc,
    Is_fusible (fold_left (fun i xComp (f x) i) xs iacc)
    
    (Is_fusible (fold_left (fun i xComp (f x) i) xs Skip)
      Is_fusible iacc).
Proof.
  Hint Constructors Is_fusible.
  induction xs as [|x xs IH]; [now intuition|].
  intros; split.
  - intro HH. simpl in HH. apply IH in HH.
    destruct HH as [Hxs Hiacc].
    split; [simpl; rewrite IH|];
    repeat progress match goal with
    | H:Is_fusible (Comp _ _) |- _inversion_clear H
    | _now intuition
    end.
  - intros [HH0 HH1].
    simpl in HH0; rewrite IH in HH0.
    simpl; rewrite IH.
    intuition.
    repeat progress match goal with
    | H:Is_fusible (Comp _ _) |- _inversion_clear H
    | _now intuition
    end.
Qed.


Fixpoint zip s1 s2 : stmt :=
  match s1, s2 with
  | Ifte e1 t1 f1, Ifte e2 t2 f2
      if exp_eqb e1 e2
      then Ifte e1 (zip t1 t2) (zip f1 f2)
      else Comp s1 s2
  | Skip, ss
  | s, Skips
  | Comp s1´ s2´, _Comp s1´ (zip s2´ s2)
  | s1, s2Comp s1 s2
  end.

Lemma Can_write_in_zip:
   s1 s2 x,
    (Can_write_in x s1 Can_write_in x s2)
     Can_write_in x (zip s1 s2).
Proof.
  Hint Constructors Can_write_in.
  induction s1, s2; simpl;
    repeat progress
           match goal with
           | H:Can_write_in _ (Comp _ _) |- _inversion H; subst; clear H
           | H:Can_write_in _ (Ifte _ _ _) |- _inversion H; subst; clear H
           | H:Can_write_in _ Skip |- _now inversion H
           | H:Can_write_in _ _ Can_write_in _ _ |- _destruct H
           | |- context [exp_eqb ?e1 ?e2] ⇒
             destruct (exp_eqb e1 e2) eqn:Heq;
               [apply exp_eqb_eq in Heq; subst|]
           | |- Can_write_in _ (Ifte _ _ _) ⇒
             (apply CWIIfteTrue; apply IHs1_1; now intuition)
             || (apply CWIIfteFalse; apply IHs1_2; now intuition)
           | H:Can_write_in _ (zip _ _) |- _
             apply IHs1_1 in H || apply IHs1_2 in H
           | |- Can_write_in _ (Comp _ (zip _ _)) ⇒
             now (apply CWIComp2; apply IHs1_2; intuition)
           | _intuition
           end.
Qed.

Lemma Cannot_write_in_zip:
   s1 s2 x,
    (¬Can_write_in x s1 ¬Can_write_in x s2)
     ¬Can_write_in x (zip s1 s2).
Proof.
  intros s1 s2 x.
  split; intro HH.
  - intro Hcan; apply Can_write_in_zip in Hcan; intuition.
  - split; intro Hcan; apply HH; apply Can_write_in_zip; intuition.
Qed.

Lemma zip_free_write:
   s1 s2,
    Is_fusible s1
    → Is_fusible s2
    → Is_fusible (zip s1 s2).
Proof.
  Hint Constructors Is_fusible Can_write_in.
  induction s1, s2;
    intros Hfree1 Hfree2;
    inversion_clear Hfree1;
    simpl;
    try now intuition.
  match goal with
  | |- context [exp_eqb ?e1 ?e2] ⇒ destruct (exp_eqb e1 e2) eqn:Heq
  end; [|intuition].
  apply exp_eqb_eq in Heq; subst.
  inversion_clear Hfree2;
  constructor;
    repeat progress
           match goal with
           | H1:Is_fusible ?s1,
             H2:Is_fusible ?s2,
             Hi:context [Is_fusible (zip _ _)]
             |- Is_fusible (zip ?s1 ?s2)
             ⇒ apply Hi with (1:=H1) (2:=H2)
           | |- x, Is_free_in_exp x ?e_
             ⇒ intros x Hfree
           | H1:context [Is_free_in_exp _ ?e_ _],
                H2:Is_free_in_exp _ ?e |- _
             ⇒ specialize (H1 _ H2)
           | |- _ _split
           | |- ¬Can_write_in _ (zip _ _)
             ⇒ apply Cannot_write_in_zip; intuition
           | _idtac
           end.
Qed.

Lemma zip_Comp´:
   s1 s2,
    Is_fusible s1
    → (stmt_eval_eq (zip s1 s2) (Comp s1 s2)).
Proof.
  induction s1, s2;
  try rewrite stmt_eval_eq_Comp_Skip1;
  try rewrite stmt_eval_eq_Comp_Skip2;
  try reflexivity;
  inversion_clear 1;
  simpl;
  repeat progress
         match goal with
         | |- context [exp_eqb ?e1 ?e2]
           ⇒ destruct (exp_eqb e1 e2) eqn:Heq;
             [apply exp_eqb_eq in Heq; subst|]
         | H:Is_fusible ?s1,
             IH:context [stmt_eval_eq (zip ?s1 _) _]
           |- context [zip ?s1 ?s2]
           ⇒ rewrite IH with (1:=H)
         end;
  try (rewrite lift_Ifte; [|assumption]);
  try rewrite Comp_assoc;
  reflexivity.
Qed.

Fixpoint fuse´ s1 s2 : stmt :=
  match s1, s2 with
  | s1, Comp s2 s3fuse´ (zip s1 s2) s3
  | s1, s2zip s1 s2
  end.

Definition fuse s : stmt :=
  match s with
  | Comp s1 s2fuse´ s1 s2
  | _s
  end.

Lemma fuse´_free_write:
   s2 s1,
    Is_fusible s1
    → Is_fusible s2
    → Is_fusible (fuse´ s1 s2).
Proof.
  induction s2;
    try (intros; apply zip_free_write; assumption).
  intros s1 Hfree1 Hfree2.
  inversion_clear Hfree2.
  apply IHs2_2; [|assumption].
  apply zip_free_write; assumption.
Qed.

fuse_eval_eq

Require Import Relations.
Require Import Morphisms.
Require Import Setoid.

Definition fuse_eval_eq s1 s2: Prop :=
  stmt_eval_eq s1 s2 (Is_fusible s1 Is_fusible s2).

Lemma fuse_eval_eq_refl:
   s,
    Is_fusible s
    → Proper fuse_eval_eq s.
Proof.
  intros s Hfree; unfold Proper, fuse_eval_eq; intuition.
Qed.

Lemma fuse_eval_eq_trans:
  transitive stmt fuse_eval_eq.
Proof.
  intros s1 s2 s3 Heq1 Heq2.
  unfold fuse_eval_eq in ×.
  split; [|now intuition].
  destruct Heq1 as [Heq1 ?].
  destruct Heq2 as [Heq2 ?].
  rewrite Heq1, Heq2; reflexivity.
Qed.

Lemma fuse_eval_eq_sym:
  symmetric stmt fuse_eval_eq.
Proof.
  intros s1 s2 Heq.
  unfold fuse_eval_eq in ×.
  split; [|now intuition].
  destruct Heq as [Heq ?].
  rewrite Heq; reflexivity.
Qed.

Add Relation stmt (fuse_eval_eq)
    symmetry proved by fuse_eval_eq_sym
    transitivity proved by fuse_eval_eq_trans
  as fuse_eval_equiv.

Instance subrelation_stmt_fuse_eval_eq:
  subrelation fuse_eval_eq stmt_eval_eq.
Proof.
  intros s1 s2 Heq x menv env menv´ env´.
  now apply Heq.
Qed.

Lemma zip_Comp:
   s1 s2,
    Is_fusible s1
    → Is_fusible s2
    → fuse_eval_eq (zip s1 s2) (Comp s1 s2).
Proof.
  intros s1 s2 Hfree1 Hfree2.
  unfold fuse_eval_eq.
  split; [|split].
  - rewrite zip_Comp´ with (1:=Hfree1); reflexivity.
  - apply zip_free_write with (1:=Hfree1) (2:=Hfree2).
  - intuition.
Qed.

Instance fuse_eval_eq_Proper:
    Proper (eq ==> eq ==> eq ==> fuse_eval_eq ==> eq ==> iff) stmt_eval.
Proof.
  intros prog´ prog HR1 menv´ menv HR2 env´ env HR3 s1 s2 Heq r HR4;
  subst; destruct r as [menv´ env´].
  unfold fuse_eval_eq in Heq.
  destruct Heq as [Heq ?].
  rewrite Heq; reflexivity.
Qed.

Instance fuse_eval_eq_Comp_Proper:
  Proper (fuse_eval_eq ==> fuse_eval_eq ==> fuse_eval_eq) Comp.
Proof.
  Hint Constructors Is_fusible.
  intros s Hseq t Hteq.
  unfold fuse_eval_eq in ×.
  destruct Hseq as [Hseq [Hfrees Hfrees´]].
  destruct Hteq as [Hteq [Hfreet Hfreet´]].
  split; [|intuition].
  rewrite Hseq, Hteq; reflexivity.
Qed.

Instance zip_fuse_eval_eq_Proper:
  Proper (fuse_eval_eq ==> fuse_eval_eq ==> fuse_eval_eq) zip.
Proof.
  intros s Hseq t Hteq.
  unfold fuse_eval_eq in ×.
  destruct Hseq as [Hseq [Hfrees Hfrees´]].
  destruct Hteq as [Hteq [Hfreet Hfreet´]].
  split; [|split]; [|apply zip_free_write with (1:=Hfrees) (2:=Hfreet)
                    |apply zip_free_write with (1:=Hfrees´) (2:=Hfreet´)].
  rewrite zip_Comp´ with (1:=Hfrees).
  rewrite zip_Comp´ with (1:=Hfrees´).
  rewrite Hseq, Hteq; reflexivity.
Qed.

Lemma fuse´_Comp:
   s2 s1,
    Is_fusible s1
    → Is_fusible s2
    → stmt_eval_eq (fuse´ s1 s2) (Comp s1 s2).
Proof.
  Hint Constructors Is_fusible.
  induction s2;
  intros s1 Hifte1 Hifte2; simpl;
  try now (rewrite zip_Comp´; intuition).
  rewrite Comp_assoc.
  inversion_clear Hifte2.
  rewrite IHs2_2;
    match goal with
    | H1:Is_fusible ?s1,
      H2:Is_fusible ?s2 |- Is_fusible (zip ?s1 ?s2)
      ⇒ apply zip_free_write with (1:=H1) (2:=H2)
    | |- Is_fusible ?sassumption
    | H:Is_fusible s2_2 |- _pose proof (fuse_eval_eq_refl _ H)
    end.
  intros prog menv env menv´ env´.
  rewrite zip_Comp; [now apply iff_refl|assumption|assumption].
Qed.

Instance fuse´_fuse_eval_eq_Proper:
  Proper (fuse_eval_eq ==> fuse_eval_eq ==> fuse_eval_eq) fuse´.
Proof.
  intros s Hseq t Hteq.
  unfold fuse_eval_eq in ×.
  destruct Hseq as [Hseq [Hfrees Hfrees´]].
  destruct Hteq as [Hteq [Hfreet Hfreet´]].
  split; [|split]; [|apply fuse´_free_write with (1:=Hfrees) (2:=Hfreet)
                    |apply fuse´_free_write with (1:=Hfrees´) (2:=Hfreet´)].
  repeat rewrite fuse´_Comp; try assumption.
  rewrite Hseq, Hteq.
  reflexivity.
Qed.

Lemma fuse_Comp:
   s,
    Is_fusible s
    → stmt_eval_eq (fuse s) s.
Proof.
  intros s Hfree prog menv env menv´ env´.
  destruct s; simpl; try reflexivity.
  inversion_clear Hfree.
  destruct s2;
  match goal with
  | H: Is_fusible ?s2 |- context [fuse´ _ ?s2]
    ⇒ pose proof (fuse_eval_eq_refl _ H)
  end;
  rewrite fuse´_Comp; auto; reflexivity.
Qed.