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/lib/stream_eq.ma".
16 include "ground/relocation/rtmap_eq.ma".
18 (* RELOCATION P-STREAM ******************************************************)
20 (* Properties (specific) ******************************************************)
22 fact eq_inv_seq_aux: ∀f1,f2,p1,p2. p1⨮f1 ≡ p2⨮f2 → p1 = p2 ∧ f1 ≡ f2.
23 #f1 #f2 #p1 #p2 @(pnat_ind_2 … p1 p2) -p1 -p2
24 [ #p2 #H elim (gr_eq_inv_push_sn … H) -H [2,3: // ]
25 #g1 #H1 #H elim (push_inv_seq_dx … H) -H /2 width=1 by conj/
26 | #p1 #_ #H elim (gr_eq_inv_next_push … H) -H //
27 | #p1 #p2 #IH #H lapply (gr_eq_inv_next_bi … H ????) -H [5:|*: // ]
28 #H elim (IH H) -IH -H /2 width=1 by conj/
32 lemma eq_inv_seq: ∀g1,g2. g1 ≡ g2 → ∀f1,f2,p1,p2. p1⨮f1 = g1 → p2⨮f2 = g2 →
34 /2 width=1 by eq_inv_seq_aux/ qed-.
36 corec lemma nstream_eq: ∀f1,f2. f1 ≡ f2 → f1 ≗ f2.
37 * #p1 #f1 * #p2 #f2 #Hf cases (gr_eq_inv_gen … Hf) -Hf *
39 [ cases (push_inv_seq_dx … H1) -H1 * -p1 #H1
40 cases (push_inv_seq_dx … H2) -H2 * -p2 #H2
41 @stream_eq_cons /2 width=1 by/
42 | cases (next_inv_seq_dx … H1) -H1 #m1 #H1 * -p1
43 cases (next_inv_seq_dx … H2) -H2 #m2 #H2 * -p2
44 cases (eq_inv_seq … Hg … H1 H2) -g1 -g2 #Hm #Hf
45 @stream_eq_cons /2 width=1 by/
49 corec lemma nstream_inv_eq: ∀f1,f2. f1 ≗ f2 → f1 ≡ f2.
50 * #p1 #f1 * #p2 #f2 #H cases (stream_eq_inv_cons ??? H) -H [|*: // ]
51 #Hf * -p2 cases p1 -p1 /3 width=5 by gr_eq_next/
52 #n @gr_eq_next /3 width=5 by stream_eq_cons/
55 lemma eq_seq_id: ∀f1,f2. f1 ≡ f2 → ∀n. n⨮f1 ≡ n⨮f2.
56 /4 width=1 by nstream_inv_eq, nstream_eq, stream_eq_cons/ qed.