X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fstatic_2%2Fstatic%2Faaa_drops.ma;h=866a0a528bf62fc0260430b4308dec0eab28ed31;hp=071e044752359299b10c14a54f807487411d0f2c;hb=bd53c4e895203eb049e75434f638f26b5a161a2b;hpb=3b7b8afcb429a60d716d5226a5b6ab0d003228b1 diff --git a/matita/matita/contribs/lambdadelta/static_2/static/aaa_drops.ma b/matita/matita/contribs/lambdadelta/static_2/static/aaa_drops.ma index 071e04475..866a0a528 100644 --- a/matita/matita/contribs/lambdadelta/static_2/static/aaa_drops.ma +++ b/matita/matita/contribs/lambdadelta/static_2/static/aaa_drops.ma @@ -22,7 +22,7 @@ include "static_2/static/aaa.ma". (* Advanced properties ******************************************************) (* Basic_2A1: was: aaa_lref *) -lemma aaa_lref_drops: ∀I,G,K,V,B,i,L. ⇩*[i] L ≘ K.ⓑ{I}V → ⦃G,K⦄ ⊢ V ⁝ B → ⦃G,L⦄ ⊢ #i ⁝ B. +lemma aaa_lref_drops: ∀I,G,K,V,B,i,L. ⇩*[i] L ≘ K.ⓑ[I]V → ❪G,K❫ ⊢ V ⁝ B → ❪G,L❫ ⊢ #i ⁝ B. #I #G #K #V #B #i elim i -i [ #L #H lapply (drops_fwd_isid … H ?) -H // #H destruct /2 width=1 by aaa_zero/ @@ -34,8 +34,8 @@ qed. (* Advanced inversion lemmas ************************************************) (* Basic_2A1: was: aaa_inv_lref *) -lemma aaa_inv_lref_drops: ∀G,A,i,L. ⦃G,L⦄ ⊢ #i ⁝ A → - ∃∃I,K,V. ⇩*[i] L ≘ K.ⓑ{I}V & ⦃G,K⦄ ⊢ V ⁝ A. +lemma aaa_inv_lref_drops: ∀G,A,i,L. ❪G,L❫ ⊢ #i ⁝ A → + ∃∃I,K,V. ⇩*[i] L ≘ K.ⓑ[I]V & ❪G,K❫ ⊢ V ⁝ A. #G #A #i elim i -i [ #L #H elim (aaa_inv_zero … H) -H /3 width=5 by drops_refl, ex2_3_intro/ | #i #IH #L #H elim (aaa_inv_lref … H) -H @@ -44,7 +44,7 @@ lemma aaa_inv_lref_drops: ∀G,A,i,L. ⦃G,L⦄ ⊢ #i ⁝ A → qed-. lemma aaa_pair_inv_lref (G) (L) (i): - ∀A. ⦃G,L⦄ ⊢ #i ⁝ A → ∀I,K,V. ⇩*[i] L ≘ K.ⓑ{I}V → ⦃G,K⦄ ⊢ V ⁝ A. + ∀A. ❪G,L❫ ⊢ #i ⁝ A → ∀I,K,V. ⇩*[i] L ≘ K.ⓑ[I]V → ❪G,K❫ ⊢ V ⁝ A. #G #L #i #A #H #I #K #V #HLK elim (aaa_inv_lref_drops … H) -H #J #Y #X #HLY #HX lapply (drops_mono … HLY … HLK) -L -i #H destruct // @@ -54,8 +54,8 @@ qed-. (* Basic_2A1: includes: aaa_lift *) (* Note: it should use drops_split_trans_pair2 *) -lemma aaa_lifts: ∀G,L1,T1,A. ⦃G,L1⦄ ⊢ T1 ⁝ A → ∀b,f,L2. ⇩*[b,f] L2 ≘ L1 → - ∀T2. ⇧*[f] T1 ≘ T2 → ⦃G,L2⦄ ⊢ T2 ⁝ A. +lemma aaa_lifts: ∀G,L1,T1,A. ❪G,L1❫ ⊢ T1 ⁝ A → ∀b,f,L2. ⇩*[b,f] L2 ≘ L1 → + ∀T2. ⇧*[f] T1 ≘ T2 → ❪G,L2❫ ⊢ T2 ⁝ A. @(fqup_wf_ind_eq (Ⓣ)) #G0 #L0 #T0 #IH #G #L1 * * [ #s #HG #HL #HT #A #H #b #f #L2 #HL21 #X #HX -b -IH lapply (aaa_inv_sort … H) -H #H destruct @@ -93,8 +93,8 @@ qed-. (* Inversion lemmas with generic slicing for local environments *************) (* Basic_2A1: includes: aaa_inv_lift *) -lemma aaa_inv_lifts: ∀G,L2,T2,A. ⦃G,L2⦄ ⊢ T2 ⁝ A → ∀b,f,L1. ⇩*[b,f] L2 ≘ L1 → - ∀T1. ⇧*[f] T1 ≘ T2 → ⦃G,L1⦄ ⊢ T1 ⁝ A. +lemma aaa_inv_lifts: ∀G,L2,T2,A. ❪G,L2❫ ⊢ T2 ⁝ A → ∀b,f,L1. ⇩*[b,f] L2 ≘ L1 → + ∀T1. ⇧*[f] T1 ≘ T2 → ❪G,L1❫ ⊢ T1 ⁝ A. @(fqup_wf_ind_eq (Ⓣ)) #G0 #L0 #T0 #IH #G #L2 * * [ #s #HG #HL #HT #A #H #b #f #L1 #HL21 #X #HX -b -IH lapply (aaa_inv_sort … H) -H #H destruct @@ -102,7 +102,7 @@ lemma aaa_inv_lifts: ∀G,L2,T2,A. ⦃G,L2⦄ ⊢ T2 ⁝ A → ∀b,f,L1. ⇩*[b | #i2 #HG #HL #HT #A #H #b #f #L1 #HL21 #X #HX elim (aaa_inv_lref_drops … H) -H #J #K2 #V2 #HLK2 #HA elim (lifts_inv_lref2 … HX) -HX #i1 #Hf #H destruct - lapply (drops_split_div … HL21 (𝐔❴i1❵) ???) -HL21 [4: * |*: // ] #Y #HLK1 #HY + lapply (drops_split_div … HL21 (𝐔❨i1❩) ???) -HL21 [4: * |*: // ] #Y #HLK1 #HY lapply (drops_conf … HLK2 … HY ??) -HY [1,2: /2 width=6 by after_uni_dx/ ] #HY lapply (drops_tls_at … Hf … HY) -HY #HY -Hf elim (drops_inv_skip1 … HY) -HY #Z #K1 #HK21 #HZ #H destruct