(**************************************************************************) (* ___ *) (* ||M|| *) (* ||A|| A project by Andrea Asperti *) (* ||T|| *) (* ||I|| Developers: *) (* ||T|| The HELM team. *) (* ||A|| http://helm.cs.unibo.it *) (* \ / *) (* \ / This file is distributed under the terms of the *) (* v GNU General Public License Version 2 *) (* *) (**************************************************************************) include "basic_2/substitution/cpys_lift.ma". include "basic_2/substitution/cofrees.ma". (* CONTEXT-SENSITIVE EXCLUSION FROM FREE VARIABLES **************************) (* Advanced inversion lemmas ************************************************) lemma cofrees_inv_lref_be: ∀L,d,i,j. L ⊢ i ~ϵ 𝐅*[d]⦃#j⦄ → d ≤ yinj j → j < i → ∀I,K,W. ⇩[j]L ≡ K.ⓑ{I}W → K ⊢ i-j-1 ~ϵ 𝐅*[yinj 0]⦃W⦄. #L #d #i #j #Hj #Hdj #Hji #I #K #W1 #HLK #W2 #HW12 elim (lift_total W2 0 (j+1)) #X2 #HWX2 elim (Hj X2) /2 width=7 by cpys_subst_Y2/ -I -L -K -W1 -d #Z2 #HZX2 elim (lift_div_le … HWX2 (i-j-1) 1 Z2) -HWX2 /2 width=2 by ex_intro/ >minus_plus minus_plus_plus_l // | #J #W #U #Hn #d #i #H1 #j #H2 #I #K #V #HLK #Hdj #Hji destruct elim (cofrees_inv_flat … H1) -H1 #HW #HU elim (nlift_inv_flat … H2) -H2 [ /3 width=9 by/ ] #HnU @(IH … HU … HnU … HLK) // (**) (* full auto fails *) ] qed-. (* Advanced properties ******************************************************) lemma cofrees_lref_skip: ∀L,d,i,j. j < i → yinj j < d → L ⊢ i ~ϵ 𝐅*[d]⦃#j⦄. #L #d #i #j #Hji #Hjd #X #H elim (cpys_inv_lref1_Y2 … H) -H [ #H destruct /3 width=2 by lift_lref_lt, ex_intro/ | * #I #K #W1 #W2 #Hdj elim (ylt_yle_false … Hdj) -i -I -L -K -W1 -W2 -X // ] qed. lemma cofrees_lref_lt: ∀L,d,i,j. i < j → L ⊢ i ~ϵ 𝐅*[d]⦃#j⦄. #L #d #i #j #Hij #X #H elim (cpys_inv_lref1_Y2 … H) -H [ #H destruct /3 width=2 by lift_lref_ge_minus, ex_intro/ | * #I #K #V1 #V2 #_ #_ #_ #H -I -L -K -V1 -d elim (lift_split … H i j) /2 width=2 by lt_to_le, ex_intro/ ] qed. lemma cofrees_lref_gt: ∀I,L,K,W,d,i,j. j < i → ⇩[j] L ≡ K.ⓑ{I}W → K ⊢ (i-j-1) ~ϵ 𝐅*[O]⦃W⦄ → L ⊢ i ~ϵ 𝐅*[d]⦃#j⦄. #I #L #K #W1 #d #i #j #Hji #HLK #HW1 #X #H elim (cpys_inv_lref1_Y2 … H) -H [ #H destruct /3 width=2 by lift_lref_lt, ex_intro/ | * #I0 #K0 #W0 #W2 #Hdj #HLK0 #HW12 #HW2 lapply (ldrop_mono … HLK0 … HLK) -L #H destruct elim (HW1 … HW12) -I -K -W1 -d #V2 #HVW2 elim (lift_trans_le … HVW2 … HW2) -W2 // >minus_plus minus_plus yplus_inj >yminus_Y_inj #T2 #HT12 lapply (cpys_weak … HT12 (d-yinj e0) (∞) ? ?) /2 width=1 by yle_plus2_to_minus_inj2/ -HT12 | elim (cpys_inv_lift1_ge … HU12 … HLK … HTU1) // #T2 ] | elim (cpys_inv_lift1_be … HU12 … HLK … HTU1) // >yminus_Y_inj #T2 #HT12 lapply (cpys_weak … HT12 (d-yinj e0) (∞) ? ?) // -HT12 ] -s -L #HT12 #HTU2 elim (HT1 … HT12) -T1 #V2 #HVT2 elim (lift_trans_le … HVT2 … HTU2 ?) //