Library Rustre.Example

Require Import PArith.
Require Import Rustre.Nelist.
Require Import List.
Import List.ListNotations.
Open Scope positive.
Open Scope list.

Require Import Rustre.Common.
Require Import Rustre.Dataflow.Syntax.
Require Import Rustre.Minimp.Syntax.
Require Import Rustre.Translation.
Require Import Rustre.Minimp.FuseIfte.

Require Import Rustre.Dataflow.Memories.
Require Import Rustre.Dataflow.WellFormed.
Require Import Rustre.Dataflow.WellFormed.Decide.

Require Import Rustre.Minimp.FuseIfte.

Class Assignment U V := {assign : identUV}.
Notation "x ´::=´ y" := (assign x y) (at level 47, no associativity, only parsing).
Class OpCall T U := {opcall : operatorTU}.
Notation "f ´<´ nel ´>´" := (opcall f nel) (at level 46, only parsing, format "f '<' nel '>'").
Notation "x :,: y" := (necons x y) (at level 30, right associativity, only parsing).
Notation "x ´§´" := (nebase x) (at level 30).

Coercion Cint : BinInt.Z >-> const.
Coercion Cbool : bool >-> const.
Coercion Evar : ident >-> lexp.
Coercion Eexp : lexp >-> cexp.
Coercion Econst : const >-> lexp.

Notation "x ´on´ ck" := (Con x ck) (at level 44).
Notation "x ´when´ ck" := (Ewhen x ck true) (at level 45, left associativity).
Notation "x ´whenot´ ck" := (Ewhen x ck false) (at level 45, left associativity).
Notation "x ´::=´ v ´fby´ y" := (EqFby x Cbase (v : const) (y : lexp))
                                  (at level 47, v at next level).
Notation "x ´::=´ f ´(|´ nel ´|)´" := (EqApp x Cbase f nel)
                                        (at level 47, f, nel at next level, format "x '::=' f '(|' nel '|)'").
Notation "x ´::=´ f ´(|´ nel ´|)´ ´@´ ck" := (EqApp x ck f nel)
                                        (at level 47, f, nel at next level).

Instance EqDef_CAssign : Assignment cexp equation := {assign := (fun iEqDef i Cbase)}.
Instance EqDef_EAssign : Assignment lexp equation := {assign := (fun i eEqDef i Cbase (Eexp e))}.
Instance Eop_OpCall : OpCall lexps lexp := {opcall := Eop}.

Coercion Const : const >-> exp.
Coercion Var : ident >-> exp.

Class IFTE U V := {ifte : UVVV}.
Notation "´If´ b ´Then´ t ´Else´ f" := (ifte b t f) (at level 47, t at level 47, f at level 47).

Instance Assign_Assign : Assignment exp stmt := {assign := Assign}.
Notation "´state(|´ x ´|)::=´ y" := (AssignSt x y) (at level 47).
Instance Op_OpCall : OpCall (nelist exp) exp := {opcall := Op}.
Notation "stmt1 ;; stmt2" := (Comp stmt1 stmt2) (at level 48, right associativity).
Instance IFTE_imp : IFTE exp stmt := {ifte := Ifte}.
Instance IFTE_impVar : IFTE ident stmt := {ifte := fun iIfte (Var i) }.

Notation "x ´::=´ class ´(|´ obj ´|).step(´ args ´)´" := (Step_ap x class obj args)
    (at level 47, class, obj, args at next level).
Notation " class ´(|´ obj ´|).reset()´" := (Reset_ap class obj)
                                                 (at level 47, obj at next level).


Open Scope Z_scope.

Examples from paper

