Library stream
Require Export Streams List Omega FunctionalExtensionality.
Set Implicit Arguments.
Unset Strict Implicit.
Section Stream.
Variable ELEM: Type.
Set Implicit Arguments.
Unset Strict Implicit.
Section Stream.
Variable ELEM: Type.
Streams are already defined in the Coq standard library, but most things such as Str_nth are very badly defined (due to hd and tl not automatically unfolding), so I redefine them here.
Fixpoint appinf (l : list ELEM) (il : Stream ELEM) {struct l} : Stream ELEM :=
match l with
| nil => il
| e :: l' => Cons e (appinf l' il)
end.
Lemma appinf_app: forall l1 l2 st,
appinf l1 (appinf l2 st) = appinf (l1 ++ l2) st.
Proof.
induction l1; simpl; trivial; intros; f_equal; eauto.
Qed.
Bisimilarity
CoInductive bisim : Stream ELEM -> Stream ELEM -> Prop :=
| bisim_intro
(l1 l2 : Stream ELEM)
(CIHbisim: bisim l1 l2)
e :
bisim (Cons e l1) (Cons e l2)
.
Lemma bisim_appinf: forall etr1 etr2
(Hbisim: bisim etr1 etr2)
l,
bisim (appinf l etr1) (appinf l etr2).
Proof.
intros.
revert etr1 etr2 Hbisim.
induction l; simpl.
tauto.
intros.
constructor.
auto.
Qed.
Lemma appinf_bisim: forall l etr1 etr2
(Hbisim: bisim (appinf l etr1) (appinf l etr2)),
bisim etr1 etr2.
Proof.
induction l; simpl; auto.
inversion 1; subst; eauto.
Qed.
Lemma bisim_appinf_strong: forall l etr1 etr
(Hbisim: bisim (appinf l etr1) etr),
exists etr2, etr = appinf l etr2.
Proof.
induction l; simpl.
eauto.
inversion 1; subst.
edestruct IHl.
eassumption.
subst.
eauto.
Qed.
Lemma appinf_le_bisim: forall l1 l2
(Hlength: length l1 <= length l2)
etr1 etr2 (Hbisim: bisim (appinf l1 etr1) (appinf l2 etr2)),
exists l, l2 = l1 ++ l /\ bisim etr1 (appinf l etr2).
Proof.
induction l1.
simpl.
intros.
esplit.
split.
reflexivity.
assumption.
destruct l2.
simpl.
intro; exfalso; omega.
simpl.
inversion 2; subst.
edestruct IHl1.
2 : eassumption.
omega.
destruct H.
subst.
esplit.
split.
reflexivity.
assumption.
Qed.
Lemma bisim_EqSt: forall s1 s2
(Hbisim: bisim s1 s2),
EqSt s1 s2.
Proof.
cofix COINDHYP.
inversion 1; subst.
econstructor; simpl; eauto.
Qed.
Lemma EqSt_bisim: forall s1 s2
(HEqSt: EqSt s1 s2),
bisim s1 s2.
Proof.
cofix COINDHYP.
destruct s1, s2.
inversion 1; subst; simpl in *; subst.
constructor. eauto.
Qed.
In some cases, streams are constructed by producing several elements at a time. The following bisim' and traceinf' type takes care of this, and helps recover the regular streams.
CoInductive bisim' : Stream ELEM -> Stream ELEM -> Prop :=
| bisim'_intro
(l1 l2 : Stream ELEM)
(CIHbisim: bisim' l1 l2)
e (He_not_nil: e <> nil)
l1' (Hl1': l1' = appinf e l1)
l2' (Hl2': l2' = appinf e l2):
bisim' l1' l2'
.
Lemma bisim_bisim': forall l1 l2, bisim l1 l2 -> bisim' l1 l2.
Proof.
cofix COINDHYP.
inversion 1; subst.
eapply bisim'_intro with (e := e :: nil).
eapply COINDHYP.
eapply CIHbisim.
clear COINDHYP; congruence.
reflexivity.
reflexivity.
Qed.
Lemma bisim'_inv: forall l1 l2, bisim' l1 l2 ->
exists e, exists l'1,
l1 = Cons e l'1 /\
exists l'2, l2 = Cons e l'2 /\
bisim' l'1 l'2.
Proof.
inversion 1; subst.
destruct e.
destruct He_not_nil. trivial.
destruct e0.
esplit.
esplit.
split.
simpl; reflexivity.
esplit.
split.
simpl; reflexivity.
assumption.
esplit.
esplit.
split.
simpl; reflexivity.
esplit.
split.
simpl; reflexivity.
eapply bisim'_intro with (e := e0 :: e1).
3: simpl; reflexivity.
3: simpl; reflexivity.
eexact CIHbisim.
congruence.
Qed.
Corollary bisim'_bisim: forall l1 l2, bisim' l1 l2 ->
bisim l1 l2.
Proof.
cofix COINDHYP.
intros.
destruct (bisim'_inv H).
destruct H0.
destruct H0.
destruct H1.
destruct H1.
subst.
econstructor.
eapply COINDHYP.
assumption.
Qed.
CoInductive traceinf': Type :=
| Econsinf': forall (t: list ELEM) (T: traceinf'), t <> nil -> traceinf'.
Implicit Arguments Econsinf' [].
Program Definition split_traceinf' (t: list ELEM) (T: traceinf') (NE: t <> nil): ELEM * traceinf' :=
match t with
| nil => _
| e :: nil => (e, T)
| e :: t' => (e, @Econsinf' t' T _)
end.
Next Obligation.
intro. subst. edestruct H. reflexivity.
Qed.
Implicit Arguments split_traceinf' [].
CoFixpoint traceinf_of_traceinf' (T': traceinf') : Stream ELEM :=
match T' with
| Econsinf' t T'' NOTEMPTY =>
let (e, tl) := split_traceinf' t T'' NOTEMPTY in
Cons e (traceinf_of_traceinf' tl)
end.
Implicit Arguments traceinf_of_traceinf' [].
Remark unroll_traceinf':
forall T, T = match T with Econsinf' t T' NE => Econsinf' t T' NE end.
Proof.
intros. destruct T; auto.
Qed.
Lemma traceinf_traceinf'_app:
forall t T NE,
traceinf_of_traceinf' (Econsinf' t T NE) = appinf t (traceinf_of_traceinf' T).
Proof.
induction t.
intros. elim NE. auto.
intros. simpl.
rewrite (unfold_Stream (traceinf_of_traceinf' (Econsinf' (a :: t) T NE))).
simpl. destruct t. auto.
simpl. f_equal. apply IHt.
Qed.
Representation of a Stream ELEM through a function nat -> ELEM
Function func_of_stream (il: Stream ELEM) (n: nat) {struct n} : ELEM :=
match il with Cons e il' =>
match n with
| O => e
| S n' => func_of_stream il' n'
end
end.
Lemma bisim_func_eq : forall tr1 tr2
(Hbisim: bisim tr1 tr2),
func_of_stream tr1 = func_of_stream tr2.
Proof.
unfold func_of_stream.
intros.
apply functional_extensionality.
intro n.
revert tr1 tr2 Hbisim.
induction n; simpl; inversion 1; subst; simpl; eauto.
Qed.
Lemma func_eq_bisim: forall tr1 tr2,
func_of_stream tr1 = func_of_stream tr2 ->
bisim tr1 tr2.
Proof.
cofix COINDHYP.
destruct tr1; destruct tr2.
intro.
replace e0 with e.
econstructor.
apply COINDHYP.
clear COINDHYP.
apply functional_extensionality.
intro x.
change (func_of_stream tr1 x) with (func_of_stream (Cons e tr1) (S x)).
change (func_of_stream tr2 x) with (func_of_stream (Cons e0 tr2) (S x)).
congruence.
change e with (func_of_stream (Cons e tr1) O).
change e0 with (func_of_stream (Cons e0 tr2) O).
congruence.
Qed.
Corollary bisim_refl: forall tr, bisim tr tr.
Proof.
intros.
eapply func_eq_bisim.
reflexivity.
Qed.
Lemma bisim_sym: forall tr1 tr2 (Hbisim: bisim tr1 tr2),
bisim tr2 tr1.
Proof.
eauto using func_eq_bisim, sym_eq, bisim_func_eq.
Qed.
Lemma bisim_trans: forall tr1 tr2 (Hbisim12: bisim tr1 tr2)
tr3 (Hbisim23: bisim tr2 tr3),
bisim tr1 tr3.
Proof.
eauto using func_eq_bisim, trans_eq, bisim_func_eq.
Qed.
Function pop (il: Stream ELEM) (n: nat) {struct n}: list ELEM :=
match n with
| O => nil
| S n' => match il with Cons e il' => e :: pop il' n' end
end.
Fixpoint discard (il: Stream ELEM) (n: nat) {struct n}: Stream ELEM :=
match n with
| O => il
| S n' => match il with Cons _ il' => discard il' n' end
end.
Lemma pop_discard: forall n il,
il = appinf (pop il n) (discard il n).
Proof.
induction n; destruct il; simpl; auto; f_equal; auto.
Qed.
CoFixpoint stream_of_func (f: nat -> ELEM): Stream ELEM :=
Cons (f O) (stream_of_func (fun n => f (S n))).
Lemma func_of_stream_of_func: forall f,
func_of_stream (stream_of_func f) = f.
Proof.
intro.
apply functional_extensionality.
intro n.
revert f.
induction n; simpl; auto.
intros.
rewrite IHn.
trivial.
Qed.
Lemma stream_of_func_of_stream: forall st,
bisim st (stream_of_func (func_of_stream st)).
Proof.
cofix COINDHYP.
destruct st.
unfold stream_of_func.
rewrite unfold_Stream.
fold stream_of_func.
econstructor.
simpl.
replace (fun n => func_of_stream st n) with (func_of_stream st).
auto.
apply functional_extensionality.
trivial.
Qed.
Fixpoint n_f (f : nat -> ELEM) (n: nat) {struct n}: list ELEM :=
match n with
| O => nil
| S n' => (n_f f n') ++ f n' :: nil
end.
Lemma n_f_app: forall n2 f n1,
n_f f (n1 + n2) = n_f f n1 ++ n_f (fun n' => f (n1 + n')) n2.
Proof.
induction n2; simpl.
intros.
rewrite <- app_nil_end.
replace (n1 + O) with (n1) by omega.
trivial.
intros.
rewrite <- app_ass.
rewrite <- IHn2.
replace (n1 + S n2) with (S (n1 + n2)) by omega.
reflexivity.
Qed.
Lemma stream_of_func_unfold': forall f,
stream_of_func f = Cons (f O) (stream_of_func (fun x => f (S x))).
Proof.
unfold stream_of_func at 1.
intro.
rewrite unfold_Stream at 1.
trivial.
Qed.
Lemma stream_of_func_appinf: forall n f,
stream_of_func f = appinf (n_f f n) (stream_of_func (fun n' => f (n + n'))).
Proof.
induction n; simpl.
intro.
replace (fun n' => f n') with f.
trivial.
apply functional_extensionality.
trivial.
intros.
rewrite <- (appinf_app).
simpl.
rewrite (IHn f) at 1.
f_equal.
rewrite stream_of_func_unfold' at 1.
replace (n+O) with n by omega.
f_equal.
f_equal.
apply functional_extensionality.
intros.
f_equal.
omega.
Qed.
End Stream.
A map operator over streams and their functional representations
Section Mapinf.
Variables (A B : Type) (f: A -> B).
CoFixpoint mapinf (il: Stream A): Stream B :=
match il with
| Cons a il' => Cons (f a) (mapinf il')
end.
Lemma mapinf_cons: forall a al,
mapinf (Cons a al) = Cons (f a) (mapinf al).
Proof.
intros.
rewrite (unfold_Stream (mapinf (Cons a al))).
reflexivity.
Qed.
Lemma mapinf_appinf: forall l al,
mapinf (appinf l al) = appinf (map f l) (mapinf al).
Proof.
induction l; simpl.
trivial.
intros.
rewrite mapinf_cons.
rewrite IHl.
reflexivity.
Qed.
Function fmapinf (fa: nat -> A) (n: nat): B :=
f (fa n).
Lemma fmapinf_func_of_stream: forall t,
fmapinf (func_of_stream t) = func_of_stream (mapinf t).
Proof.
intros.
rewrite <- (func_of_stream_of_func (fmapinf (func_of_stream t))).
apply bisim_func_eq.
revert t.
cofix COINDHYP.
destruct t.
rewrite unfold_Stream.
simpl.
rewrite (unfold_Stream (stream_of_func (fmapinf (func_of_stream (Cons a t))))).
simpl.
unfold fmapinf at 1.
unfold func_of_stream at 1.
constructor.
eapply COINDHYP.
Qed.
Corollary fmapinf_func_of_stream_bisim: forall t1 t2,
bisim t1 t2 ->
fmapinf (func_of_stream t1) = func_of_stream (mapinf t2).
Proof.
intros.
rewrite (bisim_func_eq H).
eauto using fmapinf_func_of_stream.
Qed.
Lemma func_of_stream_mapinf_of_func: forall fa,
fmapinf fa = func_of_stream (mapinf (stream_of_func fa)).
Proof.
intros.
rewrite <- (func_of_stream_of_func fa) at 1.
eapply fmapinf_func_of_stream.
Qed.
Lemma mapinf_stream_of_func : forall fa,
bisim (mapinf (stream_of_func fa)) (stream_of_func (fmapinf fa)).
Proof.
intros.
eapply func_eq_bisim.
rewrite <- func_of_stream_mapinf_of_func.
rewrite func_of_stream_of_func.
reflexivity.
Qed.
End Mapinf.
Implicit Arguments Econsinf' [ELEM].
Implicit Arguments split_traceinf' [ELEM].
Implicit Arguments traceinf_of_traceinf' [ELEM].