X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_computation%2Fcsx_simple.ma;h=95cb186f5f096b0ee9d804e72a07370d177c69c0;hp=efa69cff7a0af3861ea6579fd6166c15da866fcd;hb=4173283e148199871d787c53c0301891deb90713;hpb=a67fc50ccfda64377e2c94c18c3a0d9265f651db diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_simple.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_simple.ma index efa69cff7..95cb186f5 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_simple.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_simple.ma @@ -19,10 +19,10 @@ include "basic_2/rt_computation/csx_csx.ma". (* Properties with simple terms *********************************************) -lemma csx_appl_simple: ∀h,o,G,L,V. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃V⦄ → ∀T1. - (∀T2. ⦃G, L⦄ ⊢ T1 ⬈[h] T2 → (T1 ≛[h, o] T2 → ⊥) → ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃ⓐV.T2⦄) → - 𝐒⦃T1⦄ → ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃ⓐV.T1⦄. -#h #o #G #L #V #H @(csx_ind … H) -V +lemma csx_appl_simple: ∀h,G,L,V. ⦃G, L⦄ ⊢ ⬈*[h] 𝐒⦃V⦄ → ∀T1. + (∀T2. ⦃G, L⦄ ⊢ T1 ⬈[h] T2 → (T1 ≛ T2 → ⊥) → ⦃G, L⦄ ⊢ ⬈*[h] 𝐒⦃ⓐV.T2⦄) → + 𝐒⦃T1⦄ → ⦃G, L⦄ ⊢ ⬈*[h] 𝐒⦃ⓐV.T1⦄. +#h #G #L #V #H @(csx_ind … H) -V #V #_ #IHV #T1 #IHT1 #HT1 @csx_intro #X #H1 #H2 elim (cpx_inv_appl1_simple … H1) // -H1