X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_computation%2Fcsx_cpxs.ma;h=cfc4d6314e2e6890394e5385181c96ecfdddc299;hb=d02c188ee3d4bd9885490447e63453adb2cb4ea1;hp=87a5e890f968b644cc6e99eca8ebfa902f3c5e6c;hpb=6167cca50de37eba76a062537b24f7caef5b34f2;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_cpxs.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_cpxs.ma index 87a5e890f..cfc4d6314 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_cpxs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_cpxs.ma @@ -16,9 +16,9 @@ include "basic_2/rt_computation/cpxs_tdeq.ma". include "basic_2/rt_computation/cpxs_cpxs.ma". include "basic_2/rt_computation/csx_csx.ma". -(* UNCOUNTED CONTEXT-SENSITIVE PARALLEL RT-COMPUTATION FOR TERMS ************) +(* UNBOUND CONTEXT-SENSITIVE PARALLEL RT-COMPUTATION FOR TERMS **************) -(* Properties with uncounted context-sensitive rt-computation for terms *****) +(* Properties with unbound context-sensitive rt-computation for terms *******) (* Basic_1: was just: sn3_intro *) lemma csx_intro_cpxs: ∀h,o,G,L,T1. @@ -33,7 +33,7 @@ lemma csx_cpxs_trans: ∀h,o,G,L,T1. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃T1⦄ → /2 width=3 by csx_cpx_trans/ qed-. -(* Eliminators with uncounted context-sensitive rt-computation for terms ****) +(* Eliminators with unbound context-sensitive rt-computation for terms ******) lemma csx_ind_cpxs_tdeq: ∀h,o,G,L. ∀R:predicate term. (∀T1. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃T1⦄ → @@ -50,9 +50,9 @@ lapply (tdneq_tdeq_canc_dx … H … HV02) -H #HnTV0 elim (tdeq_dec h o T1 T0) #H [ lapply (tdeq_tdneq_trans … H … HnTV0) -H -HnTV0 #Hn10 lapply (cpxs_trans … HT10 … HTV0) -T0 #H10 - elim (cpxs_tdneq_inv_step_sn … H10 … Hn10) -H10 -Hn10 + elim (cpxs_tdneq_fwd_step_sn … H10 … Hn10) -H10 -Hn10 /3 width=8 by tdeq_trans/ -| elim (cpxs_tdneq_inv_step_sn … HT10 … H) -HT10 -H #T #V #HT1 #HnT1 #HTV #HVT0 +| elim (cpxs_tdneq_fwd_step_sn … HT10 … H) -HT10 -H #T #V #HT1 #HnT1 #HTV #HVT0 elim (tdeq_cpxs_trans … HVT0 … HTV0) -T0 /3 width=8 by cpxs_trans, tdeq_trans/ ]