X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fsubstitution%2Ffqu.ma;h=48e230f00845a38d9fbaa9976b256da7be2a6f99;hb=52e675f555f559c047d5449db7fc89a51b977d35;hp=651ca43221e15080fc45c7e001791218213ba37f;hpb=75fac6d60f67a4dfa38ea6c2cc45a18eda5d8996;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/fqu.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/fqu.ma index 651ca4322..48e230f00 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/fqu.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/substitution/fqu.ma @@ -14,7 +14,7 @@ include "basic_2/notation/relations/supterm_6.ma". include "basic_2/grammar/cl_weight.ma". -include "basic_2/substitution/ldrop.ma". +include "basic_2/substitution/drop.ma". (* SUPCLOSURE ***************************************************************) @@ -40,7 +40,7 @@ lemma fqu_drop_lt: ∀G,L,K,T,U,e. 0 < e → qed. lemma fqu_lref_S_lt: ∀I,G,L,V,i. 0 < i → ⦃G, L.ⓑ{I}V, #i⦄ ⊐ ⦃G, L, #(i-1)⦄. -/3 width=3 by fqu_drop, ldrop_drop, lift_lref_ge_minus/ +/3 width=3 by fqu_drop, drop_drop, lift_lref_ge_minus/ qed. (* Basic forward lemmas *****************************************************) @@ -48,7 +48,7 @@ qed. lemma fqu_fwd_fw: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → ♯{G2, L2, T2} < ♯{G1, L1, T1}. #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 // #G #L #K #T #U #e #HLK #HTU -lapply (ldrop_fwd_lw_lt … HLK ?) -HLK // #HKL +lapply (drop_fwd_lw_lt … HLK ?) -HLK // #HKL lapply (lift_fwd_tw … HTU) -e #H normalize in ⊢ (?%%); /2 width=1 by lt_minus_to_plus/ qed-. @@ -58,7 +58,7 @@ fact fqu_fwd_length_lref1_aux: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 [1: normalize // |3: #a -|5: /2 width=4 by ldrop_fwd_length_lt4/ +|5: /2 width=4 by drop_fwd_length_lt4/ ] #I #G #L #V #T #j #H destruct qed-.