Library Rustre.Dataflow.WellFormed

Require Import Rustre.Common.
Require Import Rustre.Dataflow.Syntax.
Require Import Rustre.Dataflow.IsFree.
Require Import Rustre.Dataflow.IsVariable.
Require Import Rustre.Dataflow.IsDefined.
Require Import Rustre.Dataflow.Memories.
Require Import Rustre.Dataflow.Ordered.
Require Import Rustre.Dataflow.NoDup.

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

Well formed CoreDF programs


Section IsWellSch.

Variable memories : PS.t.
Variable arg: Nelist.nelist ident.

A list of equations is well scheduled if
  • the free stack variables (¬PS.In _ memories) are defined before (i.e. in eqs), or they belong to the input argument
  • the free memory variables (PS.In _ memories) have not been re-assigned before (i.e. in eqs)
  • the variable defined by an equation must be defined for the first time
Remark: we assume that the list of equations is in reverse order: the first equation to execute should be the last in the list.
Inductive Is_well_sch : list equationProp :=
| WSchNil: Is_well_sch nil
| WSchEq: eq eqs,
    Is_well_sch eqs
    ( i, Is_free_in_eq i eq
          (PS.In i memories¬Is_defined_in_eqs i eqs)
        (¬PS.In i memoriesIs_variable_in_eqs i eqs Nelist.In i arg)) →
    ( i, Is_defined_in_eq i eq¬Is_defined_in_eqs i eqs) →
    Is_well_sch (eq :: eqs).

End IsWellSch.

A CoreDF program is well defined if
  • Each node is well-defined, that is
    • Each input arguments' name is distinct
    • The equations are well scheduled
    • The input variables are not redefined by the equations
    • The output variable is indeed defined by the equations
  • Every node call points to a previously-defined node
  • Each of the nodes' name is distinct

Inductive Welldef_global : list nodeProp :=
| WGnil:
    Welldef_global []
| WGcons:
     nd nds,
      Welldef_global nds
      let eqs := nd.(n_eqs) in
      let ni := nd.(n_input) in
      let no := nd.(n_output) in
        NoDup (Nelist.nelist2list ni)
      → Is_well_sch (memories eqs) ni eqs
      → ¬ List.Exists (fun niIs_defined_in_eqs ni eqs) (Nelist.nelist2list ni)
      → Is_variable_in_eqs no eqs
      → ¬Is_node_in nd.(n_name) eqs
      → ( f, Is_node_in f eqsfind_node f nds None)
      → List.Forall (fun nd´nd.(n_name) nd´.(n_name)) nds
      → Welldef_global (nd::nds).

Properties of Is_well_sch


Hint Constructors Is_well_sch.

Lemma Is_well_sch_NoDup_defs:
   mems argIn eqs,
    Is_well_sch mems argIn eqs
    → NoDup_defs eqs.
Proof.
  induction eqs as [|eq eqs IH]; [now constructor|].
  inversion_clear 1; destruct eq; constructor; try apply IH; auto.
Qed.

Lemma Is_well_sch_cons:
   m argIn eq eqs, Is_well_sch m argIn (eq :: eqs) → Is_well_sch m argIn eqs.
Proof. inversion 1; auto. Qed.

Lemma Is_well_sch_free_variable:
   argIn x eq eqs mems,
    Is_well_sch mems argIn (eq :: eqs)
    → Is_free_in_eq x eq
    → ¬ PS.In x mems
    → Is_variable_in_eqs x eqs Nelist.In x argIn.
Proof.
  intros argIn x eq eqs mems Hwsch Hfree Hnim.
  destruct eq;
    inversion_clear Hwsch;
    inversion_clear Hfree as [? ? ? ? Hc | ? ? ? ? ? Hc | ? ? ? ? ? Hc];
    subst; intuition;
    match goal with
      | H: context[ Is_variable_in_eqs _ _ Nelist.In _ _] |- _
        eapply H; eauto
    end.
