X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2FBasic_2%2Funfold%2Fltpss_ltpss.ma;h=e4291acf91e315b7342ceb50215dcb64f8616001;hb=c4ac63d7ae22b2adcc7fe7b54286a0226296eabc;hp=9a023056df89054b48f1319c4878230ad8d9ea87;hpb=78f21d7d9014e5c7655f58239e4f1a128ea2c558;p=helm.git diff --git a/matita/matita/contribs/lambda_delta/Basic_2/unfold/ltpss_ltpss.ma b/matita/matita/contribs/lambda_delta/Basic_2/unfold/ltpss_ltpss.ma index 9a023056d..e4291acf9 100644 --- a/matita/matita/contribs/lambda_delta/Basic_2/unfold/ltpss_ltpss.ma +++ b/matita/matita/contribs/lambda_delta/Basic_2/unfold/ltpss_ltpss.ma @@ -20,14 +20,14 @@ include "Basic_2/unfold/ltpss_tpss.ma". theorem ltpss_trans_eq: ∀L1,L,L2,d,e. L1 [d, e] ≫* L → L [d, e] ≫* L2 → L1 [d, e] ≫* L2. -/2/ qed. +/2 width=3/ qed. lemma ltpss_tpss2: ∀L1,L2,I,V1,V2,e. L1 [0, e] ≫* L2 → L2 ⊢ V1 [0, e] ≫* V2 → L1. 𝕓{I} V1 [0, e + 1] ≫* L2. 𝕓{I} V2. #L1 #L2 #I #V1 #V2 #e #HL12 #H @(tpss_ind … H) -V2 -[ /2/ -| #V #V2 #_ #HV2 #IHV @(ltpss_trans_eq … IHV) /2/ +[ /2 width=1/ +| #V #V2 #_ #HV2 #IHV @(ltpss_trans_eq … IHV) /2 width=1/ ] qed. @@ -35,15 +35,15 @@ lemma ltpss_tpss2_lt: ∀L1,L2,I,V1,V2,e. L1 [0, e - 1] ≫* L2 → L2 ⊢ V1 [0, e - 1] ≫* V2 → 0 < e → L1. 𝕓{I} V1 [0, e] ≫* L2. 𝕓{I} V2. #L1 #L2 #I #V1 #V2 #e #HL12 #HV12 #He ->(plus_minus_m_m e 1) /2/ +>(plus_minus_m_m e 1) // /2 width=1/ qed. lemma ltpss_tpss1: ∀L1,L2,I,V1,V2,d,e. L1 [d, e] ≫* L2 → L2 ⊢ V1 [d, e] ≫* V2 → L1. 𝕓{I} V1 [d + 1, e] ≫* L2. 𝕓{I} V2. #L1 #L2 #I #V1 #V2 #d #e #HL12 #H @(tpss_ind … H) -V2 -[ /2/ -| #V #V2 #_ #HV2 #IHV @(ltpss_trans_eq … IHV) /2/ +[ /2 width=1/ +| #V #V2 #_ #HV2 #IHV @(ltpss_trans_eq … IHV) /2 width=1/ ] qed. @@ -51,5 +51,5 @@ lemma ltpss_tpss1_lt: ∀L1,L2,I,V1,V2,d,e. L1 [d - 1, e] ≫* L2 → L2 ⊢ V1 [d - 1, e] ≫* V2 → 0 < d → L1. 𝕓{I} V1 [d, e] ≫* L2. 𝕓{I} V2. #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #Hd ->(plus_minus_m_m d 1) /2/ +>(plus_minus_m_m d 1) // /2 width=1/ qed.