X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Frelocation%2Frtmap_after.ma;h=15969a23c46edb255b980c8b2d79c6a40767a72c;hb=3ef251397627da80aeea0cf08b053a4bc781ef88;hp=222bc07f7b5f383c5ebed44e7b25921ffd390e3c;hpb=93bba1c94779e83184d111cd077d4167e42a74aa;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_after.ma b/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_after.ma index 222bc07f7..15969a23c 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_after.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_after.ma @@ -320,6 +320,24 @@ lemma after_isid_isuni: ∀f1,f2. 𝐈⦃f2⦄ → 𝐔⦃f1⦄ → f1 ⊚ ⫯f2 /5 width=7 by isid_after_dx, after_eq_repl_back_2, after_next, after_push, eq_push_inv_isid/ qed. +lemma after_uni_next2: ∀f2. 𝐔⦃f2⦄ → ∀f1,f. ⫯f2 ⊚ f1 ≡ f → f2 ⊚ ⫯f1 ≡ f. +#f2 #H elim H -f2 +[ #f2 #Hf2 #f1 #f #Hf + elim (after_inv_nxx … Hf) -Hf [2,3: // ] #g #Hg #H0 destruct + /4 width=7 by after_isid_inv_sn, isid_after_sn, after_eq_repl_back_0, eq_next/ +| #f2 #_ #g2 #H2 #IH #f1 #f #Hf + elim (after_inv_nxx … Hf) -Hf [2,3: // ] #g #Hg #H0 destruct + /3 width=5 by after_next/ +] +qed. + +(* Properties on uni ********************************************************) + +lemma after_uni: ∀n1,n2. 𝐔❴n1❵ ⊚ 𝐔❴n2❵ ≡ 𝐔❴n1+n2❵. +@nat_elim2 +/4 width=5 by after_uni_next2, isid_after_sn, isid_after_dx, after_next/ +qed. + (* Forward lemmas on at *****************************************************) lemma after_at_fwd: ∀i,i1,f. @⦃i1, f⦄ ≡ i → ∀f2,f1. f2 ⊚ f1 ≡ f → @@ -383,8 +401,8 @@ qed-. (* Properties with at *******************************************************) -lemma after_isuni_dx: ∀i2,i1,f2. @⦃i1, f2⦄ ≡ i2 → - ∀f. f2 ⊚ 𝐔❴i1❵ ≡ f → 𝐔❴i2❵ ⊚ ⫱*[i2] f2 ≡ f. +lemma after_uni_dx: ∀i2,i1,f2. @⦃i1, f2⦄ ≡ i2 → + ∀f. f2 ⊚ 𝐔❴i1❵ ≡ f → 𝐔❴i2❵ ⊚ ⫱*[i2] f2 ≡ f. #i2 elim i2 -i2 [ #i1 #f2 #Hf2 #f #Hf elim (at_inv_xxp … Hf2) -Hf2 // #g2 #H1 #H2 destruct @@ -402,8 +420,8 @@ lemma after_isuni_dx: ∀i2,i1,f2. @⦃i1, f2⦄ ≡ i2 → ] qed. -lemma after_isuni_sn: ∀i2,i1,f2. @⦃i1, f2⦄ ≡ i2 → - ∀f. 𝐔❴i2❵ ⊚ ⫱*[i2] f2 ≡ f → f2 ⊚ 𝐔❴i1❵ ≡ f. +lemma after_uni_sn: ∀i2,i1,f2. @⦃i1, f2⦄ ≡ i2 → + ∀f. 𝐔❴i2❵ ⊚ ⫱*[i2] f2 ≡ f → f2 ⊚ 𝐔❴i1❵ ≡ f. #i2 elim i2 -i2 [ #i1 #f2 #Hf2 #f #Hf elim (at_inv_xxp … Hf2) -Hf2 // #g2 #H1 #H2 destruct