X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frelocation%2Fdrops_vector.ma;h=6a7560248f692ed7c04de6d7ed21ea61f4966c62;hp=63bc71da6d090c3ad6dec5b4802891b400b5da5a;hb=222044da28742b24584549ba86b1805a87def070;hpb=952ec5aa2e9a54787acb63a5c8d6fdbf9011ab60 diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_vector.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_vector.ma index 63bc71da6..6a7560248 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_vector.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_vector.ma @@ -15,17 +15,18 @@ include "basic_2/relocation/lifts_vector.ma". include "basic_2/relocation/drops.ma". -(* GENERAL SLICING FOR LOCAL ENVIRONMENTS ***********************************) +(* GENERIC SLICING FOR LOCAL ENVIRONMENTS ***********************************) -definition d_liftable1_all: relation2 lenv term → predicate bool ≝ - λR,s. ∀L,K,t. ⬇*[s, t] L ≡ K → - ∀Ts,Us. ⬆*[t] Ts ≡ Us → - all … (R K) Ts → all … (R L) Us. +definition d_liftable1_all: predicate (relation2 lenv term) ≝ + λR. ∀K,Ts. all … (R K) Ts → + ∀b,f,L. ⬇*[b, f] L ≘ K → + ∀Us. ⬆*[f] Ts ≘ Us → all … (R L) Us. -(* Properties on general relocation for term vectors ************************) +(* Properties with generic relocation for term vectors **********************) (* Basic_2A1: was: d1_liftables_liftables_all *) -lemma d1_liftable_liftable_all: ∀R,s. d_liftable1 R s → d_liftable1_all R s. -#R #s #HR #L #K #t #HLK #Ts #Us #H elim H -Ts -Us normalize // +lemma d1_liftable_liftable_all: ∀R. d_liftable1 R → d_liftable1_all R. +#R #HR #K #Ts #HTs #b #f #L #HLK #Us #H +generalize in match HTs; -HTs elim H -Ts -Us normalize // #Ts #Us #T #U #HTU #_ #IHTUs * /3 width=7 by conj/ qed.