]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2A/substitution/cpy_nlift.ma
update in lambdadelta
[helm.git] / matita / matita / contribs / lambdadelta / basic_2A / substitution / cpy_nlift.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 include "basic_2A/substitution/lift_neg.ma".
16 include "basic_2A/substitution/lift_lift.ma".
17 include "basic_2A/substitution/cpy.ma".
18
19 (* CONTEXT-SENSITIVE EXTENDED ORDINARY SUBSTITUTION FOR TERMS ***************)
20
21 (* Inversion lemmas on negated relocation ***********************************)
22
23 lemma cpy_fwd_nlift2_ge: ∀G,L,U1,U2,l,m. ⦃G, L⦄ ⊢ U1 ▶[l, m] U2 →
24                          ∀i. l ≤ yinj i → (∀T2. ⬆[i, 1] T2 ≡ U2 → ⊥) →
25                          (∀T1. ⬆[i, 1] T1 ≡ U1 → ⊥) ∨
26                          ∃∃I,K,W,j. l ≤ yinj j & j < i & ⬇[j]L ≡ K.ⓑ{I}W &
27                                     (∀V. ⬆[i-j-1, 1] V ≡ W → ⊥) & (∀T1. ⬆[j, 1] T1 ≡ U1 → ⊥).
28 #G #L #U1 #U2 #l #m #H elim H -G -L -U1 -U2 -l -m
29 [ /3 width=2 by or_introl/
30 | #I #G #L #K #V #W #j #l #m #Hlj #Hjlm #HLK #HVW #i #Hli #HnW
31   elim (lt_or_ge j i) #Hij
32   [ @or_intror @(ex5_4_intro … HLK) // -HLK
33     [ #X #HXV elim (lift_trans_le … HXV … HVW ?) -V //
34       #Y #HXY >minus_plus <plus_minus_m_m /2 width=2 by/
35     | -HnW /2 width=7 by lift_inv_lref2_be/
36     ]
37   | elim (lift_split … HVW i j) -HVW //
38     #X #_ #H elim HnW -HnW //
39   ]
40 | #a #I #G #L #W1 #W2 #U1 #U2 #l #m #_ #_ #IHW12 #IHU12 #i #Hli #H elim (nlift_inv_bind … H) -H
41   [ #HnW2 elim (IHW12 … HnW2) -IHW12 -HnW2 -IHU12 //
42     [ /4 width=9 by nlift_bind_sn, or_introl/
43     | * /5 width=9 by nlift_bind_sn, ex5_4_intro, or_intror/
44     ]
45   | #HnU2 elim (IHU12 … HnU2) -IHU12 -HnU2 -IHW12 /2 width=1 by yle_succ/
46     [ /4 width=9 by nlift_bind_dx, or_introl/
47     | * #J #K #W #j #Hlj #Hji #HLK #HnW
48       elim (yle_inv_succ1 … Hlj) -Hlj #Hlj #Hj
49       lapply (ylt_O … Hj) -Hj #Hj
50       lapply (drop_inv_drop1_lt … HLK ?) // -HLK #HLK
51       >(plus_minus_m_m j 1) in ⊢ (%→?); [2: /3 width=3 by yle_trans, yle_inv_inj/ ]
52       #HnU1 <minus_le_minus_minus_comm in HnW;
53       /5 width=9 by nlift_bind_dx, monotonic_lt_pred, lt_plus_to_minus_r, ex5_4_intro, or_intror/
54     ]
55   ]
56 | #I #G #L #W1 #W2 #U1 #U2 #l #m #_ #_ #IHW12 #IHU12 #i #Hli #H elim (nlift_inv_flat … H) -H
57   [ #HnW2 elim (IHW12 … HnW2) -IHW12 -HnW2 -IHU12 //
58     [ /4 width=9 by nlift_flat_sn, or_introl/
59     | * /5 width=9 by nlift_flat_sn, ex5_4_intro, or_intror/
60     ]
61   | #HnU2 elim (IHU12 … HnU2) -IHU12 -HnU2 -IHW12 //
62     [ /4 width=9 by nlift_flat_dx, or_introl/
63     | * /5 width=9 by nlift_flat_dx, ex5_4_intro, or_intror/
64   ]
65 ]
66 qed-.