X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fstatic%2Ffrees_drops.ma;h=fe24eadf5d0496b0dce3a9145e2e441f55e4174b;hp=441d3c51f5670ff8298c9f8f35738024078f3102;hb=222044da28742b24584549ba86b1805a87def070;hpb=0c7cb3c503c0fcab104ad89ebc88683dc9830d06 diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/frees_drops.ma b/matita/matita/contribs/lambdadelta/basic_2/static/frees_drops.ma index 441d3c51f..fe24eadf5 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/frees_drops.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/frees_drops.ma @@ -14,127 +14,138 @@ include "ground_2/relocation/nstream_coafter.ma". include "basic_2/relocation/drops_drops.ma". -include "basic_2/static/frees_frees.ma". +include "basic_2/static/frees_fqup.ma". (* CONTEXT-SENSITIVE FREE VARIABLES *****************************************) (* Advanced properties ******************************************************) -lemma frees_lref_atom: ∀b,L,i. ⬇*[b, 𝐔❴i❵] L ≡ ⋆ → - ∀f. 𝐈⦃f⦄ → L ⊢ 𝐅*⦃#i⦄ ≡ f. +lemma frees_atom_drops: ∀b,L,i. ⬇*[b, 𝐔❴i❵] L ≘ ⋆ → + ∀f. 𝐈⦃f⦄ → L ⊢ 𝐅*⦃#i⦄ ≘ ⫯*[i]↑f. #b #L elim L -L /2 width=1 by frees_atom/ -#L #I #V #IH * +#L #I #IH * [ #H lapply (drops_fwd_isid … H ?) -H // #H destruct -| /5 width=3 by frees_eq_repl_back, frees_lref, drops_inv_drop1, eq_push_inv_isid/ +| /4 width=3 by frees_lref, drops_inv_drop1/ ] qed. -lemma frees_lref_pair: ∀f,K,V. K ⊢ 𝐅*⦃V⦄ ≡ f → - ∀i,I,L. ⬇*[i] L ≡ K.ⓑ{I}V → L ⊢ 𝐅*⦃#i⦄ ≡ ↑*[i] ⫯f. +lemma frees_pair_drops: ∀f,K,V. K ⊢ 𝐅*⦃V⦄ ≘ f → + ∀i,I,L. ⬇*[i] L ≘ K.ⓑ{I}V → L ⊢ 𝐅*⦃#i⦄ ≘ ⫯*[i] ↑f. #f #K #V #Hf #i elim i -i -[ #I #L #H lapply (drops_fwd_isid … H ?) -H /2 width=1 by frees_zero/ +[ #I #L #H lapply (drops_fwd_isid … H ?) -H /2 width=1 by frees_pair/ | #i #IH #I #L #H elim (drops_inv_succ … H) -H /3 width=2 by frees_lref/ ] qed. -lemma frees_sort_pushs: ∀f,K,s. K ⊢ 𝐅*⦃⋆s⦄ ≡ f → - ∀i,L. ⬇*[i] L ≡ K → L ⊢ 𝐅*⦃⋆s⦄ ≡ ↑*[i] f. +lemma frees_unit_drops: ∀f. 𝐈⦃f⦄ → ∀I,K,i,L. ⬇*[i] L ≘ K.ⓤ{I} → + L ⊢ 𝐅*⦃#i⦄ ≘ ⫯*[i] ↑f. +#f #Hf #I #K #i elim i -i +[ #L #H lapply (drops_fwd_isid … H ?) -H /2 width=1 by frees_unit/ +| #i #IH #Y #H elim (drops_inv_succ … H) -H + #J #L #HLK #H destruct /3 width=1 by frees_lref/ +] +qed. +(* +lemma frees_sort_pushs: ∀f,K,s. K ⊢ 𝐅*⦃⋆s⦄ ≘ f → + ∀i,L. ⬇*[i] L ≘ K → L ⊢ 𝐅*⦃⋆s⦄ ≘ ⫯*[i] f. #f #K #s #Hf #i elim i -i [ #L #H lapply (drops_fwd_isid … H ?) -H // | #i #IH #L #H elim (drops_inv_succ … H) -H /3 width=1 by frees_sort/ ] qed. - -lemma frees_lref_pushs: ∀f,K,j. K ⊢ 𝐅*⦃#j⦄ ≡ f → - ∀i,L. ⬇*[i] L ≡ K → L ⊢ 𝐅*⦃#(i+j)⦄ ≡ ↑*[i] f. +*) +lemma frees_lref_pushs: ∀f,K,j. K ⊢ 𝐅*⦃#j⦄ ≘ f → + ∀i,L. ⬇*[i] L ≘ K → L ⊢ 𝐅*⦃#(i+j)⦄ ≘ ⫯*[i] f. #f #K #j #Hf #i elim i -i [ #L #H lapply (drops_fwd_isid … H ?) -H // -| #i #IH #L #H elim (drops_inv_succ … H) -H /3 width=1 by frees_lref/ +| #i #IH #L #H elim (drops_inv_succ … H) -H + #I #Y #HYK #H destruct /3 width=1 by frees_lref/ ] qed. - -lemma frees_gref_pushs: ∀f,K,l. K ⊢ 𝐅*⦃§l⦄ ≡ f → - ∀i,L. ⬇*[i] L ≡ K → L ⊢ 𝐅*⦃§l⦄ ≡ ↑*[i] f. +(* +lemma frees_gref_pushs: ∀f,K,l. K ⊢ 𝐅*⦃§l⦄ ≘ f → + ∀i,L. ⬇*[i] L ≘ K → L ⊢ 𝐅*⦃§l⦄ ≘ ⫯*[i] f. #f #K #l #Hf #i elim i -i [ #L #H lapply (drops_fwd_isid … H ?) -H // | #i #IH #L #H elim (drops_inv_succ … H) -H /3 width=1 by frees_gref/ ] qed. - +*) (* Advanced inversion lemmas ************************************************) -lemma frees_inv_lref_drops: ∀i,f,L. L ⊢ 𝐅*⦃#i⦄ ≡ f → - (⬇*[Ⓕ, 𝐔❴i❵] L ≡ ⋆ ∧ 𝐈⦃f⦄) ∨ - ∃∃g,I,K,V. K ⊢ 𝐅*⦃V⦄ ≡ g & - ⬇*[i] L ≡ K.ⓑ{I}V & f = ↑*[i] ⫯g. -#i elim i -i -[ #f #L #H elim (frees_inv_zero … H) -H * - /4 width=7 by ex3_4_intro, or_introl, or_intror, conj, drops_refl/ -| #i #IH #f #L #H elim (frees_inv_lref … H) -H * /3 width=1 by or_introl, conj/ - #g #I #K #V #Hg #H1 #H2 destruct - elim (IH … Hg) -IH -Hg * - [ /4 width=3 by or_introl, conj, isid_push, drops_drop/ - | /4 width=7 by drops_drop, ex3_4_intro, or_intror/ +lemma frees_inv_lref_drops: ∀L,i,f. L ⊢ 𝐅*⦃#i⦄ ≘ f → + ∨∨ ∃∃g. ⬇*[Ⓕ, 𝐔❴i❵] L ≘ ⋆ & 𝐈⦃g⦄ & f = ⫯*[i] ↑g + | ∃∃g,I,K,V. K ⊢ 𝐅*⦃V⦄ ≘ g & + ⬇*[i] L ≘ K.ⓑ{I}V & f = ⫯*[i] ↑g + | ∃∃g,I,K. ⬇*[i] L ≘ K.ⓤ{I} & 𝐈⦃g⦄ & f = ⫯*[i] ↑g. +#L elim L -L +[ #i #g | #L #I #IH * [ #g cases I -I [ #I | #I #V ] -IH | #i #g ] ] #H +[ elim (frees_inv_atom … H) -H #f #Hf #H destruct + /3 width=3 by or3_intro0, ex3_intro/ +| elim (frees_inv_unit … H) -H #f #Hf #H destruct + /4 width=3 by drops_refl, or3_intro2, ex3_3_intro/ +| elim (frees_inv_pair … H) -H #f #Hf #H destruct + /4 width=7 by drops_refl, or3_intro1, ex3_4_intro/ +| elim (frees_inv_lref … H) -H #f #Hf #H destruct + elim (IH … Hf) -IH -Hf * + [ /4 width=3 by drops_drop, or3_intro0, ex3_intro/ + | /4 width=7 by drops_drop, or3_intro1, ex3_4_intro/ + | /4 width=3 by drops_drop, or3_intro2, ex3_3_intro/ ] ] qed-. (* Properties with generic slicing for local environments *******************) -lemma frees_lifts: ∀b,f1,K,T. K ⊢ 𝐅*⦃T⦄ ≡ f1 → - ∀f,L. ⬇*[b, f] L ≡ K → ∀U. ⬆*[f] T ≡ U → - ∀f2. f ~⊚ f1 ≡ f2 → L ⊢ 𝐅*⦃U⦄ ≡ f2. +lemma frees_lifts: ∀b,f1,K,T. K ⊢ 𝐅*⦃T⦄ ≘ f1 → + ∀f,L. ⬇*[b, f] L ≘ K → ∀U. ⬆*[f] T ≘ U → + ∀f2. f ~⊚ f1 ≘ f2 → L ⊢ 𝐅*⦃U⦄ ≘ f2. #b #f1 #K #T #H lapply (frees_fwd_isfin … H) elim H -f1 -K -T -[ #f1 #I #Hf1 #_ #f #L #H1 #U #H2 #f2 #H3 +[ #f1 #K #s #Hf1 #_ #f #L #HLK #U #H2 #f2 #H3 lapply (coafter_isid_inv_dx … H3 … Hf1) -f1 #Hf2 - elim (lifts_inv_atom1 … H2) -H2 * - /2 width=1 by frees_sort_gen, frees_gref_gen/ - #i #j #Hij #H #H0 destruct + >(lifts_inv_sort1 … H2) -U /2 width=1 by frees_sort/ +| #f1 #i #Hf1 #_ #f #L #H1 #U #H2 #f2 #H3 + elim (lifts_inv_lref1 … H2) -H2 #j #Hij #H destruct + elim (coafter_fwd_xnx_pushs … Hij H3) -H3 #g2 #Hg2 #H2 destruct + lapply (coafter_isid_inv_dx … Hg2 … Hf1) -f1 #Hf2 elim (drops_inv_atom2 … H1) -H1 #n #g #H1 #Hf elim (after_at_fwd … Hij … Hf) -f #x #_ #Hj -g -i lapply (at_inv_uni … Hj) -Hj #H destruct - /3 width=8 by frees_lref_atom, drops_trans/ -| #f1 #I #K #V #s #_ #IH #Hf1 #f #L #H1 #U #H2 #f2 #H3 - lapply (isfin_fwd_push … Hf1 ??) -Hf1 [3: |*: // ] #Hf1 - lapply (lifts_inv_sort1 … H2) -H2 #H destruct - elim (drops_split_trans_pair2 … H1) -H1 [ |*: // ] #Y #W #HLY #HYK #_ - elim (coafter_fwd_xpx_pushs … H3) [ |*: // ] #g2 #H2 destruct - lapply (coafter_tls_succ … H3 ??) -H3 [3: |*: // ] #H3 - lapply (IH … HYK … H3) -IH -H3 -HYK [1,3: // | skip ] - /3 width=5 by drops_isuni_fwd_drop2, frees_sort_pushs/ + /3 width=8 by frees_atom_drops, drops_trans/ | #f1 #I #K #V #_ #IH #Hf1 #f #L #H1 #U #H2 #f2 #H3 lapply (isfin_inv_next … Hf1 ??) -Hf1 [3: |*: // ] #Hf1 lapply (lifts_inv_lref1 … H2) -H2 * #j #Hf #H destruct - elim (drops_split_trans_pair2 … H1) -H1 [ |*: // ] #Y #W #HLY #HYK #HVW - elim (coafter_fwd_xnx_pushs … H3) [ |*: // ] #g2 #H2 destruct - lapply (coafter_tls_succ … H3 ??) -H3 [3: |*: // ] - plus_S1 /2 width=3 by frees_lref_pushs/ (**) (* full auto fails *) -| #f1 #I #K #V #l #_ #IH #Hf1 #f #L #H1 #U #H2 #f2 #H3 - lapply (isfin_fwd_push … Hf1 ??) -Hf1 [3: |*: // ] #Hf1 - lapply (lifts_inv_gref1 … H2) -H2 #H destruct - elim (drops_split_trans_pair2 … H1) -H1 [ |*: // ] #Y #W #HLY #HYK #_ - elim (coafter_fwd_xpx_pushs … H3) [ |*: // ] #g2 #H2 destruct - lapply (coafter_tls_succ … H3 ??) -H3 [3: |*: // ] #H3 - lapply (IH … HYK … H3) -IH -H3 -HYK [1,3: // | skip ] - /3 width=5 by drops_isuni_fwd_drop2, frees_gref_pushs/ +| #f1 #K #l #Hf1 #_ #f #L #HLK #U #H2 #f2 #H3 + lapply (coafter_isid_inv_dx … H3 … Hf1) -f1 #Hf2 + >(lifts_inv_gref1 … H2) -U /2 width=1 by frees_gref/ | #f1V #f1T #f1 #p #I #K #V #T #_ #_ #H1f1 #IHV #IHT #H2f1 #f #L #H1 #Y #H2 #f2 #H3 elim (sor_inv_isfin3 … H1f1) // #Hf1V #H lapply (isfin_inv_tl … H) -H elim (lifts_inv_bind1 … H2) -H2 #W #U #HVW #HTU #H destruct elim (coafter_sor … H3 … H1f1) /2 width=5 by coafter_isfin2_fwd/ -H3 -H1f1 #f2V #f2T #Hf2V #H - elim (coafter_inv_tl1 … H) -H /4 width=5 by frees_bind, drops_skip/ + elim (coafter_inv_tl1 … H) -H + /5 width=5 by frees_bind, drops_skip, ext2_pair/ | #f1V #f1T #f1 #I #K #V #T #_ #_ #H1f1 #IHV #IHT #H2f1 #f #L #H1 #Y #H2 #f2 #H3 elim (sor_inv_isfin3 … H1f1) // elim (lifts_inv_flat1 … H2) -H2 #W #U #HVW #HTU #H destruct @@ -143,133 +154,92 @@ lemma frees_lifts: ∀b,f1,K,T. K ⊢ 𝐅*⦃T⦄ ≡ f1 → ] qed-. +lemma frees_lifts_SO: ∀b,L,K. ⬇*[b, 𝐔❴1❵] L ≘ K → ∀T,U. ⬆*[1] T ≘ U → + ∀f. K ⊢ 𝐅*⦃T⦄ ≘ f → L ⊢ 𝐅*⦃U⦄ ≘ ⫯f. +#b #L #K #HLK #T #U #HTU #f #Hf +@(frees_lifts b … Hf … HTU) // (**) (* auto fails *) +qed. + (* Forward lemmas with generic slicing for local environments ***************) -lemma frees_fwd_coafter: ∀b,f2,L,U. L ⊢ 𝐅*⦃U⦄ ≡ f2 → - ∀f,K. ⬇*[b, f] L ≡ K → ∀T. ⬆*[f] T ≡ U → - ∀f1. K ⊢ 𝐅*⦃T⦄ ≡ f1 → f ~⊚ f1 ≡ f2. +lemma frees_fwd_coafter: ∀b,f2,L,U. L ⊢ 𝐅*⦃U⦄ ≘ f2 → + ∀f,K. ⬇*[b, f] L ≘ K → ∀T. ⬆*[f] T ≘ U → + ∀f1. K ⊢ 𝐅*⦃T⦄ ≘ f1 → f ~⊚ f1 ≘ f2. /4 width=11 by frees_lifts, frees_mono, coafter_eq_repl_back0/ qed-. (* Inversion lemmas with generic slicing for local environments *************) -lemma frees_inv_lifts: ∀b,f2,L,U. L ⊢ 𝐅*⦃U⦄ ≡ f2 → - ∀f,K. ⬇*[b, f] L ≡ K → ∀T. ⬆*[f] T ≡ U → - ∀f1. f ~⊚ f1 ≡ f2 → K ⊢ 𝐅*⦃T⦄ ≡ f1. -#b #f2 #L #U #H lapply (frees_fwd_isfin … H) elim H -f2 -L -U -[ #f2 #I #Hf2 #_ #f #K #H1 #T #H2 #f1 #H3 - lapply (coafter_fwd_isid2 … H3 … Hf2) -H3 // -Hf2 #Hf1 - elim (drops_inv_atom1 … H1) -H1 #H #_ destruct - elim (lifts_inv_atom2 … H2) -H2 * /2 width=3 by frees_atom/ -| #f2 #I #L #W #s #_ #IH #Hf2 #f #Y #H1 #T #H2 #f1 #H3 - lapply (isfin_fwd_push … Hf2 ??) -Hf2 [3: |*: // ] #Hf2 - lapply (lifts_inv_sort2 … H2) -H2 #H destruct - elim (coafter_inv_xxp … H3) -H3 [1,3: * |*: // ] - [ #g #g1 #Hf2 #H #H0 destruct - elim (drops_inv_skip1 … H1) -H1 #K #V #HLK #_ #H destruct - | #g #Hf2 #H destruct - lapply (drops_inv_drop1 … H1) -H1 - ] /3 width=4 by frees_sort/ -| #f2 #I #L #W #_ #IH #Hf2 #f #Y #H1 #T #H2 #f1 #H3 - lapply (isfin_inv_next … Hf2 ??) -Hf2 [3: |*: // ] #Hf2 - elim (lifts_inv_lref2 … H2) -H2 #i #H2 #H destruct - lapply (at_inv_xxp … H2 ?) -H2 // * #g #H #H0 destruct - elim (drops_inv_skip1 … H1) -H1 #K #V #HLK #HVW #H destruct - elim (coafter_inv_pxn … H3) -H3 [ |*: // ] #g1 #Hf2 #H destruct - /3 width=4 by frees_zero/ -| #f2 #I #L #W #j #_ #IH #Hf2 #f #Y #H1 #T #H2 #f1 #H3 - lapply (isfin_fwd_push … Hf2 ??) -Hf2 [3: |*: // ] #Hf2 - elim (lifts_inv_lref2 … H2) -H2 #x #H2 #H destruct - elim (coafter_inv_xxp … H3) -H3 [1,3: * |*: // ] - [ #g #g1 #Hf2 #H #H0 destruct - elim (drops_inv_skip1 … H1) -H1 #K #V #HLK #HVW #H destruct - elim (at_inv_xpn … H2) -H2 [ |*: // ] #j #Hg #H destruct - | #g #Hf2 #H destruct - lapply (drops_inv_drop1 … H1) -H1 - lapply (at_inv_xnn … H2 ????) -H2 [5: |*: // ] - ] /4 width=4 by lifts_lref, frees_lref/ -| #f2 #I #L #W #l #_ #IH #Hf2 #f #Y #H1 #T #H2 #f1 #H3 - lapply (isfin_fwd_push … Hf2 ??) -Hf2 [3: |*: // ] #Hf2 - lapply (lifts_inv_gref2 … H2) -H2 #H destruct - elim (coafter_inv_xxp … H3) -H3 [1,3: * |*: // ] - [ #g #g1 #Hf2 #H #H0 destruct - elim (drops_inv_skip1 … H1) -H1 #K #V #HLK #_ #H destruct - | #g #Hf2 #H destruct - lapply (drops_inv_drop1 … H1) -H1 - ] /3 width=4 by frees_gref/ -| #f2W #f2U #f2 #p #I #L #W #U #_ #_ #H1f2 #IHW #IHU #H2f2 #f #K #H1 #X #H2 #f1 #H3 - elim (sor_inv_isfin3 … H1f2) // #H1f2W #H - lapply (isfin_inv_tl … H) -H - elim (lifts_inv_bind2 … H2) -H2 #V #T #HVW #HTU #H destruct - elim (coafter_inv_sor … H3 … H1f2) -H3 -H1f2 // #f1W #f1U #H2f2W #H - elim (coafter_inv_tl0 … H) -H /4 width=5 by frees_bind, drops_skip/ -| #f2W #f2U #f2 #I #L #W #U #_ #_ #H1f2 #IHW #IHU #H2f2 #f #K #H1 #X #H2 #f1 #H3 - elim (sor_inv_isfin3 … H1f2) // - elim (lifts_inv_flat2 … H2) -H2 #V #T #HVW #HTU #H destruct - elim (coafter_inv_sor … H3 … H1f2) -H3 -H1f2 /3 width=5 by frees_flat/ -] +lemma frees_inv_lifts_ex: ∀b,f2,L,U. L ⊢ 𝐅*⦃U⦄ ≘ f2 → + ∀f,K. ⬇*[b, f] L ≘ K → ∀T. ⬆*[f] T ≘ U → + ∃∃f1. f ~⊚ f1 ≘ f2 & K ⊢ 𝐅*⦃T⦄ ≘ f1. +#b #f2 #L #U #Hf2 #f #K #HLK #T elim (frees_total K T) +/3 width=9 by frees_fwd_coafter, ex2_intro/ +qed-. + +lemma frees_inv_lifts_SO: ∀b,f,L,U. L ⊢ 𝐅*⦃U⦄ ≘ f → + ∀K. ⬇*[b, 𝐔❴1❵] L ≘ K → ∀T. ⬆*[1] T ≘ U → + K ⊢ 𝐅*⦃T⦄ ≘ ⫱f. +#b #f #L #U #H #K #HLK #T #HTU elim(frees_inv_lifts_ex … H … HLK … HTU) -b -L -U +#f1 #Hf #Hf1 elim (coafter_inv_nxx … Hf) -Hf +/3 width=5 by frees_eq_repl_back, coafter_isid_inv_sn/ qed-. -lemma frees_inv_drops: ∀f2,L,U. L ⊢ 𝐅*⦃U⦄ ≡ f2 → - ∀f,K. ⬇*[Ⓣ, f] L ≡ K → ∀f1. f ~⊚ f1 ≡ f2 → - ∃∃T. K ⊢ 𝐅*⦃T⦄ ≡ f1 & ⬆*[f] T ≡ U. -#f2 #L #U #H lapply (frees_fwd_isfin … H) elim H -f2 -L -U -[ #f2 #I #Hf2 #_ #f #K #H1 #f1 #H2 - lapply (coafter_fwd_isid2 … H2 ??) -H2 // -Hf2 #Hf1 - elim (drops_inv_atom1 … H1) -H1 #H #Hf destruct - /4 width=3 by frees_atom, lifts_refl, ex2_intro/ -| #f2 #I #L #W #s #_ #IH #Hf2 #f #Y #H1 #f1 #H2 - lapply (isfin_fwd_push … Hf2 ??) -Hf2 [3: |*: // ] #Hf2 - elim (coafter_inv_xxp … H2) -H2 [1,3: * |*: // ] - [ #g #g1 #Hf2 #H #H0 destruct - elim (drops_inv_skip1 … H1) -H1 #K #V #HLK #_ #H destruct - | #g #Hf2 #H destruct - lapply (drops_inv_drop1 … H1) -H1 #HLK +lemma frees_inv_lifts: ∀b,f2,L,U. L ⊢ 𝐅*⦃U⦄ ≘ f2 → + ∀f,K. ⬇*[b, f] L ≘ K → ∀T. ⬆*[f] T ≘ U → + ∀f1. f ~⊚ f1 ≘ f2 → K ⊢ 𝐅*⦃T⦄ ≘ f1. +#b #f2 #L #U #H #f #K #HLK #T #HTU #f1 #Hf2 elim (frees_inv_lifts_ex … H … HLK … HTU) -b -L -U +/3 width=7 by frees_eq_repl_back, coafter_inj/ +qed-. + +(* Note: this is used by rex_conf and might be modified *) +lemma frees_inv_drops_next: ∀f1,L1,T1. L1 ⊢ 𝐅*⦃T1⦄ ≘ f1 → + ∀I2,L2,V2,n. ⬇*[n] L1 ≘ L2.ⓑ{I2}V2 → + ∀g1. ↑g1 = ⫱*[n] f1 → + ∃∃g2. L2 ⊢ 𝐅*⦃V2⦄ ≘ g2 & g2 ⊆ g1. +#f1 #L1 #T1 #H elim H -f1 -L1 -T1 +[ #f1 #L1 #s #Hf1 #I2 #L2 #V2 #n #_ #g1 #H1 -I2 -L1 -s + lapply (isid_tls n … Hf1) -Hf1

