X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_computation%2Fcsx.ma;h=da0bbe1ee7ef9a563bca8e2c6d39442be9ce4bb1;hb=5c186c72f508da0849058afeecc6877cd9ed6303;hp=49f22191178bc1d6d2418a4ed8f5e1889c707ffe;hpb=6167cca50de37eba76a062537b24f7caef5b34f2;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx.ma index 49f221911..da0bbe1ee 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx.ma @@ -16,24 +16,24 @@ include "basic_2/notation/relations/predtystrong_5.ma". include "basic_2/syntax/tdeq.ma". include "basic_2/rt_transition/cpx.ma". -(* STRONGLY NORMALIZING TERMS FOR UNCOUNTED PARALLEL RT-TRANSITION **********) +(* STRONGLY NORMALIZING TERMS FOR UNBOUND PARALLEL RT-TRANSITION ************) definition csx: ∀h. sd h → relation3 genv lenv term ≝ λh,o,G,L. SN … (cpx h G L) (tdeq h o …). interpretation - "strong normalization for uncounted context-sensitive parallel rt-transition (term)" + "strong normalization for unbound context-sensitive parallel rt-transition (term)" 'PRedTyStrong h o G L T = (csx h o G L T). (* Basic eliminators ********************************************************) -lemma csx_ind: ∀h,o,G,L. ∀R:predicate term. +lemma csx_ind: ∀h,o,G,L. ∀Q:predicate term. (∀T1. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃T1⦄ → - (∀T2. ⦃G, L⦄ ⊢ T1 ⬈[h] T2 → (T1 ≛[h, o] T2 → ⊥) → R T2) → - R T1 + (∀T2. ⦃G, L⦄ ⊢ T1 ⬈[h] T2 → (T1 ≛[h, o] T2 → ⊥) → Q T2) → + Q T1 ) → - ∀T. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃T⦄ → R T. -#h #o #G #L #R #H0 #T1 #H elim H -T1 + ∀T. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃T⦄ → Q T. +#h #o #G #L #Q #H0 #T1 #H elim H -T1 /5 width=1 by SN_intro/ qed-.