X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frelocation%2Fdrops_vector.ma;h=66bfbcfd6979924f7085810a0873fe90e40018a9;hb=6f06bee3eeb4f718b273f0cca5873e9dc5e758b2;hp=9cf6e7c1fd76399eac10a7bc48d9b0a663e5e21b;hpb=11792b819aba3f464e63cb8f7834cac1652e372a;p=helm.git 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 9cf6e7c1f..66bfbcfd6 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,17 @@ 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,c. ∀L,K,f. ⬇*[c, f] L ≡ K → + λR,b. ∀f,L,K. ⬇*[b, f] L ≡ K → ∀Ts,Us. ⬆*[f] Ts ≡ Us → all … (R K) Ts → 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,c. d_liftable1 R c → d_liftable1_all R c. -#R #c #HR #L #K #f #HLK #Ts #Us #H elim H -Ts -Us normalize // +lemma d1_liftable_liftable_all: ∀R,b. d_liftable1 R b → d_liftable1_all R b. +#R #b #HR #f #L #K #HLK #Ts #Us #H elim H -Ts -Us normalize // #Ts #Us #T #U #HTU #_ #IHTUs * /3 width=7 by conj/ qed.