X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fstatic%2Ffrees_drops.ma;h=681e07d8c701474386cf75c5abb2cea7da2a7155;hb=f4e73c50acfc4ed453edd423c9bbe28af5dc9c4c;hp=aa4f34998ffe6732e4d870dec140ff248daaf5ff;hpb=926796df5884453d8f0cf9f294d7776d469ef45b;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/frees_drops.ma b/matita/matita/contribs/lambdadelta/basic_2/static/frees_drops.ma index aa4f34998..681e07d8c 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/frees_drops.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/frees_drops.ma @@ -12,7 +12,6 @@ (* *) (**************************************************************************) -include "ground_2/relocation/rtmap_pushs.ma". include "ground_2/relocation/rtmap_coafter.ma". include "basic_2/relocation/drops_drops.ma". include "basic_2/static/frees.ma". @@ -38,6 +37,30 @@ lemma frees_lref_pair: ∀f,K,V. K ⊢ 𝐅*⦃V⦄ ≡ f → ] qed. +lemma frees_sort_pushs: ∀f,K,s. K ⊢ 𝐅*⦃⋆s⦄ ≡ f → + ∀i,L. ⬇*[i] L ≡ K → L ⊢ 𝐅*⦃⋆s⦄ ≡ ↑*[i] f. +#f #K #s #Hf #i elim i -i +[ #L #H lapply (drops_fwd_isid … H ?) -H // +| #i #IH #L #H elim (drops_inv_succ … H) -H /3 width=1 by frees_sort/ +] +qed. + +lemma frees_lref_pushs: ∀f,K,j. K ⊢ 𝐅*⦃#j⦄ ≡ f → + ∀i,L. ⬇*[i] L ≡ K → L ⊢ 𝐅*⦃#(i+j)⦄ ≡ ↑*[i] f. +#f #K #j #Hf #i elim i -i +[ #L #H lapply (drops_fwd_isid … H ?) -H // +| #i #IH #L #H elim (drops_inv_succ … H) -H /3 width=1 by frees_lref/ +] +qed. + +lemma frees_gref_pushs: ∀f,K,l. K ⊢ 𝐅*⦃§l⦄ ≡ f → + ∀i,L. ⬇*[i] L ≡ K → L ⊢ 𝐅*⦃§l⦄ ≡ ↑*[i] f. +#f #K #l #Hf #i elim i -i +[ #L #H lapply (drops_fwd_isid … H ?) -H // +| #i #IH #L #H elim (drops_inv_succ … H) -H /3 width=1 by frees_gref/ +] +qed. + (* Advanced inversion lemmas ************************************************) lemma frees_inv_lref_drops: ∀i,f,L. L ⊢ 𝐅*⦃#i⦄ ≡ f → @@ -58,28 +81,6 @@ qed-. (* Properties with generic slicing for local environments *******************) -axiom coafter_inv_xpx: ∀g2,f1,g. g2 ~⊚ ↑f1 ≡ g → ∀n. @⦃0, g2⦄ ≡ n → - ∃∃f2,f. f2 ~⊚ f1 ≡ f & ⫱*[n]g2 = ↑f2 & ⫱*[n]g = ↑f. -(* -#g2 #g1 #g #Hg #n #Hg2 -lapply (coafter_tls … Hg2 … Hg) -Hg #Hg -lapply (at_pxx_tls … Hg2) -Hg2 #H -elim (at_inv_pxp … H) -H [ |*: // ] #f2 #H2 -elim (coafter_inv_pxx … Hg … H2) -Hg * #f1 #f #Hf #H1 #H0 destruct -plus_S1 /2 width=3 by frees_lref_pushs/ (**) (* full auto fails *) +| #f1 #I #K #V #l #_ #IH #Hf1 #f #L #H1 #U #H2 #f2 #H3 + lapply (isfin_fwd_push … Hf1 ??) -Hf1 [3: |*: // ] #Hf1 + lapply (lifts_inv_gref1 … H2) -H2 #H destruct + elim (drops_split_trans_pair2 … H1) -H1 [ |*: // ] #Y #W #HLY #HYK #_ + elim (coafter_fwd_xpx_pushs … H3) [ |*: // ] #g2 #H2 destruct + lapply (coafter_tls_succ … H3 ??) -H3 [3: |*: // ] #H3 + lapply (IH … HYK … H3) -IH -H3 -HYK [1,3: // | skip ] + /3 width=5 by drops_isuni_fwd_drop2, frees_gref_pushs/ | #f1V #f1T #f1 #p #I #K #V #T #_ #_ #H1f1 #IHV #IHT #H2f1 #f #L #H1 #Y #H2 #f2 #H3 elim (sor_inv_isfin3 … H1f1) // #Hf1V #H lapply (isfin_inv_tl … H) -H @@ -142,6 +141,7 @@ lemma frees_lifts: ∀b,f1,K,T. K ⊢ 𝐅*⦃T⦄ ≡ f1 → elim (coafter_sor … H3 … H1f1) /3 width=5 by coafter_isfin2_fwd, frees_flat/ ] +qed-. (* Inversion lemmas with generic slicing for local environments *************)