X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fdynamic%2Flsubsv.ma;h=537a97d1e309fc967a9dbb194de36c19d02f491c;hb=65008df95049eb835941ffea1aa682c9253c4c2b;hp=d203f0da4a3990a38ec9243932984e27f628354e;hpb=c07e9b0a3e65c28ca4154fec76a54a9a118fa7e1;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv.ma b/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv.ma index d203f0da4..537a97d1e 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv.ma @@ -12,7 +12,7 @@ (* *) (**************************************************************************) -include "basic_2/equivalence/lsubss.ma". +include "basic_2/notation/relations/crsubeqv_4.ma". include "basic_2/dynamic/snv.ma". (* LOCAL ENVIRONMENT REFINEMENT FOR STRATIFIED NATIVE VALIDITY **************) @@ -21,10 +21,10 @@ include "basic_2/dynamic/snv.ma". inductive lsubsv (h:sh) (g:sd h): relation lenv ≝ | lsubsv_atom: lsubsv h g (⋆) (⋆) | lsubsv_pair: ∀I,L1,L2,V. lsubsv h g L1 L2 → - lsubsv h g (L1. ⓑ{I} V) (L2. ⓑ{I} V) -| lsubsv_abbr: ∀L1,L2,V1,V2,W1,W2,l. ⦃h, L1⦄ ⊢ V1 ¡[g] → ⦃h, L1⦄ ⊢ V1 •[g] ⦃l+1, W1⦄ → - L1 ⊢ W1 ⬌* W2 → ⦃h, L2⦄ ⊢ W2 ¡[g] → ⦃h, L2⦄ ⊢ W2 •[g] ⦃l, V2⦄ → - lsubsv h g L1 L2 → lsubsv h g (L1. ⓓV1) (L2. ⓛW2) + lsubsv h g (L1.ⓑ{I}V) (L2.ⓑ{I}V) +| lsubsv_abbr: ∀L1,L2,W,V,W1,V2,l. ⦃h, L1⦄ ⊢ ⓝW.V ¡[g] → ⦃h, L2⦄ ⊢ W ¡[g] → + ⦃h, L1⦄ ⊢ V •[g] ⦃l+1, W1⦄ → ⦃h, L2⦄ ⊢ W •[g] ⦃l, V2⦄ → + lsubsv h g L1 L2 → lsubsv h g (L1.ⓓⓝW.V) (L2.ⓛW) . interpretation @@ -37,7 +37,7 @@ fact lsubsv_inv_atom1_aux: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[g] L2 → L1 = ⋆ → L #h #g #L1 #L2 * -L1 -L2 [ // | #I #L1 #L2 #V #_ #H destruct -| #L1 #L2 #V1 #V2 #W1 #W2 #l #_ #_ #_ #_ #_ #_ #H destruct +| #L1 #L2 #W #V #V1 #V2 #l #_ #_ #_ #_ #_ #H destruct ] qed-. @@ -45,30 +45,32 @@ lemma lsubsv_inv_atom1: ∀h,g,L2. h ⊢ ⋆ ¡⊑[g] L2 → L2 = ⋆. /2 width=5 by lsubsv_inv_atom1_aux/ qed-. fact lsubsv_inv_pair1_aux: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[g] L2 → - ∀I,K1,V1. L1 = K1. ⓑ{I} V1 → - (∃∃K2. h ⊢ K1 ¡⊑[g] K2 & L2 = K2. ⓑ{I} V1) ∨ - ∃∃K2,V2,W1,W2,l. ⦃h, K1⦄ ⊢ V1 ¡[g] & ⦃h, K1⦄ ⊢ V1 •[g] ⦃l+1, W1⦄ & - K1 ⊢ W1 ⬌* W2 & ⦃h, K2⦄ ⊢ W2 ¡[g] & ⦃h, K2⦄ ⊢ W2 •[g] ⦃l, V2⦄ & - h ⊢ K1 ¡⊑[g] K2 & L2 = K2. ⓛW2 & I = Abbr. + ∀I,K1,X. L1 = K1.ⓑ{I}X → + (∃∃K2. h ⊢ K1 ¡⊑[g] K2 & L2 = K2.ⓑ{I}X) ∨ + ∃∃K2,W,V,W1,V2,l. ⦃h, K1⦄ ⊢ X ¡[g] & ⦃h, K2⦄ ⊢ W ¡[g] & + ⦃h, K1⦄ ⊢ V •[g] ⦃l+1, W1⦄ & ⦃h, K2⦄ ⊢ W •[g] ⦃l, V2⦄ & + h ⊢ K1 ¡⊑[g] K2 & + I = Abbr & L2 = K2.ⓛW & X = ⓝW.V. #h #g #L1 #L2 * -L1 -L2 -[ #J #K1 #U1 #H destruct -| #I #L1 #L2 #V #HL12 #J #K1 #U1 #H destruct /3 width=3/ -| #L1 #L2 #V1 #V2 #W1 #W2 #l #HV1 #HVW1 #HW12 #HW2 #HWV2 #HL12 #J #K1 #U1 #H destruct /3 width=11/ +[ #J #K1 #X #H destruct +| #I #L1 #L2 #V #HL12 #J #K1 #X #H destruct /3 width=3/ +| #L1 #L2 #W #V #W1 #V2 #l #HV #HW #HW1 #HV2 #HL12 #J #K1 #X #H destruct /3 width=12/ ] qed-. -lemma lsubsv_inv_pair1: ∀h,g,I,K1,L2,V1. h ⊢ K1. ⓑ{I} V1 ¡⊑[g] L2 → - (∃∃K2. h ⊢ K1 ¡⊑[g] K2 & L2 = K2. ⓑ{I} V1) ∨ - ∃∃K2,V2,W1,W2,l. ⦃h, K1⦄ ⊢ V1 ¡[g] & ⦃h, K1⦄ ⊢ V1 •[g] ⦃l+1, W1⦄ & - K1 ⊢ W1 ⬌* W2 & ⦃h, K2⦄ ⊢ W2 ¡[g] & ⦃h, K2⦄ ⊢ W2 •[g] ⦃l, V2⦄ & - h ⊢ K1 ¡⊑[g] K2 & L2 = K2. ⓛW2 & I = Abbr. +lemma lsubsv_inv_pair1: ∀h,g,I,K1,L2,X. h ⊢ K1.ⓑ{I}X ¡⊑[g] L2 → + (∃∃K2. h ⊢ K1 ¡⊑[g] K2 & L2 = K2.ⓑ{I}X) ∨ + ∃∃K2,W,V,W1,V2,l. ⦃h, K1⦄ ⊢ X ¡[g] & ⦃h, K2⦄ ⊢ W ¡[g] & + ⦃h, K1⦄ ⊢ V •[g] ⦃l+1, W1⦄ & ⦃h, K2⦄ ⊢ W •[g] ⦃l, V2⦄ & + h ⊢ K1 ¡⊑[g] K2 & + I = Abbr & L2 = K2.ⓛW & X = ⓝW.V. /2 width=3 by lsubsv_inv_pair1_aux/ qed-. fact lsubsv_inv_atom2_aux: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[g] L2 → L2 = ⋆ → L1 = ⋆. #h #g #L1 #L2 * -L1 -L2 [ // | #I #L1 #L2 #V #_ #H destruct -| #L1 #L2 #V1 #V2 #W1 #W2 #l #_ #_ #_ #_ #_ #_ #H destruct +| #L1 #L2 #W #V #V1 #V2 #l #_ #_ #_ #_ #_ #H destruct ] qed-. @@ -76,37 +78,29 @@ lemma lsubsv_inv_atom2: ∀h,g,L1. h ⊢ L1 ¡⊑[g] ⋆ → L1 = ⋆. /2 width=5 by lsubsv_inv_atom2_aux/ qed-. fact lsubsv_inv_pair2_aux: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[g] L2 → - ∀I,K2,W2. L2 = K2. ⓑ{I} W2 → - (∃∃K1. h ⊢ K1 ¡⊑[g] K2 & L1 = K1. ⓑ{I} W2) ∨ - ∃∃K1,W1,V1,V2,l. ⦃h, K1⦄ ⊢ V1 ¡[g] & ⦃h, K1⦄ ⊢ V1 •[g] ⦃l+1, W1⦄ & - K1 ⊢ W1 ⬌* W2 & ⦃h, K2⦄ ⊢ W2 ¡[g] & ⦃h, K2⦄ ⊢ W2 •[g] ⦃l, V2⦄ & - h ⊢ K1 ¡⊑[g] K2 & L1 = K1. ⓓV1 & I = Abst. + ∀I,K2,W. L2 = K2.ⓑ{I}W → + (∃∃K1. h ⊢ K1 ¡⊑[g] K2 & L1 = K1.ⓑ{I}W) ∨ + ∃∃K1,V,W1,V2,l. ⦃h, K1⦄ ⊢ ⓝW.V ¡[g] & ⦃h, K2⦄ ⊢ W ¡[g] & + ⦃h, K1⦄ ⊢ V •[g] ⦃l+1, W1⦄ & ⦃h, K2⦄ ⊢ W •[g] ⦃l, V2⦄ & + h ⊢ K1 ¡⊑[g] K2 & I = Abst & L1 = K1. ⓓⓝW.V. #h #g #L1 #L2 * -L1 -L2 -[ #J #K2 #U2 #H destruct -| #I #L1 #L2 #V #HL12 #J #K2 #U2 #H destruct /3 width=3/ -| #L1 #L2 #V1 #V2 #W1 #W2 #l #HV #HVW1 #HW12 #HW2 #HWV2 #HL12 #J #K2 #U2 #H destruct /3 width=11/ +[ #J #K2 #U #H destruct +| #I #L1 #L2 #V #HL12 #J #K2 #U #H destruct /3 width=3/ +| #L1 #L2 #W #V #W1 #V2 #l #HV #HW #HW1 #HV2 #HL12 #J #K2 #U #H destruct /3 width=10/ ] qed-. -lemma lsubsv_inv_pair2: ∀h,g,I,L1,K2,W2. h ⊢ L1 ¡⊑[g] K2. ⓑ{I} W2 → - (∃∃K1. h ⊢ K1 ¡⊑[g] K2 & L1 = K1. ⓑ{I} W2) ∨ - ∃∃K1,W1,V1,V2,l. ⦃h, K1⦄ ⊢ V1 ¡[g] & ⦃h, K1⦄ ⊢ V1 •[g] ⦃l+1, W1⦄ & - K1 ⊢ W1 ⬌* W2 & ⦃h, K2⦄ ⊢ W2 ¡[g] & ⦃h, K2⦄ ⊢ W2 •[g] ⦃l, V2⦄ & - h ⊢ K1 ¡⊑[g] K2 & L1 = K1. ⓓV1 & I = Abst. +lemma lsubsv_inv_pair2: ∀h,g,I,L1,K2,W. h ⊢ L1 ¡⊑[g] K2.ⓑ{I}W → + (∃∃K1. h ⊢ K1 ¡⊑[g] K2 & L1 = K1.ⓑ{I}W) ∨ + ∃∃K1,V,W1,V2,l. ⦃h, K1⦄ ⊢ ⓝW.V ¡[g] & ⦃h, K2⦄ ⊢ W ¡[g] & + ⦃h, K1⦄ ⊢ V •[g] ⦃l+1, W1⦄ & ⦃h, K2⦄ ⊢ W •[g] ⦃l, V2⦄ & + h ⊢ K1 ¡⊑[g] K2 & I = Abst & L1 = K1. ⓓⓝW.V. /2 width=3 by lsubsv_inv_pair2_aux/ qed-. (* Basic_forward lemmas *****************************************************) -lemma lsubsv_fwd_lsubss: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[g] L2 → h ⊢ L1 •⊑[g] L2. -#h #g #L1 #L2 #H elim H -L1 -L2 // /2 width=1/ /2 width=6/ -qed-. - -lemma lsubsv_fwd_lsubr1: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[g] L2 → L1 ⊑ L2. -/3 width=3 by lsubsv_fwd_lsubss, lsubss_fwd_lsubr1/ -qed-. - -lemma lsubsv_fwd_lsubr2: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[g] L2 → L1 ⊑ L2. -/3 width=3 by lsubsv_fwd_lsubss, lsubss_fwd_lsubr2/ +lemma lsubsv_fwd_lsubx: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[g] L2 → L1 ⓝ⊑ L2. +#h #g #L1 #L2 #H elim H -L1 -L2 // /2 width=1/ qed-. (* Basic properties *********************************************************) @@ -117,5 +111,5 @@ qed. lemma lsubsv_cprs_trans: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[g] L2 → ∀T1,T2. L2 ⊢ T1 ➡* T2 → L1 ⊢ T1 ➡* T2. -/3 width=5 by lsubsv_fwd_lsubss, lsubss_cprs_trans/ +/3 width=5 by lsubsv_fwd_lsubx, lsubx_cprs_trans/ qed-.