X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fstatic%2Faaa_drops.ma;fp=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fstatic%2Faaa_drops.ma;h=4255a20c4dac165e5fdc3df3d5b61e7ee28f3b74;hb=384b04944ac31922ee41418b106b8e19a19ba9f0;hp=e388af445214426bf245614f35c6b766bfc08d4e;hpb=83cf6c88d2d0bd2c8af86ab7f95bf94c1ae59bc9;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/aaa_drops.ma b/matita/matita/contribs/lambdadelta/basic_2/static/aaa_drops.ma index e388af445..4255a20c4 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/aaa_drops.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/aaa_drops.ma @@ -22,7 +22,7 @@ include "basic_2/static/aaa.ma". (* Advanced properties ******************************************************) (* Basic_2A1: was: aaa_lref *) -lemma aaa_lref_gen: ∀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_gen: ∀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 @@ -53,13 +53,13 @@ lemma aaa_lifts: ∀G,L1,T1,A. ⦃G, L1⦄ ⊢ T1 ⁝ A → ∀b,f,L2. ⬇*[b, f lapply (aaa_inv_sort … H) -H #H destruct >(lifts_inv_sort1 … HX) -HX // | #i1 #HG #HL #HT #A #H #b #f #L2 #HL21 #X #HX - elim (aaa_inv_lref_gen … H) -H #J #K1 #V1 #HLK1 #HA + elim (aaa_inv_lref_drops … H) -H #J #K1 #V1 #HLK1 #HA elim (lifts_inv_lref1 … HX) -HX #i2 #Hf #H destruct lapply (drops_trans … HL21 … HLK1 ??) -HL21 [1,2: // ] #H elim (drops_split_trans … H) -H [ |*: /2 width=6 by after_uni_dx/ ] #Y #HLK2 #HY lapply (drops_inv_tls_at … Hf … HY) -HY #HY -Hf elim (drops_inv_skip2 … HY) -HY #K2 #V2 #HK21 #HV12 #H destruct - /4 width=12 by aaa_lref_gen, fqup_lref, drops_inv_gen/ + /4 width=12 by aaa_lref_drops, fqup_lref, drops_inv_gen/ | #l #HG #HL #HT #A #H #b #f #L2 #HL21 #X #HX -b -f -IH elim (aaa_inv_gref … H) | #p * #V1 #T1 #HG #HL #HT #A #H #b #f #L2 #HL21 #X #HX @@ -91,13 +91,13 @@ lemma aaa_inv_lifts: ∀G,L2,T2,A. ⦃G, L2⦄ ⊢ T2 ⁝ A → ∀b,f,L1. ⬇*[ lapply (aaa_inv_sort … H) -H #H destruct >(lifts_inv_sort2 … HX) -HX // | #i2 #HG #HL #HT #A #H #b #f #L1 #HL21 #X #HX - elim (aaa_inv_lref_gen … H) -H #J #K2 #V2 #HLK2 #HA + 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_conf … HLK2 … HY ??) -HY [1,2: /2 width=6 by after_uni_dx/ ] #HY lapply (drops_inv_tls_at … Hf … HY) -HY #HY -Hf elim (drops_inv_skip1 … HY) -HY #K1 #V1 #HK21 #HV12 #H destruct - /4 width=12 by aaa_lref_gen, fqup_lref, drops_inv_F/ + /4 width=12 by aaa_lref_drops, fqup_lref, drops_inv_F/ | #l #HG #HL #HT #A #H #b #f #L1 #HL21 #X #HX -IH -b -f elim (aaa_inv_gref … H) | #p * #V2 #T2 #HG #HL #HT #A #H #b #f #L1 #HL21 #X #HX