]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/ground_2/lib/streams_eq.ma
notational update in ground_2 and basic_2
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / lib / streams_eq.ma
index 9435b501018ff372bc89d9c2d9dfde43a8215b5d..a9940337e5a852c7560197f83ccad36851d3dec0 100644 (file)
@@ -12,7 +12,7 @@
 (*                                                                        *)
 (**************************************************************************)
 
-include "ground_2/notation/relations/exteq_3.ma".
+include "ground_2/notation/relations/ringeq_3.ma".
 include "ground_2/lib/streams.ma".
 
 (* STREAMS ******************************************************************)
@@ -22,22 +22,22 @@ coinductive eq_stream (A): relation (stream A) ≝
 .
 
 interpretation "extensional equivalence (nstream)"
-   'ExtEq A t1 t2 = (eq_stream A t1 t2).
+   'RingEq A t1 t2 = (eq_stream A t1 t2).
 
 definition eq_stream_repl (A) (R:relation …) ≝
-                          â\88\80t1,t2. t1 â\89\90â¦\8bAâ¦\8c t2 → R t1 t2.
+                          â\88\80t1,t2. t1 â\89\97{A} t2 → R t1 t2.
 
 definition eq_stream_repl_back (A) (R:predicate …) ≝
-                               â\88\80t1. R t1 â\86\92 â\88\80t2. t1 â\89\90â¦\8bAâ¦\8c t2 → R t2.
+                               â\88\80t1. R t1 â\86\92 â\88\80t2. t1 â\89\97{A} t2 → R t2.
 
 definition eq_stream_repl_fwd (A) (R:predicate …) ≝
-                              â\88\80t1. R t1 â\86\92 â\88\80t2. t2 â\89\90â¦\8bAâ¦\8c t1 → R t2.
+                              â\88\80t1. R t1 â\86\92 â\88\80t2. t2 â\89\97{A} t1 → R t2.
 
 (* Basic inversion lemmas ***************************************************)
 
-lemma eq_stream_inv_seq: â\88\80A,t1,t2. t1 â\89\90â¦\8bAâ¦\8c t2 →
+lemma eq_stream_inv_seq: â\88\80A,t1,t2. t1 â\89\97{A} t2 →
                          ∀u1,u2,a1,a2. a1@u1 = t1 → a2@u2 = t2 →
-                         u1 â\89\90 u2 ∧ a1 = a2.
+                         u1 â\89\97 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-.
@@ -64,8 +64,8 @@ corec theorem eq_stream_trans: ∀A. Transitive … (eq_stream A).
 /3 width=7 by eq_seq/
 qed-.
 
-theorem eq_stream_canc_sn: â\88\80A,t,t1,t2. t â\89\90 t1 â\86\92 t â\89\90 t2 â\86\92 t1 â\89\90â¦\8bAâ¦\8c t2.
+theorem eq_stream_canc_sn: â\88\80A,t,t1,t2. t â\89\97 t1 â\86\92 t â\89\97 t2 â\86\92 t1 â\89\97{A} t2.
 /3 width=3 by eq_stream_trans, eq_stream_sym/ qed-.
 
-theorem eq_stream_canc_dx: â\88\80A,t,t1,t2. t1 â\89\90 t â\86\92 t2 â\89\90 t â\86\92 t1 â\89\90â¦\8bAâ¦\8c t2.
+theorem eq_stream_canc_dx: â\88\80A,t,t1,t2. t1 â\89\97 t â\86\92 t2 â\89\97 t â\86\92 t1 â\89\97{A} t2.
 /3 width=3 by eq_stream_trans, eq_stream_sym/ qed-.