X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2A%2Fmultiple%2Fdrops.ma;h=0320d863924d5b4beee2e22c817e473b3c96f979;hb=2f6f2b7c01d47d23f61dd48d767bcb37aecdcfea;hp=a86afce84bf3776e8a25a182d69e27eb0b649168;hpb=d2545ffd201b1aa49887313791386add78fa8603;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2A/multiple/drops.ma b/matita/matita/contribs/lambdadelta/basic_2A/multiple/drops.ma index a86afce84..0320d8639 100644 --- a/matita/matita/contribs/lambdadelta/basic_2A/multiple/drops.ma +++ b/matita/matita/contribs/lambdadelta/basic_2A/multiple/drops.ma @@ -12,18 +12,19 @@ (* *) (**************************************************************************) +include "ground/xoa/ex_4_3.ma". +include "ground/relocation/mr2_minus.ma". include "basic_2A/notation/relations/rdropstar_3.ma". include "basic_2A/notation/relations/rdropstar_4.ma". include "basic_2A/substitution/drop.ma". -include "basic_2A/multiple/mr2_minus.ma". include "basic_2A/multiple/lifts_vector.ma". (* ITERATED LOCAL ENVIRONMENT SLICING ***************************************) -inductive drops (s:bool): list2 nat nat → relation lenv ≝ +inductive drops (s:bool): mr2 → relation lenv ≝ | drops_nil : ∀L. drops s (◊) L L | drops_cons: ∀L1,L,L2,cs,l,m. - drops s cs L1 L → ⬇[s, l, m] L ≡ L2 → drops s ({l, m} @ cs) L1 L2 + drops s cs L1 L → ⬇[s, l, m] L ≡ L2 → drops s (❨l, m❩; cs) L1 L2 . interpretation "iterated slicing (local environment) abstract" @@ -53,12 +54,11 @@ fact drops_inv_nil_aux: ∀L1,L2,s,cs. ⬇*[s, cs] L1 ≡ L2 → cs = ◊ → L1 #L1 #L #L2 #l #m #cs #_ #_ #H destruct qed-. -(* Basic_1: was: drop1_gen_pnil *) lemma drops_inv_nil: ∀L1,L2,s. ⬇*[s, ◊] L1 ≡ L2 → L1 = L2. /2 width=4 by drops_inv_nil_aux/ qed-. fact drops_inv_cons_aux: ∀L1,L2,s,cs. ⬇*[s, cs] L1 ≡ L2 → - ∀l,m,tl. cs = {l, m} @ tl → + ∀l,m,tl. cs = ❨l, m❩; tl → ∃∃L. ⬇*[s, tl] L1 ≡ L & ⬇[s, l, m] L ≡ L2. #L1 #L2 #s #cs * -L1 -L2 -cs [ #L #l #m #tl #H destruct @@ -67,14 +67,13 @@ fact drops_inv_cons_aux: ∀L1,L2,s,cs. ⬇*[s, cs] L1 ≡ L2 → ] qed-. -(* Basic_1: was: drop1_gen_pcons *) -lemma drops_inv_cons: ∀L1,L2,s,l,m,cs. ⬇*[s, {l, m} @ cs] L1 ≡ L2 → +lemma drops_inv_cons: ∀L1,L2,s,l,m,cs. ⬇*[s, ❨l, m❩; cs] L1 ≡ L2 → ∃∃L. ⬇*[s, cs] L1 ≡ L & ⬇[s, l, m] L ≡ L2. /2 width=3 by drops_inv_cons_aux/ qed-. -lemma drops_inv_skip2: ∀I,s,cs,cs2,i. cs ▭ i ≡ cs2 → +lemma drops_inv_skip2: ∀I,s,cs,cs2,i. cs ▭ i ≘ cs2 → ∀L1,K2,V2. ⬇*[s, cs2] L1 ≡ K2. ⓑ{I} V2 → - ∃∃K1,V1,cs1. cs + 1 ▭ i + 1 ≡ cs1 + 1 & + ∃∃K1,V1,cs1. cs + 1 ▭ i + 1 ≘ cs1 + 1 & ⬇*[s, cs1] K1 ≡ K2 & ⬆*[cs1] V2 ≡ V1 & L1 = K1. ⓑ{I} V1. @@ -95,7 +94,6 @@ qed-. (* Basic properties *********************************************************) -(* Basic_1: was: drop1_skip_bind *) lemma drops_skip: ∀L1,L2,s,cs. ⬇*[s, cs] L1 ≡ L2 → ∀V1,V2. ⬆*[cs] V2 ≡ V1 → ∀I. ⬇*[s, cs + 1] L1.ⓑ{I}V1 ≡ L2.ⓑ{I}V2. #L1 #L2 #s #cs #H elim H -L1 -L2 -cs @@ -118,5 +116,3 @@ lemma d1_liftables_liftables_all: ∀R,s. d_liftables1 R s → d_liftables1_all #R #s #HR #L #K #cs #HLK #Ts #Us #H elim H -Ts -Us normalize // #Ts #Us #T #U #HTU #_ #IHTUs * /3 width=7 by conj/ qed. - -(* Basic_1: removed theorems 1: drop1_getl_trans *)