X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_computation%2Fcpxs_tdeq.ma;h=10433b26797d273bbd9d27195818649a7cbbe4e6;hb=54c9014b6657403c6e235c652176218e750d4b8a;hp=c3e2e20658368c7d0185d4dbbd60f891b2b7b3dd;hpb=b11e1907f99bea1de50db890d849ba5469d2e0e7;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_tdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_tdeq.ma index c3e2e2065..10433b267 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_tdeq.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_tdeq.ma @@ -12,23 +12,23 @@ (* *) (**************************************************************************) -include "basic_2/syntax/tdeq_tdeq.ma". +include "basic_2/rt_transition/lfpx_lfdeq.ma". include "basic_2/rt_computation/cpxs.ma". -include "basic_2/rt_transition/cpx_lfdeq.ma". -include "basic_2/rt_transition/lfpx_fqup.ma". (* UNCOUNTED CONTEXT-SENSITIVE PARALLEL RT-COMPUTATION FOR TERMS ************) -lemma tdeq_cpxs_trans: ∀h,o,U1,T1. U1 ≡[h, o] T1 → ∀G,L,T2. ⦃G, L⦄ ⊢ T1 ⬈*[h] T2 → - ∃∃U2. ⦃G, L⦄ ⊢ U1 ⬈*[h] U2 & U2 ≡[h, o] T2. +(* Properties with degree-based equivalence for terms ***********************) + +lemma tdeq_cpxs_trans: ∀h,o,U1,T1. U1 ≛[h, o] T1 → ∀G,L,T2. ⦃G, L⦄ ⊢ T1 ⬈*[h] T2 → + ∃∃U2. ⦃G, L⦄ ⊢ U1 ⬈*[h] U2 & U2 ≛[h, o] T2. #h #o #U1 #T1 #HUT1 #G #L #T2 #HT12 @(cpxs_ind … HT12) -T2 /2 width=3 by ex2_intro/ #T #T2 #_ #HT2 * #U #HU1 #HUT elim (tdeq_cpx_trans … HUT … HT2) -T -T1 /3 width=3 by ex2_intro, cpxs_strap1/ qed-. (* Note: this requires tdeq to be symmetric *) -lemma cpxs_tdneq_inv_step_sn: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ⬈*[h] T2 → (T1 ≡[h, o] T2 → ⊥) → - ∃∃T,T0. ⦃G, L⦄ ⊢ T1 ⬈[h] T & T1 ≡[h, o] T → ⊥ & ⦃G, L⦄ ⊢ T ⬈*[h] T0 & T0 ≡[h, o] T2. +lemma cpxs_tdneq_inv_step_sn: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ⬈*[h] T2 → (T1 ≛[h, o] T2 → ⊥) → + ∃∃T,T0. ⦃G, L⦄ ⊢ T1 ⬈[h] T & T1 ≛[h, o] T → ⊥ & ⦃G, L⦄ ⊢ T ⬈*[h] T0 & T0 ≛[h, o] T2. #h #o #G #L #T1 #T2 #H @(cpxs_ind_dx … H) -T1 [ #H elim H -H // | #T1 #T #H1 #H2 #IH #Hn12 elim (tdeq_dec h o T1 T) #H destruct @@ -40,5 +40,3 @@ lemma cpxs_tdneq_inv_step_sn: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ⬈*[h] T2 → ] ] qed-. - -(* Basic_2A1: removed theorems 1: cpxs_neq_inv_step_sn *)