1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "ground_2/notation/relations/exteq_3.ma".
16 include "ground_2/lib/streams.ma".
18 (* STREAMS ******************************************************************)
20 coinductive eq_stream (A): relation (stream A) ≝
21 | eq_seq: ∀t1,t2,b1,b2. b1 = b2 → eq_stream A t1 t2 → eq_stream A (b1@t1) (b2@t2)
24 interpretation "extensional equivalence (nstream)"
25 'ExtEq A t1 t2 = (eq_stream A t1 t2).
27 definition eq_stream_repl (A) (R:relation …) ≝
28 ∀t1,t2. t1 ≐⦋A⦌ t2 → R t1 t2.
30 definition eq_stream_repl_back (A) (R:predicate …) ≝
31 ∀t1. R t1 → ∀t2. t1 ≐⦋A⦌ t2 → R t2.
33 definition eq_stream_repl_fwd (A) (R:predicate …) ≝
34 ∀t1. R t1 → ∀t2. t2 ≐⦋A⦌ t1 → R t2.
36 (* Basic inversion lemmas ***************************************************)
38 lemma eq_stream_inv_seq: ∀A,t1,t2. t1 ≐⦋A⦌ t2 →
39 ∀u1,u2,a1,a2. a1@u1 = t1 → a2@u2 = t2 →
42 #t1 #t2 #b1 #b2 #Hb #Ht #u1 #u2 #a1 #a2 #H1 #H2 destruct /2 width=1 by conj/
45 (* Basic properties *********************************************************)
47 corec lemma eq_stream_refl: ∀A. reflexive … (eq_stream A).
51 corec lemma eq_stream_sym: ∀A. symmetric … (eq_stream A).
53 #t1 #t2 #b1 #b2 #Hb #Ht @eq_seq /2 width=1 by/
56 lemma eq_stream_repl_sym: ∀A,R. eq_stream_repl_back A R → eq_stream_repl_fwd A R.
57 /3 width=3 by eq_stream_sym/ qed-.
59 (* Main properties **********************************************************)
61 corec theorem eq_stream_trans: ∀A. Transitive … (eq_stream A).
63 #t1 #t #b1 #b * #Ht1 * #b2 #t2 #H cases (eq_stream_inv_seq A … H) -H -b
67 theorem eq_stream_canc_sn: ∀A,t,t1,t2. t ≐ t1 → t ≐ t2 → t1 ≐⦋A⦌ t2.
68 /3 width=3 by eq_stream_trans, eq_stream_sym/ qed-.
70 theorem eq_stream_canc_dx: ∀A,t,t1,t2. t1 ≐ t → t2 ≐ t → t1 ≐⦋A⦌ t2.
71 /3 width=3 by eq_stream_trans, eq_stream_sym/ qed-.