Library Rustre.Dataflow.Semantics

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

Require Import Coq.FSets.FMapPositive.
Require Import Rustre.Common.
Require Import Rustre.Dataflow.Syntax.
Require Import Rustre.Dataflow.Ordered.
Require Import Rustre.Dataflow.Stream.

The CoreDF semantics

We provide a "standard" dataflow semantics relating an environment of streams to a stream of outputs.

Environment and history

An history maps variables to streams of values (the variables' histories). Taking a snapshot of the history at a given time yields an environment.

Definition R := PM.t value.
Definition history := PM.t (stream value).

Implicit Type R: R.
Implicit Type H: history.

Instantaneous semantics


Section InstantSemantics.

Variable base : bool.
Variable R: R.

Inductive sem_var_instant (x: ident) v: Prop :=
| Sv:
      PM.find x R = Some v
      sem_var_instant x v.

Inductive sem_clock_instant: clockboolProp :=
| Sbase:
      sem_clock_instant Cbase base
| Son_tick:
     ck x c,
      sem_clock_instant ck true
      sem_var_instant x (present (Cbool c)) →
      sem_clock_instant (Con ck x c) true
| Son_abs1:
     ck x c,
      sem_clock_instant ck false
      sem_clock_instant (Con ck x c) false
| Son_abs2:
     ck x c ,
      sem_clock_instant ck true
      sem_var_instant x (present (Cbool )) →
      ¬ (c = )
      sem_clock_instant (Con ck x c) false.

Inductive sem_lexp_instant: lexpvalueProp:=
| Sconst:
     c v,
      v = (if base then present c else absent)
      sem_lexp_instant (Econst c) v
| Svar:
     x v,
      sem_var_instant x v
      sem_lexp_instant (Evar x) v
| Swhen_eq:
     s x b v,
      sem_var_instant x (present (Cbool b)) →
      sem_lexp_instant s v
      sem_lexp_instant (Ewhen s x b) v
| Swhen_abs1:
     s x b ,
      sem_var_instant x (present (Cbool )) →
      ¬ (b = )
      
      sem_lexp_instant (Ewhen s x b) absent
| Swhen_abs2:
     s x b,
      sem_var_instant x absent
      
      sem_lexp_instant (Ewhen s x b) absent
| Sop_eq: les op cs,
    Nelist.Forall2 sem_lexp_instant les (Nelist.map present cs) →
    Valid_args (get_arity op) cs
    sem_lexp_instant (Eop op les) (option2value (apply_op op cs))
| Sop_abs: les op,
    Nelist.Forall2 sem_lexp_instant les (alls absent les) →
    sem_lexp_instant (Eop op les) absent.

Definition sem_lexps_instant (les: nelist lexp)(vs: nelist value) :=
  Nelist.Forall2 sem_lexp_instant les vs.

Inductive sem_laexp_instant: clocklexpvalueProp:=
| SLtick:
     ck ce c,
      sem_lexp_instant ce (present c) →
      sem_clock_instant ck true
      sem_laexp_instant ck ce (present c)
| SLabs:
     ck ce,
      sem_clock_instant ck false
      sem_laexp_instant ck ce absent.

Inductive sem_laexps_instant: clocklexpsnelist valueProp:=
| SLticks:
     ck ces cs vs,
      vs = Nelist.map present cs
      sem_lexps_instant ces vs
      sem_clock_instant ck true
      sem_laexps_instant ck ces vs
| SLabss:
     ck ces vs,
      vs = Nelist.map (fun _absent) ces
      sem_clock_instant ck false
      sem_laexps_instant ck ces vs.

Inductive sem_cexp_instant: cexpvalueProp :=
| Smerge_true:
     x t f v,
      sem_var_instant x (present (Cbool true)) →
      sem_cexp_instant t v
      sem_cexp_instant (Emerge x t f) v
| Smerge_false:
     x t f v,
      sem_var_instant x (present (Cbool false)) →
      sem_cexp_instant f v
      sem_cexp_instant (Emerge x t f) v
| Smerge_abs:
     x t f,
      sem_var_instant x absent
      sem_cexp_instant (Emerge x t f) absent
| Sexp:
     e v,
      sem_lexp_instant e v
      sem_cexp_instant (Eexp e) v.

Inductive sem_caexp_instant: clockcexpvalueProp :=
| SCtick:
     ck ce c,
      sem_cexp_instant ce (present c) →
      sem_clock_instant ck true
      sem_caexp_instant ck ce (present c)
| SCabs:
     ck ce,
      sem_clock_instant ck false
      sem_caexp_instant ck ce absent.

Inductive rhs_absent_instant: equationProp :=
| AEqDef:
     x ck cae,
      sem_caexp_instant ck cae absent
      rhs_absent_instant (EqDef x ck cae)
| AEqApp:
     x f ck laes vs,
      sem_laexps_instant ck laes vs
      Nelist.Forall (fun cc = absent) vs
      rhs_absent_instant (EqApp x ck f laes)
| AEqFby:
     x ck v0 lae,
      sem_laexp_instant ck lae absent
      rhs_absent_instant (EqFby x ck v0 lae).

End InstantSemantics.

Liftings of instantaneous semantics


Section LiftSemantics.

Variable bk : stream bool.

Definition restr H (n: nat): R :=
  PM.map (fun xsxs n) H.
Hint Unfold restr.

Definition lift1 {A B} (f : AB) (s : stream A) : stream B := fun nf (s n).
Hint Unfold lift1.

Definition lift {A B} (sem: boolRABProp) H x (ys: stream B): Prop :=
   n, sem (bk n) (restr H n) x (ys n).
Hint Unfold lift.

Definition sem_clock H (ck: clock)(xs: stream bool): Prop :=
  lift sem_clock_instant H ck xs.

Definition sem_var H (x: ident)(xs: stream value): Prop :=
  lift (fun basesem_var_instant) H x xs.

Definition sem_vars H (x: nelist ident)(xs: stream (nelist value)): Prop :=
  lift (fun base RNelist.Forall2 (sem_var_instant R)) H x xs.

Definition sem_laexp H ck (e: lexp)(xs: stream value): Prop :=
  lift (fun base Rsem_laexp_instant base R ck) H e xs.

Definition sem_laexps H (ck: clock)(e: lexps)(xs: stream (nelist value)): Prop :=
  lift (fun base Rsem_laexps_instant base R ck) H e xs.

Definition sem_lexp H (e: lexp)(xs: stream value): Prop :=
  lift sem_lexp_instant H e xs.

Definition sem_lexps H (e: lexps)(xs: stream (nelist value)): Prop :=
  lift sem_lexps_instant H e xs.

Definition sem_caexp H ck (c: cexp)(xs: stream value): Prop :=
  lift (fun base Rsem_caexp_instant base R ck) H c xs.

Definition sem_cexp H (c: cexp)(xs: stream value): Prop :=
  lift sem_cexp_instant H c xs.

End LiftSemantics.

Time-dependent semantics


Definition absent_list (xss: stream (nelist value))(n: nat): Prop :=
  xss n = Nelist.map (fun _absent) (xss n).

Definition present_list (xss: stream (nelist value))(n: nat)(vs: nelist const): Prop :=
  xss n = Nelist.map present vs.

Definition clock_of (xss: stream (nelist value))(bs: stream bool): Prop :=
   n,
    ( vs, present_list xss n vs) bs n = true.

Inductive sem_equation G : stream boolhistoryequationProp :=
| SEqDef:
     bk H x xs ck ce,
      sem_var bk H x xs
      sem_caexp bk H ck ce xs
      sem_equation G bk H (EqDef x ck ce)
| SEqApp:
     bk H x ck f arg ls xs,
      sem_laexps bk H ck arg ls
      sem_var bk H x xs
      sem_node G f ls xs
      sem_equation G bk H (EqApp x ck f arg)
| SEqFby:
     bk H x ls xs v0 ck le,
      sem_laexp bk H ck le ls
      sem_var bk H x xs
      xs = fby v0 ls
      sem_equation G bk H (EqFby x ck v0 le)

with sem_node G: identstream (nelist value) → stream valueProp :=
| SNode:
     bk f xss ys i o eqs,
      clock_of xss bk
      find_node f G = Some (mk_node f i o eqs) →
      ( H,
           sem_vars bk H i xss
         sem_var bk H o ys
         ( n, absent_list xss n ys n = absent)
         List.Forall (sem_equation G bk H) eqs) →
      sem_node G f xss ys.

Definition sem_nodes (G: global) : Prop :=
  List.Forall (fun no xs ys, sem_node G no.(n_name) xs ys) G.

Induction principle for sem_node and sem_equation

The automagically-generated induction principle is not strong enough: it does not support the internal fixpoint introduced by List.Forall

Section sem_node_mult.
  Variable G: global.

  Variable P : bk H (eq: equation), sem_equation G bk H eqProp.
  Variable Pn : (f: ident) xss ys, sem_node G f xss ysProp.

  Hypothesis EqDef_case :
     (bk : stream bool)
           (H : history)
           (x : ident)
           (ck : clock)
           (ce : cexp)
           (xs : stream value)
           (Hvar : sem_var bk H x xs)
           (Hexp : sem_caexp bk H ck ce xs),
      P bk H (EqDef x ck ce) (SEqDef G bk H x xs ck ce Hvar Hexp).

  Hypothesis EqApp_case :
     (bk: stream bool)
           (H : history)
           (y : ident)
           (ck : clock)
           (f : ident)
           (les : lexps)
           (ls : stream (nelist value))
           (ys : stream value)
           (Hlaes : sem_laexps bk H ck les ls)
           (Hvar : sem_var bk H y ys)
           (Hnode : sem_node G f ls ys),
      Pn f ls ys Hnode
      P bk H (EqApp y ck f les) (SEqApp G bk H y ck f les ls ys Hlaes Hvar Hnode).

  Hypothesis EqFby_case :
     (bk: stream bool)
           (H : history)
           (y : ident)
           (ls : stream value)
           (yS : stream value)
           (v0 : const)
           (ck : clock)
           (lae : lexp)
           (Hls : sem_laexp bk H ck lae ls)
           (Hys : sem_var bk H y yS)
           (Hfby: yS = fby v0 ls),
      P bk H (EqFby y ck v0 lae) (SEqFby G bk H y ls yS v0 ck lae Hls Hys Hfby).

  Hypothesis SNode_case :
     (bk: stream bool)
           (f : ident)
           (xss : stream (nelist value))
           (ys : stream value)
           (i : nelist ident)
           (o : ident)
           (eqs : list equation)
           (Hck : clock_of xss bk)
           (Hf : find_node f G = Some (mk_node f i o eqs))
           (Heqs : H,
                        sem_vars bk H i xss
                      sem_var bk H o ys
                      ( n, absent_list xss n ys n = absent)
                      List.Forall (sem_equation G bk H) eqs),
      ( H,
             sem_vars bk H i xss
           sem_var bk H o ys
           ( n, absent_list xss n ys n = absent)
           List.Forall (fun eq Hsem, P bk H eq Hsem) eqs)
      → Pn f xss ys (SNode G bk f xss ys i o eqs Hck Hf Heqs).

  Fixpoint sem_equation_mult (bk: stream bool)
                             (H : history)
                             (eq : equation)
                             (Heq : sem_equation G bk H eq) {struct Heq}
                                                                : P bk H eq Heq :=
    match Heq in (sem_equation _ bk H eq) return (P bk H eq Heq) with
    | SEqDef bk H y xs ck ce Hvar HexpEqDef_case bk H y ck ce xs Hvar Hexp
    | SEqApp bk H y ck f lae ls ys Hlae Hvar Hnode
      EqApp_case bk H y ck f lae ls ys Hlae Hvar Hnode
                 (sem_node_mult f ls ys Hnode)
    | SEqFby bk H y ls yS ck v0 lae Hls Hys HfbyEqFby_case bk H y ls yS ck v0 lae Hls Hys Hfby
    end

  with sem_node_mult (f : ident)
                     (ls : stream (nelist value))
                     (ys : stream value)
                     (Hn : sem_node G f ls ys) {struct Hn} : Pn f ls ys Hn :=
    match Hn in (sem_node _ f ls ys) return (Pn f ls ys Hn) with
    | SNode bk f ls ys i o eqs Hck Hf Hnode
        SNode_case bk f ls ys i o eqs Hck Hf Hnode
          
           (match Hnode with
            | ex_intro H (conj Hxs (conj Hys (conj Hout Heqs))) ⇒
                ex_intro _ H (conj Hxs (conj Hys (conj Hout
                  (((fix map (eqs : list equation)
                             (Heqs: List.Forall (sem_equation G bk H) eqs) :=
                       match Heqs in List.Forall _ fs
                             return (List.Forall (fun eq Hsem,
                                                        P bk H eq Hsem) fs)
                       with
                       | List.Forall_nilList.Forall_nil _
                       | List.Forall_cons eq eqs Heq Heqs´
                         List.Forall_cons eq (@ex_intro _ _ Heq
                                           (sem_equation_mult bk H eq Heq))
                                     (map eqs Heqs´)
                       end) eqs Heqs)))))
            end)
    end.

End sem_node_mult.

Determinism of the semantics

Instantaneous semantics


Section InstantDeterminism.

Variable base: bool.

Lemma sem_var_instant_det:
   x R v1 v2,
    sem_var_instant R x v1
    → sem_var_instant R x v2
    → v1 = v2.
Proof.
  intros R x v1 v2 H1 H2.
  inversion_clear H1 as [Hf1];
    inversion_clear H2 as [Hf2];
    congruence.
Qed.

Lemma sem_clock_instant_det:
   ck R v1 v2,
    sem_clock_instant base R ck v1
    → sem_clock_instant base R ck v2
    → v1 = v2.
Proof.
  induction ck; repeat inversion_clear 1; intuition;
  try match goal with
        | H1: sem_clock_instant ?bk ?R ?ck ?l,
          H2: sem_clock_instant ?bk ?R ?ck ?r |- ?l = ?r
          eapply IHck; eassumption
        | H1: sem_var_instant ?R ?i (present (Cbool ?l)),
          H2: sem_var_instant ?R ?i (present (Cbool ?r)),
          H3: ?l = ?rFalse |- _ = _
          exfalso; apply H3;
          cut (present (Cbool l) = present (Cbool r)); [injection 1; auto|];
          eapply sem_var_instant_det; eassumption
      end.
Qed.

Lemma sem_lexp_instant_det:
   R e v1 v2,
    sem_lexp_instant base R e v1
    → sem_lexp_instant base R e v2
    → v1 = v2.
Proof.
  intros R e.
  induction e using lexp_ind2;
    try now (do 2 inversion_clear 1);
    match goal with
    | H1:sem_var_instant ?R ?e (present (Cbool ?b1)),
      H2:sem_var_instant ?R ?e (present (Cbool ?b2)),
      H3: ?b1 ?b2 |- _
      exfalso; apply H3;
      cut (present (Cbool b1) = present (Cbool b2)); [injection 1; auto|];
      eapply sem_var_instant_det; eassumption
    | H1:sem_var_instant ?R ?e ?v1,
      H2:sem_var_instant ?R ?e ?v2 |- ?v1 = ?v2
      eapply sem_var_instant_det; eassumption
    | H1:sem_var_instant ?R ?e (present _),
      H2:sem_var_instant ?R ?e absent |- _
      apply (sem_var_instant_det _ _ _ _ H1) in H2;
      discriminate
    | _auto
    end.
- do 2 inversion_clear 1; destruct base; congruence.
- intros v1 v2 Hsem1 Hsem2.
inversion_clear Hsem1; inversion_clear Hsem2.
× do 2 f_equal. clear H1 H3. revert cs cs0 H0 H2.
  induction les as [| le les]; intros cs1 cs2 Hrec1 Hrec2.
  + inversion Hrec1. inversion Hrec2. subst. symmetry in H1, H4.
    apply Nelist.map_eq_nebase in H1. destruct H1 as [? [? ?]].
    apply Nelist.map_eq_nebase in H4. destruct H4 as [? [? ?]]. subst.
    f_equal. rewrite present_injection. inversion_clear H. now apply H0.
  + inversion Hrec1; subst. inversion Hrec2; subst.
    symmetry in H2, H5.
    apply Nelist.map_eq_necons in H2. destruct H2 as [x1 [cs1´ [Hcs1 [Hx1 Hmap1]]]].
    apply Nelist.map_eq_necons in H5. destruct H5 as [x2 [cs2´ [Hcs2 [Hx2 Hmap2]]]]. subst.
    assert (Hx : x1 = x2).
    { inversion_clear H. rewrite present_injection. now apply H0. }
    inversion_clear H. inversion_clear Hrec1; inversion_clear Hrec2.
    f_equal; trivial. now apply (IHles H1).
× exfalso. destruct les as [le | le les].
  + inversion H0. inversion H2. subst. symmetry in H4.
    apply Nelist.map_eq_nebase in H4. destruct H4 as [? [? ?]]. subst. simpl in ×.
    inversion_clear H. specialize (H3 _ _ H8 H5). discriminate.
  + inversion H0; subst. inversion H2; subst.
    inversion_clear H. specialize (H3 _ _ H6 H9).
    symmetry in H5. apply Nelist.map_eq_necons in H5. decompose [ex and] H5. subst. discriminate.
× exfalso. destruct les as [| le les].
  + inversion H0. inversion H1. subst. symmetry in H7.
    apply Nelist.map_eq_nebase in H7. destruct H7 as [? [? ?]]. subst. simpl in ×.
    inversion_clear H. specialize (H3 _ _ H8 H5). discriminate.
  + inversion H0; subst. inversion H1; subst.
    inversion_clear H. specialize (H3 _ _ H6 H7).
    symmetry in H5. apply Nelist.map_eq_necons in H5. decompose [ex and] H5. subst. discriminate.
× reflexivity.
Qed.

Lemma sem_laexp_instant_det:
   R ck e v1 v2,
    sem_laexp_instant base R ck e v1
    → sem_laexp_instant base R ck e v2
    → v1 = v2.
Proof.
  intros R ck e v1 v2.
  do 2 inversion_clear 1;
    match goal with
      | H1:sem_lexp_instant _ _ _ _, H2:sem_lexp_instant _ _ _ _ |- _
        eapply sem_lexp_instant_det; eassumption
      | H1:sem_clock_instant _ _ _ ?T, H2:sem_clock_instant _ _ _ ?F |- _
        assert (T = F) by (eapply sem_clock_instant_det; eassumption);
          try discriminate
    end; auto.
Qed.

Lemma sem_lexps_instant_det:
   R les cs1 cs2,
    sem_lexps_instant base R les cs1
    sem_lexps_instant base R les cs2
    cs1 = cs2.
Proof.
  intros R les cs1 cs2. apply Nelist.Forall2_det. apply sem_lexp_instant_det.
Qed.

Lemma sem_laexps_instant_det:
   R ck e v1 v2,
    sem_laexps_instant base R ck e v1
    → sem_laexps_instant base R ck e v2
    → v1 = v2.
Proof.
  intros until v2.
  do 2 inversion_clear 1;
    match goal with
      | H1: sem_lexps_instant _ _ _ _, H2: sem_lexps_instant _ _ _ _ |- _
        eapply sem_lexps_instant_det; eauto
      | H1:sem_clock_instant _ _ _ ?T, H2:sem_clock_instant _ _ _ ?F |- _
        let H := fresh in
        assert (H: T = F) by (eapply sem_clock_instant_det; eassumption);
          try discriminate H
    end; congruence.
Qed.

Lemma sem_cexp_instant_det:
   R e v1 v2,
    sem_cexp_instant base R e v1
    → sem_cexp_instant base R e v2
    → v1 = v2.
Proof.
  intros R e.
  induction e;
  do 2 inversion_clear 1;
    try match goal with
      | H1: sem_cexp_instant ?bk ?R ?e ?l,
        H2: sem_cexp_instant ?bk ?R ?e ?r
        |- ?l = ?r
        (eapply IHe1; eassumption)
     || (eapply IHe2; eassumption)
      | H1: sem_var_instant ?R ?i (present (Cbool true)),
        H2: sem_var_instant ?R ?i (present (Cbool false)) |- _
        exfalso;
          assert (present (Cbool true) = present (Cbool false))
          by (eapply sem_var_instant_det; eassumption);
          discriminate
      | H1: sem_lexp_instant ?bk ?R ?l ?v1,
        H2: sem_lexp_instant ?bk ?R ?l ?v2 |- ?v1 = ?v2
        eapply sem_lexp_instant_det; eassumption
      | H1: sem_var_instant ?R ?i (present _),
        H2: sem_var_instant ?R ?i absent |- _
        apply sem_var_instant_det with (1:=H1) in H2; discriminate
      | |- absent = absentreflexivity
    end.
Qed.

Lemma sem_caexp_instant_det:
   R ck e v1 v2,
    sem_caexp_instant base R ck e v1
    → sem_caexp_instant base R ck e v2
    → v1 = v2.
Proof.
  intros until v2.
  do 2 inversion_clear 1;
  match goal with
  | H1: sem_cexp_instant _ _ _ _,
    H2: sem_cexp_instant _ _ _ _ |- _
    eapply sem_cexp_instant_det; eassumption
  | H1:sem_clock_instant _ _ _ ?T,
    H2:sem_clock_instant _ _ _ ?F |- _
    let H := fresh in
    assert (H: T = F) by (eapply sem_clock_instant_det; eassumption);
      try discriminate H
  end; congruence.
Qed.

End InstantDeterminism.

Lifted semantics


Section LiftDeterminism.

Variable bk : stream bool.

Require Import Logic.FunctionalExtensionality.

Lemma lift_det:
   {A B} (P: boolRABProp) (bk: stream bool) H x (xs1 xs2 : stream B),
    ( b R v1 v2, P b R x v1P b R x v2v1 = v2) →
    lift bk P H x xs1lift bk P H x xs2xs1 = xs2.
Proof.
  intros ** Hpoint H1 H2.
  extensionality n. specialize (H1 n). specialize (H2 n).
  eapply Hpoint; eassumption.
Qed.

Ltac apply_lift sem_det :=
  intros; eapply lift_det; try eassumption;
  compute; intros; eapply sem_det; eauto.

Lemma sem_var_det:
   H x xs1 xs2,
    sem_var bk H x xs1sem_var bk H x xs2xs1 = xs2.
Proof.
  apply_lift sem_var_instant_det.
Qed.

Lemma sem_clock_det : H ck bs1 bs2,
  sem_clock bk H ck bs1sem_clock bk H ck bs2bs1 = bs2.
Proof.
  apply_lift sem_clock_instant_det.
Qed.

Lemma sem_lexp_det:
   H e xs1 xs2,
    sem_lexp bk H e xs1sem_lexp bk H e xs2xs1 = xs2.
Proof.
  apply_lift sem_lexp_instant_det.
Qed.

Lemma sem_lexps_det:
   H les cs1 cs2,
    sem_lexps bk H les cs1
    sem_lexps bk H les cs2
    cs1 = cs2.
Proof.
  apply_lift sem_lexps_instant_det.
Qed.

Lemma sem_laexp_det:
   H ck e xs1 xs2,
    sem_laexp bk H ck e xs1sem_laexp bk H ck e xs2xs1 = xs2.
Proof.
  apply_lift sem_laexp_instant_det.
Qed.

Lemma sem_laexps_det:
   H ck e xs1 xs2,
    sem_laexps bk H ck e xs1sem_laexps bk H ck e xs2xs1 = xs2.
Proof.
  apply_lift sem_laexps_instant_det.
Qed.

Lemma sem_cexp_det:
   H c xs1 xs2,
    sem_cexp bk H c xs1sem_cexp bk H c xs2xs1 = xs2.
Proof.
  apply_lift sem_cexp_instant_det.
Qed.

Lemma sem_caexp_det:
   H ck c xs1 xs2,
    sem_caexp bk H ck c xs1sem_caexp bk H ck c xs2xs1 = xs2.
Proof.
  apply_lift sem_caexp_instant_det.
Qed.

End LiftDeterminism.

Ltac sem_det :=
  match goal with
    | H1: sem_cexp_instant ?bk ?H ?C ?X,
      H2: sem_cexp_instant ?bk ?H ?C ?Y |- ?X = ?Y
      eapply sem_cexp_instant_det; eexact H1 || eexact H2
    | H1: sem_cexp ?bk ?H ?C ?X,
      H2: sem_cexp ?bk ?H ?C ?Y |- ?X = ?Y
      eapply sem_cexp_det; eexact H1 || eexact H2
    | H1: sem_lexps_instant ?bk ?H ?C ?X,
      H2: sem_lexps_instant ?bk ?H ?C ?Y |- ?X = ?Y
      eapply sem_lexps_instant_det; eexact H1 || eexact H2
    | H1: sem_lexps ?bk ?H ?C ?X,
      H2: sem_lexps ?bk ?H ?C ?Y |- ?X = ?Y
      eapply sem_lexps_det; eexact H1 || eexact H2
    | H1: sem_laexps_instant ?bk ?H ?ck ?C ?X,
      H2: sem_laexps_instant ?bk ?H ?ck ?C ?Y |- ?X = ?Y
      eapply sem_laexps_instant_det; eexact H1 || eexact H2
    | H1: sem_laexps ?bk ?H ?ck ?C ?X,
      H2: sem_laexps ?bk ?H ?ck ?C ?Y |- ?X = ?Y
      eapply sem_laexps_det; eexact H1 || eexact H2
    | H1: sem_lexp_instant ?bk ?H ?C ?X,
      H2: sem_lexp_instant ?bk ?H ?C ?Y |- ?X = ?Y
      eapply sem_lexp_instant_det; eexact H1 || eexact H2
    | H1: sem_lexp ?bk ?H ?C ?X,
      H2: sem_lexp ?bk ?H ?C ?Y |- ?X = ?Y
      eapply sem_lexp_det; eexact H1 || eexact H2
    | H1: sem_laexp_instant ?bk ?H ?CK ?C ?X,
      H2: sem_laexp_instant ?bk ?H ?CK ?C ?Y |- ?X = ?Y
      eapply sem_laexp_instant_det; eexact H1 || eexact H2
    | H1: sem_laexp ?bk ?H ?CK ?C ?X,
      H2: sem_laexp ?bk ?H ?CK ?C ?Y |- ?X = ?Y
      eapply sem_laexp_det; eexact H1 || eexact H2
    | H1: sem_var_instant ?H ?C ?X,
      H2: sem_var_instant ?H ?C ?Y |- ?X = ?Y
      eapply sem_var_instant_det; eexact H1 || eexact H2
    | H1: sem_var ?bk ?H ?C ?X,
      H2: sem_var ?bk ?H ?C ?Y |- ?X = ?Y
      eapply sem_var_det; eexact H1 || eexact H2
  end.

Properties of the global environment


Lemma find_node_other:
   f node G node´,
    node.(n_name) f
    → (find_node f (node::G) = Some node´
         find_node f G = Some node´).
Proof.
  intros f node G node´ Hnf.
  apply BinPos.Pos.eqb_neq in Hnf.
  simpl.
  unfold ident_eqb.
  rewrite Hnf.
  reflexivity.
Qed.

Lemma sem_node_cons:
   node G f xs ys,
    Ordered_nodes (node::G)
    → sem_node (node::G) f xs ys
    → node.(n_name) f
    → sem_node G f xs ys.
Proof.
  intros node G f xs ys Hord Hsem Hnf.
  revert Hnf.
  induction Hsem as [
                    | bk H y ck f lae ls ys Hlae Hvar Hnode IH
                    |
                    | bk f xs ys i o eqs Hbk Hf Heqs IH]
  using sem_node_mult
  with (P := fun bk H eq Hsem¬Is_node_in_eq node.(n_name) eq
                              → sem_equation G bk H eq).
  - econstructor; eassumption.
  - intro Hnin.
    eapply @SEqApp with (1:=Hlae) (2:=Hvar).
    apply IH. intro Hnf. apply Hnin. rewrite Hnf. constructor.
  - intro; eapply SEqFby; eassumption.
  - intro.
    rewrite find_node_tl with (1:=Hnf) in Hf.
    eapply SNode; eauto.
    clear Heqs.
    destruct IH as [H [Hxs [Hys [Hout Heqs]]]].
     H.
    repeat (split; eauto).
    set (cnode := {| n_name := f; n_input := i; n_output := o; n_eqs := eqs |}).
    assert (List.Forall (fun eq¬ Is_node_in_eq (n_name node) eq) (n_eqs cnode))
      by (eapply Is_node_in_Forall; try eassumption;
          eapply find_node_later_not_Is_node_in; try eassumption).
    eapply Forall_Forall in Heqs; try eauto.
    eapply Forall_impl with (2:=Heqs).
    destruct 1 as [Hnini [Hsem HH]].
    intuition.
Qed.

Lemma find_node_find_again:
   G f i o eqs g,
    Ordered_nodes G
    → find_node f G =
       Some {| n_name := f; n_input := i; n_output := o; n_eqs := eqs |}
    → Is_node_in g eqs
    → List.Exists (fun ndg = nd.(n_name)) G.
Proof.
  intros G f i o eqs g Hord Hfind Hini.
  apply find_node_split in Hfind.
  destruct Hfind as [bG [aG Hfind]].
  rewrite Hfind in ×.
  clear Hfind.
  apply Ordered_nodes_append in Hord.
  apply Exists_app.
  constructor 2.
  inversion_clear Hord as [|? ? ? HH H0]; clear H0.
  apply HH in Hini; clear HH.
  intuition.
Qed.

Lemma sem_node_cons2:
   nd G f xs ys,
    Ordered_nodes G
    → sem_node G f xs ys
    → List.Forall (fun nd´ : noden_name nd n_name nd´) G
    → sem_node (nd::G) f xs ys.
Proof.
  Hint Constructors sem_equation.
  intros nd G f xs ys Hord Hsem Hnin.
  assert (Hnin´:=Hnin).
  revert Hnin´.
  induction Hsem as [
                    | bk H y f lae ls ys Hlae Hvar Hnode IH
                    |
                    | bk f xs ys i o eqs Hbk Hfind Heqs IH]
  using sem_node_mult
  with (P := fun bk H eq Hsem¬Is_node_in_eq nd.(n_name) eq
                              → sem_equation (nd::G) bk H eq);
    try eauto; intro HH.
  clear HH.
  assert (nd.(n_name) f) as Hnf.
  { intro Hnf.
    rewrite Hnf in ×.
    apply find_node_split in Hfind.
    destruct Hfind as [bG [aG Hge]].
    rewrite Hge in Hnin.
    apply Forall_app in Hnin.
    destruct Hnin as [H0 Hfg]; clear H0.
    inversion_clear Hfg.
    match goal with H:f_ |- Falseapply H end.
    reflexivity. }
  apply find_node_other with (2:=Hfind) in Hnf.
  econstructor; eauto.
  clear Heqs.
  destruct IH as [H [Hxs [Hys [Hout Heqs]]]].
   H.
  intuition; clear Hxs Hys.
  assert ( g, Is_node_in g eqs
                    → List.Exists (fun ndg = nd.(n_name)) G)
    as Hniex
    by (intros g Hini;
        apply find_node_find_again with (1:=Hord) (2:=Hfind) in Hini;
        exact Hini).
  assert (List.Forall
            (fun eq g,
                 Is_node_in_eq g eq
                 → List.Exists (fun ndg = nd.(n_name)) G) eqs) as HH.
  {
    clear Hfind Heqs Hnf.
    induction eqs as [|eq eqs IH]; [now constructor|].
    constructor.
    - intros g Hini.
      apply Hniex.
      constructor 1; apply Hini.
    - apply IH.
      intros g Hini; apply Hniex.
      constructor 2; apply Hini.
  }
  apply Forall_Forall with (1:=HH) in Heqs.
  apply Forall_impl with (2:=Heqs).
  intros eq IH.
  destruct IH as [Hsem [IH0 IH1]].
  apply IH1.
  intro Hini.
  apply Hsem in Hini.
  apply Forall_Exists with (1:=Hnin) in Hini.
  apply List.Exists_exists in Hini.
  destruct Hini as [nd´ [Hin [Hneq Heq]]].
  intuition.
Qed.

Lemma Forall_sem_equation_global_tl:
   bk nd G H eqs,
    Ordered_nodes (nd::G)
    → ¬ Is_node_in nd.(n_name) eqs
    → List.Forall (sem_equation (nd::G) bk H) eqs
    → List.Forall (sem_equation G bk H) eqs.
Proof.
  intros bk nd G H eqs Hord.
  induction eqs as [|eq eqs IH]; [trivial|].
  intros Hnini Hsem.
  apply Forall_cons2 in Hsem; destruct Hsem as [Hseq Hseqs].
  apply IH in Hseqs.
  2:(apply not_Is_node_in_cons in Hnini;
     destruct Hnini; assumption).
  apply List.Forall_cons with (2:=Hseqs).
  inversion Hseq as [|? ? ? ? ? ? ? Hsem Hvar Hnode|]; subst.
  - econstructor; eassumption.
  - apply not_Is_node_in_cons in Hnini.
    destruct Hnini as [Hninieq Hnini].
    assert (nd.(n_name) f) as Hnf
      by (intro HH; apply Hninieq; rewrite HH; constructor).
    econstructor; eauto.
    eapply sem_node_cons; eauto.
  - econstructor; eauto.
Qed.

Clocking property


Lemma subrate_clock:
   R ck,
    sem_clock_instant false R ck false.
Proof.
  Hint Constructors sem_clock_instant.
  intros R ck.
  induction ck; eauto.
Qed.

Presence and absence in non-empty lists


Lemma not_absent_present_list:
   xss n vs,
    present_list xss n vs¬ absent_list xss n.
Proof.
  intros × Hpres Habs.
  unfold present_list in Hpres.
  unfold absent_list in Habs.
  rewrite Hpres in ×. destruct vs; inversion_clear Habs; discriminate.
Qed.