Module Infty
From Velus Require Import Common.
From Velus Require Import Lustre.Denot.Cpo.
Require Import SDfuns CommonDS.
 Infinity of streams defined in SDfuns  
Ltac solve_err :=
  
try match goal with
    | |- 
context [ 
DS_const _ ] =>
        
repeat rewrite DS_const_eq, 
rem_cons;
        
now (
apply is_ncons_DS_const || 
apply is_consn_DS_const)
        || 
now auto using is_cons_DS_const, 
is_consn_DS_const, 
is_ncons_DS_const
    end.
 Productivity of Lustre operators 
Section Ncons_ops.
Context {
A B D : 
Type}.
Lemma is_ncons_sconst :
  
forall (
c : 
A) 
bs n,
    
is_ncons n bs ->
    
is_ncons n (
sconst c bs).
Proof.
Lemma sconst_inf :
  
forall (
c : 
A) 
bs,
    
infinite bs ->
    
infinite (
sconst c bs).
Proof.
Lemma is_ncons_sunop :
  
forall (
f : 
A -> 
option B) 
s n,
    
is_ncons n s ->
    
is_ncons n (
sunop f s).
Proof.
Lemma sunop_inf :
  
forall (
op : 
A -> 
option B) 
s,
    
infinite s ->
    
infinite (
sunop op s).
Proof.
Lemma is_ncons_sbinop :
  
forall (
f : 
A -> 
B -> 
option D) 
s1 s2 n,
    
is_ncons n s1 ->
    
is_ncons n s2 ->
    
is_ncons n (
sbinop f s1 s2).
Proof.
Lemma sbinop_inf :
  
forall (
f : 
A -> 
B -> 
option D) 
s1 s2,
    
infinite s1 ->
    
infinite s2 ->
    
infinite (
sbinop f s1 s2).
Proof.
Lemma inf_sbinop :
  
forall (
f : 
A -> 
B -> 
option D) 
s1 s2,
    
infinite (
sbinop f s1 s2) ->
    
infinite s1 /\ 
infinite s2.
Proof.
Lemma is_cons_fby :
  
forall (
xs ys : 
DS (
sampl A)),
    
is_cons xs ->
    
is_cons (
fby xs ys).
Proof.
Lemma is_ncons_S_fby1 :
  
forall n v (
xs ys : 
DS (
sampl A)),
    
is_ncons (
S n) 
xs ->
    
is_ncons n ys ->
    
is_ncons (
S n) (
fby1 v xs ys).
Proof.
Lemma is_ncons_S_fby :
  
forall n (
xs ys : 
DS (
sampl A)),
    
is_ncons (
S n) 
xs ->
    
is_ncons n ys ->
    
is_ncons (
S n) (
fby xs ys).
Proof.
Lemma is_ncons_fby :
  
forall n (
xs ys : 
DS (
sampl A)),
    
is_ncons n xs ->
    
is_ncons n ys ->
    
is_ncons n (
SDfuns.fby xs ys).
Proof.
Lemma fby_inf :
  
forall (
xs ys : 
DS (
sampl A)),
    
infinite xs ->
    
infinite ys ->
    
infinite (
SDfuns.fby xs ys).
Proof.
Lemma is_cons_swhen :
  
forall T OT TB,
  
forall k xs cs,
    
is_cons xs ->
    
is_cons cs ->
    
is_cons (@
swhen A B T OT TB k xs cs).
Proof.
Lemma is_ncons_swhen :
  
forall T OT TB,
  
forall k n xs cs,
    
is_ncons n xs ->
    
is_ncons n cs ->
    
is_ncons n (@
swhen A B T OT TB k xs cs).
Proof.
Lemma swhen_inf :
  
forall T OT TB,
  
forall k xs cs,
    
infinite xs ->
    
infinite cs ->
    
infinite (@
swhen A B T OT TB k xs cs).
Proof.
Lemma is_ncons_zip3 :
  
forall A B C D (
op : 
A -> 
B -> 
C -> 
D),
  
forall xs ys zs n,
    
is_ncons n xs /\ 
is_ncons n ys /\ 
is_ncons n zs ->
    
is_ncons n (
ZIP3 op xs ys zs).
Proof.
Lemma is_ncons_smerge :
  
forall T OT TB,
  
forall l n xs cs,
    
is_ncons n cs ->
    
forall_nprod (@
is_ncons _ n) 
xs ->
    
is_ncons n (@
smerge A B T OT TB l cs xs).
Proof.
Lemma smerge_inf :
  
forall T OT TB,
  
forall l xs cs,
    
infinite cs ->
    
forall_nprod (@
infinite _) 
xs ->
    
infinite (@
smerge A B T OT TB l cs xs).
Proof.
Lemma is_ncons_scase :
  
forall T OT TB,
  
forall l n xs cs,
    
is_ncons n cs ->
    
forall_nprod (@
is_ncons _ n) 
xs ->
    
is_ncons n (@
scase A B T OT TB l cs xs).
Proof.
Lemma scase_inf :
  
forall T OT TB,
  
forall l xs cs,
    
infinite cs ->
    
forall_nprod (@
infinite _) 
xs ->
    
