]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/rt_computation/lsubsx_lfsx.ma
update in basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_computation / lsubsx_lfsx.ma
index 4ef9cf96238e41a1c213148fc011d42563e6c2fe..511e62369c283ff57dd919168505966b09066a33 100644 (file)
@@ -16,9 +16,9 @@ include "basic_2/rt_computation/lfsx_drops.ma".
 include "basic_2/rt_computation/lfsx_lfpxs.ma".
 include "basic_2/rt_computation/lsubsx.ma".
 
-(* CLEAR OF STRONGLY NORMALIZING ENTRIES FOR UNCOUNTED RT-TRANSITION ********)
+(* CLEAR OF STRONGLY NORMALIZING ENTRIES FOR UNBOUND RT-TRANSITION **********)
 
-(* Properties with strong normalizing env's for uncounted rt-transition *****)
+(* Properties with strong normalizing env's for unbound rt-transition *******)
 
 (* Basic_2A1: uses: lsx_cpx_trans_lcosx *)
 lemma lfsx_cpx_trans_lsubsx: ∀h,o,G,L0,T1,T2. ⦃G, L0⦄ ⊢ T1 ⬈[h] T2 →
@@ -62,3 +62,12 @@ qed-.
 lemma lfsx_cpx_trans: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ⬈[h] T2 →
                       G ⊢ ⬈*[h, o, T1] 𝐒⦃L⦄ → G ⊢ ⬈*[h, o, T2] 𝐒⦃L⦄.
 /3 width=6 by lfsx_cpx_trans_lsubsx, lsubsx_refl/ qed-.
+
+(* Properties with strong normalizing env's for unbound rt-computation ******)
+
+lemma lfsx_cpxs_trans: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ⬈*[h] T2 →
+                       G ⊢ ⬈*[h, o, T1] 𝐒⦃L⦄ → G ⊢ ⬈*[h, o, T2] 𝐒⦃L⦄.
+#h #o #G #L #T1 #T2 #H
+@(cpxs_ind_dx ???????? H) -T1 //
+/3 width=3 by lfsx_cpx_trans/
+qed-.