(**************************************************************************) (* ___ *) (* ||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_2A/substitution/lift_neg.ma". include "basic_2A/substitution/lift_lift.ma". include "basic_2A/substitution/cpy.ma". (* CONTEXT-SENSITIVE EXTENDED ORDINARY SUBSTITUTION FOR TERMS ***************) (* Inversion lemmas on negated relocation ***********************************) lemma cpy_fwd_nlift2_ge: ∀G,L,U1,U2,l,m. ⦃G, L⦄ ⊢ U1 ▶[l, m] U2 → ∀i. l ≤ yinj i → (∀T2. ⬆[i, 1] T2 ≡ U2 → ⊥) → (∀T1. ⬆[i, 1] T1 ≡ U1 → ⊥) ∨ ∃∃I,K,W,j. l ≤ yinj j & j < i & ⬇[j]L ≡ K.ⓑ{I}W & (∀V. ⬆[i-j-1, 1] V ≡ W → ⊥) & (∀T1. ⬆[j, 1] T1 ≡ U1 → ⊥). #G #L #U1 #U2 #l #m #H elim H -G -L -U1 -U2 -l -m [ /3 width=2 by or_introl/ | #I #G #L #K #V #W #j #l #m #Hlj #Hjlm #HLK #HVW #i #Hli #HnW elim (lt_or_ge j i) #Hij [ @or_intror @(ex5_4_intro … HLK) // -HLK [ #X #HXV elim (lift_trans_le … HXV … HVW ?) -V // #Y #HXY >minus_plus (plus_minus_m_m j 1) in ⊢ (%→?); [2: /3 width=3 by yle_trans, yle_inv_inj/ ] #HnU1