infinite (@
scase A B T OT TB l cs xs).
Proof.
Lemma is_ncons_scase_def_ :
  
forall T OT TB,
  
forall l n cs ds xs,
    
is_ncons n cs ->
    
is_ncons n ds ->
    
forall_nprod (@
is_ncons _ n) 
xs ->
    
is_ncons n (@
scase_def_ A B T OT TB l cs ds xs).
Proof.
Lemma is_ncons_scase_def :
  
forall T OT TB,
  
forall l n cs dxs,
    
is_ncons n cs ->
    
forall_nprod (@
is_ncons _ n) 
dxs ->
    
is_ncons n (@
scase_def A B T OT TB l cs dxs).
Proof.
Lemma scase_def__inf :
  
forall T OT TB,
  
forall l cs ds xs,
    
infinite cs ->
    
infinite ds ->
    
forall_nprod (@
infinite _) 
xs ->
    
infinite (@
scase_def_ A B T OT TB l cs ds xs).
Proof.
Lemma scase_def_inf :
  
forall T OT TB,
  
forall l cs dxs,
    
infinite cs ->
    
forall_nprod (@
infinite _) 
dxs ->
    
infinite (@
scase_def A B T OT TB l cs dxs).
Proof.
Lemma is_cons_sreset :
  
forall I,
  
forall (
f : 
DS_prod (
fun _ : 
I => 
sampl A) -
C-> 
DS_prod (
fun _ : 
I => 
sampl A)) 
R X x,
    (
forall envI, (
forall x, 
is_cons (
envI x)) -> (
forall x, 
is_cons (
f envI x))) ->
    
is_cons R ->
    (
forall x, 
is_cons (
X x)) ->
    
is_cons (@
sreset I A f R X x).
Proof.
Lemma is_cons_sreset_aux :
  
forall I,
  
forall (
f : 
DS_prod (
fun _ : 
I => 
sampl A) -
C-> 
DS_prod (
fun _ : 
I => 
sampl A)) 
R X Y x,
    (
forall envI, (
forall x, 
is_cons (
envI x)) -> (
forall x, 
is_cons (
f envI x))) ->
    
is_cons R ->
    (
forall x, 
is_cons (
X x)) ->
    (
forall x, 
is_cons (
Y x)) ->
    
is_cons (@
sreset_aux I A f R X Y x).
Proof.
Lemma is_ncons_sreset :
  
forall I,
  
forall (
f : 
DS_prod (
fun _ : 
I => 
sampl A) -
C-> 
DS_prod (
fun _ : 
I => 
sampl A)) 
R X x,
    (
forall n envI, (
forall x, 
is_ncons n (
envI x)) -> (
forall x, 
is_ncons n (
f envI x))) ->
    
forall n,
      
is_ncons n R ->
      (
forall x, 
is_ncons n (
X x)) ->
      
is_ncons n (@
sreset I A f R X x).
Proof.
Lemma sreset_inf :
  
forall I,
  
forall (
f : 
DS_prod (
fun _ : 
I => 
sampl A) -
C-> 
DS_prod (
fun _ : 
I => 
sampl A)) 
R X,
    (
forall envI, 
all_infinite envI -> 
all_infinite (
f envI)) ->
    
infinite R ->
    
all_infinite X ->
    
all_infinite (@
sreset I A f R X).
Proof.
Lemma sreset_inf_dom :
  
forall I,
  
forall (
f : 
DS_prod (
fun _ : 
I => 
sampl A) -
C-> 
DS_prod (
fun _ : 
I => 
sampl A)) 
R X,
  
forall ins outs,
    (
forall envI, 
infinite_dom envI ins -> 
infinite_dom (
f envI) 
outs) ->
    
infinite R ->
    
infinite_dom X ins ->
    
infinite_dom (
sreset f R X) 
outs.
Proof.
  intros * 
If Ir Ix.
  
rewrite sreset_eq.
  
assert (
infinite_dom (
f X) 
outs) 
as Iy by auto.
  
remember (
_ f X) 
as Y eqn:
HH; 
clear HH.
  
intros x Hin.
  
remember_ds (
sreset_aux _ _ _ _ _) 
as t.
  
revert Ir Ix Iy Ht Hin.
  
revert R X Y x t.
  
cofix Cof; 
intros.
  
apply infinite_decomp in Ir as (
r & 
R' & 
Hr & 
Ir').
  
rewrite <- 
PROJ_simpl, 
Hr, 
sreset_aux_eq in Ht.
  
apply REM_env_inf_dom in Ix as Irx, 
Iy as Iry.
  
destruct (
Iy x), (
If X Ix x); 
auto. 
  
apply If, 
REM_env_inf_dom in Ix as Ifx.
  
constructor.
  - 
cases; 
rewrite Ht, ?
sreset_aux_eq, 
PROJ_simpl, 
APP_env_eq;
      
eauto using app_is_cons.
  - 
apply rem_eq_compat in Ht.
    
destruct r; [
rewrite sreset_aux_eq in Ht|];
      
rewrite PROJ_simpl, 
APP_env_eq, 
rem_app in Ht; 
eauto 2.
Qed.
 
End Ncons_ops.