X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fstatic_2%2Fstatic%2Flsubf_lsubr.ma;h=0b0124fee7fed6a399a115abb3ac01bcf1321151;hp=a021ca5a4ac0af24c8252c302944310af7d535c8;hb=f308429a0fde273605a2330efc63268b4ac36c99;hpb=87f57ddc367303c33e19c83cd8989cd561f3185b diff --git a/matita/matita/contribs/lambdadelta/static_2/static/lsubf_lsubr.ma b/matita/matita/contribs/lambdadelta/static_2/static/lsubf_lsubr.ma index a021ca5a4..0b0124fee 100644 --- a/matita/matita/contribs/lambdadelta/static_2/static/lsubf_lsubr.ma +++ b/matita/matita/contribs/lambdadelta/static_2/static/lsubf_lsubr.ma @@ -19,7 +19,7 @@ include "static_2/static/lsubf_lsubf.ma". (* Forward lemmas with restricted refinement for local environments *********) -lemma lsubf_fwd_lsubr_isdiv: ∀f1,f2,L1,L2. ⦃L1, f1⦄ ⫃𝐅* ⦃L2, f2⦄ → +lemma lsubf_fwd_lsubr_isdiv: ∀f1,f2,L1,L2. ⦃L1,f1⦄ ⫃𝐅* ⦃L2,f2⦄ → 𝛀⦃f1⦄ → 𝛀⦃f2⦄ → L1 ⫃ L2. #f1 #f2 #L1 #L2 #H elim H -f1 -f2 -L1 -L2 /4 width=3 by lsubr_bind, isdiv_inv_next/ @@ -33,7 +33,7 @@ qed-. (* Properties with restricted refinement for local environments *************) lemma lsubr_lsubf_isid: ∀L1,L2. L1 ⫃ L2 → - ∀f1,f2. 𝐈⦃f1⦄ → 𝐈⦃f2⦄ → ⦃L1, f1⦄ ⫃𝐅* ⦃L2, f2⦄. + ∀f1,f2. 𝐈⦃f1⦄ → 𝐈⦃f2⦄ → ⦃L1,f1⦄ ⫃𝐅* ⦃L2,f2⦄. #L1 #L2 #H elim H -L1 -L2 [ /3 width=1 by lsubf_atom, isid_inv_eq_repl/ | #I #L1 #L2 | #L1 #L2 #V #W | #I1 #I2 #L1 #L2 #V @@ -45,7 +45,7 @@ elim (isid_inv_gen … Hf2) -Hf2 #g2 #Hg2 #H destruct qed. lemma lsubr_lsubf: ∀f2,L2,T. L2 ⊢ 𝐅*⦃T⦄ ≘ f2 → ∀L1. L1 ⫃ L2 → - ∀f1. L1 ⊢ 𝐅*⦃T⦄ ≘ f1 → ⦃L1, f1⦄ ⫃𝐅* ⦃L2, f2⦄. + ∀f1. L1 ⊢ 𝐅*⦃T⦄ ≘ f1 → ⦃L1,f1⦄ ⫃𝐅* ⦃L2,f2⦄. #f2 #L2 #T #H elim H -f2 -L2 -T [ #f2 #L2 #s #Hf2 #L1 #HL12 #f1 #Hf1 lapply (frees_inv_sort … Hf1) -Hf1 /2 width=1 by lsubr_lsubf_isid/