X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=inline;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fetc%2Fcpr%2Fcpr.etc;h=3f30e621f74f62615df3e83dff06506fd2b638f5;hb=2002da6bcdbf12203a87a7d9630d738f67ede68c;hp=d82299c382dbe0d02ffa948fa60f5d1ad396e198;hpb=713673ecf863cb6187291f016ed4490b12a03ac0;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/cpr.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/cpr.etc index d82299c38..3f30e621f 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/cpr.etc +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/cpr.etc @@ -1,111 +1,29 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/unfold/tpss.ma". -include "basic_2/reducibility/tpr.ma". - -(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS ****************************) - -(* Basic_1: includes: pr2_delta1 *) -definition cpr: lenv → relation term ≝ - λL,T1,T2. ∃∃T. T1 ➡ T & L ⊢ T ▶* [0, |L|] T2. - -interpretation - "context-sensitive parallel reduction (term)" - 'PRed L T1 T2 = (cpr L T1 T2). - -(* Basic properties *********************************************************) - -lemma cpr_intro: ∀L,T1,T,T2,d,e. T1 ➡ T → L ⊢ T ▶* [d, e] T2 → L ⊢ T1 ➡ T2. -/3 width=5/ qed-. - -(* Basic_1: was by definition: pr2_free *) -lemma cpr_tpr: ∀T1,T2. T1 ➡ T2 → ∀L. L ⊢ T1 ➡ T2. -/2 width=3/ qed. - -lemma cpr_tpss: ∀L,T1,T2,d,e. L ⊢ T1 ▶* [d, e] T2 → L ⊢ T1 ➡ T2. -/3 width=5/ qed. - -lemma cpr_refl: ∀L,T. L ⊢ T ➡ T. -/2 width=1/ qed. - -(* Note: new property *) -(* Basic_1: was only: pr2_thin_dx *) -lemma cpr_flat: ∀I,L,V1,V2,T1,T2. - L ⊢ V1 ➡ V2 → L ⊢ T1 ➡ T2 → L ⊢ ⓕ{I} V1. T1 ➡ ⓕ{I} V2. T2. -#I #L #V1 #V2 #T1 #T2 * #V #HV1 #HV2 * /3 width=5/ -qed. - -lemma cpr_cast: ∀L,V,T1,T2. - L ⊢ T1 ➡ T2 → L ⊢ ⓝV. T1 ➡ T2. -#L #V #T1 #T2 * /3 width=3/ -qed. - -(* Note: it does not hold replacing |L1| with |L2| *) -(* Basic_1: was only: pr2_change *) -lemma cpr_lsubr_trans: ∀L1,T1,T2. L1 ⊢ T1 ➡ T2 → - ∀L2. L2 ⊑ [0, |L1|] L1 → L2 ⊢ T1 ➡ T2. -#L1 #T1 #T2 * #T #HT1 #HT2 #L2 #HL12 -lapply (tpss_lsubr_trans … HT2 … HL12) -HT2 -HL12 /3 width=4/ -qed. - -(* Basic inversion lemmas ***************************************************) - -(* Basic_1: was: pr2_gen_csort *) -lemma cpr_inv_atom: ∀T1,T2. ⋆ ⊢ T1 ➡ T2 → T1 ➡ T2. -#T1 #T2 * #T #HT normalize #HT2 -<(tpss_inv_refl_O2 … HT2) -HT2 // -qed-. - -(* Basic_1: was: pr2_gen_sort *) -lemma cpr_inv_sort1: ∀L,T2,k. L ⊢ ⋆k ➡ T2 → T2 = ⋆k. -#L #T2 #k * #X #H ->(tpr_inv_atom1 … H) -H #H ->(tpss_inv_sort1 … H) -H // -qed-. - -(* Basic_1: was: pr2_gen_cast *) -lemma cpr_inv_cast1: ∀L,V1,T1,U2. L ⊢ ⓝV1. T1 ➡ U2 → ( - ∃∃V2,T2. L ⊢ V1 ➡ V2 & L ⊢ T1 ➡ T2 & - U2 = ⓝV2. T2 - ) ∨ L ⊢ T1 ➡ U2. -#L #V1 #T1 #U2 * #X #H #HU2 -elim (tpr_inv_cast1 … H) -H /3 width=3/ -* #V #T #HV1 #HT1 #H destruct -elim (tpss_inv_flat1 … HU2) -HU2 #V2 #T2 #HV2 #HT2 #H destruct /4 width=5/ -qed-. - -(* Basic forward lemmas *****************************************************) - -lemma cpr_fwd_bind1_minus: ∀I,L,V1,T1,T. L ⊢ -ⓑ{I}V1.T1 ➡ T → ∀b. - ∃∃V2,T2. L ⊢ ⓑ{b,I}V1.T1 ➡ ⓑ{b,I}V2.T2 & - T = -ⓑ{I}V2.T2. -#I #L #V1 #T1 #T * #X #H1 #H2 #b -elim (tpr_fwd_bind1_minus … H1 b) -H1 #V0 #T0 #HT10 #H destruct -elim (tpss_inv_bind1 … H2) -H2 #V2 #T2 #HV02 #HT02 #H destruct /4 width=5/ +lemma cpr_delift: ∀G,K,V,T1,L,l. ⬇[l] L ≡ (K.ⓓV) → + ∃∃T2,T. ⦃G, L⦄ ⊢ T1 ➡ T2 & ⬆[l, 1] T ≡ T2. +#G #K #V #T1 elim T1 -T1 +[ * /2 width=4 by cpr_atom, lift_sort, lift_gref, ex2_2_intro/ + #i #L #l #HLK elim (lt_or_eq_or_gt i l) + #Hil [1,3: /4 width=4 by lift_lref_ge_minus, lift_lref_lt, ylt_inj, yle_inj, ex2_2_intro/ ] + destruct + elim (lift_total V 0 (i+1)) #W #HVW + elim (lift_split … HVW i i) /3 width=6 by cpr_delta, ex2_2_intro/ +| * [ #a ] #I #W1 #U1 #IHW1 #IHU1 #L #l #HLK + elim (IHW1 … HLK) -IHW1 #W2 #W #HW12 #HW2 + [ elim (IHU1 (L. ⓑ{I}W1) (l+1)) -IHU1 /3 width=9 by drop_drop, cpr_bind, lift_bind, ex2_2_intro/ + | elim (IHU1 … HLK) -IHU1 -HLK /3 width=8 by cpr_flat, lift_flat, ex2_2_intro/ + ] +] qed-. -lemma cpr_fwd_shift1: ∀L,L1,T1,T. L ⊢ L1 @@ T1 ➡ T → - ∃∃L2,T2. |L1| = |L2| & T = L2 @@ T2. -#L #L1 #T1 #T * #X #H1 #H2 -elim (tpr_fwd_shift1 … H1) -H1 #L0 #T0 #HL10 #H destruct -elim (tpss_fwd_shift1 … H2) -H2 #L2 #T2 #HL02 #H destruct /2 width=4/ +fact lstas_cpr_aux: ∀h,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 •*[h, d] T2 → + d = 0 → ⦃G, L⦄ ⊢ T1 ➡ T2. +#h #G #L #T1 #T2 #d #H elim H -G -L -T1 -T2 -d +/3 width=1 by cpr_eps, cpr_flat, cpr_bind/ +[ #G #L #K #V1 #V2 #W2 #i #d #HLK #_ #HVW2 #IHV12 #H destruct + /3 width=6 by cpr_delta/ +| #G #L #K #V1 #V2 #W2 #i #d #_ #_ #_ #_