1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "Basic_2/substitution/ldrop_ldrop.ma".
16 include "Basic_2/unfold/ldrops.ma".
18 (* GENERIC LOCAL ENVIRONMENT SLICING ****************************************)
20 (* Properties concerning basic local environment slicing ********************)
22 lemma ldrops_ldrop_trans: ∀L1,L,des. ⇩*[des] L1 ≡ L → ∀L2,i. ⇩[0, i] L ≡ L2 →
23 ∃∃L0,des0,i0. ⇩[0, i0] L1 ≡ L0 & ⇩*[des0] L0 ≡ L2 &
24 @[i] des ≡ i0 & des ▭ i ≡ des0.
25 #L1 #L #des #H elim H -L1 -L -des
27 | #L1 #L3 #L #des3 #d #e #_ #HL3 #IHL13 #L2 #i #HL2
28 elim (lt_or_ge i d) #Hid
29 [ elim (ldrop_trans_le … HL3 … HL2 ?) -L /2 width=2/ #L #HL3 #HL2
30 elim (IHL13 … HL3) -L3 /3 width=7/
31 | lapply (ldrop_trans_ge … HL3 … HL2 ?) -L // #HL32
32 elim (IHL13 … HL32) -L3 /3 width=7/