From: Ferruccio Guidi Date: Thu, 20 Jun 2013 17:23:20 +0000 (+0000) Subject: - bug fix il ldrop (interesting), fsup, fsups X-Git-Tag: make_still_working~1132 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=bdfd9f6ada4c66f67c674abc3c7b5ed64d27add3;p=helm.git - bug fix il ldrop (interesting), fsup, fsups - now fsupq is the reflexive closure of fsup, as expected - some renaming --- diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_lift.ma index 47d66618b..52aff988f 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_lift.ma @@ -63,25 +63,13 @@ include "basic_2/substitution/fsups.ma". lemma fsupq_cpxs_trans: ∀h,g,L1,L2,T2,U2. ⦃h, L2⦄ ⊢ T2 ➡*[g] U2 → ∀T1. ⦃L1, T1⦄ ⊃⸮ ⦃L2, T2⦄ → ∃∃U1. ⦃h, L1⦄ ⊢ T1 ➡*[g] U1 & ⦃L1, U1⦄ ⊃* ⦃L2, U2⦄. -#h #g #L1 #L2 #T2 #U2 #H @(cpxs_ind_dx … H) -T2 [ (* /3 width=3/ *) | +#h #g #L1 #L2 #T2 #U2 #H @(cpxs_ind_dx … H) -T2 [ /3 width=3/ ] #T #T2 #HT2 #_ #IHTU2 #T1 #HT1 elim (fsupq_cpx_trans … HT1 … HT2) -T #T #HT1 #HT2 elim (IHTU2 … HT2) -T2 /3 width=3/ - - -(* - elim H -L1 -L2 -T1 -T2 [2,3,4,5: /3 width=5/ ] -[ #L1 #K1 #K2 #T1 #T2 #U1 #d #e #HLK1 #HTU1 #_ #IHT12 #U2 #HTU2 - elim (IHT12 … HTU2) -IHT12 -HTU2 #T #HT1 #HT2 - elim (lift_total T d e) #U #HTU - lapply (cpx_lift … HT1 … HLK1 … HTU1 … HTU) -HT1 -HTU1 /3 width=11/ -| #I #L1 #V2 #U2 #HVU2 - elim (lift_total U2 0 1) /4 width=9/ -] qed-. - + lemma fsup_ssta_trans: ∀h,g,L1,L2,T1,T2. ⦃L1, T1⦄ ⊃ ⦃L2, T2⦄ → ∀U2,l. ⦃h, L2⦄ ⊢ T2 •[g] ⦃l+1, U2⦄ → ∃∃U1. ⦃h, L1⦄ ⊢ T1 ➡[g] U1 & ⦃L1, U1⦄ ⊃⸮ ⦃L2, U2⦄. /3 width=4 by fsup_cpx_trans, ssta_cpx/ qed-. -*) \ No newline at end of file diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/lsubc_ldrop.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/lsubc_ldrop.ma index a7c7c7a99..e2e4eea04 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/lsubc_ldrop.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/lsubc_ldrop.ma @@ -24,16 +24,15 @@ include "basic_2/computation/lsubc.ma". lemma lsubc_ldrop_O1_trans: ∀RP,L1,L2. L1 ⊑[RP] L2 → ∀K2,e. ⇩[0, e] L2 ≡ K2 → ∃∃K1. ⇩[0, e] L1 ≡ K1 & K1 ⊑[RP] K2. #RP #L1 #L2 #H elim H -L1 -L2 -[ #X #e #H - >(ldrop_inv_atom1 … H) -H /2 width=3/ +[ #X #e #H elim (ldrop_inv_atom1 … H) -H /2 width=3/ | #I #L1 #L2 #V #_ #IHL12 #X #e #H - elim (ldrop_inv_O1 … H) -H * #He #H destruct - [ elim (IHL12 L2 0 ?) -IHL12 // #X #H <(ldrop_inv_refl … H) -H /3 width=3/ + elim (ldrop_inv_O1_pair1 … H) -H * #He #H destruct + [ elim (IHL12 L2 0) -IHL12 // #X #H <(ldrop_inv_O2 … H) -H /3 width=3/ | elim (IHL12 … H) -L2 /3 width=3/ ] | #L1 #L2 #V #W #A #HV #HW #_ #IHL12 #X #e #H - elim (ldrop_inv_O1 … H) -H * #He #H destruct - [ elim (IHL12 L2 0 ?) -IHL12 // #X #H <(ldrop_inv_refl … H) -H /3 width=7/ + elim (ldrop_inv_O1_pair1 … H) -H * #He #H destruct + [ elim (IHL12 L2 0) -IHL12 // #X #H <(ldrop_inv_O2 … H) -H /3 width=7/ | elim (IHL12 … H) -L2 /3 width=3/ ] qed-. @@ -44,8 +43,7 @@ lemma ldrop_lsubc_trans: ∀RR,RS,RP. ∀L1,K1,d,e. ⇩[d, e] L1 ≡ K1 → ∀K2. K1 ⊑[RP] K2 → ∃∃L2. L1 ⊑[RP] L2 & ⇩[d, e] L2 ≡ K2. #RR #RS #RP #Hacp #Hacr #L1 #K1 #d #e #H elim H -L1 -K1 -d -e -[ #d #e #X #H - >(lsubc_inv_atom1 … H) -H /2 width=3/ +[ #d #X #H elim (lsubc_inv_atom1 … H) -H /2 width=3/ | #L1 #I #V1 #X #H elim (lsubc_inv_pair1 … H) -H * [ #K1 #HLK1 #H destruct /3 width=3/ diff --git a/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_ldrop.ma b/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_ldrop.ma index 686095062..bd46208b8 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_ldrop.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_ldrop.ma @@ -25,17 +25,17 @@ lemma lsubsv_ldrop_O1_conf: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[g] L2 → #h #g #L1 #L2 #H elim H -L1 -L2 [ /2 width=3/ | #I #L1 #L2 #V #_ #IHL12 #K1 #e #H - elim (ldrop_inv_O1 … H) -H * #He #HLK1 + elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK1 [ destruct elim (IHL12 L1 0 ?) -IHL12 // #X #HL12 #H - <(ldrop_inv_refl … H) in HL12; -H /3 width=3/ + <(ldrop_inv_O2 … H) in HL12; -H /3 width=3/ | elim (IHL12 … HLK1) -L1 /3 width=3/ ] | #L1 #L2 #V1 #V2 #W1 #W2 #l #HV1 #HVW1 #HW12 #HW2 #HWV2 #_ #IHL12 #K1 #e #H - elim (ldrop_inv_O1 … H) -H * #He #HLK1 + elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK1 [ destruct elim (IHL12 L1 0 ?) -IHL12 // #X #HL12 #H - <(ldrop_inv_refl … H) in HL12; -H /3 width=6/ + <(ldrop_inv_O2 … H) in HL12; -H /3 width=6/ | elim (IHL12 … HLK1) -L1 /3 width=3/ ] ] @@ -48,17 +48,17 @@ lemma lsubsv_ldrop_O1_trans: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[g] L2 → #h #g #L1 #L2 #H elim H -L1 -L2 [ /2 width=3/ | #I #L1 #L2 #V #_ #IHL12 #K2 #e #H - elim (ldrop_inv_O1 … H) -H * #He #HLK2 + elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK2 [ destruct elim (IHL12 L2 0 ?) -IHL12 // #X #HL12 #H - <(ldrop_inv_refl … H) in HL12; -H /3 width=3/ + <(ldrop_inv_O2 … H) in HL12; -H /3 width=3/ | elim (IHL12 … HLK2) -L2 /3 width=3/ ] | #L1 #L2 #V1 #V2 #W1 #W2 #l #HV1 #HVW1 #HW12 #HW2 #HWV2 #_ #IHL12 #K2 #e #H - elim (ldrop_inv_O1 … H) -H * #He #HLK2 + elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK2 [ destruct elim (IHL12 L2 0 ?) -IHL12 // #X #HL12 #H - <(ldrop_inv_refl … H) in HL12; -H /3 width=6/ + <(ldrop_inv_O2 … H) in HL12; -H /3 width=6/ | elim (IHL12 … HLK2) -L2 /3 width=3/ ] ] diff --git a/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_lift.ma index 0a1feae91..8b68e242c 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_lift.ma @@ -85,9 +85,9 @@ lemma snv_fsup_conf: ∀h,g,L1,L2,T1,T2. ⦃L1, T1⦄ ⊃ ⦃L2, T2⦄ → #h #g #L1 #L2 #T1 #T2 #H elim H -L1 -L2 -T1 -T2 [ #I1 #L1 #V1 #H elim (snv_inv_lref … H) -H #I2 #L2 #V2 #H #HV2 - lapply (ldrop_inv_refl … H) -H #H destruct // + lapply (ldrop_inv_O2 … H) -H #H destruct // |2: * -|5: /3 width=7 by snv_inv_lift/ +|5,6: /3 width=7 by snv_inv_lift/ ] [1,3: #a #I #L1 #V1 #T1 #H elim (snv_inv_bind … H) -H // |2,4: * #L1 #V1 #T1 #H diff --git a/matita/matita/contribs/lambdadelta/basic_2/equivalence/lsubss_ldrop.ma b/matita/matita/contribs/lambdadelta/basic_2/equivalence/lsubss_ldrop.ma index 7c5e53c1d..1f19a066e 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/equivalence/lsubss_ldrop.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/equivalence/lsubss_ldrop.ma @@ -25,17 +25,17 @@ lemma lsubss_ldrop_O1_conf: ∀h,g,L1,L2. h ⊢ L1 •⊑[g] L2 → #h #g #L1 #L2 #H elim H -L1 -L2 [ /2 width=3/ | #I #L1 #L2 #V #_ #IHL12 #K1 #e #H - elim (ldrop_inv_O1 … H) -H * #He #HLK1 + elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK1 [ destruct elim (IHL12 L1 0 ?) -IHL12 // #X #HL12 #H - <(ldrop_inv_refl … H) in HL12; -H /3 width=3/ + <(ldrop_inv_O2 … H) in HL12; -H /3 width=3/ | elim (IHL12 … HLK1) -L1 /3 width=3/ ] | #L1 #L2 #V1 #V2 #W1 #W2 #l #HW12 #HVW1 #HWV2 #_ #IHL12 #K1 #e #H - elim (ldrop_inv_O1 … H) -H * #He #HLK1 + elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK1 [ destruct elim (IHL12 L1 0 ?) -IHL12 // #X #HL12 #H - <(ldrop_inv_refl … H) in HL12; -H /3 width=6/ + <(ldrop_inv_O2 … H) in HL12; -H /3 width=6/ | elim (IHL12 … HLK1) -L1 /3 width=3/ ] ] @@ -48,17 +48,17 @@ lemma lsubss_ldrop_O1_trans: ∀h,g,L1,L2. h ⊢ L1 •⊑[g] L2 → #h #g #L1 #L2 #H elim H -L1 -L2 [ /2 width=3/ | #I #L1 #L2 #V #_ #IHL12 #K2 #e #H - elim (ldrop_inv_O1 … H) -H * #He #HLK2 + elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK2 [ destruct elim (IHL12 L2 0 ?) -IHL12 // #X #HL12 #H - <(ldrop_inv_refl … H) in HL12; -H /3 width=3/ + <(ldrop_inv_O2 … H) in HL12; -H /3 width=3/ | elim (IHL12 … HLK2) -L2 /3 width=3/ ] | #L1 #L2 #V1 #V2 #W1 #W2 #l #HW12 #HVW1 #HWV2 #_ #IHL12 #K2 #e #H - elim (ldrop_inv_O1 … H) -H * #He #HLK2 + elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK2 [ destruct elim (IHL12 L2 0 ?) -IHL12 // #X #HL12 #H - <(ldrop_inv_refl … H) in HL12; -H /3 width=6/ + <(ldrop_inv_O2 … H) in HL12; -H /3 width=6/ | elim (IHL12 … HLK2) -L2 /3 width=3/ ] ] diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr.ma index 10430b3a1..d7c59c80d 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr.ma @@ -87,7 +87,7 @@ qed-. lemma cpr_append: l_appendable_sn … cpr. #K #T1 #T2 #H elim H -K -T1 -T2 // /2 width=1/ /2 width=3/ #K #K0 #V1 #V2 #W2 #i #HK0 #_ #HVW2 #IHV12 #L -lapply (ldrop_fwd_ldrop2_length … HK0) #H +lapply (ldrop_fwd_length_lt2 … HK0) #H @(cpr_delta … (L@@K0) V1 … HVW2) // @(ldrop_O1_append_sn_le … HK0) /2 width=2/ (**) (* /3/ does not work *) qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpx.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpx.ma index 33063d819..7a08bc435 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpx.ma @@ -76,7 +76,7 @@ qed-. lemma cpx_append: ∀h,g. l_appendable_sn … (cpx h g). #h #g #K #T1 #T2 #H elim H -K -T1 -T2 // /2 width=1/ /2 width=3/ #I #K #K0 #V1 #V2 #W2 #i #HK0 #_ #HVW2 #IHV12 #L -lapply (ldrop_fwd_ldrop2_length … HK0) #H +lapply (ldrop_fwd_length_lt2 … HK0) #H @(cpx_delta … I … (L@@K0) V1 … HVW2) // @(ldrop_O1_append_sn_le … HK0) /2 width=2/ (**) (* /3/ does not work *) qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/crf.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/crf.ma index 59549b04d..af60a2b0d 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/crf.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/crf.ma @@ -45,7 +45,7 @@ interpretation fact trf_inv_atom_aux: ∀I,L,T. L ⊢ 𝐑⦃T⦄ → L = ⋆ → T = ⓪{I} → ⊥. #I #L #T * -L -T [ #L #K #V #i #HLK #H1 #H2 destruct - lapply (ldrop_inv_atom1 … HLK) -HLK #H destruct + elim (ldrop_inv_atom1 … HLK) -HLK #H destruct | #L #V #T #_ #_ #H destruct | #L #V #T #_ #_ #H destruct | #J #L #V #T #_ #_ #H destruct diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/crf_append.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/crf_append.ma index 7d3ac6050..0a86278d0 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/crf_append.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/crf_append.ma @@ -22,7 +22,7 @@ include "basic_2/reduction/crf.ma". lemma crf_labst_last: ∀L,T,W. L ⊢ 𝐑⦃T⦄ → ⋆.ⓛW @@ L ⊢ 𝐑⦃T⦄. #L #T #W #H elim H -L -T /2 width=1/ #L #K #V #i #HLK -lapply (ldrop_fwd_ldrop2_length … HLK) #Hi +lapply (ldrop_fwd_length_lt2 … HLK) #Hi lapply (ldrop_O1_append_sn_le … HLK … (⋆.ⓛW)) -HLK /2 width=2/ -Hi /2 width=3/ qed. @@ -36,14 +36,14 @@ fact crf_inv_labst_last_aux: ∀L1,T,W. L1 ⊢ 𝐑⦃T⦄ → ∀L2. L1 = ⋆.ⓛW @@ L2 → L2 ⊢ 𝐑⦃T⦄. #L1 #T #W #H elim H -L1 -T /2 width=1/ /3 width=1/ [ #L1 #K1 #V1 #i #HLK1 #L2 #H destruct - lapply (ldrop_fwd_ldrop2_length … HLK1) + lapply (ldrop_fwd_length_lt2 … HLK1) >append_length >commutative_plus normalize in ⊢ (??% → ?); #H elim (le_to_or_lt_eq i (|L2|) ?) /2 width=1/ -H #Hi destruct [ elim (ldrop_O1_lt … Hi) #I2 #K2 #V2 #HLK2 lapply (ldrop_O1_inv_append1_le … HLK1 … HLK2) -HLK1 /2 width=2/ -Hi normalize #H destruct /2 width=3/ | lapply (ldrop_O1_inv_append1_ge … HLK1 ?) -HLK1 // (ldrop_inv_O2 … HLK) -L1 >(lift_inv_O2 … HTU) -T2 -d // +qed. + +(* Main properties **********************************************************) + +theorem fsupq_fsupqa: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⊃⸮ ⦃L2, T2⦄ → ⦃L1, T1⦄ ⊃⊃⸮ ⦃L2, T2⦄. +#L1 #L2 #T1 #T2 #H elim H -L1 -L2 -T1 -T2 // /2 width=1/ /2 width=7/ +qed. + +(* Main inversion properties ************************************************) + +theorem fsupqa_inv_fsupq: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⊃⊃⸮ ⦃L2, T2⦄ → ⦃L1, T1⦄ ⊃⸮ ⦃L2, T2⦄. +#L1 #L2 #T1 #T2 #H elim H -H /2 width=1/ +* #H1 #H2 destruct // +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop.ma index f51640fa1..cf7c7ab88 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop.ma @@ -20,7 +20,7 @@ include "basic_2/relocation/lift.ma". (* Basic_1: includes: drop_skip_bind *) inductive ldrop: nat → nat → relation lenv ≝ -| ldrop_atom : ∀d,e. ldrop d e (⋆) (⋆) +| ldrop_atom : ∀d. ldrop d 0 (⋆) (⋆) | ldrop_pair : ∀L,I,V. ldrop 0 0 (L. ⓑ{I} V) (L. ⓑ{I} V) | ldrop_ldrop: ∀L1,L2,I,V,e. ldrop 0 e L1 L2 → ldrop 0 (e + 1) (L1. ⓑ{I} V) L2 | ldrop_skip : ∀L1,L2,I,V1,V2,d,e. @@ -53,53 +53,40 @@ definition dropable_dx: predicate (relation lenv) ≝ (* Basic inversion lemmas ***************************************************) -fact ldrop_inv_refl_aux: ∀d,e,L1,L2. ⇩[d, e] L1 ≡ L2 → d = 0 → e = 0 → L1 = L2. -#d #e #L1 #L2 * -d -e -L1 -L2 -[ // -| // -| #L1 #L2 #I #V #e #_ #_ >commutative_plus normalize #H destruct -| #L1 #L2 #I #V1 #V2 #d #e #_ #_ >commutative_plus normalize #H destruct -] -qed. - -(* Basic_1: was: drop_gen_refl *) -lemma ldrop_inv_refl: ∀L1,L2. ⇩[0, 0] L1 ≡ L2 → L1 = L2. -/2 width=5/ qed-. - fact ldrop_inv_atom1_aux: ∀d,e,L1,L2. ⇩[d, e] L1 ≡ L2 → L1 = ⋆ → - L2 = ⋆. + L2 = ⋆ ∧ e = 0. #d #e #L1 #L2 * -d -e -L1 -L2 -[ // +[ /2 width=1/ | #L #I #V #H destruct | #L1 #L2 #I #V #e #_ #H destruct | #L1 #L2 #I #V1 #V2 #d #e #_ #_ #H destruct ] -qed. +qed-. (* Basic_1: was: drop_gen_sort *) -lemma ldrop_inv_atom1: ∀d,e,L2. ⇩[d, e] ⋆ ≡ L2 → L2 = ⋆. -/2 width=5/ qed-. +lemma ldrop_inv_atom1: ∀d,e,L2. ⇩[d, e] ⋆ ≡ L2 → L2 = ⋆ ∧ e = 0. +/2 width=4 by ldrop_inv_atom1_aux/ qed-. -fact ldrop_inv_O1_aux: ∀d,e,L1,L2. ⇩[d, e] L1 ≡ L2 → d = 0 → - ∀K,I,V. L1 = K. ⓑ{I} V → - (e = 0 ∧ L2 = K. ⓑ{I} V) ∨ - (0 < e ∧ ⇩[d, e - 1] K ≡ L2). +fact ldrop_inv_O1_pair1_aux: ∀d,e,L1,L2. ⇩[d, e] L1 ≡ L2 → d = 0 → + ∀K,I,V. L1 = K. ⓑ{I} V → + (e = 0 ∧ L2 = K. ⓑ{I} V) ∨ + (0 < e ∧ ⇩[d, e - 1] K ≡ L2). #d #e #L1 #L2 * -d -e -L1 -L2 -[ #d #e #_ #K #I #V #H destruct +[ #d #_ #K #I #V #H destruct | #L #I #V #_ #K #J #W #HX destruct /3 width=1/ | #L1 #L2 #I #V #e #HL12 #_ #K #J #W #H destruct /3 width=1/ | #L1 #L2 #I #V1 #V2 #d #e #_ #_ >commutative_plus normalize #H destruct ] -qed. +qed-. -lemma ldrop_inv_O1: ∀e,K,I,V,L2. ⇩[0, e] K. ⓑ{I} V ≡ L2 → - (e = 0 ∧ L2 = K. ⓑ{I} V) ∨ - (0 < e ∧ ⇩[0, e - 1] K ≡ L2). -/2 width=3/ qed-. +lemma ldrop_inv_O1_pair1: ∀e,K,I,V,L2. ⇩[0, e] K. ⓑ{I} V ≡ L2 → + (e = 0 ∧ L2 = K. ⓑ{I} V) ∨ + (0 < e ∧ ⇩[0, e - 1] K ≡ L2). +/2 width=3 by ldrop_inv_O1_pair1_aux/ qed-. lemma ldrop_inv_pair1: ∀K,I,V,L2. ⇩[0, 0] K. ⓑ{I} V ≡ L2 → L2 = K. ⓑ{I} V. #K #I #V #L2 #H -elim (ldrop_inv_O1 … H) -H * // #H destruct +elim (ldrop_inv_O1_pair1 … H) -H * // #H destruct elim (lt_refl_false … H) qed-. @@ -107,7 +94,7 @@ qed-. lemma ldrop_inv_ldrop1: ∀e,K,I,V,L2. ⇩[0, e] K. ⓑ{I} V ≡ L2 → 0 < e → ⇩[0, e - 1] K ≡ L2. #e #K #I #V #L2 #H #He -elim (ldrop_inv_O1 … H) -H * // #H destruct +elim (ldrop_inv_O1_pair1 … H) -H * // #H destruct elim (lt_refl_false … He) qed-. @@ -117,27 +104,27 @@ fact ldrop_inv_skip1_aux: ∀d,e,L1,L2. ⇩[d, e] L1 ≡ L2 → 0 < d → ⇧[d - 1, e] V2 ≡ V1 & L2 = K2. ⓑ{I} V2. #d #e #L1 #L2 * -d -e -L1 -L2 -[ #d #e #_ #I #K #V #H destruct +[ #d #_ #I #K #V #H destruct | #L #I #V #H elim (lt_refl_false … H) | #L1 #L2 #I #V #e #_ #H elim (lt_refl_false … H) | #X #L2 #Y #Z #V2 #d #e #HL12 #HV12 #_ #I #L1 #V1 #H destruct /2 width=5/ ] -qed. +qed-. (* Basic_1: was: drop_gen_skip_l *) lemma ldrop_inv_skip1: ∀d,e,I,K1,V1,L2. ⇩[d, e] K1. ⓑ{I} V1 ≡ L2 → 0 < d → ∃∃K2,V2. ⇩[d - 1, e] K1 ≡ K2 & ⇧[d - 1, e] V2 ≡ V1 & L2 = K2. ⓑ{I} V2. -/2 width=3/ qed-. +/2 width=3 by ldrop_inv_skip1_aux/ qed-. lemma ldrop_inv_O1_pair2: ∀I,K,V,e,L1. ⇩[0, e] L1 ≡ K. ⓑ{I} V → (e = 0 ∧ L1 = K. ⓑ{I} V) ∨ ∃∃I1,K1,V1. ⇩[0, e - 1] K1 ≡ K. ⓑ{I} V & L1 = K1.ⓑ{I1}V1 & 0 < e. #I #K #V #e * -[ #H lapply (ldrop_inv_atom1 … H) -H #H destruct +[ #H elim (ldrop_inv_atom1 … H) -H #H destruct | #L1 #I1 #V1 #H - elim (ldrop_inv_O1 … H) -H * + elim (ldrop_inv_O1_pair1 … H) -H * [ #H1 #H2 destruct /3 width=1/ | /3 width=5/ ] @@ -150,24 +137,25 @@ fact ldrop_inv_skip2_aux: ∀d,e,L1,L2. ⇩[d, e] L1 ≡ L2 → 0 < d → ⇧[d - 1, e] V2 ≡ V1 & L1 = K1. ⓑ{I} V1. #d #e #L1 #L2 * -d -e -L1 -L2 -[ #d #e #_ #I #K #V #H destruct +[ #d #_ #I #K #V #H destruct | #L #I #V #H elim (lt_refl_false … H) | #L1 #L2 #I #V #e #_ #H elim (lt_refl_false … H) | #L1 #X #Y #V1 #Z #d #e #HL12 #HV12 #_ #I #L2 #V2 #H destruct /2 width=5/ ] -qed. +qed-. (* Basic_1: was: drop_gen_skip_r *) lemma ldrop_inv_skip2: ∀d,e,I,L1,K2,V2. ⇩[d, e] L1 ≡ K2. ⓑ{I} V2 → 0 < d → ∃∃K1,V1. ⇩[d - 1, e] K1 ≡ K2 & ⇧[d - 1, e] V2 ≡ V1 & L1 = K1. ⓑ{I} V1. -/2 width=3/ qed-. +/2 width=3 by ldrop_inv_skip2_aux/ qed-. (* Basic properties *********************************************************) (* Basic_1: was by definition: drop_refl *) -lemma ldrop_refl: ∀L. ⇩[0, 0] L ≡ L. +lemma ldrop_refl: ∀L,d. ⇩[d, 0] L ≡ L. #L elim L -L // +#L #I #V #IHL #d @(nat_ind_plus … d) -d // /2 width=1/ qed. lemma ldrop_ldrop_lt: ∀L1,L2,I,V,e. @@ -181,21 +169,21 @@ lemma ldrop_skip_lt: ∀L1,L2,I,V1,V2,d,e. #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV21 #Hd >(plus_minus_m_m d 1) // /2 width=1/ qed. -lemma ldrop_O1_le: ∀i,L. i ≤ |L| → ∃K. ⇩[0, i] L ≡ K. -#i @(nat_ind_plus … i) -i /2 width=2/ -#i #IHi * +lemma ldrop_O1_le: ∀e,L. e ≤ |L| → ∃K. ⇩[0, e] L ≡ K. +#e @(nat_ind_plus … e) -e /2 width=2/ +#e #IHe * [ #H lapply (le_n_O_to_eq … H) -H >commutative_plus normalize #H destruct | #L #I #V normalize #H - elim (IHi L ?) -IHi /2 width=1/ -H /3 width=2/ + elim (IHe L) -IHe /2 width=1/ -H /3 width=2/ ] qed. -lemma ldrop_O1_lt: ∀L,i. i < |L| → ∃∃I,K,V. ⇩[0, i] L ≡ K.ⓑ{I}V. +lemma ldrop_O1_lt: ∀L,e. e < |L| → ∃∃I,K,V. ⇩[0, e] L ≡ K.ⓑ{I}V. #L elim L -L -[ #i #H elim (lt_zero_false … H) -| #L #I #V #IHL #i @(nat_ind_plus … i) -i /2 width=4/ - #i #_ normalize #H - elim (IHL i ? ) -IHL /2 width=1/ -H /3 width=4/ +[ #e #H elim (lt_zero_false … H) +| #L #I #V #IHL #e @(nat_ind_plus … e) -e /2 width=4/ + #e #_ normalize #H + elim (IHL e) -IHL /2 width=1/ -H /3 width=4/ ] qed. @@ -247,78 +235,91 @@ qed. (* Basic forvard lemmas *****************************************************) -lemma ldrop_fwd_lw: ∀L1,L2,d,e. ⇩[d, e] L1 ≡ L2 → ♯{L2} ≤ ♯{L1}. -#L1 #L2 #d #e #H elim H -L1 -L2 -d -e // normalize -[ /2 width=3/ -| #L1 #L2 #I #V1 #V2 #d #e #_ #HV21 #IHL12 - >(lift_fwd_tw … HV21) -HV21 /2 width=1/ -] -qed-. - (* Basic_1: was: drop_S *) lemma ldrop_fwd_ldrop2: ∀L1,I2,K2,V2,e. ⇩[O, e] L1 ≡ K2. ⓑ{I2} V2 → ⇩[O, e + 1] L1 ≡ K2. #L1 elim L1 -L1 -[ #I2 #K2 #V2 #e #H lapply (ldrop_inv_atom1 … H) -H #H destruct +[ #I2 #K2 #V2 #e #H lapply (ldrop_inv_atom1 … H) -H * #H destruct | #K1 #I1 #V1 #IHL1 #I2 #K2 #V2 #e #H - elim (ldrop_inv_O1 … H) -H * #He #H + elim (ldrop_inv_O1_pair1 … H) -H * #He #H [ -IHL1 destruct /2 width=1/ | @ldrop_ldrop >(plus_minus_m_m e 1) // /2 width=3/ ] ] qed-. -lemma ldrop_fwd_length: ∀L1,L2,d,e. ⇩[d, e] L1 ≡ L2 → |L2| ≤ |L1|. +lemma ldrop_fwd_length: ∀L1,L2,d,e. ⇩[d, e] L1 ≡ L2 → |L1| = |L2| + e. #L1 #L2 #d #e #H elim H -L1 -L2 -d -e // normalize /2 width=1/ qed-. -lemma ldrop_fwd_ldrop2_length: ∀L1,I2,K2,V2,e. - ⇩[0, e] L1 ≡ K2. ⓑ{I2} V2 → e < |L1|. -#L1 elim L1 -L1 -[ #I2 #K2 #V2 #e #H lapply (ldrop_inv_atom1 … H) -H #H destruct -| #K1 #I1 #V1 #IHL1 #I2 #K2 #V2 #e #H - elim (ldrop_inv_O1 … H) -H * #He #H - [ -IHL1 destruct // - | lapply (IHL1 … H) -IHL1 -H #HeK1 whd in ⊢ (? ? %); /2 width=1/ - ] -] +lemma ldrop_fwd_length_minus2: ∀L1,L2,d,e. ⇩[d, e] L1 ≡ L2 → |L2| = |L1| - e. +#L1 #L2 #d #e #H lapply (ldrop_fwd_length … H) -H /2 width=1/ qed-. -lemma ldrop_fwd_O1_length: ∀L1,L2,e. ⇩[0, e] L1 ≡ L2 → |L2| = |L1| - e. -#L1 elim L1 -L1 -[ #L2 #e #H >(ldrop_inv_atom1 … H) -H // -| #K1 #I1 #V1 #IHL1 #L2 #e #H - elim (ldrop_inv_O1 … H) -H * #He #H - [ -IHL1 destruct // - | lapply (IHL1 … H) -IHL1 -H #H >H -H normalize - >minus_le_minus_minus_comm // - ] -] +lemma ldrop_fwd_length_minus4: ∀L1,L2,d,e. ⇩[d, e] L1 ≡ L2 → e = |L1| - |L2|. +#L1 #L2 #d #e #H lapply (ldrop_fwd_length … H) -H // qed-. -lemma ldrop_fwd_lw_eq: ∀L1,L2,d,e. ⇩[d, e] L1 ≡ L2 → - |L1| = |L2| → ♯{L2} = ♯{L1}. -#L1 #L2 #d #e #H elim H -L1 -L2 -d -e // -[ #L1 #L2 #I #V #e #HL12 #_ - lapply (ldrop_fwd_O1_length … HL12) -HL12 #HL21 >HL21 -HL21 normalize #H -I - lapply (discr_plus_xy_minus_xz … H) -e #H destruct -| #L1 #L2 #I #V1 #V2 #d #e #_ #HV21 #HL12 normalize in ⊢ (??%%→??%%); #H -I - >(lift_fwd_tw … HV21) -V2 /3 width=1 by eq_f2/ (**) (* auto is a bit slow without trace *) +lemma ldrop_fwd_length_le2: ∀L1,L2,d,e. ⇩[d, e] L1 ≡ L2 → e ≤ |L1|. +#L1 #L2 #d #e #H lapply (ldrop_fwd_length … H) -H // +qed-. + +lemma ldrop_fwd_length_le4: ∀L1,L2,d,e. ⇩[d, e] L1 ≡ L2 → |L2| ≤ |L1|. +#L1 #L2 #d #e #H lapply (ldrop_fwd_length … H) -H // +qed-. + +lemma ldrop_fwd_length_lt2: ∀L1,I2,K2,V2,d,e. + ⇩[d, e] L1 ≡ K2. ⓑ{I2} V2 → e < |L1|. +#L1 #I2 #K2 #V2 #d #e #H +lapply (ldrop_fwd_length … H) normalize in ⊢ (%→?); -I2 -V2 // +qed-. + +lemma ldrop_fwd_length_lt4: ∀L1,L2,d,e. ⇩[d, e] L1 ≡ L2 → 0 < e → |L2| < |L1|. +#L1 #L2 #d #e #H lapply (ldrop_fwd_length … H) -H /2 width=1/ +qed-. + +lemma ldrop_fwd_lw: ∀L1,L2,d,e. ⇩[d, e] L1 ≡ L2 → ♯{L2} ≤ ♯{L1}. +#L1 #L2 #d #e #H elim H -L1 -L2 -d -e // normalize +[ /2 width=3/ +| #L1 #L2 #I #V1 #V2 #d #e #_ #HV21 #IHL12 + >(lift_fwd_tw … HV21) -HV21 /2 width=1/ ] qed-. -lemma ldrop_fwd_lw_lt: ∀L1,L2,d,e. ⇩[d, e] L1 ≡ L2 → - |L2| < |L1| → ♯{L2} < ♯{L1}. +lemma ldrop_fwd_lw_lt: ∀L1,L2,d,e. ⇩[d, e] L1 ≡ L2 → 0 < e → ♯{L2} < ♯{L1}. #L1 #L2 #d #e #H elim H -L1 -L2 -d -e // [ #L #I #V #H elim (lt_refl_false … H) | #L1 #L2 #I #V #e #HL12 #_ #_ lapply (ldrop_fwd_lw … HL12) -HL12 #HL12 @(le_to_lt_to_lt … HL12) -HL12 // -| #L1 #L2 #I #V1 #V2 #d #e #_ #HV21 #IHL12 normalize in ⊢ (?%%→?%%); #H -I - >(lift_fwd_tw … HV21) -V2 /4 width=2 by lt_minus_to_plus, lt_plus_to_lt_l/ (**) (* auto too slow without trace *) +| #L1 #L2 #I #V1 #V2 #d #e #_ #HV21 #IHL12 #H normalize in ⊢ (?%%); -I + >(lift_fwd_tw … HV21) -V2 /3 by lt_minus_to_plus/ (**) (* auto too slow without trace *) ] qed-. +(* Advanced inversion lemmas ************************************************) + +fact ldrop_inv_O2_aux: ∀d,e,L1,L2. ⇩[d, e] L1 ≡ L2 → e = 0 → L1 = L2. +#d #e #L1 #L2 #H elim H -d -e -L1 -L2 +[ // +| // +| #L1 #L2 #I #V #e #_ #_ >commutative_plus normalize #H destruct +| #L1 #L2 #I #V1 #V2 #d #e #_ #HV21 #IHL12 #H + >(IHL12 H) -L1 >(lift_inv_O2_aux … HV21 … H) -V2 -d -e // +] +qed-. + +(* Basic_1: was: drop_gen_refl *) +lemma ldrop_inv_O2: ∀L1,L2,d. ⇩[d, 0] L1 ≡ L2 → L1 = L2. +/2 width=4 by ldrop_inv_O2_aux/ qed-. + +lemma ldrop_inv_length_eq: ∀L1,L2,d,e. ⇩[d, e] L1 ≡ L2 → |L1| = |L2| → e = 0. +#L1 #L2 #d #e #H #HL12 lapply (ldrop_fwd_length_minus4 … H) // +qed-. + +lemma ldrop_inv_refl: ∀L,d,e. ⇩[d, e] L ≡ L → e = 0. +/2 width=5 by ldrop_inv_length_eq/ qed-. + (* Basic_1: removed theorems 50: drop_ctail drop_skip_flat cimp_flat_sx cimp_flat_dx cimp_bind cimp_getl_conf diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop_append.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop_append.ma index 6cae606c7..9b49fdd5b 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop_append.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop_append.ma @@ -23,8 +23,6 @@ fact ldrop_O1_append_sn_le_aux: ∀L1,L2,d,e. ⇩[d, e] L1 ≡ L2 → d = 0 → e ≤ |L1| → ∀L. ⇩[0, e] L @@ L1 ≡ L @@ L2. #L1 #L2 #d #e #H elim H -L1 -L2 -d -e normalize // /4 width=1/ -#d #e #_ #H #L -d -lapply (le_n_O_to_eq … H) -H // qed-. lemma ldrop_O1_append_sn_le: ∀L1,L2,e. ⇩[0, e] L1 ≡ L2 → e ≤ |L1| → @@ -37,7 +35,7 @@ lemma ldrop_O1_inv_append1_ge: ∀K,L1,L2,e. ⇩[0, e] L1 @@ L2 ≡ K → |L2| ≤ e → ⇩[0, e - |L2|] L1 ≡ K. #K #L1 #L2 elim L2 -L2 normalize // #L2 #I #V #IHL2 #e #H #H1e -elim (ldrop_inv_O1 … H) -H * #H2e #HL12 destruct +elim (ldrop_inv_O1_pair1 … H) -H * #H2e #HL12 destruct [ lapply (le_n_O_to_eq … H1e) -H1e -IHL2 >commutative_plus normalize #H destruct | minus_minus_comm /3 width=1/ @@ -49,12 +47,12 @@ lemma ldrop_O1_inv_append1_le: ∀K,L1,L2,e. ⇩[0, e] L1 @@ L2 ≡ K → e ≤ #K #L1 #L2 elim L2 -L2 normalize [ #e #H1 #H2 #K2 #H3 lapply (le_n_O_to_eq … H2) -H2 #H2 - lapply (ldrop_inv_atom1 … H3) -H3 #H3 destruct - >(ldrop_inv_refl … H1) -H1 // + elim (ldrop_inv_atom1 … H3) -H3 #H3 #_ destruct + >(ldrop_inv_O2 … H1) -H1 // | #L2 #I #V #IHL2 #e @(nat_ind_plus … e) -e [ -IHL2 ] [ #H1 #_ #K2 #H2 - lapply (ldrop_inv_refl … H1) -H1 #H1 - lapply (ldrop_inv_refl … H2) -H2 #H2 destruct // + lapply (ldrop_inv_O2 … H1) -H1 #H1 + lapply (ldrop_inv_O2 … H2) -H2 #H2 destruct // | #e #_ #H1 #H #K2 #H2 lapply (le_plus_to_le_r … H) -H lapply (ldrop_inv_ldrop1 … H1 ?) -H1 // diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop_ldrop.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop_ldrop.ma index 88f37fcfb..5a7ac853f 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop_ldrop.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop_ldrop.ma @@ -23,10 +23,8 @@ include "basic_2/relocation/ldrop.ma". theorem ldrop_mono: ∀d,e,L,L1. ⇩[d, e] L ≡ L1 → ∀L2. ⇩[d, e] L ≡ L2 → L1 = L2. #d #e #L #L1 #H elim H -d -e -L -L1 -[ #d #e #L2 #H - >(ldrop_inv_atom1 … H) -L2 // -| #K #I #V #L2 #HL12 - <(ldrop_inv_refl … HL12) -L2 // +[ #d #L2 #H elim (ldrop_inv_atom1 … H) -H // +| #K #I #V #L2 #HL12 <(ldrop_inv_O2 … HL12) -L2 // | #L #K #I #V #e #_ #IHLK #L2 #H lapply (ldrop_inv_ldrop1 … H ?) -H // /2 width=1/ | #L #K1 #I #T #V1 #d #e #_ #HVT1 #IHLK1 #X #H @@ -40,11 +38,8 @@ qed-. theorem ldrop_conf_ge: ∀d1,e1,L,L1. ⇩[d1, e1] L ≡ L1 → ∀e2,L2. ⇩[0, e2] L ≡ L2 → d1 + e1 ≤ e2 → ⇩[0, e2 - e1] L1 ≡ L2. -#d1 #e1 #L #L1 #H elim H -d1 -e1 -L -L1 -[ #d #e #e2 #L2 #H - >(ldrop_inv_atom1 … H) -L2 // -| // -| #L #K #I #V #e #_ #IHLK #e2 #L2 #H #He2 +#d1 #e1 #L #L1 #H elim H -d1 -e1 -L -L1 // +[ #L #K #I #V #e #_ #IHLK #e2 #L2 #H #He2 lapply (ldrop_inv_ldrop1 … H ?) -H /2 width=2/ #HL2 minus_minus_comm /3 width=1/ | #L #K #I #V1 #V2 #d #e #_ #_ #IHLK #e2 #L2 #H #Hdee2 @@ -60,12 +55,13 @@ theorem ldrop_conf_be: ∀L0,L1,d1,e1. ⇩[d1, e1] L0 ≡ L1 → ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → d1 ≤ e2 → e2 ≤ d1 + e1 → ∃∃L. ⇩[0, d1 + e1 - e2] L2 ≡ L & ⇩[0, d1] L1 ≡ L. #L0 #L1 #d1 #e1 #H elim H -L0 -L1 -d1 -e1 -[ #d1 #e1 #L2 #e2 #H >(ldrop_inv_atom1 … H) -H /2 width=3/ +[ #d1 #L2 #e2 #H #Hd1 #_ elim (ldrop_inv_atom1 … H) -H #H1 #H2 destruct + <(le_n_O_to_eq … Hd1) -d1 /2 width=3/ | normalize #L #I #V #L2 #e2 #HL2 #_ #He2 lapply (le_n_O_to_eq … He2) -He2 #H destruct - lapply (ldrop_inv_refl … HL2) -HL2 #H destruct /2 width=3/ + lapply (ldrop_inv_O2 … HL2) -HL2 #H destruct /2 width=3/ | normalize #L0 #K0 #I #V1 #e1 #HLK0 #IHLK0 #L2 #e2 #H #_ #He21 - lapply (ldrop_inv_O1 … H) -H * * #He2 #HL20 + lapply (ldrop_inv_O1_pair1 … H) -H * * #He2 #HL20 [ -IHLK0 -He21 destruct minus_le_minus_minus_comm // /3 width=3/ + elim (IHK01 … HK0L2) -IHK01 -HK0L2 /2 width=1/ >minus_le_minus_minus_comm // /3 width=3/ ] ] qed. @@ -103,11 +99,8 @@ qed. (* Basic_1: was: drop_trans_ge *) theorem ldrop_trans_ge: ∀d1,e1,L1,L. ⇩[d1, e1] L1 ≡ L → ∀e2,L2. ⇩[0, e2] L ≡ L2 → d1 ≤ e2 → ⇩[0, e1 + e2] L1 ≡ L2. -#d1 #e1 #L1 #L #H elim H -d1 -e1 -L1 -L -[ #d #e #e2 #L2 #H - >(ldrop_inv_atom1 … H) -H -L2 // -| // -| /3 width=1/ +#d1 #e1 #L1 #L #H elim H -d1 -e1 -L1 -L // +[ /3 width=1/ | #L1 #L2 #I #V1 #V2 #d #e #H_ #_ #IHL12 #e2 #L #H #Hde2 lapply (lt_to_le_to_lt 0 … Hde2) // #He2 lapply (lt_to_le_to_lt … (e + e2) He2 ?) // #Hee2 @@ -121,19 +114,19 @@ theorem ldrop_trans_le: ∀d1,e1,L1,L. ⇩[d1, e1] L1 ≡ L → ∀e2,L2. ⇩[0, e2] L ≡ L2 → e2 ≤ d1 → ∃∃L0. ⇩[0, e2] L1 ≡ L0 & ⇩[d1 - e2, e1] L0 ≡ L2. #d1 #e1 #L1 #L #H elim H -d1 -e1 -L1 -L -[ #d #e #e2 #L2 #H - >(ldrop_inv_atom1 … H) -L2 /2 width=3/ +[ #d #e2 #L2 #H + elim (ldrop_inv_atom1 … H) -H /2 width=3/ | #K #I #V #e2 #L2 #HL2 #H lapply (le_n_O_to_eq … H) -H #H destruct /2 width=3/ | #L1 #L2 #I #V #e #_ #IHL12 #e2 #L #HL2 #H lapply (le_n_O_to_eq … H) -H #H destruct elim (IHL12 … HL2 ?) -IHL12 -HL2 // #L0 #H #HL0 - lapply (ldrop_inv_refl … H) -H #H destruct /3 width=5/ + lapply (ldrop_inv_O2 … H) -H #H destruct /3 width=5/ | #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #IHL12 #e2 #L #H #He2d - elim (ldrop_inv_O1 … H) -H * + elim (ldrop_inv_O1_pair1 … H) -H * [ -He2d -IHL12 #H1 #H2 destruct /3 width=5/ | -HL12 -HV12 #He2 #HL2 - elim (IHL12 … HL2 ?) -L2 [ >minus_le_minus_minus_comm // /3 width=3/ | /2 width=1/ ] + elim (IHL12 … HL2) -L2 [ >minus_le_minus_minus_comm // /3 width=3/ | /2 width=1/ ] ] ] qed. @@ -149,8 +142,8 @@ lemma ldrop_conf_lt: ∀d1,e1,L,L1. ⇩[d1, e1] L ≡ L1 → ∃∃K1,V1. ⇩[0, e2] L1 ≡ K1. ⓑ{I} V1 & ⇩[d, e1] K2 ≡ K1 & ⇧[d, e1] V1 ≡ V2. #d1 #e1 #L #L1 #H1 #e2 #K2 #I #V2 #H2 #He2d1 -elim (ldrop_conf_le … H1 … H2 ?) -L [2: /2 width=2/] #K #HL1K #HK2 -elim (ldrop_inv_skip1 … HK2 ?) -HK2 [2: /2 width=1/] #K1 #V1 #HK21 #HV12 #H destruct /2 width=5/ +elim (ldrop_conf_le … H1 … H2) -L [2: /2 width=2/] #K #HL1K #HK2 +elim (ldrop_inv_skip1 … HK2) -HK2 [2: /2 width=1/] #K1 #V1 #HK21 #HV12 #H destruct /2 width=5/ qed. lemma ldrop_trans_ge_comm: ∀d1,e1,e2,L1,L2,L. @@ -167,10 +160,10 @@ elim (le_or_ge e1 e2) #He [ lapply (ldrop_conf_ge … HLK1 … HLK2 ?) | lapply (ldrop_conf_ge … HLK2 … HLK1 ?) ] -HLK1 -HLK2 // #HK -lapply (ldrop_fwd_O1_length … HK) #H +lapply (ldrop_fwd_length_minus2 … HK) #H elim (discr_minus_x_xy … H) -H [1,3: normalize H in HK; #HK -lapply (ldrop_inv_refl … HK) -HK #H destruct +lapply (ldrop_inv_O2 … HK) -HK #H destruct lapply (inv_eq_minus_O … H) -H /3 width=1/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop_lpx_sn.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop_lpx_sn.ma index 41490973f..40864c2c7 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop_lpx_sn.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop_lpx_sn.ma @@ -21,7 +21,7 @@ include "basic_2/relocation/ldrop.ma". lemma lpx_sn_deliftable_dropable: ∀R. l_deliftable_sn R → dropable_sn (lpx_sn R). #R #HR #L1 #K1 #d #e #H elim H -L1 -K1 -d -e -[ #d #e #X #H >(lpx_sn_inv_atom1 … H) -H /2 width=3/ +[ #d #X #H >(lpx_sn_inv_atom1 … H) -H /2 width=3/ | #K1 #I #V1 #X #H elim (lpx_sn_inv_pair1 … H) -H #L2 #V2 #HL12 #HV12 #H destruct /3 width=5/ | #L1 #K1 #I #V1 #e #_ #IHLK1 #X #H @@ -37,7 +37,7 @@ qed-. lemma lpx_sn_liftable_dedropable: ∀R. (∀L. reflexive ? (R L)) → l_liftable R → dedropable_sn (lpx_sn R). #R #H1R #H2R #L1 #K1 #d #e #H elim H -L1 -K1 -d -e -[ #d #e #X #H >(lpx_sn_inv_atom1 … H) -H /2 width=3/ +[ #d #X #H >(lpx_sn_inv_atom1 … H) -H /2 width=3/ | #K1 #I #V1 #X #H elim (lpx_sn_inv_pair1 … H) -H #K2 #V2 #HK12 #HV12 #H destruct /3 width=5/ | #L1 #K1 #I #V1 #e #_ #IHLK1 #K2 #HK12 @@ -53,7 +53,7 @@ qed-. fact lpx_sn_dropable_aux: ∀R,L2,K2,d,e. ⇩[d, e] L2 ≡ K2 → ∀L1. lpx_sn R L1 L2 → d = 0 → ∃∃K1. ⇩[0, e] L1 ≡ K1 & lpx_sn R K1 K2. #R #L2 #K2 #d #e #H elim H -L2 -K2 -d -e -[ #d #e #X #H >(lpx_sn_inv_atom2 … H) -H /2 width=3/ +[ #d #X #H >(lpx_sn_inv_atom2 … H) -H /2 width=3/ | #K2 #I #V2 #X #H elim (lpx_sn_inv_pair2 … H) -H #K1 #V1 #HK12 #HV12 #H destruct /3 width=5/ | #L2 #K2 #I #V2 #e #_ #IHLK2 #X #H #_ diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/lift.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/lift.ma index 909b02125..da4dd9c5a 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/lift.ma @@ -37,12 +37,12 @@ interpretation "relocation" 'RLift d e T1 T2 = (lift d e T1 T2). (* Basic inversion lemmas ***************************************************) -fact lift_inv_refl_O2_aux: ∀d,e,T1,T2. ⇧[d, e] T1 ≡ T2 → e = 0 → T1 = T2. +fact lift_inv_O2_aux: ∀d,e,T1,T2. ⇧[d, e] T1 ≡ T2 → e = 0 → T1 = T2. #d #e #T1 #T2 #H elim H -d -e -T1 -T2 // /3 width=1/ -qed. +qed-. -lemma lift_inv_refl_O2: ∀d,T1,T2. ⇧[d, 0] T1 ≡ T2 → T1 = T2. -/2 width=4/ qed-. +lemma lift_inv_O2: ∀d,T1,T2. ⇧[d, 0] T1 ≡ T2 → T1 = T2. +/2 width=4 by lift_inv_O2_aux/ qed-. fact lift_inv_sort1_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → ∀k. T1 = ⋆k → T2 = ⋆k. #d #e #T1 #T2 * -d -e -T1 -T2 // @@ -50,10 +50,10 @@ fact lift_inv_sort1_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → ∀k. T1 = ⋆k | #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct | #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct ] -qed. +qed-. lemma lift_inv_sort1: ∀d,e,T2,k. ⇧[d,e] ⋆k ≡ T2 → T2 = ⋆k. -/2 width=5/ qed-. +/2 width=5 by lift_inv_sort1_aux/ qed-. fact lift_inv_lref1_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → ∀i. T1 = #i → (i < d ∧ T2 = #i) ∨ (d ≤ i ∧ T2 = #(i + e)). @@ -65,11 +65,11 @@ fact lift_inv_lref1_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → ∀i. T1 = #i → | #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct | #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct ] -qed. +qed-. lemma lift_inv_lref1: ∀d,e,T2,i. ⇧[d,e] #i ≡ T2 → (i < d ∧ T2 = #i) ∨ (d ≤ i ∧ T2 = #(i + e)). -/2 width=3/ qed-. +/2 width=3 by lift_inv_lref1_aux/ qed-. lemma lift_inv_lref1_lt: ∀d,e,T2,i. ⇧[d,e] #i ≡ T2 → i < d → T2 = #i. #d #e #T2 #i #H elim (lift_inv_lref1 … H) -H * // @@ -89,10 +89,10 @@ fact lift_inv_gref1_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → ∀p. T1 = §p → | #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct | #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct ] -qed. +qed-. lemma lift_inv_gref1: ∀d,e,T2,p. ⇧[d,e] §p ≡ T2 → T2 = §p. -/2 width=5/ qed-. +/2 width=5 by lift_inv_gref1_aux/ qed-. fact lift_inv_bind1_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → ∀a,I,V1,U1. T1 = ⓑ{a,I} V1.U1 → @@ -106,12 +106,12 @@ fact lift_inv_bind1_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → | #b #J #W1 #W2 #T1 #T2 #d #e #HW #HT #a #I #V1 #U1 #H destruct /2 width=5/ | #J #W1 #W2 #T1 #T2 #d #e #_ #HT #a #I #V1 #U1 #H destruct ] -qed. +qed-. lemma lift_inv_bind1: ∀d,e,T2,a,I,V1,U1. ⇧[d,e] ⓑ{a,I} V1. U1 ≡ T2 → ∃∃V2,U2. ⇧[d,e] V1 ≡ V2 & ⇧[d+1,e] U1 ≡ U2 & T2 = ⓑ{a,I} V2. U2. -/2 width=3/ qed-. +/2 width=3 by lift_inv_bind1_aux/ qed-. fact lift_inv_flat1_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → ∀I,V1,U1. T1 = ⓕ{I} V1.U1 → @@ -125,12 +125,12 @@ fact lift_inv_flat1_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → | #a #J #W1 #W2 #T1 #T2 #d #e #_ #_ #I #V1 #U1 #H destruct | #J #W1 #W2 #T1 #T2 #d #e #HW #HT #I #V1 #U1 #H destruct /2 width=5/ ] -qed. +qed-. lemma lift_inv_flat1: ∀d,e,T2,I,V1,U1. ⇧[d,e] ⓕ{I} V1. U1 ≡ T2 → ∃∃V2,U2. ⇧[d,e] V1 ≡ V2 & ⇧[d,e] U1 ≡ U2 & T2 = ⓕ{I} V2. U2. -/2 width=3/ qed-. +/2 width=3 by lift_inv_flat1_aux/ qed-. fact lift_inv_sort2_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → ∀k. T2 = ⋆k → T1 = ⋆k. #d #e #T1 #T2 * -d -e -T1 -T2 // @@ -138,11 +138,11 @@ fact lift_inv_sort2_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → ∀k. T2 = ⋆k | #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct | #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct ] -qed. +qed-. (* Basic_1: was: lift_gen_sort *) lemma lift_inv_sort2: ∀d,e,T1,k. ⇧[d,e] T1 ≡ ⋆k → T1 = ⋆k. -/2 width=5/ qed-. +/2 width=5 by lift_inv_sort2_aux/ qed-. fact lift_inv_lref2_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → ∀i. T2 = #i → (i < d ∧ T1 = #i) ∨ (d + e ≤ i ∧ T1 = #(i - e)). @@ -154,12 +154,12 @@ fact lift_inv_lref2_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → ∀i. T2 = #i → | #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct | #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct ] -qed. +qed-. (* Basic_1: was: lift_gen_lref *) lemma lift_inv_lref2: ∀d,e,T1,i. ⇧[d,e] T1 ≡ #i → (i < d ∧ T1 = #i) ∨ (d + e ≤ i ∧ T1 = #(i - e)). -/2 width=3/ qed-. +/2 width=3 by lift_inv_lref2_aux/ qed-. (* Basic_1: was: lift_gen_lref_lt *) lemma lift_inv_lref2_lt: ∀d,e,T1,i. ⇧[d,e] T1 ≡ #i → i < d → T1 = #i. @@ -192,10 +192,10 @@ fact lift_inv_gref2_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → ∀p. T2 = §p → | #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct | #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct ] -qed. +qed-. lemma lift_inv_gref2: ∀d,e,T1,p. ⇧[d,e] T1 ≡ §p → T1 = §p. -/2 width=5/ qed-. +/2 width=5 by lift_inv_gref2_aux/ qed-. fact lift_inv_bind2_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → ∀a,I,V2,U2. T2 = ⓑ{a,I} V2.U2 → @@ -209,13 +209,13 @@ fact lift_inv_bind2_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → | #b #J #W1 #W2 #T1 #T2 #d #e #HW #HT #a #I #V2 #U2 #H destruct /2 width=5/ | #J #W1 #W2 #T1 #T2 #d #e #_ #_ #a #I #V2 #U2 #H destruct ] -qed. +qed-. (* Basic_1: was: lift_gen_bind *) lemma lift_inv_bind2: ∀d,e,T1,a,I,V2,U2. ⇧[d,e] T1 ≡ ⓑ{a,I} V2. U2 → ∃∃V1,U1. ⇧[d,e] V1 ≡ V2 & ⇧[d+1,e] U1 ≡ U2 & T1 = ⓑ{a,I} V1. U1. -/2 width=3/ qed-. +/2 width=3 by lift_inv_bind2_aux/ qed-. fact lift_inv_flat2_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → ∀I,V2,U2. T2 = ⓕ{I} V2.U2 → @@ -229,13 +229,13 @@ fact lift_inv_flat2_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → | #a #J #W1 #W2 #T1 #T2 #d #e #_ #_ #I #V2 #U2 #H destruct | #J #W1 #W2 #T1 #T2 #d #e #HW #HT #I #V2 #U2 #H destruct /2 width=5/ ] -qed. +qed-. (* Basic_1: was: lift_gen_flat *) lemma lift_inv_flat2: ∀d,e,T1,I,V2,U2. ⇧[d,e] T1 ≡ ⓕ{I} V2. U2 → ∃∃V1,U1. ⇧[d,e] V1 ≡ V2 & ⇧[d,e] U1 ≡ U2 & T1 = ⓕ{I} V1. U1. -/2 width=3/ qed-. +/2 width=3 by lift_inv_flat2_aux/ qed-. lemma lift_inv_pair_xy_x: ∀d,e,I,V,T. ⇧[d, e] ②{I} V. T ≡ V → ⊥. #d #e #J #V elim V -V diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lsuba_ldrop.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lsuba_ldrop.ma index 247a8b221..b32e6ba93 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lsuba_ldrop.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/lsuba_ldrop.ma @@ -24,17 +24,17 @@ lemma lsuba_ldrop_O1_conf: ∀L1,L2. L1 ⁝⊑ L2 → ∀K1,e. ⇩[0, e] L1 ≡ #L1 #L2 #H elim H -L1 -L2 [ /2 width=3/ | #I #L1 #L2 #V #_ #IHL12 #K1 #e #H - elim (ldrop_inv_O1 … H) -H * #He #HLK1 + elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK1 [ destruct elim (IHL12 L1 0 ?) -IHL12 // #X #HL12 #H - <(ldrop_inv_refl … H) in HL12; -H /3 width=3/ + <(ldrop_inv_O2 … H) in HL12; -H /3 width=3/ | elim (IHL12 … HLK1) -L1 /3 width=3/ ] | #L1 #L2 #V #W #A #HV #HW #_ #IHL12 #K1 #e #H - elim (ldrop_inv_O1 … H) -H * #He #HLK1 + elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK1 [ destruct - elim (IHL12 L1 0 ?) -IHL12 // #X #HL12 #H - <(ldrop_inv_refl … H) in HL12; -H /3 width=3/ + elim (IHL12 L1 0) -IHL12 // #X #HL12 #H + <(ldrop_inv_O2 … H) in HL12; -H /3 width=3/ | elim (IHL12 … HLK1) -L1 /3 width=3/ ] ] @@ -46,17 +46,17 @@ lemma lsuba_ldrop_O1_trans: ∀L1,L2. L1 ⁝⊑ L2 → ∀K2,e. ⇩[0, e] L2 ≡ #L1 #L2 #H elim H -L1 -L2 [ /2 width=3/ | #I #L1 #L2 #V #_ #IHL12 #K2 #e #H - elim (ldrop_inv_O1 … H) -H * #He #HLK2 + elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK2 [ destruct - elim (IHL12 L2 0 ?) -IHL12 // #X #HL12 #H - <(ldrop_inv_refl … H) in HL12; -H /3 width=3/ + elim (IHL12 L2 0) -IHL12 // #X #HL12 #H + <(ldrop_inv_O2 … H) in HL12; -H /3 width=3/ | elim (IHL12 … HLK2) -L2 /3 width=3/ ] | #L1 #L2 #V #W #A #HV #HW #_ #IHL12 #K2 #e #H - elim (ldrop_inv_O1 … H) -H * #He #HLK2 + elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK2 [ destruct - elim (IHL12 L2 0 ?) -IHL12 // #X #HL12 #H - <(ldrop_inv_refl … H) in HL12; -H /3 width=3/ + elim (IHL12 L2 0) -IHL12 // #X #HL12 #H + <(ldrop_inv_O2 … H) in HL12; -H /3 width=3/ | elim (IHL12 … HLK2) -L2 /3 width=3/ ] ] diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/cpss.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/cpss.ma index a10cc99f5..ff5a7d2d5 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/cpss.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/substitution/cpss.ma @@ -72,7 +72,7 @@ qed-. lemma cpss_append: l_appendable_sn … cpss. #K #T1 #T2 #H elim H -K -T1 -T2 // /2 width=1/ #K #K0 #V1 #V2 #W2 #i #HK0 #_ #HVW2 #IHV12 #L -lapply (ldrop_fwd_ldrop2_length … HK0) #H +lapply (ldrop_fwd_length_lt2 … HK0) #H @(cpss_delta … (L@@K0) V1 … HVW2) // @(ldrop_O1_append_sn_le … HK0) /2 width=2/ (**) (* /3/ does not work *) qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/fsups.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/fsups.ma index 07a7b836d..74fa55f3a 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/fsups.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/substitution/fsups.ma @@ -12,11 +12,11 @@ (* *) (**************************************************************************) -include "basic_2/substitution/fsupp.ma". +include "basic_2/relocation/fsupq.ma". (* STAR-ITERATED SUPCLOSURE *************************************************) -definition fsups: bi_relation lenv term ≝ bi_star … fsup. +definition fsups: bi_relation lenv term ≝ bi_TC … fsupq. interpretation "star-iterated structural successor (closure)" 'SupTermStar L1 T1 L2 T2 = (fsups L1 T1 L2 T2). @@ -24,17 +24,17 @@ interpretation "star-iterated structural successor (closure)" (* Basic eliminators ********************************************************) lemma fsups_ind: ∀L1,T1. ∀R:relation2 lenv term. R L1 T1 → - (∀L,L2,T,T2. ⦃L1, T1⦄ ⊃* ⦃L, T⦄ → ⦃L, T⦄ ⊃ ⦃L2, T2⦄ → R L T → R L2 T2) → + (∀L,L2,T,T2. ⦃L1, T1⦄ ⊃* ⦃L, T⦄ → ⦃L, T⦄ ⊃⸮ ⦃L2, T2⦄ → R L T → R L2 T2) → ∀L2,T2. ⦃L1, T1⦄ ⊃* ⦃L2, T2⦄ → R L2 T2. #L1 #T1 #R #IH1 #IH2 #L2 #T2 #H -@(bi_star_ind … IH1 IH2 ? ? H) +@(bi_TC_star_ind … IH1 IH2 ? ? H) // qed-. lemma fsups_ind_dx: ∀L2,T2. ∀R:relation2 lenv term. R L2 T2 → - (∀L1,L,T1,T. ⦃L1, T1⦄ ⊃ ⦃L, T⦄ → ⦃L, T⦄ ⊃* ⦃L2, T2⦄ → R L T → R L1 T1) → + (∀L1,L,T1,T. ⦃L1, T1⦄ ⊃⸮ ⦃L, T⦄ → ⦃L, T⦄ ⊃* ⦃L2, T2⦄ → R L T → R L1 T1) → ∀L1,T1. ⦃L1, T1⦄ ⊃* ⦃L2, T2⦄ → R L1 T1. #L2 #T2 #R #IH1 #IH2 #L1 #T1 #H -@(bi_star_ind_dx … IH1 IH2 ? ? H) +@(bi_TC_star_ind_dx … IH1 IH2 ? ? H) // qed-. (* Basic properties *********************************************************) @@ -42,33 +42,22 @@ qed-. lemma fsups_refl: bi_reflexive … fsups. /2 width=1/ qed. -lemma fsupp_fsups: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⊃+ ⦃L2, T2⦄ → ⦃L1, T1⦄ ⊃* ⦃L2, T2⦄. +lemma fsupq_fsups: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⊃⸮ ⦃L2, T2⦄ → ⦃L1, T1⦄ ⊃* ⦃L2, T2⦄. /2 width=1/ qed. -lemma fsup_fsups: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⊃ ⦃L2, T2⦄ → ⦃L1, T1⦄ ⊃* ⦃L2, T2⦄. -/2 width=1/ qed. - -lemma fsups_strap1: ∀L1,L,L2,T1,T,T2. ⦃L1, T1⦄ ⊃* ⦃L, T⦄ → ⦃L, T⦄ ⊃ ⦃L2, T2⦄ → +lemma fsups_strap1: ∀L1,L,L2,T1,T,T2. ⦃L1, T1⦄ ⊃* ⦃L, T⦄ → ⦃L, T⦄ ⊃⸮ ⦃L2, T2⦄ → ⦃L1, T1⦄ ⊃* ⦃L2, T2⦄. /2 width=4/ qed. -lemma fsups_strap2: ∀L1,L,L2,T1,T,T2. ⦃L1, T1⦄ ⊃ ⦃L, T⦄ → ⦃L, T⦄ ⊃* ⦃L2, T2⦄ → +lemma fsups_strap2: ∀L1,L,L2,T1,T,T2. ⦃L1, T1⦄ ⊃⸮ ⦃L, T⦄ → ⦃L, T⦄ ⊃* ⦃L2, T2⦄ → ⦃L1, T1⦄ ⊃* ⦃L2, T2⦄. /2 width=4/ qed. -lemma fsups_fsupp_fsupp: ∀L1,L,L2,T1,T,T2. ⦃L1, T1⦄ ⊃* ⦃L, T⦄ → - ⦃L, T⦄ ⊃+ ⦃L2, T2⦄ → ⦃L1, T1⦄ ⊃+ ⦃L2, T2⦄. -/2 width=4/ qed. - -lemma fsupp_fsups_fsupp: ∀L1,L,L2,T1,T,T2. ⦃L1, T1⦄ ⊃+ ⦃L, T⦄ → - ⦃L, T⦄ ⊃* ⦃L2, T2⦄ → ⦃L1, T1⦄ ⊃+ ⦃L2, T2⦄. -/2 width=4/ qed. - (* Basic forward lemmas *****************************************************) lemma fsups_fwd_fw: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⊃* ⦃L2, T2⦄ → ♯{L2, T2} ≤ ♯{L1, T1}. #L1 #L2 #T1 #T2 #H @(fsups_ind … H) -L2 -T2 // -/4 width=3 by fsup_fwd_fw, lt_to_le_to_lt, lt_to_le/ (**) (* slow even with trace *) +/3 width=3 by fsupq_fwd_fw, transitive_le/ (**) (* slow even with trace *) qed-. (* (* Advanced inversion lemmas on plus-iterated supclosure ********************) diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/lpss_ldrop.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/lpss_ldrop.ma index 4237c7666..7a549f143 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/lpss_ldrop.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/substitution/lpss_ldrop.ma @@ -35,9 +35,13 @@ lemma lpss_ldrop_trans_O1: dropable_dx lpss. lemma fsup_cpss_trans: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⊃ ⦃L2, T2⦄ → ∀U2. L2 ⊢ T2 ▶* U2 → ∃∃L,U1. L1 ⊢ ▶* L & L ⊢ T1 ▶* U1 & ⦃L, U1⦄ ⊃ ⦃L2, U2⦄. #L1 #L2 #T1 #T2 #H elim H -L1 -L2 -T1 -T2 [2: * ] [1,2,3,4,5: /3 width=5/ ] -#L1 #K1 #K2 #T1 #T2 #U1 #d #e #HLK1 #HTU1 #_ #IHT12 #U2 #HTU2 -elim (IHT12 … HTU2) -IHT12 -HTU2 #K #T #HK1 #HT1 #HT2 -elim (lift_total T d e) #U #HTU -elim (ldrop_lpss_trans … HLK1 … HK1) -HLK1 -HK1 #L2 #HL12 #HL2K -lapply (cpss_lift … HT1 … HL2K … HTU1 … HTU) -HT1 -HTU1 /3 width=11/ +[ #L #K #U #T #d #e #HLK #HUT #He #U2 #HU2 + elim (lift_total U2 d e) #T2 #HUT2 + lapply (cpss_lift … HU2 … HLK … HUT … HUT2) -HU2 -HUT /3 width=9/ +| #L1 #K1 #K2 #T1 #T2 #U1 #d #e #HLK1 #HTU1 #_ #IHT12 #U2 #HTU2 + elim (IHT12 … HTU2) -IHT12 -HTU2 #K #T #HK1 #HT1 #HT2 + elim (lift_total T d e) #U #HTU + elim (ldrop_lpss_trans … HLK1 … HK1) -HLK1 -HK1 #L2 #HL12 #HL2K + lapply (cpss_lift … HT1 … HL2K … HTU1 … HTU) -HT1 -HTU1 /3 width=11/ +] qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/lsubr.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/lsubr.ma index e3afc573d..eee557904 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/lsubr.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/substitution/lsubr.ma @@ -98,14 +98,14 @@ lemma lsubr_fwd_ldrop2_abbr: ∀L1,L2. L1 ⊑ L2 → ∃∃K1. K1 ⊑ K2 & ⇩[0, i] L1 ≡ K1. ⓓW. #L1 #L2 #H elim H -L1 -L2 [ #L #K2 #W #i #H - lapply (ldrop_inv_atom1 … H) -H #H destruct + elim (ldrop_inv_atom1 … H) -H #H destruct | #L1 #L2 #V #HL12 #IHL12 #K2 #W #i #H - elim (ldrop_inv_O1 … H) -H * #Hi #HLK2 destruct [ -IHL12 | -HL12 ] + elim (ldrop_inv_O1_pair1 … H) -H * #Hi #HLK2 destruct [ -IHL12 | -HL12 ] [ /2 width=3/ | elim (IHL12 … HLK2) -IHL12 -HLK2 /3 width=3/ ] | #I #L1 #L2 #V1 #V2 #_ #IHL12 #K2 #W #i #H - elim (ldrop_inv_O1 … H) -H * #Hi #HLK2 destruct + elim (ldrop_inv_O1_pair1 … H) -H * #Hi #HLK2 destruct elim (IHL12 … HLK2) -IHL12 -HLK2 /3 width=3/ ] qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/unfold/cpqs.ma b/matita/matita/contribs/lambdadelta/basic_2/unfold/cpqs.ma index cddb52b6e..fdec753dd 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/unfold/cpqs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/unfold/cpqs.ma @@ -64,7 +64,7 @@ qed-. lemma cpqs_append: l_appendable_sn … cpqs. #K #T1 #T2 #H elim H -K -T1 -T2 // /2 width=1/ /2 width=3/ #K #K0 #V1 #V2 #W2 #i #HK0 #_ #HVW2 #IHV12 #L -lapply (ldrop_fwd_ldrop2_length … HK0) #H +lapply (ldrop_fwd_length_lt2 … HK0) #H @(cpqs_delta … (L@@K0) V1 … HVW2) // @(ldrop_O1_append_sn_le … HK0) /2 width=2/ (**) (* /3/ does not work *) qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/unfold/lpqs_ldrop.ma b/matita/matita/contribs/lambdadelta/basic_2/unfold/lpqs_ldrop.ma index 854420d60..e121c0cc4 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/unfold/lpqs_ldrop.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/unfold/lpqs_ldrop.ma @@ -35,9 +35,13 @@ lemma lpqs_ldrop_trans_O1: dropable_dx lpqs. lemma fsup_cpqs_trans: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⊃ ⦃L2, T2⦄ → ∀U2. L2 ⊢ T2 ➤* U2 → ∃∃L,U1. L1 ⊢ ➤* L & L ⊢ T1 ➤* U1 & ⦃L, U1⦄ ⊃ ⦃L2, U2⦄. #L1 #L2 #T1 #T2 #H elim H -L1 -L2 -T1 -T2 [2: * ] [1,2,3,4,5: /3 width=5/ ] -#L1 #K1 #K2 #T1 #T2 #U1 #d #e #HLK1 #HTU1 #_ #IHT12 #U2 #HTU2 -elim (IHT12 … HTU2) -IHT12 -HTU2 #K #T #HK1 #HT1 #HT2 -elim (lift_total T d e) #U #HTU -elim (ldrop_lpqs_trans … HLK1 … HK1) -HLK1 -HK1 #L2 #HL12 #HL2K -lapply (cpqs_lift … HT1 … HL2K … HTU1 … HTU) -HT1 -HTU1 /3 width=11/ +[ #L #K #U #T #d #e #HLK #HUT #He #U2 #HU2 + elim (lift_total U2 d e) #T2 #HUT2 + lapply (cpqs_lift … HU2 … HLK … HUT … HUT2) -HU2 -HUT /3 width=9/ +| #L1 #K1 #K2 #T1 #T2 #U1 #d #e #HLK1 #HTU1 #_ #IHT12 #U2 #HTU2 + elim (IHT12 … HTU2) -IHT12 -HTU2 #K #T #HK1 #HT1 #HT2 + elim (lift_total T d e) #U #HTU + elim (ldrop_lpqs_trans … HLK1 … HK1) -HLK1 -HK1 #L2 #HL12 #HL2K + lapply (cpqs_lift … HT1 … HL2K … HTU1 … HTU) -HT1 -HTU1 /3 width=11/ +] qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl b/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl index 5d5b58611..714792343 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl +++ b/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl @@ -204,7 +204,7 @@ table { class "orange" [ { "relocation" * } { [ { "structural successor for closures" * } { - [ "fsup ( ⦃?,?⦄ ⊃ ⦃?,?⦄ )" "fsupq ( ⦃?,?⦄ ⊃⸮ ⦃?,?⦄ )" * ] + [ "fsup ( ⦃?,?⦄ ⊃ ⦃?,?⦄ )" "fsupq ( ⦃?,?⦄ ⊃⸮ ⦃?,?⦄ )" "fsupq_alt" * ] } ] [ { "global env. slicing" * } {