X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Frelocation%2Frtmap_after.ma;fp=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Frelocation%2Frtmap_after.ma;h=80401a3d14ab99b9e016c4cbcfb5225212db6e71;hb=9a023f554e56d6edbbb2eeaf17ce61e31857ef4a;hp=fdfefbba524dd9f9d0c75a4f3675ad63914c4f8c;hpb=e28ddccd4096c80b2090ca78af00e2590f629b71;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 fdfefbba5..80401a3d1 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_after.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_after.ma @@ -14,7 +14,7 @@ include "ground_2/notation/relations/rafter_3.ma". include "ground_2/relocation/rtmap_istot.ma". -include "ground_2/relocation/rtmap_isuni.ma". +include "ground_2/relocation/rtmap_uni.ma". (* RELOCATION MAP ***********************************************************) @@ -251,7 +251,7 @@ corec theorem after_trans2: ∀f1,f0,f4. f1 ⊚ f0 ≡ f4 → ] qed-. -(* Main inversion lemmas on after *******************************************) +(* Main inversion lemmas ****************************************************) corec theorem after_mono: ∀f1,f2,x,y. f1 ⊚ f2 ≡ x → f1 ⊚ f2 ≡ y → x ≗ y. #f1 #f2 #x #y * -f1 -f2 -x @@ -382,6 +382,43 @@ lemma after_fwd_at1: ∀i,i2,i1,f,f2. @⦃i1, f⦄ ≡ i → @⦃i2, f2⦄ ≡ i ] qed-. +(* Properties with at *******************************************************) + +lemma after_isuni_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 + lapply (after_isid_inv_dx … Hf ?) -Hf + /3 width=3 by isid_after_sn, after_eq_repl_back_0/ +| #i2 #IH #i1 #f2 #Hf2 #f #Hf + elim (at_inv_xxn … Hf2) -Hf2 [1,3: * |*: // ] + [ #g2 #j1 #Hg2 #H1 #H2 destruct + elim (after_inv_pnx … Hf) -Hf [ |*: // ] #g #Hg #H destruct + /3 width=5 by after_next/ + | #g2 #Hg2 #H2 destruct + elim (after_inv_nxx … Hf) -Hf [2,3: // ] #g #Hg #H destruct + /3 width=5 by after_next/ + ] +] +qed. + +lemma after_isuni_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 + lapply (after_isid_inv_sn … Hf ?) -Hf + /3 width=3 by isid_after_dx, after_eq_repl_back_0/ +| #i2 #IH #i1 #f2 #Hf2 #f #Hf + elim (after_inv_nxx … Hf) -Hf [2,3: // ] #g #Hg #H destruct + elim (at_inv_xxn … Hf2) -Hf2 [1,3: * |*: // ] + [ #g2 #j1 #Hg2 #H1 #H2 destruct /3 width=7 by after_push/ + | #g2 #Hg2 #H2 destruct /3 width=5 by after_next/ + ] +] +qed-. + (* Forward lemmas on istot **************************************************) lemma after_istot_fwd: ∀f2,f1,f. f2 ⊚ f1 ≡ f → 𝐓⦃f2⦄ → 𝐓⦃f1⦄ → 𝐓⦃f⦄.