]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/ground/relocation/rtmap_after_uni.ma
propagating the arithmetics library, partial commit
[helm.git] / matita / matita / contribs / lambdadelta / ground / relocation / rtmap_after_uni.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/arith/nat_plus.ma".
16 include "ground/relocation/rtmap_uni.ma".
17 include "ground/relocation/rtmap_after.ma".
18
19 (* RELOCATION MAP ***********************************************************)
20
21 (* Properties on isuni ******************************************************)
22
23 lemma after_isid_isuni: ∀f1,f2. 𝐈❪f2❫ → 𝐔❪f1❫ → f1 ⊚ ↑f2 ≘ ↑f1.
24 #f1 #f2 #Hf2 #H elim H -H
25 /5 width=7 by after_isid_dx, after_eq_repl_back2, after_next, after_push, eq_push_inv_isid/
26 qed.
27
28 lemma after_uni_next2: ∀f2. 𝐔❪f2❫ → ∀f1,f. ↑f2 ⊚ f1 ≘ f → f2 ⊚ ↑f1 ≘ f.
29 #f2 #H elim H -f2
30 [ #f2 #Hf2 #f1 #f #Hf
31   elim (after_inv_nxx … Hf) -Hf [2,3: // ] #g #Hg #H0 destruct
32   /4 width=7 by after_isid_inv_sn, after_isid_sn, after_eq_repl_back0, eq_next/
33 | #f2 #_ #g2 #H2 #IH #f1 #f #Hf
34   elim (after_inv_nxx … Hf) -Hf [2,3: // ] #g #Hg #H0 destruct
35   /3 width=5 by after_next/
36 ]
37 qed.
38
39 (* Properties on uni ********************************************************)
40
41 lemma after_uni: ∀n1,n2. 𝐔❨n1❩ ⊚ 𝐔❨n2❩ ≘ 𝐔❨n2+n1❩.
42 #n1 @(nat_ind_succ … n1) -n1
43 /3 width=5 by after_isid_sn, after_next, eq_f/
44 qed.
45
46 lemma after_uni_sn_pushs (m) (f): 𝐔❨m❩ ⊚ f ≘ ↑*[m]f.
47 #m @(nat_ind_succ … m) -m
48 /2 width=5 by after_isid_sn, after_next/
49 qed.
50
51 (* Properties with at *******************************************************)
52
53 lemma after_uni_succ_dx: ∀i2,i1,f2. @❪i1, f2❫ ≘ i2 →
54                          ∀f. f2 ⊚ 𝐔❨i1❩ ≘ f → 𝐔❨i2❩ ⊚ ⫱*[i2] f2 ≘ f.
55 #i2 elim i2 -i2
56 [ #i1 #f2 #Hf2 #f #Hf
57   elim (at_inv_xxp … Hf2) -Hf2 // #g2 #H1 #H2 destruct
58   elim (after_inv_pnx … Hf) -Hf [ |*: // ] #g #Hg #H
59   lapply (after_isid_inv_dx … Hg ?) -Hg
60   /4 width=5 by after_isid_sn, after_eq_repl_back0, after_next/
61 | #i2 #IH #i1 #f2 #Hf2 #f #Hf >nsucc_inj
62   elim (at_inv_xxn … Hf2) -Hf2 [1,3: * |*: // ]
63   [ #g2 #j1 #Hg2 #H1 #H2 destruct >nsucc_inj in Hf; #Hf
64     elim (after_inv_pnx … Hf) -Hf [ |*: // ] #g #Hg #H destruct
65     <tls_xn /3 width=5 by after_next/
66   | #g2 #Hg2 #H2 destruct
67     elim (after_inv_nxx … Hf) -Hf [2,3: // ] #g #Hg #H destruct
68     <tls_xn /3 width=5 by after_next/
69   ]
70 ]
71 qed.
72
73 lemma after_uni_succ_sn: ∀i2,i1,f2. @❪i1, f2❫ ≘ i2 →
74                          ∀f. 𝐔❨i2❩ ⊚ ⫱*[i2] f2 ≘ f → f2 ⊚ 𝐔❨i1❩ ≘ f.
75 #i2 elim i2 -i2
76 [ #i1 #f2 #Hf2 #f #Hf
77   elim (at_inv_xxp … Hf2) -Hf2 // #g2 #H1 #H2 destruct
78   elim (after_inv_nxx … Hf) -Hf [ |*: // ] #g #Hg #H destruct
79   lapply (after_isid_inv_sn … Hg ?) -Hg
80   /4 width=7 by after_isid_dx, after_eq_repl_back0, after_push/
81 | #i2 #IH #i1 #f2 #Hf2 #f >nsucc_inj #Hf
82   elim (after_inv_nxx … Hf) -Hf [2,3: // ] #g #Hg #H destruct
83   elim (at_inv_xxn … Hf2) -Hf2 [1,3: * |*: // ]
84   [ #g2 #j1 #Hg2 #H1 #H2 destruct <tls_xn in Hg; /3 width=7 by after_push/
85   | #g2 #Hg2 #H2 destruct <tls_xn in Hg; /3 width=5 by after_next/
86   ]
87 ]
88 qed-.
89
90 lemma after_uni_one_dx: ∀f2,f. ⫯f2 ⊚ 𝐔❨𝟏❩ ≘ f → 𝐔❨𝟏❩ ⊚ f2 ≘ f.
91 #f2 #f #H @(after_uni_succ_dx … (⫯f2)) /2 width=3 by at_refl/
92 qed.
93
94 lemma after_uni_one_sn: ∀f1,f. 𝐔❨𝟏❩ ⊚ f1 ≘ f → ⫯f1 ⊚ 𝐔❨𝟏❩ ≘ f.
95 /3 width=3 by after_uni_succ_sn, at_refl/ qed-.