X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Freduction%2Fcpr.ma;h=ad4fcb3135663f1c02fd4ed5f5d0a825be5ab376;hb=52e675f555f559c047d5449db7fc89a51b977d35;hp=7fcfa3f323aed2e9eecdcad047978d4a41f32d23;hpb=2ba2dc23443ad764adab652e06d6f5ed10bd912d;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr.ma index 7fcfa3f32..ad4fcb313 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr.ma @@ -14,9 +14,7 @@ include "basic_2/notation/relations/pred_4.ma". include "basic_2/grammar/genv.ma". -include "basic_2/grammar/cl_shift.ma". -include "basic_2/relocation/ldrop_append.ma". -include "basic_2/relocation/lsubr.ma". +include "basic_2/static/lsubr.ma". (* CONTEXT-SENSITIVE PARALLEL REDUCTION FOR TERMS ***************************) @@ -36,7 +34,7 @@ inductive cpr: relation4 genv lenv term term ≝ cpr G L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2) | cpr_zeta : ∀G,L,V,T1,T,T2. cpr G (L.ⓓV) T1 T → ⇧[0, 1] T2 ≡ T → cpr G L (+ⓓV.T1) T2 -| cpr_tau : ∀G,L,V,T1,T2. cpr G L T1 T2 → cpr G L (ⓝV.T1) T2 +| cpr_eps : ∀G,L,V,T1,T2. cpr G L T1 T2 → cpr G L (ⓝV.T1) T2 | cpr_beta : ∀a,G,L,V1,V2,W1,W2,T1,T2. cpr G L V1 V2 → cpr G L W1 W2 → cpr G (L.ⓛW1) T1 T2 → cpr G L (ⓐV1.ⓛ{a}W1.T1) (ⓓ{a}ⓝW2.V2.T2) @@ -54,10 +52,10 @@ lemma lsubr_cpr_trans: ∀G. lsub_trans … (cpr G) lsubr. #G #L1 #T1 #T2 #H elim H -G -L1 -T1 -T2 [ // | #G #L1 #K1 #V1 #V2 #W2 #i #HLK1 #_ #HVW2 #IHV12 #L2 #HL12 - elim (lsubr_fwd_ldrop2_abbr … HL12 … HLK1) -L1 * + elim (lsubr_fwd_drop2_abbr … HL12 … HLK1) -L1 * /3 width=6 by cpr_delta/ |3,7: /4 width=1 by lsubr_bind, cpr_bind, cpr_beta/ -|4,6: /3 width=1 by cpr_flat, cpr_tau/ +|4,6: /3 width=1 by cpr_flat, cpr_eps/ |5,8: /4 width=3 by lsubr_bind, cpr_zeta, cpr_theta/ ] qed-. @@ -89,21 +87,12 @@ lemma cpr_delift: ∀G,K,V,T1,L,d. ⇩[d] L ≡ (K.ⓓV) → elim (lift_split … HVW i i) /3 width=6 by cpr_delta, ex2_2_intro/ | * [ #a ] #I #W1 #U1 #IHW1 #IHU1 #L #d #HLK elim (IHW1 … HLK) -IHW1 #W2 #W #HW12 #HW2 - [ elim (IHU1 (L. ⓑ{I}W1) (d+1)) -IHU1 /3 width=9 by ldrop_drop, cpr_bind, lift_bind, ex2_2_intro/ + [ elim (IHU1 (L. ⓑ{I}W1) (d+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_append: ∀G. l_appendable_sn … (cpr G). -#G #K #T1 #T2 #H elim H -G -K -T1 -T2 -/2 width=3 by cpr_bind, cpr_flat, cpr_zeta, cpr_tau, cpr_beta, cpr_theta/ -#G #K #K0 #V1 #V2 #W2 #i #HK0 #_ #HVW2 #IHV12 #L -lapply (ldrop_fwd_length_lt2 … HK0) #H -@(cpr_delta … (L@@K0) V1 … HVW2) // -@(ldrop_O1_append_sn_le … HK0) /2 width=2 by lt_to_le/ (**) (* /3/ does not work *) -qed. - (* Basic inversion lemmas ***************************************************) fact cpr_inv_atom1_aux: ∀G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ∀I. T1 = ⓪{I} → @@ -294,23 +283,6 @@ elim (cpr_inv_bind1 … H) -H * ] qed-. -lemma cpr_fwd_shift1: ∀G,L1,L,T1,T. ⦃G, L⦄ ⊢ L1 @@ T1 ➡ T → - ∃∃L2,T2. |L1| = |L2| & T = L2 @@ T2. -#G #L1 @(lenv_ind_dx … L1) -L1 normalize -[ #L #T1 #T #HT1 - @(ex2_2_intro … (⋆)) // (**) (* explicit constructor *) -| #I #L1 #V1 #IH #L #T1 #X - >shift_append_assoc normalize #H - elim (cpr_inv_bind1 … H) -H * - [ #V0 #T0 #_ #HT10 #H destruct - elim (IH … HT10) -IH -HT10 #L2 #T2 #HL12 #H destruct - >append_length >HL12 -HL12 - @(ex2_2_intro … (⋆.ⓑ{I}V0@@L2) T2) [ >append_length ] /2 width=3 by trans_eq/ (**) (* explicit constructor *) - | #T #_ #_ #H destruct - ] -] -qed-. - (* Basic_1: removed theorems 11: pr0_subst0_back pr0_subst0_fwd pr0_subst0 pr2_head_2 pr2_cflat clear_pr2_trans @@ -318,6 +290,6 @@ qed-. pr2_gen_ctail pr2_ctail *) (* Basic_1: removed local theorems 4: - pr0_delta_tau pr0_cong_delta + pr0_delta_eps pr0_cong_delta pr2_free_free pr2_free_delta *)