(injective_next … Hgf1) -g1 + /2 width=3 by ex2_intro/ + | -Hf1 #n #HL12 lapply (drops_inv_drop1 … HL12) -HL12 + #HL12 #g1 tls_xn #H2 elim (IHT1 … H2) -IHT1 -H2 + /3 width=6 by drops_drop, sor_inv_sle_dx_trans, ex2_intro/ ] - elim (IH … HLK … Hf2) -L // -f2 #X #Hg1 #HX - elim (lifts_inv_lref2 … HX) -HX #i #Hij #H destruct - /4 width=7 by frees_lref, lifts_lref, at_S1, at_next, ex2_intro/ -| #f2 #I #L #W #l #_ #IH #Hf2 #f #Y #H1 #f1 #H2 - lapply (isfin_fwd_push … Hf2 ??) -Hf2 [3: |*: // ] #Hf2 - elim (coafter_inv_xxp … H2) -H2 [1,3: * |*: // ] - [ #g #g1 #Hf2 #H #H0 destruct - elim (drops_inv_skip1 … H1) -H1 #K #V #HLK #_ #H destruct - | #g #Hf2 #H destruct - lapply (drops_inv_drop1 … H1) -H1 #HLK +| #fV1 #fT1 #f1 #I1 #L1 #V1 #T1 #_ #_ #Hf1 #IHV1 #IHT1 #I2 #L2 #V2 #n #HL12 #g1 #Hgf1 + lapply (sor_tls … Hf1 n) -Hf1