Qed.

Lemma Is_well_sch_free_variable_in_mems:
   argIn y eq eqs mems,
    Is_well_sch mems argIn (eq :: eqs)
    → Is_free_in_eq y eq
    → PS.In y mems
    → ¬Is_defined_in_eqs y eqs.
Proof.
  intros argIn x eq eqs mems Hwsch Hfree Hnim.
  destruct eq;
    inversion_clear Hwsch;
    inversion_clear Hfree as [? ? ? ? Hc | ? ? ? ? ? Hc | ? ? ? ? ? Hc];
    match goal with
      | H: context[ Nelist.In _ _ ] |- _
        eapply H
    end;
    auto.
Qed.

Lemma Is_wsch_is_defined_in:
   x eq eqs mems argIn,
    Is_well_sch mems argIn (eq :: eqs) →
    Is_defined_in_eqs x (eq :: eqs) →
    Is_defined_in_eq x eq
     (¬Is_defined_in_eq x eq Is_defined_in_eqs x eqs).
Proof.
  intros x eq eqs mems argIn Hwsch Hdef.
  apply List.Exists_cons in Hdef.
  destruct (Is_defined_in_eq_dec x eq); intuition.
Qed.

Properties of Welldef_global


Lemma Welldef_global_cons:
   node G,
    Welldef_global (node::G) → Welldef_global G.
Proof.
  inversion 1; assumption.
Qed.

Lemma Welldef_global_Ordered_nodes:
   G, Welldef_global GOrdered_nodes G.
Proof.
  induction G as [|node G IH]; [constructor|].
  intro Hwdef.
  constructor.
  - apply IH; apply Welldef_global_cons with (1:=Hwdef).
  - intros f Hini.
    inversion Hwdef.
    split; [ intro; subst | ];
    repeat match goal with
           | eqs:=n_eqs node |- _unfold eqs in *; clear eqs
           | H1:¬Is_node_in _ _, H2:Is_node_in _ _ |- Falsecontradiction
           | H1: Is_node_in _ _,
             H2: context[Is_node_in _ _find_node _ _ None] |- _
             apply H2 in H1; apply find_node_Exists in H1; exact H1
           end.
  - inversion Hwdef; assumption.
Qed.

Lemma Welldef_global_app:
   G , Welldef_global (G ++ ) → Welldef_global .
Proof.
  intros G Hwdef.
  induction G as [|g G IH]; [now apply Hwdef|].
  rewrite <- List.app_comm_cons in Hwdef.
  apply Welldef_global_cons in Hwdef.
  apply IH.
  apply Hwdef.
Qed.

Lemma Welldef_global_input_not_Is_defined_in:
   f G fnode,
    Welldef_global G
    → find_node f G = Some fnode
    → ¬ Nelist.Exists (fun niIs_defined_in_eqs ni fnode.(n_eqs)) fnode.(n_input).
Proof.
  induction G as [|node G IH]; [inversion_clear 2|].
  intros fnode HWdef Hfnode.
  apply find_node_split in Hfnode.
  destruct Hfnode as [bG [aG HnG]].
  rewrite HnG in HWdef; clear HnG.
  apply Welldef_global_app in HWdef.
  inversion_clear HWdef.
  now rewrite <- Nelist.nelist2list_Exists.
Qed.

Lemma Welldef_global_output_Is_variable_in:
   f G fnode,
    Welldef_global G
    → find_node f G = Some fnode
    → Is_variable_in_eqs fnode.(n_output) fnode.(n_eqs).
Proof.
  induction G as [|node G IH]; [inversion_clear 2|].
  intros fnode HWdef Hfnode.
  apply find_node_split in Hfnode.
  destruct Hfnode as [bG [aG HnG]].
  rewrite HnG in HWdef; clear HnG.
  apply Welldef_global_app in HWdef.
  inversion_clear HWdef.
  assumption.
Qed.