]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/ground/relocation/tr_pap_eq.ma
update in ground
[helm.git] / matita / matita / contribs / lambdadelta / ground / relocation / tr_pap_eq.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 include "ground/lib/stream_eq.ma".
16 include "ground/relocation/tr_pap.ma".
17
18 (* POSITIVE APPLICATION FOR TOTAL RELOCATION MAPS ***************************)
19
20 (* Main constructions with stream_eq ****************************************)
21
22 (*** apply_eq_repl *)
23 theorem tr_pap_eq_repl (i):
24         stream_eq_repl … (λf1,f2. f1@⧣❨i❩ = f2@⧣❨i❩).
25 #i elim i -i [2: #i #IH ] * #p1 #f1 * #p2 #f2 #H
26 elim (stream_eq_inv_cons_bi … H) -H [1,8: |*: // ] #Hp #Hf //
27 <tr_cons_pap_succ <tr_cons_pap_succ /3 width=1 by eq_f2/
28 qed.
29
30 (* Main inversions with stream_eq *******************************************)
31
32 corec theorem nstream_eq_inv_ext:
33               ∀f1,f2. (∀i. f1@⧣❨i❩ = f2@⧣❨i❩) → f1 ≗ f2.
34 * #p1 #f1 * #p2 #f2 #Hf @stream_eq_cons
35 [ @(Hf (𝟏))
36 | @nstream_eq_inv_ext -nstream_eq_inv_ext #i
37   lapply (Hf (𝟏)) <tr_cons_pap_unit <tr_cons_pap_unit #H destruct
38   lapply (Hf (↑i)) <tr_cons_pap_succ <tr_cons_pap_succ #H
39   /3 width=2 by eq_inv_pplus_bi_dx, eq_inv_psucc_bi/
40 ]
41 qed-.