X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fsubstitution%2Fcpys_alt.ma;h=3fb179ed2922a458f7caf116f0086eca03b050ca;hb=e5378812c068074f0ac627541d3f066e135c8824;hp=929007118c8ec26e2eaee90a189f6fc3a7c3462d;hpb=928cfe1ebf2fbd31731c8851cdec70802596016d;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/cpys_alt.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/cpys_alt.ma index 929007118..3fb179ed2 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/cpys_alt.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/substitution/cpys_alt.ma @@ -21,7 +21,7 @@ include "basic_2/substitution/cpys_lift.ma". inductive cpysa: ynat → ynat → relation4 genv lenv term term ≝ | cpysa_atom : ∀I,G,L,d,e. cpysa d e G L (⓪{I}) (⓪{I}) | cpysa_subst: ∀I,G,L,K,V1,V2,W2,i,d,e. d ≤ yinj i → i < d+e → - ⇩[0, i] L ≡ K.ⓑ{I}V1 → cpysa 0 (⫰(d+e-i)) G K V1 V2 → + ⇩[i] L ≡ K.ⓑ{I}V1 → cpysa 0 (⫰(d+e-i)) G K V1 V2 → ⇧[0, i+1] V2 ≡ W2 → cpysa d e G L (#i) W2 | cpysa_bind : ∀a,I,G,L,V1,V2,T1,T2,d,e. cpysa d e G L V1 V2 → cpysa (⫯d) e G (L.ⓑ{I}V2) T1 T2 → @@ -58,7 +58,7 @@ lemma cpysa_cpy_trans: ∀G,L,T1,T,d,e. ⦃G, L⦄ ⊢ T1 ▶▶*×[d, e] T → [ #I #G #L #d #e #X #H elim (cpy_inv_atom1 … H) -H // * /2 width=7 by cpysa_subst/ | #I #G #L #K #V1 #V2 #W2 #i #d #e #Hdi #Hide #HLK #_ #HVW2 #IHV12 #T2 #H - lapply (ldrop_fwd_ldrop2 … HLK) #H0LK + lapply (ldrop_fwd_drop2 … HLK) #H0LK lapply (cpy_weak … H 0 (d+e) ? ?) -H // #H elim (cpy_inv_lift1_be … H … H0LK … HVW2) -H -H0LK -HVW2 /3 width=7 by cpysa_subst, ylt_fwd_le_succ/ @@ -86,7 +86,7 @@ qed-. lemma cpys_ind_alt: ∀R:ynat→ynat→relation4 genv lenv term term. (∀I,G,L,d,e. R d e G L (⓪{I}) (⓪{I})) → (∀I,G,L,K,V1,V2,W2,i,d,e. d ≤ yinj i → i < d + e → - ⇩[O, i] L ≡ K.ⓑ{I}V1 → ⦃G, K⦄ ⊢ V1 ▶*×[O, ⫰(d+e-i)] V2 → + ⇩[i] L ≡ K.ⓑ{I}V1 → ⦃G, K⦄ ⊢ V1 ▶*×[O, ⫰(d+e-i)] V2 → ⇧[O, i+1] V2 ≡ W2 → R O (⫰(d+e-i)) G K V1 V2 → R d e G L (#i) W2 ) → (∀a,I,G,L,V1,V2,T1,T2,d,e. ⦃G, L⦄ ⊢ V1 ▶*×[d, e] V2 →