X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fstatic_2%2Fstatic%2Flsubf_lsubr.ma;fp=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fstatic_2%2Fstatic%2Flsubf_lsubr.ma;h=0ad58737df796940b5be39133f0c593e9677f5f1;hb=8ec019202bff90959cf1a7158b309e7f83fa222e;hp=3d4448a68811cf0c3f1c18b105d794ccce90a48e;hpb=33d0a7a9029859be79b25b5a495e0f30dab11f37;p=helm.git 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 3d4448a68..0ad58737d 100644 --- a/matita/matita/contribs/lambdadelta/static_2/static/lsubf_lsubr.ma +++ b/matita/matita/contribs/lambdadelta/static_2/static/lsubf_lsubr.ma @@ -20,7 +20,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❫ → 𝛀❪f1❫ → 𝛀❪f2❫ → L1 ⫃ L2. + ∀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, pr_isd_inv_next/ [ #f1 #f2 #I1 #I2 #L1 #L2 #_ #_ #H @@ -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❫. + ∀L1,L2. L1 ⫃ L2 → ∀f1,f2. 𝐈❨f1❩ → 𝐈❨f2❩ → ❨L1,f1❩ ⫃𝐅+ ❨L2,f2❩. #L1 #L2 #H elim H -L1 -L2 [ /3 width=1 by lsubf_atom, pr_isi_inv_eq_repl/ | #I #L1 #L2 | #L1 #L2 #V #W | #I1 #I2 #L1 #L2 #V @@ -45,8 +45,8 @@ elim (pr_isi_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❫. + ∀f2,L2,T. L2 ⊢ 𝐅+❨T❩ ≘ f2 → ∀L1. L1 ⫃ L2 → + ∀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/