From 5d5461cffbec45aebc0db0fe2fcb072592a90efe Mon Sep 17 00:00:00 2001 From: Ferruccio Guidi Date: Fri, 17 Jan 2014 19:10:33 +0000 Subject: [PATCH] commit of the "static" component --- .../lambdadelta/basic_2/static/aaa.ma | 16 +-- .../lambdadelta/basic_2/static/aaa_fqus.ma | 2 +- .../lambdadelta/basic_2/static/aaa_lift.ma | 53 ++++----- .../lambdadelta/basic_2/static/aaa_lifts.ma | 6 +- .../contribs/lambdadelta/basic_2/static/da.ma | 16 +-- .../lambdadelta/basic_2/static/da_lift.ma | 66 ++++++------ .../lambdadelta/basic_2/static/lsuba_ldrop.ma | 40 +++---- .../lambdadelta/basic_2/static/lsubd.ma | 56 +++++----- .../lambdadelta/basic_2/static/ssta.ma | 40 +++---- .../lambdadelta/basic_2/static/ssta_lift.ma | 101 ++++++++++-------- 10 files changed, 206 insertions(+), 190 deletions(-) diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/aaa.ma b/matita/matita/contribs/lambdadelta/basic_2/static/aaa.ma index cc1ed23c2..5e93f55f1 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/aaa.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/aaa.ma @@ -22,7 +22,7 @@ include "basic_2/relocation/ldrop.ma". (* activate genv *) inductive aaa: relation4 genv lenv term aarity ≝ | aaa_sort: ∀G,L,k. aaa G L (⋆k) (⓪) -| aaa_lref: ∀I,G,L,K,V,B,i. ⇩[0, i] L ≡ K. ⓑ{I}V → aaa G K V B → aaa G L (#i) B +| aaa_lref: ∀I,G,L,K,V,B,i. ⇩[i] L ≡ K. ⓑ{I}V → aaa G K V B → aaa G L (#i) B | aaa_abbr: ∀a,G,L,V,T,B,A. aaa G L V B → aaa G (L.ⓓV) T A → aaa G L (ⓓ{a}V.T) A | aaa_abst: ∀a,G,L,V,T,B,A. @@ -51,10 +51,10 @@ lemma aaa_inv_sort: ∀G,L,A,k. ⦃G, L⦄ ⊢ ⋆k ⁝ A → A = ⓪. /2 width=6 by aaa_inv_sort_aux/ qed-. fact aaa_inv_lref_aux: ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀i. T = #i → - ∃∃I,K,V. ⇩[0, i] L ≡ K.ⓑ{I} V & ⦃G, K⦄ ⊢ V ⁝ A. + ∃∃I,K,V. ⇩[i] L ≡ K.ⓑ{I} V & ⦃G, K⦄ ⊢ V ⁝ A. #G #L #T #A * -G -L -T -A [ #G #L #k #i #H destruct -| #I #G #L #K #V #B #j #HLK #HB #i #H destruct /2 width=5/ +| #I #G #L #K #V #B #j #HLK #HB #i #H destruct /2 width=5 by ex2_3_intro/ | #a #G #L #V #T #B #A #_ #_ #i #H destruct | #a #G #L #V #T #B #A #_ #_ #i #H destruct | #G #L #V #T #B #A #_ #_ #i #H destruct @@ -63,7 +63,7 @@ fact aaa_inv_lref_aux: ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀i. T = #i → qed-. lemma aaa_inv_lref: ∀G,L,A,i. ⦃G, L⦄ ⊢ #i ⁝ A → - ∃∃I,K,V. ⇩[0, i] L ≡ K. ⓑ{I} V & ⦃G, K⦄ ⊢ V ⁝ A. + ∃∃I,K,V. ⇩[i] L ≡ K. ⓑ{I} V & ⦃G, K⦄ ⊢ V ⁝ A. /2 width=3 by aaa_inv_lref_aux/ qed-. fact aaa_inv_gref_aux: ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀p. T = §p → ⊥. @@ -85,7 +85,7 @@ fact aaa_inv_abbr_aux: ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀a,W,U. T = ⓓ{ #G #L #T #A * -G -L -T -A [ #G #L #k #a #W #U #H destruct | #I #G #L #K #V #B #i #_ #_ #a #W #U #H destruct -| #b #G #L #V #T #B #A #HV #HT #a #W #U #H destruct /2 width=2/ +| #b #G #L #V #T #B #A #HV #HT #a #W #U #H destruct /2 width=2 by ex2_intro/ | #b #G #L #V #T #B #A #_ #_ #a #W #U #H destruct | #G #L #V #T #B #A #_ #_ #a #W #U #H destruct | #G #L #V #T #A #_ #_ #a #W #U #H destruct @@ -102,7 +102,7 @@ fact aaa_inv_abst_aux: ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀a,W,U. T = ⓛ{ [ #G #L #k #a #W #U #H destruct | #I #G #L #K #V #B #i #_ #_ #a #W #U #H destruct | #b #G #L #V #T #B #A #_ #_ #a #W #U #H destruct -| #b #G #L #V #T #B #A #HV #HT #a #W #U #H destruct /2 width=5/ +| #b #G #L #V #T #B #A #HV #HT #a #W #U #H destruct /2 width=5 by ex3_2_intro/ | #G #L #V #T #B #A #_ #_ #a #W #U #H destruct | #G #L #V #T #A #_ #_ #a #W #U #H destruct ] @@ -119,7 +119,7 @@ fact aaa_inv_appl_aux: ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀W,U. T = ⓐW.U | #I #G #L #K #V #B #i #_ #_ #W #U #H destruct | #a #G #L #V #T #B #A #_ #_ #W #U #H destruct | #a #G #L #V #T #B #A #_ #_ #W #U #H destruct -| #G #L #V #T #B #A #HV #HT #W #U #H destruct /2 width=3/ +| #G #L #V #T #B #A #HV #HT #W #U #H destruct /2 width=3 by ex2_intro/ | #G #L #V #T #A #_ #_ #W #U #H destruct ] qed-. @@ -136,7 +136,7 @@ fact aaa_inv_cast_aux: ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀W,U. T = ⓝW. | #a #G #L #V #T #B #A #_ #_ #W #U #H destruct | #a #G #L #V #T #B #A #_ #_ #W #U #H destruct | #G #L #V #T #B #A #_ #_ #W #U #H destruct -| #G #L #V #T #A #HV #HT #W #U #H destruct /2 width=1/ +| #G #L #V #T #A #HV #HT #W #U #H destruct /2 width=1 by conj/ ] qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/aaa_fqus.ma b/matita/matita/contribs/lambdadelta/basic_2/static/aaa_fqus.ma index cc3777262..f9bd61ada 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/aaa_fqus.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/aaa_fqus.ma @@ -39,7 +39,7 @@ lemma aaa_fqu_conf: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊃ ⦃G2, L2, T2⦄ [ elim (aaa_inv_appl … H) | elim (aaa_inv_cast … H) ] -H /2 width=2 by ex_intro/ -| /3 width=8 by aaa_inv_lift, ex_intro/ +| /3 width=9 by aaa_inv_lift, ex_intro/ ] qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/aaa_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/static/aaa_lift.ma index ea039002f..f0b5f8a40 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/aaa_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/aaa_lift.ma @@ -19,54 +19,55 @@ include "basic_2/static/aaa.ma". (* Properties on basic relocation *******************************************) -lemma aaa_lift: ∀G,L1,T1,A. ⦃G, L1⦄ ⊢ T1 ⁝ A → ∀L2,d,e. ⇩[d, e] L2 ≡ L1 → +lemma aaa_lift: ∀G,L1,T1,A. ⦃G, L1⦄ ⊢ T1 ⁝ A → ∀L2,s,d,e. ⇩[s, d, e] L2 ≡ L1 → ∀T2. ⇧[d, e] T1 ≡ T2 → ⦃G, L2⦄ ⊢ T2 ⁝ A. #G #L1 #T1 #A #H elim H -G -L1 -T1 -A -[ #G #L1 #k #L2 #d #e #_ #T2 #H +[ #G #L1 #k #L2 #s #d #e #_ #T2 #H >(lift_inv_sort1 … H) -H // -| #I #G #L1 #K1 #V1 #B #i #HLK1 #_ #IHB #L2 #d #e #HL21 #T2 #H +| #I #G #L1 #K1 #V1 #B #i #HLK1 #_ #IHB #L2 #s #d #e #HL21 #T2 #H elim (lift_inv_lref1 … H) -H * #Hid #H destruct [ elim (ldrop_trans_le … HL21 … HLK1) -L1 /2 width=2 by lt_to_le/ #X #HLK2 #H elim (ldrop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid #K2 #V2 #HK21 #HV12 #H destruct - /3 width=8/ - | lapply (ldrop_trans_ge … HL21 … HLK1 ?) -L1 /3 width=8 by aaa_lref/ + /3 width=9 by aaa_lref/ + | lapply (ldrop_trans_ge … HL21 … HLK1 ?) -L1 + /3 width=9 by aaa_lref, ldrop_inv_gen/ ] -| #a #G #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #d #e #HL21 #X #H +| #a #G #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #s #d #e #HL21 #X #H elim (lift_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct - /4 width=4 by aaa_abbr, ldrop_skip/ -| #a #G #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #d #e #HL21 #X #H + /4 width=5 by aaa_abbr, ldrop_skip/ +| #a #G #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #s #d #e #HL21 #X #H elim (lift_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct - /4 width=4 by aaa_abst, ldrop_skip/ -| #G #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #d #e #HL21 #X #H + /4 width=5 by aaa_abst, ldrop_skip/ +| #G #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #s #d #e #HL21 #X #H elim (lift_inv_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct - /3 width=4 by aaa_appl/ -| #G #L1 #V1 #T1 #A #_ #_ #IH1 #IH2 #L2 #d #e #HL21 #X #H + /3 width=5 by aaa_appl/ +| #G #L1 #V1 #T1 #A #_ #_ #IH1 #IH2 #L2 #s #d #e #HL21 #X #H elim (lift_inv_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct - /3 width=4 by aaa_cast/ + /3 width=5 by aaa_cast/ ] qed. -lemma aaa_inv_lift: ∀G,L2,T2,A. ⦃G, L2⦄ ⊢ T2 ⁝ A → ∀L1,d,e. ⇩[d, e] L2 ≡ L1 → +lemma aaa_inv_lift: ∀G,L2,T2,A. ⦃G, L2⦄ ⊢ T2 ⁝ A → ∀L1,s,d,e. ⇩[s, d, e] L2 ≡ L1 → ∀T1. ⇧[d, e] T1 ≡ T2 → ⦃G, L1⦄ ⊢ T1 ⁝ A. #G #L2 #T2 #A #H elim H -G -L2 -T2 -A -[ #G #L2 #k #L1 #d #e #_ #T1 #H +[ #G #L2 #k #L1 #s #d #e #_ #T1 #H >(lift_inv_sort2 … H) -H // -| #I #G #L2 #K2 #V2 #B #i #HLK2 #_ #IHB #L1 #d #e #HL21 #T1 #H +| #I #G #L2 #K2 #V2 #B #i #HLK2 #_ #IHB #L1 #s #d #e #HL21 #T1 #H elim (lift_inv_lref2 … H) -H * #Hid #H destruct - [ elim (ldrop_conf_lt … HL21 … HLK2) -L2 /3 width=8 by aaa_lref/ - | lapply (ldrop_conf_ge … HL21 … HLK2 ?) -L2 /3 width=8 by aaa_lref/ + [ elim (ldrop_conf_lt … HL21 … HLK2) -L2 /3 width=9 by aaa_lref/ + | lapply (ldrop_conf_ge … HL21 … HLK2 ?) -L2 /3 width=9 by aaa_lref/ ] -| #a #G #L2 #V2 #T2 #B #A #_ #_ #IHB #IHA #L1 #d #e #HL21 #X #H +| #a #G #L2 #V2 #T2 #B #A #_ #_ #IHB #IHA #L1 #s #d #e #HL21 #X #H elim (lift_inv_bind2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct - /4 width=4 by aaa_abbr, ldrop_skip/ -| #a #G #L2 #V2 #T2 #B #A #_ #_ #IHB #IHA #L1 #d #e #HL21 #X #H + /4 width=5 by aaa_abbr, ldrop_skip/ +| #a #G #L2 #V2 #T2 #B #A #_ #_ #IHB #IHA #L1 #s #d #e #HL21 #X #H elim (lift_inv_bind2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct - /4 width=4 by aaa_abst, ldrop_skip/ -| #G #L2 #V2 #T2 #B #A #_ #_ #IHB #IHA #L1 #d #e #HL21 #X #H + /4 width=5 by aaa_abst, ldrop_skip/ +| #G #L2 #V2 #T2 #B #A #_ #_ #IHB #IHA #L1 #s #d #e #HL21 #X #H elim (lift_inv_flat2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct - /3 width=4 by aaa_appl/ -| #G #L2 #V2 #T2 #A #_ #_ #IH1 #IH2 #L1 #d #e #HL21 #X #H + /3 width=5 by aaa_appl/ +| #G #L2 #V2 #T2 #A #_ #_ #IH1 #IH2 #L1 #s #d #e #HL21 #X #H elim (lift_inv_flat2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct - /3 width=4 by aaa_cast/ + /3 width=5 by aaa_cast/ ] qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/aaa_lifts.ma b/matita/matita/contribs/lambdadelta/basic_2/static/aaa_lifts.ma index 6cd3855c5..b803d42ae 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/aaa_lifts.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/aaa_lifts.ma @@ -19,12 +19,12 @@ include "basic_2/static/aaa_lift.ma". (* Properties concerning generic relocation *********************************) -lemma aaa_lifts: ∀G,L1,L2,T2,A,des. ⇩*[des] L2 ≡ L1 → ∀T1. ⇧*[des] T1 ≡ T2 → +lemma aaa_lifts: ∀G,L1,L2,T2,A,s,des. ⇩*[s, des] L2 ≡ L1 → ∀T1. ⇧*[des] T1 ≡ T2 → ⦃G, L1⦄ ⊢ T1 ⁝ A → ⦃G, L2⦄ ⊢ T2 ⁝ A. -#G #L1 #L2 #T2 #A #des #H elim H -L1 -L2 -des +#G #L1 #L2 #T2 #A #s #des #H elim H -L1 -L2 -des [ #L #T1 #H #HT1 <(lifts_inv_nil … H) -H // | #L1 #L #L2 #des #d #e #_ #HL2 #IHL1 #T1 #H #HT1 - elim (lifts_inv_cons … H) -H /3 width=9/ + elim (lifts_inv_cons … H) -H /3 width=10 by aaa_lift/ ] qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/da.ma b/matita/matita/contribs/lambdadelta/basic_2/static/da.ma index dc594f68d..ad4d8c182 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/da.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/da.ma @@ -22,8 +22,8 @@ include "basic_2/static/sd.ma". (* activate genv *) inductive da (h:sh) (g:sd h): relation4 genv lenv term nat ≝ | da_sort: ∀G,L,k,l. deg h g k l → da h g G L (⋆k) l -| da_ldef: ∀G,L,K,V,i,l. ⇩[0, i] L ≡ K.ⓓV → da h g G K V l → da h g G L (#i) l -| da_ldec: ∀G,L,K,W,i,l. ⇩[0, i] L ≡ K.ⓛW → da h g G K W l → da h g G L (#i) (l+1) +| da_ldef: ∀G,L,K,V,i,l. ⇩[i] L ≡ K.ⓓV → da h g G K V l → da h g G L (#i) l +| da_ldec: ∀G,L,K,W,i,l. ⇩[i] L ≡ K.ⓛW → da h g G K W l → da h g G L (#i) (l+1) | da_bind: ∀a,I,G,L,V,T,l. da h g G (L.ⓑ{I}V) T l → da h g G L (ⓑ{a,I}V.T) l | da_flat: ∀I,G,L,V,T,l. da h g G L T l → da h g G L (ⓕ{I}V.T) l . @@ -48,22 +48,22 @@ lemma da_inv_sort: ∀h,g,G,L,k,l. ⦃G, L⦄ ⊢ ⋆k ▪[h, g] l → deg h g k /2 width=5 by da_inv_sort_aux/ qed-. fact da_inv_lref_aux: ∀h,g,G,L,T,l. ⦃G, L⦄ ⊢ T ▪[h, g] l → ∀j. T = #j → - (∃∃K,V. ⇩[0, j] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V ▪[h, g] l) ∨ - (∃∃K,W,l0. ⇩[0, j] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W ▪[h, g] l0 & + (∃∃K,V. ⇩[j] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V ▪[h, g] l) ∨ + (∃∃K,W,l0. ⇩[j] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W ▪[h, g] l0 & l = l0 + 1 ). #h #g #G #L #T #l * -G -L -T -l [ #G #L #k #l #_ #j #H destruct -| #G #L #K #V #i #l #HLK #HV #j #H destruct /3 width=4/ -| #G #L #K #W #i #l #HLK #HW #j #H destruct /3 width=6/ +| #G #L #K #V #i #l #HLK #HV #j #H destruct /3 width=4 by ex2_2_intro, or_introl/ +| #G #L #K #W #i #l #HLK #HW #j #H destruct /3 width=6 by ex3_3_intro, or_intror/ | #a #I #G #L #V #T #l #_ #j #H destruct | #I #G #L #V #T #l #_ #j #H destruct ] qed-. lemma da_inv_lref: ∀h,g,G,L,j,l. ⦃G, L⦄ ⊢ #j ▪[h, g] l → - (∃∃K,V. ⇩[0, j] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V ▪[h, g] l) ∨ - (∃∃K,W,l0. ⇩[0, j] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W ▪[h, g] l0 & l = l0+1). + (∃∃K,V. ⇩[j] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V ▪[h, g] l) ∨ + (∃∃K,W,l0. ⇩[j] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W ▪[h, g] l0 & l = l0+1). /2 width=3 by da_inv_lref_aux/ qed-. fact da_inv_gref_aux: ∀h,g,G,L,T,l. ⦃G, L⦄ ⊢ T ▪[h, g] l → ∀p0. T = §p0 → ⊥. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/da_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/static/da_lift.ma index c0f13ceb7..78d575b8d 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/da_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/da_lift.ma @@ -20,53 +20,59 @@ include "basic_2/static/da.ma". (* Properties on relocation *************************************************) lemma da_lift: ∀h,g,G,L1,T1,l. ⦃G, L1⦄ ⊢ T1 ▪[h, g] l → - ∀L2,d,e. ⇩[d, e] L2 ≡ L1 → ∀T2. ⇧[d, e] T1 ≡ T2 → + ∀L2,s,d,e. ⇩[s, d, e] L2 ≡ L1 → ∀T2. ⇧[d, e] T1 ≡ T2 → ⦃G, L2⦄ ⊢ T2 ▪[h, g] l. #h #g #G #L1 #T1 #l #H elim H -G -L1 -T1 -l -[ #G #L1 #k #l #Hkl #L2 #d #e #_ #X #H - >(lift_inv_sort1 … H) -X /2 width=1/ -| #G #L1 #K1 #V1 #i #l #HLK1 #_ #IHV1 #L2 #d #e #HL21 #X #H +[ #G #L1 #k #l #Hkl #L2 #s #d #e #_ #X #H + >(lift_inv_sort1 … H) -X /2 width=1 by da_sort/ +| #G #L1 #K1 #V1 #i #l #HLK1 #_ #IHV1 #L2 #s #d #e #HL21 #X #H elim (lift_inv_lref1 … H) * #Hid #H destruct - [ elim (ldrop_trans_le … HL21 … HLK1) -L1 /2 width=2/ #X #HLK2 #H - elim (ldrop_inv_skip2 … H) -H /2 width=1/ -Hid #K2 #V2 #HK21 #HV12 #H destruct - /3 width=7/ - | lapply (ldrop_trans_ge … HL21 … HLK1 ?) -L1 // -Hid /3 width=7/ + [ elim (ldrop_trans_le … HL21 … HLK1) -L1 /2 width=2 by lt_to_le/ #X #HLK2 #H + elim (ldrop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid #K2 #V2 #HK21 #HV12 #H destruct + /3 width=9 by da_ldef/ + | lapply (ldrop_trans_ge … HL21 … HLK1 ?) -L1 + /3 width=8 by da_ldef, ldrop_inv_gen/ ] -| #G #L1 #K1 #W1 #i #l #HLK1 #_ #IHW1 #L2 #d #e #HL21 #X #H +| #G #L1 #K1 #W1 #i #l #HLK1 #_ #IHW1 #L2 #s #d #e #HL21 #X #H elim (lift_inv_lref1 … H) * #Hid #H destruct - [ elim (ldrop_trans_le … HL21 … HLK1) -L1 /2 width=2/ #X #HLK2 #H - elim (ldrop_inv_skip2 … H) -H /2 width=1/ -Hid #K2 #W2 #HK21 #HW12 #H destruct - /3 width=7/ - | lapply (ldrop_trans_ge … HL21 … HLK1 ?) -L1 // -Hid /3 width=7/ + [ elim (ldrop_trans_le … HL21 … HLK1) -L1 /2 width=2 by lt_to_le/ #X #HLK2 #H + elim (ldrop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid #K2 #W2 #HK21 #HW12 #H destruct + /3 width=8 by da_ldec/ + | lapply (ldrop_trans_ge … HL21 … HLK1 ?) -L1 + /3 width=8 by da_ldec, ldrop_inv_gen/ ] -| #a #I #G #L1 #V1 #T1 #l #_ #IHT1 #L2 #d #e #HL21 #X #H - elim (lift_inv_bind1 … H) -H #V2 #T2 #HV12 #HU12 #H destruct /4 width=4/ -| #I #G #L1 #V1 #T1 #l #_ #IHT1 #L2 #d #e #HL21 #X #H - elim (lift_inv_flat1 … H) -H #V2 #T2 #HV12 #HU12 #H destruct /3 width=4/ +| #a #I #G #L1 #V1 #T1 #l #_ #IHT1 #L2 #s #d #e #HL21 #X #H + elim (lift_inv_bind1 … H) -H #V2 #T2 #HV12 #HU12 #H destruct + /4 width=5 by da_bind, ldrop_skip/ +| #I #G #L1 #V1 #T1 #l #_ #IHT1 #L2 #s #d #e #HL21 #X #H + elim (lift_inv_flat1 … H) -H #V2 #T2 #HV12 #HU12 #H destruct + /3 width=5 by da_flat/ ] qed. (* Inversion lemmas on relocation *******************************************) lemma da_inv_lift: ∀h,g,G,L2,T2,l. ⦃G, L2⦄ ⊢ T2 ▪[h, g] l → - ∀L1,d,e. ⇩[d, e] L2 ≡ L1 → ∀T1. ⇧[d, e] T1 ≡ T2 → + ∀L1,s,d,e. ⇩[s, d, e] L2 ≡ L1 → ∀T1. ⇧[d, e] T1 ≡ T2 → ⦃G, L1⦄ ⊢ T1 ▪[h, g] l. #h #g #G #L2 #T2 #l #H elim H -G -L2 -T2 -l -[ #G #L2 #k #l #Hkl #L1 #d #e #_ #X #H - >(lift_inv_sort2 … H) -X /2 width=1/ -| #G #L2 #K2 #V2 #i #l #HLK2 #HV2 #IHV2 #L1 #d #e #HL21 #X #H +[ #G #L2 #k #l #Hkl #L1 #s #d #e #_ #X #H + >(lift_inv_sort2 … H) -X /2 width=1 by da_sort/ +| #G #L2 #K2 #V2 #i #l #HLK2 #HV2 #IHV2 #L1 #s #d #e #HL21 #X #H elim (lift_inv_lref2 … H) * #Hid #H destruct [ -HV2 | -IHV2 ] - [ elim (ldrop_conf_lt … HL21 … HLK2) -L2 // /3 width=7/ - | lapply (ldrop_conf_ge … HL21 … HLK2 ?) -L2 // /2 width=4/ + [ elim (ldrop_conf_lt … HL21 … HLK2) -L2 /3 width=8 by da_ldef/ + | lapply (ldrop_conf_ge … HL21 … HLK2 ?) -L2 /2 width=4 by da_ldef/ ] -| #G #L2 #K2 #W2 #i #l #HLK2 #HW2 #IHW2 #L1 #d #e #HL21 #X #H +| #G #L2 #K2 #W2 #i #l #HLK2 #HW2 #IHW2 #L1 #s #d #e #HL21 #X #H elim (lift_inv_lref2 … H) * #Hid #H destruct [ -HW2 | -IHW2 ] - [ elim (ldrop_conf_lt … HL21 … HLK2) -L2 // /3 width=7/ - | lapply (ldrop_conf_ge … HL21 … HLK2 ?) -L2 // /2 width=4/ + [ elim (ldrop_conf_lt … HL21 … HLK2) -L2 /3 width=8 by da_ldec/ + | lapply (ldrop_conf_ge … HL21 … HLK2 ?) -L2 /2 width=4 by da_ldec/ ] -| #a #I #G #L2 #V2 #T2 #l #_ #IHT2 #L1 #d #e #HL21 #X #H - elim (lift_inv_bind2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct /4 width=4/ -| #I #G #L2 #V2 #T2 #l #_ #IHT2 #L1 #d #e #HL21 #X #H - elim (lift_inv_flat2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct /3 width=4/ +| #a #I #G #L2 #V2 #T2 #l #_ #IHT2 #L1 #s #d #e #HL21 #X #H + elim (lift_inv_bind2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct + /4 width=5 by da_bind, ldrop_skip/ +| #I #G #L2 #V2 #T2 #l #_ #IHT2 #L1 #s #d #e #HL21 #X #H + elim (lift_inv_flat2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct + /3 width=5 by da_flat/ ] qed-. 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 05afc6ebd..43de659bf 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lsuba_ldrop.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/lsuba_ldrop.ma @@ -19,45 +19,45 @@ include "basic_2/static/lsuba.ma". (* Properties concerning basic local environment slicing ********************) (* Note: the constant 0 cannot be generalized *) -lemma lsuba_ldrop_O1_conf: ∀G,L1,L2. G ⊢ L1 ⁝⊑ L2 → ∀K1,e. ⇩[0, e] L1 ≡ K1 → - ∃∃K2. G ⊢ K1 ⁝⊑ K2 & ⇩[0, e] L2 ≡ K2. +lemma lsuba_ldrop_O1_conf: ∀G,L1,L2. G ⊢ L1 ⁝⊑ L2 → ∀K1,s,e. ⇩[s, 0, e] L1 ≡ K1 → + ∃∃K2. G ⊢ K1 ⁝⊑ K2 & ⇩[s, 0, e] L2 ≡ K2. #G #L1 #L2 #H elim H -L1 -L2 [ /2 width=3/ -| #I #L1 #L2 #V #_ #IHL12 #K1 #e #H +| #I #L1 #L2 #V #_ #IHL12 #K1 #s #e #H elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK1 [ destruct - 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/ + elim (IHL12 L1 s 0) -IHL12 // #X #HL12 #H + <(ldrop_inv_O2 … H) in HL12; -H /3 width=3 by lsuba_pair, ldrop_pair, ex2_intro/ + | elim (IHL12 … HLK1) -L1 /3 width=3 by ldrop_drop_lt, ex2_intro/ ] -| #L1 #L2 #W #V #A #HV #HW #_ #IHL12 #K1 #e #H +| #L1 #L2 #W #V #A #HV #HW #_ #IHL12 #K1 #s #e #H elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK1 [ destruct - 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/ + elim (IHL12 L1 s 0) -IHL12 // #X #HL12 #H + <(ldrop_inv_O2 … H) in HL12; -H /3 width=3 by lsuba_abbr, ldrop_pair, ex2_intro/ + | elim (IHL12 … HLK1) -L1 /3 width=3 by ldrop_drop_lt, ex2_intro/ ] ] qed-. (* Note: the constant 0 cannot be generalized *) -lemma lsuba_ldrop_O1_trans: ∀G,L1,L2. G ⊢ L1 ⁝⊑ L2 → ∀K2,e. ⇩[0, e] L2 ≡ K2 → - ∃∃K1. G ⊢ K1 ⁝⊑ K2 & ⇩[0, e] L1 ≡ K1. +lemma lsuba_ldrop_O1_trans: ∀G,L1,L2. G ⊢ L1 ⁝⊑ L2 → ∀K2,s,e. ⇩[s, 0, e] L2 ≡ K2 → + ∃∃K1. G ⊢ K1 ⁝⊑ K2 & ⇩[s, 0, e] L1 ≡ K1. #G #L1 #L2 #H elim H -L1 -L2 [ /2 width=3/ -| #I #L1 #L2 #V #_ #IHL12 #K2 #e #H +| #I #L1 #L2 #V #_ #IHL12 #K2 #s #e #H elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK2 [ destruct - 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/ + elim (IHL12 L2 s 0) -IHL12 // #X #HL12 #H + <(ldrop_inv_O2 … H) in HL12; -H /3 width=3 by lsuba_pair, ldrop_pair, ex2_intro/ + | elim (IHL12 … HLK2) -L2 /3 width=3 by ldrop_drop_lt, ex2_intro/ ] -| #L1 #L2 #W #V #A #HV #HW #_ #IHL12 #K2 #e #H +| #L1 #L2 #W #V #A #HV #HW #_ #IHL12 #K2 #s #e #H elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK2 [ destruct - 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/ + elim (IHL12 L2 s 0) -IHL12 // #X #HL12 #H + <(ldrop_inv_O2 … H) in HL12; -H /3 width=3 by lsuba_abbr, ldrop_pair, ex2_intro/ + | elim (IHL12 … HLK2) -L2 /3 width=3 by ldrop_drop_lt, ex2_intro/ ] ] qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lsubd.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lsubd.ma index 0644771eb..32ca7ad3d 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lsubd.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/lsubd.ma @@ -33,7 +33,7 @@ interpretation (* Basic_forward lemmas *****************************************************) lemma lsubd_fwd_lsubr: ∀h,g,G,L1,L2. G ⊢ L1 ▪⊑[h, g] L2 → L1 ⊑ L2. -#h #g #G #L1 #L2 #H elim H -L1 -L2 // /2 width=1/ +#h #g #G #L1 #L2 #H elim H -L1 -L2 /2 width=1 by lsubr_bind, lsubr_abst/ qed-. (* Basic inversion lemmas ***************************************************) @@ -57,8 +57,8 @@ fact lsubd_inv_pair1_aux: ∀h,g,G,L1,L2. G ⊢ L1 ▪⊑[h, g] L2 → I = Abbr & L2 = K2.ⓛW & X = ⓝW.V. #h #g #G #L1 #L2 * -L1 -L2 [ #J #K1 #X #H destruct -| #I #L1 #L2 #V #HL12 #J #K1 #X #H destruct /3 width=3/ -| #L1 #L2 #W #V #l #HV #HW #HL12 #J #K1 #X #H destruct /3 width=9/ +| #I #L1 #L2 #V #HL12 #J #K1 #X #H destruct /3 width=3 by ex2_intro, or_introl/ +| #L1 #L2 #W #V #l #HV #HW #HL12 #J #K1 #X #H destruct /3 width=9 by ex6_4_intro, or_intror/ ] qed-. @@ -87,8 +87,8 @@ fact lsubd_inv_pair2_aux: ∀h,g,G,L1,L2. G ⊢ L1 ▪⊑[h, g] L2 → G ⊢ K1 ▪⊑[h, g] K2 & I = Abst & L1 = K1. ⓓⓝW.V. #h #g #G #L1 #L2 * -L1 -L2 [ #J #K2 #U #H destruct -| #I #L1 #L2 #V #HL12 #J #K2 #U #H destruct /3 width=3/ -| #L1 #L2 #W #V #l #HV #HW #HL12 #J #K2 #U #H destruct /3 width=7/ +| #I #L1 #L2 #V #HL12 #J #K2 #U #H destruct /3 width=3 by ex2_intro, or_introl/ +| #L1 #L2 #W #V #l #HV #HW #HL12 #J #K2 #U #H destruct /3 width=7 by ex5_3_intro, or_intror/ ] qed-. @@ -101,51 +101,51 @@ lemma lsubd_inv_pair2: ∀h,g,I,G,L1,K2,W. G ⊢ L1 ▪⊑[h, g] K2.ⓑ{I}W → (* Basic properties *********************************************************) lemma lsubd_refl: ∀h,g,G,L. G ⊢ L ▪⊑[h, g] L. -#h #g #G #L elim L -L // /2 width=1/ +#h #g #G #L elim L -L /2 width=1 by lsubd_pair/ qed. (* Note: the constant 0 cannot be generalized *) lemma lsubd_ldrop_O1_conf: ∀h,g,G,L1,L2. G ⊢ L1 ▪⊑[h, g] L2 → - ∀K1,e. ⇩[0, e] L1 ≡ K1 → - ∃∃K2. G ⊢ K1 ▪⊑[h, g] K2 & ⇩[0, e] L2 ≡ K2. + ∀K1,s,e. ⇩[s, 0, e] L1 ≡ K1 → + ∃∃K2. G ⊢ K1 ▪⊑[h, g] K2 & ⇩[s, 0, e] L2 ≡ K2. #h #g #G #L1 #L2 #H elim H -L1 -L2 -[ /2 width=3/ -| #I #L1 #L2 #V #_ #IHL12 #K1 #e #H +[ /2 width=3 by ex2_intro/ +| #I #L1 #L2 #V #_ #IHL12 #K1 #s #e #H elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK1 [ destruct - 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/ + elim (IHL12 L1 s 0) -IHL12 // #X #HL12 #H + <(ldrop_inv_O2 … H) in HL12; -H /3 width=3 by lsubd_pair, ldrop_pair, ex2_intro/ + | elim (IHL12 … HLK1) -L1 /3 width=3 by ldrop_drop_lt, ex2_intro/ ] -| #L1 #L2 #W #V #l #HV #HW #_ #IHL12 #K1 #e #H +| #L1 #L2 #W #V #l #HV #HW #_ #IHL12 #K1 #s #e #H elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK1 [ destruct - 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/ + elim (IHL12 L1 s 0) -IHL12 // #X #HL12 #H + <(ldrop_inv_O2 … H) in HL12; -H /3 width=3 by lsubd_abbr, ldrop_pair, ex2_intro/ + | elim (IHL12 … HLK1) -L1 /3 width=3 by ldrop_drop_lt, ex2_intro/ ] ] qed-. (* Note: the constant 0 cannot be generalized *) lemma lsubd_ldrop_O1_trans: ∀h,g,G,L1,L2. G ⊢ L1 ▪⊑[h, g] L2 → - ∀K2,e. ⇩[0, e] L2 ≡ K2 → - ∃∃K1. G ⊢ K1 ▪⊑[h, g] K2 & ⇩[0, e] L1 ≡ K1. + ∀K2,s,e. ⇩[s, 0, e] L2 ≡ K2 → + ∃∃K1. G ⊢ K1 ▪⊑[h, g] K2 & ⇩[s, 0, e] L1 ≡ K1. #h #g #G #L1 #L2 #H elim H -L1 -L2 -[ /2 width=3/ -| #I #L1 #L2 #V #_ #IHL12 #K2 #e #H +[ /2 width=3 by ex2_intro/ +| #I #L1 #L2 #V #_ #IHL12 #K2 #s #e #H elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK2 [ destruct - 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/ + elim (IHL12 L2 s 0) -IHL12 // #X #HL12 #H + <(ldrop_inv_O2 … H) in HL12; -H /3 width=3 by lsubd_pair, ldrop_pair, ex2_intro/ + | elim (IHL12 … HLK2) -L2 /3 width=3 by ldrop_drop_lt, ex2_intro/ ] -| #L1 #L2 #W #V #l #HV #HW #_ #IHL12 #K2 #e #H +| #L1 #L2 #W #V #l #HV #HW #_ #IHL12 #K2 #s #e #H elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK2 [ destruct - 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/ + elim (IHL12 L2 s 0) -IHL12 // #X #HL12 #H + <(ldrop_inv_O2 … H) in HL12; -H /3 width=3 by lsubd_abbr, ldrop_pair, ex2_intro/ + | elim (IHL12 … HLK2) -L2 /3 width=3 by ldrop_drop_lt, ex2_intro/ ] ] qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/ssta.ma b/matita/matita/contribs/lambdadelta/basic_2/static/ssta.ma index eede2acc8..f09d03b86 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/ssta.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/ssta.ma @@ -20,9 +20,9 @@ include "basic_2/static/da.ma". (* activate genv *) inductive ssta (h) (g): relation4 genv lenv term term ≝ | ssta_sort: ∀G,L,k. ssta h g G L (⋆k) (⋆(next h k)) -| ssta_ldef: ∀G,L,K,V,U,W,i. ⇩[0, i] L ≡ K.ⓓV → ssta h g G K V W → +| ssta_ldef: ∀G,L,K,V,U,W,i. ⇩[i] L ≡ K.ⓓV → ssta h g G K V W → ⇧[0, i + 1] W ≡ U → ssta h g G L (#i) U -| ssta_ldec: ∀G,L,K,W,U,l,i. ⇩[0, i] L ≡ K.ⓛW → ⦃G, K⦄ ⊢ W ▪[h, g] l → +| ssta_ldec: ∀G,L,K,W,U,l,i. ⇩[i] L ≡ K.ⓛW → ⦃G, K⦄ ⊢ W ▪[h, g] l → ⇧[0, i + 1] W ≡ U → ssta h g G L (#i) U | ssta_bind: ∀a,I,G,L,V,T,U. ssta h g G (L.ⓑ{I}V) T U → ssta h g G L (ⓑ{a,I}V.T) (ⓑ{a,I}V.U) @@ -51,16 +51,16 @@ lemma ssta_inv_sort1: ∀h,g,G,L,U,k. ⦃G, L⦄ ⊢ ⋆k •[h, g] U → U = /2 width=6 by ssta_inv_sort1_aux/ qed-. fact ssta_inv_lref1_aux: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T •[h, g] U → ∀j. T = #j → - (∃∃K,V,W. ⇩[0, j] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V •[h, g] W & + (∃∃K,V,W. ⇩[j] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V •[h, g] W & ⇧[0, j + 1] W ≡ U ) ∨ - (∃∃K,W,l. ⇩[0, j] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W ▪[h, g] l & + (∃∃K,W,l. ⇩[j] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W ▪[h, g] l & ⇧[0, j + 1] W ≡ U ). #h #g #G #L #T #U * -G -L -T -U [ #G #L #k #j #H destruct -| #G #L #K #V #U #W #i #HLK #HVW #HWU #j #H destruct /3 width=6/ -| #G #L #K #W #U #l #i #HLK #HWl #HWU #j #H destruct /3 width=6/ +| #G #L #K #V #U #W #i #HLK #HVW #HWU #j #H destruct /3 width=6 by ex3_3_intro, or_introl/ +| #G #L #K #W #U #l #i #HLK #HWl #HWU #j #H destruct /3 width=6 by ex3_3_intro, or_intror/ | #a #I #G #L #V #T #U #_ #j #H destruct | #G #L #V #T #U #_ #j #H destruct | #G #L #W #T #U #_ #j #H destruct @@ -68,10 +68,10 @@ fact ssta_inv_lref1_aux: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T •[h, g] U → ∀j. qed-. lemma ssta_inv_lref1: ∀h,g,G,L,U,i. ⦃G, L⦄ ⊢ #i •[h, g] U → - (∃∃K,V,W. ⇩[0, i] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V •[h, g] W & + (∃∃K,V,W. ⇩[i] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V •[h, g] W & ⇧[0, i + 1] W ≡ U ) ∨ - (∃∃K,W,l. ⇩[0, i] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W ▪[h, g] l & + (∃∃K,W,l. ⇩[i] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W ▪[h, g] l & ⇧[0, i + 1] W ≡ U ). /2 width=3 by ssta_inv_lref1_aux/ qed-. @@ -97,7 +97,7 @@ fact ssta_inv_bind1_aux: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T •[h, g] U → [ #G #L #k #b #J #X #Y #H destruct | #G #L #K #V #U #W #i #_ #_ #_ #b #J #X #Y #H destruct | #G #L #K #W #U #l #i #_ #_ #_ #b #J #X #Y #H destruct -| #a #I #G #L #V #T #U #HTU #b #J #X #Y #H destruct /2 width=3/ +| #a #I #G #L #V #T #U #HTU #b #J #X #Y #H destruct /2 width=3 by ex2_intro/ | #G #L #V #T #U #_ #b #J #X #Y #H destruct | #G #L #W #T #U #_ #b #J #X #Y #H destruct ] @@ -114,7 +114,7 @@ fact ssta_inv_appl1_aux: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T •[h, g] U → ∀X,Y | #G #L #K #V #U #W #i #_ #_ #_ #X #Y #H destruct | #G #L #K #W #U #l #i #_ #_ #_ #X #Y #H destruct | #a #I #G #L #V #T #U #_ #X #Y #H destruct -| #G #L #V #T #U #HTU #X #Y #H destruct /2 width=3/ +| #G #L #V #T #U #HTU #X #Y #H destruct /2 width=3 by ex2_intro/ | #G #L #W #T #U #_ #X #Y #H destruct ] qed-. @@ -143,12 +143,12 @@ lemma ssta_inv_cast1: ∀h,g,G,L,X,Y,U. ⦃G, L⦄ ⊢ ⓝY.X •[h, g] U → lemma ssta_inv_da: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T •[h, g] U → ∃l. ⦃G, L⦄ ⊢ T ▪[h, g] l. #h #g #G #L #T #U #H elim H -G -L -T -U -[ #G #L #k elim (deg_total h g k) /3 width=2/ -| #G #L #K #V #U #W #i #HLK #_ #_ * /3 width=5/ -| #G #L #K #W #U #l #i #HLK #HWl #_ /3 width=5/ -| #a #I #G #L #V #T #U #_ * /3 width=2/ -| #G #L #V #T #U #_ * /3 width=2/ -| #G #L #W #T #U #_ * /3 width=2/ +[ #G #L #k elim (deg_total h g k) /3 width=2 by da_sort, ex_intro/ +| #G #L #K #V #U #W #i #HLK #_ #_ * /3 width=5 by da_ldef, ex_intro/ +| #G #L #K #W #U #l #i #HLK #HWl #_ /3 width=5 by da_ldec, ex_intro/ +| #a #I #G #L #V #T #U #_ * /3 width=2 by da_bind, ex_intro/ +| #G #L #V #T #U #_ * /3 width=2 by da_flat, ex_intro/ +| #G #L #W #T #U #_ * /3 width=2 by da_flat, ex_intro/ ] qed-. @@ -159,11 +159,11 @@ lemma da_ssta: ∀h,g,G,L,T,l. ⦃G, L⦄ ⊢ T ▪[h, g] l → #h #g #G #L #T #l #H elim H -G -L -T -l [ /2 width=2/ | #G #L #K #V #i #l #HLK #_ * #W #HVW - elim (lift_total W 0 (i+1)) /3 width=7/ + elim (lift_total W 0 (i+1)) /3 width=7 by ssta_ldef, ex_intro/ | #G #L #K #W #i #l #HLK #HW #_ - elim (lift_total W 0 (i+1)) /3 width=7/ -| #a #I #G #L #V #T #l #_ * /3 width=2/ -| * #G #L #V #T #l #_ * /3 width=2/ + elim (lift_total W 0 (i+1)) /3 width=7 by ssta_ldec, ex_intro/ +| #a #I #G #L #V #T #l #_ * /3 width=2 by ssta_bind, ex_intro/ +| * #G #L #V #T #l #_ * /3 width=2 by ssta_appl, ssta_cast, ex_intro/ ] qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/ssta_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/static/ssta_lift.ma index ff975d25a..843bc3a1e 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/ssta_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/ssta_lift.ma @@ -21,38 +21,40 @@ include "basic_2/static/ssta.ma". lemma ssta_lift: ∀h,g,G. l_liftable (ssta h g G). #h #g #G #L1 #T1 #U1 #H elim H -G -L1 -T1 -U1 -[ #G #L1 #k #L2 #d #e #HL21 #X1 #H1 #X2 #H2 +[ #G #L1 #k #L2 #s #d #e #HL21 #X1 #H1 #X2 #H2 >(lift_inv_sort1 … H1) -X1 >(lift_inv_sort1 … H2) -X2 // -| #G #L1 #K1 #V1 #U1 #W1 #i #HLK1 #_ #HWU1 #IHVW1 #L2 #d #e #HL21 #X #H #U2 #HU12 +| #G #L1 #K1 #V1 #U1 #W1 #i #HLK1 #_ #HWU1 #IHVW1 #L2 #s #d #e #HL21 #X #H #U2 #HU12 elim (lift_inv_lref1 … H) * #Hid #H destruct [ elim (lift_trans_ge … HWU1 … HU12) -U1 // #W2 #HW12 #HWU2 - elim (ldrop_trans_le … HL21 … HLK1) -L1 /2 width=2/ #X #HLK2 #H - elim (ldrop_inv_skip2 … H) -H /2 width=1/ -Hid #K2 #V2 #HK21 #HV12 #H destruct - /3 width=8/ - | lapply (lift_trans_be … HWU1 … HU12 ? ?) -U1 // /2 width=1/ #HW1U2 - lapply (ldrop_trans_ge … HL21 … HLK1 ?) -L1 // -Hid /3 width=8/ + elim (ldrop_trans_le … HL21 … HLK1) -L1 /2 width=2 by lt_to_le/ #X #HLK2 #H + elim (ldrop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid #K2 #V2 #HK21 #HV12 #H destruct + /3 width=9 by ssta_ldef/ + | lapply (lift_trans_be … HWU1 … HU12 ? ?) -U1 /2 width=1 by le_S/ #HW1U2 + lapply (ldrop_trans_ge … HL21 … HLK1 ?) -L1 // -Hid + /3 width=9 by ssta_ldef, ldrop_inv_gen/ ] -| #G #L1 #K1 #W1 #U1 #l #i #HLK1 #HW1l #HWU1 #L2 #d #e #HL21 #X #H #U2 #HU12 +| #G #L1 #K1 #W1 #U1 #l #i #HLK1 #HW1l #HWU1 #L2 #s #d #e #HL21 #X #H #U2 #HU12 elim (lift_inv_lref1 … H) * #Hid #H destruct [ elim (lift_trans_ge … HWU1 … HU12) -U1 // (lift_inv_sort2 … H) -X /2 width=3/ -| #G #L2 #K2 #V2 #U2 #W2 #i #HLK2 #HVW2 #HWU2 #IHVW2 #L1 #d #e #HL21 #X #H +[ #G #L2 #k #L1 #s #d #e #_ #X #H + >(lift_inv_sort2 … H) -X /2 width=3 by ssta_sort, lift_sort, ex2_intro/ +| #G #L2 #K2 #V2 #U2 #W2 #i #HLK2 #HVW2 #HWU2 #IHVW2 #L1 #s #d #e #HL21 #X #H elim (lift_inv_lref2 … H) * #Hid #H destruct [ -HVW2 | -IHVW2 ] [ elim (ldrop_conf_lt … HL21 … HLK2) -L2 // #K1 #V1 #HLK1 #HK21 #HV12 elim (IHVW2 … HK21 … HV12) -K2 -V2 #W1 #HW12 #HVW1 - elim (lift_trans_le … HW12 … HWU2) -W2 // >minus_plus minus_plus minus_minus_m_m /2 width=1/ /3 width=8/ + elim (lift_split … HWU2 d (i-e+1)) -HWU2 // [3: /2 width=1 by le_S/ ] + [ #W0 #HW20 minus_minus_m_m + /3 width=8 by ssta_ldef, le_S, ex2_intro/ | minus_plus minus_plus minus_minus_m_m /2 width=1/ /3 width=8/ + elim (lift_split … HWU2 d (i-e+1)) -HWU2 // [3: /2 width=1 by le_S/ ] + [ #W0 #HW20 minus_minus_m_m + /3 width=8 by ssta_ldec, le_S, ex2_intro/ |