Section CodegenPaper.

  Require Import Nelist.

  Definition Plus : operator := existT arrows (Tcons Tint (Tcons Tint (Tout Tint))) BinInt.Z.add.
  Opaque Plus.

  Class NPlus U := {plus : UUU}.
  Notation "x ´:+´ y" := (plus x y) (at level 47).

  Instance NPlus_lexp : NPlus lexp := {plus := fun x yEop Plus (necons x (nebase y))}.
  Instance NPlus_exp : NPlus exp := {plus := fun x yOp Plus (necons x (nebase y))}.

  Definition Ifte_int : operator :=
    existT arrows (Tcons Tbool (Tcons Tint (Tcons Tint (Tout Tint))))
             (fun (x : bool) t fif x then t else f).
  Definition op_ifte (x: lexp) (t: lexp) (f: lexp) : lexp :=
    Eop Ifte_int (necons x (necons t (nebase f))).
  Opaque Ifte_int.

  Instance IFTE_lustre : IFTE lexp lexp := {ifte := op_ifte}.
  Instance IFTE_impOp : IFTE exp exp := {ifte := fun x t fOp Ifte_int (necons x (necons t (nebase f)))}.

  Definition Disj : operator :=
    existT arrows (Tcons Tbool (Tcons Tbool (Tout Tbool))) orb.
  Opaque Disj.

  Class NDisj U := {disj : UUU}.
  Notation "x ´:||´ y" := (disj x y) (at level 47).

  Instance NDisj_lexp : NDisj lexp := {disj := fun x yEop Disj (necons x (nebase y))}.
  Instance NDisj_exp : NDisj exp := {disj := fun x yOp Disj (necons x (nebase y))}.

  Definition n_count : ident := 1%positive.
  Definition n_avgvelocity : ident := n_count + 1.


  Definition ini : ident := 1%positive.
  Definition inc : ident := 2%positive.
  Definition restart : ident := 3%positive.
  Definition n : ident := 4%positive.
  Definition f : ident := 5%positive.
  Definition c : ident := 6%positive.

  Example count_eqns : list equation :=
    [
      c ::= 0 fby n ;
      f ::= true fby false;
      n ::= If ((f : lexp) :|| restart)
            Then (ini : lexp)
            Else ((c : lexp) :+ inc)
    ].

  Definition Div : operator := existT arrows (Tcons Tint (Tcons Tint (Tout Tint))) BinInt.Z.div.
  Opaque Div.

  Class NDiv U := {div : UUU}.
  Notation "x ´:/´ y" := (div x y) (at level 47).

  Instance NDiv_lexp : NDiv lexp := {div := fun x yEop Div (necons x (nebase y))}.
  Instance NDiv_imp : NDiv exp := {div := fun x yOp Div (necons x (nebase y))}.

  Lemma count_eqns_well_sch :
    Is_well_sch (PS.add f (PS.singleton c))
                (ini :,: inc :,: restart§) count_eqns.
  Proof. apply Is_well_sch_by_refl; reflexivity. Qed.

  Example count : node :=
    mk_node n_count (necons ini (necons inc (nebase restart))) n count_eqns.

  Example count_prog :=
    {|
      c_name := n_count;
      c_input := ini :,: inc :,: restart §;
      c_output := n;
      c_mems := [f; c];
      c_objs := [];
      c_step := n ::=
                  (If (State f) :|| restart
                  Then (ini : exp)
                  Else ((State c) :+ inc));;
                state(| f |)::= false;;
                state(| c |)::= n;;
                Skip;
      c_reset := state(| f |)::= true;;
                 state(| c |)::= 0;;
                 Skip
    |}.

  Remark count_prog_good: translate_node count = count_prog.
  Proof eq_refl.

  Remark count_prog_step_fuse:
    fuse (c_step count_prog) =
         n ::= (If (State f) :|| restart
                Then (ini : exp)
                Else ((State c) :+ inc));;
         state(| f |)::= false;;
         state(| c |)::= n.
  Proof eq_refl.

  Remark count_prog_reset_fuse:
    fuse (c_reset count_prog) = state(| f |)::= true;; state(| c |)::= 0.
  Proof eq_refl.


  Definition delta : ident := 1%positive.
  Definition sec : ident := 2%positive.
  Definition r : ident := 3%positive.
  Definition t : ident := 4%positive.
  Definition v : ident := 5%positive.
  Definition h : ident := 6%positive.

  Example avgvelocity_eqns : list equation :=
    [
      h ::= 0 fby v;
      v ::= (Emerge sec
                    (Ewhen r sec true :/ t)
                    (Ewhen h sec false));
      t ::= n_count (| (Ewhen 1 sec true) :,:
                       (Ewhen 1 sec true) :,:
                       (Ewhen false sec true) § |)
        @ (Con Cbase sec true);
      r ::= n_count (| (0 : lexp) :,: (delta : lexp) :,: (false : lexp) § |)
    ].

  Lemma avgvelocity_eqns_Well_sch :
    Is_well_sch (PS.singleton h) (delta :,: sec§) avgvelocity_eqns.
  Proof. apply Is_well_sch_by_refl; reflexivity. Qed.

  Example avgvelocity : node :=
    mk_node n_avgvelocity (delta :,: sec§) v avgvelocity_eqns.

  Example avgvelocity_prog :=
    {|
      c_name := n_avgvelocity;
      c_input := delta :,: sec §;
      c_output := v;
      c_mems := [h];
      c_objs := [{| obj_inst := r; obj_class := n_count |};
                 {| obj_inst := t; obj_class := n_count |}];
      c_step := (r ::= n_count (| r |).step((0 : exp) :,:
                                            (delta : exp) :,:
                                            (false : exp) §));;
                If sec
                Then (t ::= n_count (| t |).step((1 : exp) :,:
                                                 (1 : exp) :,:
                                                 (false : exp) §))
                Else Skip;;
                If sec
                Then v ::= ((r : exp) :/ (t : exp))
                 Else v ::= (State h);;
                state(| h |)::= v;;
                Skip;
      c_reset := n_count (| r |).reset() ;;
                 n_count (| t |).reset();;
                 state(| h |)::= 0;;
                 Skip
    |}.

  Remark avgvelocity_prog_good: translate_node avgvelocity = avgvelocity_prog.
  Proof eq_refl.

  Remark avgvelocity_prog_step_fuse:
    fuse (c_step avgvelocity_prog) =
         r ::= n_count (| r |).step((0 : exp) :,:
                                    (delta : exp) :,:
                                    (false : exp) §);;
         If sec
         Then (t ::= n_count (| t |).step((1 : exp) :,:
                                          (1 : exp) :,:
                                          (false : exp) §);;
              v ::= ((r : exp) :/ (t : exp)))
         Else v ::= State h;;
         state(| h |)::= v.
  Proof eq_refl.

  Remark avgvelocity_prog_reset_fuse:
    fuse (c_reset avgvelocity_prog) = n_count (| r |).reset();;
                                      n_count (| t |).reset();;
                                      state(| h |)::= 0.
  Proof eq_refl.

End CodegenPaper.