∃∃V2,T2. V1 ➡ V2 & T1 ➡ T2 & U2 = ⓛV2. T2.
/2 width=3/ qed-.
+(* Basic_1: was pr2_gen_appl *)
+lemma cpr_inv_appl1: ∀L,V1,U0,U2. L ⊢ ⓐV1. U0 ➡ U2 →
+ ∨∨ ∃∃V2,T2. L ⊢ V1 ➡ V2 & L ⊢ U0 ➡ T2 &
+ U2 = ⓐV2. T2
+ | ∃∃V2,W,T1,T2. L ⊢ V1 ➡ V2 & L. ⓓV2 ⊢ T1 ➡ T2 &
+ U0 = ⓛW. T1 &
+ U2 = ⓓV2. T2
+ | ∃∃V2,V,W1,W2,T1,T2. L ⊢ V1 ➡ V2 & L ⊢ W1 ➡ W2 & L. ⓓW2 ⊢ T1 ➡ T2 &
+ ⇧[0,1] V2 ≡ V &
+ U0 = ⓓW1. T1 &
+ U2 = ⓓW2. ⓐV. T2.
+#L #V1 #U0 #Y * #X #H1 #H2
+elim (tpr_inv_appl1 … H1) -H1 *
+[ #V #U #HV1 #HU0 #H destruct
+ elim (tpss_inv_flat1 … H2) -H2 #V2 #U2 #HV2 #HU2 #H destruct /4 width=5/
+| #V #W #T0 #T #HV1 #HT0 #H #H1 destruct
+ elim (tpss_inv_bind1 … H2) -H2 #V2 #T2 #HV2 #HT2 #H destruct
+ lapply (tpss_weak … HT2 0 (|L|+1) ? ?) -HT2 // /4 width=8/
+| #V0 #V #W #W0 #T #T0 #HV10 #HW0 #HT0 #HV0 #H #H1 destruct
+ elim (tpss_inv_bind1 … H2) -H2 #W2 #X #HW02 #HX #HY destruct
+ elim (tpss_inv_flat1 … HX) -HX #V2 #T2 #HV2 #HT2 #H destruct
+ elim (tpss_inv_lift1_ge … HV2 … HV0 ?) -V // [3: /2 width=1/ |2: skip ] #V <minus_plus_m_m
+ lapply (tpss_weak … HT2 0 (|L|+1) ? ?) -HT2 // /4 width=12/
+]
+qed-.
+
+(* Note: the main property of simple terms *)
+lemma cpr_inv_appl1_simple: ∀L,V1,T1,U. L ⊢ ⓐV1. T1 ➡ U → 𝐒[T1] →
+ ∃∃V2,T2. L ⊢ V1 ➡ V2 & L ⊢ T1 ➡ T2 &
+ U = ⓐV2. T2.
+#L #V1 #T1 #U #H #HT1
+elim (cpr_inv_appl1 … H) -H *
+[ /2 width=5/
+| #V2 #W #W1 #W2 #_ #_ #H #_ destruct
+ elim (simple_inv_bind … HT1)
+| #V2 #V #W1 #W2 #U1 #U2 #_ #_ #_ #_ #H #_ destruct
+ elim (simple_inv_bind … HT1)
+]
+qed-.
+
(* Relocation properties ****************************************************)
(* Basic_1: was: pr2_lift *)
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||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/unfold/ltpss_ltpss.ma".
+include "Basic_2/reducibility/cpr.ma".
+
+(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS ****************************)
+
+(* Advanced inversion lemmas ************************************************)
+
+(* Basic_1: was pr2_gen_abbr *)
+lemma cpr_inv_abbr1: ∀L,V1,T1,U2. L ⊢ ⓓV1. T1 ➡ U2 →
+ (∃∃V2,T2,T. L ⊢ V1 ➡ V2 & L. ⓓV2 ⊢ T1 ➡ T2 &
+ L. ⓓV2 ⊢ T [1, |L|] ▶* T2 &
+ U2 = ⓓV2. T
+ ) ∨
+ ∃∃T. ⇧[0,1] T ≡ T1 & L ⊢ T ➡ U2.
+#L #V1 #T1 #Y * #X #H1 #H2
+elim (tpr_inv_abbr1 … H1) -H1 *
+[ #V #T0 #T #HV1 #HT10 #HT0 #H destruct
+ elim (tpss_inv_bind1 … H2) -H2 #V2 #T2 #HV2 #HT2 #H destruct
+ lapply (tps_lsubs_conf … HT0 (L. ⓓV) ?) -HT0 /2 width=1/ #HT0
+ elim (ltpss_tps_conf … HT0 (L. ⓓV2) 1 (|L|) ?) -HT0 /2 width=1/ #V0 #HV20 #HV0
+ lapply (tpss_weak_all … HV20) -HV20 #HV20
+ lapply (tpss_lsubs_conf … HV0 (L. ⓓV2) ?) -HV0 /2 width=1/ #HV0
+ elim (tpss_conf_eq … HT2 … HV0) -T #T3 #HT23 #HVT3
+ lapply (tpss_weak_all … HVT3) -HVT3 #HVT3
+ lapply (tpss_trans_eq … HV20 … HVT3) -V0 /4 width=7/
+| /4 width=5/
+]
+qed-.