X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fcomputation%2Ffpbu_lift.ma;h=a76af75e4a65051997dc4a1716e7bb1d259db49f;hb=ff7754f834f937bfe2384c7703cf63f552885395;hp=d7b48650e3fd18853aa929839feb2b04e3306adf;hpb=4720368dcf18593959c6d21484f62fb5b61f3d26;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/fpbu_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/fpbu_lift.ma index d7b48650e..a76af75e4 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/fpbu_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/fpbu_lift.ma @@ -12,7 +12,7 @@ (* *) (**************************************************************************) -include "basic_2/static/ssta_ssta.ma". +include "basic_2/static/sta_sta.ma". include "basic_2/computation/cpxs_lift.ma". include "basic_2/computation/fpbu.ma". @@ -20,13 +20,13 @@ include "basic_2/computation/fpbu.ma". (* Advanced properties ******************************************************) -lemma lsstas_fpbu: ∀h,g,G,L,T1,T2,l2. ⦃G, L⦄ ⊢ T1 •*[h, g, l2] T2 → (T1 = T2 → ⊥) → - ∀l1. l2 ≤ l1 → ⦃G, L⦄ ⊢ T1 ▪[h, g] l1 → ⦃G, L, T1⦄ ≻[h, g] ⦃G, L, T2⦄. -/4 width=5 by fpbu_cpxs, lsstas_cpxs/ qed. +lemma lstas_fpbu: ∀h,g,G,L,T1,T2,l2. ⦃G, L⦄ ⊢ T1 •*[h, l2] T2 → (T1 = T2 → ⊥) → + ∀l1. l2 ≤ l1 → ⦃G, L⦄ ⊢ T1 ▪[h, g] l1 → ⦃G, L, T1⦄ ≻[h, g] ⦃G, L, T2⦄. +/4 width=5 by fpbu_cpxs, lstas_cpxs/ qed. -lemma ssta_fpbu: ∀h,g,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 ▪[h, g] l+1 → - ⦃G, L⦄ ⊢ T1 •[h, g] T2 → ⦃G, L, T1⦄ ≻[h, g] ⦃G, L, T2⦄. +lemma sta_fpbu: ∀h,g,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 ▪[h, g] l+1 → + ⦃G, L⦄ ⊢ T1 •[h] T2 → ⦃G, L, T1⦄ ≻[h, g] ⦃G, L, T2⦄. #h #g #G #L #T1 #T2 #l #HT1 #HT12 elim (eq_term_dec T1 T2) -/3 width=5 by ssta_lsstas, lsstas_fpbu/ #H destruct -elim (ssta_inv_refl_pos … HT1 … HT12) +/3 width=5 by sta_lstas, lstas_fpbu/ #H destruct +elim (sta_inv_refl_pos … HT1 … HT12) qed.