]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfsx.ma
partial update update in basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_computation / lfsx.ma
index bf17640ddde9733852d02b281e361b83a2082825..955d70edb036ff21b1f56179a62b0f52971328b1 100644 (file)
@@ -16,13 +16,13 @@ include "basic_2/notation/relations/predtysnstrong_5.ma".
 include "basic_2/static/lfdeq.ma".
 include "basic_2/rt_transition/lfpx.ma".
 
-(* STRONGLY NORMALIZING LOCAL ENV.S FOR UNCOUNTED PARALLEL RT-TRANSITION ****)
+(* STRONGLY NORMALIZING LOCAL ENV.S FOR UNBOUND PARALLEL RT-TRANSITION ******)
 
 definition lfsx: ∀h. sd h → relation3 term genv lenv ≝
                  λh,o,T,G. SN … (lfpx h G T) (lfdeq h o T).
 
 interpretation
-   "strong normalization for uncounted context-sensitive parallel rt-transition on referred entries (local environment)"
+   "strong normalization for unbound context-sensitive parallel rt-transition on referred entries (local environment)"
    'PRedTySNStrong h o T G L = (lfsx h o T G L).
 
 (* Basic eliminators ********************************************************)
@@ -35,7 +35,7 @@ lemma lfsx_ind: ∀h,o,G,T. ∀R:predicate lenv.
                 ) →
                 ∀L. G ⊢ ⬈*[h, o, T] 𝐒⦃L⦄ → R L.
 #h #o #G #T #R #H0 #L1 #H elim H -L1
-/5 width=1 by lfdeq_sym, SN_intro/
+/5 width=1 by SN_intro/
 qed-.
 
 (* Basic properties *********************************************************)
@@ -44,7 +44,7 @@ qed-.
 lemma lfsx_intro: ∀h,o,G,L1,T.
                   (∀L2. ⦃G, L1⦄ ⊢ ⬈[h, T] L2 → (L1 ≛[h, o, T] L2 → ⊥) → G ⊢ ⬈*[h, o, T] 𝐒⦃L2⦄) →
                   G ⊢ ⬈*[h, o, T] 𝐒⦃L1⦄.
-/5 width=1 by lfdeq_sym, SN_intro/ qed.
+/5 width=1 by SN_intro/ qed.
 
 (* Basic_2A1: uses: lsx_sort *)
 lemma lfsx_sort: ∀h,o,G,L,s. G ⊢ ⬈*[h, o, ⋆s] 𝐒⦃L⦄.