]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2A/multiple/llpx_sn_drop.ma
milestone update in ground_2 and basic_2A
[helm.git] / matita / matita / contribs / lambdadelta / basic_2A / multiple / llpx_sn_drop.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 "ground_2/xoa/ex_4_2.ma".
16 include "basic_2A/substitution/drop_drop.ma".
17 include "basic_2A/multiple/llpx_sn_lreq.ma".
18
19 (* LAZY SN POINTWISE EXTENSION OF A CONTEXT-SENSITIVE REALTION FOR TERMS ****)
20
21 (* Advanced forward lemmas **************************************************)
22
23 lemma llpx_sn_fwd_lref_dx: ∀R,L1,L2,l,i. llpx_sn R l (#i) L1 L2 →
24                            ∀I,K2,V2. ⬇[i] L2 ≡ K2.ⓑ{I}V2 →
25                            i < l ∨
26                            ∃∃K1,V1. ⬇[i] L1 ≡ K1.ⓑ{I}V1 & llpx_sn R 0 V1 K1 K2 &
27                                     R K1 V1 V2 & l ≤ i.
28 #R #L1 #L2 #l #i #H #I #K2 #V2 #HLK2 elim (llpx_sn_fwd_lref … H) -H [ * || * ]
29 [ #_ #H elim (lt_refl_false i)
30   lapply (drop_fwd_length_lt2 … HLK2) -HLK2
31   /2 width=3 by lt_to_le_to_lt/ (**) (* full auto too slow *)
32 | /2 width=1 by or_introl/
33 | #I #K11 #K22 #V11 #V22 #HLK11 #HLK22 #HK12 #HV12 #Hli
34   lapply (drop_mono … HLK22 … HLK2) -L2 #H destruct
35   /3 width=5 by ex4_2_intro, or_intror/
36 ]
37 qed-.
38
39 lemma llpx_sn_fwd_lref_sn: ∀R,L1,L2,l,i. llpx_sn R l (#i) L1 L2 →
40                            ∀I,K1,V1. ⬇[i] L1 ≡ K1.ⓑ{I}V1 →
41                            i < l ∨
42                            ∃∃K2,V2. ⬇[i] L2 ≡ K2.ⓑ{I}V2 & llpx_sn R 0 V1 K1 K2 &
43                                     R K1 V1 V2 & l ≤ i.
44 #R #L1 #L2 #l #i #H #I #K1 #V1 #HLK1 elim (llpx_sn_fwd_lref … H) -H [ * || * ]
45 [ #H #_ elim (lt_refl_false i)
46   lapply (drop_fwd_length_lt2 … HLK1) -HLK1
47   /2 width=3 by lt_to_le_to_lt/ (**) (* full auto too slow *)
48 | /2 width=1 by or_introl/
49 | #I #K11 #K22 #V11 #V22 #HLK11 #HLK22 #HK12 #HV12 #Hli
50   lapply (drop_mono … HLK11 … HLK1) -L1 #H destruct
51   /3 width=5 by ex4_2_intro, or_intror/
52 ]
53 qed-.
54
55 (* Advanced inversion lemmas ************************************************)
56
57 lemma llpx_sn_inv_lref_ge_dx: ∀R,L1,L2,l,i. llpx_sn R l (#i) L1 L2 → l ≤ i →
58                               ∀I,K2,V2. ⬇[i] L2 ≡ K2.ⓑ{I}V2 →
59                               ∃∃K1,V1. ⬇[i] L1 ≡ K1.ⓑ{I}V1 &
60                                        llpx_sn R 0 V1 K1 K2 & R K1 V1 V2.
61 #R #L1 #L2 #l #i #H #Hli #I #K2 #V2 #HLK2 elim (llpx_sn_fwd_lref_dx … H … HLK2) -L2
62 [ #H elim (ylt_yle_false … H Hli)
63 | * /2 width=5 by ex3_2_intro/
64 ]
65 qed-.
66
67 lemma llpx_sn_inv_lref_ge_sn: ∀R,L1,L2,l,i. llpx_sn R l (#i) L1 L2 → l ≤ i →
68                               ∀I,K1,V1. ⬇[i] L1 ≡ K1.ⓑ{I}V1 →
69                               ∃∃K2,V2. ⬇[i] L2 ≡ K2.ⓑ{I}V2 &
70                                        llpx_sn R 0 V1 K1 K2 & R K1 V1 V2.
71 #R #L1 #L2 #l #i #H #Hli #I #K1 #V1 #HLK1 elim (llpx_sn_fwd_lref_sn … H … HLK1) -L1
72 [ #H elim (ylt_yle_false … H Hli)
73 | * /2 width=5 by ex3_2_intro/
74 ]
75 qed-.
76
77 lemma llpx_sn_inv_lref_ge_bi: ∀R,L1,L2,l,i. llpx_sn R l (#i) L1 L2 → l ≤ i →
78                               ∀I1,I2,K1,K2,V1,V2.
79                               ⬇[i] L1 ≡ K1.ⓑ{I1}V1 → ⬇[i] L2 ≡ K2.ⓑ{I2}V2 →
80                               ∧∧ I1 = I2 & llpx_sn R 0 V1 K1 K2 & R K1 V1 V2.
81 #R #L1 #L2 #l #i #HL12 #Hli #I1 #I2 #K1 #K2 #V1 #V2 #HLK1 #HLK2
82 elim (llpx_sn_inv_lref_ge_sn … HL12 … HLK1) // -L1 -l
83 #J #Y #HY lapply (drop_mono … HY … HLK2) -L2 -i #H destruct /2 width=1 by and3_intro/
84 qed-.
85
86 fact llpx_sn_inv_S_aux: ∀R,L1,L2,T,l0. llpx_sn R l0 T L1 L2 → ∀l. l0 = l + 1 →
87                         ∀K1,K2,I,V1,V2. ⬇[l] L1 ≡ K1.ⓑ{I}V1 → ⬇[l] L2 ≡ K2.ⓑ{I}V2 →
88                         llpx_sn R 0 V1 K1 K2 → R K1 V1 V2 → llpx_sn R l T L1 L2.
89 #R #L1 #L2 #T #l0 #H elim H -L1 -L2 -T -l0
90 /2 width=1 by llpx_sn_gref, llpx_sn_free, llpx_sn_sort/
91 [ #L1 #L2 #l0 #i #HL12 #Hil #l #H #K1 #K2 #I #V1 #V2 #HLK1 #HLK2 #HK12 #HV12 destruct
92   elim (yle_split_eq i l) /2 width=1 by llpx_sn_skip, ylt_fwd_succ2/ -HL12 -Hil
93   #H destruct /2 width=9 by llpx_sn_lref/
94 | #I #L1 #L2 #K11 #K22 #V1 #V2 #l0 #i #Hl0i #HLK11 #HLK22 #HK12 #HV12 #_ #l #H #K1 #K2 #J #W1 #W2 #_ #_ #_ #_ destruct
95   /3 width=9 by llpx_sn_lref, yle_pred_sn/
96 | #a #I #L1 #L2 #V #T #l0 #_ #_ #IHV #IHT #l #H #K1 #K2 #J #W1 #W2 #HLK1 #HLK2 #HK12 #HW12 destruct
97   /4 width=9 by llpx_sn_bind, drop_drop/
98 | #I #L1 #L2 #V #T #l0 #_ #_ #IHV #IHT #l #H #K1 #K2 #J #W1 #W2 #HLK1 #HLK2 #HK12 #HW12 destruct
99   /3 width=9 by llpx_sn_flat/
100 ]
101 qed-.
102
103 lemma llpx_sn_inv_S: ∀R,L1,L2,T,l. llpx_sn R (l + 1) T L1 L2 →
104                      ∀K1,K2,I,V1,V2. ⬇[l] L1 ≡ K1.ⓑ{I}V1 → ⬇[l] L2 ≡ K2.ⓑ{I}V2 →
105                      llpx_sn R 0 V1 K1 K2 → R K1 V1 V2 → llpx_sn R l T L1 L2.
106 /2 width=9 by llpx_sn_inv_S_aux/ qed-.
107
108 lemma llpx_sn_inv_bind_O: ∀R. (∀L. reflexive … (R L)) →
109                           ∀a,I,L1,L2,V,T. llpx_sn R 0 (ⓑ{a,I}V.T) L1 L2 →
110                           llpx_sn R 0 V L1 L2 ∧ llpx_sn R 0 T (L1.ⓑ{I}V) (L2.ⓑ{I}V).
111 #R #HR #a #I #L1 #L2 #V #T #H elim (llpx_sn_inv_bind … H) -H
112 /3 width=9 by drop_pair, conj, llpx_sn_inv_S/
113 qed-.
114
115 (* More advanced forward lemmas *********************************************)
116
117 lemma llpx_sn_fwd_bind_O_dx: ∀R. (∀L. reflexive … (R L)) →
118                              ∀a,I,L1,L2,V,T. llpx_sn R 0 (ⓑ{a,I}V.T) L1 L2 →
119                              llpx_sn R 0 T (L1.ⓑ{I}V) (L2.ⓑ{I}V).
120 #R #HR #a #I #L1 #L2 #V #T #H elim (llpx_sn_inv_bind_O … H) -H //
121 qed-.
122
123 (* Advanced properties ******************************************************)
124
125 lemma llpx_sn_bind_repl_O: ∀R,I,L1,L2,V1,V2,T. llpx_sn R 0 T (L1.ⓑ{I}V1) (L2.ⓑ{I}V2) →
126                            ∀J,W1,W2. llpx_sn R 0 W1 L1 L2 → R L1 W1 W2 → llpx_sn R 0 T (L1.ⓑ{J}W1) (L2.ⓑ{J}W2).
127 /3 width=9 by llpx_sn_bind_repl_SO, llpx_sn_inv_S/ qed-.
128
129 lemma llpx_sn_dec: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) →
130                    ∀T,L1,L2,l. Decidable (llpx_sn R l T L1 L2).
131 #R #HR #T #L1 @(f2_ind … rfw … L1 T) -L1 -T
132 #x #IH #L1 * *
133 [ #k #Hx #L2 elim (eq_nat_dec (|L1|) (|L2|)) /3 width=1 by or_introl, llpx_sn_sort/
134 | #i #Hx #L2 elim (eq_nat_dec (|L1|) (|L2|))
135   [ #HL12 #l elim (ylt_split i l) /3 width=1 by llpx_sn_skip, or_introl/
136     #Hli elim (lt_or_ge i (|L1|)) #HiL1
137     elim (lt_or_ge i (|L2|)) #HiL2 /3 width=1 by or_introl, llpx_sn_free/
138     elim (drop_O1_lt (Ⓕ) … HiL2) #I2 #K2 #V2 #HLK2
139     elim (drop_O1_lt (Ⓕ) … HiL1) #I1 #K1 #V1 #HLK1
140     elim (eq_bind2_dec I2 I1)
141     [ #H2 elim (HR K1 V1 V2) -HR
142       [ #H3 elim (IH K1 V1 … K2 0) destruct
143         /3 width=9 by llpx_sn_lref, drop_fwd_rfw, or_introl/
144       ]
145     ]
146     -IH #H3 @or_intror
147     #H elim (llpx_sn_fwd_lref … H) -H [1,3,4,6,7,9: * ]
148     [1,3,5: /3 width=4 by lt_to_le_to_lt, lt_refl_false/
149     |7,8,9: /2 width=4 by ylt_yle_false/
150     ]
151     #Z #Y1 #Y2 #X1 #X2 #HLY1 #HLY2 #HY12 #HX12
152     lapply (drop_mono … HLY1 … HLK1) -HLY1 -HLK1
153     lapply (drop_mono … HLY2 … HLK2) -HLY2 -HLK2
154     #H #H0 destruct /2 width=1 by/
155   ]
156 | #p #Hx #L2 elim (eq_nat_dec (|L1|) (|L2|)) /3 width=1 by or_introl, llpx_sn_gref/
157 | #a #I #V #T #Hx #L2 #l destruct
158   elim (IH L1 V … L2 l) /2 width=1 by/
159   elim (IH (L1.ⓑ{I}V) T … (L2.ⓑ{I}V) (↑l)) -IH /3 width=1 by or_introl, llpx_sn_bind/
160   #H1 #H2 @or_intror
161   #H elim (llpx_sn_inv_bind … H) -H /2 width=1 by/
162 | #I #V #T #Hx #L2 #l destruct
163   elim (IH L1 V … L2 l) /2 width=1 by/
164   elim (IH L1 T … L2 l) -IH /3 width=1 by or_introl, llpx_sn_flat/
165   #H1 #H2 @or_intror
166   #H elim (llpx_sn_inv_flat … H) -H /2 width=1 by/
167 ]
168 -x /4 width=4 by llpx_sn_fwd_length, or_intror/
169 qed-.
170
171 (* Properties on relocation *************************************************)
172
173 lemma llpx_sn_lift_le: ∀R. d_liftable R →
174                        ∀K1,K2,T,l0. llpx_sn R l0 T K1 K2 →
175                        ∀L1,L2,l,m. ⬇[Ⓕ, l, m] L1 ≡ K1 → ⬇[Ⓕ, l, m] L2 ≡ K2 →
176                        ∀U. ⬆[l, m] T ≡ U → l0 ≤ l → llpx_sn R l0 U L1 L2.
177 #R #HR #K1 #K2 #T #l0 #H elim H -K1 -K2 -T -l0
178 [ #K1 #K2 #l0 #k #HK12 #L1 #L2 #l #m #HLK1 #HLK2 #X #H #_ >(lift_inv_sort1 … H) -X
179   lapply (drop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 -l
180   /2 width=1 by llpx_sn_sort/
181 | #K1 #K2 #l0 #i #HK12 #Hil0 #L1 #L2 #l #m #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_lref1 … H) -H
182   * #Hli #H destruct
183   [ lapply (drop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 -l
184     /2 width=1 by llpx_sn_skip/
185   | elim (ylt_yle_false … Hil0) -L1 -L2 -K1 -K2 -m -Hil0
186     /3 width=3 by yle_trans, yle_inj/
187   ]
188 | #I #K1 #K2 #K11 #K22 #V1 #V2 #l0 #i #Hil0 #HK11 #HK22 #HK12 #HV12 #IHK12 #L1 #L2 #l #m #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_lref1 … H) -H
189   * #Hli #H destruct [ -HK12 | -IHK12 ]
190   [ elim (drop_trans_lt … HLK1 … HK11) // -K1
191     elim (drop_trans_lt … HLK2 … HK22) // -Hli -K2
192     /3 width=18 by llpx_sn_lref/
193   | lapply (drop_trans_ge_comm … HLK1 … HK11 ?) // -K1
194     lapply (drop_trans_ge_comm … HLK2 … HK22 ?) // -Hli -Hl0 -K2
195     /3 width=9 by llpx_sn_lref, yle_plus_dx1_trans/
196   ]
197 | #K1 #K2 #l0 #i #HK1 #HK2 #HK12 #L1 #L2 #l #m #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_lref1 … H) -H
198   * #Hil #H destruct
199   lapply (drop_fwd_length_eq2 … HLK1 HLK2 HK12) -HK12
200   [ /3 width=7 by llpx_sn_free, drop_fwd_be/
201   | lapply (drop_fwd_length … HLK1) -HLK1 #HLK1
202     lapply (drop_fwd_length … HLK2) -HLK2 #HLK2
203     @llpx_sn_free [ >HLK1 | >HLK2 ] -Hil -HLK1 -HLK2 /2 width=1 by monotonic_le_plus_r/ (**) (* explicit constructor *)
204   ]
205 | #K1 #K2 #l0 #p #HK12 #L1 #L2 #l #m #HLK1 #HLK2 #X #H #_ >(lift_inv_gref1 … H) -X
206   lapply (drop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 -l -m
207   /2 width=1 by llpx_sn_gref/
208 | #a #I #K1 #K2 #V #T #l0 #_ #_ #IHV #IHT #L1 #L2 #l #m #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_bind1 … H) -H
209   #W #U #HVW #HTU #H destruct /4 width=6 by llpx_sn_bind, drop_skip, yle_succ/
210 | #I #K1 #K2 #V #T #l0 #_ #_ #IHV #IHT #L1 #L2 #l #m #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_flat1 … H) -H
211   #W #U #HVW #HTU #H destruct /3 width=6 by llpx_sn_flat/
212 ]
213 qed-.
214
215 lemma llpx_sn_lift_ge: ∀R,K1,K2,T,l0. llpx_sn R l0 T K1 K2 →
216                        ∀L1,L2,l,m. ⬇[Ⓕ, l, m] L1 ≡ K1 → ⬇[Ⓕ, l, m] L2 ≡ K2 →
217                        ∀U. ⬆[l, m] T ≡ U → l ≤ l0 → llpx_sn R (l0+m) U L1 L2.
218 #R #K1 #K2 #T #l0 #H elim H -K1 -K2 -T -l0
219 [ #K1 #K2 #l0 #k #HK12 #L1 #L2 #l #m #HLK1 #HLK2 #X #H #_ >(lift_inv_sort1 … H) -X
220   lapply (drop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 -l
221   /2 width=1 by llpx_sn_sort/
222 | #K1 #K2 #l0 #i #HK12 #Hil0 #L1 #L2 #l #m #HLK1 #HLK2 #X #H #_ elim (lift_inv_lref1 … H) -H
223   * #_ #H destruct
224   lapply (drop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2
225   [ /3 width=3 by llpx_sn_skip, ylt_plus_dx1_trans/
226   | /3 width=3 by llpx_sn_skip, monotonic_ylt_plus_dx/
227   ]
228 | #I #K1 #K2 #K11 #K22 #V1 #V2 #l0 #i #Hil0 #HK11 #HK22 #HK12 #HV12 #_ #L1 #L2 #l #m #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_lref1 … H) -H
229   * #Hil #H destruct
230   [ elim (ylt_yle_false … Hil0) -I -L1 -L2 -K1 -K2 -K11 -K22 -V1 -V2 -m -Hil0
231     /3 width=3 by ylt_yle_trans, ylt_inj/
232   | lapply (drop_trans_ge_comm … HLK1 … HK11 ?) // -K1
233     lapply (drop_trans_ge_comm … HLK2 … HK22 ?) // -Hil -Hl0 -K2
234     /3 width=9 by llpx_sn_lref, monotonic_yle_plus_dx/
235   ]
236 | #K1 #K2 #l0 #i #HK1 #HK2 #HK12 #L1 #L2 #l #m #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_lref1 … H) -H
237   * #Hil #H destruct
238   lapply (drop_fwd_length_eq2 … HLK1 HLK2 HK12) -HK12
239   [ /3 width=7 by llpx_sn_free, drop_fwd_be/
240   | lapply (drop_fwd_length … HLK1) -HLK1 #HLK1
241     lapply (drop_fwd_length … HLK2) -HLK2 #HLK2
242     @llpx_sn_free [ >HLK1 | >HLK2 ] -Hil -HLK1 -HLK2 /2 width=1 by monotonic_le_plus_r/ (**) (* explicit constructor *)
243   ]
244 | #K1 #K2 #l0 #p #HK12 #L1 #L2 #l #m #HLK1 #HLK2 #X #H #_ >(lift_inv_gref1 … H) -X
245   lapply (drop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 -l
246   /2 width=1 by llpx_sn_gref/
247 | #a #I #K1 #K2 #V #T #l0 #_ #_ #IHV #IHT #L1 #L2 #l #m #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_bind1 … H) -H
248   #W #U #HVW #HTU #H destruct /4 width=5 by llpx_sn_bind, drop_skip, yle_succ/
249 | #I #K1 #K2 #V #T #l0 #_ #_ #IHV #IHT #L1 #L2 #l #m #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_flat1 … H) -H
250   #W #U #HVW #HTU #H destruct /3 width=5 by llpx_sn_flat/
251 ]
252 qed-.
253
254 (* Inversion lemmas on relocation *******************************************)
255
256 lemma llpx_sn_inv_lift_le: ∀R. d_deliftable_sn R →
257                            ∀L1,L2,U,l0. llpx_sn R l0 U L1 L2 →
258                            ∀K1,K2,l,m. ⬇[Ⓕ, l, m] L1 ≡ K1 → ⬇[Ⓕ, l, m] L2 ≡ K2 →
259                            ∀T. ⬆[l, m] T ≡ U → l0 ≤ l → llpx_sn R l0 T K1 K2.
260 #R #HR #L1 #L2 #U #l0 #H elim H -L1 -L2 -U -l0
261 [ #L1 #L2 #l0 #k #HL12 #K1 #K2 #l #m #HLK1 #HLK2 #X #H #_ >(lift_inv_sort2 … H) -X
262   lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -l -m
263   /2 width=1 by llpx_sn_sort/
264 | #L1 #L2 #l0 #i #HL12 #Hil0 #K1 #K2 #l #m #HLK1 #HLK2 #X #H #_ elim (lift_inv_lref2 … H) -H
265   * #_ #H destruct
266   lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2
267   [ /2 width=1 by llpx_sn_skip/
268   | /3 width=3 by llpx_sn_skip, yle_ylt_trans/
269   ]
270 | #I #L1 #L2 #K11 #K22 #W1 #W2 #l0 #i #Hil0 #HLK11 #HLK22 #HK12 #HW12 #IHK12 #K1 #K2 #l #m #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_lref2 … H) -H
271   * #Hil #H destruct [ -HK12 | -IHK12 ]
272   [ elim (drop_conf_lt … HLK1 … HLK11) // -L1 #L1 #V1 #HKL1 #HKL11 #HVW1
273     elim (drop_conf_lt … HLK2 … HLK22) // -Hil -L2 #L2 #V2 #HKL2 #HKL22 #HVW2
274     elim (HR … HW12 … HKL11 … HVW1) -HR #V0 #HV0 #HV12
275     lapply (lift_inj … HV0 … HVW2) -HV0 -HVW2 #H destruct
276     /3 width=10 by llpx_sn_lref/
277   | lapply (drop_conf_ge … HLK1 … HLK11 ?) // -L1
278     lapply (drop_conf_ge … HLK2 … HLK22 ?) // -L2 -Hil0
279     elim (le_inv_plus_l … Hil) -Hil /4 width=9 by llpx_sn_lref, yle_trans, yle_inj/ (**) (* slow *)
280   ]
281 | #L1 #L2 #l0 #i #HL1 #HL2 #HL12 #K1 #K2 #l #m #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_lref2 … H) -H
282   * #_ #H destruct
283   lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12)
284   [ lapply (drop_fwd_length_le4 … HLK1) -HLK1
285     lapply (drop_fwd_length_le4 … HLK2) -HLK2
286     #HKL2 #HKL1 #HK12 @llpx_sn_free // /2 width=3 by transitive_le/ (**) (* full auto too slow *)
287   | lapply (drop_fwd_length … HLK1) -HLK1 #H >H in HL1; -H
288     lapply (drop_fwd_length … HLK2) -HLK2 #H >H in HL2; -H
289     /3 width=1 by llpx_sn_free, le_plus_to_minus_r/
290   ]
291 | #L1 #L2 #l0 #p #HL12 #K1 #K2 #l #m #HLK1 #HLK2 #X #H #_ >(lift_inv_gref2 … H) -X
292   lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -l -m
293   /2 width=1 by llpx_sn_gref/
294 | #a #I #L1 #L2 #W #U #l0 #_ #_ #IHW #IHU #K1 #K2 #l #m #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_bind2 … H) -H
295   #V #T #HVW #HTU #H destruct /4 width=6 by llpx_sn_bind, drop_skip, yle_succ/
296 | #I #L1 #L2 #W #U #l0 #_ #_ #IHW #IHU #K1 #K2 #l #m #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_flat2 … H) -H
297   #V #T #HVW #HTU #H destruct /3 width=6 by llpx_sn_flat/
298 ]
299 qed-.
300
301 lemma llpx_sn_inv_lift_be: ∀R,L1,L2,U,l0. llpx_sn R l0 U L1 L2 →
302                            ∀K1,K2,l,m. ⬇[Ⓕ, l, m] L1 ≡ K1 → ⬇[Ⓕ, l, m] L2 ≡ K2 →
303                            ∀T. ⬆[l, m] T ≡ U → l ≤ l0 → l0 ≤ yinj l + m → llpx_sn R l T K1 K2.
304 #R #L1 #L2 #U #l0 #H elim H -L1 -L2 -U -l0
305 [ #L1 #L2 #l0 #k #HL12 #K1 #K2 #l #m #HLK1 #HLK2 #X #H #_ #_ >(lift_inv_sort2 … H) -X
306   lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -l0 -m
307   /2 width=1 by llpx_sn_sort/
308 | #L1 #L2 #l0 #i #HL12 #Hil0 #K1 #K2 #l #m #HLK1 #HLK2 #X #H #Hl0 #Hl0m elim (lift_inv_lref2 … H) -H
309   * #Hil #H destruct
310   [ lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2
311     -Hil0 /3 width=1 by llpx_sn_skip, ylt_inj/
312   | elim (ylt_yle_false … Hil0) -L1 -L2 -Hl0 -Hil0
313     /3 width=3 by yle_trans, yle_inj/ (**) (* slow *)
314   ]
315 | #I #L1 #L2 #K11 #K22 #W1 #W2 #l0 #i #Hil0 #HLK11 #HLK22 #HK12 #HW12 #_ #K1 #K2 #l #m #HLK1 #HLK2 #X #H #Hl0 #Hl0m elim (lift_inv_lref2 … H) -H
316   * #Hil #H destruct
317   [ elim (ylt_yle_false … Hil0) -I -L1 -L2 -K11 -K22 -W1 -W2 -Hl0m -Hil0
318     /3 width=3 by ylt_yle_trans, ylt_inj/
319   | lapply (drop_conf_ge … HLK1 … HLK11 ?) // -L1
320     lapply (drop_conf_ge … HLK2 … HLK22 ?) // -L2 -Hil0 -Hl0 -Hl0m
321     elim (le_inv_plus_l … Hil) -Hil /3 width=9 by llpx_sn_lref, yle_inj/
322   ]
323 | #L1 #L2 #l0 #i #HL1 #HL2 #HL12 #K1 #K2 #l #m #HLK1 #HLK2 #X #H #Hl0 #Hl0m elim (lift_inv_lref2 … H) -H
324   * #_ #H destruct
325   lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12)
326   [ lapply (drop_fwd_length_le4 … HLK1) -HLK1
327     lapply (drop_fwd_length_le4 … HLK2) -HLK2
328     #HKL2 #HKL1 #HK12 @llpx_sn_free // /2 width=3 by transitive_le/ (**) (* full auto too slow *)
329   | lapply (drop_fwd_length … HLK1) -HLK1 #H >H in HL1; -H
330     lapply (drop_fwd_length … HLK2) -HLK2 #H >H in HL2; -H
331     /3 width=1 by llpx_sn_free, le_plus_to_minus_r/
332   ]
333 | #L1 #L2 #l0 #p #HL12 #K1 #K2 #l #m #HLK1 #HLK2 #X #H #_ #_ >(lift_inv_gref2 … H) -X
334   lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -l0 -m
335   /2 width=1 by llpx_sn_gref/
336 | #a #I #L1 #L2 #W #U #l0 #_ #_ #IHW #IHU #K1 #K2 #l #m #HLK1 #HLK2 #X #H #Hl0 #Hl0m elim (lift_inv_bind2 … H) -H
337   >commutative_plus #V #T #HVW #HTU #H destruct
338   @llpx_sn_bind [ /2 width=5 by/ ] -IHW (**) (* explicit constructor *)
339   @(IHU … HTU) -IHU -HTU /2 width=1 by drop_skip, yle_succ/
340 | #I #L1 #L2 #W #U #l0 #_ #_ #IHW #IHU #K1 #K2 #l #m #HLK1 #HLK2 #X #H #Hl0 #Hl0m elim (lift_inv_flat2 … H) -H
341   #V #T #HVW #HTU #H destruct /3 width=6 by llpx_sn_flat/
342 ]
343 qed-.
344
345 lemma llpx_sn_inv_lift_ge: ∀R,L1,L2,U,l0. llpx_sn R l0 U L1 L2 →
346                            ∀K1,K2,l,m. ⬇[Ⓕ, l, m] L1 ≡ K1 → ⬇[Ⓕ, l, m] L2 ≡ K2 →
347                            ∀T. ⬆[l, m] T ≡ U → yinj l + m ≤ l0 → llpx_sn R (l0-m) T K1 K2.
348 #R #L1 #L2 #U #l0 #H elim H -L1 -L2 -U -l0
349 [ #L1 #L2 #l0 #k #HL12 #K1 #K2 #l #m #HLK1 #HLK2 #X #H #_ >(lift_inv_sort2 … H) -X
350   lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -l
351   /2 width=1 by llpx_sn_sort/
352 | #L1 #L2 #l0 #i #HL12 #Hil0 #K1 #K2 #l #m #HLK1 #HLK2 #X #H #Hlml0 elim (lift_inv_lref2 … H) -H
353   * #Hil #H destruct [ -Hil0 | -Hlml0 ]
354   lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2
355   [ /4 width=3 by llpx_sn_skip, yle_plus1_to_minus_inj2, ylt_yle_trans, ylt_inj/
356   | elim (le_inv_plus_l … Hil) -Hil #_
357     /4 width=1 by llpx_sn_skip, monotonic_ylt_minus_dx, yle_inj/
358   ]
359 | #I #L1 #L2 #K11 #K22 #W1 #W2 #l0 #i #Hil0 #HLK11 #HLK22 #HK12 #HW12 #_ #K1 #K2 #l #m #HLK1 #HLK2 #X #H #Hlml0 elim (lift_inv_lref2 … H) -H
360   * #Hil #H destruct
361   [ elim (ylt_yle_false … Hil0) -I -L1 -L2 -K11 -K22 -W1 -W2 -Hil0
362     /3 width=3 by yle_fwd_plus_sn1, ylt_yle_trans, ylt_inj/
363   | lapply (drop_conf_ge … HLK1 … HLK11 ?) // -L1
364     lapply (drop_conf_ge … HLK2 … HLK22 ?) // -L2 -Hlml0 -Hil
365     /3 width=9 by llpx_sn_lref, monotonic_yle_minus_dx/
366   ]
367 | #L1 #L2 #l0 #i #HL1 #HL2 #HL12 #K1 #K2 #l #m #HLK1 #HLK2 #X #H #Hlml0 elim (lift_inv_lref2 … H) -H
368   * #_ #H destruct
369   lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12)
370   [ lapply (drop_fwd_length_le4 … HLK1) -HLK1
371     lapply (drop_fwd_length_le4 … HLK2) -HLK2
372     #HKL2 #HKL1 #HK12 @llpx_sn_free // /2 width=3 by transitive_le/ (**) (* full auto too slow *)
373   | lapply (drop_fwd_length … HLK1) -HLK1 #H >H in HL1; -H
374     lapply (drop_fwd_length … HLK2) -HLK2 #H >H in HL2; -H
375     /3 width=1 by llpx_sn_free, le_plus_to_minus_r/
376   ]
377 | #L1 #L2 #l0 #p #HL12 #K1 #K2 #l #m #HLK1 #HLK2 #X #H #_ >(lift_inv_gref2 … H) -X
378   lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -l
379   /2 width=1 by llpx_sn_gref/
380 | #a #I #L1 #L2 #W #U #l0 #_ #_ #IHW #IHU #K1 #K2 #l #m #HLK1 #HLK2 #X #H #Hlml0 elim (lift_inv_bind2 … H) -H
381   #V #T #HVW #HTU #H destruct
382   @llpx_sn_bind [ /2 width=5 by/ ] -IHW (**) (* explicit constructor *)
383   <yminus_succ1_inj /2 width=2 by yle_fwd_plus_sn2/
384   @(IHU … HTU) -IHU -HTU /2 width=1 by drop_skip, yle_succ/
385 | #I #L1 #L2 #W #U #l0 #_ #_ #IHW #IHU #K1 #K2 #l #m #HLK1 #HLK2 #X #H #Hlml0 elim (lift_inv_flat2 … H) -H
386   #V #T #HVW #HTU #H destruct /3 width=5 by llpx_sn_flat/
387 ]
388 qed-.
389
390 (* Advanced inversion lemmas on relocation **********************************)
391
392 lemma llpx_sn_inv_lift_O: ∀R,L1,L2,U. llpx_sn R 0 U L1 L2 →
393                           ∀K1,K2,m. ⬇[m] L1 ≡ K1 → ⬇[m] L2 ≡ K2 →
394                           ∀T. ⬆[0, m] T ≡ U → llpx_sn R 0 T K1 K2.
395 /2 width=11 by llpx_sn_inv_lift_be/ qed-.
396
397 lemma llpx_sn_drop_conf_O: ∀R,L1,L2,U. llpx_sn R 0 U L1 L2 →
398                            ∀K1,m. ⬇[m] L1 ≡ K1 → ∀T. ⬆[0, m] T ≡ U →
399                            ∃∃K2. ⬇[m] L2 ≡ K2 & llpx_sn R 0 T K1 K2.
400 #R #L1 #L2 #U #HU #K1 #m #HLK1 #T #HTU elim (llpx_sn_fwd_drop_sn … HU … HLK1)
401 /3 width=10 by llpx_sn_inv_lift_O, ex2_intro/
402 qed-.
403
404 lemma llpx_sn_drop_trans_O: ∀R,L1,L2,U. llpx_sn R 0 U L1 L2 →
405                             ∀K2,m. ⬇[m] L2 ≡ K2 → ∀T. ⬆[0, m] T ≡ U →
406                             ∃∃K1. ⬇[m] L1 ≡ K1 & llpx_sn R 0 T K1 K2.
407 #R #L1 #L2 #U #HU #K2 #m #HLK2 #T #HTU elim (llpx_sn_fwd_drop_dx … HU … HLK2)
408 /3 width=10 by llpx_sn_inv_lift_O, ex2_intro/
409 qed-.
410
411 (* Inversion lemmas on negated lazy pointwise extension *********************)
412
413 lemma nllpx_sn_inv_bind: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) →
414                          ∀a,I,L1,L2,V,T,l. (llpx_sn R l (ⓑ{a,I}V.T) L1 L2 → ⊥) →
415                          (llpx_sn R l V L1 L2 → ⊥) ∨ (llpx_sn R (↑l) T (L1.ⓑ{I}V) (L2.ⓑ{I}V) → ⊥).
416 #R #HR #a #I #L1 #L2 #V #T #l #H elim (llpx_sn_dec … HR V L1 L2 l)
417 /4 width=1 by llpx_sn_bind, or_intror, or_introl/
418 qed-.
419
420 lemma nllpx_sn_inv_flat: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) →
421                          ∀I,L1,L2,V,T,l. (llpx_sn R l (ⓕ{I}V.T) L1 L2 → ⊥) →
422                          (llpx_sn R l V L1 L2 → ⊥) ∨ (llpx_sn R l T L1 L2 → ⊥).
423 #R #HR #I #L1 #L2 #V #T #l #H elim (llpx_sn_dec … HR V L1 L2 l)
424 /4 width=1 by llpx_sn_flat, or_intror, or_introl/
425 qed-.
426
427 lemma nllpx_sn_inv_bind_O: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) →
428                            ∀a,I,L1,L2,V,T. (llpx_sn R 0 (ⓑ{a,I}V.T) L1 L2 → ⊥) →
429                            (llpx_sn R 0 V L1 L2 → ⊥) ∨ (llpx_sn R 0 T (L1.ⓑ{I}V) (L2.ⓑ{I}V) → ⊥).
430 #R #HR #a #I #L1 #L2 #V #T #H elim (llpx_sn_dec … HR V L1 L2 0)
431 /4 width=1 by llpx_sn_bind_O, or_intror, or_introl/
432 qed-.