X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fstatic%2Ffrees_drops.ma;h=096683f573a8076cbf70e5b635516d4ab5424979;hb=990f97071a9939d47be16b36f6045d3b23f218e0;hp=0e1ae0e21c115e56f95f077c4c707164e7da37c3;hpb=21827c7db90ddb4fd30adec6becd94e201898f9c;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 0e1ae0e21..096683f57 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/frees_drops.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/frees_drops.ma @@ -154,6 +154,12 @@ lemma frees_lifts: ∀b,f1,K,T. K ⊢ 𝐅*⦃T⦄ ≡ f1 → ] qed-. +lemma frees_lifts_SO: ∀b,L,K. ⬇*[b, 𝐔❴1❵] L ≡ K → ∀T,U. ⬆*[1] T ≡ U → + ∀f. K ⊢ 𝐅*⦃T⦄ ≡ f → L ⊢ 𝐅*⦃U⦄ ≡ ↑f. +#b #L #K #HLK #T #U #HTU #f #Hf +@(frees_lifts b … Hf … HTU) // (**) (* auto fails *) +qed. + (* Forward lemmas with generic slicing for local environments ***************) lemma frees_fwd_coafter: ∀b,f2,L,U. L ⊢ 𝐅*⦃U⦄ ≡ f2 → @@ -185,6 +191,7 @@ lemma frees_inv_lifts: ∀b,f2,L,U. L ⊢ 𝐅*⦃U⦄ ≡ f2 → /3 width=7 by frees_eq_repl_back, coafter_inj/ qed-. +(* Note: this is used by lfxs_conf and might be modified *) lemma frees_inv_drops_next: ∀f1,L1,T1. L1 ⊢ 𝐅*⦃T1⦄ ≡ f1 → ∀I2,L2,V2,n. ⬇*[n] L1 ≡ L2.ⓑ{I2}V2 → ∀g1. ⫯g1 = ⫱*[n] f1 →