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/tr_pap.ma".
18 (* POSITIVE APPLICATION FOR TOTAL RELOCATION MAPS ***************************)
20 (* Main constructions with stream_eq ****************************************)
22 (* Note: a better statement would be: tr_eq_repl … (λf1,f2. f1@❨i❩ = f2@❨i❩) *)
24 theorem apply_eq_repl (i):
25 ∀f1,f2. f1 ≗ f2 → f1@❨i❩ = f2@❨i❩.
26 #i elim i -i [2: #i #IH ] * #p1 #f1 * #p2 #f2 #H
27 elim (stream_eq_inv_cons_bi … H) -H [1,8: |*: // ] #Hp #Hf //
28 <tr_pap_succ <tr_pap_succ /3 width=1 by eq_f2/
31 (* Main inversions with stream_eq *******************************************)
33 corec theorem nstream_eq_inv_ext:
34 ∀f1,f2. (∀i. f1@❨i❩ = f2@❨i❩) → f1 ≗ f2.
35 * #p1 #f1 * #p2 #f2 #Hf @stream_eq_cons
37 | @nstream_eq_inv_ext -nstream_eq_inv_ext #i
38 lapply (Hf (𝟏)) <tr_pap_unit <tr_pap_unit #H destruct
39 lapply (Hf (↑i)) <tr_pap_succ <tr_pap_succ #H
40 /3 width=2 by eq_inv_pplus_bi_dx, eq_inv_psucc_bi/