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_lsubr.ma;h=b969425727c0842eb4c1d6f8fa4fb63e5e0134bb;hp=169970ed12ca49d1f68b9353850544fdcd89a787;hb=4173283e148199871d787c53c0301891deb90713;hpb=a67fc50ccfda64377e2c94c18c3a0d9265f651db diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_lsubr.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_lsubr.ma index 169970ed1..b96942572 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_lsubr.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_lsubr.ma @@ -19,9 +19,9 @@ include "basic_2/rt_computation/csx_csx.ma". (* Advanced properties ******************************************************) -fact csx_appl_beta_aux: ∀h,o,p,G,L,U1. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃U1⦄ → - ∀V,W,T1. U1 = ⓓ{p}ⓝW.V.T1 → ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃ⓐV.ⓛ{p}W.T1⦄. -#h #o #p #G #L #X #H @(csx_ind … H) -X +fact csx_appl_beta_aux: ∀h,p,G,L,U1. ⦃G, L⦄ ⊢ ⬈*[h] 𝐒⦃U1⦄ → + ∀V,W,T1. U1 = ⓓ{p}ⓝW.V.T1 → ⦃G, L⦄ ⊢ ⬈*[h] 𝐒⦃ⓐV.ⓛ{p}W.T1⦄. +#h #p #G #L #X #H @(csx_ind … H) -X #X #HT1 #IHT1 #V #W #T1 #H1 destruct @csx_intro #X #H1 #H2 elim (cpx_inv_appl1 … H1) -H1 * @@ -42,23 +42,23 @@ elim (cpx_inv_appl1 … H1) -H1 * qed-. (* Basic_1: was just: sn3_beta *) -lemma csx_appl_beta: ∀h,o,p,G,L,V,W,T. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃ⓓ{p}ⓝW.V.T⦄ → ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃ⓐV.ⓛ{p}W.T⦄. +lemma csx_appl_beta: ∀h,p,G,L,V,W,T. ⦃G, L⦄ ⊢ ⬈*[h] 𝐒⦃ⓓ{p}ⓝW.V.T⦄ → ⦃G, L⦄ ⊢ ⬈*[h] 𝐒⦃ⓐV.ⓛ{p}W.T⦄. /2 width=3 by csx_appl_beta_aux/ qed. (* Advanced forward lemmas **************************************************) -fact csx_fwd_bind_dx_unit_aux: ∀h,o,G,L,U. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃U⦄ → - ∀p,I,J,V,T. U = ⓑ{p,I}V.T → ⦃G, L.ⓤ{J}⦄ ⊢ ⬈*[h, o] 𝐒⦃T⦄. -#h #o #G #L #U #H elim H -H #U0 #_ #IH #p #I #J #V #T #H destruct +fact csx_fwd_bind_dx_unit_aux: ∀h,G,L,U. ⦃G, L⦄ ⊢ ⬈*[h] 𝐒⦃U⦄ → + ∀p,I,J,V,T. U = ⓑ{p,I}V.T → ⦃G, L.ⓤ{J}⦄ ⊢ ⬈*[h] 𝐒⦃T⦄. +#h #G #L #U #H elim H -H #U0 #_ #IH #p #I #J #V #T #H destruct @csx_intro #T2 #HLT2 #HT2 @(IH (ⓑ{p,I}V.T2)) -IH /2 width=4 by cpx_bind_unit/ -HLT2 #H elim (tdeq_inv_pair … H) -H /2 width=1 by/ qed-. -lemma csx_fwd_bind_dx_unit: ∀h,o,p,I,G,L,V,T. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃ⓑ{p,I}V.T⦄ → - ∀J. ⦃G, L.ⓤ{J}⦄ ⊢ ⬈*[h, o] 𝐒⦃T⦄. +lemma csx_fwd_bind_dx_unit: ∀h,p,I,G,L,V,T. ⦃G, L⦄ ⊢ ⬈*[h] 𝐒⦃ⓑ{p,I}V.T⦄ → + ∀J. ⦃G, L.ⓤ{J}⦄ ⊢ ⬈*[h] 𝐒⦃T⦄. /2 width=6 by csx_fwd_bind_dx_unit_aux/ qed-. -lemma csx_fwd_bind_unit: ∀h,o,p,I,G,L,V,T. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃ⓑ{p,I}V.T⦄ → - ∀J. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃V⦄ ∧ ⦃G, L.ⓤ{J}⦄ ⊢ ⬈*[h, o] 𝐒⦃T⦄. +lemma csx_fwd_bind_unit: ∀h,p,I,G,L,V,T. ⦃G, L⦄ ⊢ ⬈*[h] 𝐒⦃ⓑ{p,I}V.T⦄ → + ∀J. ⦃G, L⦄ ⊢ ⬈*[h] 𝐒⦃V⦄ ∧ ⦃G, L.ⓤ{J}⦄ ⊢ ⬈*[h] 𝐒⦃T⦄. /3 width=4 by csx_fwd_pair_sn, csx_fwd_bind_dx_unit, conj/ qed-.