X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=sidebyside;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fsubstitution%2Flleq_llor.ma;h=8cbff4cd9bd01bc47e071619ccd4445403242bfc;hb=c671743a83bbc7fff114e3e3694f628c0ec6685b;hp=80bc596f2d94589a11c3c49f7e55ba15b744ca31;hpb=a8e31c02eefecdcd7d8a893c9f0a036a30fa57e4;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/lleq_llor.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/lleq_llor.ma index 80bc596f2..8cbff4cd9 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/lleq_llor.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/substitution/lleq_llor.ma @@ -18,19 +18,17 @@ include "basic_2/substitution/lleq_alt.ma". (* LAZY EQUIVALENCE FOR LOCAL ENVIRONMENTS **********************************) (* Properties on poinwise union for local environments **********************) -(* -lemma llor_lleq_O: ∀R,L1,L2,T. llpx_sn R 0 T L1 L2 → - ∀L. L1 ⩖[T] L2 ≡ L → L2 ≡[T, 0] L. -#R #L1 #L2 #T #H1 #L #H2 + +lemma llpx_sn_llor_dx: ∀R,L1,L2. + (∀U,i. L2 ⊢ i ϵ 𝐅*[0]⦃U⦄ → L1 ⊢ i ϵ 𝐅*[0]⦃U⦄) → + ∀T. llpx_sn R 0 T L1 L2 → ∀L. L1 ⩖[T] L2 ≡ L → L2 ≡[T, 0] L. +#R #L1 #L2 #HR #T #H1 #L #H2 elim (llpx_sn_llpx_sn_alt … H1) -H1 #HL12 #IH1 elim H2 -H2 #_ #HL1 #IH2 @lleq_intro_alt // #I2 #I #K2 #K #V2 #V #i #Hi #HnT #HLK2 #HLK lapply (ldrop_fwd_length_lt2 … HLK) #HiL elim (ldrop_O1_lt (Ⓕ) L1 i) // -HiL #I1 #K1 #V1 #HLK1 -elim (IH1 … HLK1 HLK2) // -IH1 -[ #H #_ destruct - elim (IH2 … HLK1 HLK2 HLK) * /2 width=1 by conj/ - #H elim HnT -HnT -| #H @HnT -HnT -] -*) +elim (IH1 … HLK1 HLK2) -IH1 /2 width=1 by/ #H #_ destruct +elim (IH2 … HLK1 HLK2 HLK) -IH2 -HLK1 -HLK2 -HLK * /2 width=1 by conj/ #H +elim H -H /2 width=1 by/ +qed.