X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fs_computation%2Ffqup_drops.ma;h=f2db58c837d1aaae401ac8d7798f01a70dcbb97a;hp=f669ff80e2ad3c30d2740678edfd12905a709ff1;hb=222044da28742b24584549ba86b1805a87def070;hpb=98fbba1b68d457807c73ebf70eb2a48696381da4 diff --git a/matita/matita/contribs/lambdadelta/basic_2/s_computation/fqup_drops.ma b/matita/matita/contribs/lambdadelta/basic_2/s_computation/fqup_drops.ma index f669ff80e..f2db58c83 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/s_computation/fqup_drops.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/s_computation/fqup_drops.ma @@ -19,7 +19,7 @@ include "basic_2/s_computation/fqup.ma". (* Properties with generic slicing for local environments *******************) -lemma fqup_drops_succ: ∀b,G,K,T,i,L,U. ⬇*[⫯i] L ≡ K → ⬆*[⫯i] T ≡ U → +lemma fqup_drops_succ: ∀b,G,K,T,i,L,U. ⬇*[↑i] L ≘ K → ⬆*[↑i] T ≘ U → ⦃G, L, U⦄ ⊐+[b] ⦃G, K, T⦄. #b #G #K #T #i elim i -i [ #L #U #HLK #HTU elim (drops_inv_succ … HLK) -HLK @@ -27,12 +27,12 @@ lemma fqup_drops_succ: ∀b,G,K,T,i,L,U. ⬇*[⫯i] L ≡ K → ⬆*[⫯i] T ≡ /3 width=2 by fqu_fqup, fqu_drop/ | #l #IH #L #U #HLK #HTU elim (drops_inv_succ … HLK) -HLK #I #Y #HY #H destruct - elim (lifts_split_trans … HTU … (𝐔❴⫯l❵) (𝐔❴1❵)) -HTU + elim (lifts_split_trans … HTU … (𝐔❴↑l❵) (𝐔❴1❵)) -HTU /4 width=5 by fqup_strap2, fqu_drop/ ] qed. -lemma fqup_drops_strap1: ∀b,G1,G2,L1,K1,K2,T1,T2,U1,i. ⬇*[i] L1 ≡ K1 → ⬆*[i] T1 ≡ U1 → +lemma fqup_drops_strap1: ∀b,G1,G2,L1,K1,K2,T1,T2,U1,i. ⬇*[i] L1 ≘ K1 → ⬆*[i] T1 ≘ U1 → ⦃G1, K1, T1⦄ ⊐[b] ⦃G2, K2, T2⦄ → ⦃G1, L1, U1⦄ ⊐+[b] ⦃G2, K2, T2⦄. #b #G1 #G2 #L1 #K1 #K2 #T1 #T2 #U1 * [ #HLK1 #HTU1 #HT12 @@ -42,5 +42,5 @@ lemma fqup_drops_strap1: ∀b,G1,G2,L1,K1,K2,T1,T2,U1,i. ⬇*[i] L1 ≡ K1 → ] qed-. -lemma fqup_lref: ∀b,I,G,L,K,V,i. ⬇*[i] L ≡ K.ⓑ{I}V → ⦃G, L, #i⦄ ⊐+[b] ⦃G, K, V⦄. +lemma fqup_lref: ∀b,I,G,L,K,V,i. ⬇*[i] L ≘ K.ⓑ{I}V → ⦃G, L, #i⦄ ⊐+[b] ⦃G, K, V⦄. /2 width=6 by fqup_drops_strap1/ qed.