]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/ground_2/lib/streams.ma
syntactic components detached from basic_2 become static_2
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / lib / streams.ma
index ec7194229483c5b3642aeecc17d22bccb27b6d8b..e1b3a68124c4ae14047c5ef8f75d17bf7257e3ae 100644 (file)
@@ -12,9 +12,8 @@
 (*                                                                        *)
 (**************************************************************************)
 
-include "ground_2/notation/constructors/cons_2.ma".
-include "ground_2/notation/relations/exteq_3.ma".
-include "ground_2/lib/star.ma".
+include "ground_2/notation/functions/oplusright_3.ma".
+include "ground_2/lib/relations.ma".
 
 (* STREAMS ******************************************************************)
 
@@ -22,61 +21,10 @@ coinductive stream (A:Type[0]): Type[0] ≝
 | seq: A → stream A → stream A
 .
 
-interpretation "cons (nstream)" 'Cons b t = (seq ? b t).
-
-coinductive eq_stream (A): relation (stream A) ≝
-| eq_seq: ∀t1,t2,b1,b2. b1 = b2 → eq_stream A t1 t2 → eq_stream A (b1@t1) (b2@t2)
-.
-
-interpretation "extensional equivalence (nstream)"
-   'ExtEq A t1 t2 = (eq_stream A t1 t2).
-
-definition eq_stream_repl (A) (R:relation …) ≝
-                          ∀t1,t2. t1 ≐⦋A⦌ t2 → R t1 t2.
-
-definition eq_stream_repl_back (A) (R:predicate …) ≝
-                               ∀t1. R t1 → ∀t2. t1 ≐⦋A⦌ t2 → R t2.
-
-definition eq_stream_repl_fwd (A) (R:predicate …) ≝
-                              ∀t1. R t1 → ∀t2. t2 ≐⦋A⦌ t1 → R t2.
-
-(* Basic inversion lemmas ***************************************************)
-
-lemma eq_stream_inv_seq: ∀A,t1,t2. t1 ≐⦋A⦌ t2 →
-                         ∀u1,u2,a1,a2. a1@u1 = t1 → a2@u2 = t2 →
-                         u1 ≐ u2 ∧ a1 = a2.
-#A #t1 #t2 * -t1 -t2
-#t1 #t2 #b1 #b2 #Hb #Ht #u1 #u2 #a1 #a2 #H1 #H2 destruct /2 width=1 by conj/
-qed-.
+interpretation "cons (nstream)" 'OPlusRight A a u = (seq A a u).
 
 (* Basic properties *********************************************************)
 
-lemma stream_rew (A) (t:stream A): match t with [ seq a u ⇒ a @ u ] = t.
+lemma stream_rew (A) (t:stream A): match t with [ seq a u ⇒ a  u ] = t.
 #A * //
 qed.
-
-let corec eq_stream_refl: ∀A. reflexive … (eq_stream A) ≝ ?.
-#A * #b #t @eq_seq //
-qed.
-
-let corec eq_stream_sym: ∀A. symmetric … (eq_stream A) ≝ ?.
-#A #t1 #t2 * -t1 -t2
-#t1 #t2 #b1 #b2 #Hb #Ht @eq_seq /2 width=1 by/
-qed-.
-
-lemma eq_stream_repl_sym: ∀A,R. eq_stream_repl_back A R → eq_stream_repl_fwd A R.
-/3 width=3 by eq_stream_sym/ qed-.
-
-(* Main properties **********************************************************)
-
-let corec eq_stream_trans: ∀A. Transitive … (eq_stream A) ≝ ?.
-#A #t1 #t * -t1 -t
-#t1 #t #b1 #b * #Ht1 * #b2 #t2 #H cases (eq_stream_inv_seq A … H) -H -b
-/3 width=7 by eq_seq/
-qed-.
-
-theorem eq_stream_canc_sn: ∀A,t,t1,t2. t ≐ t1 → t ≐ t2 → t1 ≐⦋A⦌ t2.
-/3 width=3 by eq_stream_trans, eq_stream_sym/ qed-.
-
-theorem eq_stream_canc_dx: ∀A,t,t1,t2. t1 ≐ t → t2 ≐ t → t1 ≐⦋A⦌ t2.
-/3 width=3 by eq_stream_trans, eq_stream_sym/ qed-.