(* transitions *)
inductive rtm_step: relation rtm ≝
-| rtm_ldrop : ∀G,u,E,I,t,D,V,S,i.
+| rtm_drop : ∀G,u,E,I,t,D,V,S,i.
rtm_step (mk_rtm G u (E. ④{I} {t, D, V}) S (#(i + 1)))
(mk_rtm G u E S (#i))
| rtm_ldelta: ∀G,u,E,t,D,V,S.
(**************************************************************************)
include "basic_2/grammar/genv.ma".
-include "basic_2/multiple/ldrops.ma".
+include "basic_2/multiple/drops.ma".
(* ABSTRACT COMPUTATION PROPERTIES ******************************************)
(**************************************************************************)
include "basic_2/multiple/lifts_lifts.ma".
-include "basic_2/multiple/ldrops_ldrops.ma".
+include "basic_2/multiple/drops_drops.ma".
include "basic_2/static/aaa_lifts.ma".
include "basic_2/static/aaa_aaa.ma".
-include "basic_2/computation/lsubc_ldrops.ma".
+include "basic_2/computation/lsubc_drops.ma".
(* ABSTRACT COMPUTATION PROPERTIES ******************************************)
| #I #G #L1 #K1 #V1 #B #i #HLK1 #HKV1B #IHB #L0 #des #HL01 #X #H #L2 #HL20
lapply (aacr_acr … H1RP H2RP B) #HB
elim (lifts_inv_lref1 … H) -H #i1 #Hi1 #H destruct
- lapply (ldrop_fwd_drop2 … HLK1) #HK1b
- elim (ldrops_ldrop_trans … HL01 … HLK1) #X #des1 #i0 #HL0 #H #Hi0 #Hdes1
+ lapply (drop_fwd_drop2 … HLK1) #HK1b
+ elim (drops_drop_trans … HL01 … HLK1) #X #des1 #i0 #HL0 #H #Hi0 #Hdes1
>(at_mono … Hi1 … Hi0) -i1
- elim (ldrops_inv_skip2 … Hdes1 … H) -des1 #K0 #V0 #des0 #Hdes0 #HK01 #HV10 #H destruct
- elim (lsubc_ldrop_O1_trans … HL20 … HL0) -HL0 #X #HLK2 #H
+ elim (drops_inv_skip2 … Hdes1 … H) -des1 #K0 #V0 #des0 #Hdes0 #HK01 #HV10 #H destruct
+ elim (lsubc_drop_O1_trans … HL20 … HL0) -HL0 #X #HLK2 #H
elim (lsubc_inv_pair2 … H) -H *
[ #K2 #HK20 #H destruct
elim (lift_total V0 0 (i0 +1)) #V #HV0
elim (lifts_lift_trans … Hi0 … Hdes0 … HV10 … HV0) -HV10 #V2 #HV12 #HV2
- @(s5 … HB … (◊) … HV0 HLK2) /3 width=7 by ldrops_cons, lifts_cons/ (* Note: uses IHB HL20 V2 HV0 *)
+ @(s5 … HB … (◊) … HV0 HLK2) /3 width=7 by drops_cons, lifts_cons/ (* Note: uses IHB HL20 V2 HV0 *)
| -HLK1 -IHB -HL01 -HL20 -HK1b -Hi0 -Hdes0
#K2 #V2 #A2 #HKV2A #H1KV0A #H2KV0A #_ #H1 #H2 destruct
- lapply (ldrop_fwd_drop2 … HLK2) #HLK2b
+ lapply (drop_fwd_drop2 … HLK2) #HLK2b
lapply (aaa_lifts … HK01 … HV10 HKV1B) -HKV1B -HK01 -HV10 #HKV0B
lapply (aaa_mono … H2KV0A … HKV0B) #H destruct -H2KV0A -HKV0B
elim (lift_total V0 0 (i0 +1)) #V3 #HV03
lapply (aacr_acr … H1RP H2RP A) #HA
lapply (aacr_acr … H1RP H2RP B) #HB
lapply (s1 … HB) -HB #HB
- @(s6 … HA … (◊) (◊)) /3 width=5 by lsubc_pair, ldrops_skip, liftv_nil/
+ @(s6 … HA … (◊) (◊)) /3 width=5 by lsubc_pair, drops_skip, liftv_nil/
| #a #G #L #W #T #B #A #HLWB #_ #IHB #IHA #L0 #des #HL0 #X #H #L2 #HL02
elim (lifts_inv_bind1 … H) -H #W0 #T0 #HW0 #HT0 #H destruct
@(aacr_abst … H1RP H2RP) [ /2 width=5 by/ ]
#L3 #V3 #W3 #T3 #des3 #HL32 #HW03 #HT03 #H1B #H2B
- elim (ldrops_lsubc_trans … H1RP H2RP … HL32 … HL02) -L2 #L2 #HL32 #HL20
- lapply (aaa_lifts … L2 W3 … (des @@ des3) … HLWB) -HLWB /2 width=4 by ldrops_trans, lifts_trans/ #HLW2B
+ elim (drops_lsubc_trans … H1RP H2RP … HL32 … HL02) -L2 #L2 #HL32 #HL20
+ lapply (aaa_lifts … L2 W3 … (des @@ des3) … HLWB) -HLWB /2 width=4 by drops_trans, lifts_trans/ #HLW2B
@(IHA (L2. ⓛW3) … (des + 1 @@ des3 + 1)) -IHA
- /3 width=5 by lsubc_abbr, ldrops_trans, ldrops_skip, lifts_trans/
+ /3 width=5 by lsubc_abbr, drops_trans, drops_skip, lifts_trans/
| #G #L #V #T #B #A #_ #_ #IHB #IHA #L0 #des #HL0 #X #H #L2 #HL20
elim (lifts_inv_flat1 … H) -H #V0 #T0 #HV0 #HT0 #H destruct
- /3 width=10 by ldrops_nil, lifts_nil/
+ /3 width=10 by drops_nil, lifts_nil/
| #G #L #V #T #A #_ #_ #IH1A #IH2A #L0 #des #HL0 #X #H #L2 #HL20
elim (lifts_inv_flat1 … H) -H #V0 #T0 #HV0 #HT0 #H destruct
lapply (aacr_acr … H1RP H2RP A) #HA
(* Basic_1: was: sc3_arity *)
lemma aacr_aaa: ∀RR,RS,RP. acp RR RS RP → acr RR RS RP (λG,L,T. RP G L T) →
∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ⦃G, L, T⦄ ϵ[RP] 〚A〛.
-/2 width=8 by ldrops_nil, lifts_nil, aacr_aaa_csubc_lifts/ qed.
+/2 width=8 by drops_nil, lifts_nil, aacr_aaa_csubc_lifts/ qed.
lemma acp_aaa: ∀RR,RS,RP. acp RR RS RP → acr RR RS RP (λG,L,T. RP G L T) →
∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → RP G L T.
include "basic_2/grammar/aarity.ma".
include "basic_2/multiple/gr2_gr2.ma".
include "basic_2/multiple/lifts_lift_vector.ma".
-include "basic_2/multiple/ldrops_ldrop.ma".
+include "basic_2/multiple/drops_drop.ma".
include "basic_2/computation/acp.ma".
(* ABSTRACT COMPUTATION PROPERTIES ******************************************)
∀A. acr RR RS RP (aacr RP A).
#RR #RS #RP #H1RP #H2RP #A elim A -A normalize //
#B #A #IHB #IHA @mk_acr normalize
-[ /3 width=7 by ldrops_cons, lifts_cons/
+[ /3 width=7 by drops_cons, lifts_cons/
| #G #L #T #H
elim (cp1 … H1RP G L) #k #HK
lapply (H ? (⋆k) ? (⟠) ? ? ?) -H
| #I #G #L #K #Vs #V1 #V2 #i #HA #HV12 #HLK #L0 #V0 #X #des #HB #HL0 #H
elim (lifts_inv_applv1 … H) -H #V0s #Y #HV0s #HY #H destruct
elim (lifts_inv_lref1 … HY) -HY #i0 #Hi0 #H destruct
- elim (ldrops_ldrop_trans … HL0 … HLK) #X #des0 #i1 #HL02 #H #Hi1 #Hdes0
+ elim (drops_drop_trans … HL0 … HLK) #X #des0 #i1 #HL02 #H #Hi1 #Hdes0
>(at_mono … Hi1 … Hi0) in HL02; -i1 #HL02
- elim (ldrops_inv_skip2 … Hdes0 … H) -H -des0 #L2 #W1 #des0 #Hdes0 #HLK #HVW1 #H destruct
+ elim (drops_inv_skip2 … Hdes0 … H) -H -des0 #L2 #W1 #des0 #Hdes0 #HLK #HVW1 #H destruct
elim (lift_total W1 0 (i0 + 1)) #W2 #HW12
elim (lifts_lift_trans … Hdes0 … HVW1 … HW12) // -Hdes0 -Hi0 #V3 #HV13 #HVW2
>(lift_mono … HV13 … HV12) in HVW2; -V3 #HVW2
elim (lift_total V10 0 1) #V20 #HV120
elim (liftv_total 0 1 V10s) #V20s #HV120s
@(s6 … IHA … (V10 @ V10s) (V20 @ V20s)) /3 width=7 by rp_lifts, liftv_cons/
- @(HA … (des + 1)) /2 width=2 by ldrops_skip/
- [ @(s0 … IHB … HB … HV120) /2 width=2 by ldrop_drop/
+ @(HA … (des + 1)) /2 width=2 by drops_skip/
+ [ @(s0 … IHB … HB … HV120) /2 width=2 by drop_drop/
| @lifts_applv //
elim (liftsv_liftv_trans_le … HV10s … HV120s) -V10s #V10s #HV10s #HV120s
>(liftv_mono … HV12s … HV10s) -V1s //
#G #L2 #T1 #T2 #HT12 elim HT12 -G -L2 -T1 -T2
[ /2 width=3 by/
| #G #L2 #K2 #V0 #V2 #W2 #i #HLK2 #_ #HVW2 #IHV02 #L1 #HL12
- elim (lpr_ldrop_trans_O1 … HL12 … HLK2) -L2 #X #HLK1 #H
+ elim (lpr_drop_trans_O1 … HL12 … HLK2) -L2 #X #HLK1 #H
elim (lpr_inv_pair2 … H) -H #K1 #V1 #HK12 #HV10 #H destruct
/4 width=6 by cprs_strap2, cprs_delta/
|3,7: /4 width=1 by lpr_pair, cprs_bind, cprs_beta/
∀W2. ⇧[0, i + 1] V2 ≡ W2 → ⦃G, L⦄ ⊢ #i ➡* W2.
#G #L #K #V #V2 #i #HLK #H elim H -V2 [ /3 width=6 by cpr_cprs, cpr_delta/ ]
#V1 #V2 #_ #HV12 #IHV1 #W2 #HVW2
-lapply (ldrop_fwd_drop2 … HLK) -HLK #HLK
+lapply (drop_fwd_drop2 … HLK) -HLK #HLK
elim (lift_total V1 0 (i+1)) /4 width=12 by cpr_lift, cprs_strap1/
qed.
elim (cpr_inv_lref1 … HT2) -HT2 /2 width=1 by or_introl/
* /4 width=6 by cpr_cprs, ex3_3_intro, or_intror/
| * #K #V1 #T1 #HLK #HVT1 #HT1
- lapply (ldrop_fwd_drop2 … HLK) #H0LK
+ lapply (drop_fwd_drop2 … HLK) #H0LK
elim (cpr_inv_lift1 … HT2 … H0LK … HT1) -H0LK -T
/4 width=6 by cprs_strap1, ex3_3_intro, or_intror/
]
(* *)
(**************************************************************************)
-include "basic_2/reduction/lpx_ldrop.ma".
+include "basic_2/reduction/lpx_drop.ma".
include "basic_2/computation/cpxs_lift.ma".
(* CONTEXT-SENSITIVE EXTENDED PARALLEL COMPUTATION ON TERMS *****************)
[ /2 width=3 by/
| /3 width=2 by cpx_cpxs, cpx_st/
| #I #G #L2 #K2 #V0 #V2 #W2 #i #HLK2 #_ #HVW2 #IHV02 #L1 #HL12
- elim (lpx_ldrop_trans_O1 … HL12 … HLK2) -L2 #X #HLK1 #H
+ elim (lpx_drop_trans_O1 … HL12 … HLK2) -L2 #X #HLK1 #H
elim (lpx_inv_pair2 … H) -H #K1 #V1 #HK12 #HV10 #H destruct
/4 width=7 by cpxs_delta, cpxs_strap2/
|4,9: /4 width=1 by cpxs_beta, cpxs_bind, lpx_pair/
#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
[ #I #G #L #V1 #V2 #HV12 #_ elim (lift_total V2 0 1)
#U2 #HVU2 @(ex3_intro … U2)
- [1,3: /3 width=7 by fqu_drop, cpxs_delta, ldrop_pair, ldrop_drop/
+ [1,3: /3 width=7 by fqu_drop, cpxs_delta, drop_pair, drop_drop/
| #H destruct /2 width=7 by lift_inv_lref2_be/
]
| #I #G #L #V1 #T #V2 #HV12 #H @(ex3_intro … (②{I}V2.T))
∀W2. ⇧[0, i+1] V2 ≡ W2 → ⦃G, L⦄ ⊢ #i ➡*[h, g] W2.
#h #g #I #G #L #K #V #V2 #i #HLK #H elim H -V2
[ /3 width=9 by cpx_cpxs, cpx_delta/
-| #V1 lapply (ldrop_fwd_drop2 … HLK) -HLK
+| #V1 lapply (drop_fwd_drop2 … HLK) -HLK
elim (lift_total V1 0 (i+1)) /4 width=12 by cpx_lift, cpxs_strap1/
]
qed.
elim (cpx_inv_lref1 … HT2) -HT2 /2 width=1 by or_introl/
* /4 width=7 by cpx_cpxs, ex3_4_intro, or_intror/
| * #I #K #V1 #T1 #HLK #HVT1 #HT1
- lapply (ldrop_fwd_drop2 … HLK) #H0LK
+ lapply (drop_fwd_drop2 … HLK) #H0LK
elim (cpx_inv_lift1 … HT2 … H0LK … HT1) -H0LK -T
/4 width=7 by cpxs_strap1, ex3_4_intro, or_intror/
]
(* Forward lemmas involving same top term constructor ***********************)
-lemma cpxs_fwd_cnx: â\88\80h,g,G,L,T. â¦\83G, Lâ¦\84 â\8a¢ â\9e¡[h, g] ð\9d\90\8dâ¦\83Tâ¦\84 â\86\92 â\88\80U. â¦\83G, Lâ¦\84 â\8a¢ T â\9e¡*[h, g] U â\86\92 T â\89\83 U.
+lemma cpxs_fwd_cnx: â\88\80h,g,G,L,T. â¦\83G, Lâ¦\84 â\8a¢ â\9e¡[h, g] ð\9d\90\8dâ¦\83Tâ¦\84 â\86\92 â\88\80U. â¦\83G, Lâ¦\84 â\8a¢ T â\9e¡*[h, g] U â\86\92 T â\89\82 U.
#h #g #G #L #T #HT #U #H
>(cpxs_inv_cnx1 … H HT) -G -L -T //
qed-.
lemma cpxs_fwd_sort: ∀h,g,G,L,U,k. ⦃G, L⦄ ⊢ ⋆k ➡*[h, g] U →
- â\8b\86k â\89\83 U ∨ ⦃G, L⦄ ⊢ ⋆(next h k) ➡*[h, g] U.
+ â\8b\86k â\89\82 U ∨ ⦃G, L⦄ ⊢ ⋆(next h k) ➡*[h, g] U.
#h #g #G #L #U #k #H
elim (cpxs_inv_sort1 … H) -H #n #l generalize in match k; -k @(nat_ind_plus … n) -n
[ #k #_ #H -l destruct /2 width=1 by or_introl/
(* Basic_1: was just: pr3_iso_beta *)
lemma cpxs_fwd_beta: ∀h,g,a,G,L,V,W,T,U. ⦃G, L⦄ ⊢ ⓐV.ⓛ{a}W.T ➡*[h, g] U →
- â\93\90V.â\93\9b{a}W.T â\89\83 U ∨ ⦃G, L⦄ ⊢ ⓓ{a}ⓝW.V.T ➡*[h, g] U.
+ â\93\90V.â\93\9b{a}W.T â\89\82 U ∨ ⦃G, L⦄ ⊢ ⓓ{a}ⓝW.V.T ➡*[h, g] U.
#h #g #a #G #L #V #W #T #U #H
elim (cpxs_inv_appl1 … H) -H *
[ #V0 #T0 #_ #_ #H destruct /2 width=1 by tstc_pair, or_introl/
lemma cpxs_fwd_delta: ∀h,g,I,G,L,K,V1,i. ⇩[i] L ≡ K.ⓑ{I}V1 →
∀V2. ⇧[0, i + 1] V1 ≡ V2 →
∀U. ⦃G, L⦄ ⊢ #i ➡*[h, g] U →
- #i â\89\83 U ∨ ⦃G, L⦄ ⊢ V2 ➡*[h, g] U.
+ #i â\89\82 U ∨ ⦃G, L⦄ ⊢ V2 ➡*[h, g] U.
#h #g #I #G #L #K #V1 #i #HLK #V2 #HV12 #U #H
elim (cpxs_inv_lref1 … H) -H /2 width=1 by or_introl/
* #I0 #K0 #V0 #U0 #HLK0 #HVU0 #HU0
-lapply (ldrop_mono … HLK0 … HLK) -HLK0 #H destruct
-/4 width=10 by cpxs_lift, ldrop_fwd_drop2, or_intror/
+lapply (drop_mono … HLK0 … HLK) -HLK0 #H destruct
+/4 width=10 by cpxs_lift, drop_fwd_drop2, or_intror/
qed-.
lemma cpxs_fwd_theta: ∀h,g,a,G,L,V1,V,T,U. ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}V.T ➡*[h, g] U →
- â\88\80V2. â\87§[0, 1] V1 â\89¡ V2 â\86\92 â\93\90V1.â\93\93{a}V.T â\89\83 U ∨
+ â\88\80V2. â\87§[0, 1] V1 â\89¡ V2 â\86\92 â\93\90V1.â\93\93{a}V.T â\89\82 U ∨
⦃G, L⦄ ⊢ ⓓ{a}V.ⓐV2.T ➡*[h, g] U.
#h #g #a #G #L #V1 #V #T #U #H #V2 #HV12
elim (cpxs_inv_appl1 … H) -H *
elim (cpxs_inv_abbr1 … HT0) -HT0 *
[ #V5 #T5 #HV5 #HT5 #H destruct
lapply (cpxs_lift … HV13 (L.ⓓV) … HV12 … HV34) -V1 -V3
- /3 width=2 by cpxs_flat, cpxs_bind, ldrop_drop/
+ /3 width=2 by cpxs_flat, cpxs_bind, drop_drop/
| #X #HT1 #H #H0 destruct
elim (lift_inv_bind1 … H) -H #V5 #T5 #HV05 #HT05 #H destruct
- lapply (cpxs_lift … HV13 (L.ⓓV0) … HV12 … HV34) -V3 /2 width=2 by ldrop_drop/ #HV24
+ lapply (cpxs_lift … HV13 (L.ⓓV0) … HV12 … HV34) -V3 /2 width=2 by drop_drop/ #HV24
@(cpxs_trans … (+ⓓV.ⓐV2.ⓓ{b}V5.T5)) [ /3 width=1 by cpxs_flat_dx, cpxs_bind_dx/ ] -T
@(cpxs_strap2 … (ⓐV1.ⓓ{b}V0.T0)) [ /4 width=7 by cpx_zeta, lift_bind, lift_flat/ ] -V -V5 -T5
@(cpxs_strap2 … (ⓓ{b}V0.ⓐV2.T0)) /3 width=3 by cpxs_pair_sn, cpxs_bind_dx, cpr_cpx, cpr_theta/
qed-.
lemma cpxs_fwd_cast: ∀h,g,G,L,W,T,U. ⦃G, L⦄ ⊢ ⓝW.T ➡*[h, g] U →
- â\88¨â\88¨ â\93\9dW. T â\89\83 U | ⦃G, L⦄ ⊢ T ➡*[h, g] U | ⦃G, L⦄ ⊢ W ➡*[h, g] U.
+ â\88¨â\88¨ â\93\9dW. T â\89\82 U | ⦃G, L⦄ ⊢ T ➡*[h, g] U | ⦃G, L⦄ ⊢ W ➡*[h, g] U.
#h #g #G #L #W #T #U #H
elim (cpxs_inv_cast1 … H) -H /2 width=1 by or3_intro1, or3_intro2/ *
#W0 #T0 #_ #_ #H destruct /2 width=1 by tstc_pair, or3_intro0/
(* Basic_1: was just: nf2_iso_appls_lref *)
lemma cpxs_fwd_cnx_vector: ∀h,g,G,L,T. 𝐒⦃T⦄ → ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃T⦄ →
- â\88\80Vs,U. â¦\83G, Lâ¦\84 â\8a¢ â\92¶Vs.T â\9e¡*[h, g] U â\86\92 â\92¶Vs.T â\89\83 U.
+ â\88\80Vs,U. â¦\83G, Lâ¦\84 â\8a¢ â\92¶Vs.T â\9e¡*[h, g] U â\86\92 â\92¶Vs.T â\89\82 U.
#h #g #G #L #T #H1T #H2T #Vs elim Vs -Vs [ @(cpxs_fwd_cnx … H2T) ] (**) (* /2 width=3 by cpxs_fwd_cnx/ does not work *)
#V #Vs #IHVs #U #H
elim (cpxs_inv_appl1 … H) -H *
qed-.
lemma cpxs_fwd_sort_vector: ∀h,g,G,L,k,Vs,U. ⦃G, L⦄ ⊢ ⒶVs.⋆k ➡*[h, g] U →
- â\92¶Vs.â\8b\86k â\89\83 U ∨ ⦃G, L⦄ ⊢ ⒶVs.⋆(next h k) ➡*[h, g] U.
+ â\92¶Vs.â\8b\86k â\89\82 U ∨ ⦃G, L⦄ ⊢ ⒶVs.⋆(next h k) ➡*[h, g] U.
#h #g #G #L #k #Vs elim Vs -Vs /2 width=1 by cpxs_fwd_sort/
#V #Vs #IHVs #U #H
elim (cpxs_inv_appl1 … H) -H *
(* Basic_1: was just: pr3_iso_appls_beta *)
lemma cpxs_fwd_beta_vector: ∀h,g,a,G,L,Vs,V,W,T,U. ⦃G, L⦄ ⊢ ⒶVs.ⓐV.ⓛ{a}W.T ➡*[h, g] U →
- â\92¶Vs. â\93\90V. â\93\9b{a}W. T â\89\83 U ∨ ⦃G, L⦄ ⊢ ⒶVs.ⓓ{a}ⓝW.V.T ➡*[h, g] U.
+ â\92¶Vs. â\93\90V. â\93\9b{a}W. T â\89\82 U ∨ ⦃G, L⦄ ⊢ ⒶVs.ⓓ{a}ⓝW.V.T ➡*[h, g] U.
#h #g #a #G #L #Vs elim Vs -Vs /2 width=1 by cpxs_fwd_beta/
#V0 #Vs #IHVs #V #W #T #U #H
elim (cpxs_inv_appl1 … H) -H *
lemma cpxs_fwd_delta_vector: ∀h,g,I,G,L,K,V1,i. ⇩[i] L ≡ K.ⓑ{I}V1 →
∀V2. ⇧[0, i + 1] V1 ≡ V2 →
∀Vs,U. ⦃G, L⦄ ⊢ ⒶVs.#i ➡*[h, g] U →
- â\92¶Vs.#i â\89\83 U ∨ ⦃G, L⦄ ⊢ ⒶVs.V2 ➡*[h, g] U.
+ â\92¶Vs.#i â\89\82 U ∨ ⦃G, L⦄ ⊢ ⒶVs.V2 ➡*[h, g] U.
#h #g #I #G #L #K #V1 #i #HLK #V2 #HV12 #Vs elim Vs -Vs /2 width=5 by cpxs_fwd_delta/
#V #Vs #IHVs #U #H -K -V1
elim (cpxs_inv_appl1 … H) -H *
(* Basic_1: was just: pr3_iso_appls_abbr *)
lemma cpxs_fwd_theta_vector: ∀h,g,G,L,V1s,V2s. ⇧[0, 1] V1s ≡ V2s →
∀a,V,T,U. ⦃G, L⦄ ⊢ ⒶV1s.ⓓ{a}V.T ➡*[h, g] U →
- â\92¶V1s. â\93\93{a}V. T â\89\83 U ∨ ⦃G, L⦄ ⊢ ⓓ{a}V.ⒶV2s.T ➡*[h, g] U.
+ â\92¶V1s. â\93\93{a}V. T â\89\82 U ∨ ⦃G, L⦄ ⊢ ⓓ{a}V.ⒶV2s.T ➡*[h, g] U.
#h #g #G #L #V1s #V2s * -V1s -V2s /3 width=1 by or_intror/
#V1s #V2s #V1a #V2a #HV12a #HV12s #a
generalize in match HV12a; -HV12a
@(cpxs_trans … HU) -U
elim (cpxs_inv_abbr1 … HT0) -HT0 *
[ #V1 #T1 #HV1 #HT1 #H destruct
- lapply (cpxs_lift … HV10a (L.ⓓV) (Ⓣ) … HV12a … HV0a) -V1a -V0a [ /2 width=1 by ldrop_drop/ ] #HV2a
+ lapply (cpxs_lift … HV10a (L.ⓓV) (Ⓣ) … HV12a … HV0a) -V1a -V0a [ /2 width=1 by drop_drop/ ] #HV2a
@(cpxs_trans … (ⓓ{a}V.ⓐV2a.T1)) /3 width=1 by cpxs_bind, cpxs_pair_sn, cpxs_flat_dx, cpxs_bind_dx/
| #X #HT1 #H #H0 destruct
elim (lift_inv_bind1 … H) -H #V1 #T1 #HW01 #HT01 #H destruct
- lapply (cpxs_lift … HV10a (L.ⓓV0) (Ⓣ) … HV12a … HV0a) -V0a [ /2 width=1 by ldrop_drop/ ] #HV2a
+ lapply (cpxs_lift … HV10a (L.ⓓV0) (Ⓣ) … HV12a … HV0a) -V0a [ /2 width=1 by drop_drop/ ] #HV2a
@(cpxs_trans … (+ⓓV.ⓐV2a.ⓓ{b}V1.T1)) [ /3 width=1 by cpxs_flat_dx, cpxs_bind_dx/ ] -T -V2b -V2s
@(cpxs_strap2 … (ⓐV1a.ⓓ{b}V0.T0)) [ /4 width=7 by cpx_zeta, lift_bind, lift_flat/ ] -V -V1 -T1
@(cpxs_strap2 … (ⓓ{b}V0.ⓐV2a.T0)) /3 width=3 by cpxs_pair_sn, cpxs_bind_dx, cpr_cpx, cpr_theta/
(* Basic_1: was just: pr3_iso_appls_cast *)
lemma cpxs_fwd_cast_vector: ∀h,g,G,L,Vs,W,T,U. ⦃G, L⦄ ⊢ ⒶVs.ⓝW.T ➡*[h, g] U →
- â\88¨â\88¨ â\92¶Vs. â\93\9dW. T â\89\83 U
+ â\88¨â\88¨ â\92¶Vs. â\93\9dW. T â\89\82 U
| ⦃G, L⦄ ⊢ ⒶVs.T ➡*[h, g] U
| ⦃G, L⦄ ⊢ ⒶVs.W ➡*[h, g] U.
#h #g #G #L #Vs elim Vs -Vs /2 width=1 by cpxs_fwd_cast/
⦃G, L⦄ ⊢ ⬊*[h, g] #i → ⦃G, K⦄ ⊢ ⬊*[h, g] V.
#h #g #I #G #L #K #V #i #HLK #Hi
elim (lift_total V 0 (i+1))
-/4 width=9 by csx_inv_lift, csx_cpx_trans, cpx_delta, ldrop_fwd_drop2/
+/4 width=9 by csx_inv_lift, csx_cpx_trans, cpx_delta, drop_fwd_drop2/
qed-.
(* Advanced properties ******************************************************)
elim (cpx_inv_lref1 … H) -H
[ #H destruct elim Hi //
| -Hi * #I0 #K0 #V0 #V1 #HLK0 #HV01 #HV1
- lapply (ldrop_mono … HLK0 … HLK) -HLK #H destruct
- /3 width=8 by csx_lift, csx_cpx_trans, ldrop_fwd_drop2/
+ lapply (drop_mono … HLK0 … HLK) -HLK #H destruct
+ /3 width=8 by csx_lift, csx_cpx_trans, drop_fwd_drop2/
]
qed.
| -IHV -HLV1 * #H destruct /3 width=1 by cpx_cpxs/
]
| -IHV -IHT -H2
- /3 width=8 by csx_cpx_trans, csx_inv_lift, ldrop_drop/
+ /3 width=8 by csx_cpx_trans, csx_inv_lift, drop_drop/
]
qed.
| * #_ #H elim H //
]
| -H -HVT #H
- lapply (cpx_lift … HLV10 (L.ⓓV) (Ⓣ) … HV12 … HV04) -HLV10 -HV12 /2 width=1 by ldrop_drop/ #HV24
+ lapply (cpx_lift … HLV10 (L.ⓓV) (Ⓣ) … HV12 … HV04) -HLV10 -HV12 /2 width=1 by drop_drop/ #HV24
@(IHVT … H … HV04) -IHVT /4 width=1 by cpx_cpxs, cpx_bind, cpx_flat/
]
| -H -IHVT #T0 #HLT0 #HT0 #H0 destruct
lapply (csx_cpx_trans … HVT (ⓐV2.T0) ?) /2 width=1 by cpx_flat/ -T #HVT0
lapply (csx_inv_lift … L … (Ⓣ) … 1 HVT0 ? ? ?) -HVT0
- /3 width=5 by csx_cpx_trans, cpx_pair_sn, ldrop_drop, lift_flat/
+ /3 width=5 by csx_cpx_trans, cpx_pair_sn, drop_drop, lift_flat/
]
| -HV -HV12 -HVT -IHVT -H #b #V0 #W0 #W1 #T0 #T1 #_ #_ #_ #H destruct
| -IHVT -H #b #V0 #V3 #W0 #W1 #T0 #T1 #HLV10 #HV03 #HLW01 #HLT01 #H1 #H2 destruct
- lapply (cpx_lift … HLV10 (L. ⓓW0) … HV12 … HV03) -HLV10 -HV12 -HV03 /2 width=2 by ldrop_drop/ #HLV23
+ lapply (cpx_lift … HLV10 (L. ⓓW0) … HV12 … HV03) -HLV10 -HV12 -HV03 /2 width=2 by drop_drop/ #HLV23
@csx_abbr /2 width=3 by csx_cpx_trans/ -HV
@(csx_lpx_conf … (L.ⓓW0)) /2 width=1 by lpx_pair/ -W1
/4 width=5 by csx_cpxs_trans, cpx_cpxs, cpx_flat/
(* Basic_1: was just: sn3_appl_appl *)
lemma csx_appl_simple_tstc: ∀h,g,G,L,V. ⦃G, L⦄ ⊢ ⬊*[h, g] V → ∀T1. ⦃G, L⦄ ⊢ ⬊*[h, g] T1 →
- (â\88\80T2. â¦\83G, Lâ¦\84 â\8a¢ T1 â\9e¡*[h, g] T2 â\86\92 (T1 â\89\83 T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊*[h, g] ⓐV.T2) →
+ (â\88\80T2. â¦\83G, Lâ¦\84 â\8a¢ T1 â\9e¡*[h, g] T2 â\86\92 (T1 â\89\82 T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊*[h, g] ⓐV.T2) →
𝐒⦃T1⦄ → ⦃G, L⦄ ⊢ ⬊*[h, g] ⓐV.T1.
#h #g #G #L #V #H @(csx_ind … H) -V #V #_ #IHV #T1 #H @(csx_ind … H) -T1 #T1 #H1T1 #IHT1 #H2T1 #H3T1
@csx_intro #X #HL #H
∀V2. ⇧[0, i + 1] V1 ≡ V2 →
∀Vs. ⦃G, L⦄ ⊢ ⬊*[h, g] (ⒶVs.V2) → ⦃G, L⦄ ⊢ ⬊*[h, g] (ⒶVs.#i).
#h #g #I #G #L #K #V1 #i #HLK #V2 #HV12 #Vs elim Vs -Vs
-[ /4 width=12 by csx_inv_lift, csx_lref_bind, ldrop_fwd_drop2/
+[ /4 width=12 by csx_inv_lift, csx_lref_bind, drop_fwd_drop2/
| #V #Vs #IHV #H1T
lapply (csx_fwd_pair_sn … H1T) #HV
lapply (csx_fwd_flat_dx … H1T) #H2T
#h #g #G #L elim L /2 width=1 by lcosx_skip/
qed.
-lemma lcosx_ldrop_trans_lt: ∀h,g,G,L,d. G ⊢ ~⬊*[h, g, d] L →
+lemma lcosx_drop_trans_lt: ∀h,g,G,L,d. G ⊢ ~⬊*[h, g, d] L →
∀I,K,V,i. ⇩[i] L ≡ K.ⓑ{I}V → i < d →
G ⊢ ~⬊*[h, g, ⫰(d-i)] K ∧ G ⊢ ⬊*[h, g, V, ⫰(d-i)] K.
#h #g #G #L #d #H elim H -L -d
-[ #d #J #K #V #i #H elim (ldrop_inv_atom1 … H) -H #H destruct
+[ #d #J #K #V #i #H elim (drop_inv_atom1 … H) -H #H destruct
| #I #L #T #_ #_ #J #K #V #i #_ #H elim (ylt_yle_false … H) -H //
| #I #L #T #d #HT #HL #IHL #J #K #V #i #H #Hid
- elim (ldrop_inv_O1_pair1 … H) -H * #Hi #HLK destruct
+ elim (drop_inv_O1_pair1 … H) -H * #Hi #HLK destruct
[ >ypred_succ /2 width=1 by conj/
| lapply (ylt_pred … Hid ?) -Hid /2 width=1 by ylt_inj/ >ypred_succ #Hid
elim (IHL … HLK ?) -IHL -HLK <yminus_inj >yminus_SO2 //
(**************************************************************************)
include "ground_2/ynat/ynat_max.ma".
-include "basic_2/computation/lsx_ldrop.ma".
+include "basic_2/computation/lsx_drop.ma".
include "basic_2/computation/lsx_lpx.ma".
include "basic_2/computation/lsx_lpxs.ma".
include "basic_2/computation/lcosx.ma".
[ #I #G #L #K #V1 #V2 #W2 #i #HLK #_ #HVW2 #IHV12 #d #HL #H
elim (ylt_split i d) #Hdi [ -H | -HL ]
[ <(ymax_pre_sn d (⫯i)) /2 width=1 by ylt_fwd_le_succ/
- elim (lcosx_ldrop_trans_lt … HL … HLK) // -HL -Hdi
- lapply (ldrop_fwd_drop2 … HLK) -HLK /3 width=7 by lsx_lift_ge/
+ elim (lcosx_drop_trans_lt … HL … HLK) // -HL -Hdi
+ lapply (drop_fwd_drop2 … HLK) -HLK /3 width=7 by lsx_lift_ge/
| lapply (lsx_fwd_lref_be … H … HLK) // -H -Hdi
- lapply (ldrop_fwd_drop2 … HLK) -HLK
+ lapply (drop_fwd_drop2 … HLK) -HLK
/4 width=10 by lsx_ge, lsx_lift_le/
]
| #a #I #G #L #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #d #HL #H
elim (lsx_inv_flat … H) -H /3 width=1 by lsx_flat/
| #G #L #V #U1 #U2 #T2 #_ #HTU2 #IHU12 #d #HL #H
elim (lsx_inv_bind … H) -H
- /4 width=9 by lcosx_pair, lsx_inv_lift_ge, ldrop_drop/
+ /4 width=9 by lcosx_pair, lsx_inv_lift_ge, drop_drop/
| #G #L #V #T1 #T2 #_ #IHT12 #d #HL #H
elim (lsx_inv_flat … H) -H /2 width=1 by/
| #G #L #V1 #V2 #T #_ #IHV12 #d #HL #H
elim (lsx_inv_flat … H) -H #HV1 #H
elim (lsx_inv_bind … H) -H #HW1 #HT1
@lsx_bind /2 width=1 by/ (**) (* explicit constructor *)
- @lsx_flat [ /3 width=7 by lsx_lift_ge, ldrop_drop/ ]
+ @lsx_flat [ /3 width=7 by lsx_lift_ge, drop_drop/ ]
@(lsx_leq_conf … (L.ⓓW1)) /3 width=1 by lcosx_pair, leq_succ/
]
qed-.
]
| #U1 #HTU1 #HU01 elim (lift_total U2 0 1)
#U #HU2 lapply (cpr_lift … HU02 (L.ⓓV1) … HU01 … HU2) -U0
- /4 width=3 by cprs_strap1, ldrop_drop, ex3_intro, or_intror/
+ /4 width=3 by cprs_strap1, drop_drop, ex3_intro, or_intror/
]
qed-.
--- /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/reduction/lpr_drop.ma".
+include "basic_2/computation/lprs.ma".
+
+(* SN PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS ****************************)
+
+(* Properies on local environment slicing ***********************************)
+
+lemma lprs_drop_conf: ∀G. dropable_sn (lprs G).
+/3 width=3 by dropable_sn_TC, lpr_drop_conf/ qed-.
+
+lemma drop_lprs_trans: ∀G. dedropable_sn (lprs G).
+/3 width=3 by dedropable_sn_TC, drop_lpr_trans/ qed-.
+
+lemma lprs_drop_trans_O1: ∀G. dropable_dx (lprs G).
+/3 width=3 by dropable_dx_TC, lpr_drop_trans_O1/ qed-.
+++ /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/reduction/lpr_ldrop.ma".
-include "basic_2/computation/lprs.ma".
-
-(* SN PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS ****************************)
-
-(* Properies on local environment slicing ***********************************)
-
-lemma lprs_ldrop_conf: ∀G. dropable_sn (lprs G).
-/3 width=3 by dropable_sn_TC, lpr_ldrop_conf/ qed-.
-
-lemma ldrop_lprs_trans: ∀G. dedropable_sn (lprs G).
-/3 width=3 by dedropable_sn_TC, ldrop_lpr_trans/ qed-.
-
-lemma lprs_ldrop_trans_O1: ∀G. dropable_dx (lprs G).
-/3 width=3 by dropable_dx_TC, lpr_ldrop_trans_O1/ qed-.
]
| #U1 #HTU1 #HU01
elim (lift_total U2 0 1) #U #HU2
- /6 width=12 by cpxs_strap1, cpx_lift, ldrop_drop, ex3_intro, or_intror/
+ /6 width=12 by cpxs_strap1, cpx_lift, drop_drop, ex3_intro, or_intror/
]
qed-.
--- /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/reduction/lpx_drop.ma".
+include "basic_2/computation/lpxs.ma".
+
+(* SN EXTENDED PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS *******************)
+
+(* Properies on local environment slicing ***********************************)
+
+lemma lpxs_drop_conf: ∀h,g,G. dropable_sn (lpxs h g G).
+/3 width=3 by dropable_sn_TC, lpx_drop_conf/ qed-.
+
+lemma drop_lpxs_trans: ∀h,g,G. dedropable_sn (lpxs h g G).
+/3 width=3 by dedropable_sn_TC, drop_lpx_trans/ qed-.
+
+lemma lpxs_drop_trans_O1: ∀h,g,G. dropable_dx (lpxs h g G).
+/3 width=3 by dropable_dx_TC, lpx_drop_trans_O1/ qed-.
+++ /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/reduction/lpx_ldrop.ma".
-include "basic_2/computation/lpxs.ma".
-
-(* SN EXTENDED PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS *******************)
-
-(* Properies on local environment slicing ***********************************)
-
-lemma lpxs_ldrop_conf: ∀h,g,G. dropable_sn (lpxs h g G).
-/3 width=3 by dropable_sn_TC, lpx_ldrop_conf/ qed-.
-
-lemma ldrop_lpxs_trans: ∀h,g,G. dedropable_sn (lpxs h g G).
-/3 width=3 by dedropable_sn_TC, ldrop_lpx_trans/ qed-.
-
-lemma lpxs_ldrop_trans_O1: ∀h,g,G. dropable_dx (lpxs h g G).
-/3 width=3 by dropable_dx_TC, lpx_ldrop_trans_O1/ qed-.
include "basic_2/reduction/lpx_lleq.ma".
include "basic_2/computation/cpxs_leq.ma".
-include "basic_2/computation/lpxs_ldrop.ma".
+include "basic_2/computation/lpxs_drop.ma".
include "basic_2/computation/lpxs_cpxs.ma".
(* SN EXTENDED PARALLEL COMPUTATION FOR LOCAL ENVIRONMENTS ******************)
[ #I #G1 #L1 #V1 #X #H1 #H2 elim (lpxs_inv_pair2 … H1) -H1
#K0 #V0 #H1KL1 #_ #H destruct
elim (lleq_inv_lref_ge_dx … H2 ? I L1 V1) -H2 //
- #K1 #H #H2KL1 lapply (ldrop_inv_O2 … H) -H #H destruct
+ #K1 #H #H2KL1 lapply (drop_inv_O2 … H) -H #H destruct
/2 width=4 by fqu_lref_O, ex3_intro/
| * [ #a ] #I #G1 #L1 #V1 #T1 #K1 #HLK1 #H
[ elim (lleq_inv_bind … H)
| #I #G1 #L1 #V1 #T1 #K1 #HLK1 #H elim (lleq_inv_flat … H) -H
/2 width=4 by fqu_flat_dx, ex3_intro/
| #G1 #L1 #L #T1 #U1 #e #HL1 #HTU1 #K1 #H1KL1 #H2KL1
- elim (ldrop_O1_le (Ⓕ) (e+1) K1)
+ elim (drop_O1_le (Ⓕ) (e+1) K1)
[ #K #HK1 lapply (lleq_inv_lift_le … H2KL1 … HK1 HL1 … HTU1 ?) -H2KL1 //
- #H2KL elim (lpxs_ldrop_trans_O1 … H1KL1 … HL1) -L1
- #K0 #HK10 #H1KL lapply (ldrop_mono … HK10 … HK1) -HK10 #H destruct
+ #H2KL elim (lpxs_drop_trans_O1 … H1KL1 … HL1) -L1
+ #K0 #HK10 #H1KL lapply (drop_mono … HK10 … HK1) -HK10 #H destruct
/3 width=4 by fqu_drop, ex3_intro/
- | lapply (ldrop_fwd_length_le2 … HL1) -L -T1 -g
+ | lapply (drop_fwd_length_le2 … HL1) -L -T1 -g
lapply (lleq_fwd_length … H2KL1) //
]
]
]
qed-.
-fact leq_lpxs_trans_lleq_aux: â\88\80h,g,G,L1,L0,d,e. L1 â\89\83[d, e] L0 → e = ∞ →
+fact leq_lpxs_trans_lleq_aux: â\88\80h,g,G,L1,L0,d,e. L1 ⩬[d, e] L0 → e = ∞ →
∀L2. ⦃G, L0⦄ ⊢ ➡*[h, g] L2 →
- â\88\83â\88\83L. L â\89\83[d, e] L2 & ⦃G, L1⦄ ⊢ ➡*[h, g] L &
+ â\88\83â\88\83L. L ⩬[d, e] L2 & ⦃G, L1⦄ ⊢ ➡*[h, g] L &
(∀T. L0 ≡[T, d] L2 ↔ L1 ≡[T, d] L).
#h #g #G #L1 #L0 #d #e #H elim H -L1 -L0 -d -e
[ #d #e #_ #L2 #H >(lpxs_inv_atom1 … H) -H
]
qed-.
-lemma leq_lpxs_trans_lleq: â\88\80h,g,G,L1,L0,d. L1 â\89\83[d, ∞] L0 →
+lemma leq_lpxs_trans_lleq: â\88\80h,g,G,L1,L0,d. L1 ⩬[d, ∞] L0 →
∀L2. ⦃G, L0⦄ ⊢ ➡*[h, g] L2 →
- â\88\83â\88\83L. L â\89\83[d, ∞] L2 & ⦃G, L1⦄ ⊢ ➡*[h, g] L &
+ â\88\83â\88\83L. L ⩬[d, ∞] L2 & ⦃G, L1⦄ ⊢ ➡*[h, g] L &
(∀T. L0 ≡[T, d] L2 ↔ L1 ≡[T, d] L).
/2 width=1 by leq_lpxs_trans_lleq_aux/ qed-.
--- /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/static/aaa_lift.ma".
+include "basic_2/computation/lsubc.ma".
+
+(* LOCAL ENVIRONMENT REFINEMENT FOR ABSTRACT CANDIDATES OF REDUCIBILITY *****)
+
+(* Properties concerning basic local environment slicing ********************)
+
+(* Basic_1: was: csubc_drop_conf_O *)
+(* Note: the constant 0 can not be generalized *)
+lemma lsubc_drop_O1_trans: ∀RP,G,L1,L2. G ⊢ L1 ⫃[RP] L2 → ∀K2,s,e. ⇩[s, 0, e] L2 ≡ K2 →
+ ∃∃K1. ⇩[s, 0, e] L1 ≡ K1 & G ⊢ K1 ⫃[RP] K2.
+#RP #G #L1 #L2 #H elim H -L1 -L2
+[ #X #s #e #H elim (drop_inv_atom1 … H) -H /4 width=3 by drop_atom, ex2_intro/
+| #I #L1 #L2 #V #_ #IHL12 #X #s #e #H
+ elim (drop_inv_O1_pair1 … H) -H * #He #H destruct
+ [ elim (IHL12 L2 s 0) -IHL12 // #X #H <(drop_inv_O2 … H) -H
+ /3 width=3 by lsubc_pair, drop_pair, ex2_intro/
+ | elim (IHL12 … H) -L2 /3 width=3 by drop_drop_lt, ex2_intro/
+ ]
+| #L1 #L2 #V #W #A #HV #H1W #H2W #_ #IHL12 #X #s #e #H
+ elim (drop_inv_O1_pair1 … H) -H * #He #H destruct
+ [ elim (IHL12 L2 s 0) -IHL12 // #X #H <(drop_inv_O2 … H) -H
+ /3 width=8 by lsubc_abbr, drop_pair, ex2_intro/
+ | elim (IHL12 … H) -L2 /3 width=3 by drop_drop_lt, ex2_intro/
+ ]
+]
+qed-.
+
+(* Basic_1: was: csubc_drop_conf_rev *)
+lemma drop_lsubc_trans: ∀RR,RS,RP.
+ acp RR RS RP → acr RR RS RP (λG,L,T. RP G L T) →
+ ∀G,L1,K1,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ∀K2. G ⊢ K1 ⫃[RP] K2 →
+ ∃∃L2. G ⊢ L1 ⫃[RP] L2 & ⇩[Ⓕ, d, e] L2 ≡ K2.
+#RR #RS #RP #Hacp #Hacr #G #L1 #K1 #d #e #H elim H -L1 -K1 -d -e
+[ #d #e #He #X #H elim (lsubc_inv_atom1 … H) -H
+ >He /2 width=3 by ex2_intro/
+| #L1 #I #V1 #X #H
+ elim (lsubc_inv_pair1 … H) -H *
+ [ #K1 #HLK1 #H destruct /3 width=3 by lsubc_pair, drop_pair, ex2_intro/
+ | #K1 #V #W1 #A #HV1 #H1W1 #H2W1 #HLK1 #H1 #H2 #H3 destruct
+ /3 width=4 by lsubc_abbr, drop_pair, ex2_intro/
+ ]
+| #I #L1 #K1 #V1 #e #_ #IHLK1 #K2 #HK12
+ elim (IHLK1 … HK12) -K1 /3 width=5 by lsubc_pair, drop_drop, ex2_intro/
+| #I #L1 #K1 #V1 #V2 #d #e #HLK1 #HV21 #IHLK1 #X #H
+ elim (lsubc_inv_pair1 … H) -H *
+ [ #K2 #HK12 #H destruct
+ elim (IHLK1 … HK12) -K1 /3 width=5 by lsubc_pair, drop_skip, ex2_intro/
+ | #K2 #V #W2 #A #HV2 #H1W2 #H2W2 #HK12 #H1 #H2 #H3 destruct
+ elim (lift_inv_flat1 … HV21) -HV21 #W3 #V3 #HW23 #HV3 #H destruct
+ elim (IHLK1 … HK12) #K #HL1K #HK2
+ lapply (aacr_acr … Hacp Hacr A) -Hacp -Hacr #HA
+ lapply (s0 … HA … HV2 … HLK1 HV3) -HV2
+ lapply (s0 … HA … H1W2 … HLK1 HW23) -H1W2
+ /4 width=11 by lsubc_abbr, aaa_lift, drop_skip, ex2_intro/
+ ]
+]
+qed-.
--- /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/computation/lsubc_drop.ma".
+
+(* LOCAL ENVIRONMENT REFINEMENT FOR ABSTRACT CANDIDATES OF REDUCIBILITY *****)
+
+(* Properties concerning generic local environment slicing ******************)
+
+(* Basic_1: was: csubc_drop1_conf_rev *)
+lemma drops_lsubc_trans: ∀RR,RS,RP.
+ acp RR RS RP → acr RR RS RP (λG,L,T. RP G L T) →
+ ∀G,L1,K1,des. ⇩*[Ⓕ, des] L1 ≡ K1 → ∀K2. G ⊢ K1 ⫃[RP] K2 →
+ ∃∃L2. G ⊢ L1 ⫃[RP] L2 & ⇩*[Ⓕ, des] L2 ≡ K2.
+#RR #RS #RP #Hacp #Hacr #G #L1 #K1 #des #H elim H -L1 -K1 -des
+[ /2 width=3 by drops_nil, ex2_intro/
+| #L1 #L #K1 #des #d #e #_ #HLK1 #IHL #K2 #HK12
+ elim (drop_lsubc_trans … Hacp Hacr … HLK1 … HK12) -Hacp -Hacr -K1 #K #HLK #HK2
+ elim (IHL … HLK) -IHL -HLK /3 width=5 by drops_cons, ex2_intro/
+]
+qed-.
+++ /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/static/aaa_lift.ma".
-include "basic_2/computation/lsubc.ma".
-
-(* LOCAL ENVIRONMENT REFINEMENT FOR ABSTRACT CANDIDATES OF REDUCIBILITY *****)
-
-(* Properties concerning basic local environment slicing ********************)
-
-(* Basic_1: was: csubc_drop_conf_O *)
-(* Note: the constant 0 can not be generalized *)
-lemma lsubc_ldrop_O1_trans: ∀RP,G,L1,L2. G ⊢ L1 ⫃[RP] L2 → ∀K2,s,e. ⇩[s, 0, e] L2 ≡ K2 →
- ∃∃K1. ⇩[s, 0, e] L1 ≡ K1 & G ⊢ K1 ⫃[RP] K2.
-#RP #G #L1 #L2 #H elim H -L1 -L2
-[ #X #s #e #H elim (ldrop_inv_atom1 … H) -H /4 width=3 by ldrop_atom, ex2_intro/
-| #I #L1 #L2 #V #_ #IHL12 #X #s #e #H
- elim (ldrop_inv_O1_pair1 … H) -H * #He #H destruct
- [ elim (IHL12 L2 s 0) -IHL12 // #X #H <(ldrop_inv_O2 … H) -H
- /3 width=3 by lsubc_pair, ldrop_pair, ex2_intro/
- | elim (IHL12 … H) -L2 /3 width=3 by ldrop_drop_lt, ex2_intro/
- ]
-| #L1 #L2 #V #W #A #HV #H1W #H2W #_ #IHL12 #X #s #e #H
- elim (ldrop_inv_O1_pair1 … H) -H * #He #H destruct
- [ elim (IHL12 L2 s 0) -IHL12 // #X #H <(ldrop_inv_O2 … H) -H
- /3 width=8 by lsubc_abbr, ldrop_pair, ex2_intro/
- | elim (IHL12 … H) -L2 /3 width=3 by ldrop_drop_lt, ex2_intro/
- ]
-]
-qed-.
-
-(* Basic_1: was: csubc_drop_conf_rev *)
-lemma ldrop_lsubc_trans: ∀RR,RS,RP.
- acp RR RS RP → acr RR RS RP (λG,L,T. RP G L T) →
- ∀G,L1,K1,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ∀K2. G ⊢ K1 ⫃[RP] K2 →
- ∃∃L2. G ⊢ L1 ⫃[RP] L2 & ⇩[Ⓕ, d, e] L2 ≡ K2.
-#RR #RS #RP #Hacp #Hacr #G #L1 #K1 #d #e #H elim H -L1 -K1 -d -e
-[ #d #e #He #X #H elim (lsubc_inv_atom1 … H) -H
- >He /2 width=3 by ex2_intro/
-| #L1 #I #V1 #X #H
- elim (lsubc_inv_pair1 … H) -H *
- [ #K1 #HLK1 #H destruct /3 width=3 by lsubc_pair, ldrop_pair, ex2_intro/
- | #K1 #V #W1 #A #HV1 #H1W1 #H2W1 #HLK1 #H1 #H2 #H3 destruct
- /3 width=4 by lsubc_abbr, ldrop_pair, ex2_intro/
- ]
-| #I #L1 #K1 #V1 #e #_ #IHLK1 #K2 #HK12
- elim (IHLK1 … HK12) -K1 /3 width=5 by lsubc_pair, ldrop_drop, ex2_intro/
-| #I #L1 #K1 #V1 #V2 #d #e #HLK1 #HV21 #IHLK1 #X #H
- elim (lsubc_inv_pair1 … H) -H *
- [ #K2 #HK12 #H destruct
- elim (IHLK1 … HK12) -K1 /3 width=5 by lsubc_pair, ldrop_skip, ex2_intro/
- | #K2 #V #W2 #A #HV2 #H1W2 #H2W2 #HK12 #H1 #H2 #H3 destruct
- elim (lift_inv_flat1 … HV21) -HV21 #W3 #V3 #HW23 #HV3 #H destruct
- elim (IHLK1 … HK12) #K #HL1K #HK2
- lapply (aacr_acr … Hacp Hacr A) -Hacp -Hacr #HA
- lapply (s0 … HA … HV2 … HLK1 HV3) -HV2
- lapply (s0 … HA … H1W2 … HLK1 HW23) -H1W2
- /4 width=11 by lsubc_abbr, aaa_lift, ldrop_skip, ex2_intro/
- ]
-]
-qed-.
+++ /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/computation/lsubc_ldrop.ma".
-
-(* LOCAL ENVIRONMENT REFINEMENT FOR ABSTRACT CANDIDATES OF REDUCIBILITY *****)
-
-(* Properties concerning generic local environment slicing ******************)
-
-(* Basic_1: was: csubc_drop1_conf_rev *)
-lemma ldrops_lsubc_trans: ∀RR,RS,RP.
- acp RR RS RP → acr RR RS RP (λG,L,T. RP G L T) →
- ∀G,L1,K1,des. ⇩*[Ⓕ, des] L1 ≡ K1 → ∀K2. G ⊢ K1 ⫃[RP] K2 →
- ∃∃L2. G ⊢ L1 ⫃[RP] L2 & ⇩*[Ⓕ, des] L2 ≡ K2.
-#RR #RS #RP #Hacp #Hacr #G #L1 #K1 #des #H elim H -L1 -K1 -des
-[ /2 width=3 by ldrops_nil, ex2_intro/
-| #L1 #L #K1 #des #d #e #_ #HLK1 #IHL #K2 #HK12
- elim (ldrop_lsubc_trans … Hacp Hacr … HLK1 … HK12) -Hacp -Hacr -K1 #K #HLK #HK2
- elim (IHL … HLK) -IHL -HLK /3 width=5 by ldrops_cons, ex2_intro/
-]
-qed-.
#h #g #I #G #K1 #V #i #d #Hdi #H @(csx_ind_alt … H) -V
#V0 #_ #IHV0 #K2 #H @(lsx_ind … H) -K2
#K0 #HK0 #IHK0 #HK10 #L0 #HLK0 @lsx_intro
-#L2 #HL02 #HnL02 elim (lpx_ldrop_conf … HLK0 … HL02) -HL02
+#L2 #HL02 #HnL02 elim (lpx_drop_conf … HLK0 … HL02) -HL02
#Y #H #HLK2 elim (lpx_inv_pair1 … H) -H
#K2 #V2 #HK02 #HV02 #H destruct
elim (eq_term_dec V0 V2) #HnV02 destruct [ -IHV0 -HV02 -HK0 | -IHK0 -HnL02 -HLK0 ]
[ #i #HG #HL #HT #H #d destruct
elim (lt_or_ge i (|L|)) /2 width=1 by lsx_lref_free/
elim (ylt_split i d) /2 width=1 by lsx_lref_skip/
- #Hdi #Hi elim (ldrop_O1_lt (Ⓕ) … Hi) -Hi
+ #Hdi #Hi elim (drop_O1_lt (Ⓕ) … Hi) -Hi
#I #K #V #HLK lapply (csx_inv_lref_bind … HLK … H) -H
/4 width=6 by lsx_lref_be, fqup_lref/
| #a #I #V #T #HG #HL #HT #H #d destruct
--- /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/multiple/lleq_drop.ma".
+include "basic_2/reduction/lpx_drop.ma".
+include "basic_2/computation/lsx.ma".
+
+(* SN EXTENDED STRONGLY NORMALIZING LOCAL ENVIRONMENTS **********************)
+
+(* Advanced properties ******************************************************)
+
+lemma lsx_lref_free: ∀h,g,G,L,d,i. |L| ≤ i → G ⊢ ⬊*[h, g, #i, d] L.
+#h #g #G #L1 #d #i #HL1 @lsx_intro
+#L2 #HL12 #H elim H -H
+/4 width=6 by lpx_fwd_length, lleq_free, le_repl_sn_conf_aux/
+qed.
+
+lemma lsx_lref_skip: ∀h,g,G,L,d,i. yinj i < d → G ⊢ ⬊*[h, g, #i, d] L.
+#h #g #G #L1 #d #i #HL1 @lsx_intro
+#L2 #HL12 #H elim H -H
+/3 width=4 by lpx_fwd_length, lleq_skip/
+qed.
+
+(* Advanced forward lemmas **************************************************)
+
+lemma lsx_fwd_lref_be: ∀h,g,I,G,L,d,i. d ≤ yinj i → G ⊢ ⬊*[h, g, #i, d] L →
+ ∀K,V. ⇩[i] L ≡ K.ⓑ{I}V → G ⊢ ⬊*[h, g, V, 0] K.
+#h #g #I #G #L #d #i #Hdi #H @(lsx_ind … H) -L
+#L1 #_ #IHL1 #K1 #V #HLK1 @lsx_intro
+#K2 #HK12 #HnK12 lapply (drop_fwd_drop2 … HLK1)
+#H2LK1 elim (drop_lpx_trans … H2LK1 … HK12) -H2LK1 -HK12
+#L2 #HL12 #H2LK2 #H elim (leq_drop_conf_be … H … HLK1) -H /2 width=1 by ylt_inj/
+#Y #_ #HLK2 lapply (drop_fwd_drop2 … HLK2)
+#HY lapply (drop_mono … HY … H2LK2) -HY -H2LK2 #H destruct
+/4 width=10 by lleq_inv_lref_ge/
+qed-.
+
+(* Properties on relocation *************************************************)
+
+lemma lsx_lift_le: ∀h,g,G,K,T,U,dt,d,e. dt ≤ yinj d →
+ ⇧[d, e] T ≡ U → G ⊢ ⬊*[h, g, T, dt] K →
+ ∀L. ⇩[Ⓕ, d, e] L ≡ K → G ⊢ ⬊*[h, g, U, dt] L.
+#h #g #G #K #T #U #dt #d #e #Hdtd #HTU #H @(lsx_ind … H) -K
+#K1 #_ #IHK1 #L1 #HLK1 @lsx_intro
+#L2 #HL12 #HnU elim (lpx_drop_conf … HLK1 … HL12) -HL12
+/4 width=10 by lleq_lift_le/
+qed-.
+
+lemma lsx_lift_ge: ∀h,g,G,K,T,U,dt,d,e. yinj d ≤ dt →
+ ⇧[d, e] T ≡ U → G ⊢ ⬊*[h, g, T, dt] K →
+ ∀L. ⇩[Ⓕ, d, e] L ≡ K → G ⊢ ⬊*[h, g, U, dt + e] L.
+#h #g #G #K #T #U #dt #d #e #Hddt #HTU #H @(lsx_ind … H) -K
+#K1 #_ #IHK1 #L1 #HLK1 @lsx_intro
+#L2 #HL12 #HnU elim (lpx_drop_conf … HLK1 … HL12) -HL12
+/4 width=9 by lleq_lift_ge/
+qed-.
+
+(* Inversion lemmas on relocation *******************************************)
+
+lemma lsx_inv_lift_le: ∀h,g,G,L,T,U,dt,d,e. dt ≤ yinj d →
+ ⇧[d, e] T ≡ U → G ⊢ ⬊*[h, g, U, dt] L →
+ ∀K. ⇩[Ⓕ, d, e] L ≡ K → G ⊢ ⬊*[h, g, T, dt] K.
+#h #g #G #L #T #U #dt #d #e #Hdtd #HTU #H @(lsx_ind … H) -L
+#L1 #_ #IHL1 #K1 #HLK1 @lsx_intro
+#K2 #HK12 #HnT elim (drop_lpx_trans … HLK1 … HK12) -HK12
+/4 width=10 by lleq_inv_lift_le/
+qed-.
+
+lemma lsx_inv_lift_be: ∀h,g,G,L,T,U,dt,d,e. yinj d ≤ dt → dt ≤ d + e →
+ ⇧[d, e] T ≡ U → G ⊢ ⬊*[h, g, U, dt] L →
+ ∀K. ⇩[Ⓕ, d, e] L ≡ K → G ⊢ ⬊*[h, g, T, d] K.
+#h #g #G #L #T #U #dt #d #e #Hddt #Hdtde #HTU #H @(lsx_ind … H) -L
+#L1 #_ #IHL1 #K1 #HLK1 @lsx_intro
+#K2 #HK12 #HnT elim (drop_lpx_trans … HLK1 … HK12) -HK12
+/4 width=11 by lleq_inv_lift_be/
+qed-.
+
+lemma lsx_inv_lift_ge: ∀h,g,G,L,T,U,dt,d,e. yinj d + yinj e ≤ dt →
+ ⇧[d, e] T ≡ U → G ⊢ ⬊*[h, g, U, dt] L →
+ ∀K. ⇩[Ⓕ, d, e] L ≡ K → G ⊢ ⬊*[h, g, T, dt-e] K.
+#h #g #G #L #T #U #dt #d #e #Hdedt #HTU #H @(lsx_ind … H) -L
+#L1 #_ #IHL1 #K1 #HLK1 @lsx_intro
+#K2 #HK12 #HnT elim (drop_lpx_trans … HLK1 … HK12) -HK12
+/4 width=9 by lleq_inv_lift_ge/
+qed-.
+++ /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/multiple/lleq_ldrop.ma".
-include "basic_2/reduction/lpx_ldrop.ma".
-include "basic_2/computation/lsx.ma".
-
-(* SN EXTENDED STRONGLY NORMALIZING LOCAL ENVIRONMENTS **********************)
-
-(* Advanced properties ******************************************************)
-
-lemma lsx_lref_free: ∀h,g,G,L,d,i. |L| ≤ i → G ⊢ ⬊*[h, g, #i, d] L.
-#h #g #G #L1 #d #i #HL1 @lsx_intro
-#L2 #HL12 #H elim H -H
-/4 width=6 by lpx_fwd_length, lleq_free, le_repl_sn_conf_aux/
-qed.
-
-lemma lsx_lref_skip: ∀h,g,G,L,d,i. yinj i < d → G ⊢ ⬊*[h, g, #i, d] L.
-#h #g #G #L1 #d #i #HL1 @lsx_intro
-#L2 #HL12 #H elim H -H
-/3 width=4 by lpx_fwd_length, lleq_skip/
-qed.
-
-(* Advanced forward lemmas **************************************************)
-
-lemma lsx_fwd_lref_be: ∀h,g,I,G,L,d,i. d ≤ yinj i → G ⊢ ⬊*[h, g, #i, d] L →
- ∀K,V. ⇩[i] L ≡ K.ⓑ{I}V → G ⊢ ⬊*[h, g, V, 0] K.
-#h #g #I #G #L #d #i #Hdi #H @(lsx_ind … H) -L
-#L1 #_ #IHL1 #K1 #V #HLK1 @lsx_intro
-#K2 #HK12 #HnK12 lapply (ldrop_fwd_drop2 … HLK1)
-#H2LK1 elim (ldrop_lpx_trans … H2LK1 … HK12) -H2LK1 -HK12
-#L2 #HL12 #H2LK2 #H elim (leq_ldrop_conf_be … H … HLK1) -H /2 width=1 by ylt_inj/
-#Y #_ #HLK2 lapply (ldrop_fwd_drop2 … HLK2)
-#HY lapply (ldrop_mono … HY … H2LK2) -HY -H2LK2 #H destruct
-/4 width=10 by lleq_inv_lref_ge/
-qed-.
-
-(* Properties on relocation *************************************************)
-
-lemma lsx_lift_le: ∀h,g,G,K,T,U,dt,d,e. dt ≤ yinj d →
- ⇧[d, e] T ≡ U → G ⊢ ⬊*[h, g, T, dt] K →
- ∀L. ⇩[Ⓕ, d, e] L ≡ K → G ⊢ ⬊*[h, g, U, dt] L.
-#h #g #G #K #T #U #dt #d #e #Hdtd #HTU #H @(lsx_ind … H) -K
-#K1 #_ #IHK1 #L1 #HLK1 @lsx_intro
-#L2 #HL12 #HnU elim (lpx_ldrop_conf … HLK1 … HL12) -HL12
-/4 width=10 by lleq_lift_le/
-qed-.
-
-lemma lsx_lift_ge: ∀h,g,G,K,T,U,dt,d,e. yinj d ≤ dt →
- ⇧[d, e] T ≡ U → G ⊢ ⬊*[h, g, T, dt] K →
- ∀L. ⇩[Ⓕ, d, e] L ≡ K → G ⊢ ⬊*[h, g, U, dt + e] L.
-#h #g #G #K #T #U #dt #d #e #Hddt #HTU #H @(lsx_ind … H) -K
-#K1 #_ #IHK1 #L1 #HLK1 @lsx_intro
-#L2 #HL12 #HnU elim (lpx_ldrop_conf … HLK1 … HL12) -HL12
-/4 width=9 by lleq_lift_ge/
-qed-.
-
-(* Inversion lemmas on relocation *******************************************)
-
-lemma lsx_inv_lift_le: ∀h,g,G,L,T,U,dt,d,e. dt ≤ yinj d →
- ⇧[d, e] T ≡ U → G ⊢ ⬊*[h, g, U, dt] L →
- ∀K. ⇩[Ⓕ, d, e] L ≡ K → G ⊢ ⬊*[h, g, T, dt] K.
-#h #g #G #L #T #U #dt #d #e #Hdtd #HTU #H @(lsx_ind … H) -L
-#L1 #_ #IHL1 #K1 #HLK1 @lsx_intro
-#K2 #HK12 #HnT elim (ldrop_lpx_trans … HLK1 … HK12) -HK12
-/4 width=10 by lleq_inv_lift_le/
-qed-.
-
-lemma lsx_inv_lift_be: ∀h,g,G,L,T,U,dt,d,e. yinj d ≤ dt → dt ≤ d + e →
- ⇧[d, e] T ≡ U → G ⊢ ⬊*[h, g, U, dt] L →
- ∀K. ⇩[Ⓕ, d, e] L ≡ K → G ⊢ ⬊*[h, g, T, d] K.
-#h #g #G #L #T #U #dt #d #e #Hddt #Hdtde #HTU #H @(lsx_ind … H) -L
-#L1 #_ #IHL1 #K1 #HLK1 @lsx_intro
-#K2 #HK12 #HnT elim (ldrop_lpx_trans … HLK1 … HK12) -HK12
-/4 width=11 by lleq_inv_lift_be/
-qed-.
-
-lemma lsx_inv_lift_ge: ∀h,g,G,L,T,U,dt,d,e. yinj d + yinj e ≤ dt →
- ⇧[d, e] T ≡ U → G ⊢ ⬊*[h, g, U, dt] L →
- ∀K. ⇩[Ⓕ, d, e] L ≡ K → G ⊢ ⬊*[h, g, T, dt-e] K.
-#h #g #G #L #T #U #dt #d #e #Hdedt #HTU #H @(lsx_ind … H) -L
-#L1 #_ #IHL1 #K1 #HLK1 @lsx_intro
-#K2 #HK12 #HnT elim (ldrop_lpx_trans … HLK1 … HK12) -HK12
-/4 width=9 by lleq_inv_lift_ge/
-qed-.
qed-.
lemma lsx_leq_conf: ∀h,g,G,L1,T,d. G ⊢ ⬊*[h, g, T, d] L1 →
- â\88\80L2. L1 â\89\83[d, ∞] L2 → G ⊢ ⬊*[h, g, T, d] L2.
+ â\88\80L2. L1 ⩬[d, ∞] L2 → G ⊢ ⬊*[h, g, T, d] L2.
#h #g #G #L1 #T #d #H @(lsx_ind … H) -L1
#L1 #_ #IHL1 #L2 #HL12 @lsx_intro
#L3 #HL23 #HnL23 elim (leq_lpx_trans_lleq … HL12 … HL23) -HL12 -HL23
qed-.
(* Note: the constant 0 cannot be generalized *)
-lemma lsubsv_ldrop_O1_conf: ∀h,g,G,L1,L2. G ⊢ L1 ¡⫃[h, g] L2 →
- ∀K1,s,e. ⇩[s, 0, e] L1 ≡ K1 →
- ∃∃K2. G ⊢ K1 ¡⫃[h, g] K2 & ⇩[s, 0, e] L2 ≡ K2.
+lemma lsubsv_drop_O1_conf: ∀h,g,G,L1,L2. G ⊢ L1 ¡⫃[h, g] L2 →
+ ∀K1,s,e. ⇩[s, 0, e] L1 ≡ K1 →
+ ∃∃K2. G ⊢ K1 ¡⫃[h, g] K2 & ⇩[s, 0, e] L2 ≡ K2.
#h #g #G #L1 #L2 #H elim H -L1 -L2
[ /2 width=3 by ex2_intro/
| #I #L1 #L2 #V #_ #IHL12 #K1 #s #e #H
- elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK1
+ elim (drop_inv_O1_pair1 … H) -H * #He #HLK1
[ destruct
elim (IHL12 L1 s 0) -IHL12 // #X #HL12 #H
- <(ldrop_inv_O2 … H) in HL12; -H /3 width=3 by lsubsv_pair, ldrop_pair, ex2_intro/
- | elim (IHL12 … HLK1) -L1 /3 width=3 by ldrop_drop_lt, ex2_intro/
+ <(drop_inv_O2 … H) in HL12; -H /3 width=3 by lsubsv_pair, drop_pair, ex2_intro/
+ | elim (IHL12 … HLK1) -L1 /3 width=3 by drop_drop_lt, ex2_intro/
]
| #L1 #L2 #W #V #l #H1W #HV #HVW #H2W #H1l #H2l #_ #IHL12 #K1 #s #e #H
- elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK1
+ elim (drop_inv_O1_pair1 … H) -H * #He #HLK1
[ destruct
elim (IHL12 L1 s 0) -IHL12 // #X #HL12 #H
- <(ldrop_inv_O2 … H) in HL12; -H /3 width=4 by lsubsv_abbr, ldrop_pair, ex2_intro/
- | elim (IHL12 … HLK1) -L1 /3 width=3 by ldrop_drop_lt, ex2_intro/
+ <(drop_inv_O2 … H) in HL12; -H /3 width=4 by lsubsv_abbr, drop_pair, ex2_intro/
+ | elim (IHL12 … HLK1) -L1 /3 width=3 by drop_drop_lt, ex2_intro/
]
]
qed-.
(* Note: the constant 0 cannot be generalized *)
-lemma lsubsv_ldrop_O1_trans: ∀h,g,G,L1,L2. G ⊢ L1 ¡⫃[h, g] L2 →
- ∀K2,s, e. ⇩[s, 0, e] L2 ≡ K2 →
- ∃∃K1. G ⊢ K1 ¡⫃[h, g] K2 & ⇩[s, 0, e] L1 ≡ K1.
+lemma lsubsv_drop_O1_trans: ∀h,g,G,L1,L2. G ⊢ L1 ¡⫃[h, g] L2 →
+ ∀K2,s, e. ⇩[s, 0, e] L2 ≡ K2 →
+ ∃∃K1. G ⊢ K1 ¡⫃[h, g] K2 & ⇩[s, 0, e] L1 ≡ K1.
#h #g #G #L1 #L2 #H elim H -L1 -L2
[ /2 width=3 by ex2_intro/
| #I #L1 #L2 #V #_ #IHL12 #K2 #s #e #H
- elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK2
+ elim (drop_inv_O1_pair1 … H) -H * #He #HLK2
[ destruct
elim (IHL12 L2 s 0) -IHL12 // #X #HL12 #H
- <(ldrop_inv_O2 … H) in HL12; -H /3 width=3 by lsubsv_pair, ldrop_pair, ex2_intro/
- | elim (IHL12 … HLK2) -L2 /3 width=3 by ldrop_drop_lt, ex2_intro/
+ <(drop_inv_O2 … H) in HL12; -H /3 width=3 by lsubsv_pair, drop_pair, ex2_intro/
+ | elim (IHL12 … HLK2) -L2 /3 width=3 by drop_drop_lt, ex2_intro/
]
| #L1 #L2 #W #V #l #H1W #HV #HVW #H2W #H1l #H2l #_ #IHL12 #K2 #s #e #H
- elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK2
+ elim (drop_inv_O1_pair1 … H) -H * #He #HLK2
[ destruct
elim (IHL12 L2 s 0) -IHL12 // #X #HL12 #H
- <(ldrop_inv_O2 … H) in HL12; -H /3 width=4 by lsubsv_abbr, ldrop_pair, ex2_intro/
- | elim (IHL12 … HLK2) -L2 /3 width=3 by ldrop_drop_lt, ex2_intro/
+ <(drop_inv_O2 … H) in HL12; -H /3 width=4 by lsubsv_abbr, drop_pair, ex2_intro/
+ | elim (IHL12 … HLK2) -L2 /3 width=3 by drop_drop_lt, ex2_intro/
]
]
qed-.
[1,2: /2 width=3 by ex2_intro/
| #G #L2 #K2 #X #Y #U #i #l1 #HLK2 #_ #HYU #IHXY #l2 #Hl12 #Hl2 #L1 #HL12
elim (da_inv_lref … Hl2) -Hl2 * #K0 #V0 [| #l0 ] #HK0 #HV0
- lapply (ldrop_mono … HK0 … HLK2) -HK0 #H destruct
- elim (lsubsv_ldrop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1
+ lapply (drop_mono … HK0 … HLK2) -HK0 #H destruct
+ elim (lsubsv_drop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1
elim (lsubsv_inv_pair2 … H) -H * #K1 [ | -HYU -IHXY -HLK1 ]
[ #HK12 #H destruct
elim (IHXY … Hl12 HV0 … HK12) -K2 -l2 #T #HXT #HTY
- lapply (ldrop_fwd_drop2 … HLK1) #H
+ lapply (drop_fwd_drop2 … HLK1) #H
elim (lift_total T 0 (i+1))
/3 width=12 by lstas_ldef, cpcs_lift, ex2_intro/
| #V #l0 #_ #_ #_ #_ #_ #_ #_ #H destruct
]
| #G #L2 #K2 #X #Y #Y0 #U #i #l1 #HLK2 #HXY0 #_ #HYU #IHXY #l2 #Hl12 #Hl2 #L1 #HL12
elim (da_inv_lref … Hl2) -Hl2 * #K0 #V0 [| #l0 ] #HK0 #HV0 [| #H1 ]
- lapply (ldrop_mono … HK0 … HLK2) -HK0 #H2 destruct
+ lapply (drop_mono … HK0 … HLK2) -HK0 #H2 destruct
lapply (le_plus_to_le_r … Hl12) -Hl12 #Hl12
- elim (lsubsv_ldrop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1
+ elim (lsubsv_drop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1
elim (lsubsv_inv_pair2 … H) -H * #K1
[ #HK12 #H destruct
lapply (lsubsv_fwd_lsubd … HK12) #H
lapply (lsubd_da_trans … HV0 … H) -H #H
elim (da_inv_sta … H) -H
elim (IHXY … Hl12 HV0 … HK12) -K2 -Hl12 #Y1
- lapply (ldrop_fwd_drop2 … HLK1)
+ lapply (drop_fwd_drop2 … HLK1)
elim (lift_total Y1 0 (i+1))
/3 width=12 by lstas_ldec, cpcs_lift, ex2_intro/
| #V #l #_ #_ #HVX #_ #HV #HX #HK12 #_ #H destruct
elim (lstas_total … HVW (l1+1)) -W #W #HVW
lapply (HVX … Hl12 HVW HXY0) -HVX -Hl12 -HXY0 #HWY0
lapply (cpcs_trans … HWY0 … HY0) -Y0
- lapply (ldrop_fwd_drop2 … HLK1)
+ lapply (drop_fwd_drop2 … HLK1)
elim (lift_total W 0 (i+1))
/4 width=12 by lstas_ldef, lstas_cast, cpcs_lift, ex2_intro/
]
∀L1. G ⊢ L1 ¡⫃[h, g] L2 → ⦃G, L1⦄ ⊢ T ¡[h, g].
#h #g #G #L2 #T #H elim H -G -L2 -T //
[ #I #G #L2 #K2 #V #i #HLK2 #_ #IHV #L1 #HL12
- elim (lsubsv_ldrop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1
+ elim (lsubsv_drop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1
elim (lsubsv_inv_pair2 … H) -H * #K1
[ #HK12 #H destruct /3 width=5 by snv_lref/
| #W #l #H1V #H1W #HWV #_ #HWl #_ #_ #H1 #H2 destruct -IHV
| #i #HG0 #HL0 #HT0 #H1 #l #H2 #X3 #H3 #L2 #HL12 destruct -IH3 -IH2
elim (snv_inv_lref … H1) -H1 #I0 #K0 #X0 #H #HX0
elim (da_inv_lref … H2) -H2 * #K1 [ #V1 | #W1 #l1 ] #HLK1 [ #HV1 | #HW1 #H ] destruct
- lapply (ldrop_mono … H … HLK1) -H #H destruct
+ lapply (drop_mono … H … HLK1) -H #H destruct
elim (cpr_inv_lref1 … H3) -H3
[1,3: #H destruct
lapply (fqup_lref … G1 … HLK1)
- elim (lpr_ldrop_conf … HLK1 … HL12) -HLK1 -HL12 #X #H #HLK2
+ elim (lpr_drop_conf … HLK1 … HL12) -HLK1 -HL12 #X #H #HLK2
elim (lpr_inv_pair1 … H) -H #K2 #V2 #HK12 #HV12 #H destruct
/4 width=10 by da_ldef, da_ldec, fqup_fpbg/
|2,4: * #K0 #V0 #W0 #H #HVW0 #HW0
- lapply (ldrop_mono … H … HLK1) -H #H destruct
+ lapply (drop_mono … H … HLK1) -H #H destruct
lapply (fqup_lref … G1 … HLK1)
- elim (lpr_ldrop_conf … HLK1 … HL12) -HLK1 -HL12 #X #H #HLK2
+ elim (lpr_drop_conf … HLK1 … HL12) -HLK1 -HL12 #X #H #HLK2
elim (lpr_inv_pair1 … H) -H #K2 #V2 #HK12 #_ #H destruct
- lapply (ldrop_fwd_drop2 … HLK2) -V2
+ lapply (drop_fwd_drop2 … HLK2) -V2
/4 width=8 by da_lift, fqup_fpbg/
]
| #p #_ #_ #HT0 #H1 destruct -IH3 -IH2 -IH1
[ #V2 #T2 #HV12 #HT12 #H destruct
/4 width=9 by da_bind, fqup_fpbg, lpr_pair/
| #T2 #HT12 #HT2 #H1 #H2 destruct
- /4 width=11 by da_inv_lift, fqup_fpbg, lpr_pair, ldrop_drop/
+ /4 width=11 by da_inv_lift, fqup_fpbg, lpr_pair, drop_drop/
]
| #V1 #T1 #HG0 #HL0 #HT0 #H1 #l #H2 #X3 #H3 #L2 #HL12 destruct
elim (snv_inv_appl … H1) -H1 #b0 #W1 #W0 #T0 #l0 #HV1 #HT1 #Hl0 #HVW1 #HW10 #HT10
>(lift_inv_sort1 … H) -X -K -d -e //
| #I #G #K #K0 #V #i #HK0 #_ #IHV #L #s #d #e #HLK #X #H
elim (lift_inv_lref1 … H) * #Hid #H destruct
- [ elim (ldrop_trans_le … HLK … HK0) -K /2 width=2 by lt_to_le/ #X #HL0 #H
- elim (ldrop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid #L0 #W #HLK0 #HVW #H destruct
+ [ elim (drop_trans_le … HLK … HK0) -K /2 width=2 by lt_to_le/ #X #HL0 #H
+ elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid #L0 #W #HLK0 #HVW #H destruct
/3 width=9 by snv_lref/
- | lapply (ldrop_trans_ge … HLK … HK0 ?) -K
- /3 width=9 by snv_lref, ldrop_inv_gen/
+ | lapply (drop_trans_ge … HLK … HK0 ?) -K
+ /3 width=9 by snv_lref, drop_inv_gen/
]
| #a #I #G #K #V #T #_ #_ #IHV #IHT #L #s #d #e #HLK #X #H
elim (lift_inv_bind1 … H) -H #W #U #HVW #HTU #H destruct
- /4 width=5 by snv_bind, ldrop_skip/
+ /4 width=5 by snv_bind, drop_skip/
| #a #G #K #V #V0 #V1 #T #T1 #l #_ #_ #Hl #HV0 #HV01 #HT1 #IHV #IHT #L #s #d #e #HLK #X #H
elim (lift_inv_flat1 … H) -H #W #U #HVW #HTU #H destruct
elim (lift_total V0 d e) #W0 #HVW0
>(lift_inv_sort2 … H) -X -L -d -e //
| #I #G #L #L0 #W #i #HL0 #_ #IHW #K #s #d #e #HLK #X #H
elim (lift_inv_lref2 … H) * #Hid #H destruct
- [ elim (ldrop_conf_le … HLK … HL0) -L /2 width=2 by lt_to_le/ #X #HK0 #H
- elim (ldrop_inv_skip1 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid #K0 #V #HLK0 #HVW #H destruct
+ [ elim (drop_conf_le … HLK … HL0) -L /2 width=2 by lt_to_le/ #X #HK0 #H
+ elim (drop_inv_skip1 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid #K0 #V #HLK0 #HVW #H destruct
/3 width=12 by snv_lref/
- | lapply (ldrop_conf_ge … HLK … HL0 ?) -L /3 width=9 by snv_lref/
+ | lapply (drop_conf_ge … HLK … HL0 ?) -L /3 width=9 by snv_lref/
]
| #a #I #G #L #W #U #_ #_ #IHW #IHU #K #s #d #e #HLK #X #H
elim (lift_inv_bind2 … H) -H #V #T #HVW #HTU #H destruct
- /4 width=5 by snv_bind, ldrop_skip/
+ /4 width=5 by snv_bind, drop_skip/
| #a #G #L #W #W0 #W1 #U #U1 #l #_ #_ #Hl #HW0 #HW01 #HU1 #IHW #IHU #K #s #d #e #HLK #X #H
elim (lift_inv_flat2 … H) -H #V #T #HVW #HTU #H destruct
lapply (da_inv_lift … Hl … HLK … HVW) -Hl #Hl
#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
[ #I1 #G1 #L1 #V1 #H
elim (snv_inv_lref … H) -H #I2 #L2 #V2 #H #HV2
- lapply (ldrop_inv_O2 … H) -H #H destruct //
+ lapply (drop_inv_O2 … H) -H #H destruct //
|2: *
|5,6: /3 width=8 by snv_inv_lift/
]
>(cpr_inv_sort1 … H2) -X //
| #i #HG0 #HL0 #HT0 #H1 #X #H2 #L2 #HL12 destruct -IH4 -IH3 -IH2
elim (snv_inv_lref … H1) -H1 #I #K1 #V1 #HLK1 #HV1
- elim (lpr_ldrop_conf … HLK1 … HL12) -HL12 #X #H #HLK2
+ elim (lpr_drop_conf … HLK1 … HL12) -HL12 #X #H #HLK2
elim (lpr_inv_pair1 … H) -H #K2 #V2 #HK12 #HV12 #H destruct
lapply (fqup_lref … G1 … HLK1) #HKL
elim (cpr_inv_lref1 … H2) -H2
[ #H destruct -HLK1 /4 width=10 by fqup_fpbg, snv_lref/
| * #K0 #V0 #W0 #H #HVW0 #W0 -HV12
- lapply (ldrop_mono … H … HLK1) -HLK1 -H #H destruct
- lapply (ldrop_fwd_drop2 … HLK2) -HLK2 /4 width=8 by fqup_fpbg, snv_lift/
+ lapply (drop_mono … H … HLK1) -HLK1 -H #H destruct
+ lapply (drop_fwd_drop2 … HLK2) -HLK2 /4 width=8 by fqup_fpbg, snv_lift/
]
| #p #HG0 #HL0 #HT0 #H1 #X #H2 #L2 #HL12 destruct -IH4 -IH3 -IH2 -IH1
elim (snv_inv_gref … H1)
elim (cpr_inv_bind1 … H2) -H2 *
[ #V2 #T2 #HV12 #HT12 #H destruct /4 width=8 by fqup_fpbg, snv_bind, lpr_pair/
| #T2 #HT12 #HXT2 #H1 #H2 destruct -HV1
- /4 width=10 by fqup_fpbg, snv_inv_lift, lpr_pair, ldrop_drop/
+ /4 width=10 by fqup_fpbg, snv_inv_lift, lpr_pair, drop_drop/
]
| #V1 #T1 #HG0 #HL0 #HT0 #H1 #X #H2 #L2 #HL12 destruct
elim (snv_inv_appl … H1) -H1 #a #W10 #W1 #U1 #l0 #HV1 #HT1 #Hl0 #HVW1 #HW10 #HTU1
elim (sta_cpr_lpr_aux … IH3 … HVW1 … HV10 … HL12) /2 width=2 by fqup_fpbg/ -IH3 -HVW1 #X #H1 #H2
lapply (cpcs_canc_sn … H2 HW10) -W10 #H2
elim (lift_total X 0 1) #W20 #H3
- lapply (sta_lift … H1 (L2.ⓓW2) … HV02 … H3) /2 width=2 by ldrop_drop/ -H1 #HVW20
- lapply (cpcs_lift … (L2.ⓓW2) … H3 … HW13 H2) /2 width=2 by ldrop_drop/ -HW13 -H3 -H2 #HW320
+ lapply (sta_lift … H1 (L2.ⓓW2) … HV02 … H3) /2 width=2 by drop_drop/ -H1 #HVW20
+ lapply (cpcs_lift … (L2.ⓓW2) … H3 … HW13 H2) /2 width=2 by drop_drop/ -HW13 -H3 -H2 #HW320
lapply (cpcs_cprs_strap1 … HW320 … HW1) -W3 #HW20
elim (cpcs_inv_cprs … HW20) -HW20 #W3 #HW203 #HW3
lapply (cpds_cprs_trans … (ⓛ{a}W3.U2) HTU2 ?) /2 width=1 by cprs_bind/ -HW3 -HTU2 #HTU2
lapply (IH2 … Hl0 … HV10 … HL12) /2 width=1 by fqup_fpbg/ -IH2 -Hl0 #Hl0
- lapply (da_lift … Hl0 (L2.ⓓW2) … HV02) /2 width=2 by ldrop_drop/ -Hl0 #Hl0
+ lapply (da_lift … Hl0 (L2.ⓓW2) … HV02) /2 width=2 by drop_drop/ -Hl0 #Hl0
lapply (IH1 … HW02 … HL12) /2 width=1 by fqup_fpbg/ -HW0 #HW2
lapply (IH1 … HV10 … HL12) /2 width=1 by fqup_fpbg/ -HV1 -HV10 #HV0
lapply (IH1 … HT02 (L2.ⓓW2) ?) /2 width=1 by fqup_fpbg, lpr_pair/ -L1 #HT2
- lapply (snv_lift … HV0 (L2.ⓓW2) … HV02) /3 width=7 by snv_bind, snv_appl, ldrop_drop/
+ lapply (snv_lift … HV0 (L2.ⓓW2) … HV02) /3 width=7 by snv_bind, snv_appl, drop_drop/
]
| #W1 #T1 #HG0 #HL0 #HT0 #H1 #X #H2 #L2 #HL12 destruct -IH4
elim (snv_inv_cast … H1) -H1 #U1 #l0 #HW1 #HT1 #Hl0 #HTU1 #HUW1
[ lapply (lstas_inv_O … H2) -H2 #H destruct // ]
elim (snv_inv_lref … H1) -H1 #I0 #K0 #X0 #HLK0 #HX0
elim (da_inv_lref … Hl1) -Hl1 * #K1 [ #V1 | #W1 #l ] #HLK1 [ #Hl1 | #Hl #H ]
- lapply (ldrop_mono … HLK0 … HLK1) -HLK0 #H0 destruct
+ lapply (drop_mono … HLK0 … HLK1) -HLK0 #H0 destruct
elim (lstas_inv_lref1 … H2) -H2 * #K0 #Y0 #X0 [2,4: #Y1 ] #HLK0 [1,2: #HY01 ] #HYX0 #HX0
- lapply (ldrop_mono … HLK0 … HLK1) -HLK0 #H destruct
+ lapply (drop_mono … HLK0 … HLK1) -HLK0 #H destruct
[ lapply (le_plus_to_le_r … Hl21) -Hl21 #Hl21 ]
lapply (fqup_lref … G1 … HLK1) #H
- lapply (ldrop_fwd_drop2 … HLK1) -HLK1 /4 width=8 by fqup_fpbg, snv_lift/
+ lapply (drop_fwd_drop2 … HLK1) -HLK1 /4 width=8 by fqup_fpbg, snv_lift/
| #p #HG0 #HL0 #HT0 #H1 #l1 #l2 #Hl21 #Hl1 #X #H2 destruct -IH4 -IH3 -IH2 -IH1
elim (snv_inv_gref … H1)
| #a #I #V1 #T1 #HG0 #HL0 #HT0 #H1 #l1 #l2 #Hl21 #Hl1 #X #H2 destruct -IH4 -IH3 -IH2
[ lapply (lstas_inv_O … H2) -H2 #H destruct -IH1 -H1 -l1 /4 width=5 by lpr_cpcs_conf, cpr_cpcs_dx, ex2_intro/ ]
elim (snv_inv_lref … H1) -H1 #I0 #K0 #X0 #HK0 #HX0
elim (da_inv_lref … Hl1) -Hl1 * #K1 [ #V1 | #W1 #l0 ] #HLK1 [ #HVl1 | #HWl1 #H destruct ]
- lapply (ldrop_mono … HK0 … HLK1) -HK0 #H destruct
+ lapply (drop_mono … HK0 … HLK1) -HK0 #H destruct
elim (lstas_inv_lref1 … H2) -H2 * #K0 #V0 #W0 [2,4: #X0 ] #HK0 [1,2: #_ -X0 ] #HVW0 #HX2
- lapply (ldrop_mono … HK0 … HLK1) -HK0 #H destruct
+ lapply (drop_mono … HK0 … HLK1) -HK0 #H destruct
[ lapply (le_plus_to_le_r … Hl21) -Hl21 #Hl21 ]
lapply (fqup_lref … G1 … HLK1) #HKV1
- elim (lpr_ldrop_conf … HLK1 … HL12) -HL12 #X #H #HLK2
+ elim (lpr_drop_conf … HLK1 … HL12) -HL12 #X #H #HLK2
elim (lpr_inv_pair1 … H) -H #K2 [ #W2 | #V2 ] #HK12 [ #HW12 | #HV12 ] #H destruct
- lapply (ldrop_fwd_drop2 … HLK2) #H2
+ lapply (drop_fwd_drop2 … HLK2) #H2
elim (cpr_inv_lref1 … H3) -H3
[1,3: #H destruct -HLK1
|2,4: * #K0 #V0 #X0 #H #HVX0 #HX0
- lapply (ldrop_mono … H … HLK1) -H -HLK1 #H destruct
+ lapply (drop_mono … H … HLK1) -H -HLK1 #H destruct
]
[ lapply (IH2 … HWl1 … HW12 … HK12) /2 width=1 by fqup_fpbg/ -IH2 #H
elim (da_inv_sta … H) -H
| #T3 #HT13 #HXT3 #H1 #H2 destruct
elim (IH1 … Hl1 … HTU1 … HT13 (L2.ⓓV1)) -IH1 -Hl1 -HTU1 -HT13 /2 width=1 by fqup_fpbg, lpr_pair/ -T1 -HL12 #U3 #HTU3 #HU13
elim (lstas_inv_lift1 … HTU3 L2 … HXT3) -T3
- /5 width=8 by cpcs_cpr_strap1, cpcs_bind1, cpr_zeta, ldrop_drop, ex2_intro/
+ /5 width=8 by cpcs_cpr_strap1, cpcs_bind1, cpr_zeta, drop_drop, ex2_intro/
]
| #V1 #T1 #HG0 #HL0 #HT0 #H1 #l1 #l2 #Hl21 #Hl1 #X2 #H2 #X3 #H3 #L2 #HL12 destruct
elim (snv_inv_appl … H1) -H1 #a #W1 #W10 #U10 #l0 #HV1 #HT1 #Hl0 #HVW1 #HW10 #HTU10
(**************************************************************************)
include "ground_2/ynat/ynat_lt.ma".
-include "basic_2/notation/relations/iso_4.ma".
+include "basic_2/notation/relations/midiso_4.ma".
include "basic_2/grammar/lenv_length.ma".
(* EQUIVALENCE FOR LOCAL ENVIRONMENTS ***************************************)
| leq_pair: ∀I,L1,L2,V,e. leq 0 e L1 L2 →
leq 0 (⫯e) (L1.ⓑ{I}V) (L2.ⓑ{I}V)
| leq_succ: ∀I1,I2,L1,L2,V1,V2,d,e.
- leq d e L1 L2 → leq (⫯d) e (L1. ⓑ{I1}V1) (L2. ⓑ{I2} V2)
+ leq d e L1 L2 → leq (⫯d) e (L1.ⓑ{I1}V1) (L2.ⓑ{I2}V2)
.
interpretation
"equivalence (local environment)"
- 'Iso d e L1 L2 = (leq d e L1 L2).
+ 'MidIso d e L1 L2 = (leq d e L1 L2).
(* Basic properties *********************************************************)
-lemma leq_pair_lt: â\88\80I,L1,L2,V,e. L1 â\89\83[0, ⫰e] L2 → 0 < e →
- L1.â\93\91{I}V â\89\83[0, e] L2.ⓑ{I}V.
+lemma leq_pair_lt: â\88\80I,L1,L2,V,e. L1 ⩬[0, ⫰e] L2 → 0 < e →
+ L1.â\93\91{I}V ⩬[0, e] L2.ⓑ{I}V.
#I #L1 #L2 #V #e #HL12 #He <(ylt_inv_O1 … He) /2 width=1 by leq_pair/
qed.
-lemma leq_succ_lt: â\88\80I1,I2,L1,L2,V1,V2,d,e. L1 â\89\83[⫰d, e] L2 → 0 < d →
- L1.â\93\91{I1}V1 â\89\83[d, e] L2. ⓑ{I2}V2.
+lemma leq_succ_lt: â\88\80I1,I2,L1,L2,V1,V2,d,e. L1 ⩬[⫰d, e] L2 → 0 < d →
+ L1.â\93\91{I1}V1 ⩬[d, e] L2. ⓑ{I2}V2.
#I1 #I2 #L1 #L2 #V1 #V2 #d #e #HL12 #Hd <(ylt_inv_O1 … Hd) /2 width=1 by leq_succ/
qed.
-lemma leq_pair_O_Y: â\88\80L1,L2. L1 â\89\83[0, ∞] L2 →
- â\88\80I,V. L1.â\93\91{I}V â\89\83[0, ∞] L2.ⓑ{I}V.
+lemma leq_pair_O_Y: â\88\80L1,L2. L1 ⩬[0, ∞] L2 →
+ â\88\80I,V. L1.â\93\91{I}V ⩬[0, ∞] L2.ⓑ{I}V.
#L1 #L2 #HL12 #I #V lapply (leq_pair I … V … HL12) -HL12 //
qed.
-lemma leq_refl: â\88\80L,d,e. L â\89\83[d, e] L.
+lemma leq_refl: â\88\80L,d,e. L ⩬[d, e] L.
#L elim L -L //
#L #I #V #IHL #d elim (ynat_cases … d) [| * #x ]
#Hd destruct /2 width=1 by leq_succ/
#He destruct /2 width=1 by leq_zero, leq_pair/
qed.
-lemma leq_O2: â\88\80L1,L2,d. |L1| = |L2| â\86\92 L1 â\89\83[d, yinj 0] L2.
+lemma leq_O2: â\88\80L1,L2,d. |L1| = |L2| â\86\92 L1 ⩬[d, yinj 0] L2.
#L1 elim L1 -L1 [| #L1 #I1 #V1 #IHL1 ]
* // [1,3: #L2 #I2 #V2 ] #d normalize
[1,3: <plus_n_Sm #H destruct ]
(* Basic inversion lemmas ***************************************************)
-fact leq_inv_atom1_aux: â\88\80L1,L2,d,e. L1 â\89\83[d, e] L2 → L1 = ⋆ → L2 = ⋆.
+fact leq_inv_atom1_aux: â\88\80L1,L2,d,e. L1 ⩬[d, e] L2 → L1 = ⋆ → L2 = ⋆.
#L1 #L2 #d #e * -L1 -L2 -d -e //
[ #I1 #I2 #L1 #L2 #V1 #V2 #_ #H destruct
| #I #L1 #L2 #V #e #_ #H destruct
]
qed-.
-lemma leq_inv_atom1: â\88\80L2,d,e. â\8b\86 â\89\83[d, e] L2 → L2 = ⋆.
+lemma leq_inv_atom1: â\88\80L2,d,e. â\8b\86 ⩬[d, e] L2 → L2 = ⋆.
/2 width=5 by leq_inv_atom1_aux/ qed-.
-fact leq_inv_zero1_aux: â\88\80L1,L2,d,e. L1 â\89\83[d, e] L2 →
+fact leq_inv_zero1_aux: â\88\80L1,L2,d,e. L1 ⩬[d, e] L2 →
∀J1,K1,W1. L1 = K1.ⓑ{J1}W1 → d = 0 → e = 0 →
- â\88\83â\88\83J2,K2,W2. K1 â\89\83[0, 0] K2 & L2 = K2.ⓑ{J2}W2.
+ â\88\83â\88\83J2,K2,W2. K1 ⩬[0, 0] K2 & L2 = K2.ⓑ{J2}W2.
#L1 #L2 #d #e * -L1 -L2 -d -e
[ #d #e #J1 #K1 #W1 #H destruct
| #I1 #I2 #L1 #L2 #V1 #V2 #HL12 #J1 #K1 #W1 #H #_ #_ destruct
]
qed-.
-lemma leq_inv_zero1: â\88\80I1,K1,L2,V1. K1.â\93\91{I1}V1 â\89\83[0, 0] L2 →
- â\88\83â\88\83I2,K2,V2. K1 â\89\83[0, 0] K2 & L2 = K2.ⓑ{I2}V2.
+lemma leq_inv_zero1: â\88\80I1,K1,L2,V1. K1.â\93\91{I1}V1 ⩬[0, 0] L2 →
+ â\88\83â\88\83I2,K2,V2. K1 ⩬[0, 0] K2 & L2 = K2.ⓑ{I2}V2.
/2 width=9 by leq_inv_zero1_aux/ qed-.
-fact leq_inv_pair1_aux: â\88\80L1,L2,d,e. L1 â\89\83[d, e] L2 →
+fact leq_inv_pair1_aux: â\88\80L1,L2,d,e. L1 ⩬[d, e] L2 →
∀J,K1,W. L1 = K1.ⓑ{J}W → d = 0 → 0 < e →
- â\88\83â\88\83K2. K1 â\89\83[0, ⫰e] K2 & L2 = K2.ⓑ{J}W.
+ â\88\83â\88\83K2. K1 ⩬[0, ⫰e] K2 & L2 = K2.ⓑ{J}W.
#L1 #L2 #d #e * -L1 -L2 -d -e
[ #d #e #J #K1 #W #H destruct
| #I1 #I2 #L1 #L2 #V1 #V2 #_ #J #K1 #W #_ #_ #H
]
qed-.
-lemma leq_inv_pair1: â\88\80I,K1,L2,V,e. K1.â\93\91{I}V â\89\83[0, e] L2 → 0 < e →
- â\88\83â\88\83K2. K1 â\89\83[0, ⫰e] K2 & L2 = K2.ⓑ{I}V.
+lemma leq_inv_pair1: â\88\80I,K1,L2,V,e. K1.â\93\91{I}V ⩬[0, e] L2 → 0 < e →
+ â\88\83â\88\83K2. K1 ⩬[0, ⫰e] K2 & L2 = K2.ⓑ{I}V.
/2 width=6 by leq_inv_pair1_aux/ qed-.
-lemma leq_inv_pair: â\88\80I1,I2,L1,L2,V1,V2,e. L1.â\93\91{I1}V1 â\89\83[0, e] L2.ⓑ{I2}V2 → 0 < e →
- â\88§â\88§ L1 â\89\83[0, ⫰e] L2 & I1 = I2 & V1 = V2.
+lemma leq_inv_pair: â\88\80I1,I2,L1,L2,V1,V2,e. L1.â\93\91{I1}V1 ⩬[0, e] L2.ⓑ{I2}V2 → 0 < e →
+ â\88§â\88§ L1 ⩬[0, ⫰e] L2 & I1 = I2 & V1 = V2.
#I1 #I2 #L1 #L2 #V1 #V2 #e #H #He elim (leq_inv_pair1 … H) -H //
#Y #HL12 #H destruct /2 width=1 by and3_intro/
qed-.
-fact leq_inv_succ1_aux: â\88\80L1,L2,d,e. L1 â\89\83[d, e] L2 →
+fact leq_inv_succ1_aux: â\88\80L1,L2,d,e. L1 ⩬[d, e] L2 →
∀J1,K1,W1. L1 = K1.ⓑ{J1}W1 → 0 < d →
- â\88\83â\88\83J2,K2,W2. K1 â\89\83[⫰d, e] K2 & L2 = K2.ⓑ{J2}W2.
+ â\88\83â\88\83J2,K2,W2. K1 ⩬[⫰d, e] K2 & L2 = K2.ⓑ{J2}W2.
#L1 #L2 #d #e * -L1 -L2 -d -e
[ #d #e #J1 #K1 #W1 #H destruct
| #I1 #I2 #L1 #L2 #V1 #V2 #_ #J1 #K1 #W1 #_ #H
]
qed-.
-lemma leq_inv_succ1: â\88\80I1,K1,L2,V1,d,e. K1.â\93\91{I1}V1 â\89\83[d, e] L2 → 0 < d →
- â\88\83â\88\83I2,K2,V2. K1 â\89\83[⫰d, e] K2 & L2 = K2.ⓑ{I2}V2.
+lemma leq_inv_succ1: â\88\80I1,K1,L2,V1,d,e. K1.â\93\91{I1}V1 ⩬[d, e] L2 → 0 < d →
+ â\88\83â\88\83I2,K2,V2. K1 ⩬[⫰d, e] K2 & L2 = K2.ⓑ{I2}V2.
/2 width=5 by leq_inv_succ1_aux/ qed-.
-lemma leq_inv_atom2: â\88\80L1,d,e. L1 â\89\83[d, e] ⋆ → L1 = ⋆.
+lemma leq_inv_atom2: â\88\80L1,d,e. L1 ⩬[d, e] ⋆ → L1 = ⋆.
/3 width=3 by leq_inv_atom1, leq_sym/
qed-.
-lemma leq_inv_succ: â\88\80I1,I2,L1,L2,V1,V2,d,e. L1.â\93\91{I1}V1 â\89\83[d, e] L2.ⓑ{I2}V2 → 0 < d →
- L1 â\89\83[⫰d, e] L2.
+lemma leq_inv_succ: â\88\80I1,I2,L1,L2,V1,V2,d,e. L1.â\93\91{I1}V1 ⩬[d, e] L2.ⓑ{I2}V2 → 0 < d →
+ L1 ⩬[⫰d, e] L2.
#I1 #I2 #L1 #L2 #V1 #V2 #d #e #H #Hd elim (leq_inv_succ1 … H) -H //
#Z #Y #X #HL12 #H destruct //
qed-.
-lemma leq_inv_zero2: â\88\80I2,K2,L1,V2. L1 â\89\83[0, 0] K2.ⓑ{I2}V2 →
- â\88\83â\88\83I1,K1,V1. K1 â\89\83[0, 0] K2 & L1 = K1.ⓑ{I1}V1.
+lemma leq_inv_zero2: â\88\80I2,K2,L1,V2. L1 ⩬[0, 0] K2.ⓑ{I2}V2 →
+ â\88\83â\88\83I1,K1,V1. K1 ⩬[0, 0] K2 & L1 = K1.ⓑ{I1}V1.
#I2 #K2 #L1 #V2 #H elim (leq_inv_zero1 … (leq_sym … H)) -H
/3 width=5 by leq_sym, ex2_3_intro/
qed-.
-lemma leq_inv_pair2: â\88\80I,K2,L1,V,e. L1 â\89\83[0, e] K2.ⓑ{I}V → 0 < e →
- â\88\83â\88\83K1. K1 â\89\83[0, ⫰e] K2 & L1 = K1.ⓑ{I}V.
+lemma leq_inv_pair2: â\88\80I,K2,L1,V,e. L1 ⩬[0, e] K2.ⓑ{I}V → 0 < e →
+ â\88\83â\88\83K1. K1 ⩬[0, ⫰e] K2 & L1 = K1.ⓑ{I}V.
#I #K2 #L1 #V #e #H #He elim (leq_inv_pair1 … (leq_sym … H)) -H
/3 width=3 by leq_sym, ex2_intro/
qed-.
-lemma leq_inv_succ2: â\88\80I2,K2,L1,V2,d,e. L1 â\89\83[d, e] K2.ⓑ{I2}V2 → 0 < d →
- â\88\83â\88\83I1,K1,V1. K1 â\89\83[⫰d, e] K2 & L1 = K1.ⓑ{I1}V1.
+lemma leq_inv_succ2: â\88\80I2,K2,L1,V2,d,e. L1 ⩬[d, e] K2.ⓑ{I2}V2 → 0 < d →
+ â\88\83â\88\83I1,K1,V1. K1 ⩬[⫰d, e] K2 & L1 = K1.ⓑ{I1}V1.
#I2 #K2 #L1 #V2 #d #e #H #Hd elim (leq_inv_succ1 … (leq_sym … H)) -H
/3 width=5 by leq_sym, ex2_3_intro/
qed-.
(* Basic forward lemmas *****************************************************)
-lemma leq_fwd_length: â\88\80L1,L2,d,e. L1 â\89\83[d, e] L2 → |L1| = |L2|.
+lemma leq_fwd_length: â\88\80L1,L2,d,e. L1 ⩬[d, e] L2 → |L1| = |L2|.
#L1 #L2 #d #e #H elim H -L1 -L2 -d -e normalize //
qed-.
(* Advanced inversion lemmas ************************************************)
-fact leq_inv_O_Y_aux: â\88\80L1,L2,d,e. L1 â\89\83[d, e] L2 → d = 0 → e = ∞ → L1 = L2.
+fact leq_inv_O_Y_aux: â\88\80L1,L2,d,e. L1 ⩬[d, e] L2 → d = 0 → e = ∞ → L1 = L2.
#L1 #L2 #d #e #H elim H -L1 -L2 -d -e //
[ #I1 #I2 #L1 #L2 #V1 #V2 #_ #_ #_ #H destruct
| /4 width=1 by eq_f3, ysucc_inv_Y_dx/
]
qed-.
-lemma leq_inv_O_Y: â\88\80L1,L2. L1 â\89\83[0, ∞] L2 → L1 = L2.
+lemma leq_inv_O_Y: â\88\80L1,L2. L1 ⩬[0, ∞] L2 → L1 = L2.
/2 width=5 by leq_inv_O_Y_aux/ qed-.
]
qed-.
-theorem leq_canc_sn: â\88\80d,e,L,L1,L2. L â\89\83[d, e] L1 â\86\92 L â\89\83[d, e] L2 â\86\92 L1 â\89\83[d, e] L2.
+theorem leq_canc_sn: â\88\80d,e,L,L1,L2. L ⩬[d, e] L1 â\86\92 L ⩬[d, e] L2 â\86\92 L1 ⩬[d, e] L2.
/3 width=3 by leq_trans, leq_sym/ qed-.
-theorem leq_canc_dx: â\88\80d,e,L,L1,L2. L1 â\89\83[d, e] L â\86\92 L2 â\89\83[d, e] L â\86\92 L1 â\89\83[d, e] L2.
+theorem leq_canc_dx: â\88\80d,e,L,L1,L2. L1 ⩬[d, e] L â\86\92 L2 ⩬[d, e] L â\86\92 L1 ⩬[d, e] L2.
/3 width=3 by leq_trans, leq_sym/ qed-.
-theorem leq_join: â\88\80L1,L2,d,i. L1 â\89\83[d, i] L2 →
- â\88\80e. L1 â\89\83[i+d, e] L2 â\86\92 L1 â\89\83[d, i+e] L2.
+theorem leq_join: â\88\80L1,L2,d,i. L1 ⩬[d, i] L2 →
+ â\88\80e. L1 ⩬[i+d, e] L2 â\86\92 L1 ⩬[d, i+e] L2.
#L1 #L2 #d #i #H elim H -L1 -L2 -d -i //
[ #I #L1 #L2 #V #e #_ #IHL12 #e #H
lapply (leq_inv_succ … H ?) -H // >ypred_succ /3 width=1 by leq_pair/
(* *)
(**************************************************************************)
-include "basic_2/notation/relations/iso_2.ma".
+include "basic_2/notation/relations/topiso_2.ma".
include "basic_2/grammar/term_simple.ma".
(* SAME TOP TERM CONSTRUCTOR ************************************************)
inductive tstc: relation term ≝
| tstc_atom: ∀I. tstc (⓪{I}) (⓪{I})
- | tstc_pair: ∀I,V1,V2,T1,T2. tstc (②{I} V1. T1) (②{I} V2. T2)
+ | tstc_pair: ∀I,V1,V2,T1,T2. tstc (②{I}V1.T1) (②{I}V2.T2)
.
-interpretation "same top constructor (term)" 'Iso T1 T2 = (tstc T1 T2).
+interpretation "same top constructor (term)" 'TopIso T1 T2 = (tstc T1 T2).
(* Basic inversion lemmas ***************************************************)
-fact tstc_inv_atom1_aux: â\88\80T1,T2. T1 â\89\83 T2 → ∀I. T1 = ⓪{I} → T2 = ⓪{I}.
+fact tstc_inv_atom1_aux: â\88\80T1,T2. T1 â\89\82 T2 → ∀I. T1 = ⓪{I} → T2 = ⓪{I}.
#T1 #T2 * -T1 -T2 //
#J #V1 #V2 #T1 #T2 #I #H destruct
-qed.
+qed-.
(* Basic_1: was: iso_gen_sort iso_gen_lref *)
-lemma tstc_inv_atom1: â\88\80I,T2. â\93ª{I} â\89\83 T2 → T2 = ⓪{I}.
-/2 width=3/ qed-.
+lemma tstc_inv_atom1: â\88\80I,T2. â\93ª{I} â\89\82 T2 → T2 = ⓪{I}.
+/2 width=3 by tstc_inv_atom1_aux/ qed-.
-fact tstc_inv_pair1_aux: â\88\80T1,T2. T1 â\89\83 T2 → ∀I,W1,U1. T1 = ②{I}W1.U1 →
+fact tstc_inv_pair1_aux: â\88\80T1,T2. T1 â\89\82 T2 → ∀I,W1,U1. T1 = ②{I}W1.U1 →
∃∃W2,U2. T2 = ②{I}W2. U2.
#T1 #T2 * -T1 -T2
[ #J #I #W1 #U1 #H destruct
-| #J #V1 #V2 #T1 #T2 #I #W1 #U1 #H destruct /2 width=3/
+| #J #V1 #V2 #T1 #T2 #I #W1 #U1 #H destruct /2 width=3 by ex1_2_intro/
]
-qed.
+qed-.
(* Basic_1: was: iso_gen_head *)
-lemma tstc_inv_pair1: â\88\80I,W1,U1,T2. â\91¡{I}W1.U1 â\89\83 T2 →
+lemma tstc_inv_pair1: â\88\80I,W1,U1,T2. â\91¡{I}W1.U1 â\89\82 T2 →
∃∃W2,U2. T2 = ②{I}W2. U2.
-/2 width=5/ qed-.
+/2 width=5 by tstc_inv_pair1_aux/ qed-.
-fact tstc_inv_atom2_aux: â\88\80T1,T2. T1 â\89\83 T2 → ∀I. T2 = ⓪{I} → T1 = ⓪{I}.
+fact tstc_inv_atom2_aux: â\88\80T1,T2. T1 â\89\82 T2 → ∀I. T2 = ⓪{I} → T1 = ⓪{I}.
#T1 #T2 * -T1 -T2 //
#J #V1 #V2 #T1 #T2 #I #H destruct
-qed.
+qed-.
-lemma tstc_inv_atom2: â\88\80I,T1. T1 â\89\83 ⓪{I} → T1 = ⓪{I}.
-/2 width=3/ qed-.
+lemma tstc_inv_atom2: â\88\80I,T1. T1 â\89\82 ⓪{I} → T1 = ⓪{I}.
+/2 width=3 by tstc_inv_atom2_aux/ qed-.
-fact tstc_inv_pair2_aux: â\88\80T1,T2. T1 â\89\83 T2 → ∀I,W2,U2. T2 = ②{I}W2.U2 →
- ∃∃W1,U1. T1 = ②{I}W1. U1.
+fact tstc_inv_pair2_aux: â\88\80T1,T2. T1 â\89\82 T2 → ∀I,W2,U2. T2 = ②{I}W2.U2 →
+ ∃∃W1,U1. T1 = ②{I}W1.U1.
#T1 #T2 * -T1 -T2
[ #J #I #W2 #U2 #H destruct
-| #J #V1 #V2 #T1 #T2 #I #W2 #U2 #H destruct /2 width=3/
+| #J #V1 #V2 #T1 #T2 #I #W2 #U2 #H destruct /2 width=3 by ex1_2_intro/
]
-qed.
+qed-.
-lemma tstc_inv_pair2: â\88\80I,T1,W2,U2. T1 â\89\83 ②{I}W2.U2 →
- ∃∃W1,U1. T1 = ②{I}W1. U1.
-/2 width=5/ qed-.
+lemma tstc_inv_pair2: â\88\80I,T1,W2,U2. T1 â\89\82 ②{I}W2.U2 →
+ ∃∃W1,U1. T1 = ②{I}W1.U1.
+/2 width=5 by tstc_inv_pair2_aux/ qed-.
(* Basic properties *********************************************************)
(* Basic_1: was: iso_refl *)
-lemma tstc_refl: ∀T. T ≃ T.
+lemma tstc_refl: reflexive … tstc.
#T elim T -T //
qed.
-lemma tstc_sym: ∀T1,T2. T1 ≃ T2 → T2 ≃ T1.
+lemma tstc_sym: symmetric … tstc.
#T1 #T2 #H elim H -T1 -T2 //
-qed.
+qed-.
-lemma tstc_dec: â\88\80T1,T2. Decidable (T1 â\89\83 T2).
+lemma tstc_dec: â\88\80T1,T2. Decidable (T1 â\89\82 T2).
* #I1 [2: #V1 #T1 ] * #I2 [2,4: #V2 #T2 ]
[ elim (eq_item2_dec I1 I2) #HI12
- [ destruct /2 width=1/
+ [ destruct /2 width=1 by tstc_pair, or_introl/
| @or_intror #H
- elim (tstc_inv_pair1 … H) -H #V #T #H destruct /2 width=1/
+ elim (tstc_inv_pair1 … H) -H #V #T #H destruct /2 width=1 by/
]
| @or_intror #H
lapply (tstc_inv_atom1 … H) -H #H destruct
| @or_intror #H
lapply (tstc_inv_atom2 … H) -H #H destruct
| elim (eq_item0_dec I1 I2) #HI12
- [ destruct /2 width=1/
+ [ destruct /2 width=1 by or_introl/
| @or_intror #H
- lapply (tstc_inv_atom2 … H) -H #H destruct /2 width=1/
+ lapply (tstc_inv_atom2 … H) -H #H destruct /2 width=1 by/
]
]
qed.
-lemma simple_tstc_repl_dx: â\88\80T1,T2. T1 â\89\83 T2 → 𝐒⦃T1⦄ → 𝐒⦃T2⦄.
+lemma simple_tstc_repl_dx: â\88\80T1,T2. T1 â\89\82 T2 → 𝐒⦃T1⦄ → 𝐒⦃T2⦄.
#T1 #T2 * -T1 -T2 //
#I #V1 #V2 #T1 #T2 #H
elim (simple_inv_pair … H) -H #J #H destruct //
-qed. (**) (* remove from index *)
+qed-.
-lemma simple_tstc_repl_sn: â\88\80T1,T2. T1 â\89\83 T2 → 𝐒⦃T2⦄ → 𝐒⦃T1⦄.
-/3 width=3/ qed-.
+lemma simple_tstc_repl_sn: â\88\80T1,T2. T1 â\89\82 T2 → 𝐒⦃T2⦄ → 𝐒⦃T1⦄.
+/3 width=3 by simple_tstc_repl_dx, tstc_sym/ qed-.
(* Main properties **********************************************************)
(* Basic_1: was: iso_trans *)
-theorem tstc_trans: ∀T1,T. T1 ≃ T → ∀T2. T ≃ T2 → T1 ≃ T2.
+theorem tstc_trans: Transitive … tstc.
#T1 #T * -T1 -T //
#I #V1 #V #T1 #T #X #H
elim (tstc_inv_pair1 … H) -H #V2 #T2 #H destruct //
-qed.
+qed-.
-theorem tstc_canc_sn: â\88\80T,T1. T â\89\83 T1 â\86\92 â\88\80T2. T â\89\83 T2 â\86\92 T1 â\89\83 T2.
-/3 width=3/ qed.
+theorem tstc_canc_sn: â\88\80T,T1. T â\89\82 T1 â\86\92 â\88\80T2. T â\89\82 T2 â\86\92 T1 â\89\82 T2.
+/3 width=3 by tstc_trans, tstc_sym/ qed-.
-theorem tstc_canc_dx: â\88\80T1,T. T1 â\89\83 T â\86\92 â\88\80T2. T2 â\89\83 T â\86\92 T1 â\89\83 T2.
-/3 width=3/ qed.
+theorem tstc_canc_dx: â\88\80T1,T. T1 â\89\82 T â\86\92 â\88\80T2. T2 â\89\82 T â\86\92 T1 â\89\82 T2.
+/3 width=3 by tstc_trans, tstc_sym/ qed-.
(* Advanced inversion lemmas ************************************************)
(* Basic_1: was only: iso_flats_lref_bind_false iso_flats_flat_bind_false *)
-lemma tstc_inv_bind_appls_simple: â\88\80a,I,Vs,V2,T1,T2. â\92¶Vs.T1 â\89\83 ⓑ{a,I} V2. T2 →
+lemma tstc_inv_bind_appls_simple: â\88\80a,I,Vs,V2,T1,T2. â\92¶Vs.T1 â\89\82 ⓑ{a,I} V2. T2 →
𝐒⦃T1⦄ → ⊥.
#a #I #Vs #V2 #T1 #T2 #H
elim (tstc_inv_pair2 … H) -H #V0 #T0
elim Vs -Vs normalize
-[ #H destruct #H
- @(simple_inv_bind … H)
+[ #H destruct #H /2 width=5 by simple_inv_bind/
| #V #Vs #_ #H destruct
]
-qed.
-
+qed-.
#G #d #e #L1 #T1 #T2 #H elim H -G -L1 -T1 -T2 -d -e
[ //
| #I #G #L1 #K1 #V1 #V2 #W2 #i #d #e #Hdi #Hide #HLK1 #_ #HVW2 #IHV12 #L2 #HL12
- elim (lsuby_ldrop_trans_be … HL12 … HLK1) -HL12 -HLK1 /3 width=7 by cpysa_subst/
+ elim (lsuby_drop_trans_be … HL12 … HLK1) -HL12 -HLK1 /3 width=7 by cpysa_subst/
| /4 width=1 by lsuby_succ, cpysa_bind/
| /3 width=1 by cpysa_flat/
]
[ #I #G #L #d #e #X #H
elim (cpy_inv_atom1 … H) -H // * /2 width=7 by cpysa_subst/
| #I #G #L #K #V1 #V2 #W2 #i #d #e #Hdi #Hide #HLK #_ #HVW2 #IHV12 #T2 #H
- lapply (ldrop_fwd_drop2 … HLK) #H0LK
+ lapply (drop_fwd_drop2 … HLK) #H0LK
lapply (cpy_weak … H 0 (d+e) ? ?) -H // #H
elim (cpy_inv_lift1_be … H … H0LK … HVW2) -H -H0LK -HVW2
/3 width=7 by cpysa_subst, ylt_fwd_le_succ/
lapply (cpys_weak … HW2 0 (i+1) ? ?) -HW2 //
[ >yplus_O1 >yplus_O1 /3 width=1 by ylt_fwd_le, ylt_inj/ ] -Hi
#HW2 >(cpys_inv_lift1_eq … HW2) -HW2 //
- | elim (ldrop_O1_le (Ⓕ) … Hi) -Hi #K2 #HLK2
+ | elim (drop_O1_le (Ⓕ) … Hi) -Hi #K2 #HLK2
elim (cpys_inv_lift1_ge_up … HW2 … HLK2 … HVW2 ? ? ?) -HW2 -HLK2 -HVW2
/2 width=1 by ylt_fwd_le_succ, yle_succ_dx/ -Hdi -Hide
#X #_ #H elim (lift_inv_lref2_be … H) -H //
| #U #U1 #_ #HU1 #IHU #U2 #HU12
elim (lift_total U 0 (i+1)) #U0 #HU0
lapply (IHU … HU0) -IHU #H
- lapply (ldrop_fwd_drop2 … HLK) -HLK #HLK
+ lapply (drop_fwd_drop2 … HLK) -HLK #HLK
lapply (cpy_lift_ge … HU1 … HLK HU0 HU12 ?) -HU1 -HLK -HU0 -HU12 // #HU02
lapply (cpy_weak … HU02 d e ? ?) -HU02
[2,3: /2 width=3 by cpys_strap1, yle_succ_dx/ ]
[ #H destruct
elim (cpy_inv_atom1 … HT2) -HT2 [ /2 width=1 by or_introl/ | * /3 width=11 by ex6_5_intro, or_intror/ ]
| * #J #K #V1 #V #i #Hdi #Hide #HLK #HV1 #HVT #HI
- lapply (ldrop_fwd_drop2 … HLK) #H
+ lapply (drop_fwd_drop2 … HLK) #H
elim (cpy_inv_lift1_ge_up … HT2 … H … HVT) -HT2 -H -HVT
[2,3,4: /2 width=1 by ylt_fwd_le_succ, yle_succ_dx/ ]
/4 width=11 by cpys_strap1, ex6_5_intro, or_intror/
* >yminus_Y_inj /3 width=7 by or_intror, ex4_4_intro/
qed-.
-lemma cpys_inv_lref1_ldrop: ∀G,L,T2,i,d,e. ⦃G, L⦄ ⊢ #i ▶*[d, e] T2 →
+lemma cpys_inv_lref1_drop: ∀G,L,T2,i,d,e. ⦃G, L⦄ ⊢ #i ▶*[d, e] T2 →
∀I,K,V1. ⇩[i] L ≡ K.ⓑ{I}V1 →
∀V2. ⇧[O, i+1] V2 ≡ T2 →
∧∧ ⦃G, K⦄ ⊢ V1 ▶*[0, ⫰(d+e-i)] V2
[ #H destruct elim (lift_inv_lref2_be … HVT2) -HVT2 -HLK //
| * #Z #Y #X1 #X2 #Hdi #Hide #HLY #HX12 #HXT2
lapply (lift_inj … HXT2 … HVT2) -T2 #H destruct
- lapply (ldrop_mono … HLY … HLK) -L #H destruct
+ lapply (drop_mono … HLY … HLK) -L #H destruct
/2 width=1 by and3_intro/
]
qed-.
--- /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/notation/relations/rdropstar_3.ma".
+include "basic_2/notation/relations/rdropstar_4.ma".
+include "basic_2/substitution/drop.ma".
+include "basic_2/multiple/gr2_minus.ma".
+include "basic_2/multiple/lifts.ma".
+
+(* ITERATED LOCAL ENVIRONMENT SLICING ***************************************)
+
+inductive drops (s:bool): list2 nat nat → relation lenv ≝
+| drops_nil : ∀L. drops s (⟠) L L
+| drops_cons: ∀L1,L,L2,des,d,e.
+ drops s des L1 L → ⇩[s, d, e] L ≡ L2 → drops s ({d, e} @ des) L1 L2
+.
+
+interpretation "iterated slicing (local environment) abstract"
+ 'RDropStar s des T1 T2 = (drops s des T1 T2).
+(*
+interpretation "iterated slicing (local environment) general"
+ 'RDropStar des T1 T2 = (drops true des T1 T2).
+*)
+
+(* Basic inversion lemmas ***************************************************)
+
+fact drops_inv_nil_aux: ∀L1,L2,s,des. ⇩*[s, des] L1 ≡ L2 → des = ⟠ → L1 = L2.
+#L1 #L2 #s #des * -L1 -L2 -des //
+#L1 #L #L2 #d #e #des #_ #_ #H destruct
+qed-.
+
+(* Basic_1: was: drop1_gen_pnil *)
+lemma drops_inv_nil: ∀L1,L2,s. ⇩*[s, ⟠] L1 ≡ L2 → L1 = L2.
+/2 width=4 by drops_inv_nil_aux/ qed-.
+
+fact drops_inv_cons_aux: ∀L1,L2,s,des. ⇩*[s, des] L1 ≡ L2 →
+ ∀d,e,tl. des = {d, e} @ tl →
+ ∃∃L. ⇩*[s, tl] L1 ≡ L & ⇩[s, d, e] L ≡ L2.
+#L1 #L2 #s #des * -L1 -L2 -des
+[ #L #d #e #tl #H destruct
+| #L1 #L #L2 #des #d #e #HT1 #HT2 #hd #he #tl #H destruct
+ /2 width=3 by ex2_intro/
+]
+qed-.
+
+(* Basic_1: was: drop1_gen_pcons *)
+lemma drops_inv_cons: ∀L1,L2,s,d,e,des. ⇩*[s, {d, e} @ des] L1 ≡ L2 →
+ ∃∃L. ⇩*[s, des] L1 ≡ L & ⇩[s, d, e] L ≡ L2.
+/2 width=3 by drops_inv_cons_aux/ qed-.
+
+lemma drops_inv_skip2: ∀I,s,des,des2,i. des ▭ i ≡ des2 →
+ ∀L1,K2,V2. ⇩*[s, des2] L1 ≡ K2. ⓑ{I} V2 →
+ ∃∃K1,V1,des1. des + 1 ▭ i + 1 ≡ des1 + 1 &
+ ⇩*[s, des1] K1 ≡ K2 &
+ ⇧*[des1] V2 ≡ V1 &
+ L1 = K1. ⓑ{I} V1.
+#I #s #des #des2 #i #H elim H -des -des2 -i
+[ #i #L1 #K2 #V2 #H
+ >(drops_inv_nil … H) -L1 /2 width=7 by lifts_nil, minuss_nil, ex4_3_intro, drops_nil/
+| #des #des2 #d #e #i #Hid #_ #IHdes2 #L1 #K2 #V2 #H
+ elim (drops_inv_cons … H) -H #L #HL1 #H
+ elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ #K #V >minus_plus #HK2 #HV2 #H destruct
+ elim (IHdes2 … HL1) -IHdes2 -HL1 #K1 #V1 #des1 #Hdes1 #HK1 #HV1 #X destruct
+ @(ex4_3_intro … K1 V1 … ) // [3,4: /2 width=7 by lifts_cons, drops_cons/ | skip ]
+ normalize >plus_minus /3 width=1 by minuss_lt, lt_minus_to_plus/ (**) (* explicit constructors *)
+| #des #des2 #d #e #i #Hid #_ #IHdes2 #L1 #K2 #V2 #H
+ elim (IHdes2 … H) -IHdes2 -H #K1 #V1 #des1 #Hdes1 #HK1 #HV1 #X destruct
+ /4 width=7 by minuss_ge, ex4_3_intro, le_S_S/
+]
+qed-.
+
+(* Basic properties *********************************************************)
+
+(* Basic_1: was: drop1_skip_bind *)
+lemma drops_skip: ∀L1,L2,s,des. ⇩*[s, des] L1 ≡ L2 → ∀V1,V2. ⇧*[des] V2 ≡ V1 →
+ ∀I. ⇩*[s, des + 1] L1.ⓑ{I}V1 ≡ L2.ⓑ{I}V2.
+#L1 #L2 #s #des #H elim H -L1 -L2 -des
+[ #L #V1 #V2 #HV12 #I
+ >(lifts_inv_nil … HV12) -HV12 //
+| #L1 #L #L2 #des #d #e #_ #HL2 #IHL #V1 #V2 #H #I
+ elim (lifts_inv_cons … H) -H /3 width=5 by drop_skip, drops_cons/
+].
+qed.
+
+(* Basic_1: removed theorems 1: drop1_getl_trans *)
--- /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/substitution/drop_drop.ma".
+include "basic_2/multiple/drops.ma".
+
+(* ITERATED LOCAL ENVIRONMENT SLICING ***************************************)
+
+(* Properties concerning basic local environment slicing ********************)
+
+lemma drops_drop_trans: ∀L1,L,des. ⇩*[Ⓕ, des] L1 ≡ L → ∀L2,i. ⇩[i] L ≡ L2 →
+ ∃∃L0,des0,i0. ⇩[i0] L1 ≡ L0 & ⇩*[Ⓕ, des0] L0 ≡ L2 &
+ @⦃i, des⦄ ≡ i0 & des ▭ i ≡ des0.
+#L1 #L #des #H elim H -L1 -L -des
+[ /2 width=7 by drops_nil, minuss_nil, at_nil, ex4_3_intro/
+| #L1 #L0 #L #des #d #e #_ #HL0 #IHL0 #L2 #i #HL2
+ elim (lt_or_ge i d) #Hid
+ [ elim (drop_trans_le … HL0 … HL2) -L /2 width=2 by lt_to_le/
+ #L #HL0 #HL2 elim (IHL0 … HL0) -L0 /3 width=7 by drops_cons, minuss_lt, at_lt, ex4_3_intro/
+ | lapply (drop_trans_ge … HL0 … HL2 ?) -L // #HL02
+ elim (IHL0 … HL02) -L0 /3 width=7 by minuss_ge, at_ge, ex4_3_intro/
+ ]
+]
+qed-.
--- /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/multiple/drops_drop.ma".
+
+(* ITERATED LOCAL ENVIRONMENT SLICING ***************************************)
+
+(* Main properties **********************************************************)
+
+(* Basic_1: was: drop1_trans *)
+theorem drops_trans: ∀L,L2,s,des2. ⇩*[s, des2] L ≡ L2 → ∀L1,des1. ⇩*[s, des1] L1 ≡ L →
+ ⇩*[s, des2 @@ des1] L1 ≡ L2.
+#L #L2 #s #des2 #H elim H -L -L2 -des2 /3 width=3 by drops_cons/
+qed-.
⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄.
/2 width=5 by tri_TC_strap/ qed.
-lemma fqup_ldrop: ∀G1,G2,L1,K1,K2,T1,T2,U1,e. ⇩[e] L1 ≡ K1 → ⇧[0, e] T1 ≡ U1 →
+lemma fqup_drop: ∀G1,G2,L1,K1,K2,T1,T2,U1,e. ⇩[e] L1 ≡ K1 → ⇧[0, e] T1 ≡ U1 →
⦃G1, K1, T1⦄ ⊐+ ⦃G2, K2, T2⦄ → ⦃G1, L1, U1⦄ ⊐+ ⦃G2, K2, T2⦄.
#G1 #G2 #L1 #K1 #K2 #T1 #T2 #U1 #e #HLK1 #HTU1 #HT12 elim (eq_or_gt … e) #H destruct
-[ >(ldrop_inv_O2 … HLK1) -L1 <(lift_inv_O2 … HTU1) -U1 //
+[ >(drop_inv_O2 … HLK1) -L1 <(lift_inv_O2 … HTU1) -U1 //
| /3 width=5 by fqup_strap2, fqu_drop_lt/
]
qed-.
lemma fqup_lref: ∀I,G,L,K,V,i. ⇩[i] L ≡ K.ⓑ{I}V → ⦃G, L, #i⦄ ⊐+ ⦃G, K, V⦄.
-/3 width=6 by fqu_lref_O, fqu_fqup, lift_lref_ge, fqup_ldrop/ qed.
+/3 width=6 by fqu_lref_O, fqu_fqup, lift_lref_ge, fqup_drop/ qed.
lemma fqup_pair_sn: ∀I,G,L,V,T. ⦃G, L, ②{I}V.T⦄ ⊐+ ⦃G, L, V⦄.
/2 width=1 by fqu_pair_sn, fqu_fqup/ qed.
⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄.
/2 width=5 by tri_TC_strap/ qed-.
-lemma fqus_ldrop: ∀G1,G2,K1,K2,T1,T2. ⦃G1, K1, T1⦄ ⊐* ⦃G2, K2, T2⦄ →
+lemma fqus_drop: ∀G1,G2,K1,K2,T1,T2. ⦃G1, K1, T1⦄ ⊐* ⦃G2, K2, T2⦄ →
∀L1,U1,e. ⇩[e] L1 ≡ K1 → ⇧[0, e] T1 ≡ U1 →
⦃G1, L1, U1⦄ ⊐* ⦃G2, K2, T2⦄.
#G1 #G2 #K1 #K2 #T1 #T2 #H @(fqus_ind … H) -G2 -K2 -T2
include "ground_2/ynat/ynat_plus.ma".
include "basic_2/notation/relations/freestar_4.ma".
include "basic_2/substitution/lift_neg.ma".
-include "basic_2/substitution/ldrop.ma".
+include "basic_2/substitution/drop.ma".
(* CONTEXT-SENSITIVE FREE VARIABLES *****************************************)
lemma frees_inv_lref_free: ∀L,d,j,i. L ⊢ i ϵ 𝐅*[d]⦃#j⦄ → |L| ≤ j → j = i.
#L #d #j #i #H #Hj elim (frees_inv_lref … H) -H //
-* #I #K #W #_ #_ #HLK lapply (ldrop_fwd_length_lt2 … HLK) -I
+* #I #K #W #_ #_ #HLK lapply (drop_fwd_length_lt2 … HLK) -I
#H elim (lt_refl_false j) /2 width=3 by lt_to_le_to_lt/
qed-.
| * #I #K #W #j #Hdj #Hji #HnX #HLK #HW elim (nlift_inv_bind … HnX) -HnX
[ /4 width=9 by frees_be, or_introl/
| #HnT @or_intror @(frees_be … HnT) -HnT
- [4,5,6: /2 width=1 by ldrop_drop, yle_succ, lt_minus_to_plus/
+ [4,5,6: /2 width=1 by drop_drop, yle_succ, lt_minus_to_plus/
|7: >minus_plus_plus_l //
|*: skip
]
elim (yle_inv_succ1 … Hdj) -Hdj <yminus_SO2 #Hyj #H
lapply (ylt_O … H) -H #Hj
>(plus_minus_m_m j 1) in HnU; // <minus_le_minus_minus_comm in HW;
- /4 width=9 by frees_be, nlift_bind_dx, ldrop_inv_drop1_lt, lt_plus_to_minus/
+ /4 width=9 by frees_be, nlift_bind_dx, drop_inv_drop1_lt, lt_plus_to_minus/
]
qed.
(* *)
(**************************************************************************)
-include "basic_2/substitution/ldrop_append.ma".
+include "basic_2/substitution/drop_append.ma".
include "basic_2/multiple/frees.ma".
(* CONTEXT-SENSITIVE FREE VARIABLES *****************************************)
∀L1. L1 @@ L2 ⊢ i ϵ 𝐅*[d]⦃U⦄.
#L2 #U #d #i #H elim H -L2 -U -d -i /3 width=2 by frees_eq/
#I #L2 #K2 #U #W #d #i #j #Hdj #Hji #HnU #HLK2 #_ #IHW #Hi #L1
-lapply (ldrop_fwd_length_minus2 … HLK2) normalize #H0
-lapply (ldrop_O1_append_sn_le … HLK2 … L1) -HLK2
+lapply (drop_fwd_length_minus2 … HLK2) normalize #H0
+lapply (drop_O1_append_sn_le … HLK2 … L1) -HLK2
[ -I -L1 -K2 -U -W -d /3 width=3 by lt_to_le, lt_to_le_to_lt/
| #HLK2 @(frees_be … HnU HLK2) // -HnU -HLK2 @IHW -IHW
>(minus_plus_m_m (|K2|) 1) >H0 -H0 /2 width=1 by monotonic_le_minus_l2/
i ≤ |L2| → L2 ⊢ i ϵ 𝐅*[d]⦃U⦄.
#L #U #d #i #H elim H -L -U -d -i /3 width=2 by frees_eq/
#Z #L #Y #U #X #d #i #j #Hdj #Hji #HnU #HLY #_ #IHW #L1 #L2 #H #Hi destruct
-elim (ldrop_O1_lt (Ⓕ) L2 j) [2: -Z -Y -L1 -X -U -d /2 width=3 by lt_to_le_to_lt/ ]
-#I #K2 #W #HLK2 lapply (ldrop_fwd_length_minus2 … HLK2) normalize #H0
-lapply (ldrop_O1_inv_append1_le … HLY … HLK2) -HLY
+elim (drop_O1_lt (Ⓕ) L2 j) [2: -Z -Y -L1 -X -U -d /2 width=3 by lt_to_le_to_lt/ ]
+#I #K2 #W #HLK2 lapply (drop_fwd_length_minus2 … HLK2) normalize #H0
+lapply (drop_O1_inv_append1_le … HLY … HLK2) -HLY
[ -Z -I -Y -K2 -L1 -X -U -W -d /3 width=3 by lt_to_le, lt_to_le_to_lt/
| normalize #H destruct
@(frees_be … HnU HLK2) -HnU -HLK2 // @IHW -IHW //
(* *)
(**************************************************************************)
-include "basic_2/substitution/ldrop_leq.ma".
+include "basic_2/substitution/drop_leq.ma".
include "basic_2/multiple/frees.ma".
(* CONTEXT-SENSITIVE FREE VARIABLES *****************************************)
(* Properties on equivalence for local environments *************************)
lemma leq_frees_trans: ∀L2,U,d,i. L2 ⊢ i ϵ 𝐅*[d]⦃U⦄ →
- â\88\80L1. L1 â\89\83[d, ∞] L2 → L1 ⊢ i ϵ 𝐅*[d]⦃U⦄.
+ â\88\80L1. L1 ⩬[d, ∞] L2 → L1 ⊢ i ϵ 𝐅*[d]⦃U⦄.
#L2 #U #d #i #H elim H -L2 -U -d -i /3 width=2 by frees_eq/
#I2 #L2 #K2 #U #W2 #d #i #j #Hdj #Hji #HnU #HLK2 #_ #IHW2 #L1 #HL12
-elim (leq_ldrop_trans_be … HL12 … HLK2) -L2 // >yminus_Y_inj #K1 #HK12 #HLK1
+elim (leq_drop_trans_be … HL12 … HLK2) -L2 // >yminus_Y_inj #K1 #HK12 #HLK1
lapply (leq_inv_O_Y … HK12) -HK12 #H destruct /3 width=9 by frees_be/
qed-.
lemma frees_leq_conf: ∀L1,U,d,i. L1 ⊢ i ϵ 𝐅*[d]⦃U⦄ →
- â\88\80L2. L1 â\89\83[d, ∞] L2 → L2 ⊢ i ϵ 𝐅*[d]⦃U⦄.
+ â\88\80L2. L1 ⩬[d, ∞] L2 → L2 ⊢ i ϵ 𝐅*[d]⦃U⦄.
/3 width=3 by leq_sym, leq_frees_trans/ qed-.
(* *)
(**************************************************************************)
-include "basic_2/substitution/ldrop_ldrop.ma".
+include "basic_2/substitution/drop_drop.ma".
include "basic_2/multiple/frees.ma".
(* CONTEXT-SENSITIVE FREE VARIABLES *****************************************)
[ -n @or_intror #H elim (lt_refl_false i)
lapply (frees_inv_lref_skip … H ?) -L //
| elim (lt_or_ge j (|L|)) #Hj
- [ elim (ldrop_O1_lt (Ⓕ) L j) // -Hj #I #K #W #HLK destruct
- elim (IH K W … 0 (i-j-1)) -IH [1,3: /3 width=5 by frees_lref_be, ldrop_fwd_rfw, or_introl/ ] #HnW
+ [ elim (drop_O1_lt (Ⓕ) L j) // -Hj #I #K #W #HLK destruct
+ elim (IH K W … 0 (i-j-1)) -IH [1,3: /3 width=5 by frees_lref_be, drop_fwd_rfw, or_introl/ ] #HnW
@or_intror #H elim (frees_inv_lref_lt … H) // #Z #Y #X #_ #HLY -d
- lapply (ldrop_mono … HLY … HLK) -L #H destruct /2 width=1 by/
+ lapply (drop_mono … HLY … HLK) -L #H destruct /2 width=1 by/
| -n @or_intror #H elim (lt_refl_false i)
lapply (frees_inv_lref_free … H ?) -d //
]
elim (le_to_or_lt_eq … Hdj) -Hdj
[ -I0 -K0 -W0 /3 width=9 by frees_be, yle_inj/
| -Hji -HnU #H destruct
- lapply (ldrop_mono … HLK0 … HLK) #H destruct -I
+ lapply (drop_mono … HLK0 … HLK) #H destruct -I
elim HnW0 -L -U -HnW0 //
]
qed.
@frees_eq #X #HXU elim (lift_div_le … HTU … HXU) -U /2 width=2 by/
| #I #K #K0 #T #V #d #i #j #Hdj #Hji #HnT #HK0 #HV #IHV #L #s #d0 #e0 #HLK #U #HTU #Hd0i
elim (lt_or_ge j d0) #H1
- [ elim (ldrop_trans_lt … HLK … HK0) // -K #L0 #W #HL0 #HLK0 #HVW
+ [ elim (drop_trans_lt … HLK … HK0) // -K #L0 #W #HL0 #HLK0 #HVW
@(frees_be … HL0) -HL0 -HV
/3 width=3 by lt_plus_to_minus_r, lt_to_le_to_lt/
[ #X #HXU >(plus_minus_m_m d0 1) in HTU; /2 width=2 by ltn_to_ltO/ #HTU
| >minus_plus <plus_minus // <minus_plus
/3 width=5 by monotonic_le_minus_l2/
]
- | lapply (ldrop_trans_ge … HLK … HK0 ?) // -K #HLK0
- lapply (ldrop_inv_gen … HLK0) >commutative_plus -HLK0 #HLK0
+ | lapply (drop_trans_ge … HLK … HK0 ?) // -K #HLK0
+ lapply (drop_inv_gen … HLK0) >commutative_plus -HLK0 #HLK0
@(frees_be … HLK0) -HLK0 -IHV
/2 width=1 by yle_plus_dx1_trans, lt_minus_to_plus/
#X #HXU elim (lift_div_le … HTU … HXU) -U /2 width=2 by/
elim (lift_split … HTU i e0) -HTU /2 width=2 by/
| #I #L #K0 #U #W #d #i #j #Hdi #Hij #HnU #HLK0 #_ #IHW #K #s #d0 #e0 #HLK #T #HTU #Hd0i #Hide0
elim (lt_or_ge j d0) #H1
- [ elim (ldrop_conf_lt … HLK … HLK0) -L // #L0 #V #H #HKL0 #HVW
+ [ elim (drop_conf_lt … HLK … HLK0) -L // #L0 #V #H #HKL0 #HVW
@(IHW … HKL0 … HVW)
[ /2 width=1 by monotonic_le_minus_l2/
| >minus_plus >minus_plus >plus_minus /2 width=1 by monotonic_le_minus_l/
elim (lift_trans_le … HXT … HTU) -T // <plus_minus_m_m /2 width=2 by/
| #I #L #K0 #U #W #d #i #j #Hdi #Hij #HnU #HLK0 #_ #IHW #K #s #d0 #e0 #HLK #T #HTU #Hde0i
elim (lt_or_ge j d0) #H1
- [ elim (ldrop_conf_lt … HLK … HLK0) -L // #L0 #V #H #HKL0 #HVW
+ [ elim (drop_conf_lt … HLK … HLK0) -L // #L0 #V #H #HKL0 #HVW
elim (le_inv_plus_l … Hde0i) #H0 #He0i
@(frees_be … H) -H
[ /3 width=1 by yle_plus_dx1_trans, monotonic_yle_minus_dx/
#X #_ #H elim (HnU … H)
| >commutative_plus /3 width=1 by le_minus_to_plus, monotonic_pred/
]
- | lapply (ldrop_conf_ge … HLK … HLK0 ?) // -L #HK0
+ | lapply (drop_conf_ge … HLK … HLK0 ?) // -L #HK0
elim (le_inv_plus_l … H2) -H2 #H2 #He0j
@(frees_be … HK0)
[ /2 width=1 by monotonic_yle_minus_dx/
+++ /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/notation/relations/rdropstar_3.ma".
-include "basic_2/notation/relations/rdropstar_4.ma".
-include "basic_2/substitution/ldrop.ma".
-include "basic_2/multiple/gr2_minus.ma".
-include "basic_2/multiple/lifts.ma".
-
-(* ITERATED LOCAL ENVIRONMENT SLICING ***************************************)
-
-inductive ldrops (s:bool): list2 nat nat → relation lenv ≝
-| ldrops_nil : ∀L. ldrops s (⟠) L L
-| ldrops_cons: ∀L1,L,L2,des,d,e.
- ldrops s des L1 L → ⇩[s, d, e] L ≡ L2 → ldrops s ({d, e} @ des) L1 L2
-.
-
-interpretation "iterated slicing (local environment) abstract"
- 'RDropStar s des T1 T2 = (ldrops s des T1 T2).
-(*
-interpretation "iterated slicing (local environment) general"
- 'RDropStar des T1 T2 = (ldrops true des T1 T2).
-*)
-
-(* Basic inversion lemmas ***************************************************)
-
-fact ldrops_inv_nil_aux: ∀L1,L2,s,des. ⇩*[s, des] L1 ≡ L2 → des = ⟠ → L1 = L2.
-#L1 #L2 #s #des * -L1 -L2 -des //
-#L1 #L #L2 #d #e #des #_ #_ #H destruct
-qed-.
-
-(* Basic_1: was: drop1_gen_pnil *)
-lemma ldrops_inv_nil: ∀L1,L2,s. ⇩*[s, ⟠] L1 ≡ L2 → L1 = L2.
-/2 width=4 by ldrops_inv_nil_aux/ qed-.
-
-fact ldrops_inv_cons_aux: ∀L1,L2,s,des. ⇩*[s, des] L1 ≡ L2 →
- ∀d,e,tl. des = {d, e} @ tl →
- ∃∃L. ⇩*[s, tl] L1 ≡ L & ⇩[s, d, e] L ≡ L2.
-#L1 #L2 #s #des * -L1 -L2 -des
-[ #L #d #e #tl #H destruct
-| #L1 #L #L2 #des #d #e #HT1 #HT2 #hd #he #tl #H destruct
- /2 width=3 by ex2_intro/
-]
-qed-.
-
-(* Basic_1: was: drop1_gen_pcons *)
-lemma ldrops_inv_cons: ∀L1,L2,s,d,e,des. ⇩*[s, {d, e} @ des] L1 ≡ L2 →
- ∃∃L. ⇩*[s, des] L1 ≡ L & ⇩[s, d, e] L ≡ L2.
-/2 width=3 by ldrops_inv_cons_aux/ qed-.
-
-lemma ldrops_inv_skip2: ∀I,s,des,des2,i. des ▭ i ≡ des2 →
- ∀L1,K2,V2. ⇩*[s, des2] L1 ≡ K2. ⓑ{I} V2 →
- ∃∃K1,V1,des1. des + 1 ▭ i + 1 ≡ des1 + 1 &
- ⇩*[s, des1] K1 ≡ K2 &
- ⇧*[des1] V2 ≡ V1 &
- L1 = K1. ⓑ{I} V1.
-#I #s #des #des2 #i #H elim H -des -des2 -i
-[ #i #L1 #K2 #V2 #H
- >(ldrops_inv_nil … H) -L1 /2 width=7 by lifts_nil, minuss_nil, ex4_3_intro, ldrops_nil/
-| #des #des2 #d #e #i #Hid #_ #IHdes2 #L1 #K2 #V2 #H
- elim (ldrops_inv_cons … H) -H #L #HL1 #H
- elim (ldrop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ #K #V >minus_plus #HK2 #HV2 #H destruct
- elim (IHdes2 … HL1) -IHdes2 -HL1 #K1 #V1 #des1 #Hdes1 #HK1 #HV1 #X destruct
- @(ex4_3_intro … K1 V1 … ) // [3,4: /2 width=7 by lifts_cons, ldrops_cons/ | skip ]
- normalize >plus_minus /3 width=1 by minuss_lt, lt_minus_to_plus/ (**) (* explicit constructors *)
-| #des #des2 #d #e #i #Hid #_ #IHdes2 #L1 #K2 #V2 #H
- elim (IHdes2 … H) -IHdes2 -H #K1 #V1 #des1 #Hdes1 #HK1 #HV1 #X destruct
- /4 width=7 by minuss_ge, ex4_3_intro, le_S_S/
-]
-qed-.
-
-(* Basic properties *********************************************************)
-
-(* Basic_1: was: drop1_skip_bind *)
-lemma ldrops_skip: ∀L1,L2,s,des. ⇩*[s, des] L1 ≡ L2 → ∀V1,V2. ⇧*[des] V2 ≡ V1 →
- ∀I. ⇩*[s, des + 1] L1.ⓑ{I}V1 ≡ L2.ⓑ{I}V2.
-#L1 #L2 #s #des #H elim H -L1 -L2 -des
-[ #L #V1 #V2 #HV12 #I
- >(lifts_inv_nil … HV12) -HV12 //
-| #L1 #L #L2 #des #d #e #_ #HL2 #IHL #V1 #V2 #H #I
- elim (lifts_inv_cons … H) -H /3 width=5 by ldrop_skip, ldrops_cons/
-].
-qed.
-
-(* Basic_1: removed theorems 1: drop1_getl_trans *)
+++ /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/substitution/ldrop_ldrop.ma".
-include "basic_2/multiple/ldrops.ma".
-
-(* ITERATED LOCAL ENVIRONMENT SLICING ***************************************)
-
-(* Properties concerning basic local environment slicing ********************)
-
-lemma ldrops_ldrop_trans: ∀L1,L,des. ⇩*[Ⓕ, des] L1 ≡ L → ∀L2,i. ⇩[i] L ≡ L2 →
- ∃∃L0,des0,i0. ⇩[i0] L1 ≡ L0 & ⇩*[Ⓕ, des0] L0 ≡ L2 &
- @⦃i, des⦄ ≡ i0 & des ▭ i ≡ des0.
-#L1 #L #des #H elim H -L1 -L -des
-[ /2 width=7 by ldrops_nil, minuss_nil, at_nil, ex4_3_intro/
-| #L1 #L0 #L #des #d #e #_ #HL0 #IHL0 #L2 #i #HL2
- elim (lt_or_ge i d) #Hid
- [ elim (ldrop_trans_le … HL0 … HL2) -L /2 width=2 by lt_to_le/
- #L #HL0 #HL2 elim (IHL0 … HL0) -L0 /3 width=7 by ldrops_cons, minuss_lt, at_lt, ex4_3_intro/
- | lapply (ldrop_trans_ge … HL0 … HL2 ?) -L // #HL02
- elim (IHL0 … HL02) -L0 /3 width=7 by minuss_ge, at_ge, ex4_3_intro/
- ]
-]
-qed-.
-
+++ /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/multiple/ldrops_ldrop.ma".
-
-(* ITERATED LOCAL ENVIRONMENT SLICING ***************************************)
-
-(* Main properties **********************************************************)
-
-(* Basic_1: was: drop1_trans *)
-theorem ldrops_trans: ∀L,L2,s,des2. ⇩*[s, des2] L ≡ L2 → ∀L1,des1. ⇩*[s, des1] L1 ≡ L →
- ⇩*[s, des2 @@ des1] L1 ≡ L2.
-#L #L2 #s #des2 #H elim H -L -L2 -des2 /3 width=3 by ldrops_cons/
-qed-.
* /3 width=7 by or3_intro2, ex4_4_intro/
qed-.
-lemma lleq_fwd_ldrop_sn: ∀L1,L2,T,d. L1 ≡[d, T] L2 → ∀K1,i. ⇩[i] L1 ≡ K1 →
+lemma lleq_fwd_drop_sn: ∀L1,L2,T,d. L1 ≡[d, T] L2 → ∀K1,i. ⇩[i] L1 ≡ K1 →
∃K2. ⇩[i] L2 ≡ K2.
-/2 width=7 by llpx_sn_fwd_ldrop_sn/ qed-.
+/2 width=7 by llpx_sn_fwd_drop_sn/ qed-.
-lemma lleq_fwd_ldrop_dx: ∀L1,L2,T,d. L1 ≡[d, T] L2 → ∀K2,i. ⇩[i] L2 ≡ K2 →
+lemma lleq_fwd_drop_dx: ∀L1,L2,T,d. L1 ≡[d, T] L2 → ∀K2,i. ⇩[i] L2 ≡ K2 →
∃K1. ⇩[i] L1 ≡ K1.
-/2 width=7 by llpx_sn_fwd_ldrop_dx/ qed-.
+/2 width=7 by llpx_sn_fwd_drop_dx/ qed-.
lemma lleq_fwd_bind_sn: ∀a,I,L1,L2,V,T,d.
L1 ≡[ⓑ{a,I}V.T, d] L2 → L1 ≡[V, d] L2.
--- /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/multiple/llpx_sn_drop.ma".
+include "basic_2/multiple/lleq.ma".
+
+(* LAZY EQUIVALENCE FOR LOCAL ENVIRONMENTS **********************************)
+
+(* Advanced properties ******************************************************)
+
+lemma lleq_bind_repl_O: ∀I,L1,L2,V,T. L1.ⓑ{I}V ≡[T, 0] L2.ⓑ{I}V →
+ ∀J,W. L1 ≡[W, 0] L2 → L1.ⓑ{J}W ≡[T, 0] L2.ⓑ{J}W.
+/2 width=7 by llpx_sn_bind_repl_O/ qed-.
+
+lemma lleq_dec: ∀T,L1,L2,d. Decidable (L1 ≡[T, d] L2).
+/3 width=1 by llpx_sn_dec, eq_term_dec/ qed-.
+
+lemma lleq_llpx_sn_trans: ∀R. lleq_transitive R →
+ ∀L1,L2,T,d. L1 ≡[T, d] L2 →
+ ∀L. llpx_sn R d T L2 L → llpx_sn R d T L1 L.
+#R #HR #L1 #L2 #T #d #H @(lleq_ind … H) -L1 -L2 -T -d
+[1,2,5: /4 width=6 by llpx_sn_fwd_length, llpx_sn_gref, llpx_sn_skip, llpx_sn_sort, trans_eq/
+|4: /4 width=6 by llpx_sn_fwd_length, llpx_sn_free, le_repl_sn_conf_aux, trans_eq/
+| #I #L1 #L2 #K1 #K2 #V #d #i #Hdi #HLK1 #HLK2 #HK12 #IHK12 #L #H elim (llpx_sn_inv_lref_ge_sn … H … HLK2) -H -HLK2
+ /3 width=11 by llpx_sn_lref/
+| #a #I #L1 #L2 #V #T #d #_ #_ #IHV #IHT #L #H elim (llpx_sn_inv_bind … H) -H
+ /3 width=1 by llpx_sn_bind/
+| #I #L1 #L2 #V #T #d #_ #_ #IHV #IHT #L #H elim (llpx_sn_inv_flat … H) -H
+ /3 width=1 by llpx_sn_flat/
+]
+qed-.
+
+lemma lleq_llpx_sn_conf: ∀R. lleq_transitive R →
+ ∀L1,L2,T,d. L1 ≡[T, d] L2 →
+ ∀L. llpx_sn R d T L1 L → llpx_sn R d T L2 L.
+/3 width=3 by lleq_llpx_sn_trans, lleq_sym/ qed-.
+
+(* Advanced inversion lemmas ************************************************)
+
+lemma lleq_inv_lref_ge_dx: ∀L1,L2,d,i. L1 ≡[#i, d] L2 → d ≤ i →
+ ∀I,K2,V. ⇩[i] L2 ≡ K2.ⓑ{I}V →
+ ∃∃K1. ⇩[i] L1 ≡ K1.ⓑ{I}V & K1 ≡[V, 0] K2.
+#L1 #L2 #d #i #H #Hdi #I #K2 #V #HLK2 elim (llpx_sn_inv_lref_ge_dx … H … HLK2) -L2
+/2 width=3 by ex2_intro/
+qed-.
+
+lemma lleq_inv_lref_ge_sn: ∀L1,L2,d,i. L1 ≡[#i, d] L2 → d ≤ i →
+ ∀I,K1,V. ⇩[i] L1 ≡ K1.ⓑ{I}V →
+ ∃∃K2. ⇩[i] L2 ≡ K2.ⓑ{I}V & K1 ≡[V, 0] K2.
+#L1 #L2 #d #i #H #Hdi #I1 #K1 #V #HLK1 elim (llpx_sn_inv_lref_ge_sn … H … HLK1) -L1
+/2 width=3 by ex2_intro/
+qed-.
+
+lemma lleq_inv_lref_ge_bi: ∀L1,L2,d,i. L1 ≡[#i, d] L2 → d ≤ i →
+ ∀I1,I2,K1,K2,V1,V2.
+ ⇩[i] L1 ≡ K1.ⓑ{I1}V1 → ⇩[i] L2 ≡ K2.ⓑ{I2}V2 →
+ ∧∧ I1 = I2 & K1 ≡[V1, 0] K2 & V1 = V2.
+/2 width=8 by llpx_sn_inv_lref_ge_bi/ qed-.
+
+lemma lleq_inv_lref_ge: ∀L1,L2,d,i. L1 ≡[#i, d] L2 → d ≤ i →
+ ∀I,K1,K2,V. ⇩[i] L1 ≡ K1.ⓑ{I}V → ⇩[i] L2 ≡ K2.ⓑ{I}V →
+ K1 ≡[V, 0] K2.
+#L1 #L2 #d #i #HL12 #Hdi #I #K1 #K2 #V #HLK1 #HLK2
+elim (lleq_inv_lref_ge_bi … HL12 … HLK1 HLK2) //
+qed-.
+
+lemma lleq_inv_S: ∀L1,L2,T,d. L1 ≡[T, d+1] L2 →
+ ∀I,K1,K2,V. ⇩[d] L1 ≡ K1.ⓑ{I}V → ⇩[d] L2 ≡ K2.ⓑ{I}V →
+ K1 ≡[V, 0] K2 → L1 ≡[T, d] L2.
+/2 width=9 by llpx_sn_inv_S/ qed-.
+
+lemma lleq_inv_bind_O: ∀a,I,L1,L2,V,T. L1 ≡[ⓑ{a,I}V.T, 0] L2 →
+ L1 ≡[V, 0] L2 ∧ L1.ⓑ{I}V ≡[T, 0] L2.ⓑ{I}V.
+/2 width=2 by llpx_sn_inv_bind_O/ qed-.
+
+(* Advanced forward lemmas **************************************************)
+
+lemma lleq_fwd_lref_dx: ∀L1,L2,d,i. L1 ≡[#i, d] L2 →
+ ∀I,K2,V. ⇩[i] L2 ≡ K2.ⓑ{I}V →
+ i < d ∨
+ ∃∃K1. ⇩[i] L1 ≡ K1.ⓑ{I}V & K1 ≡[V, 0] K2 & d ≤ i.
+#L1 #L2 #d #i #H #I #K2 #V #HLK2 elim (llpx_sn_fwd_lref_dx … H … HLK2) -L2
+[ | * ] /3 width=3 by ex3_intro, or_intror, or_introl/
+qed-.
+
+lemma lleq_fwd_lref_sn: ∀L1,L2,d,i. L1 ≡[#i, d] L2 →
+ ∀I,K1,V. ⇩[i] L1 ≡ K1.ⓑ{I}V →
+ i < d ∨
+ ∃∃K2. ⇩[i] L2 ≡ K2.ⓑ{I}V & K1 ≡[V, 0] K2 & d ≤ i.
+#L1 #L2 #d #i #H #I #K1 #V #HLK1 elim (llpx_sn_fwd_lref_sn … H … HLK1) -L1
+[ | * ] /3 width=3 by ex3_intro, or_intror, or_introl/
+qed-.
+
+lemma lleq_fwd_bind_O_dx: ∀a,I,L1,L2,V,T. L1 ≡[ⓑ{a,I}V.T, 0] L2 →
+ L1.ⓑ{I}V ≡[T, 0] L2.ⓑ{I}V.
+/2 width=2 by llpx_sn_fwd_bind_O_dx/ qed-.
+
+(* Properties on relocation *************************************************)
+
+lemma lleq_lift_le: ∀K1,K2,T,dt. K1 ≡[T, dt] K2 →
+ ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 →
+ ∀U. ⇧[d, e] T ≡ U → dt ≤ d → L1 ≡[U, dt] L2.
+/3 width=10 by llpx_sn_lift_le, lift_mono/ qed-.
+
+lemma lleq_lift_ge: ∀K1,K2,T,dt. K1 ≡[T, dt] K2 →
+ ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 →
+ ∀U. ⇧[d, e] T ≡ U → d ≤ dt → L1 ≡[U, dt+e] L2.
+/2 width=9 by llpx_sn_lift_ge/ qed-.
+
+(* Inversion lemmas on relocation *******************************************)
+
+lemma lleq_inv_lift_le: ∀L1,L2,U,dt. L1 ≡[U, dt] L2 →
+ ∀K1,K2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 →
+ ∀T. ⇧[d, e] T ≡ U → dt ≤ d → K1 ≡[T, dt] K2.
+/3 width=10 by llpx_sn_inv_lift_le, ex2_intro/ qed-.
+
+lemma lleq_inv_lift_be: ∀L1,L2,U,dt. L1 ≡[U, dt] L2 →
+ ∀K1,K2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 →
+ ∀T. ⇧[d, e] T ≡ U → d ≤ dt → dt ≤ yinj d + e → K1 ≡[T, d] K2.
+/2 width=11 by llpx_sn_inv_lift_be/ qed-.
+
+lemma lleq_inv_lift_ge: ∀L1,L2,U,dt. L1 ≡[U, dt] L2 →
+ ∀K1,K2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 →
+ ∀T. ⇧[d, e] T ≡ U → yinj d + e ≤ dt → K1 ≡[T, dt-e] K2.
+/2 width=9 by llpx_sn_inv_lift_ge/ qed-.
+
+(* Inversion lemmas on negated lazy quivalence for local environments *******)
+
+lemma nlleq_inv_bind: ∀a,I,L1,L2,V,T,d. (L1 ≡[ⓑ{a,I}V.T, d] L2 → ⊥) →
+ (L1 ≡[V, d] L2 → ⊥) ∨ (L1.ⓑ{I}V ≡[T, ⫯d] L2.ⓑ{I}V → ⊥).
+/3 width=2 by nllpx_sn_inv_bind, eq_term_dec/ qed-.
+
+lemma nlleq_inv_flat: ∀I,L1,L2,V,T,d. (L1 ≡[ⓕ{I}V.T, d] L2 → ⊥) →
+ (L1 ≡[V, d] L2 → ⊥) ∨ (L1 ≡[T, d] L2 → ⊥).
+/3 width=2 by nllpx_sn_inv_flat, eq_term_dec/ qed-.
+
+lemma nlleq_inv_bind_O: ∀a,I,L1,L2,V,T. (L1 ≡[ⓑ{a,I}V.T, 0] L2 → ⊥) →
+ (L1 ≡[V, 0] L2 → ⊥) ∨ (L1.ⓑ{I}V ≡[T, 0] L2.ⓑ{I}V → ⊥).
+/3 width=2 by nllpx_sn_inv_bind_O, eq_term_dec/ qed-.
(**************************************************************************)
include "basic_2/multiple/fqus_alt.ma".
-include "basic_2/multiple/lleq_ldrop.ma".
+include "basic_2/multiple/lleq_drop.ma".
(* LAZY EQUIVALENCE FOR LOCAL ENVIRONMENTS **********************************)
∃∃K1. ⦃G1, L1, T⦄ ⊐ ⦃G2, K1, U⦄ & K1 ≡[U, 0] K2.
#G1 #G2 #L2 #K2 #T #U #H elim H -G1 -G2 -L2 -K2 -T -U
[ #I #G #L2 #V #L1 #H elim (lleq_inv_lref_ge_dx … H … I L2 V) -H //
- #K1 #H1 #H2 lapply (ldrop_inv_O2 … H1) -H1
+ #K1 #H1 #H2 lapply (drop_inv_O2 … H1) -H1
#H destruct /2 width=3 by fqu_lref_O, ex2_intro/
| * [ #a ] #I #G #L2 #V #T #L1 #H
[ elim (lleq_inv_bind … H)
| #I #G #L2 #V #T #L1 #H elim (lleq_inv_flat … H) -H
/2 width=3 by fqu_flat_dx, ex2_intro/
| #G #L2 #K2 #T #U #e #HLK2 #HTU #L1 #HL12
- elim (ldrop_O1_le (Ⓕ) (e+1) L1)
+ elim (drop_O1_le (Ⓕ) (e+1) L1)
[ /3 width=12 by fqu_drop, lleq_inv_lift_le, ex2_intro/
- | lapply (ldrop_fwd_length_le2 … HLK2) -K2
+ | lapply (drop_fwd_length_le2 … HLK2) -K2
lapply (lleq_fwd_length … HL12) -T -U //
]
]
+++ /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/multiple/llpx_sn_ldrop.ma".
-include "basic_2/multiple/lleq.ma".
-
-(* LAZY EQUIVALENCE FOR LOCAL ENVIRONMENTS **********************************)
-
-(* Advanced properties ******************************************************)
-
-lemma lleq_bind_repl_O: ∀I,L1,L2,V,T. L1.ⓑ{I}V ≡[T, 0] L2.ⓑ{I}V →
- ∀J,W. L1 ≡[W, 0] L2 → L1.ⓑ{J}W ≡[T, 0] L2.ⓑ{J}W.
-/2 width=7 by llpx_sn_bind_repl_O/ qed-.
-
-lemma lleq_dec: ∀T,L1,L2,d. Decidable (L1 ≡[T, d] L2).
-/3 width=1 by llpx_sn_dec, eq_term_dec/ qed-.
-
-lemma lleq_llpx_sn_trans: ∀R. lleq_transitive R →
- ∀L1,L2,T,d. L1 ≡[T, d] L2 →
- ∀L. llpx_sn R d T L2 L → llpx_sn R d T L1 L.
-#R #HR #L1 #L2 #T #d #H @(lleq_ind … H) -L1 -L2 -T -d
-[1,2,5: /4 width=6 by llpx_sn_fwd_length, llpx_sn_gref, llpx_sn_skip, llpx_sn_sort, trans_eq/
-|4: /4 width=6 by llpx_sn_fwd_length, llpx_sn_free, le_repl_sn_conf_aux, trans_eq/
-| #I #L1 #L2 #K1 #K2 #V #d #i #Hdi #HLK1 #HLK2 #HK12 #IHK12 #L #H elim (llpx_sn_inv_lref_ge_sn … H … HLK2) -H -HLK2
- /3 width=11 by llpx_sn_lref/
-| #a #I #L1 #L2 #V #T #d #_ #_ #IHV #IHT #L #H elim (llpx_sn_inv_bind … H) -H
- /3 width=1 by llpx_sn_bind/
-| #I #L1 #L2 #V #T #d #_ #_ #IHV #IHT #L #H elim (llpx_sn_inv_flat … H) -H
- /3 width=1 by llpx_sn_flat/
-]
-qed-.
-
-lemma lleq_llpx_sn_conf: ∀R. lleq_transitive R →
- ∀L1,L2,T,d. L1 ≡[T, d] L2 →
- ∀L. llpx_sn R d T L1 L → llpx_sn R d T L2 L.
-/3 width=3 by lleq_llpx_sn_trans, lleq_sym/ qed-.
-
-(* Advanced inversion lemmas ************************************************)
-
-lemma lleq_inv_lref_ge_dx: ∀L1,L2,d,i. L1 ≡[#i, d] L2 → d ≤ i →
- ∀I,K2,V. ⇩[i] L2 ≡ K2.ⓑ{I}V →
- ∃∃K1. ⇩[i] L1 ≡ K1.ⓑ{I}V & K1 ≡[V, 0] K2.
-#L1 #L2 #d #i #H #Hdi #I #K2 #V #HLK2 elim (llpx_sn_inv_lref_ge_dx … H … HLK2) -L2
-/2 width=3 by ex2_intro/
-qed-.
-
-lemma lleq_inv_lref_ge_sn: ∀L1,L2,d,i. L1 ≡[#i, d] L2 → d ≤ i →
- ∀I,K1,V. ⇩[i] L1 ≡ K1.ⓑ{I}V →
- ∃∃K2. ⇩[i] L2 ≡ K2.ⓑ{I}V & K1 ≡[V, 0] K2.
-#L1 #L2 #d #i #H #Hdi #I1 #K1 #V #HLK1 elim (llpx_sn_inv_lref_ge_sn … H … HLK1) -L1
-/2 width=3 by ex2_intro/
-qed-.
-
-lemma lleq_inv_lref_ge_bi: ∀L1,L2,d,i. L1 ≡[#i, d] L2 → d ≤ i →
- ∀I1,I2,K1,K2,V1,V2.
- ⇩[i] L1 ≡ K1.ⓑ{I1}V1 → ⇩[i] L2 ≡ K2.ⓑ{I2}V2 →
- ∧∧ I1 = I2 & K1 ≡[V1, 0] K2 & V1 = V2.
-/2 width=8 by llpx_sn_inv_lref_ge_bi/ qed-.
-
-lemma lleq_inv_lref_ge: ∀L1,L2,d,i. L1 ≡[#i, d] L2 → d ≤ i →
- ∀I,K1,K2,V. ⇩[i] L1 ≡ K1.ⓑ{I}V → ⇩[i] L2 ≡ K2.ⓑ{I}V →
- K1 ≡[V, 0] K2.
-#L1 #L2 #d #i #HL12 #Hdi #I #K1 #K2 #V #HLK1 #HLK2
-elim (lleq_inv_lref_ge_bi … HL12 … HLK1 HLK2) //
-qed-.
-
-lemma lleq_inv_S: ∀L1,L2,T,d. L1 ≡[T, d+1] L2 →
- ∀I,K1,K2,V. ⇩[d] L1 ≡ K1.ⓑ{I}V → ⇩[d] L2 ≡ K2.ⓑ{I}V →
- K1 ≡[V, 0] K2 → L1 ≡[T, d] L2.
-/2 width=9 by llpx_sn_inv_S/ qed-.
-
-lemma lleq_inv_bind_O: ∀a,I,L1,L2,V,T. L1 ≡[ⓑ{a,I}V.T, 0] L2 →
- L1 ≡[V, 0] L2 ∧ L1.ⓑ{I}V ≡[T, 0] L2.ⓑ{I}V.
-/2 width=2 by llpx_sn_inv_bind_O/ qed-.
-
-(* Advanced forward lemmas **************************************************)
-
-lemma lleq_fwd_lref_dx: ∀L1,L2,d,i. L1 ≡[#i, d] L2 →
- ∀I,K2,V. ⇩[i] L2 ≡ K2.ⓑ{I}V →
- i < d ∨
- ∃∃K1. ⇩[i] L1 ≡ K1.ⓑ{I}V & K1 ≡[V, 0] K2 & d ≤ i.
-#L1 #L2 #d #i #H #I #K2 #V #HLK2 elim (llpx_sn_fwd_lref_dx … H … HLK2) -L2
-[ | * ] /3 width=3 by ex3_intro, or_intror, or_introl/
-qed-.
-
-lemma lleq_fwd_lref_sn: ∀L1,L2,d,i. L1 ≡[#i, d] L2 →
- ∀I,K1,V. ⇩[i] L1 ≡ K1.ⓑ{I}V →
- i < d ∨
- ∃∃K2. ⇩[i] L2 ≡ K2.ⓑ{I}V & K1 ≡[V, 0] K2 & d ≤ i.
-#L1 #L2 #d #i #H #I #K1 #V #HLK1 elim (llpx_sn_fwd_lref_sn … H … HLK1) -L1
-[ | * ] /3 width=3 by ex3_intro, or_intror, or_introl/
-qed-.
-
-lemma lleq_fwd_bind_O_dx: ∀a,I,L1,L2,V,T. L1 ≡[ⓑ{a,I}V.T, 0] L2 →
- L1.ⓑ{I}V ≡[T, 0] L2.ⓑ{I}V.
-/2 width=2 by llpx_sn_fwd_bind_O_dx/ qed-.
-
-(* Properties on relocation *************************************************)
-
-lemma lleq_lift_le: ∀K1,K2,T,dt. K1 ≡[T, dt] K2 →
- ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 →
- ∀U. ⇧[d, e] T ≡ U → dt ≤ d → L1 ≡[U, dt] L2.
-/3 width=10 by llpx_sn_lift_le, lift_mono/ qed-.
-
-lemma lleq_lift_ge: ∀K1,K2,T,dt. K1 ≡[T, dt] K2 →
- ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 →
- ∀U. ⇧[d, e] T ≡ U → d ≤ dt → L1 ≡[U, dt+e] L2.
-/2 width=9 by llpx_sn_lift_ge/ qed-.
-
-(* Inversion lemmas on relocation *******************************************)
-
-lemma lleq_inv_lift_le: ∀L1,L2,U,dt. L1 ≡[U, dt] L2 →
- ∀K1,K2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 →
- ∀T. ⇧[d, e] T ≡ U → dt ≤ d → K1 ≡[T, dt] K2.
-/3 width=10 by llpx_sn_inv_lift_le, ex2_intro/ qed-.
-
-lemma lleq_inv_lift_be: ∀L1,L2,U,dt. L1 ≡[U, dt] L2 →
- ∀K1,K2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 →
- ∀T. ⇧[d, e] T ≡ U → d ≤ dt → dt ≤ yinj d + e → K1 ≡[T, d] K2.
-/2 width=11 by llpx_sn_inv_lift_be/ qed-.
-
-lemma lleq_inv_lift_ge: ∀L1,L2,U,dt. L1 ≡[U, dt] L2 →
- ∀K1,K2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 →
- ∀T. ⇧[d, e] T ≡ U → yinj d + e ≤ dt → K1 ≡[T, dt-e] K2.
-/2 width=9 by llpx_sn_inv_lift_ge/ qed-.
-
-(* Inversion lemmas on negated lazy quivalence for local environments *******)
-
-lemma nlleq_inv_bind: ∀a,I,L1,L2,V,T,d. (L1 ≡[ⓑ{a,I}V.T, d] L2 → ⊥) →
- (L1 ≡[V, d] L2 → ⊥) ∨ (L1.ⓑ{I}V ≡[T, ⫯d] L2.ⓑ{I}V → ⊥).
-/3 width=2 by nllpx_sn_inv_bind, eq_term_dec/ qed-.
-
-lemma nlleq_inv_flat: ∀I,L1,L2,V,T,d. (L1 ≡[ⓕ{I}V.T, d] L2 → ⊥) →
- (L1 ≡[V, d] L2 → ⊥) ∨ (L1 ≡[T, d] L2 → ⊥).
-/3 width=2 by nllpx_sn_inv_flat, eq_term_dec/ qed-.
-
-lemma nlleq_inv_bind_O: ∀a,I,L1,L2,V,T. (L1 ≡[ⓑ{a,I}V.T, 0] L2 → ⊥) →
- (L1 ≡[V, 0] L2 → ⊥) ∨ (L1.ⓑ{I}V ≡[T, 0] L2.ⓑ{I}V → ⊥).
-/3 width=2 by nllpx_sn_inv_bind_O, eq_term_dec/ qed-.
(* Properties on equivalence for local environments *************************)
lemma leq_lleq_trans: ∀L2,L,T,d. L2 ≡[T, d] L →
- â\88\80L1. L1 â\89\83[d, ∞] L2 → L1 ≡[T, d] L.
+ â\88\80L1. L1 ⩬[d, ∞] L2 → L1 ≡[T, d] L.
/2 width=3 by leq_llpx_sn_trans/ qed-.
lemma lleq_leq_trans: ∀L,L1,T,d. L ≡[T, d] L1 →
- â\88\80L2. L1 â\89\83[d, ∞] L2 → L ≡[T, d] L2.
+ â\88\80L2. L1 ⩬[d, ∞] L2 → L ≡[T, d] L2.
/2 width=3 by llpx_sn_leq_trans/ qed-.
-lemma lleq_leq_repl: â\88\80L1,L2,T,d. L1 â\89¡[T, d] L2 â\86\92 â\88\80K1. K1 â\89\83[d, ∞] L1 →
- â\88\80K2. L2 â\89\83[d, ∞] K2 → K1 ≡[T, d] K2.
+lemma lleq_leq_repl: â\88\80L1,L2,T,d. L1 â\89¡[T, d] L2 â\86\92 â\88\80K1. K1 ⩬[d, ∞] L1 →
+ â\88\80K2. L2 ⩬[d, ∞] K2 → K1 ≡[T, d] K2.
/2 width=5 by llpx_sn_leq_repl/ qed-.
lemma lleq_bind_repl_SO: ∀I1,I2,L1,L2,V1,V2,T. L1.ⓑ{I1}V1 ≡[T, 0] L2.ⓑ{I2}V2 →
(* *)
(**************************************************************************)
-include "basic_2/multiple/lleq_ldrop.ma".
+include "basic_2/multiple/lleq_drop.ma".
(* Main properties **********************************************************)
elim (llpx_sn_llpx_sn_alt … H1) -H1 #HL12 #IH1
elim H2 -H2 #_ #HL1 #IH2
@lleq_intro_alt // #I2 #I #K2 #K #V2 #V #i #Hi #HnT #HLK2 #HLK
-lapply (ldrop_fwd_length_lt2 … HLK) #HiL
-elim (ldrop_O1_lt (Ⓕ) L1 i) // -HiL #I1 #K1 #V1 #HLK1
+lapply (drop_fwd_length_lt2 … HLK) #HiL
+elim (drop_O1_lt (Ⓕ) L1 i) // -HiL #I1 #K1 #V1 #HLK1
elim (IH1 … HLK1 HLK2) -IH1 /2 width=1 by/ #H #_ destruct
elim (IH2 … HLK1 HLK2 HLK) -IH2 -HLK1 -HLK2 -HLK * /2 width=1 by conj/ #H
[ elim (ylt_yle_false … H) -H //
lemma llor_atom: ∀T,d. ⋆ ⩖[T, d] ⋆ ≡ ⋆.
#T #d @and3_intro //
#I1 #I2 #I #K1 #K2 #K #V1 #V2 #V #i #HLK1
-elim (ldrop_inv_atom1 … HLK1) -HLK1 #H destruct
+elim (drop_inv_atom1 … HLK1) -HLK1 #H destruct
qed.
#L1 #L2 #L #U #d * #HL12 #HL1 #IH #Hd #I1 #W1 #HU #I2 #W2
@and3_intro [1,2: >ltail_length /2 width=1 by le_S_S/ ]
#J1 #J2 #J #K1 #K2 #K #V1 #V2 #V #i #HLK1 #HLK2 #HLK
-lapply (ldrop_fwd_length_lt2 … HLK1) >ltail_length #H
+lapply (drop_fwd_length_lt2 … HLK1) >ltail_length #H
lapply (lt_plus_SO_to_le … H) -H #H
elim (le_to_or_lt_eq … H) -H #H
-[ elim (ldrop_O1_lt (Ⓕ) … H) #Z1 #Y1 #X1 #HLY1
- elim (ldrop_O1_lt (Ⓕ) L2 i) // #Z2 #Y2 #X2 #HLY2
- elim (ldrop_O1_lt (Ⓕ) L i) // #Z #Y #X #HLY
- lapply (ldrop_O1_inv_append1_le … HLK1 … HLY1) /2 width=1 by lt_to_le/ -HLK1 normalize #H destruct
- lapply (ldrop_O1_inv_append1_le … HLK2 … HLY2) /2 width=1 by lt_to_le/ -HLK2 normalize #H destruct
- lapply (ldrop_O1_inv_append1_le … HLK … HLY) /2 width=1 by lt_to_le/ -HLK normalize #H destruct
+[ elim (drop_O1_lt (Ⓕ) … H) #Z1 #Y1 #X1 #HLY1
+ elim (drop_O1_lt (Ⓕ) L2 i) // #Z2 #Y2 #X2 #HLY2
+ elim (drop_O1_lt (Ⓕ) L i) // #Z #Y #X #HLY
+ lapply (drop_O1_inv_append1_le … HLK1 … HLY1) /2 width=1 by lt_to_le/ -HLK1 normalize #H destruct
+ lapply (drop_O1_inv_append1_le … HLK2 … HLY2) /2 width=1 by lt_to_le/ -HLK2 normalize #H destruct
+ lapply (drop_O1_inv_append1_le … HLK … HLY) /2 width=1 by lt_to_le/ -HLK normalize #H destruct
elim (IH … HLY1 HLY2 HLY) -IH -HLY1 -HLY2 -HLY *
[ /3 width=1 by and3_intro, or3_intro0/
| /6 width=2 by frees_inv_append, lt_to_le, or3_intro1, and3_intro/
| /5 width=1 by frees_append, lt_to_le, or3_intro2, and4_intro/
]
| -IH -HLK1 destruct
- lapply (ldrop_O1_inv_append1_le … HLK2 … (⋆) ?) // -HLK2 normalize #H destruct
- lapply (ldrop_O1_inv_append1_le … HLK … (⋆) ?) // -HLK normalize #H destruct
+ lapply (drop_O1_inv_append1_le … HLK2 … (⋆) ?) // -HLK2 normalize #H destruct
+ lapply (drop_O1_inv_append1_le … HLK … (⋆) ?) // -HLK normalize #H destruct
/3 width=1 by or3_intro2, and4_intro/
]
qed.
#L1 #L2 #L #U #d * #HL12 #HL1 #IH #I1 #W1 #HU #I2 #W2
@and3_intro [1,2: >ltail_length /2 width=1 by le_S_S/ ]
#J1 #J2 #J #K1 #K2 #K #V1 #V2 #V #i #HLK1 #HLK2 #HLK
-lapply (ldrop_fwd_length_lt2 … HLK1) >ltail_length #H
+lapply (drop_fwd_length_lt2 … HLK1) >ltail_length #H
lapply (lt_plus_SO_to_le … H) -H #H
elim (le_to_or_lt_eq … H) -H #H
-[ elim (ldrop_O1_lt (Ⓕ) … H) #Z1 #Y1 #X1 #HLY1
- elim (ldrop_O1_lt (Ⓕ) L2 i) // #Z2 #Y2 #X2 #HLY2
- elim (ldrop_O1_lt (Ⓕ) L i) // #Z #Y #X #HLY
- lapply (ldrop_O1_inv_append1_le … HLK1 … HLY1) /2 width=1 by lt_to_le/ -HLK1 normalize #H destruct
- lapply (ldrop_O1_inv_append1_le … HLK2 … HLY2) /2 width=1 by lt_to_le/ -HLK2 normalize #H destruct
- lapply (ldrop_O1_inv_append1_le … HLK … HLY) /2 width=1 by lt_to_le/ -HLK normalize #H destruct
+[ elim (drop_O1_lt (Ⓕ) … H) #Z1 #Y1 #X1 #HLY1
+ elim (drop_O1_lt (Ⓕ) L2 i) // #Z2 #Y2 #X2 #HLY2
+ elim (drop_O1_lt (Ⓕ) L i) // #Z #Y #X #HLY
+ lapply (drop_O1_inv_append1_le … HLK1 … HLY1) /2 width=1 by lt_to_le/ -HLK1 normalize #H destruct
+ lapply (drop_O1_inv_append1_le … HLK2 … HLY2) /2 width=1 by lt_to_le/ -HLK2 normalize #H destruct
+ lapply (drop_O1_inv_append1_le … HLK … HLY) /2 width=1 by lt_to_le/ -HLK normalize #H destruct
elim (IH … HLY1 HLY2 HLY) -IH -HLY1 -HLY2 -HLY *
[ /3 width=1 by and3_intro, or3_intro0/
| /6 width=2 by frees_inv_append, lt_to_le, or3_intro1, and3_intro/
| /5 width=1 by frees_append, lt_to_le, or3_intro2, and4_intro/
]
| -IH -HLK2 destruct
- lapply (ldrop_O1_inv_append1_le … HLK1 … (⋆) ?) // -HLK1 normalize #H destruct
- lapply (ldrop_O1_inv_append1_le … HLK … (⋆) ?) // -HLK normalize #H destruct
+ lapply (drop_O1_inv_append1_le … HLK1 … (⋆) ?) // -HLK1 normalize #H destruct
+ lapply (drop_O1_inv_append1_le … HLK … (⋆) ?) // -HLK normalize #H destruct
/4 width=1 by or3_intro1, and3_intro/
]
qed.
--- /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/multiple/frees_lift.ma".
+include "basic_2/multiple/llor_alt.ma".
+
+(* POINTWISE UNION FOR LOCAL ENVIRONMENTS ***********************************)
+
+(* Advanced properties ******************************************************)
+
+lemma llor_skip: ∀L1,L2,U,d. |L1| = |L2| → yinj (|L1|) ≤ d → L1 ⩖[U, d] L2 ≡ L1.
+#L1 #L2 #U #d #HL12 #Hd @and3_intro // -HL12
+#I1 #I2 #I #K1 #K2 #K #W1 #W2 #W #i #HLK1 #_ #HLK -L2 -K2
+lapply (drop_mono … HLK … HLK1) -HLK #H destruct
+lapply (drop_fwd_length_lt2 … HLK1) -K1
+/5 width=3 by ylt_yle_trans, ylt_inj, or3_intro0, and3_intro/
+qed.
+
+(* Note: lemma 1400 concludes the "big tree" theorem *)
+lemma llor_total: ∀L1,L2,T,d. |L1| = |L2| → ∃L. L1 ⩖[T, d] L2 ≡ L.
+#L1 @(lenv_ind_alt … L1) -L1
+[ #L2 #T #d #H >(length_inv_zero_sn … H) -L2 /2 width=2 by ex_intro/
+| #I1 #L1 #V1 #IHL1 #Y #T #d >ltail_length #H
+ elim (length_inv_pos_sn_ltail … H) -H #I2 #L2 #V2 #HL12 #H destruct
+ elim (ylt_split d (|ⓑ{I1}V1.L1|))
+ [ elim (frees_dec (ⓑ{I1}V1.L1) T d (|L1|)) #HnU
+ elim (IHL1 L2 T d) // -IHL1 -HL12
+ [ #L #HL12 >ltail_length /4 width=2 by llor_tail_frees, ylt_fwd_succ2, ex_intro/
+ | /4 width=2 by llor_tail_cofrees, ex_intro/
+ ]
+ | -IHL1 /4 width=3 by llor_skip, plus_to_minus, ex_intro/
+ ]
+]
+qed-.
+++ /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/multiple/frees_lift.ma".
-include "basic_2/multiple/llor_alt.ma".
-
-(* POINTWISE UNION FOR LOCAL ENVIRONMENTS ***********************************)
-
-(* Advanced properties ******************************************************)
-
-lemma llor_skip: ∀L1,L2,U,d. |L1| = |L2| → yinj (|L1|) ≤ d → L1 ⩖[U, d] L2 ≡ L1.
-#L1 #L2 #U #d #HL12 #Hd @and3_intro // -HL12
-#I1 #I2 #I #K1 #K2 #K #W1 #W2 #W #i #HLK1 #_ #HLK -L2 -K2
-lapply (ldrop_mono … HLK … HLK1) -HLK #H destruct
-lapply (ldrop_fwd_length_lt2 … HLK1) -K1
-/5 width=3 by ylt_yle_trans, ylt_inj, or3_intro0, and3_intro/
-qed.
-
-(* Note: lemma 1400 concludes the "big tree" theorem *)
-lemma llor_total: ∀L1,L2,T,d. |L1| = |L2| → ∃L. L1 ⩖[T, d] L2 ≡ L.
-#L1 @(lenv_ind_alt … L1) -L1
-[ #L2 #T #d #H >(length_inv_zero_sn … H) -L2 /2 width=2 by ex_intro/
-| #I1 #L1 #V1 #IHL1 #Y #T #d >ltail_length #H
- elim (length_inv_pos_sn_ltail … H) -H #I2 #L2 #V2 #HL12 #H destruct
- elim (ylt_split d (|ⓑ{I1}V1.L1|))
- [ elim (frees_dec (ⓑ{I1}V1.L1) T d (|L1|)) #HnU
- elim (IHL1 L2 T d) // -IHL1 -HL12
- [ #L #HL12 >ltail_length /4 width=2 by llor_tail_frees, ylt_fwd_succ2, ex_intro/
- | /4 width=2 by llor_tail_cofrees, ex_intro/
- ]
- | -IHL1 /4 width=3 by llor_skip, plus_to_minus, ex_intro/
- ]
-]
-qed-.
(**************************************************************************)
include "ground_2/ynat/ynat_plus.ma".
-include "basic_2/substitution/ldrop.ma".
+include "basic_2/substitution/drop.ma".
(* LAZY SN POINTWISE EXTENSION OF A CONTEXT-SENSITIVE REALTION FOR TERMS ****)
lemma llpx_sn_fwd_length: ∀R,L1,L2,T,d. llpx_sn R d T L1 L2 → |L1| = |L2|.
#R #L1 #L2 #T #d #H elim H -L1 -L2 -T -d //
#I #L1 #L2 #K1 #K2 #V1 #V2 #d #i #_ #HLK1 #HLK2 #_ #_ #HK12
-lapply (ldrop_fwd_length … HLK1) -HLK1
-lapply (ldrop_fwd_length … HLK2) -HLK2
+lapply (drop_fwd_length … HLK1) -HLK1
+lapply (drop_fwd_length … HLK2) -HLK2
normalize //
qed-.
-lemma llpx_sn_fwd_ldrop_sn: ∀R,L1,L2,T,d. llpx_sn R d T L1 L2 →
+lemma llpx_sn_fwd_drop_sn: ∀R,L1,L2,T,d. llpx_sn R d T L1 L2 →
∀K1,i. ⇩[i] L1 ≡ K1 → ∃K2. ⇩[i] L2 ≡ K2.
#R #L1 #L2 #T #d #H #K1 #i #HLK1 lapply (llpx_sn_fwd_length … H) -H
-#HL12 lapply (ldrop_fwd_length_le2 … HLK1) -HLK1 /2 width=1 by ldrop_O1_le/
+#HL12 lapply (drop_fwd_length_le2 … HLK1) -HLK1 /2 width=1 by drop_O1_le/
qed-.
-lemma llpx_sn_fwd_ldrop_dx: ∀R,L1,L2,T,d. llpx_sn R d T L1 L2 →
+lemma llpx_sn_fwd_drop_dx: ∀R,L1,L2,T,d. llpx_sn R d T L1 L2 →
∀K2,i. ⇩[i] L2 ≡ K2 → ∃K1. ⇩[i] L1 ≡ K1.
#R #L1 #L2 #T #d #H #K2 #i #HLK2 lapply (llpx_sn_fwd_length … H) -H
-#HL12 lapply (ldrop_fwd_length_le2 … HLK2) -HLK2 /2 width=1 by ldrop_O1_le/
+#HL12 lapply (drop_fwd_length_le2 … HLK2) -HLK2 /2 width=1 by drop_O1_le/
qed-.
fact llpx_sn_fwd_lref_aux: ∀R,L1,L2,X,d. llpx_sn R d X L1 L2 → ∀i. X = #i →
#n #IH #L * * /3 width=1 by llpx_sn_sort, llpx_sn_gref, llpx_sn_bind, llpx_sn_flat/
#i #Hn elim (lt_or_ge i (|L|)) /2 width=1 by llpx_sn_free/
#HiL #d elim (ylt_split i d) /2 width=1 by llpx_sn_skip/
-elim (ldrop_O1_lt … HiL) -HiL destruct /4 width=9 by llpx_sn_lref, ldrop_fwd_rfw/
+elim (drop_O1_lt … HiL) -HiL destruct /4 width=9 by llpx_sn_lref, drop_fwd_rfw/
qed-.
lemma llpx_sn_Y: ∀R,T,L1,L2. |L1| = |L2| → llpx_sn R (∞) T L1 L2.
| #I #L1 #L2 #K1 #K2 #W1 #W2 #dt #i #Hdti #HLK1 #HLK2 #HW1 #HW12 #_ #X #d #e #H #_
elim (lift_inv_lref2 … H) -H * #Hid #H destruct
[ lapply (llpx_sn_fwd_length … HW1) -HW1 #HK12
- lapply (ldrop_fwd_length … HLK1) lapply (ldrop_fwd_length … HLK2)
+ lapply (drop_fwd_length … HLK1) lapply (drop_fwd_length … HLK2)
normalize in ⊢ (%→%→?); -I -W1 -W2 -dt /3 width=1 by llpx_sn_skip, ylt_inj/
| /4 width=9 by llpx_sn_lref, yle_inj, le_plus_b/
]
#I1 #I2 #K1 #K2 #V1 #V2 #i #Hdi #H #HLK1 #HLK2 elim (frees_inv … H) -H
[ -n #HnU elim (IHU … HnU HLK1 HLK2) -IHU -HnU -HLK1 -HLK2 /2 width=1 by conj/
| * #J1 #K10 #W10 #j #Hdj #Hji #HnU #HLK10 #HnW10 destruct
- lapply (ldrop_fwd_drop2 … HLK10) #H
- lapply (ldrop_conf_ge … H … HLK1 ?) -H /2 width=1 by lt_to_le/ <minus_plus #HK10
- elim (ldrop_O1_lt (Ⓕ) L2 j) [2: <HL12 /2 width=5 by ldrop_fwd_length_lt2/ ] #J2 #K20 #W20 #HLK20
- lapply (ldrop_fwd_drop2 … HLK20) #H
- lapply (ldrop_conf_ge … H … HLK2 ?) -H /2 width=1 by lt_to_le/ <minus_plus #HK20
- elim (IHn K10 W10 … K20 0) -IHn -HL12 /3 width=6 by ldrop_fwd_rfw/
+ lapply (drop_fwd_drop2 … HLK10) #H
+ lapply (drop_conf_ge … H … HLK1 ?) -H /2 width=1 by lt_to_le/ <minus_plus #HK10
+ elim (drop_O1_lt (Ⓕ) L2 j) [2: <HL12 /2 width=5 by drop_fwd_length_lt2/ ] #J2 #K20 #W20 #HLK20
+ lapply (drop_fwd_drop2 … HLK20) #H
+ lapply (drop_conf_ge … H … HLK2 ?) -H /2 width=1 by lt_to_le/ <minus_plus #HK20
+ elim (IHn K10 W10 … K20 0) -IHn -HL12 /3 width=6 by drop_fwd_rfw/
elim (IHU … HnU HLK10 HLK20) -IHU -HnU -HLK10 -HLK20 //
]
qed.
#n #IHn #L1 #U #Hn #L2 #d * #HL12 #IHU @llpx_sn_intro_alt_r //
#I1 #I2 #K1 #K2 #V1 #V2 #i #Hdi #HnU #HLK1 #HLK2 destruct
elim (IHU … HLK1 HLK2) /3 width=2 by frees_eq/
-#H #HV12 @and3_intro // @IHn -IHn /3 width=6 by ldrop_fwd_rfw/
-lapply (ldrop_fwd_drop2 … HLK1) #H1
-lapply (ldrop_fwd_drop2 … HLK2) -HLK2 #H2
-@conj [ @(ldrop_fwd_length_eq1 … H1 H2) // ] -HL12
+#H #HV12 @and3_intro // @IHn -IHn /3 width=6 by drop_fwd_rfw/
+lapply (drop_fwd_drop2 … HLK1) #H1
+lapply (drop_fwd_drop2 … HLK2) -HLK2 #H2
+@conj [ @(drop_fwd_length_eq1 … H1 H2) // ] -HL12
#Z1 #Z2 #Y1 #Y2 #X1 #X2 #j #_
>(minus_plus_m_m j (i+1)) in ⊢ (%→?); >commutative_plus <minus_plus
#HnV1 #HKY1 #HKY2 (**) (* full auto too slow *)
-lapply (ldrop_trans_ge … H1 … HKY1 ?) -H1 -HKY1 // #HLY1
-lapply (ldrop_trans_ge … H2 … HKY2 ?) -H2 -HKY2 // #HLY2
+lapply (drop_trans_ge … H1 … HKY1 ?) -H1 -HKY1 // #HLY1
+lapply (drop_trans_ge … H2 … HKY2 ?) -H2 -HKY2 // #HLY2
/4 width=14 by frees_be, yle_plus_dx2_trans, yle_succ_dx/
qed-.
(**************************************************************************)
include "basic_2/substitution/lift_neg.ma".
-include "basic_2/substitution/ldrop_ldrop.ma".
+include "basic_2/substitution/drop_drop.ma".
include "basic_2/multiple/llpx_sn.ma".
(* LAZY SN POINTWISE EXTENSION OF A CONTEXT-SENSITIVE REALTION FOR TERMS ****)
[ elim (IH … HLK1 HLK2) -IH -HLK1 -HLK2
/3 width=9 by nlift_bind_sn, and3_intro/
| lapply (yle_inv_succ1 … Hdi) -Hdi * #Hdi #Hi
- lapply (ldrop_inv_drop1_lt … HLK1 ?) -HLK1 /2 width=1 by ylt_O/ #HLK1
- lapply (ldrop_inv_drop1_lt … HLK2 ?) -HLK2 /2 width=1 by ylt_O/ #HLK2
+ lapply (drop_inv_drop1_lt … HLK1 ?) -HLK1 /2 width=1 by ylt_O/ #HLK1
+ lapply (drop_inv_drop1_lt … HLK2 ?) -HLK2 /2 width=1 by ylt_O/ #HLK2
elim (IH … HLK1 HLK2) -IH -HLK1 -HLK2 /2 width=1 by and3_intro/
@nlift_bind_dx <plus_minus_m_m /2 width=2 by ylt_O/
]
#R #L1 #L2 #d #i #H elim (llpx_sn_alt_r_inv_alt … H) -H
#HL12 #IH elim (lt_or_ge i (|L1|)) /3 width=1 by or3_intro0, conj/
elim (ylt_split i d) /3 width=1 by or3_intro1/
-#Hdi #HL1 elim (ldrop_O1_lt (Ⓕ) … HL1)
-#I1 #K1 #V1 #HLK1 elim (ldrop_O1_lt (Ⓕ) L2 i) //
+#Hdi #HL1 elim (drop_O1_lt (Ⓕ) … HL1)
+#I1 #K1 #V1 #HLK1 elim (drop_O1_lt (Ⓕ) L2 i) //
#I2 #K2 #V2 #HLK2 elim (IH … HLK1 HLK2) -IH
/3 width=9 by nlift_lref_be_SO, or3_intro2, ex5_5_intro/
qed-.
llpx_sn_alt_r R d (#i) L1 L2.
#R #L1 #L2 #d #i #HL1 #_ #HL12 @llpx_sn_alt_r_intro_alt // -HL12
#I1 #I2 #K1 #K2 #V1 #V2 #j #_ #H #HLK1 elim (H (#(i-1))) -H
-lapply (ldrop_fwd_length_lt2 … HLK1) -HLK1
+lapply (drop_fwd_length_lt2 … HLK1) -HLK1
/3 width=3 by lift_lref_ge_minus, lt_to_le_to_lt/
qed.
llpx_sn_alt_r R d (#i) L1 L2.
#R #I #L1 #L2 #K1 #K2 #V1 #V2 #d #i #Hdi #HLK1 #HLK2 #HK12 #HV12 @llpx_sn_alt_r_intro_alt
[ lapply (llpx_sn_alt_r_fwd_length … HK12) -HK12 #HK12
- @(ldrop_fwd_length_eq2 … HLK1 HLK2) normalize //
+ @(drop_fwd_length_eq2 … HLK1 HLK2) normalize //
| #Z1 #Z2 #Y1 #Y2 #X1 #X2 #j #Hdj #H #HLY1 #HLY2
elim (lt_or_eq_or_gt i j) #Hij destruct
[ elim (H (#i)) -H /2 width=1 by lift_lref_lt/
- | lapply (ldrop_mono … HLY1 … HLK1) -HLY1 -HLK1 #H destruct
- lapply (ldrop_mono … HLY2 … HLK2) -HLY2 -HLK2 #H destruct /2 width=1 by and3_intro/
+ | lapply (drop_mono … HLY1 … HLK1) -HLY1 -HLK1 #H destruct
+ lapply (drop_mono … HLY2 … HLK2) -HLY2 -HLK2 #H destruct /2 width=1 by and3_intro/
| elim (H (#(i-1))) -H /2 width=1 by lift_lref_ge_minus/
]
]
#I1 #I2 #K1 #K2 #V1 #V2 #i #Hdi #HnVT #HLK1 #HLK2
elim (nlift_inv_bind … HnVT) -HnVT #H
[ elim (IHV … HLK1 … HLK2) -IHV /2 width=2 by and3_intro/
-| elim IHT -IHT /2 width=12 by ldrop_drop, yle_succ, and3_intro/
+| elim IHT -IHT /2 width=12 by drop_drop, yle_succ, and3_intro/
]
qed.
#HL12 elim (llpx_sn_alt_r_fwd_lref … H) -H
[ * /2 width=1 by llpx_sn_free/
| /2 width=1 by llpx_sn_skip/
- | * /4 width=9 by llpx_sn_lref, ldrop_fwd_rfw/
+ | * /4 width=9 by llpx_sn_lref, drop_fwd_rfw/
]
| #a #I #V #T #Hn #L2 #d #H elim (llpx_sn_alt_r_inv_bind … H) -H
/3 width=1 by llpx_sn_bind/
--- /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/substitution/drop_drop.ma".
+include "basic_2/multiple/llpx_sn_leq.ma".
+
+(* LAZY SN POINTWISE EXTENSION OF A CONTEXT-SENSITIVE REALTION FOR TERMS ****)
+
+(* Advanced forward lemmas **************************************************)
+
+lemma llpx_sn_fwd_lref_dx: ∀R,L1,L2,d,i. llpx_sn R d (#i) L1 L2 →
+ ∀I,K2,V2. ⇩[i] L2 ≡ K2.ⓑ{I}V2 →
+ i < d ∨
+ ∃∃K1,V1. ⇩[i] L1 ≡ K1.ⓑ{I}V1 & llpx_sn R 0 V1 K1 K2 &
+ R K1 V1 V2 & d ≤ i.
+#R #L1 #L2 #d #i #H #I #K2 #V2 #HLK2 elim (llpx_sn_fwd_lref … H) -H [ * || * ]
+[ #_ #H elim (lt_refl_false i)
+ lapply (drop_fwd_length_lt2 … HLK2) -HLK2
+ /2 width=3 by lt_to_le_to_lt/ (**) (* full auto too slow *)
+| /2 width=1 by or_introl/
+| #I #K11 #K22 #V11 #V22 #HLK11 #HLK22 #HK12 #HV12 #Hdi
+ lapply (drop_mono … HLK22 … HLK2) -L2 #H destruct
+ /3 width=5 by ex4_2_intro, or_intror/
+]
+qed-.
+
+lemma llpx_sn_fwd_lref_sn: ∀R,L1,L2,d,i. llpx_sn R d (#i) L1 L2 →
+ ∀I,K1,V1. ⇩[i] L1 ≡ K1.ⓑ{I}V1 →
+ i < d ∨
+ ∃∃K2,V2. ⇩[i] L2 ≡ K2.ⓑ{I}V2 & llpx_sn R 0 V1 K1 K2 &
+ R K1 V1 V2 & d ≤ i.
+#R #L1 #L2 #d #i #H #I #K1 #V1 #HLK1 elim (llpx_sn_fwd_lref … H) -H [ * || * ]
+[ #H #_ elim (lt_refl_false i)
+ lapply (drop_fwd_length_lt2 … HLK1) -HLK1
+ /2 width=3 by lt_to_le_to_lt/ (**) (* full auto too slow *)
+| /2 width=1 by or_introl/
+| #I #K11 #K22 #V11 #V22 #HLK11 #HLK22 #HK12 #HV12 #Hdi
+ lapply (drop_mono … HLK11 … HLK1) -L1 #H destruct
+ /3 width=5 by ex4_2_intro, or_intror/
+]
+qed-.
+
+(* Advanced inversion lemmas ************************************************)
+
+lemma llpx_sn_inv_lref_ge_dx: ∀R,L1,L2,d,i. llpx_sn R d (#i) L1 L2 → d ≤ i →
+ ∀I,K2,V2. ⇩[i] L2 ≡ K2.ⓑ{I}V2 →
+ ∃∃K1,V1. ⇩[i] L1 ≡ K1.ⓑ{I}V1 &
+ llpx_sn R 0 V1 K1 K2 & R K1 V1 V2.
+#R #L1 #L2 #d #i #H #Hdi #I #K2 #V2 #HLK2 elim (llpx_sn_fwd_lref_dx … H … HLK2) -L2
+[ #H elim (ylt_yle_false … H Hdi)
+| * /2 width=5 by ex3_2_intro/
+]
+qed-.
+
+lemma llpx_sn_inv_lref_ge_sn: ∀R,L1,L2,d,i. llpx_sn R d (#i) L1 L2 → d ≤ i →
+ ∀I,K1,V1. ⇩[i] L1 ≡ K1.ⓑ{I}V1 →
+ ∃∃K2,V2. ⇩[i] L2 ≡ K2.ⓑ{I}V2 &
+ llpx_sn R 0 V1 K1 K2 & R K1 V1 V2.
+#R #L1 #L2 #d #i #H #Hdi #I #K1 #V1 #HLK1 elim (llpx_sn_fwd_lref_sn … H … HLK1) -L1
+[ #H elim (ylt_yle_false … H Hdi)
+| * /2 width=5 by ex3_2_intro/
+]
+qed-.
+
+lemma llpx_sn_inv_lref_ge_bi: ∀R,L1,L2,d,i. llpx_sn R d (#i) L1 L2 → d ≤ i →
+ ∀I1,I2,K1,K2,V1,V2.
+ ⇩[i] L1 ≡ K1.ⓑ{I1}V1 → ⇩[i] L2 ≡ K2.ⓑ{I2}V2 →
+ ∧∧ I1 = I2 & llpx_sn R 0 V1 K1 K2 & R K1 V1 V2.
+#R #L1 #L2 #d #i #HL12 #Hdi #I1 #I2 #K1 #K2 #V1 #V2 #HLK1 #HLK2
+elim (llpx_sn_inv_lref_ge_sn … HL12 … HLK1) // -L1 -d
+#J #Y #HY lapply (drop_mono … HY … HLK2) -L2 -i #H destruct /2 width=1 by and3_intro/
+qed-.
+
+fact llpx_sn_inv_S_aux: ∀R,L1,L2,T,d0. llpx_sn R d0 T L1 L2 → ∀d. d0 = d + 1 →
+ ∀K1,K2,I,V1,V2. ⇩[d] L1 ≡ K1.ⓑ{I}V1 → ⇩[d] L2 ≡ K2.ⓑ{I}V2 →
+ llpx_sn R 0 V1 K1 K2 → R K1 V1 V2 → llpx_sn R d T L1 L2.
+#R #L1 #L2 #T #d0 #H elim H -L1 -L2 -T -d0
+/2 width=1 by llpx_sn_gref, llpx_sn_free, llpx_sn_sort/
+[ #L1 #L2 #d0 #i #HL12 #Hid #d #H #K1 #K2 #I #V1 #V2 #HLK1 #HLK2 #HK12 #HV12 destruct
+ elim (yle_split_eq i d) /2 width=1 by llpx_sn_skip, ylt_fwd_succ2/ -HL12 -Hid
+ #H destruct /2 width=9 by llpx_sn_lref/
+| #I #L1 #L2 #K11 #K22 #V1 #V2 #d0 #i #Hd0i #HLK11 #HLK22 #HK12 #HV12 #_ #d #H #K1 #K2 #J #W1 #W2 #_ #_ #_ #_ destruct
+ /3 width=9 by llpx_sn_lref, yle_pred_sn/
+| #a #I #L1 #L2 #V #T #d0 #_ #_ #IHV #IHT #d #H #K1 #K2 #J #W1 #W2 #HLK1 #HLK2 #HK12 #HW12 destruct
+ /4 width=9 by llpx_sn_bind, drop_drop/
+| #I #L1 #L2 #V #T #d0 #_ #_ #IHV #IHT #d #H #K1 #K2 #J #W1 #W2 #HLK1 #HLK2 #HK12 #HW12 destruct
+ /3 width=9 by llpx_sn_flat/
+]
+qed-.
+
+lemma llpx_sn_inv_S: ∀R,L1,L2,T,d. llpx_sn R (d + 1) T L1 L2 →
+ ∀K1,K2,I,V1,V2. ⇩[d] L1 ≡ K1.ⓑ{I}V1 → ⇩[d] L2 ≡ K2.ⓑ{I}V2 →
+ llpx_sn R 0 V1 K1 K2 → R K1 V1 V2 → llpx_sn R d T L1 L2.
+/2 width=9 by llpx_sn_inv_S_aux/ qed-.
+
+lemma llpx_sn_inv_bind_O: ∀R. (∀L. reflexive … (R L)) →
+ ∀a,I,L1,L2,V,T. llpx_sn R 0 (ⓑ{a,I}V.T) L1 L2 →
+ llpx_sn R 0 V L1 L2 ∧ llpx_sn R 0 T (L1.ⓑ{I}V) (L2.ⓑ{I}V).
+#R #HR #a #I #L1 #L2 #V #T #H elim (llpx_sn_inv_bind … H) -H
+/3 width=9 by drop_pair, conj, llpx_sn_inv_S/
+qed-.
+
+(* More advanced forward lemmas *********************************************)
+
+lemma llpx_sn_fwd_bind_O_dx: ∀R. (∀L. reflexive … (R L)) →
+ ∀a,I,L1,L2,V,T. llpx_sn R 0 (ⓑ{a,I}V.T) L1 L2 →
+ llpx_sn R 0 T (L1.ⓑ{I}V) (L2.ⓑ{I}V).
+#R #HR #a #I #L1 #L2 #V #T #H elim (llpx_sn_inv_bind_O … H) -H //
+qed-.
+
+(* Advanced properties ******************************************************)
+
+lemma llpx_sn_bind_repl_O: ∀R,I,L1,L2,V1,V2,T. llpx_sn R 0 T (L1.ⓑ{I}V1) (L2.ⓑ{I}V2) →
+ ∀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).
+/3 width=9 by llpx_sn_bind_repl_SO, llpx_sn_inv_S/ qed-.
+
+lemma llpx_sn_dec: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) →
+ ∀T,L1,L2,d. Decidable (llpx_sn R d T L1 L2).
+#R #HR #T #L1 @(f2_ind … rfw … L1 T) -L1 -T
+#n #IH #L1 * *
+[ #k #Hn #L2 elim (eq_nat_dec (|L1|) (|L2|)) /3 width=1 by or_introl, llpx_sn_sort/
+| #i #Hn #L2 elim (eq_nat_dec (|L1|) (|L2|))
+ [ #HL12 #d elim (ylt_split i d) /3 width=1 by llpx_sn_skip, or_introl/
+ #Hdi elim (lt_or_ge i (|L1|)) #HiL1
+ elim (lt_or_ge i (|L2|)) #HiL2 /3 width=1 by or_introl, llpx_sn_free/
+ elim (drop_O1_lt (Ⓕ) … HiL2) #I2 #K2 #V2 #HLK2
+ elim (drop_O1_lt (Ⓕ) … HiL1) #I1 #K1 #V1 #HLK1
+ elim (eq_bind2_dec I2 I1)
+ [ #H2 elim (HR K1 V1 V2) -HR
+ [ #H3 elim (IH K1 V1 … K2 0) destruct
+ /3 width=9 by llpx_sn_lref, drop_fwd_rfw, or_introl/
+ ]
+ ]
+ -IH #H3 @or_intror
+ #H elim (llpx_sn_fwd_lref … H) -H [1,3,4,6,7,9: * ]
+ [1,3,5: /3 width=4 by lt_to_le_to_lt, lt_refl_false/
+ |7,8,9: /2 width=4 by ylt_yle_false/
+ ]
+ #Z #Y1 #Y2 #X1 #X2 #HLY1 #HLY2 #HY12 #HX12
+ lapply (drop_mono … HLY1 … HLK1) -HLY1 -HLK1
+ lapply (drop_mono … HLY2 … HLK2) -HLY2 -HLK2
+ #H #H0 destruct /2 width=1 by/
+ ]
+| #p #Hn #L2 elim (eq_nat_dec (|L1|) (|L2|)) /3 width=1 by or_introl, llpx_sn_gref/
+| #a #I #V #T #Hn #L2 #d destruct
+ elim (IH L1 V … L2 d) /2 width=1 by/
+ elim (IH (L1.ⓑ{I}V) T … (L2.ⓑ{I}V) (⫯d)) -IH /3 width=1 by or_introl, llpx_sn_bind/
+ #H1 #H2 @or_intror
+ #H elim (llpx_sn_inv_bind … H) -H /2 width=1 by/
+| #I #V #T #Hn #L2 #d destruct
+ elim (IH L1 V … L2 d) /2 width=1 by/
+ elim (IH L1 T … L2 d) -IH /3 width=1 by or_introl, llpx_sn_flat/
+ #H1 #H2 @or_intror
+ #H elim (llpx_sn_inv_flat … H) -H /2 width=1 by/
+]
+-n /4 width=4 by llpx_sn_fwd_length, or_intror/
+qed-.
+
+(* Properties on relocation *************************************************)
+
+lemma llpx_sn_lift_le: ∀R. l_liftable R →
+ ∀K1,K2,T,d0. llpx_sn R d0 T K1 K2 →
+ ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 →
+ ∀U. ⇧[d, e] T ≡ U → d0 ≤ d → llpx_sn R d0 U L1 L2.
+#R #HR #K1 #K2 #T #d0 #H elim H -K1 -K2 -T -d0
+[ #K1 #K2 #d0 #k #HK12 #L1 #L2 #d #e #HLK1 #HLK2 #X #H #_ >(lift_inv_sort1 … H) -X
+ lapply (drop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 -d
+ /2 width=1 by llpx_sn_sort/
+| #K1 #K2 #d0 #i #HK12 #Hid0 #L1 #L2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_lref1 … H) -H
+ * #Hdi #H destruct
+ [ lapply (drop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 -d
+ /2 width=1 by llpx_sn_skip/
+ | elim (ylt_yle_false … Hid0) -L1 -L2 -K1 -K2 -e -Hid0
+ /3 width=3 by yle_trans, yle_inj/
+ ]
+| #I #K1 #K2 #K11 #K22 #V1 #V2 #d0 #i #Hid0 #HK11 #HK22 #HK12 #HV12 #IHK12 #L1 #L2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_lref1 … H) -H
+ * #Hdi #H destruct [ -HK12 | -IHK12 ]
+ [ elim (drop_trans_lt … HLK1 … HK11) // -K1
+ elim (drop_trans_lt … HLK2 … HK22) // -Hdi -K2
+ /3 width=18 by llpx_sn_lref/
+ | lapply (drop_trans_ge_comm … HLK1 … HK11 ?) // -K1
+ lapply (drop_trans_ge_comm … HLK2 … HK22 ?) // -Hdi -Hd0 -K2
+ /3 width=9 by llpx_sn_lref, yle_plus_dx1_trans/
+ ]
+| #K1 #K2 #d0 #i #HK1 #HK2 #HK12 #L1 #L2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_lref1 … H) -H
+ * #Hid #H destruct
+ lapply (drop_fwd_length_eq2 … HLK1 HLK2 HK12) -HK12
+ [ /3 width=7 by llpx_sn_free, drop_fwd_be/
+ | lapply (drop_fwd_length … HLK1) -HLK1 #HLK1
+ lapply (drop_fwd_length … HLK2) -HLK2 #HLK2
+ @llpx_sn_free [ >HLK1 | >HLK2 ] -Hid -HLK1 -HLK2 /2 width=1 by monotonic_le_plus_r/ (**) (* explicit constructor *)
+ ]
+| #K1 #K2 #d0 #p #HK12 #L1 #L2 #d #e #HLK1 #HLK2 #X #H #_ >(lift_inv_gref1 … H) -X
+ lapply (drop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 -d -e
+ /2 width=1 by llpx_sn_gref/
+| #a #I #K1 #K2 #V #T #d0 #_ #_ #IHV #IHT #L1 #L2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_bind1 … H) -H
+ #W #U #HVW #HTU #H destruct /4 width=6 by llpx_sn_bind, drop_skip, yle_succ/
+| #I #K1 #K2 #V #T #d0 #_ #_ #IHV #IHT #L1 #L2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_flat1 … H) -H
+ #W #U #HVW #HTU #H destruct /3 width=6 by llpx_sn_flat/
+]
+qed-.
+
+lemma llpx_sn_lift_ge: ∀R,K1,K2,T,d0. llpx_sn R d0 T K1 K2 →
+ ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 →
+ ∀U. ⇧[d, e] T ≡ U → d ≤ d0 → llpx_sn R (d0+e) U L1 L2.
+#R #K1 #K2 #T #d0 #H elim H -K1 -K2 -T -d0
+[ #K1 #K2 #d0 #k #HK12 #L1 #L2 #d #e #HLK1 #HLK2 #X #H #_ >(lift_inv_sort1 … H) -X
+ lapply (drop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 -d
+ /2 width=1 by llpx_sn_sort/
+| #K1 #K2 #d0 #i #HK12 #Hid0 #L1 #L2 #d #e #HLK1 #HLK2 #X #H #_ elim (lift_inv_lref1 … H) -H
+ * #_ #H destruct
+ lapply (drop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2
+ [ /3 width=3 by llpx_sn_skip, ylt_plus_dx2_trans/
+ | /3 width=3 by llpx_sn_skip, monotonic_ylt_plus_dx/
+ ]
+| #I #K1 #K2 #K11 #K22 #V1 #V2 #d0 #i #Hid0 #HK11 #HK22 #HK12 #HV12 #_ #L1 #L2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_lref1 … H) -H
+ * #Hid #H destruct
+ [ elim (ylt_yle_false … Hid0) -I -L1 -L2 -K1 -K2 -K11 -K22 -V1 -V2 -e -Hid0
+ /3 width=3 by ylt_yle_trans, ylt_inj/
+ | lapply (drop_trans_ge_comm … HLK1 … HK11 ?) // -K1
+ lapply (drop_trans_ge_comm … HLK2 … HK22 ?) // -Hid -Hd0 -K2
+ /3 width=9 by llpx_sn_lref, monotonic_yle_plus_dx/
+ ]
+| #K1 #K2 #d0 #i #HK1 #HK2 #HK12 #L1 #L2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_lref1 … H) -H
+ * #Hid #H destruct
+ lapply (drop_fwd_length_eq2 … HLK1 HLK2 HK12) -HK12
+ [ /3 width=7 by llpx_sn_free, drop_fwd_be/
+ | lapply (drop_fwd_length … HLK1) -HLK1 #HLK1
+ lapply (drop_fwd_length … HLK2) -HLK2 #HLK2
+ @llpx_sn_free [ >HLK1 | >HLK2 ] -Hid -HLK1 -HLK2 /2 width=1 by monotonic_le_plus_r/ (**) (* explicit constructor *)
+ ]
+| #K1 #K2 #d0 #p #HK12 #L1 #L2 #d #e #HLK1 #HLK2 #X #H #_ >(lift_inv_gref1 … H) -X
+ lapply (drop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 -d
+ /2 width=1 by llpx_sn_gref/
+| #a #I #K1 #K2 #V #T #d0 #_ #_ #IHV #IHT #L1 #L2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_bind1 … H) -H
+ #W #U #HVW #HTU #H destruct /4 width=5 by llpx_sn_bind, drop_skip, yle_succ/
+| #I #K1 #K2 #V #T #d0 #_ #_ #IHV #IHT #L1 #L2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_flat1 … H) -H
+ #W #U #HVW #HTU #H destruct /3 width=5 by llpx_sn_flat/
+]
+qed-.
+
+(* Inversion lemmas on relocation *******************************************)
+
+lemma llpx_sn_inv_lift_le: ∀R. l_deliftable_sn R →
+ ∀L1,L2,U,d0. llpx_sn R d0 U L1 L2 →
+ ∀K1,K2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 →
+ ∀T. ⇧[d, e] T ≡ U → d0 ≤ d → llpx_sn R d0 T K1 K2.
+#R #HR #L1 #L2 #U #d0 #H elim H -L1 -L2 -U -d0
+[ #L1 #L2 #d0 #k #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #_ >(lift_inv_sort2 … H) -X
+ lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -d -e
+ /2 width=1 by llpx_sn_sort/
+| #L1 #L2 #d0 #i #HL12 #Hid0 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #_ elim (lift_inv_lref2 … H) -H
+ * #_ #H destruct
+ lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2
+ [ /2 width=1 by llpx_sn_skip/
+ | /3 width=3 by llpx_sn_skip, yle_ylt_trans/
+ ]
+| #I #L1 #L2 #K11 #K22 #W1 #W2 #d0 #i #Hid0 #HLK11 #HLK22 #HK12 #HW12 #IHK12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_lref2 … H) -H
+ * #Hid #H destruct [ -HK12 | -IHK12 ]
+ [ elim (drop_conf_lt … HLK1 … HLK11) // -L1 #L1 #V1 #HKL1 #HKL11 #HVW1
+ elim (drop_conf_lt … HLK2 … HLK22) // -Hid -L2 #L2 #V2 #HKL2 #HKL22 #HVW2
+ elim (HR … HW12 … HKL11 … HVW1) -HR #V0 #HV0 #HV12
+ lapply (lift_inj … HV0 … HVW2) -HV0 -HVW2 #H destruct
+ /3 width=10 by llpx_sn_lref/
+ | lapply (drop_conf_ge … HLK1 … HLK11 ?) // -L1
+ lapply (drop_conf_ge … HLK2 … HLK22 ?) // -L2 -Hid0
+ elim (le_inv_plus_l … Hid) -Hid /4 width=9 by llpx_sn_lref, yle_trans, yle_inj/ (**) (* slow *)
+ ]
+| #L1 #L2 #d0 #i #HL1 #HL2 #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_lref2 … H) -H
+ * #_ #H destruct
+ lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12)
+ [ lapply (drop_fwd_length_le4 … HLK1) -HLK1
+ lapply (drop_fwd_length_le4 … HLK2) -HLK2
+ #HKL2 #HKL1 #HK12 @llpx_sn_free // /2 width=3 by transitive_le/ (**) (* full auto too slow *)
+ | lapply (drop_fwd_length … HLK1) -HLK1 #H >H in HL1; -H
+ lapply (drop_fwd_length … HLK2) -HLK2 #H >H in HL2; -H
+ /3 width=1 by llpx_sn_free, le_plus_to_minus_r/
+ ]
+| #L1 #L2 #d0 #p #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #_ >(lift_inv_gref2 … H) -X
+ lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -d -e
+ /2 width=1 by llpx_sn_gref/
+| #a #I #L1 #L2 #W #U #d0 #_ #_ #IHW #IHU #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_bind2 … H) -H
+ #V #T #HVW #HTU #H destruct /4 width=6 by llpx_sn_bind, drop_skip, yle_succ/
+| #I #L1 #L2 #W #U #d0 #_ #_ #IHW #IHU #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_flat2 … H) -H
+ #V #T #HVW #HTU #H destruct /3 width=6 by llpx_sn_flat/
+]
+qed-.
+
+lemma llpx_sn_inv_lift_be: ∀R,L1,L2,U,d0. llpx_sn R d0 U L1 L2 →
+ ∀K1,K2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 →
+ ∀T. ⇧[d, e] T ≡ U → d ≤ d0 → d0 ≤ yinj d + e → llpx_sn R d T K1 K2.
+#R #L1 #L2 #U #d0 #H elim H -L1 -L2 -U -d0
+[ #L1 #L2 #d0 #k #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #_ #_ >(lift_inv_sort2 … H) -X
+ lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -d0 -e
+ /2 width=1 by llpx_sn_sort/
+| #L1 #L2 #d0 #i #HL12 #Hid0 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hd0 #Hd0e elim (lift_inv_lref2 … H) -H
+ * #Hid #H destruct
+ [ lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2
+ -Hid0 /3 width=1 by llpx_sn_skip, ylt_inj/
+ | elim (ylt_yle_false … Hid0) -L1 -L2 -Hd0 -Hid0
+ /3 width=3 by yle_trans, yle_inj/ (**) (* slow *)
+ ]
+| #I #L1 #L2 #K11 #K22 #W1 #W2 #d0 #i #Hid0 #HLK11 #HLK22 #HK12 #HW12 #_ #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hd0 #Hd0e elim (lift_inv_lref2 … H) -H
+ * #Hid #H destruct
+ [ elim (ylt_yle_false … Hid0) -I -L1 -L2 -K11 -K22 -W1 -W2 -Hd0e -Hid0
+ /3 width=3 by ylt_yle_trans, ylt_inj/
+ | lapply (drop_conf_ge … HLK1 … HLK11 ?) // -L1
+ lapply (drop_conf_ge … HLK2 … HLK22 ?) // -L2 -Hid0 -Hd0 -Hd0e
+ elim (le_inv_plus_l … Hid) -Hid /3 width=9 by llpx_sn_lref, yle_inj/
+ ]
+| #L1 #L2 #d0 #i #HL1 #HL2 #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hd0 #Hd0e elim (lift_inv_lref2 … H) -H
+ * #_ #H destruct
+ lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12)
+ [ lapply (drop_fwd_length_le4 … HLK1) -HLK1
+ lapply (drop_fwd_length_le4 … HLK2) -HLK2
+ #HKL2 #HKL1 #HK12 @llpx_sn_free // /2 width=3 by transitive_le/ (**) (* full auto too slow *)
+ | lapply (drop_fwd_length … HLK1) -HLK1 #H >H in HL1; -H
+ lapply (drop_fwd_length … HLK2) -HLK2 #H >H in HL2; -H
+ /3 width=1 by llpx_sn_free, le_plus_to_minus_r/
+ ]
+| #L1 #L2 #d0 #p #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #_ #_ >(lift_inv_gref2 … H) -X
+ lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -d0 -e
+ /2 width=1 by llpx_sn_gref/
+| #a #I #L1 #L2 #W #U #d0 #_ #_ #IHW #IHU #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hd0 #Hd0e elim (lift_inv_bind2 … H) -H
+ >commutative_plus #V #T #HVW #HTU #H destruct
+ @llpx_sn_bind [ /2 width=5 by/ ] -IHW (**) (* explicit constructor *)
+ @(IHU … HTU) -IHU -HTU /2 width=1 by drop_skip, yle_succ/
+| #I #L1 #L2 #W #U #d0 #_ #_ #IHW #IHU #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hd0 #Hd0e elim (lift_inv_flat2 … H) -H
+ #V #T #HVW #HTU #H destruct /3 width=6 by llpx_sn_flat/
+]
+qed-.
+
+lemma llpx_sn_inv_lift_ge: ∀R,L1,L2,U,d0. llpx_sn R d0 U L1 L2 →
+ ∀K1,K2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 →
+ ∀T. ⇧[d, e] T ≡ U → yinj d + e ≤ d0 → llpx_sn R (d0-e) T K1 K2.
+#R #L1 #L2 #U #d0 #H elim H -L1 -L2 -U -d0
+[ #L1 #L2 #d0 #k #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #_ >(lift_inv_sort2 … H) -X
+ lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -d
+ /2 width=1 by llpx_sn_sort/
+| #L1 #L2 #d0 #i #HL12 #Hid0 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hded0 elim (lift_inv_lref2 … H) -H
+ * #Hid #H destruct [ -Hid0 | -Hded0 ]
+ lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2
+ [ /4 width=3 by llpx_sn_skip, yle_plus1_to_minus_inj2, ylt_yle_trans, ylt_inj/
+ | elim (le_inv_plus_l … Hid) -Hid #_
+ /4 width=1 by llpx_sn_skip, monotonic_ylt_minus_dx, yle_inj/
+ ]
+| #I #L1 #L2 #K11 #K22 #W1 #W2 #d0 #i #Hid0 #HLK11 #HLK22 #HK12 #HW12 #_ #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hded0 elim (lift_inv_lref2 … H) -H
+ * #Hid #H destruct
+ [ elim (ylt_yle_false … Hid0) -I -L1 -L2 -K11 -K22 -W1 -W2 -Hid0
+ /3 width=3 by yle_fwd_plus_sn1, ylt_yle_trans, ylt_inj/
+ | lapply (drop_conf_ge … HLK1 … HLK11 ?) // -L1
+ lapply (drop_conf_ge … HLK2 … HLK22 ?) // -L2 -Hded0 -Hid
+ /3 width=9 by llpx_sn_lref, monotonic_yle_minus_dx/
+ ]
+| #L1 #L2 #d0 #i #HL1 #HL2 #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hded0 elim (lift_inv_lref2 … H) -H
+ * #_ #H destruct
+ lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12)
+ [ lapply (drop_fwd_length_le4 … HLK1) -HLK1
+ lapply (drop_fwd_length_le4 … HLK2) -HLK2
+ #HKL2 #HKL1 #HK12 @llpx_sn_free // /2 width=3 by transitive_le/ (**) (* full auto too slow *)
+ | lapply (drop_fwd_length … HLK1) -HLK1 #H >H in HL1; -H
+ lapply (drop_fwd_length … HLK2) -HLK2 #H >H in HL2; -H
+ /3 width=1 by llpx_sn_free, le_plus_to_minus_r/
+ ]
+| #L1 #L2 #d0 #p #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #_ >(lift_inv_gref2 … H) -X
+ lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -d
+ /2 width=1 by llpx_sn_gref/
+| #a #I #L1 #L2 #W #U #d0 #_ #_ #IHW #IHU #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hded0 elim (lift_inv_bind2 … H) -H
+ #V #T #HVW #HTU #H destruct
+ @llpx_sn_bind [ /2 width=5 by/ ] -IHW (**) (* explicit constructor *)
+ <yminus_succ1_inj /2 width=2 by yle_fwd_plus_sn2/
+ @(IHU … HTU) -IHU -HTU /2 width=1 by drop_skip, yle_succ/
+| #I #L1 #L2 #W #U #d0 #_ #_ #IHW #IHU #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hded0 elim (lift_inv_flat2 … H) -H
+ #V #T #HVW #HTU #H destruct /3 width=5 by llpx_sn_flat/
+]
+qed-.
+
+(* Advanced inversion lemmas on relocation **********************************)
+
+lemma llpx_sn_inv_lift_O: ∀R,L1,L2,U. llpx_sn R 0 U L1 L2 →
+ ∀K1,K2,e. ⇩[e] L1 ≡ K1 → ⇩[e] L2 ≡ K2 →
+ ∀T. ⇧[0, e] T ≡ U → llpx_sn R 0 T K1 K2.
+/2 width=11 by llpx_sn_inv_lift_be/ qed-.
+
+lemma llpx_sn_drop_conf_O: ∀R,L1,L2,U. llpx_sn R 0 U L1 L2 →
+ ∀K1,e. ⇩[e] L1 ≡ K1 → ∀T. ⇧[0, e] T ≡ U →
+ ∃∃K2. ⇩[e] L2 ≡ K2 & llpx_sn R 0 T K1 K2.
+#R #L1 #L2 #U #HU #K1 #e #HLK1 #T #HTU elim (llpx_sn_fwd_drop_sn … HU … HLK1)
+/3 width=10 by llpx_sn_inv_lift_O, ex2_intro/
+qed-.
+
+lemma llpx_sn_drop_trans_O: ∀R,L1,L2,U. llpx_sn R 0 U L1 L2 →
+ ∀K2,e. ⇩[e] L2 ≡ K2 → ∀T. ⇧[0, e] T ≡ U →
+ ∃∃K1. ⇩[e] L1 ≡ K1 & llpx_sn R 0 T K1 K2.
+#R #L1 #L2 #U #HU #K2 #e #HLK2 #T #HTU elim (llpx_sn_fwd_drop_dx … HU … HLK2)
+/3 width=10 by llpx_sn_inv_lift_O, ex2_intro/
+qed-.
+
+(* Inversion lemmas on negated lazy pointwise extension *********************)
+
+lemma nllpx_sn_inv_bind: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) →
+ ∀a,I,L1,L2,V,T,d. (llpx_sn R d (ⓑ{a,I}V.T) L1 L2 → ⊥) →
+ (llpx_sn R d V L1 L2 → ⊥) ∨ (llpx_sn R (⫯d) T (L1.ⓑ{I}V) (L2.ⓑ{I}V) → ⊥).
+#R #HR #a #I #L1 #L2 #V #T #d #H elim (llpx_sn_dec … HR V L1 L2 d)
+/4 width=1 by llpx_sn_bind, or_intror, or_introl/
+qed-.
+
+lemma nllpx_sn_inv_flat: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) →
+ ∀I,L1,L2,V,T,d. (llpx_sn R d (ⓕ{I}V.T) L1 L2 → ⊥) →
+ (llpx_sn R d V L1 L2 → ⊥) ∨ (llpx_sn R d T L1 L2 → ⊥).
+#R #HR #I #L1 #L2 #V #T #d #H elim (llpx_sn_dec … HR V L1 L2 d)
+/4 width=1 by llpx_sn_flat, or_intror, or_introl/
+qed-.
+
+lemma nllpx_sn_inv_bind_O: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) →
+ ∀a,I,L1,L2,V,T. (llpx_sn R 0 (ⓑ{a,I}V.T) L1 L2 → ⊥) →
+ (llpx_sn R 0 V L1 L2 → ⊥) ∨ (llpx_sn R 0 T (L1.ⓑ{I}V) (L2.ⓑ{I}V) → ⊥).
+#R #HR #a #I #L1 #L2 #V #T #H elim (llpx_sn_dec … HR V L1 L2 0)
+/4 width=1 by llpx_sn_bind_O, or_intror, or_introl/
+qed-.
#R #H1R #H2R #L2 #U #d #i #H elim H -L2 -U -d -i /3 width=2 by frees_eq/
#I2 #L2 #K2 #U #W2 #d #i #j #Hdj #Hji #HnU #HLK2 #_ #IHW2 #L1 #HL12
elim (llpx_sn_inv_alt_r … HL12) -HL12 #HL12 #IH
-lapply (ldrop_fwd_length_lt2 … HLK2) #Hj
-elim (ldrop_O1_lt (Ⓕ) L1 j) // -Hj -HL12 #I1 #K1 #W1 #HLK1
+lapply (drop_fwd_length_lt2 … HLK2) #Hj
+elim (drop_O1_lt (Ⓕ) L1 j) // -Hj -HL12 #I1 #K1 #W1 #HLK1
elim (IH … HnU HLK1 HLK2) // -IH -HLK2 /5 width=11 by frees_be/
qed-.
+++ /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/substitution/ldrop_ldrop.ma".
-include "basic_2/multiple/llpx_sn_leq.ma".
-
-(* LAZY SN POINTWISE EXTENSION OF A CONTEXT-SENSITIVE REALTION FOR TERMS ****)
-
-(* Advanced forward lemmas **************************************************)
-
-lemma llpx_sn_fwd_lref_dx: ∀R,L1,L2,d,i. llpx_sn R d (#i) L1 L2 →
- ∀I,K2,V2. ⇩[i] L2 ≡ K2.ⓑ{I}V2 →
- i < d ∨
- ∃∃K1,V1. ⇩[i] L1 ≡ K1.ⓑ{I}V1 & llpx_sn R 0 V1 K1 K2 &
- R K1 V1 V2 & d ≤ i.
-#R #L1 #L2 #d #i #H #I #K2 #V2 #HLK2 elim (llpx_sn_fwd_lref … H) -H [ * || * ]
-[ #_ #H elim (lt_refl_false i)
- lapply (ldrop_fwd_length_lt2 … HLK2) -HLK2
- /2 width=3 by lt_to_le_to_lt/ (**) (* full auto too slow *)
-| /2 width=1 by or_introl/
-| #I #K11 #K22 #V11 #V22 #HLK11 #HLK22 #HK12 #HV12 #Hdi
- lapply (ldrop_mono … HLK22 … HLK2) -L2 #H destruct
- /3 width=5 by ex4_2_intro, or_intror/
-]
-qed-.
-
-lemma llpx_sn_fwd_lref_sn: ∀R,L1,L2,d,i. llpx_sn R d (#i) L1 L2 →
- ∀I,K1,V1. ⇩[i] L1 ≡ K1.ⓑ{I}V1 →
- i < d ∨
- ∃∃K2,V2. ⇩[i] L2 ≡ K2.ⓑ{I}V2 & llpx_sn R 0 V1 K1 K2 &
- R K1 V1 V2 & d ≤ i.
-#R #L1 #L2 #d #i #H #I #K1 #V1 #HLK1 elim (llpx_sn_fwd_lref … H) -H [ * || * ]
-[ #H #_ elim (lt_refl_false i)
- lapply (ldrop_fwd_length_lt2 … HLK1) -HLK1
- /2 width=3 by lt_to_le_to_lt/ (**) (* full auto too slow *)
-| /2 width=1 by or_introl/
-| #I #K11 #K22 #V11 #V22 #HLK11 #HLK22 #HK12 #HV12 #Hdi
- lapply (ldrop_mono … HLK11 … HLK1) -L1 #H destruct
- /3 width=5 by ex4_2_intro, or_intror/
-]
-qed-.
-
-(* Advanced inversion lemmas ************************************************)
-
-lemma llpx_sn_inv_lref_ge_dx: ∀R,L1,L2,d,i. llpx_sn R d (#i) L1 L2 → d ≤ i →
- ∀I,K2,V2. ⇩[i] L2 ≡ K2.ⓑ{I}V2 →
- ∃∃K1,V1. ⇩[i] L1 ≡ K1.ⓑ{I}V1 &
- llpx_sn R 0 V1 K1 K2 & R K1 V1 V2.
-#R #L1 #L2 #d #i #H #Hdi #I #K2 #V2 #HLK2 elim (llpx_sn_fwd_lref_dx … H … HLK2) -L2
-[ #H elim (ylt_yle_false … H Hdi)
-| * /2 width=5 by ex3_2_intro/
-]
-qed-.
-
-lemma llpx_sn_inv_lref_ge_sn: ∀R,L1,L2,d,i. llpx_sn R d (#i) L1 L2 → d ≤ i →
- ∀I,K1,V1. ⇩[i] L1 ≡ K1.ⓑ{I}V1 →
- ∃∃K2,V2. ⇩[i] L2 ≡ K2.ⓑ{I}V2 &
- llpx_sn R 0 V1 K1 K2 & R K1 V1 V2.
-#R #L1 #L2 #d #i #H #Hdi #I #K1 #V1 #HLK1 elim (llpx_sn_fwd_lref_sn … H … HLK1) -L1
-[ #H elim (ylt_yle_false … H Hdi)
-| * /2 width=5 by ex3_2_intro/
-]
-qed-.
-
-lemma llpx_sn_inv_lref_ge_bi: ∀R,L1,L2,d,i. llpx_sn R d (#i) L1 L2 → d ≤ i →
- ∀I1,I2,K1,K2,V1,V2.
- ⇩[i] L1 ≡ K1.ⓑ{I1}V1 → ⇩[i] L2 ≡ K2.ⓑ{I2}V2 →
- ∧∧ I1 = I2 & llpx_sn R 0 V1 K1 K2 & R K1 V1 V2.
-#R #L1 #L2 #d #i #HL12 #Hdi #I1 #I2 #K1 #K2 #V1 #V2 #HLK1 #HLK2
-elim (llpx_sn_inv_lref_ge_sn … HL12 … HLK1) // -L1 -d
-#J #Y #HY lapply (ldrop_mono … HY … HLK2) -L2 -i #H destruct /2 width=1 by and3_intro/
-qed-.
-
-fact llpx_sn_inv_S_aux: ∀R,L1,L2,T,d0. llpx_sn R d0 T L1 L2 → ∀d. d0 = d + 1 →
- ∀K1,K2,I,V1,V2. ⇩[d] L1 ≡ K1.ⓑ{I}V1 → ⇩[d] L2 ≡ K2.ⓑ{I}V2 →
- llpx_sn R 0 V1 K1 K2 → R K1 V1 V2 → llpx_sn R d T L1 L2.
-#R #L1 #L2 #T #d0 #H elim H -L1 -L2 -T -d0
-/2 width=1 by llpx_sn_gref, llpx_sn_free, llpx_sn_sort/
-[ #L1 #L2 #d0 #i #HL12 #Hid #d #H #K1 #K2 #I #V1 #V2 #HLK1 #HLK2 #HK12 #HV12 destruct
- elim (yle_split_eq i d) /2 width=1 by llpx_sn_skip, ylt_fwd_succ2/ -HL12 -Hid
- #H destruct /2 width=9 by llpx_sn_lref/
-| #I #L1 #L2 #K11 #K22 #V1 #V2 #d0 #i #Hd0i #HLK11 #HLK22 #HK12 #HV12 #_ #d #H #K1 #K2 #J #W1 #W2 #_ #_ #_ #_ destruct
- /3 width=9 by llpx_sn_lref, yle_pred_sn/
-| #a #I #L1 #L2 #V #T #d0 #_ #_ #IHV #IHT #d #H #K1 #K2 #J #W1 #W2 #HLK1 #HLK2 #HK12 #HW12 destruct
- /4 width=9 by llpx_sn_bind, ldrop_drop/
-| #I #L1 #L2 #V #T #d0 #_ #_ #IHV #IHT #d #H #K1 #K2 #J #W1 #W2 #HLK1 #HLK2 #HK12 #HW12 destruct
- /3 width=9 by llpx_sn_flat/
-]
-qed-.
-
-lemma llpx_sn_inv_S: ∀R,L1,L2,T,d. llpx_sn R (d + 1) T L1 L2 →
- ∀K1,K2,I,V1,V2. ⇩[d] L1 ≡ K1.ⓑ{I}V1 → ⇩[d] L2 ≡ K2.ⓑ{I}V2 →
- llpx_sn R 0 V1 K1 K2 → R K1 V1 V2 → llpx_sn R d T L1 L2.
-/2 width=9 by llpx_sn_inv_S_aux/ qed-.
-
-lemma llpx_sn_inv_bind_O: ∀R. (∀L. reflexive … (R L)) →
- ∀a,I,L1,L2,V,T. llpx_sn R 0 (ⓑ{a,I}V.T) L1 L2 →
- llpx_sn R 0 V L1 L2 ∧ llpx_sn R 0 T (L1.ⓑ{I}V) (L2.ⓑ{I}V).
-#R #HR #a #I #L1 #L2 #V #T #H elim (llpx_sn_inv_bind … H) -H
-/3 width=9 by ldrop_pair, conj, llpx_sn_inv_S/
-qed-.
-
-(* More advanced forward lemmas *********************************************)
-
-lemma llpx_sn_fwd_bind_O_dx: ∀R. (∀L. reflexive … (R L)) →
- ∀a,I,L1,L2,V,T. llpx_sn R 0 (ⓑ{a,I}V.T) L1 L2 →
- llpx_sn R 0 T (L1.ⓑ{I}V) (L2.ⓑ{I}V).
-#R #HR #a #I #L1 #L2 #V #T #H elim (llpx_sn_inv_bind_O … H) -H //
-qed-.
-
-(* Advanced properties ******************************************************)
-
-lemma llpx_sn_bind_repl_O: ∀R,I,L1,L2,V1,V2,T. llpx_sn R 0 T (L1.ⓑ{I}V1) (L2.ⓑ{I}V2) →
- ∀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).
-/3 width=9 by llpx_sn_bind_repl_SO, llpx_sn_inv_S/ qed-.
-
-lemma llpx_sn_dec: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) →
- ∀T,L1,L2,d. Decidable (llpx_sn R d T L1 L2).
-#R #HR #T #L1 @(f2_ind … rfw … L1 T) -L1 -T
-#n #IH #L1 * *
-[ #k #Hn #L2 elim (eq_nat_dec (|L1|) (|L2|)) /3 width=1 by or_introl, llpx_sn_sort/
-| #i #Hn #L2 elim (eq_nat_dec (|L1|) (|L2|))
- [ #HL12 #d elim (ylt_split i d) /3 width=1 by llpx_sn_skip, or_introl/
- #Hdi elim (lt_or_ge i (|L1|)) #HiL1
- elim (lt_or_ge i (|L2|)) #HiL2 /3 width=1 by or_introl, llpx_sn_free/
- elim (ldrop_O1_lt (Ⓕ) … HiL2) #I2 #K2 #V2 #HLK2
- elim (ldrop_O1_lt (Ⓕ) … HiL1) #I1 #K1 #V1 #HLK1
- elim (eq_bind2_dec I2 I1)
- [ #H2 elim (HR K1 V1 V2) -HR
- [ #H3 elim (IH K1 V1 … K2 0) destruct
- /3 width=9 by llpx_sn_lref, ldrop_fwd_rfw, or_introl/
- ]
- ]
- -IH #H3 @or_intror
- #H elim (llpx_sn_fwd_lref … H) -H [1,3,4,6,7,9: * ]
- [1,3,5: /3 width=4 by lt_to_le_to_lt, lt_refl_false/
- |7,8,9: /2 width=4 by ylt_yle_false/
- ]
- #Z #Y1 #Y2 #X1 #X2 #HLY1 #HLY2 #HY12 #HX12
- lapply (ldrop_mono … HLY1 … HLK1) -HLY1 -HLK1
- lapply (ldrop_mono … HLY2 … HLK2) -HLY2 -HLK2
- #H #H0 destruct /2 width=1 by/
- ]
-| #p #Hn #L2 elim (eq_nat_dec (|L1|) (|L2|)) /3 width=1 by or_introl, llpx_sn_gref/
-| #a #I #V #T #Hn #L2 #d destruct
- elim (IH L1 V … L2 d) /2 width=1 by/
- elim (IH (L1.ⓑ{I}V) T … (L2.ⓑ{I}V) (⫯d)) -IH /3 width=1 by or_introl, llpx_sn_bind/
- #H1 #H2 @or_intror
- #H elim (llpx_sn_inv_bind … H) -H /2 width=1 by/
-| #I #V #T #Hn #L2 #d destruct
- elim (IH L1 V … L2 d) /2 width=1 by/
- elim (IH L1 T … L2 d) -IH /3 width=1 by or_introl, llpx_sn_flat/
- #H1 #H2 @or_intror
- #H elim (llpx_sn_inv_flat … H) -H /2 width=1 by/
-]
--n /4 width=4 by llpx_sn_fwd_length, or_intror/
-qed-.
-
-(* Properties on relocation *************************************************)
-
-lemma llpx_sn_lift_le: ∀R. l_liftable R →
- ∀K1,K2,T,d0. llpx_sn R d0 T K1 K2 →
- ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 →
- ∀U. ⇧[d, e] T ≡ U → d0 ≤ d → llpx_sn R d0 U L1 L2.
-#R #HR #K1 #K2 #T #d0 #H elim H -K1 -K2 -T -d0
-[ #K1 #K2 #d0 #k #HK12 #L1 #L2 #d #e #HLK1 #HLK2 #X #H #_ >(lift_inv_sort1 … H) -X
- lapply (ldrop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 -d
- /2 width=1 by llpx_sn_sort/
-| #K1 #K2 #d0 #i #HK12 #Hid0 #L1 #L2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_lref1 … H) -H
- * #Hdi #H destruct
- [ lapply (ldrop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 -d
- /2 width=1 by llpx_sn_skip/
- | elim (ylt_yle_false … Hid0) -L1 -L2 -K1 -K2 -e -Hid0
- /3 width=3 by yle_trans, yle_inj/
- ]
-| #I #K1 #K2 #K11 #K22 #V1 #V2 #d0 #i #Hid0 #HK11 #HK22 #HK12 #HV12 #IHK12 #L1 #L2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_lref1 … H) -H
- * #Hdi #H destruct [ -HK12 | -IHK12 ]
- [ elim (ldrop_trans_lt … HLK1 … HK11) // -K1
- elim (ldrop_trans_lt … HLK2 … HK22) // -Hdi -K2
- /3 width=18 by llpx_sn_lref/
- | lapply (ldrop_trans_ge_comm … HLK1 … HK11 ?) // -K1
- lapply (ldrop_trans_ge_comm … HLK2 … HK22 ?) // -Hdi -Hd0 -K2
- /3 width=9 by llpx_sn_lref, yle_plus_dx1_trans/
- ]
-| #K1 #K2 #d0 #i #HK1 #HK2 #HK12 #L1 #L2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_lref1 … H) -H
- * #Hid #H destruct
- lapply (ldrop_fwd_length_eq2 … HLK1 HLK2 HK12) -HK12
- [ /3 width=7 by llpx_sn_free, ldrop_fwd_be/
- | lapply (ldrop_fwd_length … HLK1) -HLK1 #HLK1
- lapply (ldrop_fwd_length … HLK2) -HLK2 #HLK2
- @llpx_sn_free [ >HLK1 | >HLK2 ] -Hid -HLK1 -HLK2 /2 width=1 by monotonic_le_plus_r/ (**) (* explicit constructor *)
- ]
-| #K1 #K2 #d0 #p #HK12 #L1 #L2 #d #e #HLK1 #HLK2 #X #H #_ >(lift_inv_gref1 … H) -X
- lapply (ldrop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 -d -e
- /2 width=1 by llpx_sn_gref/
-| #a #I #K1 #K2 #V #T #d0 #_ #_ #IHV #IHT #L1 #L2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_bind1 … H) -H
- #W #U #HVW #HTU #H destruct /4 width=6 by llpx_sn_bind, ldrop_skip, yle_succ/
-| #I #K1 #K2 #V #T #d0 #_ #_ #IHV #IHT #L1 #L2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_flat1 … H) -H
- #W #U #HVW #HTU #H destruct /3 width=6 by llpx_sn_flat/
-]
-qed-.
-
-lemma llpx_sn_lift_ge: ∀R,K1,K2,T,d0. llpx_sn R d0 T K1 K2 →
- ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 →
- ∀U. ⇧[d, e] T ≡ U → d ≤ d0 → llpx_sn R (d0+e) U L1 L2.
-#R #K1 #K2 #T #d0 #H elim H -K1 -K2 -T -d0
-[ #K1 #K2 #d0 #k #HK12 #L1 #L2 #d #e #HLK1 #HLK2 #X #H #_ >(lift_inv_sort1 … H) -X
- lapply (ldrop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 -d
- /2 width=1 by llpx_sn_sort/
-| #K1 #K2 #d0 #i #HK12 #Hid0 #L1 #L2 #d #e #HLK1 #HLK2 #X #H #_ elim (lift_inv_lref1 … H) -H
- * #_ #H destruct
- lapply (ldrop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2
- [ /3 width=3 by llpx_sn_skip, ylt_plus_dx2_trans/
- | /3 width=3 by llpx_sn_skip, monotonic_ylt_plus_dx/
- ]
-| #I #K1 #K2 #K11 #K22 #V1 #V2 #d0 #i #Hid0 #HK11 #HK22 #HK12 #HV12 #_ #L1 #L2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_lref1 … H) -H
- * #Hid #H destruct
- [ elim (ylt_yle_false … Hid0) -I -L1 -L2 -K1 -K2 -K11 -K22 -V1 -V2 -e -Hid0
- /3 width=3 by ylt_yle_trans, ylt_inj/
- | lapply (ldrop_trans_ge_comm … HLK1 … HK11 ?) // -K1
- lapply (ldrop_trans_ge_comm … HLK2 … HK22 ?) // -Hid -Hd0 -K2
- /3 width=9 by llpx_sn_lref, monotonic_yle_plus_dx/
- ]
-| #K1 #K2 #d0 #i #HK1 #HK2 #HK12 #L1 #L2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_lref1 … H) -H
- * #Hid #H destruct
- lapply (ldrop_fwd_length_eq2 … HLK1 HLK2 HK12) -HK12
- [ /3 width=7 by llpx_sn_free, ldrop_fwd_be/
- | lapply (ldrop_fwd_length … HLK1) -HLK1 #HLK1
- lapply (ldrop_fwd_length … HLK2) -HLK2 #HLK2
- @llpx_sn_free [ >HLK1 | >HLK2 ] -Hid -HLK1 -HLK2 /2 width=1 by monotonic_le_plus_r/ (**) (* explicit constructor *)
- ]
-| #K1 #K2 #d0 #p #HK12 #L1 #L2 #d #e #HLK1 #HLK2 #X #H #_ >(lift_inv_gref1 … H) -X
- lapply (ldrop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 -d
- /2 width=1 by llpx_sn_gref/
-| #a #I #K1 #K2 #V #T #d0 #_ #_ #IHV #IHT #L1 #L2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_bind1 … H) -H
- #W #U #HVW #HTU #H destruct /4 width=5 by llpx_sn_bind, ldrop_skip, yle_succ/
-| #I #K1 #K2 #V #T #d0 #_ #_ #IHV #IHT #L1 #L2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_flat1 … H) -H
- #W #U #HVW #HTU #H destruct /3 width=5 by llpx_sn_flat/
-]
-qed-.
-
-(* Inversion lemmas on relocation *******************************************)
-
-lemma llpx_sn_inv_lift_le: ∀R. l_deliftable_sn R →
- ∀L1,L2,U,d0. llpx_sn R d0 U L1 L2 →
- ∀K1,K2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 →
- ∀T. ⇧[d, e] T ≡ U → d0 ≤ d → llpx_sn R d0 T K1 K2.
-#R #HR #L1 #L2 #U #d0 #H elim H -L1 -L2 -U -d0
-[ #L1 #L2 #d0 #k #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #_ >(lift_inv_sort2 … H) -X
- lapply (ldrop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -d -e
- /2 width=1 by llpx_sn_sort/
-| #L1 #L2 #d0 #i #HL12 #Hid0 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #_ elim (lift_inv_lref2 … H) -H
- * #_ #H destruct
- lapply (ldrop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2
- [ /2 width=1 by llpx_sn_skip/
- | /3 width=3 by llpx_sn_skip, yle_ylt_trans/
- ]
-| #I #L1 #L2 #K11 #K22 #W1 #W2 #d0 #i #Hid0 #HLK11 #HLK22 #HK12 #HW12 #IHK12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_lref2 … H) -H
- * #Hid #H destruct [ -HK12 | -IHK12 ]
- [ elim (ldrop_conf_lt … HLK1 … HLK11) // -L1 #L1 #V1 #HKL1 #HKL11 #HVW1
- elim (ldrop_conf_lt … HLK2 … HLK22) // -Hid -L2 #L2 #V2 #HKL2 #HKL22 #HVW2
- elim (HR … HW12 … HKL11 … HVW1) -HR #V0 #HV0 #HV12
- lapply (lift_inj … HV0 … HVW2) -HV0 -HVW2 #H destruct
- /3 width=10 by llpx_sn_lref/
- | lapply (ldrop_conf_ge … HLK1 … HLK11 ?) // -L1
- lapply (ldrop_conf_ge … HLK2 … HLK22 ?) // -L2 -Hid0
- elim (le_inv_plus_l … Hid) -Hid /4 width=9 by llpx_sn_lref, yle_trans, yle_inj/ (**) (* slow *)
- ]
-| #L1 #L2 #d0 #i #HL1 #HL2 #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_lref2 … H) -H
- * #_ #H destruct
- lapply (ldrop_fwd_length_eq1 … HLK1 HLK2 HL12)
- [ lapply (ldrop_fwd_length_le4 … HLK1) -HLK1
- lapply (ldrop_fwd_length_le4 … HLK2) -HLK2
- #HKL2 #HKL1 #HK12 @llpx_sn_free // /2 width=3 by transitive_le/ (**) (* full auto too slow *)
- | lapply (ldrop_fwd_length … HLK1) -HLK1 #H >H in HL1; -H
- lapply (ldrop_fwd_length … HLK2) -HLK2 #H >H in HL2; -H
- /3 width=1 by llpx_sn_free, le_plus_to_minus_r/
- ]
-| #L1 #L2 #d0 #p #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #_ >(lift_inv_gref2 … H) -X
- lapply (ldrop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -d -e
- /2 width=1 by llpx_sn_gref/
-| #a #I #L1 #L2 #W #U #d0 #_ #_ #IHW #IHU #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_bind2 … H) -H
- #V #T #HVW #HTU #H destruct /4 width=6 by llpx_sn_bind, ldrop_skip, yle_succ/
-| #I #L1 #L2 #W #U #d0 #_ #_ #IHW #IHU #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_flat2 … H) -H
- #V #T #HVW #HTU #H destruct /3 width=6 by llpx_sn_flat/
-]
-qed-.
-
-lemma llpx_sn_inv_lift_be: ∀R,L1,L2,U,d0. llpx_sn R d0 U L1 L2 →
- ∀K1,K2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 →
- ∀T. ⇧[d, e] T ≡ U → d ≤ d0 → d0 ≤ yinj d + e → llpx_sn R d T K1 K2.
-#R #L1 #L2 #U #d0 #H elim H -L1 -L2 -U -d0
-[ #L1 #L2 #d0 #k #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #_ #_ >(lift_inv_sort2 … H) -X
- lapply (ldrop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -d0 -e
- /2 width=1 by llpx_sn_sort/
-| #L1 #L2 #d0 #i #HL12 #Hid0 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hd0 #Hd0e elim (lift_inv_lref2 … H) -H
- * #Hid #H destruct
- [ lapply (ldrop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2
- -Hid0 /3 width=1 by llpx_sn_skip, ylt_inj/
- | elim (ylt_yle_false … Hid0) -L1 -L2 -Hd0 -Hid0
- /3 width=3 by yle_trans, yle_inj/ (**) (* slow *)
- ]
-| #I #L1 #L2 #K11 #K22 #W1 #W2 #d0 #i #Hid0 #HLK11 #HLK22 #HK12 #HW12 #_ #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hd0 #Hd0e elim (lift_inv_lref2 … H) -H
- * #Hid #H destruct
- [ elim (ylt_yle_false … Hid0) -I -L1 -L2 -K11 -K22 -W1 -W2 -Hd0e -Hid0
- /3 width=3 by ylt_yle_trans, ylt_inj/
- | lapply (ldrop_conf_ge … HLK1 … HLK11 ?) // -L1
- lapply (ldrop_conf_ge … HLK2 … HLK22 ?) // -L2 -Hid0 -Hd0 -Hd0e
- elim (le_inv_plus_l … Hid) -Hid /3 width=9 by llpx_sn_lref, yle_inj/
- ]
-| #L1 #L2 #d0 #i #HL1 #HL2 #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hd0 #Hd0e elim (lift_inv_lref2 … H) -H
- * #_ #H destruct
- lapply (ldrop_fwd_length_eq1 … HLK1 HLK2 HL12)
- [ lapply (ldrop_fwd_length_le4 … HLK1) -HLK1
- lapply (ldrop_fwd_length_le4 … HLK2) -HLK2
- #HKL2 #HKL1 #HK12 @llpx_sn_free // /2 width=3 by transitive_le/ (**) (* full auto too slow *)
- | lapply (ldrop_fwd_length … HLK1) -HLK1 #H >H in HL1; -H
- lapply (ldrop_fwd_length … HLK2) -HLK2 #H >H in HL2; -H
- /3 width=1 by llpx_sn_free, le_plus_to_minus_r/
- ]
-| #L1 #L2 #d0 #p #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #_ #_ >(lift_inv_gref2 … H) -X
- lapply (ldrop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -d0 -e
- /2 width=1 by llpx_sn_gref/
-| #a #I #L1 #L2 #W #U #d0 #_ #_ #IHW #IHU #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hd0 #Hd0e elim (lift_inv_bind2 … H) -H
- >commutative_plus #V #T #HVW #HTU #H destruct
- @llpx_sn_bind [ /2 width=5 by/ ] -IHW (**) (* explicit constructor *)
- @(IHU … HTU) -IHU -HTU /2 width=1 by ldrop_skip, yle_succ/
-| #I #L1 #L2 #W #U #d0 #_ #_ #IHW #IHU #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hd0 #Hd0e elim (lift_inv_flat2 … H) -H
- #V #T #HVW #HTU #H destruct /3 width=6 by llpx_sn_flat/
-]
-qed-.
-
-lemma llpx_sn_inv_lift_ge: ∀R,L1,L2,U,d0. llpx_sn R d0 U L1 L2 →
- ∀K1,K2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 →
- ∀T. ⇧[d, e] T ≡ U → yinj d + e ≤ d0 → llpx_sn R (d0-e) T K1 K2.
-#R #L1 #L2 #U #d0 #H elim H -L1 -L2 -U -d0
-[ #L1 #L2 #d0 #k #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #_ >(lift_inv_sort2 … H) -X
- lapply (ldrop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -d
- /2 width=1 by llpx_sn_sort/
-| #L1 #L2 #d0 #i #HL12 #Hid0 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hded0 elim (lift_inv_lref2 … H) -H
- * #Hid #H destruct [ -Hid0 | -Hded0 ]
- lapply (ldrop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2
- [ /4 width=3 by llpx_sn_skip, yle_plus1_to_minus_inj2, ylt_yle_trans, ylt_inj/
- | elim (le_inv_plus_l … Hid) -Hid #_
- /4 width=1 by llpx_sn_skip, monotonic_ylt_minus_dx, yle_inj/
- ]
-| #I #L1 #L2 #K11 #K22 #W1 #W2 #d0 #i #Hid0 #HLK11 #HLK22 #HK12 #HW12 #_ #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hded0 elim (lift_inv_lref2 … H) -H
- * #Hid #H destruct
- [ elim (ylt_yle_false … Hid0) -I -L1 -L2 -K11 -K22 -W1 -W2 -Hid0
- /3 width=3 by yle_fwd_plus_sn1, ylt_yle_trans, ylt_inj/
- | lapply (ldrop_conf_ge … HLK1 … HLK11 ?) // -L1
- lapply (ldrop_conf_ge … HLK2 … HLK22 ?) // -L2 -Hded0 -Hid
- /3 width=9 by llpx_sn_lref, monotonic_yle_minus_dx/
- ]
-| #L1 #L2 #d0 #i #HL1 #HL2 #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hded0 elim (lift_inv_lref2 … H) -H
- * #_ #H destruct
- lapply (ldrop_fwd_length_eq1 … HLK1 HLK2 HL12)
- [ lapply (ldrop_fwd_length_le4 … HLK1) -HLK1
- lapply (ldrop_fwd_length_le4 … HLK2) -HLK2
- #HKL2 #HKL1 #HK12 @llpx_sn_free // /2 width=3 by transitive_le/ (**) (* full auto too slow *)
- | lapply (ldrop_fwd_length … HLK1) -HLK1 #H >H in HL1; -H
- lapply (ldrop_fwd_length … HLK2) -HLK2 #H >H in HL2; -H
- /3 width=1 by llpx_sn_free, le_plus_to_minus_r/
- ]
-| #L1 #L2 #d0 #p #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #_ >(lift_inv_gref2 … H) -X
- lapply (ldrop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -d
- /2 width=1 by llpx_sn_gref/
-| #a #I #L1 #L2 #W #U #d0 #_ #_ #IHW #IHU #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hded0 elim (lift_inv_bind2 … H) -H
- #V #T #HVW #HTU #H destruct
- @llpx_sn_bind [ /2 width=5 by/ ] -IHW (**) (* explicit constructor *)
- <yminus_succ1_inj /2 width=2 by yle_fwd_plus_sn2/
- @(IHU … HTU) -IHU -HTU /2 width=1 by ldrop_skip, yle_succ/
-| #I #L1 #L2 #W #U #d0 #_ #_ #IHW #IHU #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hded0 elim (lift_inv_flat2 … H) -H
- #V #T #HVW #HTU #H destruct /3 width=5 by llpx_sn_flat/
-]
-qed-.
-
-(* Advanced inversion lemmas on relocation **********************************)
-
-lemma llpx_sn_inv_lift_O: ∀R,L1,L2,U. llpx_sn R 0 U L1 L2 →
- ∀K1,K2,e. ⇩[e] L1 ≡ K1 → ⇩[e] L2 ≡ K2 →
- ∀T. ⇧[0, e] T ≡ U → llpx_sn R 0 T K1 K2.
-/2 width=11 by llpx_sn_inv_lift_be/ qed-.
-
-lemma llpx_sn_ldrop_conf_O: ∀R,L1,L2,U. llpx_sn R 0 U L1 L2 →
- ∀K1,e. ⇩[e] L1 ≡ K1 → ∀T. ⇧[0, e] T ≡ U →
- ∃∃K2. ⇩[e] L2 ≡ K2 & llpx_sn R 0 T K1 K2.
-#R #L1 #L2 #U #HU #K1 #e #HLK1 #T #HTU elim (llpx_sn_fwd_ldrop_sn … HU … HLK1)
-/3 width=10 by llpx_sn_inv_lift_O, ex2_intro/
-qed-.
-
-lemma llpx_sn_ldrop_trans_O: ∀R,L1,L2,U. llpx_sn R 0 U L1 L2 →
- ∀K2,e. ⇩[e] L2 ≡ K2 → ∀T. ⇧[0, e] T ≡ U →
- ∃∃K1. ⇩[e] L1 ≡ K1 & llpx_sn R 0 T K1 K2.
-#R #L1 #L2 #U #HU #K2 #e #HLK2 #T #HTU elim (llpx_sn_fwd_ldrop_dx … HU … HLK2)
-/3 width=10 by llpx_sn_inv_lift_O, ex2_intro/
-qed-.
-
-(* Inversion lemmas on negated lazy pointwise extension *********************)
-
-lemma nllpx_sn_inv_bind: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) →
- ∀a,I,L1,L2,V,T,d. (llpx_sn R d (ⓑ{a,I}V.T) L1 L2 → ⊥) →
- (llpx_sn R d V L1 L2 → ⊥) ∨ (llpx_sn R (⫯d) T (L1.ⓑ{I}V) (L2.ⓑ{I}V) → ⊥).
-#R #HR #a #I #L1 #L2 #V #T #d #H elim (llpx_sn_dec … HR V L1 L2 d)
-/4 width=1 by llpx_sn_bind, or_intror, or_introl/
-qed-.
-
-lemma nllpx_sn_inv_flat: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) →
- ∀I,L1,L2,V,T,d. (llpx_sn R d (ⓕ{I}V.T) L1 L2 → ⊥) →
- (llpx_sn R d V L1 L2 → ⊥) ∨ (llpx_sn R d T L1 L2 → ⊥).
-#R #HR #I #L1 #L2 #V #T #d #H elim (llpx_sn_dec … HR V L1 L2 d)
-/4 width=1 by llpx_sn_flat, or_intror, or_introl/
-qed-.
-
-lemma nllpx_sn_inv_bind_O: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) →
- ∀a,I,L1,L2,V,T. (llpx_sn R 0 (ⓑ{a,I}V.T) L1 L2 → ⊥) →
- (llpx_sn R 0 V L1 L2 → ⊥) ∨ (llpx_sn R 0 T (L1.ⓑ{I}V) (L2.ⓑ{I}V) → ⊥).
-#R #HR #a #I #L1 #L2 #V #T #H elim (llpx_sn_dec … HR V L1 L2 0)
-/4 width=1 by llpx_sn_bind_O, or_intror, or_introl/
-qed-.
(* *)
(**************************************************************************)
-include "basic_2/substitution/ldrop_leq.ma".
+include "basic_2/substitution/drop_leq.ma".
include "basic_2/multiple/llpx_sn.ma".
(* LAZY SN POINTWISE EXTENSION OF A CONTEXT-SENSITIVE REALTION FOR TERMS ****)
(* Properties on equivalence for local environments *************************)
lemma leq_llpx_sn_trans: ∀R,L2,L,T,d. llpx_sn R d T L2 L →
- â\88\80L1. L1 â\89\83[d, ∞] L2 → llpx_sn R d T L1 L.
+ â\88\80L1. L1 ⩬[d, ∞] L2 → llpx_sn R d T L1 L.
#R #L2 #L #T #d #H elim H -L2 -L -T -d
/4 width=5 by llpx_sn_flat, llpx_sn_gref, llpx_sn_skip, llpx_sn_sort, leq_fwd_length, trans_eq/
[ #I #L2 #L #K2 #K #V2 #V #d #i #Hdi #HLK2 #HLK #HK2 #HV2 #_ #L1 #HL12
- elim (leq_ldrop_trans_be … HL12 … HLK2) -L2 // >yminus_Y_inj #K1 #HK12 #HLK1
+ elim (leq_drop_trans_be … HL12 … HLK2) -L2 // >yminus_Y_inj #K1 #HK12 #HLK1
lapply (leq_inv_O_Y … HK12) -HK12 #H destruct /2 width=9 by llpx_sn_lref/
| /4 width=5 by llpx_sn_free, leq_fwd_length, le_repl_sn_trans_aux, trans_eq/
| /4 width=1 by llpx_sn_bind, leq_succ/
qed-.
lemma llpx_sn_leq_trans: ∀R,L,L1,T,d. llpx_sn R d T L L1 →
- â\88\80L2. L1 â\89\83[d, ∞] L2 → llpx_sn R d T L L2.
+ â\88\80L2. L1 ⩬[d, ∞] L2 → llpx_sn R d T L L2.
#R #L #L1 #T #d #H elim H -L -L1 -T -d
/4 width=5 by llpx_sn_flat, llpx_sn_gref, llpx_sn_skip, llpx_sn_sort, leq_fwd_length, trans_eq/
[ #I #L #L1 #K #K1 #V #V1 #d #i #Hdi #HLK #HLK1 #HK1 #HV1 #_ #L2 #HL12
- elim (leq_ldrop_conf_be … HL12 … HLK1) -L1 // >yminus_Y_inj #K2 #HK12 #HLK2
+ elim (leq_drop_conf_be … HL12 … HLK1) -L1 // >yminus_Y_inj #K2 #HK12 #HLK2
lapply (leq_inv_O_Y … HK12) -HK12 #H destruct /2 width=9 by llpx_sn_lref/
| /4 width=5 by llpx_sn_free, leq_fwd_length, le_repl_sn_conf_aux, trans_eq/
| /4 width=1 by llpx_sn_bind, leq_succ/
]
qed-.
-lemma llpx_sn_leq_repl: â\88\80R,L1,L2,T,d. llpx_sn R d T L1 L2 â\86\92 â\88\80K1. K1 â\89\83[d, ∞] L1 →
- â\88\80K2. L2 â\89\83[d, ∞] K2 → llpx_sn R d T K1 K2.
+lemma llpx_sn_leq_repl: â\88\80R,L1,L2,T,d. llpx_sn R d T L1 L2 â\86\92 â\88\80K1. K1 ⩬[d, ∞] L1 →
+ â\88\80K2. L2 ⩬[d, ∞] K2 → llpx_sn R d T K1 K2.
/3 width=4 by llpx_sn_leq_trans, leq_llpx_sn_trans/ qed-.
lemma llpx_sn_bind_repl_SO: ∀R,I1,I2,L1,L2,V1,V2,T. llpx_sn R 0 T (L1.ⓑ{I1}V1) (L2.ⓑ{I2}V2) →
elim (llpx_sn_llpx_sn_alt … H1) -H1 #HL12 #IH1
elim H2 -H2 #_ #HL1 #IH2
@lpx_sn_intro_alt // #I1 #I #K1 #K #V1 #V #i #HLK1 #HLK
-lapply (ldrop_fwd_length_lt2 … HLK) #HiL
-elim (ldrop_O1_lt (Ⓕ) L2 i) // -HiL -HL1 -HL12 #I2 #K2 #V2 #HLK2
+lapply (drop_fwd_length_lt2 … HLK) #HiL
+elim (drop_O1_lt (Ⓕ) L2 i) // -HiL -HL1 -HL12 #I2 #K2 #V2 #HLK2
elim (IH2 … HLK1 HLK2 HLK) -IH2 -HLK * /2 width=1 by conj/
#HnT #H1 #H2 elim (IH1 … HnT … HLK1 HLK2) -IH1 -HnT -HLK1 -HLK2 /2 width=1 by conj/
qed-.
(* *)
(**************************************************************************)
-include "basic_2/substitution/lpx_sn_ldrop.ma".
+include "basic_2/substitution/lpx_sn_drop.ma".
include "basic_2/multiple/llpx_sn.ma".
(* LAZY SN POINTWISE EXTENSION OF A CONTEXT-SENSITIVE REALTION FOR TERMS ****)
[2: -IH /4 width=4 by lpx_sn_fwd_length, llpx_sn_free, le_repl_sn_conf_aux/ ]
#Hi #Hn #L2 #d elim (ylt_split i d)
[ -n /3 width=2 by llpx_sn_skip, lpx_sn_fwd_length/ ]
- #Hdi #HL12 elim (ldrop_O1_lt (Ⓕ) L1 i) //
- #I #K1 #V1 #HLK1 elim (lpx_sn_ldrop_conf … HL12 … HLK1) -HL12
- /4 width=9 by llpx_sn_lref, ldrop_fwd_rfw/
+ #Hdi #HL12 elim (drop_O1_lt (Ⓕ) L1 i) //
+ #I #K1 #V1 #HLK1 elim (lpx_sn_drop_conf … HL12 … HLK1) -HL12
+ /4 width=9 by llpx_sn_lref, drop_fwd_rfw/
| -HR -IH /4 width=2 by lpx_sn_fwd_length, llpx_sn_gref/
| /4 width=1 by llpx_sn_bind, lpx_sn_pair/
| -HR /3 width=1 by llpx_sn_flat/
(* *)
(**************************************************************************)
-include "basic_2/multiple/llpx_sn_ldrop.ma".
+include "basic_2/multiple/llpx_sn_drop.ma".
(* LAZY SN POINTWISE EXTENSION OF A CONTEXT-SENSITIVE REALTION FOR TERMS ****)
+++ /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 *)
-(* *)
-(**************************************************************************)
-
-(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-
-notation "hvbox( T1 ≃ break term 46 T2 )"
- non associative with precedence 45
- for @{ 'Iso $T1 $T2 }.
+++ /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 *)
-(* *)
-(**************************************************************************)
-
-(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-
-notation "hvbox( L1 ≃ break [ term 46 d , break term 46 e ] break term 46 L2 )"
- non associative with precedence 45
- for @{ 'Iso $d $e $L1 $L2 }.
--- /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 *)
+(* *)
+(**************************************************************************)
+
+(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
+
+notation "hvbox( L1 ⩬ break [ term 46 d , break term 46 e ] break term 46 L2 )"
+ non associative with precedence 45
+ for @{ 'MidIso $d $e $L1 $L2 }.
--- /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 *)
+(* *)
+(**************************************************************************)
+
+(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
+
+notation "hvbox( T1 ≂ break term 46 T2 )"
+ non associative with precedence 45
+ for @{ 'TopIso $T1 $T2 }.
lemma cix_lref: ∀h,g,G,L,i. ⇩[i] L ≡ ⋆ → ⦃G, L⦄ ⊢ ➡[h, g] 𝐈⦃#i⦄.
#h #g #G #L #i #HL #H elim (crx_inv_lref … H) -h #I #K #V #HLK
-lapply (ldrop_mono … HLK … HL) -L -i #H destruct
+lapply (drop_mono … HLK … HL) -L -i #H destruct
qed.
(* Properties on relocation *************************************************)
lemma cnr_lref_free: ∀G,L,i. |L| ≤ i → ⦃G, L⦄ ⊢ ➡ 𝐍⦃#i⦄.
#G #L #i #Hi #X #H elim (cpr_inv_lref1 … H) -H // *
-#K #V1 #V2 #HLK lapply (ldrop_fwd_length_lt2 … HLK) -HLK
+#K #V1 #V2 #HLK lapply (drop_fwd_length_lt2 … HLK) -HLK
#H elim (lt_refl_false i) /2 width=3 by lt_to_le_to_lt/
qed.
(* Basic_1: was only: nf2_csort_lref *)
lemma cnr_lref_atom: ∀G,L,i. ⇩[i] L ≡ ⋆ → ⦃G, L⦄ ⊢ ➡ 𝐍⦃#i⦄.
-#G #L #i #HL @cnr_lref_free >(ldrop_fwd_length … HL) -HL //
+#G #L #i #HL @cnr_lref_free >(drop_fwd_length … HL) -HL //
qed.
(* Basic_1: was: nf2_abst *)
#G #L #K #V #i #HLK #X #H
elim (cpr_inv_lref1 … H) -H // *
#K0 #V1 #V2 #HLK0 #_ #_
-lapply (ldrop_mono … HLK … HLK0) -L #H destruct
+lapply (drop_mono … HLK … HLK0) -L #H destruct
qed.
(* Relocation properties ****************************************************)
lemma cnx_lref_free: ∀h,g,G,L,i. |L| ≤ i → ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃#i⦄.
#h #g #G #L #i #Hi #X #H elim (cpx_inv_lref1 … H) -H // *
-#I #K #V1 #V2 #HLK lapply (ldrop_fwd_length_lt2 … HLK) -HLK
+#I #K #V1 #V2 #HLK lapply (drop_fwd_length_lt2 … HLK) -HLK
#H elim (lt_refl_false i) /2 width=3 by lt_to_le_to_lt/
qed.
lemma cnx_lref_atom: ∀h,g,G,L,i. ⇩[i] L ≡ ⋆ → ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃#i⦄.
-#h #g #G #L #i #HL @cnx_lref_free >(ldrop_fwd_length … HL) -HL //
+#h #g #G #L #i #HL @cnx_lref_free >(drop_fwd_length … HL) -HL //
qed.
lemma cnx_abst: ∀h,g,a,G,L,W,T. ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃W⦄ → ⦃G, L.ⓛW⦄ ⊢ ➡[h, g] 𝐍⦃T⦄ →
#G #L1 #T1 #T2 #H elim H -G -L1 -T1 -T2
[ //
| #G #L1 #K1 #V1 #V2 #W2 #i #HLK1 #_ #HVW2 #IHV12 #L2 #HL12
- elim (lsubr_fwd_ldrop2_abbr … HL12 … HLK1) -L1 *
+ elim (lsubr_fwd_drop2_abbr … HL12 … HLK1) -L1 *
/3 width=6 by cpr_delta/
|3,7: /4 width=1 by lsubr_bind, cpr_bind, cpr_beta/
|4,6: /3 width=1 by cpr_flat, cpr_eps/
elim (lift_split … HVW i i) /3 width=6 by cpr_delta, ex2_2_intro/
| * [ #a ] #I #W1 #U1 #IHW1 #IHU1 #L #d #HLK
elim (IHW1 … HLK) -IHW1 #W2 #W #HW12 #HW2
- [ elim (IHU1 (L. ⓑ{I}W1) (d+1)) -IHU1 /3 width=9 by ldrop_drop, cpr_bind, lift_bind, ex2_2_intro/
+ [ elim (IHU1 (L. ⓑ{I}W1) (d+1)) -IHU1 /3 width=9 by drop_drop, cpr_bind, lift_bind, ex2_2_intro/
| elim (IHU1 … HLK) -IHU1 -HLK /3 width=8 by cpr_flat, lift_flat, ex2_2_intro/
]
]
(* *)
(**************************************************************************)
-include "basic_2/substitution/ldrop_ldrop.ma".
+include "basic_2/substitution/drop_drop.ma".
include "basic_2/reduction/cpr.ma".
(* CONTEXT-SENSITIVE PARALLEL REDUCTION FOR TERMS ***************************)
| #G #K #KV #V #V2 #W2 #i #HKV #HV2 #HVW2 #IHV2 #L #s #d #e #HLK #U1 #H #U2 #HWU2
elim (lift_inv_lref1 … H) * #Hid #H destruct
[ elim (lift_trans_ge … HVW2 … HWU2) -W2 // <minus_plus #W2 #HVW2 #HWU2
- elim (ldrop_trans_le … HLK … HKV) -K /2 width=2 by lt_to_le/ #X #HLK #H
- elim (ldrop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid
+ elim (drop_trans_le … HLK … HKV) -K /2 width=2 by lt_to_le/ #X #HLK #H
+ elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid
#K #Y #HKV #HVY #H destruct /3 width=9 by cpr_delta/
| lapply (lift_trans_be … HVW2 … HWU2 ? ?) -W2 /2 width=1 by le_S/ >plus_plus_comm_23 #HVU2
- lapply (ldrop_trans_ge_comm … HLK … HKV ?) -K // -Hid /3 width=6 by cpr_delta, ldrop_inv_gen/
+ lapply (drop_trans_ge_comm … HLK … HKV ?) -K // -Hid /3 width=6 by cpr_delta, drop_inv_gen/
]
| #a #I #G #K #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L #s #d #e #HLK #U1 #H1 #U2 #H2
elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 destruct
- elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /4 width=6 by cpr_bind, ldrop_skip/
+ elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /4 width=6 by cpr_bind, drop_skip/
| #I #G #K #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L #s #d #e #HLK #U1 #H1 #U2 #H2
elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 destruct
elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /3 width=6 by cpr_flat/
| #G #K #V #T1 #T #T2 #_ #HT2 #IHT1 #L #s #d #e #HLK #U1 #H #U2 #HTU2
elim (lift_inv_bind1 … H) -H #VV1 #TT1 #HVV1 #HTT1 #H destruct
- elim (lift_conf_O1 … HTU2 … HT2) -T2 /4 width=6 by cpr_zeta, ldrop_skip/
+ elim (lift_conf_O1 … HTU2 … HT2) -T2 /4 width=6 by cpr_zeta, drop_skip/
| #G #K #V #T1 #T2 #_ #IHT12 #L #s #d #e #HLK #U1 #H #U2 #HTU2
elim (lift_inv_flat1 … H) -H #VV1 #TT1 #HVV1 #HTT1 #H destruct /3 width=6 by cpr_eps/
| #a #G #K #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #L #s #d #e #HLK #X1 #HX1 #X2 #HX2
elim (lift_inv_flat1 … HX1) -HX1 #V0 #X #HV10 #HX #HX1 destruct
elim (lift_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT10 #HX destruct
elim (lift_inv_bind1 … HX2) -HX2 #X #T3 #HX #HT23 #HX2 destruct
- elim (lift_inv_flat1 … HX) -HX #W3 #V3 #HW23 #HV23 #HX destruct /4 width=6 by cpr_beta, ldrop_skip/
+ elim (lift_inv_flat1 … HX) -HX #W3 #V3 #HW23 #HV23 #HX destruct /4 width=6 by cpr_beta, drop_skip/
| #a #G #K #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV1 #IHW12 #IHT12 #L #s #d #e #HLK #X1 #HX1 #X2 #HX2
elim (lift_inv_flat1 … HX1) -HX1 #V0 #X #HV10 #HX #HX1 destruct
elim (lift_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT10 #HX destruct
elim (lift_inv_bind1 … HX2) -HX2 #W3 #X #HW23 #HX #HX2 destruct
elim (lift_inv_flat1 … HX) -HX #V3 #T3 #HV3 #HT23 #HX destruct
- elim (lift_trans_ge … HV2 … HV3) -V2 /4 width=6 by cpr_theta, ldrop_skip/
+ elim (lift_trans_ge … HV2 … HV3) -V2 /4 width=6 by cpr_theta, drop_skip/
]
qed.
]
| #G #L #LV #V #V2 #W2 #i #HLV #HV2 #HVW2 #IHV2 #K #s #d #e #HLK #T1 #H
elim (lift_inv_lref2 … H) -H * #Hid #H destruct
- [ elim (ldrop_conf_lt … HLK … HLV) -L // #L #U #HKL #HLV #HUV
+ [ elim (drop_conf_lt … HLK … HLV) -L // #L #U #HKL #HLV #HUV
elim (IHV2 … HLV … HUV) -V #U2 #HUV2 #HU2
elim (lift_trans_le … HUV2 … HVW2) -V2 // >minus_plus <plus_minus_m_m /3 width=8 by cpr_delta, ex2_intro/
| elim (le_inv_plus_l … Hid) #Hdie #Hei
- lapply (ldrop_conf_ge … HLK … HLV ?) -L // #HKLV
+ lapply (drop_conf_ge … HLK … HLV ?) -L // #HKLV
elim (lift_split … HVW2 d (i - e + 1)) -HVW2 [2,3,4: /2 width=1 by le_S_S, le_S/ ] -Hid -Hdie
#V1 #HV1 >plus_minus // <minus_minus /2 width=1 by le_S/ <minus_n_n <plus_n_O
/3 width=8 by cpr_delta, ex2_intro/
| #a #I #G #L #V1 #V2 #U1 #U2 #_ #_ #IHV12 #IHU12 #K #s #d #e #HLK #X #H
elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct
elim (IHV12 … HLK … HWV1) -IHV12 #W2 #HW12 #HWV2
- elim (IHU12 … HTU1) -IHU12 -HTU1 /3 width=6 by cpr_bind, ldrop_skip, lift_bind, ex2_intro/
+ elim (IHU12 … HTU1) -IHU12 -HTU1 /3 width=6 by cpr_bind, drop_skip, lift_bind, ex2_intro/
| #I #G #L #V1 #V2 #U1 #U2 #_ #_ #IHV12 #IHU12 #K #s #d #e #HLK #X #H
elim (lift_inv_flat2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct
elim (IHV12 … HLK … HWV1) -V1
elim (IHU12 … HLK … HTU1) -U1 -HLK /3 width=6 by cpr_flat, lift_flat, ex2_intro/
| #G #L #V #U1 #U #U2 #_ #HU2 #IHU1 #K #s #d #e #HLK #X #H
elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct
- elim (IHU1 (K.ⓓW1) s … HTU1) /2 width=1 by ldrop_skip/ -L -U1 #T #HTU #HT1
+ elim (IHU1 (K.ⓓW1) s … HTU1) /2 width=1 by drop_skip/ -L -U1 #T #HTU #HT1
elim (lift_div_le … HU2 … HTU) -U /3 width=6 by cpr_zeta, ex2_intro/
| #G #L #V #U1 #U2 #_ #IHU12 #K #s #d #e #HLK #X #H
elim (lift_inv_flat2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct
elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct
elim (lift_inv_bind2 … HY) -HY #W0 #T0 #HW01 #HT01 #HY destruct
elim (IHV12 … HLK … HV01) -V1
- elim (IHT12 (K.ⓛW0) s … HT01) -T1 /2 width=1 by ldrop_skip/
+ elim (IHT12 (K.ⓛW0) s … HT01) -T1 /2 width=1 by drop_skip/
elim (IHW12 … HLK … HW01) -W1 /4 width=7 by cpr_beta, lift_flat, lift_bind, ex2_intro/
| #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV1 #IHW12 #IHT12 #K #s #d #e #HLK #X #HX
elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct
elim (lift_inv_bind2 … HY) -HY #W0 #T0 #HW01 #HT01 #HY destruct
elim (IHV1 … HLK … HV01) -V1 #V3 #HV3 #HV03
- elim (IHT12 (K.ⓓW0) s … HT01) -T1 /2 width=1 by ldrop_skip/ #T3 #HT32 #HT03
+ elim (IHT12 (K.ⓓW0) s … HT01) -T1 /2 width=1 by drop_skip/ #T3 #HT32 #HT03
elim (IHW12 … HLK … HW01) -W1
elim (lift_trans_le … HV3 … HV2) -V
/4 width=9 by cpr_theta, lift_flat, lift_bind, ex2_intro/
(* *)
(**************************************************************************)
-include "basic_2/multiple/llpx_sn_ldrop.ma".
+include "basic_2/multiple/llpx_sn_drop.ma".
include "basic_2/reduction/cpr.ma".
(* CONTEXT-SENSITIVE PARALLEL REDUCTION FOR TERMS ***************************)
[ //
| #G #Ls #Ks #V1s #V2s #W2s #i #HLKs #_ #HVW2s #IHV12s #Ld #H elim (llpx_sn_inv_lref_ge_sn … H … HLKs) // -H
#Kd #V1d #HLKd #HV1s #HV1sd
- lapply (ldrop_fwd_drop2 … HLKs) -HLKs #HLKs
- lapply (ldrop_fwd_drop2 … HLKd) -HLKd #HLKd
+ lapply (drop_fwd_drop2 … HLKs) -HLKs #HLKs
+ lapply (drop_fwd_drop2 … HLKd) -HLKd #HLKd
@(llpx_sn_lift_le … HLKs HLKd … HVW2s) -HLKs -HLKd -HVW2s /2 width=1 by/ (**) (* full auto too slow *)
| #a #I #G #Ls #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #Ld #H elim (llpx_sn_inv_bind_O … H) -H
/4 width=5 by llpx_sn_bind_repl_SO, llpx_sn_bind/
| #I #G #Ls #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
/3 width=1 by llpx_sn_flat/
| #G #Ls #V #T1 #T2 #T #_ #HT2 #IHT12 #Ld #H elim (llpx_sn_inv_bind_O … H) -H
- /3 width=10 by llpx_sn_inv_lift_le, ldrop_drop/
+ /3 width=10 by llpx_sn_inv_lift_le, drop_drop/
| #G #Ls #V #T1 #T2 #_ #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H /2 width=1 by/
| #a #G #Ls #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
#HV1 #H elim (llpx_sn_inv_bind_O … H) -H
| #a #G #Ls #V1 #V2 #V #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
#HV1 #H elim (llpx_sn_inv_bind_O … H) -H //
#HW1 #HT1 @llpx_sn_bind_O /2 width=1 by/ @llpx_sn_flat (**) (* full auto too slow *)
- [ /3 width=10 by llpx_sn_lift_le, ldrop_drop/
+ [ /3 width=10 by llpx_sn_lift_le, drop_drop/
| /3 width=4 by llpx_sn_bind_repl_O/
]
qed-.
[ //
| /2 width=2 by cpx_st/
| #I #G #L1 #K1 #V1 #V2 #W2 #i #HLK1 #_ #HVW2 #IHV12 #L2 #HL12
- elim (lsubr_fwd_ldrop2_bind … HL12 … HLK1) -HL12 -HLK1 *
+ elim (lsubr_fwd_drop2_bind … HL12 … HLK1) -HL12 -HLK1 *
/4 width=7 by cpx_delta, cpx_ct/
|4,9: /4 width=1 by cpx_bind, cpx_beta, lsubr_bind/
|5,7,8: /3 width=1 by cpx_flat, cpx_eps, cpx_ct/
elim (lift_split … HVW i i) /3 width=7 by cpx_delta, ex2_2_intro/
| * [ #a ] #I #W1 #U1 #IHW1 #IHU1 #L #d #HLK
elim (IHW1 … HLK) -IHW1 #W2 #W #HW12 #HW2
- [ elim (IHU1 (L. ⓑ{I} W1) (d+1)) -IHU1 /3 width=9 by cpx_bind, ldrop_drop, lift_bind, ex2_2_intro/
+ [ elim (IHU1 (L. ⓑ{I} W1) (d+1)) -IHU1 /3 width=9 by cpx_bind, drop_drop, lift_bind, ex2_2_intro/
| elim (IHU1 … HLK) -IHU1 -HLK /3 width=8 by cpx_flat, lift_flat, ex2_2_intro/
]
]
lemma cpx_inv_lref1_ge: ∀h,g,G,L,T2,i. ⦃G, L⦄ ⊢ #i ➡[h, g] T2 → |L| ≤ i → T2 = #i.
#h #g #G #L #T2 #i #H elim (cpx_inv_lref1 … H) -H // *
-#I #K #V1 #V2 #HLK #_ #_ #HL -h -G -V2 lapply (ldrop_fwd_length_lt2 … HLK) -K -I -V1
+#I #K #V1 #V2 #HLK #_ #_ #HL -h -G -V2 lapply (drop_fwd_length_lt2 … HLK) -K -I -V1
#H elim (lt_refl_false i) /2 width=3 by lt_to_le_to_lt/
qed-.
(* *)
(**************************************************************************)
-include "basic_2/substitution/ldrop_leq.ma".
+include "basic_2/substitution/drop_leq.ma".
include "basic_2/reduction/cpx.ma".
(* CONTEXT-SENSITIVE EXTENDED PARALLEL REDUCTION FOR TERMS ******************)
[ //
| /2 width=2 by cpx_st/
| #I #G #L1 #K1 #V1 #V2 #W2 #i #HLK1 #_ #HVW2 #IHV12 #L2 #HL12
- elim (leq_ldrop_trans_be … HL12 … HLK1) // -HL12 -HLK1 /3 width=7 by cpx_delta/
+ elim (leq_drop_trans_be … HL12 … HLK1) // -HL12 -HLK1 /3 width=7 by cpx_delta/
|4,9: /4 width=1 by cpx_bind, cpx_beta, leq_pair_O_Y/
|5,7,8: /3 width=1 by cpx_flat, cpx_eps, cpx_ct/
|6,10: /4 width=3 by cpx_zeta, cpx_theta, leq_pair_O_Y/
(* *)
(**************************************************************************)
-include "basic_2/substitution/ldrop_ldrop.ma".
+include "basic_2/substitution/drop_drop.ma".
include "basic_2/multiple/fqus_alt.ma".
include "basic_2/static/sta.ma".
include "basic_2/static/da.ma".
[ /3 width=4 by cpx_st, da_inv_sort/
| #G #L #K #V1 #V2 #W2 #i #HLK #_ #HVW2 #IHV12 #H
elim (da_inv_lref … H) -H * #K0 #V0 [| #l0 ] #HLK0
- lapply (ldrop_mono … HLK0 … HLK) -HLK0 #H destruct /3 width=7 by cpx_delta/
+ lapply (drop_mono … HLK0 … HLK) -HLK0 #H destruct /3 width=7 by cpx_delta/
| #G #L #K #W1 #W2 #V1 #i #HLK #_ #HWV1 #IHW12 #H
elim (da_inv_lref … H) -H * #K0 #W0 [| #l1 ] #HLK0
- lapply (ldrop_mono … HLK0 … HLK) -HLK0 #H destruct /3 width=7 by cpx_delta/
+ lapply (drop_mono … HLK0 … HLK) -HLK0 #H destruct /3 width=7 by cpx_delta/
| /4 width=2 by cpx_bind, da_inv_bind/
| /4 width=3 by cpx_flat, da_inv_flat/
| /4 width=3 by cpx_eps, da_inv_flat/
| #I #G #K #KV #V #V2 #W2 #i #HKV #HV2 #HVW2 #IHV2 #L #s #d #e #HLK #U1 #H #U2 #HWU2
elim (lift_inv_lref1 … H) * #Hid #H destruct
[ elim (lift_trans_ge … HVW2 … HWU2) -W2 // <minus_plus #W2 #HVW2 #HWU2
- elim (ldrop_trans_le … HLK … HKV) -K /2 width=2 by lt_to_le/ #X #HLK #H
- elim (ldrop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid
+ elim (drop_trans_le … HLK … HKV) -K /2 width=2 by lt_to_le/ #X #HLK #H
+ elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid
#K #Y #HKV #HVY #H destruct /3 width=10 by cpx_delta/
| lapply (lift_trans_be … HVW2 … HWU2 ? ?) -W2 /2 width=1 by le_S/ >plus_plus_comm_23 #HVU2
- lapply (ldrop_trans_ge_comm … HLK … HKV ?) -K /3 width=7 by cpx_delta, ldrop_inv_gen/
+ lapply (drop_trans_ge_comm … HLK … HKV ?) -K /3 width=7 by cpx_delta, drop_inv_gen/
]
| #a #I #G #K #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L #s #d #e #HLK #U1 #H1 #U2 #H2
elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 destruct
- elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /4 width=6 by cpx_bind, ldrop_skip/
+ elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /4 width=6 by cpx_bind, drop_skip/
| #I #G #K #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L #s #d #e #HLK #U1 #H1 #U2 #H2
elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 destruct
elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /3 width=6 by cpx_flat/
| #G #K #V #T1 #T #T2 #_ #HT2 #IHT1 #L #s #d #e #HLK #U1 #H #U2 #HTU2
elim (lift_inv_bind1 … H) -H #VV1 #TT1 #HVV1 #HTT1 #H destruct
- elim (lift_conf_O1 … HTU2 … HT2) -T2 /4 width=6 by cpx_zeta, ldrop_skip/
+ elim (lift_conf_O1 … HTU2 … HT2) -T2 /4 width=6 by cpx_zeta, drop_skip/
| #G #K #V #T1 #T2 #_ #IHT12 #L #s #d #e #HLK #U1 #H #U2 #HTU2
elim (lift_inv_flat1 … H) -H #VV1 #TT1 #HVV1 #HTT1 #H destruct /3 width=6 by cpx_eps/
| #G #K #V1 #V2 #T #_ #IHV12 #L #s #d #e #HLK #U1 #H #U2 #HVU2
elim (lift_inv_flat1 … HX1) -HX1 #V0 #X #HV10 #HX #HX1 destruct
elim (lift_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT10 #HX destruct
elim (lift_inv_bind1 … HX2) -HX2 #X #T3 #HX #HT23 #HX2 destruct
- elim (lift_inv_flat1 … HX) -HX #W3 #V3 #HW23 #HV23 #HX destruct /4 width=6 by cpx_beta, ldrop_skip/
+ elim (lift_inv_flat1 … HX) -HX #W3 #V3 #HW23 #HV23 #HX destruct /4 width=6 by cpx_beta, drop_skip/
| #a #G #K #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV1 #IHW12 #IHT12 #L #s #d #e #HLK #X1 #HX1 #X2 #HX2
elim (lift_inv_flat1 … HX1) -HX1 #V0 #X #HV10 #HX #HX1 destruct
elim (lift_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT10 #HX destruct
elim (lift_inv_bind1 … HX2) -HX2 #W3 #X #HW23 #HX #HX2 destruct
elim (lift_inv_flat1 … HX) -HX #V3 #T3 #HV3 #HT23 #HX destruct
- elim (lift_trans_ge … HV2 … HV3) -V2 /4 width=6 by cpx_theta, ldrop_skip/
+ elim (lift_trans_ge … HV2 … HV3) -V2 /4 width=6 by cpx_theta, drop_skip/
]
qed.
lapply (lift_inv_sort2 … H) -H #H destruct /3 width=3 by cpx_st, lift_sort, ex2_intro/
| #I #G #L #LV #V #V2 #W2 #i #HLV #HV2 #HVW2 #IHV2 #K #s #d #e #HLK #T1 #H
elim (lift_inv_lref2 … H) -H * #Hid #H destruct
- [ elim (ldrop_conf_lt … HLK … HLV) -L // #L #U #HKL #HLV #HUV
+ [ elim (drop_conf_lt … HLK … HLV) -L // #L #U #HKL #HLV #HUV
elim (IHV2 … HLV … HUV) -V #U2 #HUV2 #HU2
elim (lift_trans_le … HUV2 … HVW2) -V2 // >minus_plus <plus_minus_m_m /3 width=9 by cpx_delta, ex2_intro/
| elim (le_inv_plus_l … Hid) #Hdie #Hei
- lapply (ldrop_conf_ge … HLK … HLV ?) -L // #HKLV
+ lapply (drop_conf_ge … HLK … HLV ?) -L // #HKLV
elim (lift_split … HVW2 d (i - e + 1)) -HVW2 /3 width=1 by le_S, le_S_S/ -Hid -Hdie
#V1 #HV1 >plus_minus // <minus_minus /2 width=1 by le_S/ <minus_n_n <plus_n_O /3 width=9 by cpx_delta, ex2_intro/
]
| #a #I #G #L #V1 #V2 #U1 #U2 #_ #_ #IHV12 #IHU12 #K #s #d #e #HLK #X #H
elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct
elim (IHV12 … HLK … HWV1) -IHV12 #W2 #HW12 #HWV2
- elim (IHU12 … HTU1) -IHU12 -HTU1 /3 width=6 by cpx_bind, ldrop_skip, lift_bind, ex2_intro/
+ elim (IHU12 … HTU1) -IHU12 -HTU1 /3 width=6 by cpx_bind, drop_skip, lift_bind, ex2_intro/
| #I #G #L #V1 #V2 #U1 #U2 #_ #_ #IHV12 #IHU12 #K #s #d #e #HLK #X #H
elim (lift_inv_flat2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct
elim (IHV12 … HLK … HWV1) -V1
elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct
elim (lift_inv_bind2 … HY) -HY #W0 #T0 #HW01 #HT01 #HY destruct
elim (IHV12 … HLK … HV01) -V1 #V3 #HV32 #HV03
- elim (IHT12 (K.ⓛW0) s … HT01) -T1 /2 width=1 by ldrop_skip/ #T3 #HT32 #HT03
+ elim (IHT12 (K.ⓛW0) s … HT01) -T1 /2 width=1 by drop_skip/ #T3 #HT32 #HT03
elim (IHW12 … HLK … HW01) -W1
/4 width=7 by cpx_beta, lift_bind, lift_flat, ex2_intro/
| #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV1 #IHW12 #IHT12 #K #s #d #e #HLK #X #HX
elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct
elim (lift_inv_bind2 … HY) -HY #W0 #T0 #HW01 #HT01 #HY destruct
elim (IHV1 … HLK … HV01) -V1 #V3 #HV3 #HV03
- elim (IHT12 (K.ⓓW0) s … HT01) -T1 /2 width=1 by ldrop_skip/ #T3 #HT32 #HT03
+ elim (IHT12 (K.ⓓW0) s … HT01) -T1 /2 width=1 by drop_skip/ #T3 #HT32 #HT03
elim (IHW12 … HLK … HW01) -W1 #W3 #HW32 #HW03
elim (lift_trans_le … HV3 … HV2) -V
/4 width=9 by cpx_theta, lift_bind, lift_flat, ex2_intro/
/3 width=3 by fqu_pair_sn, fqu_bind_dx, fqu_flat_dx, cpx_pair_sn, cpx_bind, cpx_flat, ex2_intro/
[ #I #G #L #V2 #U2 #HVU2
elim (lift_total U2 0 1)
- /4 width=7 by fqu_drop, cpx_delta, ldrop_pair, ldrop_drop, ex2_intro/
+ /4 width=7 by fqu_drop, cpx_delta, drop_pair, drop_drop, ex2_intro/
| #G #L #K #T1 #U1 #e #HLK1 #HTU1 #T2 #HTU2
elim (lift_total T2 0 (e+1))
/3 width=11 by cpx_lift, fqu_drop, ex2_intro/
#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
[ #I #G #L #V1 #V2 #HV12 #_ elim (lift_total V2 0 1)
#U2 #HVU2 @(ex3_intro … U2)
- [1,3: /3 width=7 by fqu_drop, cpx_delta, ldrop_pair, ldrop_drop/
+ [1,3: /3 width=7 by fqu_drop, cpx_delta, drop_pair, drop_drop/
| #H destruct /2 width=7 by lift_inv_lref2_be/
]
| #I #G #L #V1 #T #V2 #HV12 #H @(ex3_intro … (②{I}V2.T))
(* *)
(**************************************************************************)
-include "basic_2/multiple/lleq_ldrop.ma".
+include "basic_2/multiple/lleq_drop.ma".
include "basic_2/reduction/cpx_llpx_sn.ma".
(* CONTEXT-SENSITIVE EXTENDED PARALLEL REDUCTION FOR TERMS ******************)
(* *)
(**************************************************************************)
-include "basic_2/multiple/llpx_sn_ldrop.ma".
+include "basic_2/multiple/llpx_sn_drop.ma".
include "basic_2/reduction/cpx.ma".
(* CONTEXT-SENSITIVE EXTENDED PARALLEL REDUCTION FOR TERMS ******************)
| /3 width=4 by llpx_sn_fwd_length, llpx_sn_sort/
| #I #G #Ls #Ks #V1s #V2s #W2s #i #HLKs #_ #HVW2s #IHV12s #Ld #H elim (llpx_sn_inv_lref_ge_sn … H … HLKs) // -H
#Kd #V1d #HLKd #HV1s #HV1sd
- lapply (ldrop_fwd_drop2 … HLKs) -HLKs #HLKs
- lapply (ldrop_fwd_drop2 … HLKd) -HLKd #HLKd
+ lapply (drop_fwd_drop2 … HLKs) -HLKs #HLKs
+ lapply (drop_fwd_drop2 … HLKd) -HLKd #HLKd
@(llpx_sn_lift_le … HLKs HLKd … HVW2s) -HLKs -HLKd -HVW2s /2 width=1 by/ (**) (* full auto too slow *)
| #a #I #G #Ls #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #Ld #H elim (llpx_sn_inv_bind_O … H) -H
/4 width=5 by llpx_sn_bind_repl_SO, llpx_sn_bind/
| #I #G #Ls #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
/3 width=1 by llpx_sn_flat/
| #G #Ls #V #T1 #T2 #T #_ #HT2 #IHT12 #Ld #H elim (llpx_sn_inv_bind_O … H) -H
- /3 width=10 by llpx_sn_inv_lift_le, ldrop_drop/
+ /3 width=10 by llpx_sn_inv_lift_le, drop_drop/
| #G #Ls #V #T1 #T2 #_ #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H /2 width=1 by/
| #G #Ls #V1 #V2 #T #_ #IHV12 #Ld #H elim (llpx_sn_inv_flat … H) -H /2 width=1 by/
| #a #G #Ls #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
| #a #G #Ls #V1 #V2 #V #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
#HV1 #H elim (llpx_sn_inv_bind_O … H) -H //
#HW1 #HT1 @llpx_sn_bind_O /2 width=1 by/ @llpx_sn_flat (**) (* full auto too slow *)
- [ /3 width=10 by llpx_sn_lift_le, ldrop_drop/
+ [ /3 width=10 by llpx_sn_lift_le, drop_drop/
| /3 width=4 by llpx_sn_bind_repl_O/
]
qed-.
include "basic_2/notation/relations/predreducible_3.ma".
include "basic_2/grammar/genv.ma".
-include "basic_2/substitution/ldrop.ma".
+include "basic_2/substitution/drop.ma".
(* REDUCIBLE TERMS FOR CONTEXT-SENSITIVE REDUCTION **************************)
#G * #i #H
[ elim (crr_inv_sort … H)
| elim (crr_inv_lref … H) -H #L #V #H
- elim (ldrop_inv_atom1 … H) -H #H destruct
+ elim (drop_inv_atom1 … H) -H #H destruct
| elim (crr_inv_gref … H)
]
qed-.
(* *)
(**************************************************************************)
-include "basic_2/substitution/ldrop_ldrop.ma".
+include "basic_2/substitution/drop_drop.ma".
include "basic_2/reduction/crr.ma".
(* REDUCIBLE TERMS FOR CONTEXT-SENSITIVE REDUCTION **************************)
#G #K #T #H elim H -K -T
[ #K #K0 #V #i #HK0 #L #s #d #e #HLK #X #H
elim (lift_inv_lref1 … H) -H * #Hid #H destruct
- [ elim (ldrop_trans_lt … HLK … HK0) -K /2 width=4 by crr_delta/
- | lapply (ldrop_trans_ge … HLK … HK0 ?) -K /3 width=4 by crr_delta, ldrop_inv_gen/
+ [ elim (drop_trans_lt … HLK … HK0) -K /2 width=4 by crr_delta/
+ | lapply (drop_trans_ge … HLK … HK0 ?) -K /3 width=4 by crr_delta, drop_inv_gen/
]
| #K #V #T #_ #IHV #L #s #d #e #HLK #X #H
elim (lift_inv_flat1 … H) -H #W #U #HVW #_ #H destruct /3 width=5 by crr_appl_sn/
| #a #I #K #V #T #HI #_ #IHV #L #s #d #e #HLK #X #H
elim (lift_inv_bind1 … H) -H #W #U #HVW #_ #H destruct /3 width=5 by crr_ib2_sn/
| #a #I #K #V #T #HI #_ #IHT #L #s #d #e #HLK #X #H
- elim (lift_inv_bind1 … H) -H #W #U #HVW #HTU #H destruct /4 width=5 by crr_ib2_dx, ldrop_skip/
+ elim (lift_inv_bind1 … H) -H #W #U #HVW #HTU #H destruct /4 width=5 by crr_ib2_dx, drop_skip/
| #a #K #V #V0 #T #L #s #d #e #_ #X #H
elim (lift_inv_flat1 … H) -H #W #X0 #_ #H0 #H destruct
elim (lift_inv_bind1 … H0) -H0 #W0 #U #_ #_ #H0 destruct /2 width=1 by crr_beta/
#G #L #U #H elim H -L -U
[ #L #L0 #W #i #HK0 #K #s #d #e #HLK #X #H
elim (lift_inv_lref2 … H) -H * #Hid #H destruct
- [ elim (ldrop_conf_lt … HLK … HK0) -L /2 width=4 by crr_delta/
- | lapply (ldrop_conf_ge … HLK … HK0 ?) -L /2 width=4 by crr_delta/
+ [ elim (drop_conf_lt … HLK … HK0) -L /2 width=4 by crr_delta/
+ | lapply (drop_conf_ge … HLK … HK0 ?) -L /2 width=4 by crr_delta/
]
| #L #W #U #_ #IHW #K #s #d #e #HLK #X #H
elim (lift_inv_flat2 … H) -H #V #T #HVW #_ #H destruct /3 width=5 by crr_appl_sn/
| #a #I #L #W #U #HI #_ #IHW #K #s #d #e #HLK #X #H
elim (lift_inv_bind2 … H) -H #V #T #HVW #_ #H destruct /3 width=5 by crr_ib2_sn/
| #a #I #L #W #U #HI #_ #IHU #K #s #d #e #HLK #X #H
- elim (lift_inv_bind2 … H) -H #V #T #HVW #HTU #H destruct /4 width=5 by crr_ib2_dx, ldrop_skip/
+ elim (lift_inv_bind2 … H) -H #V #T #HVW #HTU #H destruct /4 width=5 by crr_ib2_dx, drop_skip/
| #a #L #W #W0 #U #K #s #d #e #_ #X #H
elim (lift_inv_flat2 … H) -H #V #X0 #_ #H0 #H destruct
elim (lift_inv_bind2 … H0) -H0 #V0 #T #_ #_ #H0 destruct /2 width=1 by crr_beta/
#h #g * #i #G #H
[ elim (crx_inv_sort … H) -H /2 width=4 by ex2_2_intro/
| elim (crx_inv_lref … H) -H #I #L #V #H
- elim (ldrop_inv_atom1 … H) -H #H destruct
+ elim (drop_inv_atom1 … H) -H #H destruct
| elim (crx_inv_gref … H)
]
qed-.
(* *)
(**************************************************************************)
-include "basic_2/substitution/ldrop_ldrop.ma".
+include "basic_2/substitution/drop_drop.ma".
include "basic_2/reduction/crx.ma".
(* REDUCIBLE TERMS FOR CONTEXT-SENSITIVE EXTENDED REDUCTION *****************)
>(lift_inv_sort1 … H) -X /2 width=2 by crx_sort/
| #I #K #K0 #V #i #HK0 #L #s #d #e #HLK #X #H
elim (lift_inv_lref1 … H) -H * #Hid #H destruct
- [ elim (ldrop_trans_lt … HLK … HK0) -K /2 width=4 by crx_delta/
- | lapply (ldrop_trans_ge … HLK … HK0 ?) -K /3 width=5 by crx_delta, ldrop_inv_gen/
+ [ elim (drop_trans_lt … HLK … HK0) -K /2 width=4 by crx_delta/
+ | lapply (drop_trans_ge … HLK … HK0 ?) -K /3 width=5 by crx_delta, drop_inv_gen/
]
| #K #V #T #_ #IHV #L #s #d #e #HLK #X #H
elim (lift_inv_flat1 … H) -H #W #U #HVW #_ #H destruct /3 width=5 by crx_appl_sn/
| #a #I #K #V #T #HI #_ #IHV #L #s #d #e #HLK #X #H
elim (lift_inv_bind1 … H) -H #W #U #HVW #_ #H destruct /3 width=5 by crx_ib2_sn/
| #a #I #K #V #T #HI #_ #IHT #L #s #d #e #HLK #X #H
- elim (lift_inv_bind1 … H) -H #W #U #HVW #HTU #H destruct /4 width=5 by crx_ib2_dx, ldrop_skip/
+ elim (lift_inv_bind1 … H) -H #W #U #HVW #HTU #H destruct /4 width=5 by crx_ib2_dx, drop_skip/
| #a #K #V #V0 #T #L #s #d #e #_ #X #H
elim (lift_inv_flat1 … H) -H #W #X0 #_ #H0 #H destruct
elim (lift_inv_bind1 … H0) -H0 #W0 #U #_ #_ #H0 destruct /2 width=1 by crx_beta/
>(lift_inv_sort2 … H) -X /2 width=2 by crx_sort/
| #I #L #L0 #W #i #HK0 #K #s #d #e #HLK #X #H
elim (lift_inv_lref2 … H) -H * #Hid #H destruct
- [ elim (ldrop_conf_lt … HLK … HK0) -L /2 width=4 by crx_delta/
- | lapply (ldrop_conf_ge … HLK … HK0 ?) -L /2 width=4 by crx_delta/
+ [ elim (drop_conf_lt … HLK … HK0) -L /2 width=4 by crx_delta/
+ | lapply (drop_conf_ge … HLK … HK0 ?) -L /2 width=4 by crx_delta/
]
| #L #W #U #_ #IHW #K #s #d #e #HLK #X #H
elim (lift_inv_flat2 … H) -H #V #T #HVW #_ #H destruct /3 width=5 by crx_appl_sn/
| #a #I #L #W #U #HI #_ #IHW #K #s #d #e #HLK #X #H
elim (lift_inv_bind2 … H) -H #V #T #HVW #_ #H destruct /3 width=5 by crx_ib2_sn/
| #a #I #L #W #U #HI #_ #IHU #K #s #d #e #HLK #X #H
- elim (lift_inv_bind2 … H) -H #V #T #HVW #HTU #H destruct /4 width=5 by crx_ib2_dx, ldrop_skip/
+ elim (lift_inv_bind2 … H) -H #V #T #HVW #HTU #H destruct /4 width=5 by crx_ib2_dx, drop_skip/
| #a #L #W #W0 #U #K #s #d #e #_ #X #H
elim (lift_inv_flat2 … H) -H #V #X0 #_ #H0 #H destruct
elim (lift_inv_bind2 … H0) -H0 #V0 #T #_ #_ #H0 destruct /2 width=1 by crx_beta/
--- /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/substitution/lpx_sn_drop.ma".
+include "basic_2/substitution/fquq_alt.ma".
+include "basic_2/reduction/cpr_lift.ma".
+include "basic_2/reduction/lpr.ma".
+
+(* SN PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS *****************************)
+
+(* Properies on local environment slicing ***********************************)
+
+(* Basic_1: includes: wcpr0_drop *)
+lemma lpr_drop_conf: ∀G. dropable_sn (lpr G).
+/3 width=6 by lpx_sn_deliftable_dropable, cpr_inv_lift1/ qed-.
+
+(* Basic_1: includes: wcpr0_drop_back *)
+lemma drop_lpr_trans: ∀G. dedropable_sn (lpr G).
+/3 width=10 by lpx_sn_liftable_dedropable, cpr_lift/ qed-.
+
+lemma lpr_drop_trans_O1: ∀G. dropable_dx (lpr G).
+/2 width=3 by lpx_sn_dropable/ qed-.
+
+(* Properties on context-sensitive parallel reduction for terms *************)
+
+lemma fqu_cpr_trans_dx: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡ U2 →
+ ∃∃L,U1. ⦃G1, L1⦄ ⊢ ➡ L & ⦃G1, L⦄ ⊢ T1 ➡ U1 & ⦃G1, L, U1⦄ ⊐ ⦃G2, L2, U2⦄.
+#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
+/3 width=5 by fqu_lref_O, fqu_pair_sn, fqu_bind_dx, fqu_flat_dx, lpr_pair, cpr_pair_sn, cpr_atom, cpr_bind, cpr_flat, ex3_2_intro/
+#G #L #K #U #T #e #HLK #HUT #U2 #HU2
+elim (lift_total U2 0 (e+1)) #T2 #HUT2
+lapply (cpr_lift … HU2 … HLK … HUT … HUT2) -HU2 -HUT /3 width=9 by fqu_drop, ex3_2_intro/
+qed-.
+
+lemma fquq_cpr_trans_dx: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡ U2 →
+ ∃∃L,U1. ⦃G1, L1⦄ ⊢ ➡ L & ⦃G1, L⦄ ⊢ T1 ➡ U1 & ⦃G1, L, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
+#G1 #G2 #L1 #L2 #T1 #T2 #H #U2 #HTU2 elim (fquq_inv_gen … H) -H
+[ #HT12 elim (fqu_cpr_trans_dx … HT12 … HTU2) /3 width=5 by fqu_fquq, ex3_2_intro/
+| * #H1 #H2 #H3 destruct /2 width=5 by ex3_2_intro/
+]
+qed-.
+
+lemma fqu_cpr_trans_sn: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡ U2 →
+ ∃∃L,U1. ⦃G1, L1⦄ ⊢ ➡ L & ⦃G1, L1⦄ ⊢ T1 ➡ U1 & ⦃G1, L, U1⦄ ⊐ ⦃G2, L2, U2⦄.
+#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
+/3 width=5 by fqu_lref_O, fqu_pair_sn, fqu_bind_dx, fqu_flat_dx, lpr_pair, cpr_pair_sn, cpr_atom, cpr_bind, cpr_flat, ex3_2_intro/
+#G #L #K #U #T #e #HLK #HUT #U2 #HU2
+elim (lift_total U2 0 (e+1)) #T2 #HUT2
+lapply (cpr_lift … HU2 … HLK … HUT … HUT2) -HU2 -HUT /3 width=9 by fqu_drop, ex3_2_intro/
+qed-.
+
+lemma fquq_cpr_trans_sn: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡ U2 →
+ ∃∃L,U1. ⦃G1, L1⦄ ⊢ ➡ L & ⦃G1, L1⦄ ⊢ T1 ➡ U1 & ⦃G1, L, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
+#G1 #G2 #L1 #L2 #T1 #T2 #H #U2 #HTU2 elim (fquq_inv_gen … H) -H
+[ #HT12 elim (fqu_cpr_trans_sn … HT12 … HTU2) /3 width=5 by fqu_fquq, ex3_2_intro/
+| * #H1 #H2 #H3 destruct /2 width=5 by ex3_2_intro/
+]
+qed-.
+
+lemma fqu_lpr_trans: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+ ∀K2. ⦃G2, L2⦄ ⊢ ➡ K2 →
+ ∃∃K1,T. ⦃G1, L1⦄ ⊢ ➡ K1 & ⦃G1, L1⦄ ⊢ T1 ➡ T & ⦃G1, K1, T⦄ ⊐ ⦃G2, K2, T2⦄.
+#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
+/3 width=5 by fqu_lref_O, fqu_pair_sn, fqu_flat_dx, lpr_pair, ex3_2_intro/
+[ #a #I #G2 #L2 #V2 #T2 #X #H elim (lpr_inv_pair1 … H) -H
+ #K2 #W2 #HLK2 #HVW2 #H destruct
+ /3 width=5 by fqu_fquq, cpr_pair_sn, fqu_bind_dx, ex3_2_intro/
+| #G #L1 #K1 #T1 #U1 #e #HLK1 #HTU1 #K2 #HK12
+ elim (drop_lpr_trans … HLK1 … HK12) -HK12
+ /3 width=7 by fqu_drop, ex3_2_intro/
+]
+qed-.
+
+lemma fquq_lpr_trans: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
+ ∀K2. ⦃G2, L2⦄ ⊢ ➡ K2 →
+ ∃∃K1,T. ⦃G1, L1⦄ ⊢ ➡ K1 & ⦃G1, L1⦄ ⊢ T1 ➡ T & ⦃G1, K1, T⦄ ⊐⸮ ⦃G2, K2, T2⦄.
+#G1 #G2 #L1 #L2 #T1 #T2 #H #K2 #HLK2 elim (fquq_inv_gen … H) -H
+[ #HT12 elim (fqu_lpr_trans … HT12 … HLK2) /3 width=5 by fqu_fquq, ex3_2_intro/
+| * #H1 #H2 #H3 destruct /2 width=5 by ex3_2_intro/
+]
+qed-.
+++ /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/substitution/lpx_sn_ldrop.ma".
-include "basic_2/substitution/fquq_alt.ma".
-include "basic_2/reduction/cpr_lift.ma".
-include "basic_2/reduction/lpr.ma".
-
-(* SN PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS *****************************)
-
-(* Properies on local environment slicing ***********************************)
-
-(* Basic_1: includes: wcpr0_drop *)
-lemma lpr_ldrop_conf: ∀G. dropable_sn (lpr G).
-/3 width=6 by lpx_sn_deliftable_dropable, cpr_inv_lift1/ qed-.
-
-(* Basic_1: includes: wcpr0_drop_back *)
-lemma ldrop_lpr_trans: ∀G. dedropable_sn (lpr G).
-/3 width=10 by lpx_sn_liftable_dedropable, cpr_lift/ qed-.
-
-lemma lpr_ldrop_trans_O1: ∀G. dropable_dx (lpr G).
-/2 width=3 by lpx_sn_dropable/ qed-.
-
-(* Properties on context-sensitive parallel reduction for terms *************)
-
-lemma fqu_cpr_trans_dx: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
- ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡ U2 →
- ∃∃L,U1. ⦃G1, L1⦄ ⊢ ➡ L & ⦃G1, L⦄ ⊢ T1 ➡ U1 & ⦃G1, L, U1⦄ ⊐ ⦃G2, L2, U2⦄.
-#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
-/3 width=5 by fqu_lref_O, fqu_pair_sn, fqu_bind_dx, fqu_flat_dx, lpr_pair, cpr_pair_sn, cpr_atom, cpr_bind, cpr_flat, ex3_2_intro/
-#G #L #K #U #T #e #HLK #HUT #U2 #HU2
-elim (lift_total U2 0 (e+1)) #T2 #HUT2
-lapply (cpr_lift … HU2 … HLK … HUT … HUT2) -HU2 -HUT /3 width=9 by fqu_drop, ex3_2_intro/
-qed-.
-
-lemma fquq_cpr_trans_dx: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡ U2 →
- ∃∃L,U1. ⦃G1, L1⦄ ⊢ ➡ L & ⦃G1, L⦄ ⊢ T1 ➡ U1 & ⦃G1, L, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
-#G1 #G2 #L1 #L2 #T1 #T2 #H #U2 #HTU2 elim (fquq_inv_gen … H) -H
-[ #HT12 elim (fqu_cpr_trans_dx … HT12 … HTU2) /3 width=5 by fqu_fquq, ex3_2_intro/
-| * #H1 #H2 #H3 destruct /2 width=5 by ex3_2_intro/
-]
-qed-.
-
-lemma fqu_cpr_trans_sn: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
- ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡ U2 →
- ∃∃L,U1. ⦃G1, L1⦄ ⊢ ➡ L & ⦃G1, L1⦄ ⊢ T1 ➡ U1 & ⦃G1, L, U1⦄ ⊐ ⦃G2, L2, U2⦄.
-#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
-/3 width=5 by fqu_lref_O, fqu_pair_sn, fqu_bind_dx, fqu_flat_dx, lpr_pair, cpr_pair_sn, cpr_atom, cpr_bind, cpr_flat, ex3_2_intro/
-#G #L #K #U #T #e #HLK #HUT #U2 #HU2
-elim (lift_total U2 0 (e+1)) #T2 #HUT2
-lapply (cpr_lift … HU2 … HLK … HUT … HUT2) -HU2 -HUT /3 width=9 by fqu_drop, ex3_2_intro/
-qed-.
-
-lemma fquq_cpr_trans_sn: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡ U2 →
- ∃∃L,U1. ⦃G1, L1⦄ ⊢ ➡ L & ⦃G1, L1⦄ ⊢ T1 ➡ U1 & ⦃G1, L, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
-#G1 #G2 #L1 #L2 #T1 #T2 #H #U2 #HTU2 elim (fquq_inv_gen … H) -H
-[ #HT12 elim (fqu_cpr_trans_sn … HT12 … HTU2) /3 width=5 by fqu_fquq, ex3_2_intro/
-| * #H1 #H2 #H3 destruct /2 width=5 by ex3_2_intro/
-]
-qed-.
-
-lemma fqu_lpr_trans: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
- ∀K2. ⦃G2, L2⦄ ⊢ ➡ K2 →
- ∃∃K1,T. ⦃G1, L1⦄ ⊢ ➡ K1 & ⦃G1, L1⦄ ⊢ T1 ➡ T & ⦃G1, K1, T⦄ ⊐ ⦃G2, K2, T2⦄.
-#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
-/3 width=5 by fqu_lref_O, fqu_pair_sn, fqu_flat_dx, lpr_pair, ex3_2_intro/
-[ #a #I #G2 #L2 #V2 #T2 #X #H elim (lpr_inv_pair1 … H) -H
- #K2 #W2 #HLK2 #HVW2 #H destruct
- /3 width=5 by fqu_fquq, cpr_pair_sn, fqu_bind_dx, ex3_2_intro/
-| #G #L1 #K1 #T1 #U1 #e #HLK1 #HTU1 #K2 #HK12
- elim (ldrop_lpr_trans … HLK1 … HK12) -HK12
- /3 width=7 by fqu_drop, ex3_2_intro/
-]
-qed-.
-
-lemma fquq_lpr_trans: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ∀K2. ⦃G2, L2⦄ ⊢ ➡ K2 →
- ∃∃K1,T. ⦃G1, L1⦄ ⊢ ➡ K1 & ⦃G1, L1⦄ ⊢ T1 ➡ T & ⦃G1, K1, T⦄ ⊐⸮ ⦃G2, K2, T2⦄.
-#G1 #G2 #L1 #L2 #T1 #T2 #H #K2 #HLK2 elim (fquq_inv_gen … H) -H
-[ #HT12 elim (fqu_lpr_trans … HT12 … HLK2) /3 width=5 by fqu_fquq, ex3_2_intro/
-| * #H1 #H2 #H3 destruct /2 width=5 by ex3_2_intro/
-]
-qed-.
include "basic_2/substitution/lpx_sn_lpx_sn.ma".
include "basic_2/multiple/fqup.ma".
-include "basic_2/reduction/lpr_ldrop.ma".
+include "basic_2/reduction/lpr_drop.ma".
(* SN PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS *****************************)
∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
∃∃T. ⦃G, L1⦄ ⊢ #i ➡ T & ⦃G, L2⦄ ⊢ T2 ➡ T.
#G #L0 #i #IH #K0 #V0 #HLK0 #V2 #HV02 #T2 #HVT2 #L1 #HL01 #L2 #HL02
-elim (lpr_ldrop_conf … HLK0 … HL01) -HL01 #X1 #H1 #HLK1
+elim (lpr_drop_conf … HLK0 … HL01) -HL01 #X1 #H1 #HLK1
elim (lpr_inv_pair1 … H1) -H1 #K1 #V1 #HK01 #HV01 #H destruct
-elim (lpr_ldrop_conf … HLK0 … HL02) -HL02 #X2 #H2 #HLK2
+elim (lpr_drop_conf … HLK0 … HL02) -HL02 #X2 #H2 #HLK2
elim (lpr_inv_pair1 … H2) -H2 #K2 #W2 #HK02 #_ #H destruct
-lapply (ldrop_fwd_drop2 … HLK2) -W2 #HLK2
+lapply (drop_fwd_drop2 … HLK2) -W2 #HLK2
lapply (fqup_lref … G … HLK0) -HLK0 #HLK0
elim (IH … HLK0 … HV01 … HV02 … HK01 … HK02) -L0 -K0 -V0 #V #HV1 #HV2
elim (lift_total V 0 (i+1))
∃∃T. ⦃G, L1⦄ ⊢ T1 ➡ T & ⦃G, L2⦄ ⊢ T2 ➡ T.
#G #L0 #i #IH #K0 #V0 #HLK0 #V1 #HV01 #T1 #HVT1
#KX #VX #H #V2 #HV02 #T2 #HVT2 #L1 #HL01 #L2 #HL02
-lapply (ldrop_mono … H … HLK0) -H #H destruct
-elim (lpr_ldrop_conf … HLK0 … HL01) -HL01 #X1 #H1 #HLK1
+lapply (drop_mono … H … HLK0) -H #H destruct
+elim (lpr_drop_conf … HLK0 … HL01) -HL01 #X1 #H1 #HLK1
elim (lpr_inv_pair1 … H1) -H1 #K1 #W1 #HK01 #_ #H destruct
-lapply (ldrop_fwd_drop2 … HLK1) -W1 #HLK1
-elim (lpr_ldrop_conf … HLK0 … HL02) -HL02 #X2 #H2 #HLK2
+lapply (drop_fwd_drop2 … HLK1) -W1 #HLK1
+elim (lpr_drop_conf … HLK0 … HL02) -HL02 #X2 #H2 #HLK2
elim (lpr_inv_pair1 … H2) -H2 #K2 #W2 #HK02 #_ #H destruct
-lapply (ldrop_fwd_drop2 … HLK2) -W2 #HLK2
+lapply (drop_fwd_drop2 … HLK2) -W2 #HLK2
lapply (fqup_lref … G … HLK0) -HLK0 #HLK0
elim (IH … HLK0 … HV01 … HV02 … HK01 … HK02) -L0 -K0 -V0 #V #HV1 #HV2
elim (lift_total V 0 (i+1)) /3 width=12 by cpr_lift, ex2_intro/
#G #L0 #V0 #T0 #IH #V1 #HV01 #T1 #HT01
#T2 #HT02 #X2 #HXT2 #L1 #HL01 #L2 #HL02
elim (IH … HT01 … HT02 (L1.ⓓV1) … (L2.ⓓV1)) -IH -HT01 -HT02 /2 width=1 by lpr_pair/ -L0 -V0 -T0 #T #HT1 #HT2
-elim (cpr_inv_lift1 … HT2 L2 … HXT2) -T2 /3 width=3 by cpr_zeta, ldrop_drop, ex2_intro/
+elim (cpr_inv_lift1 … HT2 L2 … HXT2) -T2 /3 width=3 by cpr_zeta, drop_drop, ex2_intro/
qed-.
fact cpr_conf_lpr_zeta_zeta:
#G #L0 #V0 #T0 #IH #T1 #HT01 #X1 #HXT1
#T2 #HT02 #X2 #HXT2 #L1 #HL01 #L2 #HL02
elim (IH … HT01 … HT02 (L1.ⓓV0) … (L2.ⓓV0)) -IH -HT01 -HT02 /2 width=1 by lpr_pair/ -L0 -T0 #T #HT1 #HT2
-elim (cpr_inv_lift1 … HT1 L1 … HXT1) -T1 /2 width=2 by ldrop_drop/ #T1 #HT1 #HXT1
-elim (cpr_inv_lift1 … HT2 L2 … HXT2) -T2 /2 width=2 by ldrop_drop/ #T2 #HT2 #HXT2
+elim (cpr_inv_lift1 … HT1 L1 … HXT1) -T1 /2 width=2 by drop_drop/ #T1 #HT1 #HXT1
+elim (cpr_inv_lift1 … HT2 L2 … HXT2) -T2 /2 width=2 by drop_drop/ #T2 #HT2 #HXT2
lapply (lift_inj … HT2 … HT1) -T #H destruct /2 width=3 by ex2_intro/
qed-.
#V2 #HV02 #U2 #HVU2 #W2 #HW02 #T2 #HT02 #L1 #HL01 #L2 #HL02
elim (IH … HV01 … HV02 … HL01 … HL02) -HV01 -HV02 /2 width=1 by/ #V #HV1 #HV2
elim (lift_total V 0 1) #U #HVU
-lapply (cpr_lift … HV2 (L2.ⓓW2) … HVU2 … HVU) -HVU2 /2 width=2 by ldrop_drop/ #HU2
+lapply (cpr_lift … HV2 (L2.ⓓW2) … HVU2 … HVU) -HVU2 /2 width=2 by drop_drop/ #HU2
elim (cpr_inv_abbr1 … H) -H *
[ #W1 #T1 #HW01 #HT01 #H destruct
elim (IH … HW01 … HW02 … HL01 … HL02) /2 width=1 by/
| #T1 #HT01 #HXT1 #H destruct
elim (IH … HT01 … HT02 (L1.ⓓW2) … (L2.ⓓW2)) /2 width=1 by lpr_pair/ -L0 -V0 -W0 -T0 #T #HT1 #HT2
elim (cpr_inv_lift1 … HT1 L1 … HXT1) -HXT1
- /4 width=9 by cpr_flat, cpr_zeta, ldrop_drop, lift_flat, ex2_intro/
+ /4 width=9 by cpr_flat, cpr_zeta, drop_drop, lift_flat, ex2_intro/
]
qed-.
elim (IH … HW01 … HW02 … HL01 … HL02) /2 width=1 by/
elim (IH … HT01 … HT02 (L1.ⓓW1) … (L2.ⓓW2)) /2 width=1 by lpr_pair/ -L0 -V0 -W0 -T0
elim (lift_total V 0 1) #U #HVU
-lapply (cpr_lift … HV1 (L1.ⓓW1) … HVU1 … HVU) -HVU1 /2 width=2 by ldrop_drop/
-lapply (cpr_lift … HV2 (L2.ⓓW2) … HVU2 … HVU) -HVU2 /2 width=2 by ldrop_drop/
+lapply (cpr_lift … HV1 (L1.ⓓW1) … HVU1 … HVU) -HVU1 /2 width=2 by drop_drop/
+lapply (cpr_lift … HV2 (L2.ⓓW2) … HVU2 … HVU) -HVU2 /2 width=2 by drop_drop/
/4 width=7 by cpr_bind, cpr_flat, ex2_intro/ (**) (* full auto not tried *)
qed-.
include "basic_2/static/aaa_lift.ma".
include "basic_2/static/lsuba_aaa.ma".
-include "basic_2/reduction/lpx_ldrop.ma".
+include "basic_2/reduction/lpx_drop.ma".
(* SN EXTENDED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS ********************)
| #I #G #L1 #K1 #V1 #B #i #HLK1 #_ #IHV1 #X #H #L2 #HL12
elim (cpx_inv_lref1 … H) -H
[ #H destruct
- elim (lpx_ldrop_conf … HLK1 … HL12) -L1 #X #H #HLK2
+ elim (lpx_drop_conf … HLK1 … HL12) -L1 #X #H #HLK2
elim (lpx_inv_pair1 … H) -H
#K2 #V2 #HK12 #HV12 #H destruct /3 width=6 by aaa_lref/
| * #J #Y #Z #V2 #H #HV12 #HV2
- lapply (ldrop_mono … H … HLK1) -H #H destruct
- elim (lpx_ldrop_conf … HLK1 … HL12) -L1 #Z #H #HLK2
+ lapply (drop_mono … H … HLK1) -H #H destruct
+ elim (lpx_drop_conf … HLK1 … HL12) -L1 #Z #H #HLK2
elim (lpx_inv_pair1 … H) -H #K2 #V0 #HK12 #_ #H destruct
- /3 width=8 by aaa_lift, ldrop_fwd_drop2/
+ /3 width=8 by aaa_lift, drop_fwd_drop2/
]
| #a #G #L1 #V1 #T1 #B #A #_ #_ #IHV1 #IHT1 #X #H #L2 #HL12
elim (cpx_inv_abbr1 … H) -H *
[ #V2 #T2 #HV12 #HT12 #H destruct /4 width=2 by lpx_pair, aaa_abbr/
| #T2 #HT12 #HT2 #H destruct -IHV1
- /4 width=8 by lpx_pair, aaa_inv_lift, ldrop_drop/
+ /4 width=8 by lpx_pair, aaa_inv_lift, drop_drop/
]
| #a #G #L1 #V1 #T1 #B #A #_ #_ #IHV1 #IHT1 #X #H #L2 #HL12
elim (cpx_inv_abst1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct
elim (aaa_inv_abst … H) -H #B0 #A0 #HW1 #HU2 #H destruct
/5 width=6 by lsuba_aaa_trans, lsuba_abbr, aaa_abbr, aaa_cast/
| #b #V #V2 #W1 #W2 #U1 #U2 #HV1 #HV2 #HW12 #HU12 #H1 #H2 destruct
- lapply (aaa_lift G L2 … B … (L2.ⓓW2) … HV2) -HV2 /2 width=2 by ldrop_drop/ #HV2
+ lapply (aaa_lift G L2 … B … (L2.ⓓW2) … HV2) -HV2 /2 width=2 by drop_drop/ #HV2
lapply (IHT1 (ⓓ{b}W2.U2) … HL12) -IHT1 /2 width=1 by cpx_bind/ -L1 #H
elim (aaa_inv_abbr … H) -H /3 width=3 by aaa_abbr, aaa_appl/
]
--- /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/substitution/lpx_sn_drop.ma".
+include "basic_2/reduction/cpx_lift.ma".
+include "basic_2/reduction/lpx.ma".
+
+(* SN EXTENDED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS ********************)
+
+(* Properies on local environment slicing ***********************************)
+
+lemma lpx_drop_conf: ∀h,g,G. dropable_sn (lpx h g G).
+/3 width=6 by lpx_sn_deliftable_dropable, cpx_inv_lift1/ qed-.
+
+lemma drop_lpx_trans: ∀h,g,G. dedropable_sn (lpx h g G).
+/3 width=10 by lpx_sn_liftable_dedropable, cpx_lift/ qed-.
+
+lemma lpx_drop_trans_O1: ∀h,g,G. dropable_dx (lpx h g G).
+/2 width=3 by lpx_sn_dropable/ qed-.
+
+(* Properties on supclosure *************************************************)
+
+lemma fqu_lpx_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+ ∀K2. ⦃G2, L2⦄ ⊢ ➡[h, g] K2 →
+ ∃∃K1,T. ⦃G1, L1⦄ ⊢ ➡[h, g] K1 & ⦃G1, L1⦄ ⊢ T1 ➡[h, g] T & ⦃G1, K1, T⦄ ⊐ ⦃G2, K2, T2⦄.
+#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
+/3 width=5 by fqu_lref_O, fqu_pair_sn, fqu_flat_dx, lpx_pair, ex3_2_intro/
+[ #a #I #G2 #L2 #V2 #T2 #X #H elim (lpx_inv_pair1 … H) -H
+ #K2 #W2 #HLK2 #HVW2 #H destruct
+ /3 width=5 by fqu_fquq, cpx_pair_sn, fqu_bind_dx, ex3_2_intro/
+| #G #L1 #K1 #T1 #U1 #e #HLK1 #HTU1 #K2 #HK12
+ elim (drop_lpx_trans … HLK1 … HK12) -HK12
+ /3 width=7 by fqu_drop, ex3_2_intro/
+]
+qed-.
+
+lemma fquq_lpx_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
+ ∀K2. ⦃G2, L2⦄ ⊢ ➡[h, g] K2 →
+ ∃∃K1,T. ⦃G1, L1⦄ ⊢ ➡[h, g] K1 & ⦃G1, L1⦄ ⊢ T1 ➡[h, g] T & ⦃G1, K1, T⦄ ⊐⸮ ⦃G2, K2, T2⦄.
+#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H #K2 #HLK2 elim (fquq_inv_gen … H) -H
+[ #HT12 elim (fqu_lpx_trans … HT12 … HLK2) /3 width=5 by fqu_fquq, ex3_2_intro/
+| * #H1 #H2 #H3 destruct /2 width=5 by ex3_2_intro/
+]
+qed-.
+
+lemma lpx_fqu_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+ ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, g] L1 →
+ ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡[h, g] T & ⦃G1, K1, T⦄ ⊐ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, g] L2.
+#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
+/3 width=7 by fqu_pair_sn, fqu_bind_dx, fqu_flat_dx, lpx_pair, ex3_2_intro/
+[ #I #G1 #L1 #V1 #X #H elim (lpx_inv_pair2 … H) -H
+ #K1 #W1 #HKL1 #HWV1 #H destruct elim (lift_total V1 0 1)
+ /4 width=7 by cpx_delta, fqu_drop, drop_drop, ex3_2_intro/
+| #G #L1 #K1 #T1 #U1 #e #HLK1 #HTU1 #L0 #HL01
+ elim (lpx_drop_trans_O1 … HL01 … HLK1) -L1
+ /3 width=5 by fqu_drop, ex3_2_intro/
+]
+qed-.
+
+lemma lpx_fquq_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
+ ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, g] L1 →
+ ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡[h, g] T & ⦃G1, K1, T⦄ ⊐⸮ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, g] L2.
+#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H #K1 #HKL1 elim (fquq_inv_gen … H) -H
+[ #HT12 elim (lpx_fqu_trans … HT12 … HKL1) /3 width=5 by fqu_fquq, ex3_2_intro/
+| * #H1 #H2 #H3 destruct /2 width=5 by ex3_2_intro/
+]
+qed-.
include "basic_2/multiple/frees_leq.ma".
include "basic_2/multiple/frees_lift.ma".
-include "basic_2/reduction/lpx_ldrop.ma".
+include "basic_2/reduction/lpx_drop.ma".
(* SN EXTENDED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS ********************)
| #j #HG #HL #HU #U2 #H1 #L2 #HL12 #i #H2 elim (cpx_inv_lref1 … H1) -H1
[ #H destruct elim (frees_inv_lref … H2) -H2 //
* #I #K2 #W2 #Hj #Hji #HLK2 #HW2
- elim (lpx_ldrop_trans_O1 … HL12 … HLK2) -HL12 #Y #HLK1 #H
+ elim (lpx_drop_trans_O1 … HL12 … HLK2) -HL12 #Y #HLK1 #H
elim (lpx_inv_pair2 … H) -H #K1 #W1 #HK12 #HW12 #H destruct
/4 width=11 by frees_lref_be, fqup_lref/
| * #I #K1 #W1 #W0 #HLK1 #HW10 #HW0U2
- lapply (ldrop_fwd_drop2 … HLK1) #H0
- elim (lpx_ldrop_conf … H0 … HL12) -H0 -HL12 #K2 #HK12 #HLK2
+ lapply (drop_fwd_drop2 … HLK1) #H0
+ elim (lpx_drop_conf … H0 … HL12) -H0 -HL12 #K2 #HK12 #HLK2
elim (lt_or_ge i (j+1)) #Hji
[ -IH elim (frees_inv_lift_be … H2 … HLK2 … HW0U2) /2 width=1 by monotonic_pred/
| lapply (frees_inv_lift_ge … H2 … HLK2 … HW0U2 ?) -L2 -U2 // <minus_plus destruct
/4 width=7 by frees_bind_dx_O, frees_bind_sn, lpx_pair/
| #U2 #HU12 #HXU2 #H1 #H2 destruct
lapply (frees_lift_ge … Hi (L2.ⓓW1) (Ⓕ) … HXU2 ?)
- /4 width=7 by frees_bind_dx_O, lpx_pair, ldrop_drop/
+ /4 width=7 by frees_bind_dx_O, lpx_pair, drop_drop/
]
| #I #W1 #X1 #HG #HL #HU #X2 #HX2 #L2 #HL12 #i #Hi destruct
elim (cpx_inv_flat1 … HX2) -HX2 *
[ /4 width=7 by frees_flat_dx, frees_bind_sn/
| elim (frees_inv_flat … Hi) -Hi
[ #HW0 lapply (frees_inv_lift_ge … HW0 L2 (Ⓕ) … HW20 ?) -W0
- /3 width=7 by frees_flat_sn, ldrop_drop/
+ /3 width=7 by frees_flat_sn, drop_drop/
| /5 width=7 by frees_bind_dx_O, frees_flat_dx, lpx_pair/
]
]
+++ /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/substitution/lpx_sn_ldrop.ma".
-include "basic_2/reduction/cpx_lift.ma".
-include "basic_2/reduction/lpx.ma".
-
-(* SN EXTENDED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS ********************)
-
-(* Properies on local environment slicing ***********************************)
-
-lemma lpx_ldrop_conf: ∀h,g,G. dropable_sn (lpx h g G).
-/3 width=6 by lpx_sn_deliftable_dropable, cpx_inv_lift1/ qed-.
-
-lemma ldrop_lpx_trans: ∀h,g,G. dedropable_sn (lpx h g G).
-/3 width=10 by lpx_sn_liftable_dedropable, cpx_lift/ qed-.
-
-lemma lpx_ldrop_trans_O1: ∀h,g,G. dropable_dx (lpx h g G).
-/2 width=3 by lpx_sn_dropable/ qed-.
-
-(* Properties on supclosure *************************************************)
-
-lemma fqu_lpx_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
- ∀K2. ⦃G2, L2⦄ ⊢ ➡[h, g] K2 →
- ∃∃K1,T. ⦃G1, L1⦄ ⊢ ➡[h, g] K1 & ⦃G1, L1⦄ ⊢ T1 ➡[h, g] T & ⦃G1, K1, T⦄ ⊐ ⦃G2, K2, T2⦄.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
-/3 width=5 by fqu_lref_O, fqu_pair_sn, fqu_flat_dx, lpx_pair, ex3_2_intro/
-[ #a #I #G2 #L2 #V2 #T2 #X #H elim (lpx_inv_pair1 … H) -H
- #K2 #W2 #HLK2 #HVW2 #H destruct
- /3 width=5 by fqu_fquq, cpx_pair_sn, fqu_bind_dx, ex3_2_intro/
-| #G #L1 #K1 #T1 #U1 #e #HLK1 #HTU1 #K2 #HK12
- elim (ldrop_lpx_trans … HLK1 … HK12) -HK12
- /3 width=7 by fqu_drop, ex3_2_intro/
-]
-qed-.
-
-lemma fquq_lpx_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ∀K2. ⦃G2, L2⦄ ⊢ ➡[h, g] K2 →
- ∃∃K1,T. ⦃G1, L1⦄ ⊢ ➡[h, g] K1 & ⦃G1, L1⦄ ⊢ T1 ➡[h, g] T & ⦃G1, K1, T⦄ ⊐⸮ ⦃G2, K2, T2⦄.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H #K2 #HLK2 elim (fquq_inv_gen … H) -H
-[ #HT12 elim (fqu_lpx_trans … HT12 … HLK2) /3 width=5 by fqu_fquq, ex3_2_intro/
-| * #H1 #H2 #H3 destruct /2 width=5 by ex3_2_intro/
-]
-qed-.
-
-lemma lpx_fqu_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
- ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, g] L1 →
- ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡[h, g] T & ⦃G1, K1, T⦄ ⊐ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, g] L2.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
-/3 width=7 by fqu_pair_sn, fqu_bind_dx, fqu_flat_dx, lpx_pair, ex3_2_intro/
-[ #I #G1 #L1 #V1 #X #H elim (lpx_inv_pair2 … H) -H
- #K1 #W1 #HKL1 #HWV1 #H destruct elim (lift_total V1 0 1)
- /4 width=7 by cpx_delta, fqu_drop, ldrop_drop, ex3_2_intro/
-| #G #L1 #K1 #T1 #U1 #e #HLK1 #HTU1 #L0 #HL01
- elim (lpx_ldrop_trans_O1 … HL01 … HLK1) -L1
- /3 width=5 by fqu_drop, ex3_2_intro/
-]
-qed-.
-
-lemma lpx_fquq_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, g] L1 →
- ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡[h, g] T & ⦃G1, K1, T⦄ ⊐⸮ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, g] L2.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H #K1 #HKL1 elim (fquq_inv_gen … H) -H
-[ #HT12 elim (lpx_fqu_trans … HT12 … HKL1) /3 width=5 by fqu_fquq, ex3_2_intro/
-| * #H1 #H2 #H3 destruct /2 width=5 by ex3_2_intro/
-]
-qed-.
(* *)
(**************************************************************************)
-include "basic_2/multiple/llor_ldrop.ma".
+include "basic_2/multiple/llor_drop.ma".
include "basic_2/multiple/llpx_sn_llor.ma".
include "basic_2/multiple/llpx_sn_lpx_sn.ma".
include "basic_2/multiple/lleq_leq.ma".
[ #I #G1 #L1 #V1 #X #H1 #H2 elim (lpx_inv_pair2 … H1) -H1
#K0 #V0 #H1KL1 #_ #H destruct
elim (lleq_inv_lref_ge_dx … H2 ? I L1 V1) -H2 //
- #K1 #H #H2KL1 lapply (ldrop_inv_O2 … H) -H #H destruct
+ #K1 #H #H2KL1 lapply (drop_inv_O2 … H) -H #H destruct
/2 width=4 by fqu_lref_O, ex3_intro/
| * [ #a ] #I #G1 #L1 #V1 #T1 #K1 #HLK1 #H
[ elim (lleq_inv_bind … H)
| #I #G1 #L1 #V1 #T1 #K1 #HLK1 #H elim (lleq_inv_flat … H) -H
/2 width=4 by fqu_flat_dx, ex3_intro/
| #G1 #L1 #L #T1 #U1 #e #HL1 #HTU1 #K1 #H1KL1 #H2KL1
- elim (ldrop_O1_le (Ⓕ) (e+1) K1)
+ elim (drop_O1_le (Ⓕ) (e+1) K1)
[ #K #HK1 lapply (lleq_inv_lift_le … H2KL1 … HK1 HL1 … HTU1 ?) -H2KL1 //
- #H2KL elim (lpx_ldrop_trans_O1 … H1KL1 … HL1) -L1
- #K0 #HK10 #H1KL lapply (ldrop_mono … HK10 … HK1) -HK10 #H destruct
+ #H2KL elim (lpx_drop_trans_O1 … H1KL1 … HL1) -L1
+ #K0 #HK10 #H1KL lapply (drop_mono … HK10 … HK1) -HK10 #H destruct
/3 width=4 by fqu_drop, ex3_intro/
- | lapply (ldrop_fwd_length_le2 … HL1) -L -T1 -g
+ | lapply (drop_fwd_length_le2 … HL1) -L -T1 -g
lapply (lleq_fwd_length … H2KL1) //
]
]
]
qed-.
-fact leq_lpx_trans_lleq_aux: â\88\80h,g,G,L1,L0,d,e. L1 â\89\83[d, e] L0 → e = ∞ →
+fact leq_lpx_trans_lleq_aux: â\88\80h,g,G,L1,L0,d,e. L1 ⩬[d, e] L0 → e = ∞ →
∀L2. ⦃G, L0⦄ ⊢ ➡[h, g] L2 →
- â\88\83â\88\83L. L â\89\83[d, e] L2 & ⦃G, L1⦄ ⊢ ➡[h, g] L &
+ â\88\83â\88\83L. L ⩬[d, e] L2 & ⦃G, L1⦄ ⊢ ➡[h, g] L &
(∀T. L0 ≡[T, d] L2 ↔ L1 ≡[T, d] L).
#h #g #G #L1 #L0 #d #e #H elim H -L1 -L0 -d -e
[ #d #e #_ #L2 #H >(lpx_inv_atom1 … H) -H
]
qed-.
-lemma leq_lpx_trans_lleq: â\88\80h,g,G,L1,L0,d. L1 â\89\83[d, ∞] L0 →
+lemma leq_lpx_trans_lleq: â\88\80h,g,G,L1,L0,d. L1 ⩬[d, ∞] L0 →
∀L2. ⦃G, L0⦄ ⊢ ➡[h, g] L2 →
- â\88\83â\88\83L. L â\89\83[d, ∞] L2 & ⦃G, L1⦄ ⊢ ➡[h, g] L &
+ â\88\83â\88\83L. L ⩬[d, ∞] L2 & ⦃G, L1⦄ ⊢ ➡[h, g] L &
(∀T. L0 ≡[T, d] L2 ↔ L1 ≡[T, d] L).
/2 width=1 by leq_lpx_trans_lleq_aux/ qed-.
include "basic_2/notation/relations/atomicarity_4.ma".
include "basic_2/grammar/aarity.ma".
include "basic_2/grammar/genv.ma".
-include "basic_2/substitution/ldrop.ma".
+include "basic_2/substitution/drop.ma".
(* ATONIC ARITY ASSIGNMENT ON TERMS *****************************************)
(* *)
(**************************************************************************)
-include "basic_2/substitution/ldrop_ldrop.ma".
+include "basic_2/substitution/drop_drop.ma".
include "basic_2/static/aaa.ma".
(* ATONIC ARITY ASSIGNMENT ON TERMS *****************************************)
>(aaa_inv_sort … H) -H //
| #I1 #G #L #K1 #V1 #B #i #HLK1 #_ #IHA1 #A2 #H
elim (aaa_inv_lref … H) -H #I2 #K2 #V2 #HLK2 #HA2
- lapply (ldrop_mono … HLK1 … HLK2) -L #H destruct /2 width=1/
+ lapply (drop_mono … HLK1 … HLK2) -L #H destruct /2 width=1/
| #a #G #L #V #T #B1 #A1 #_ #_ #_ #IHA1 #A2 #H
elim (aaa_inv_abbr … H) -H /2 width=1/
| #a #G #L #V1 #T1 #B1 #A1 #_ #_ #IHB1 #IHA1 #X #H
∀A1. ⦃G1, L1⦄ ⊢ T1 ⁝ A1 → ∃A2. ⦃G2, L2⦄ ⊢ T2 ⁝ A2.
#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
[ #I #G #L #T #A #H elim (aaa_inv_lref … H) -H
- #J #K #V #H #HA lapply (ldrop_inv_O2 … H) -H
+ #J #K #V #H #HA lapply (drop_inv_O2 … H) -H
#H destruct /2 width=2 by ex_intro/
| * [ #a ] * #G #L #V #T #X #H
[ elim (aaa_inv_abbr … H)
(* *)
(**************************************************************************)
-include "basic_2/substitution/ldrop_ldrop.ma".
+include "basic_2/substitution/drop_drop.ma".
include "basic_2/static/aaa.ma".
(* ATONIC ARITY ASSIGNMENT ON TERMS *****************************************)
>(lift_inv_sort1 … H) -H //
| #I #G #L1 #K1 #V1 #B #i #HLK1 #_ #IHB #L2 #s #d #e #HL21 #T2 #H
elim (lift_inv_lref1 … H) -H * #Hid #H destruct
- [ elim (ldrop_trans_le … HL21 … HLK1) -L1 /2 width=2 by lt_to_le/ #X #HLK2 #H
- elim (ldrop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid #K2 #V2 #HK21 #HV12 #H destruct
+ [ elim (drop_trans_le … HL21 … HLK1) -L1 /2 width=2 by lt_to_le/ #X #HLK2 #H
+ elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid #K2 #V2 #HK21 #HV12 #H destruct
/3 width=9 by aaa_lref/
- | lapply (ldrop_trans_ge … HL21 … HLK1 ?) -L1
- /3 width=9 by aaa_lref, ldrop_inv_gen/
+ | lapply (drop_trans_ge … HL21 … HLK1 ?) -L1
+ /3 width=9 by aaa_lref, drop_inv_gen/
]
| #a #G #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #s #d #e #HL21 #X #H
elim (lift_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct
- /4 width=5 by aaa_abbr, ldrop_skip/
+ /4 width=5 by aaa_abbr, drop_skip/
| #a #G #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #s #d #e #HL21 #X #H
elim (lift_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct
- /4 width=5 by aaa_abst, ldrop_skip/
+ /4 width=5 by aaa_abst, drop_skip/
| #G #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #s #d #e #HL21 #X #H
elim (lift_inv_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct
/3 width=5 by aaa_appl/
>(lift_inv_sort2 … H) -H //
| #I #G #L2 #K2 #V2 #B #i #HLK2 #_ #IHB #L1 #s #d #e #HL21 #T1 #H
elim (lift_inv_lref2 … H) -H * #Hid #H destruct
- [ elim (ldrop_conf_lt … HL21 … HLK2) -L2 /3 width=9 by aaa_lref/
- | lapply (ldrop_conf_ge … HL21 … HLK2 ?) -L2 /3 width=9 by aaa_lref/
+ [ elim (drop_conf_lt … HL21 … HLK2) -L2 /3 width=9 by aaa_lref/
+ | lapply (drop_conf_ge … HL21 … HLK2 ?) -L2 /3 width=9 by aaa_lref/
]
| #a #G #L2 #V2 #T2 #B #A #_ #_ #IHB #IHA #L1 #s #d #e #HL21 #X #H
elim (lift_inv_bind2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct
- /4 width=5 by aaa_abbr, ldrop_skip/
+ /4 width=5 by aaa_abbr, drop_skip/
| #a #G #L2 #V2 #T2 #B #A #_ #_ #IHB #IHA #L1 #s #d #e #HL21 #X #H
elim (lift_inv_bind2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct
- /4 width=5 by aaa_abst, ldrop_skip/
+ /4 width=5 by aaa_abst, drop_skip/
| #G #L2 #V2 #T2 #B #A #_ #_ #IHB #IHA #L1 #s #d #e #HL21 #X #H
elim (lift_inv_flat2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct
/3 width=5 by aaa_appl/
(* *)
(**************************************************************************)
-include "basic_2/multiple/ldrops.ma".
+include "basic_2/multiple/drops.ma".
include "basic_2/static/aaa_lift.ma".
(* ATONIC ARITY ASSIGNMENT ON TERMS *****************************************)
(* *)
(**************************************************************************)
-include "basic_2/multiple/lleq_ldrop.ma".
+include "basic_2/multiple/lleq_drop.ma".
include "basic_2/static/aaa.ma".
(* ATONIC ARITY ASSIGNMENT ON TERMS *****************************************)
include "basic_2/notation/relations/degree_6.ma".
include "basic_2/grammar/genv.ma".
-include "basic_2/substitution/ldrop.ma".
+include "basic_2/substitution/drop.ma".
include "basic_2/static/sd.ma".
(* DEGREE ASSIGNMENT FOR TERMS **********************************************)
>(deg_mono … Hkl2 … Hkl1) -h -k -l2 //
| #G #L #K #V #i #l1 #HLK #_ #IHV #l2 #H
elim (da_inv_lref … H) -H * #K0 #V0 [| #l0 ] #HLK0 #HV0 [| #Hl0 ]
- lapply (ldrop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct /2 width=1/
+ lapply (drop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct /2 width=1/
| #G #L #K #W #i #l1 #HLK #_ #IHW #l2 #H
elim (da_inv_lref … H) -H * #K0 #W0 [| #l0 ] #HLK0 #HW0 [| #Hl0 ]
- lapply (ldrop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct /3 width=1/
+ lapply (drop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct /3 width=1/
| #a #I #G #L #V #T #l1 #_ #IHT #l2 #H
lapply (da_inv_bind … H) -H /2 width=1/
| #I #G #L #V #T #l1 #_ #IHT #l2 #H
(* *)
(**************************************************************************)
-include "basic_2/substitution/ldrop_ldrop.ma".
+include "basic_2/substitution/drop_drop.ma".
include "basic_2/static/da.ma".
(* DEGREE ASSIGNMENT FOR TERMS **********************************************)
>(lift_inv_sort1 … H) -X /2 width=1 by da_sort/
| #G #L1 #K1 #V1 #i #l #HLK1 #_ #IHV1 #L2 #s #d #e #HL21 #X #H
elim (lift_inv_lref1 … H) * #Hid #H destruct
- [ elim (ldrop_trans_le … HL21 … HLK1) -L1 /2 width=2 by lt_to_le/ #X #HLK2 #H
- elim (ldrop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid #K2 #V2 #HK21 #HV12 #H destruct
+ [ elim (drop_trans_le … HL21 … HLK1) -L1 /2 width=2 by lt_to_le/ #X #HLK2 #H
+ elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid #K2 #V2 #HK21 #HV12 #H destruct
/3 width=9 by da_ldef/
- | lapply (ldrop_trans_ge … HL21 … HLK1 ?) -L1
- /3 width=8 by da_ldef, ldrop_inv_gen/
+ | lapply (drop_trans_ge … HL21 … HLK1 ?) -L1
+ /3 width=8 by da_ldef, drop_inv_gen/
]
| #G #L1 #K1 #W1 #i #l #HLK1 #_ #IHW1 #L2 #s #d #e #HL21 #X #H
elim (lift_inv_lref1 … H) * #Hid #H destruct
- [ elim (ldrop_trans_le … HL21 … HLK1) -L1 /2 width=2 by lt_to_le/ #X #HLK2 #H
- elim (ldrop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid #K2 #W2 #HK21 #HW12 #H destruct
+ [ elim (drop_trans_le … HL21 … HLK1) -L1 /2 width=2 by lt_to_le/ #X #HLK2 #H
+ elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid #K2 #W2 #HK21 #HW12 #H destruct
/3 width=8 by da_ldec/
- | lapply (ldrop_trans_ge … HL21 … HLK1 ?) -L1
- /3 width=8 by da_ldec, ldrop_inv_gen/
+ | lapply (drop_trans_ge … HL21 … HLK1 ?) -L1
+ /3 width=8 by da_ldec, drop_inv_gen/
]
| #a #I #G #L1 #V1 #T1 #l #_ #IHT1 #L2 #s #d #e #HL21 #X #H
elim (lift_inv_bind1 … H) -H #V2 #T2 #HV12 #HU12 #H destruct
- /4 width=5 by da_bind, ldrop_skip/
+ /4 width=5 by da_bind, drop_skip/
| #I #G #L1 #V1 #T1 #l #_ #IHT1 #L2 #s #d #e #HL21 #X #H
elim (lift_inv_flat1 … H) -H #V2 #T2 #HV12 #HU12 #H destruct
/3 width=5 by da_flat/
>(lift_inv_sort2 … H) -X /2 width=1 by da_sort/
| #G #L2 #K2 #V2 #i #l #HLK2 #HV2 #IHV2 #L1 #s #d #e #HL21 #X #H
elim (lift_inv_lref2 … H) * #Hid #H destruct [ -HV2 | -IHV2 ]
- [ elim (ldrop_conf_lt … HL21 … HLK2) -L2 /3 width=8 by da_ldef/
- | lapply (ldrop_conf_ge … HL21 … HLK2 ?) -L2 /2 width=4 by da_ldef/
+ [ elim (drop_conf_lt … HL21 … HLK2) -L2 /3 width=8 by da_ldef/
+ | lapply (drop_conf_ge … HL21 … HLK2 ?) -L2 /2 width=4 by da_ldef/
]
| #G #L2 #K2 #W2 #i #l #HLK2 #HW2 #IHW2 #L1 #s #d #e #HL21 #X #H
elim (lift_inv_lref2 … H) * #Hid #H destruct [ -HW2 | -IHW2 ]
- [ elim (ldrop_conf_lt … HL21 … HLK2) -L2 /3 width=8 by da_ldec/
- | lapply (ldrop_conf_ge … HL21 … HLK2 ?) -L2 /2 width=4 by da_ldec/
+ [ elim (drop_conf_lt … HL21 … HLK2) -L2 /3 width=8 by da_ldec/
+ | lapply (drop_conf_ge … HL21 … HLK2 ?) -L2 /2 width=4 by da_ldec/
]
| #a #I #G #L2 #V2 #T2 #l #_ #IHT2 #L1 #s #d #e #HL21 #X #H
elim (lift_inv_bind2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct
- /4 width=5 by da_bind, ldrop_skip/
+ /4 width=5 by da_bind, drop_skip/
| #I #G #L2 #V2 #T2 #l #_ #IHT2 #L1 #s #d #e #HL21 #X #H
elim (lift_inv_flat2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct
/3 width=5 by da_flat/
lapply (da_inv_sort … H) -H /3 width=1 by da_sort, deg_next/
| #G #L #K #V #U #W #i #HLK #_ #HWU #IHVW #l #H
elim (da_inv_lref … H) -H * #K0 #V0 [| #l0] #HLK0 #HV0
- lapply (ldrop_mono … HLK0 … HLK) -HLK0 #H destruct
- lapply (ldrop_fwd_drop2 … HLK) -HLK /3 width=8 by da_lift/
+ lapply (drop_mono … HLK0 … HLK) -HLK0 #H destruct
+ lapply (drop_fwd_drop2 … HLK) -HLK /3 width=8 by da_lift/
| #G #L #K #W #V #U #i #HLK #_ #HWU #IHWV #l #H
elim (da_inv_lref … H) -H * #K0 #V0 [| #l0] #HLK0 #HV0 [| #H0 ]
- lapply (ldrop_mono … HLK0 … HLK) -HLK0 #H destruct
- lapply (ldrop_fwd_drop2 … HLK) -HLK /3 width=8 by da_lift/
+ lapply (drop_mono … HLK0 … HLK) -HLK0 #H destruct
+ lapply (drop_fwd_drop2 … HLK) -HLK /3 width=8 by da_lift/
| #a #I #G #L #V #T #U #_ #IHTU #l #H
lapply (da_inv_bind … H) -H /3 width=1 by da_bind/
| #G #L #V #T #U #_ #IHTU #l #H
qed.
(* Note: the constant 0 cannot be generalized *)
-lemma lsuba_ldrop_O1_conf: ∀G,L1,L2. G ⊢ L1 ⁝⫃ L2 → ∀K1,s,e. ⇩[s, 0, e] L1 ≡ K1 →
- ∃∃K2. G ⊢ K1 ⁝⫃ K2 & ⇩[s, 0, e] L2 ≡ K2.
+lemma lsuba_drop_O1_conf: ∀G,L1,L2. G ⊢ L1 ⁝⫃ L2 → ∀K1,s,e. ⇩[s, 0, e] L1 ≡ K1 →
+ ∃∃K2. G ⊢ K1 ⁝⫃ K2 & ⇩[s, 0, e] L2 ≡ K2.
#G #L1 #L2 #H elim H -L1 -L2
[ /2 width=3 by ex2_intro/
| #I #L1 #L2 #V #_ #IHL12 #K1 #s #e #H
- elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK1
+ elim (drop_inv_O1_pair1 … H) -H * #He #HLK1
[ destruct
elim (IHL12 L1 s 0) -IHL12 // #X #HL12 #H
- <(ldrop_inv_O2 … H) in HL12; -H /3 width=3 by lsuba_pair, ldrop_pair, ex2_intro/
- | elim (IHL12 … HLK1) -L1 /3 width=3 by ldrop_drop_lt, ex2_intro/
+ <(drop_inv_O2 … H) in HL12; -H /3 width=3 by lsuba_pair, drop_pair, ex2_intro/
+ | elim (IHL12 … HLK1) -L1 /3 width=3 by drop_drop_lt, ex2_intro/
]
| #L1 #L2 #W #V #A #HV #HW #_ #IHL12 #K1 #s #e #H
- elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK1
+ elim (drop_inv_O1_pair1 … H) -H * #He #HLK1
[ destruct
elim (IHL12 L1 s 0) -IHL12 // #X #HL12 #H
- <(ldrop_inv_O2 … H) in HL12; -H /3 width=3 by lsuba_abbr, ldrop_pair, ex2_intro/
- | elim (IHL12 … HLK1) -L1 /3 width=3 by ldrop_drop_lt, ex2_intro/
+ <(drop_inv_O2 … H) in HL12; -H /3 width=3 by lsuba_abbr, drop_pair, ex2_intro/
+ | elim (IHL12 … HLK1) -L1 /3 width=3 by drop_drop_lt, ex2_intro/
]
]
qed-.
(* Note: the constant 0 cannot be generalized *)
-lemma lsuba_ldrop_O1_trans: ∀G,L1,L2. G ⊢ L1 ⁝⫃ L2 → ∀K2,s,e. ⇩[s, 0, e] L2 ≡ K2 →
- ∃∃K1. G ⊢ K1 ⁝⫃ K2 & ⇩[s, 0, e] L1 ≡ K1.
+lemma lsuba_drop_O1_trans: ∀G,L1,L2. G ⊢ L1 ⁝⫃ L2 → ∀K2,s,e. ⇩[s, 0, e] L2 ≡ K2 →
+ ∃∃K1. G ⊢ K1 ⁝⫃ K2 & ⇩[s, 0, e] L1 ≡ K1.
#G #L1 #L2 #H elim H -L1 -L2
[ /2 width=3 by ex2_intro/
| #I #L1 #L2 #V #_ #IHL12 #K2 #s #e #H
- elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK2
+ elim (drop_inv_O1_pair1 … H) -H * #He #HLK2
[ destruct
elim (IHL12 L2 s 0) -IHL12 // #X #HL12 #H
- <(ldrop_inv_O2 … H) in HL12; -H /3 width=3 by lsuba_pair, ldrop_pair, ex2_intro/
- | elim (IHL12 … HLK2) -L2 /3 width=3 by ldrop_drop_lt, ex2_intro/
+ <(drop_inv_O2 … H) in HL12; -H /3 width=3 by lsuba_pair, drop_pair, ex2_intro/
+ | elim (IHL12 … HLK2) -L2 /3 width=3 by drop_drop_lt, ex2_intro/
]
| #L1 #L2 #W #V #A #HV #HW #_ #IHL12 #K2 #s #e #H
- elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK2
+ elim (drop_inv_O1_pair1 … H) -H * #He #HLK2
[ destruct
elim (IHL12 L2 s 0) -IHL12 // #X #HL12 #H
- <(ldrop_inv_O2 … H) in HL12; -H /3 width=3 by lsuba_abbr, ldrop_pair, ex2_intro/
- | elim (IHL12 … HLK2) -L2 /3 width=3 by ldrop_drop_lt, ex2_intro/
+ <(drop_inv_O2 … H) in HL12; -H /3 width=3 by lsuba_abbr, drop_pair, ex2_intro/
+ | elim (IHL12 … HLK2) -L2 /3 width=3 by drop_drop_lt, ex2_intro/
]
]
qed-.
#G #L1 #V #A #H elim H -G -L1 -V -A
[ //
| #I #G #L1 #K1 #V #A #i #HLK1 #HV #IHV #L2 #HL12
- elim (lsuba_ldrop_O1_conf … HL12 … HLK1) -L1 #X #H #HLK2
+ elim (lsuba_drop_O1_conf … HL12 … HLK1) -L1 #X #H #HLK2
elim (lsuba_inv_pair1 … H) -H * #K2
[ #HK12 #H destruct /3 width=5 by aaa_lref/
| #W0 #V0 #A0 #HV0 #HW0 #_ #H1 #H2 #H3 destruct
#G #L2 #V #A #H elim H -G -L2 -V -A
[ //
| #I #G #L2 #K2 #V #A #i #HLK2 #H1V #IHV #L1 #HL12
- elim (lsuba_ldrop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1
+ elim (lsuba_drop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1
elim (lsuba_inv_pair2 … H) -H * #K1
[ #HK12 #H destruct /3 width=5 by aaa_lref/
| #V0 #A0 #HV0 #H2V #_ #H1 #H2 destruct
qed.
(* Note: the constant 0 cannot be generalized *)
-lemma lsubd_ldrop_O1_conf: ∀h,g,G,L1,L2. G ⊢ L1 ▪⫃[h, g] L2 →
- ∀K1,s,e. ⇩[s, 0, e] L1 ≡ K1 →
- ∃∃K2. G ⊢ K1 ▪⫃[h, g] K2 & ⇩[s, 0, e] L2 ≡ K2.
+lemma lsubd_drop_O1_conf: ∀h,g,G,L1,L2. G ⊢ L1 ▪⫃[h, g] L2 →
+ ∀K1,s,e. ⇩[s, 0, e] L1 ≡ K1 →
+ ∃∃K2. G ⊢ K1 ▪⫃[h, g] K2 & ⇩[s, 0, e] L2 ≡ K2.
#h #g #G #L1 #L2 #H elim H -L1 -L2
[ /2 width=3 by ex2_intro/
| #I #L1 #L2 #V #_ #IHL12 #K1 #s #e #H
- elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK1
+ elim (drop_inv_O1_pair1 … H) -H * #He #HLK1
[ destruct
elim (IHL12 L1 s 0) -IHL12 // #X #HL12 #H
- <(ldrop_inv_O2 … H) in HL12; -H /3 width=3 by lsubd_pair, ldrop_pair, ex2_intro/
- | elim (IHL12 … HLK1) -L1 /3 width=3 by ldrop_drop_lt, ex2_intro/
+ <(drop_inv_O2 … H) in HL12; -H /3 width=3 by lsubd_pair, drop_pair, ex2_intro/
+ | elim (IHL12 … HLK1) -L1 /3 width=3 by drop_drop_lt, ex2_intro/
]
| #L1 #L2 #W #V #l #HV #HW #_ #IHL12 #K1 #s #e #H
- elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK1
+ elim (drop_inv_O1_pair1 … H) -H * #He #HLK1
[ destruct
elim (IHL12 L1 s 0) -IHL12 // #X #HL12 #H
- <(ldrop_inv_O2 … H) in HL12; -H /3 width=3 by lsubd_abbr, ldrop_pair, ex2_intro/
- | elim (IHL12 … HLK1) -L1 /3 width=3 by ldrop_drop_lt, ex2_intro/
+ <(drop_inv_O2 … H) in HL12; -H /3 width=3 by lsubd_abbr, drop_pair, ex2_intro/
+ | elim (IHL12 … HLK1) -L1 /3 width=3 by drop_drop_lt, ex2_intro/
]
]
qed-.
(* Note: the constant 0 cannot be generalized *)
-lemma lsubd_ldrop_O1_trans: ∀h,g,G,L1,L2. G ⊢ L1 ▪⫃[h, g] L2 →
- ∀K2,s,e. ⇩[s, 0, e] L2 ≡ K2 →
- ∃∃K1. G ⊢ K1 ▪⫃[h, g] K2 & ⇩[s, 0, e] L1 ≡ K1.
+lemma lsubd_drop_O1_trans: ∀h,g,G,L1,L2. G ⊢ L1 ▪⫃[h, g] L2 →
+ ∀K2,s,e. ⇩[s, 0, e] L2 ≡ K2 →
+ ∃∃K1. G ⊢ K1 ▪⫃[h, g] K2 & ⇩[s, 0, e] L1 ≡ K1.
#h #g #G #L1 #L2 #H elim H -L1 -L2
[ /2 width=3 by ex2_intro/
| #I #L1 #L2 #V #_ #IHL12 #K2 #s #e #H
- elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK2
+ elim (drop_inv_O1_pair1 … H) -H * #He #HLK2
[ destruct
elim (IHL12 L2 s 0) -IHL12 // #X #HL12 #H
- <(ldrop_inv_O2 … H) in HL12; -H /3 width=3 by lsubd_pair, ldrop_pair, ex2_intro/
- | elim (IHL12 … HLK2) -L2 /3 width=3 by ldrop_drop_lt, ex2_intro/
+ <(drop_inv_O2 … H) in HL12; -H /3 width=3 by lsubd_pair, drop_pair, ex2_intro/
+ | elim (IHL12 … HLK2) -L2 /3 width=3 by drop_drop_lt, ex2_intro/
]
| #L1 #L2 #W #V #l #HV #HW #_ #IHL12 #K2 #s #e #H
- elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK2
+ elim (drop_inv_O1_pair1 … H) -H * #He #HLK2
[ destruct
elim (IHL12 L2 s 0) -IHL12 // #X #HL12 #H
- <(ldrop_inv_O2 … H) in HL12; -H /3 width=3 by lsubd_abbr, ldrop_pair, ex2_intro/
- | elim (IHL12 … HLK2) -L2 /3 width=3 by ldrop_drop_lt, ex2_intro/
+ <(drop_inv_O2 … H) in HL12; -H /3 width=3 by lsubd_abbr, drop_pair, ex2_intro/
+ | elim (IHL12 … HLK2) -L2 /3 width=3 by drop_drop_lt, ex2_intro/
]
]
qed-.
#h #g #G #L2 #T #l #H elim H -G -L2 -T -l
[ /2 width=1/
| #G #L2 #K2 #V #i #l #HLK2 #_ #IHV #L1 #HL12
- elim (lsubd_ldrop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1
+ elim (lsubd_drop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1
elim (lsubd_inv_pair2 … H) -H * #K1 [ | -IHV -HLK1 ]
[ #HK12 #H destruct /3 width=4/
| #W #l0 #_ #_ #_ #H destruct
]
| #G #L2 #K2 #W #i #l #HLK2 #HW #IHW #L1 #HL12
- elim (lsubd_ldrop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1
+ elim (lsubd_drop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1
elim (lsubd_inv_pair2 … H) -H * #K1 [ -HW | -IHW ]
[ #HK12 #H destruct /3 width=4/
| #V #l0 #HV #H0W #_ #_ #H destruct
#h #g #G #L1 #T #l #H elim H -G -L1 -T -l
[ /2 width=1/
| #G #L1 #K1 #V #i #l #HLK1 #HV #IHV #L2 #HL12
- elim (lsubd_ldrop_O1_conf … HL12 … HLK1) -L1 #X #H #HLK2
+ elim (lsubd_drop_O1_conf … HL12 … HLK1) -L1 #X #H #HLK2
elim (lsubd_inv_pair1 … H) -H * #K2 [ -HV | -IHV ]
[ #HK12 #H destruct /3 width=4/
| #W0 #V0 #l0 #HV0 #HW0 #_ #_ #H1 #H2 destruct
lapply (da_mono … H0V0 … HV0) -H0V0 -HV0 #H destruct /2 width=4/
]
| #G #L1 #K1 #W #i #l #HLK1 #HW #IHW #L2 #HL12
- elim (lsubd_ldrop_O1_conf … HL12 … HLK1) -L1 #X #H #HLK2
+ elim (lsubd_drop_O1_conf … HL12 … HLK1) -L1 #X #H #HLK2
elim (lsubd_inv_pair1 … H) -H * #K2 [ -HW | -IHW ]
[ #HK12 #H destruct /3 width=4/
| #W0 #V0 #l0 #HV0 #HW0 #_ #H destruct
(**************************************************************************)
include "basic_2/notation/relations/lrsubeqc_2.ma".
-include "basic_2/substitution/ldrop.ma".
+include "basic_2/substitution/drop.ma".
(* RESTRICTED LOCAL ENVIRONMENT REFINEMENT **********************************)
#L1 #L2 #H elim H -L1 -L2 /2 width=1 by monotonic_le_plus_l/
qed-.
-lemma lsubr_fwd_ldrop2_bind: ∀L1,L2. L1 ⫃ L2 →
- ∀I,K2,W,s,i. ⇩[s, 0, i] L2 ≡ K2.ⓑ{I}W →
- (∃∃K1. K1 ⫃ K2 & ⇩[s, 0, i] L1 ≡ K1.ⓑ{I}W) ∨
- ∃∃K1,V. K1 ⫃ K2 & ⇩[s, 0, i] L1 ≡ K1.ⓓⓝW.V & I = Abst.
+lemma lsubr_fwd_drop2_bind: ∀L1,L2. L1 ⫃ L2 →
+ ∀I,K2,W,s,i. ⇩[s, 0, i] L2 ≡ K2.ⓑ{I}W →
+ (∃∃K1. K1 ⫃ K2 & ⇩[s, 0, i] L1 ≡ K1.ⓑ{I}W) ∨
+ ∃∃K1,V. K1 ⫃ K2 & ⇩[s, 0, i] L1 ≡ K1.ⓓⓝW.V & I = Abst.
#L1 #L2 #H elim H -L1 -L2
[ #L #I #K2 #W #s #i #H
- elim (ldrop_inv_atom1 … H) -H #H destruct
+ elim (drop_inv_atom1 … H) -H #H destruct
| #J #L1 #L2 #V #HL12 #IHL12 #I #K2 #W #s #i #H
- elim (ldrop_inv_O1_pair1 … H) -H * #Hi #HLK2 destruct [ -IHL12 | -HL12 ]
- [ /3 width=3 by ldrop_pair, ex2_intro, or_introl/
+ elim (drop_inv_O1_pair1 … H) -H * #Hi #HLK2 destruct [ -IHL12 | -HL12 ]
+ [ /3 width=3 by drop_pair, ex2_intro, or_introl/
| elim (IHL12 … HLK2) -IHL12 -HLK2 *
- /4 width=4 by ldrop_drop_lt, ex3_2_intro, ex2_intro, or_introl, or_intror/
+ /4 width=4 by drop_drop_lt, ex3_2_intro, ex2_intro, or_introl, or_intror/
]
| #L1 #L2 #V1 #V2 #HL12 #IHL12 #I #K2 #W #s #i #H
- elim (ldrop_inv_O1_pair1 … H) -H * #Hi #HLK2 destruct [ -IHL12 | -HL12 ]
- [ /3 width=4 by ldrop_pair, ex3_2_intro, or_intror/
+ elim (drop_inv_O1_pair1 … H) -H * #Hi #HLK2 destruct [ -IHL12 | -HL12 ]
+ [ /3 width=4 by drop_pair, ex3_2_intro, or_intror/
| elim (IHL12 … HLK2) -IHL12 -HLK2 *
- /4 width=4 by ldrop_drop_lt, ex3_2_intro, ex2_intro, or_introl, or_intror/
+ /4 width=4 by drop_drop_lt, ex3_2_intro, ex2_intro, or_introl, or_intror/
]
]
qed-.
-lemma lsubr_fwd_ldrop2_abbr: ∀L1,L2. L1 ⫃ L2 →
- ∀K2,V,s,i. ⇩[s, 0, i] L2 ≡ K2.ⓓV →
- ∃∃K1. K1 ⫃ K2 & ⇩[s, 0, i] L1 ≡ K1.ⓓV.
-#L1 #L2 #HL12 #K2 #V #s #i #HLK2 elim (lsubr_fwd_ldrop2_bind … HL12 … HLK2) -L2 // *
+lemma lsubr_fwd_drop2_abbr: ∀L1,L2. L1 ⫃ L2 →
+ ∀K2,V,s,i. ⇩[s, 0, i] L2 ≡ K2.ⓓV →
+ ∃∃K1. K1 ⫃ K2 & ⇩[s, 0, i] L1 ≡ K1.ⓓV.
+#L1 #L2 #HL12 #K2 #V #s #i #HLK2 elim (lsubr_fwd_drop2_bind … HL12 … HLK2) -L2 // *
#K1 #W #_ #_ #H destruct
qed-.
include "basic_2/notation/relations/statictype_5.ma".
include "basic_2/grammar/genv.ma".
-include "basic_2/substitution/ldrop.ma".
+include "basic_2/substitution/drop.ma".
include "basic_2/static/sh.ma".
(* STATIC TYPE ASSIGNMENT ON TERMS ******************************************)
lapply (sta_inv_sort1 … H) -H #H destruct //
| #I #G #L #K #V #B #i #HLK #HV #IHV #U #H
elim (sta_inv_lref1 … H) -H * #K0 #V0 #W0 #HLK0 #HVW0 #HU
- lapply (ldrop_mono … HLK0 … HLK) -HLK0 #H0 destruct
- lapply (ldrop_fwd_drop2 … HLK) -HLK #HLK
+ lapply (drop_mono … HLK0 … HLK) -HLK0 #H0 destruct
+ lapply (drop_fwd_drop2 … HLK) -HLK #HLK
@(aaa_lift … HLK … HU) -HU -L /2 width=2 by/
| #a #G #L #V #T #B #A #HV #_ #_ #IHT #X #H
elim (sta_inv_bind1 … H) -H #U #HTU #H destruct /3 width=2 by aaa_abbr/
(* *)
(**************************************************************************)
-include "basic_2/substitution/ldrop_ldrop.ma".
+include "basic_2/substitution/drop_drop.ma".
include "basic_2/static/sta.ma".
(* STATIC TYPE ASSIGNMENT ON TERMS ******************************************)
| #G #L1 #K1 #V1 #W1 #W #i #HLK1 #_ #HW1 #IHVW1 #L2 #s #d #e #HL21 #X #H #U2 #HWU2
elim (lift_inv_lref1 … H) * #Hid #H destruct
[ elim (lift_trans_ge … HW1 … HWU2) -W // #W2 #HW12 #HWU2
- elim (ldrop_trans_le … HL21 … HLK1) -L1 /2 width=2 by lt_to_le/ #X #HLK2 #H
- elim (ldrop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid #K2 #V2 #HK21 #HV12 #H destruct
+ elim (drop_trans_le … HL21 … HLK1) -L1 /2 width=2 by lt_to_le/ #X #HLK2 #H
+ elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid #K2 #V2 #HK21 #HV12 #H destruct
/3 width=9 by sta_ldef/
| lapply (lift_trans_be … HW1 … HWU2 ? ?) -W /2 width=1 by le_S/ #HW1U2
- lapply (ldrop_trans_ge … HL21 … HLK1 ?) -L1 /3 width=9 by sta_ldef, ldrop_inv_gen/
+ lapply (drop_trans_ge … HL21 … HLK1 ?) -L1 /3 width=9 by sta_ldef, drop_inv_gen/
]
| #G #L1 #K1 #W1 #V1 #W #i #HLK1 #_ #HW1 #IHWV1 #L2 #s #d #e #HL21 #X #H #U2 #HWU2
elim (lift_inv_lref1 … H) * #Hid #H destruct
[ elim (lift_trans_ge … HW1 … HWU2) -W // <minus_plus #W #HW1 #HWU2
- elim (ldrop_trans_le … HL21 … HLK1) -L1 /2 width=2 by lt_to_le/ #X #HLK2 #H
- elim (ldrop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid #K2 #W2 #HK21 #HW12 #H destruct
+ elim (drop_trans_le … HL21 … HLK1) -L1 /2 width=2 by lt_to_le/ #X #HLK2 #H
+ elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid #K2 #W2 #HK21 #HW12 #H destruct
lapply (lift_mono … HW1 … HW12) -HW1 #H destruct
elim (lift_total V1 (d-i-1) e) /3 width=9 by sta_ldec/
| lapply (lift_trans_be … HW1 … HWU2 ? ?) -W /2 width=1 by le_S/ #HW1U2
- lapply (ldrop_trans_ge … HL21 … HLK1 ?) -L1 /3 width=9 by sta_ldec, ldrop_inv_gen/
+ lapply (drop_trans_ge … HL21 … HLK1 ?) -L1 /3 width=9 by sta_ldec, drop_inv_gen/
]
| #a #I #G #L1 #V1 #T1 #U1 #_ #IHTU1 #L2 #s #d #e #HL21 #X1 #H1 #X2 #H2
elim (lift_inv_bind1 … H1) -H1 #V2 #T2 #HV12 #HT12 #H destruct
elim (lift_inv_bind1 … H2) -H2 #X #U2 #H1 #HU12 #H2 destruct
- lapply (lift_mono … H1 … HV12) -H1 #H destruct /4 width=6 by sta_bind, ldrop_skip/
+ lapply (lift_mono … H1 … HV12) -H1 #H destruct /4 width=6 by sta_bind, drop_skip/
| #G #L1 #V1 #T1 #U1 #_ #IHTU1 #L2 #s #d #e #HL21 #X1 #H1 #X2 #H2
elim (lift_inv_flat1 … H1) -H1 #V2 #T2 #HV12 #HT12 #H destruct
elim (lift_inv_flat1 … H2) -H2 #X #U2 #H1 #HU12 #H2 destruct
>(lift_inv_sort2 … H) -X /2 width=3 by sta_sort, lift_sort, ex2_intro/
| #G #L2 #K2 #V2 #W2 #W #i #HLK2 #HVW2 #HW2 #IHVW2 #L1 #s #d #e #HL21 #X #H
elim (lift_inv_lref2 … H) * #Hid #H destruct [ -HVW2 | -IHVW2 ]
- [ elim (ldrop_conf_lt … HL21 … HLK2) -L2 // #K1 #V1 #HLK1 #HK21 #HV12
+ [ elim (drop_conf_lt … HL21 … HLK2) -L2 // #K1 #V1 #HLK1 #HK21 #HV12
elim (IHVW2 … HK21 … HV12) -K2 -V2 #W1 #HW12 #HVW1
elim (lift_trans_le … HW12 … HW2) -W2 // >minus_plus <plus_minus_m_m /3 width=8 by sta_ldef, ex2_intro/
- | lapply (ldrop_conf_ge … HL21 … HLK2 ?) -L2 // #HL1K2
+ | lapply (drop_conf_ge … HL21 … HLK2 ?) -L2 // #HL1K2
elim (le_inv_plus_l … Hid) -Hid #Hdie #ei
elim (lift_split … HW2 d (i-e+1)) -HW2 /2 width=1 by le_S_S, le_S/
#W0 #HW20 <le_plus_minus_comm // >minus_minus_m_m /3 width=8 by sta_ldef, le_S, ex2_intro/
]
| #G #L2 #K2 #W2 #V2 #W #i #HLK2 #HWV2 #HW2 #IHWV2 #L1 #s #d #e #HL21 #X #H
elim (lift_inv_lref2 … H) * #Hid #H destruct [ -HWV2 | -IHWV2 ]
- [ elim (ldrop_conf_lt … HL21 … HLK2) -L2 // #K1 #W1 #HLK1 #HK21 #HW12
+ [ elim (drop_conf_lt … HL21 … HLK2) -L2 // #K1 #W1 #HLK1 #HK21 #HW12
elim (IHWV2 … HK21 … HW12) -K2 #V1 #_ #HWV1
elim (lift_trans_le … HW12 … HW2) -W2 // >minus_plus <plus_minus_m_m /3 width=8 by sta_ldec, ex2_intro/
- | lapply (ldrop_conf_ge … HL21 … HLK2 ?) -L2 // #HL1K2
+ | lapply (drop_conf_ge … HL21 … HLK2 ?) -L2 // #HL1K2
elim (le_inv_plus_l … Hid) -Hid #Hdie #ei
elim (lift_split … HW2 d (i-e+1)) -HW2 /2 width=1 by le_S_S, le_S/
#W0 #HW20 <le_plus_minus_comm // >minus_minus_m_m /3 width=8 by sta_ldec, le_S, ex2_intro/
]
| #a #I #G #L2 #V2 #T2 #U2 #_ #IHTU2 #L1 #s #d #e #HL21 #X #H
elim (lift_inv_bind2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct
- elim (IHTU2 (L1.ⓑ{I}V1) … HT12) -IHTU2 -HT12 /3 width=5 by sta_bind, ldrop_skip, lift_bind, ex2_intro/
+ elim (IHTU2 (L1.ⓑ{I}V1) … HT12) -IHTU2 -HT12 /3 width=5 by sta_bind, drop_skip, lift_bind, ex2_intro/
| #G #L2 #V2 #T2 #U2 #_ #IHTU2 #L1 #s #d #e #HL21 #X #H
elim (lift_inv_flat2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct
elim (IHTU2 … HL21 … HT12) -L2 -HT12 /3 width=5 by sta_appl, lift_flat, ex2_intro/
#h #G #L #T #U #H elim H -G -L -T -U
[ /2 width=2/
| #G #L #K #V #W #W0 #i #HLK #_ #HW0 * #V0 #HWV0
- lapply (ldrop_fwd_drop2 … HLK) -HLK #HLK
+ lapply (drop_fwd_drop2 … HLK) -HLK #HLK
elim (lift_total V0 0 (i+1)) /3 width=11 by ex_intro, sta_lift/
| #G #L #K #W #V #V0 #i #HLK #HWV #HWV0 #_
- lapply (ldrop_fwd_drop2 … HLK) -HLK #HLK
+ lapply (drop_fwd_drop2 … HLK) -HLK #HLK
elim (lift_total V 0 (i+1)) /3 width=11 by ex_intro, sta_lift/
| #a #I #G #L #V #T #U #_ * /3 width=2 by sta_bind, ex_intro/
| #G #L #V #T #U #_ * #T0 #HUT0 /3 width=2 by sta_appl, ex_intro/
(* *)
(**************************************************************************)
-include "basic_2/multiple/llpx_sn_ldrop.ma".
+include "basic_2/multiple/llpx_sn_drop.ma".
include "basic_2/static/sta.ma".
(* STRATIFIED STATIC TYPE ASSIGNMENT FOR TERMS ******************************)
[ /3 width=4 by llpx_sn_fwd_length, llpx_sn_sort/
| #G #Ls #Ks #V1s #W2s #V2s #i #HLKs #_ #HVW2s #IHV12s #Ld #H elim (llpx_sn_inv_lref_ge_sn … H … HLKs) // -H
#Kd #V1d #HLKd #HV1s #HV1sd
- lapply (ldrop_fwd_drop2 … HLKs) -HLKs #HLKs
- lapply (ldrop_fwd_drop2 … HLKd) -HLKd #HLKd
+ lapply (drop_fwd_drop2 … HLKs) -HLKs #HLKs
+ lapply (drop_fwd_drop2 … HLKd) -HLKd #HLKd
@(llpx_sn_lift_le … HLKs HLKd … HVW2s) -HLKs -HLKd -HVW2s /2 width=1 by/ (**) (* full auto too slow *)
| #G #Ls #Ks #V1s #W1s #V2s #i #HLKs #_ #HV12s #IHVW1s #Ld #H elim (llpx_sn_inv_lref_ge_sn … H … HLKs) // -H
#Kd #V1d #HLKd #HV1s #HV1sd
- lapply (ldrop_fwd_drop2 … HLKs) -HLKs #HLKs
- lapply (ldrop_fwd_drop2 … HLKd) -HLKd #HLKd
+ lapply (drop_fwd_drop2 … HLKs) -HLKs #HLKs
+ lapply (drop_fwd_drop2 … HLKd) -HLKd #HLKd
@(llpx_sn_lift_le … HLKs HLKd … HV12s) -HLKs -HLKd -HV12s /2 width=1 by/ (**) (* full auto too slow *)
| #a #I #G #Ls #V #T1 #T2 #_ #IHT12 #Ld #H elim (llpx_sn_inv_bind_O … H) -H
/4 width=5 by llpx_sn_bind_repl_SO, llpx_sn_bind/
(* *)
(**************************************************************************)
-include "basic_2/substitution/ldrop_ldrop.ma".
+include "basic_2/substitution/drop_drop.ma".
include "basic_2/static/sta.ma".
(* STATIC TYPE ASSIGNMENT ON TERMS ******************************************)
[ #G #L #k #X #H >(sta_inv_sort1 … H) -X //
| #G #L #K #V #W #U1 #i #HLK #_ #HWU1 #IHVW #U2 #H
elim (sta_inv_lref1 … H) -H * #K0 #V0 #W0 #HLK0 #HVW0 #HW0U2
- lapply (ldrop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct
+ lapply (drop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct
lapply (IHVW … HVW0) -IHVW -HVW0 #H destruct
>(lift_mono … HWU1 … HW0U2) -W0 -U1 //
| #G #L #K #W #V #U1 #i #HLK #_ #HWU1 #IHWV #U2 #H
elim (sta_inv_lref1 … H) -H * #K0 #W0 #V0 #HLK0 #HWV0 #HV0U2
- lapply (ldrop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct
+ lapply (drop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct
lapply (IHWV … HWV0) -IHWV -HWV0 #H destruct
>(lift_mono … HWU1 … HV0U2) -W -U1 //
| #a #I #G #L #V #T #U1 #_ #IHTU1 #X #H
#G #d #e #L1 #T1 #T2 #H elim H -G -L1 -T1 -T2 -d -e
[ //
| #I #G #L1 #K1 #V #W #i #d #e #Hdi #Hide #HLK1 #HVW #L2 #HL12
- elim (lsuby_ldrop_trans_be … HL12 … HLK1) -HL12 -HLK1 /2 width=5 by cpy_subst/
+ elim (lsuby_drop_trans_be … HL12 … HLK1) -HL12 -HLK1 /2 width=5 by cpy_subst/
| /4 width=1 by lsuby_succ, cpy_bind/
| /3 width=1 by cpy_flat/
]
| * [ #a ] #J #W1 #U1 #IHW1 #IHU1 #L #d #HLK
elim (IHW1 … HLK) -IHW1 #W2 #W #HW12 #HW2
[ elim (IHU1 (L.ⓑ{J}W1) (d+1)) -IHU1
- /3 width=9 by cpy_bind, ldrop_drop, lift_bind, ex2_2_intro/
+ /3 width=9 by cpy_bind, drop_drop, lift_bind, ex2_2_intro/
| elim (IHU1 … HLK) -IHU1 -HLK
/3 width=8 by cpy_flat, lift_flat, ex2_2_intro/
]
⦃G, L⦄ ⊢ T1 ▶[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶[d, |L| - d] T2.
#G #L #T1 #T2 #d #e #H elim H -G -L -T1 -T2 -d -e //
[ #I #G #L #K #V #W #i #d #e #Hdi #_ #HLK #HVW
- lapply (ldrop_fwd_length_lt2 … HLK)
+ lapply (drop_fwd_length_lt2 … HLK)
/4 width=5 by cpy_subst, ylt_yle_trans, ylt_inj/
| #a #I #G #L #V1 #V2 normalize in match (|L.ⓑ{I}V2|); (**) (* |?| does not work *)
/2 width=1 by cpy_bind/
elim (cpy_inv_lref1 … H) -H
[ #HX destruct /3 width=7 by cpy_subst, ex2_intro/
| -Hd1 -Hde1 * #I2 #K2 #V2 #_ #_ #HLK2 #HVT2
- lapply (ldrop_mono … HLK1 … HLK2) -HLK1 -HLK2 #H destruct
+ lapply (drop_mono … HLK1 … HLK2) -HLK1 -HLK2 #H destruct
>(lift_mono … HVT1 … HVT2) -HVT1 -HVT2 /2 width=3 by ex2_intro/
]
| #a #I #G #L #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #X #d2 #e2 #HX
(* *)
(**************************************************************************)
-include "basic_2/substitution/ldrop_ldrop.ma".
+include "basic_2/substitution/drop_drop.ma".
include "basic_2/substitution/cpy.ma".
(* CONTEXT-SENSITIVE EXTENDED ORDINARY SUBSTITUTION FOR TERMS ***************)
lapply (ylt_inv_inj … Hid) -Hid #Hid
lapply (lift_inv_lref1_lt … H … Hid) -H #H destruct
elim (lift_trans_ge … HVW … HWU2) -W // <minus_plus #W #HVW #HWU2
- elim (ldrop_trans_le … HLK … HKV) -K /2 width=2 by lt_to_le/ #X #HLK #H
- elim (ldrop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid #K #Y #_ #HVY
+ elim (drop_trans_le … HLK … HKV) -K /2 width=2 by lt_to_le/ #X #HLK #H
+ elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid #K #Y #_ #HVY
>(lift_mono … HVY … HVW) -Y -HVW #H destruct /2 width=5 by cpy_subst/
| #a #I #G #K #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #d #e #HLK #H1 #H2 #Hdetd
elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1
elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct
- /4 width=7 by cpy_bind, ldrop_skip, yle_succ/
+ /4 width=7 by cpy_bind, drop_skip, yle_succ/
| #G #I #K #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #d #e #HLK #H1 #H2 #Hdetd
elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1
elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct
[ -Hdtd
lapply (ylt_yle_trans … (dt+et+e) … Hidet) // -Hidet #Hidete
elim (lift_trans_ge … HVW … HWU2) -W // <minus_plus #W #HVW #HWU2
- elim (ldrop_trans_le … HLK … HKV) -K /2 width=2 by lt_to_le/ #X #HLK #H
- elim (ldrop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid #K #Y #_ #HVY
+ elim (drop_trans_le … HLK … HKV) -K /2 width=2 by lt_to_le/ #X #HLK #H
+ elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid #K #Y #_ #HVY
>(lift_mono … HVY … HVW) -V #H destruct /2 width=5 by cpy_subst/
| -Hdti
elim (yle_inv_inj2 … Hdtd) -Hdtd #dtt #Hdtd #H destruct
lapply (transitive_le … Hdtd Hid) -Hdtd #Hdti
lapply (lift_trans_be … HVW … HWU2 ? ?) -W /2 width=1 by le_S/ >plus_plus_comm_23 #HVU2
- lapply (ldrop_trans_ge_comm … HLK … HKV ?) -K // -Hid
- /4 width=5 by cpy_subst, ldrop_inv_gen, monotonic_ylt_plus_dx, yle_plus_dx1_trans, yle_inj/
+ lapply (drop_trans_ge_comm … HLK … HKV ?) -K // -Hid
+ /4 width=5 by cpy_subst, drop_inv_gen, monotonic_ylt_plus_dx, yle_plus_dx1_trans, yle_inj/
]
| #a #I #G #K #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #d #e #HLK #H1 #H2 #Hdtd #Hddet
elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1
elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct
- /4 width=7 by cpy_bind, ldrop_skip, yle_succ/
+ /4 width=7 by cpy_bind, drop_skip, yle_succ/
| #I #G #K #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #d #e #HLK #H1 #H2 #Hdetd
elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1
elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct
elim (yle_inv_inj2 … Hid) -Hid #dd #Hddi #H0 destruct
lapply (lift_inv_lref1_ge … H … Hddi) -H #H destruct
lapply (lift_trans_be … HVW … HWU2 ? ?) -W /2 width=1 by le_S/ >plus_plus_comm_23 #HVU2
- lapply (ldrop_trans_ge_comm … HLK … HKV ?) -K // -Hddi
- /3 width=5 by cpy_subst, ldrop_inv_gen, monotonic_ylt_plus_dx, monotonic_yle_plus_dx/
+ lapply (drop_trans_ge_comm … HLK … HKV ?) -K // -Hddi
+ /3 width=5 by cpy_subst, drop_inv_gen, monotonic_ylt_plus_dx, monotonic_yle_plus_dx/
| #a #I #G #K #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #d #e #HLK #H1 #H2 #Hddt
elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1
elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct
- /4 width=6 by cpy_bind, ldrop_skip, yle_succ/
+ /4 width=6 by cpy_bind, drop_skip, yle_succ/
| #I #G #K #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #d #e #HLK #H1 #H2 #Hddt
elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1
elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct
lapply (ylt_yle_trans … Hdetd … Hidet) -Hdetd #Hid
lapply (ylt_inv_inj … Hid) -Hid #Hid
lapply (lift_inv_lref2_lt … H … Hid) -H #H destruct
- elim (ldrop_conf_lt … HLK … HLKV) -L // #L #U #HKL #_ #HUV
+ elim (drop_conf_lt … HLK … HLKV) -L // #L #U #HKL #_ #HUV
elim (lift_trans_le … HUV … HVW) -V // >minus_plus <plus_minus_m_m // -Hid /3 width=5 by cpy_subst, ex2_intro/
| #a #I #G #L #W1 #W2 #U1 #U2 #dt #et #_ #_ #IHW12 #IHU12 #K #s #d #e #HLK #X #H #Hdetd
elim (lift_inv_bind2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct
elim (IHW12 … HLK … HVW1) -IHW12 // #V2 #HV12 #HVW2
elim (IHU12 … HTU1) -IHU12 -HTU1
- /3 width=6 by cpy_bind, yle_succ, ldrop_skip, lift_bind, ex2_intro/
+ /3 width=6 by cpy_bind, yle_succ, drop_skip, lift_bind, ex2_intro/
| #I #G #L #W1 #W2 #U1 #U2 #dt #et #_ #_ #IHW12 #IHU12 #K #s #d #e #HLK #X #H #Hdetd
elim (lift_inv_flat2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct
elim (IHW12 … HLK … HVW1) -W1 //
elim (lift_inv_lref2 … H) -H * #Hid #H destruct [ -Hdtd -Hidet | -Hdti -Hdedet ]
[ lapply (ylt_yle_trans i d (dt+(et-e)) ? ?) /2 width=1 by ylt_inj/
[ >yplus_minus_assoc_inj /2 width=1 by yle_plus1_to_minus_inj2/ ] -Hdedet #Hidete
- elim (ldrop_conf_lt … HLK … HLKV) -L // #L #U #HKL #_ #HUV
+ elim (drop_conf_lt … HLK … HLKV) -L // #L #U #HKL #_ #HUV
elim (lift_trans_le … HUV … HVW) -V // >minus_plus <plus_minus_m_m // -Hid
/3 width=5 by cpy_subst, ex2_intro/
| elim (le_inv_plus_l … Hid) #Hdie #Hei
lapply (yle_trans … Hdtd (i-e) ?) /2 width=1 by yle_inj/ -Hdtd #Hdtie
- lapply (ldrop_conf_ge … HLK … HLKV ?) -L // #HKV
+ lapply (drop_conf_ge … HLK … HLKV ?) -L // #HKV
elim (lift_split … HVW d (i-e+1)) -HVW [2,3,4: /2 width=1 by le_S_S, le_S/ ] -Hid -Hdie
#V1 #HV1 >plus_minus // <minus_minus /2 width=1 by le_S/ <minus_n_n <plus_n_O #H
@(ex2_intro … H) @(cpy_subst … HKV HV1) // (**) (* explicit constructor *)
elim (lift_inv_bind2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct
elim (IHW12 … HLK … HVW1) -IHW12 // #V2 #HV12 #HVW2
elim (IHU12 … HTU1) -U1
- /3 width=6 by cpy_bind, ldrop_skip, lift_bind, yle_succ, ex2_intro/
+ /3 width=6 by cpy_bind, drop_skip, lift_bind, yle_succ, ex2_intro/
| #I #G #L #W1 #W2 #U1 #U2 #dt #et #_ #_ #IHW12 #IHU12 #K #s #d #e #HLK #X #H #Hdtd #Hdedet
elim (lift_inv_flat2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct
elim (IHW12 … HLK … HVW1) -W1 //
elim (yle_inv_plus_inj2 … Hdedt) -Hdedt #_ #Hedt
elim (yle_inv_plus_inj2 … Hdei) #Hdie #Hei
lapply (lift_inv_lref2_ge … H ?) -H /2 width=1 by yle_inv_inj/ #H destruct
- lapply (ldrop_conf_ge … HLK … HLKV ?) -L /2 width=1 by yle_inv_inj/ #HKV
+ lapply (drop_conf_ge … HLK … HLKV ?) -L /2 width=1 by yle_inv_inj/ #HKV
elim (lift_split … HVW d (i-e+1)) -HVW [2,3,4: /3 width=1 by yle_inv_inj, le_S_S, le_S/ ] -Hdei -Hdie
#V0 #HV10 >plus_minus /2 width=1 by yle_inv_inj/ <minus_minus /3 width=1 by yle_inv_inj, le_S/ <minus_n_n <plus_n_O #H
@(ex2_intro … H) @(cpy_subst … HKV HV10) (**) (* explicit constructor *)
elim (lift_inv_bind2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct
elim (yle_inv_plus_inj2 … Hdetd) #_ #Hedt
elim (IHW12 … HLK … HVW1) -IHW12 // #V2 #HV12 #HVW2
- elim (IHU12 … HTU1) -U1 [4: @ldrop_skip // |2,5: skip |3: /2 width=1 by yle_succ/ ]
+ elim (IHU12 … HTU1) -U1 [4: @drop_skip // |2,5: skip |3: /2 width=1 by yle_succ/ ]
>yminus_succ1_inj /3 width=5 by cpy_bind, lift_bind, ex2_intro/
| #I #G #L #W1 #W2 #U1 #U2 #dt #et #_ #_ #IHW12 #IHU12 #K #s #d #e #HLK #X #H #Hdetd
elim (lift_inv_flat2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct
| * #J #K #W #j #Hdj #Hji #HLK #HnW
elim (yle_inv_succ1 … Hdj) -Hdj #Hdj #Hj
lapply (ylt_O … Hj) -Hj #Hj
- lapply (ldrop_inv_drop1_lt … HLK ?) // -HLK #HLK
+ lapply (drop_inv_drop1_lt … HLK ?) // -HLK #HLK
>(plus_minus_m_m j 1) in ⊢ (%→?); [2: /3 width=3 by yle_trans, yle_inv_inj/ ]
#HnU1 <minus_le_minus_minus_comm in HnW;
/5 width=9 by nlift_bind_dx, monotonic_lt_pred, lt_plus_to_minus_r, ex5_4_intro, or_intror/
--- /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 "ground_2/lib/bool.ma".
+include "ground_2/lib/lstar.ma".
+include "basic_2/notation/relations/rdrop_5.ma".
+include "basic_2/notation/relations/rdrop_4.ma".
+include "basic_2/notation/relations/rdrop_3.ma".
+include "basic_2/grammar/lenv_length.ma".
+include "basic_2/grammar/cl_restricted_weight.ma".
+include "basic_2/substitution/lift.ma".
+
+(* BASIC SLICING FOR LOCAL ENVIRONMENTS *************************************)
+
+(* Basic_1: includes: drop_skip_bind *)
+inductive drop (s:bool): relation4 nat nat lenv lenv ≝
+| drop_atom: ∀d,e. (s = Ⓕ → e = 0) → drop s d e (⋆) (⋆)
+| drop_pair: ∀I,L,V. drop s 0 0 (L.ⓑ{I}V) (L.ⓑ{I}V)
+| drop_drop: ∀I,L1,L2,V,e. drop s 0 e L1 L2 → drop s 0 (e+1) (L1.ⓑ{I}V) L2
+| drop_skip: ∀I,L1,L2,V1,V2,d,e.
+ drop s d e L1 L2 → ⇧[d, e] V2 ≡ V1 →
+ drop s (d+1) e (L1.ⓑ{I}V1) (L2.ⓑ{I}V2)
+.
+
+interpretation
+ "basic slicing (local environment) abstract"
+ 'RDrop s d e L1 L2 = (drop s d e L1 L2).
+(*
+interpretation
+ "basic slicing (local environment) general"
+ 'RDrop d e L1 L2 = (drop true d e L1 L2).
+*)
+interpretation
+ "basic slicing (local environment) lget"
+ 'RDrop e L1 L2 = (drop false O e L1 L2).
+
+definition l_liftable: predicate (lenv → relation term) ≝
+ λR. ∀K,T1,T2. R K T1 T2 → ∀L,s,d,e. ⇩[s, d, e] L ≡ K →
+ ∀U1. ⇧[d, e] T1 ≡ U1 → ∀U2. ⇧[d, e] T2 ≡ U2 → R L U1 U2.
+
+definition l_deliftable_sn: predicate (lenv → relation term) ≝
+ λR. ∀L,U1,U2. R L U1 U2 → ∀K,s,d,e. ⇩[s, d, e] L ≡ K →
+ ∀T1. ⇧[d, e] T1 ≡ U1 →
+ ∃∃T2. ⇧[d, e] T2 ≡ U2 & R K T1 T2.
+
+definition dropable_sn: predicate (relation lenv) ≝
+ λR. ∀L1,K1,s,d,e. ⇩[s, d, e] L1 ≡ K1 → ∀L2. R L1 L2 →
+ ∃∃K2. R K1 K2 & ⇩[s, d, e] L2 ≡ K2.
+
+definition dropable_dx: predicate (relation lenv) ≝
+ λR. ∀L1,L2. R L1 L2 → ∀K2,s,e. ⇩[s, 0, e] L2 ≡ K2 →
+ ∃∃K1. ⇩[s, 0, e] L1 ≡ K1 & R K1 K2.
+
+(* Basic inversion lemmas ***************************************************)
+
+fact drop_inv_atom1_aux: ∀L1,L2,s,d,e. ⇩[s, d, e] L1 ≡ L2 → L1 = ⋆ →
+ L2 = ⋆ ∧ (s = Ⓕ → e = 0).
+#L1 #L2 #s #d #e * -L1 -L2 -d -e
+[ /3 width=1 by conj/
+| #I #L #V #H destruct
+| #I #L1 #L2 #V #e #_ #H destruct
+| #I #L1 #L2 #V1 #V2 #d #e #_ #_ #H destruct
+]
+qed-.
+
+(* Basic_1: was: drop_gen_sort *)
+lemma drop_inv_atom1: ∀L2,s,d,e. ⇩[s, d, e] ⋆ ≡ L2 → L2 = ⋆ ∧ (s = Ⓕ → e = 0).
+/2 width=4 by drop_inv_atom1_aux/ qed-.
+
+fact drop_inv_O1_pair1_aux: ∀L1,L2,s,d,e. ⇩[s, d, e] L1 ≡ L2 → d = 0 →
+ ∀K,I,V. L1 = K.ⓑ{I}V →
+ (e = 0 ∧ L2 = K.ⓑ{I}V) ∨
+ (0 < e ∧ ⇩[s, d, e-1] K ≡ L2).
+#L1 #L2 #s #d #e * -L1 -L2 -d -e
+[ #d #e #_ #_ #K #J #W #H destruct
+| #I #L #V #_ #K #J #W #HX destruct /3 width=1 by or_introl, conj/
+| #I #L1 #L2 #V #e #HL12 #_ #K #J #W #H destruct /3 width=1 by or_intror, conj/
+| #I #L1 #L2 #V1 #V2 #d #e #_ #_ >commutative_plus normalize #H destruct
+]
+qed-.
+
+lemma drop_inv_O1_pair1: ∀I,K,L2,V,s,e. ⇩[s, 0, e] K. ⓑ{I} V ≡ L2 →
+ (e = 0 ∧ L2 = K.ⓑ{I}V) ∨
+ (0 < e ∧ ⇩[s, 0, e-1] K ≡ L2).
+/2 width=3 by drop_inv_O1_pair1_aux/ qed-.
+
+lemma drop_inv_pair1: ∀I,K,L2,V,s. ⇩[s, 0, 0] K.ⓑ{I}V ≡ L2 → L2 = K.ⓑ{I}V.
+#I #K #L2 #V #s #H
+elim (drop_inv_O1_pair1 … H) -H * // #H destruct
+elim (lt_refl_false … H)
+qed-.
+
+(* Basic_1: was: drop_gen_drop *)
+lemma drop_inv_drop1_lt: ∀I,K,L2,V,s,e.
+ ⇩[s, 0, e] K.ⓑ{I}V ≡ L2 → 0 < e → ⇩[s, 0, e-1] K ≡ L2.
+#I #K #L2 #V #s #e #H #He
+elim (drop_inv_O1_pair1 … H) -H * // #H destruct
+elim (lt_refl_false … He)
+qed-.
+
+lemma drop_inv_drop1: ∀I,K,L2,V,s,e.
+ ⇩[s, 0, e+1] K.ⓑ{I}V ≡ L2 → ⇩[s, 0, e] K ≡ L2.
+#I #K #L2 #V #s #e #H lapply (drop_inv_drop1_lt … H ?) -H //
+qed-.
+
+fact drop_inv_skip1_aux: ∀L1,L2,s,d,e. ⇩[s, d, e] L1 ≡ L2 → 0 < d →
+ ∀I,K1,V1. L1 = K1.ⓑ{I}V1 →
+ ∃∃K2,V2. ⇩[s, d-1, e] K1 ≡ K2 &
+ ⇧[d-1, e] V2 ≡ V1 &
+ L2 = K2.ⓑ{I}V2.
+#L1 #L2 #s #d #e * -L1 -L2 -d -e
+[ #d #e #_ #_ #J #K1 #W1 #H destruct
+| #I #L #V #H elim (lt_refl_false … H)
+| #I #L1 #L2 #V #e #_ #H elim (lt_refl_false … H)
+| #I #L1 #L2 #V1 #V2 #d #e #HL12 #HV21 #_ #J #K1 #W1 #H destruct /2 width=5 by ex3_2_intro/
+]
+qed-.
+
+(* Basic_1: was: drop_gen_skip_l *)
+lemma drop_inv_skip1: ∀I,K1,V1,L2,s,d,e. ⇩[s, d, e] K1.ⓑ{I}V1 ≡ L2 → 0 < d →
+ ∃∃K2,V2. ⇩[s, d-1, e] K1 ≡ K2 &
+ ⇧[d-1, e] V2 ≡ V1 &
+ L2 = K2.ⓑ{I}V2.
+/2 width=3 by drop_inv_skip1_aux/ qed-.
+
+lemma drop_inv_O1_pair2: ∀I,K,V,s,e,L1. ⇩[s, 0, e] L1 ≡ K.ⓑ{I}V →
+ (e = 0 ∧ L1 = K.ⓑ{I}V) ∨
+ ∃∃I1,K1,V1. ⇩[s, 0, e-1] K1 ≡ K.ⓑ{I}V & L1 = K1.ⓑ{I1}V1 & 0 < e.
+#I #K #V #s #e *
+[ #H elim (drop_inv_atom1 … H) -H #H destruct
+| #L1 #I1 #V1 #H
+ elim (drop_inv_O1_pair1 … H) -H *
+ [ #H1 #H2 destruct /3 width=1 by or_introl, conj/
+ | /3 width=5 by ex3_3_intro, or_intror/
+ ]
+]
+qed-.
+
+fact drop_inv_skip2_aux: ∀L1,L2,s,d,e. ⇩[s, d, e] L1 ≡ L2 → 0 < d →
+ ∀I,K2,V2. L2 = K2.ⓑ{I}V2 →
+ ∃∃K1,V1. ⇩[s, d-1, e] K1 ≡ K2 &
+ ⇧[d-1, e] V2 ≡ V1 &
+ L1 = K1.ⓑ{I}V1.
+#L1 #L2 #s #d #e * -L1 -L2 -d -e
+[ #d #e #_ #_ #J #K2 #W2 #H destruct
+| #I #L #V #H elim (lt_refl_false … H)
+| #I #L1 #L2 #V #e #_ #H elim (lt_refl_false … H)
+| #I #L1 #L2 #V1 #V2 #d #e #HL12 #HV21 #_ #J #K2 #W2 #H destruct /2 width=5 by ex3_2_intro/
+]
+qed-.
+
+(* Basic_1: was: drop_gen_skip_r *)
+lemma drop_inv_skip2: ∀I,L1,K2,V2,s,d,e. ⇩[s, d, e] L1 ≡ K2.ⓑ{I}V2 → 0 < d →
+ ∃∃K1,V1. ⇩[s, d-1, e] K1 ≡ K2 & ⇧[d-1, e] V2 ≡ V1 &
+ L1 = K1.ⓑ{I}V1.
+/2 width=3 by drop_inv_skip2_aux/ qed-.
+
+lemma drop_inv_O1_gt: ∀L,K,e,s. ⇩[s, 0, e] L ≡ K → |L| < e →
+ s = Ⓣ ∧ K = ⋆.
+#L elim L -L [| #L #Z #X #IHL ] #K #e #s #H normalize in ⊢ (?%?→?); #H1e
+[ elim (drop_inv_atom1 … H) -H elim s -s /2 width=1 by conj/
+ #_ #Hs lapply (Hs ?) // -Hs #H destruct elim (lt_zero_false … H1e)
+| elim (drop_inv_O1_pair1 … H) -H * #H2e #HLK destruct
+ [ elim (lt_zero_false … H1e)
+ | elim (IHL … HLK) -IHL -HLK /2 width=1 by lt_plus_to_minus_r, conj/
+ ]
+]
+qed-.
+
+(* Basic properties *********************************************************)
+
+lemma drop_refl_atom_O2: ∀s,d. ⇩[s, d, O] ⋆ ≡ ⋆.
+/2 width=1 by drop_atom/ qed.
+
+(* Basic_1: was by definition: drop_refl *)
+lemma drop_refl: ∀L,d,s. ⇩[s, d, 0] L ≡ L.
+#L elim L -L //
+#L #I #V #IHL #d #s @(nat_ind_plus … d) -d /2 width=1 by drop_pair, drop_skip/
+qed.
+
+lemma drop_drop_lt: ∀I,L1,L2,V,s,e.
+ ⇩[s, 0, e-1] L1 ≡ L2 → 0 < e → ⇩[s, 0, e] L1.ⓑ{I}V ≡ L2.
+#I #L1 #L2 #V #s #e #HL12 #He >(plus_minus_m_m e 1) /2 width=1 by drop_drop/
+qed.
+
+lemma drop_skip_lt: ∀I,L1,L2,V1,V2,s,d,e.
+ ⇩[s, d-1, e] L1 ≡ L2 → ⇧[d-1, e] V2 ≡ V1 → 0 < d →
+ ⇩[s, d, e] L1. ⓑ{I} V1 ≡ L2.ⓑ{I}V2.
+#I #L1 #L2 #V1 #V2 #s #d #e #HL12 #HV21 #Hd >(plus_minus_m_m d 1) /2 width=1 by drop_skip/
+qed.
+
+lemma drop_O1_le: ∀s,e,L. e ≤ |L| → ∃K. ⇩[s, 0, e] L ≡ K.
+#s #e @(nat_ind_plus … e) -e /2 width=2 by ex_intro/
+#e #IHe *
+[ #H elim (le_plus_xSy_O_false … H)
+| #L #I #V normalize #H elim (IHe L) -IHe /3 width=2 by drop_drop, monotonic_pred, ex_intro/
+]
+qed-.
+
+lemma drop_O1_lt: ∀s,L,e. e < |L| → ∃∃I,K,V. ⇩[s, 0, e] L ≡ K.ⓑ{I}V.
+#s #L elim L -L
+[ #e #H elim (lt_zero_false … H)
+| #L #I #V #IHL #e @(nat_ind_plus … e) -e /2 width=4 by drop_pair, ex1_3_intro/
+ #e #_ normalize #H elim (IHL e) -IHL /3 width=4 by drop_drop, lt_plus_to_minus_r, lt_plus_to_lt_l, ex1_3_intro/
+]
+qed-.
+
+lemma drop_O1_pair: ∀L,K,e,s. ⇩[s, 0, e] L ≡ K → e ≤ |L| → ∀I,V.
+ ∃∃J,W. ⇩[s, 0, e] L.ⓑ{I}V ≡ K.ⓑ{J}W.
+#L elim L -L [| #L #Z #X #IHL ] #K #e #s #H normalize #He #I #V
+[ elim (drop_inv_atom1 … H) -H #H <(le_n_O_to_eq … He) -e
+ #Hs destruct /2 width=3 by ex1_2_intro/
+| elim (drop_inv_O1_pair1 … H) -H * #He #HLK destruct /2 width=3 by ex1_2_intro/
+ elim (IHL … HLK … Z X) -IHL -HLK
+ /3 width=3 by drop_drop_lt, le_plus_to_minus, ex1_2_intro/
+]
+qed-.
+
+lemma drop_O1_ge: ∀L,e. |L| ≤ e → ⇩[Ⓣ, 0, e] L ≡ ⋆.
+#L elim L -L [ #e #_ @drop_atom #H destruct ]
+#L #I #V #IHL #e @(nat_ind_plus … e) -e [ #H elim (le_plus_xSy_O_false … H) ]
+normalize /4 width=1 by drop_drop, monotonic_pred/
+qed.
+
+lemma drop_O1_eq: ∀L,s. ⇩[s, 0, |L|] L ≡ ⋆.
+#L elim L -L /2 width=1 by drop_drop, drop_atom/
+qed.
+
+lemma drop_split: ∀L1,L2,d,e2,s. ⇩[s, d, e2] L1 ≡ L2 → ∀e1. e1 ≤ e2 →
+ ∃∃L. ⇩[s, d, e2 - e1] L1 ≡ L & ⇩[s, d, e1] L ≡ L2.
+#L1 #L2 #d #e2 #s #H elim H -L1 -L2 -d -e2
+[ #d #e2 #Hs #e1 #He12 @(ex2_intro … (⋆))
+ @drop_atom #H lapply (Hs H) -s #H destruct /2 width=1 by le_n_O_to_eq/
+| #I #L1 #V #e1 #He1 lapply (le_n_O_to_eq … He1) -He1
+ #H destruct /2 width=3 by ex2_intro/
+| #I #L1 #L2 #V #e2 #HL12 #IHL12 #e1 @(nat_ind_plus … e1) -e1
+ [ /3 width=3 by drop_drop, ex2_intro/
+ | -HL12 #e1 #_ #He12 lapply (le_plus_to_le_r … He12) -He12
+ #He12 elim (IHL12 … He12) -IHL12 >minus_plus_plus_l
+ #L #HL1 #HL2 elim (lt_or_ge (|L1|) (e2-e1)) #H0
+ [ elim (drop_inv_O1_gt … HL1 H0) -HL1 #H1 #H2 destruct
+ elim (drop_inv_atom1 … HL2) -HL2 #H #_ destruct
+ @(ex2_intro … (⋆)) [ @drop_O1_ge normalize // ]
+ @drop_atom #H destruct
+ | elim (drop_O1_pair … HL1 H0 I V) -HL1 -H0 /3 width=5 by drop_drop, ex2_intro/
+ ]
+ ]
+| #I #L1 #L2 #V1 #V2 #d #e2 #_ #HV21 #IHL12 #e1 #He12 elim (IHL12 … He12) -IHL12
+ #L #HL1 #HL2 elim (lift_split … HV21 d e1) -HV21 /3 width=5 by drop_skip, ex2_intro/
+]
+qed-.
+
+lemma drop_FT: ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ L2 → ⇩[Ⓣ, d, e] L1 ≡ L2.
+#L1 #L2 #d #e #H elim H -L1 -L2 -d -e
+/3 width=1 by drop_atom, drop_drop, drop_skip/
+qed.
+
+lemma drop_gen: ∀L1,L2,s,d,e. ⇩[Ⓕ, d, e] L1 ≡ L2 → ⇩[s, d, e] L1 ≡ L2.
+#L1 #L2 * /2 width=1 by drop_FT/
+qed-.
+
+lemma drop_T: ∀L1,L2,s,d,e. ⇩[s, d, e] L1 ≡ L2 → ⇩[Ⓣ, d, e] L1 ≡ L2.
+#L1 #L2 * /2 width=1 by drop_FT/
+qed-.
+
+lemma l_liftable_LTC: ∀R. l_liftable R → l_liftable (LTC … R).
+#R #HR #K #T1 #T2 #H elim H -T2
+[ /3 width=10 by inj/
+| #T #T2 #_ #HT2 #IHT1 #L #s #d #e #HLK #U1 #HTU1 #U2 #HTU2
+ elim (lift_total T d e) /4 width=12 by step/
+]
+qed-.
+
+lemma l_deliftable_sn_LTC: ∀R. l_deliftable_sn R → l_deliftable_sn (LTC … R).
+#R #HR #L #U1 #U2 #H elim H -U2
+[ #U2 #HU12 #K #s #d #e #HLK #T1 #HTU1
+ elim (HR … HU12 … HLK … HTU1) -HR -L -U1 /3 width=3 by inj, ex2_intro/
+| #U #U2 #_ #HU2 #IHU1 #K #s #d #e #HLK #T1 #HTU1
+ elim (IHU1 … HLK … HTU1) -IHU1 -U1 #T #HTU #HT1
+ elim (HR … HU2 … HLK … HTU) -HR -L -U /3 width=5 by step, ex2_intro/
+]
+qed-.
+
+lemma dropable_sn_TC: ∀R. dropable_sn R → dropable_sn (TC … R).
+#R #HR #L1 #K1 #s #d #e #HLK1 #L2 #H elim H -L2
+[ #L2 #HL12 elim (HR … HLK1 … HL12) -HR -L1
+ /3 width=3 by inj, ex2_intro/
+| #L #L2 #_ #HL2 * #K #HK1 #HLK elim (HR … HLK … HL2) -HR -L
+ /3 width=3 by step, ex2_intro/
+]
+qed-.
+
+lemma dropable_dx_TC: ∀R. dropable_dx R → dropable_dx (TC … R).
+#R #HR #L1 #L2 #H elim H -L2
+[ #L2 #HL12 #K2 #s #e #HLK2 elim (HR … HL12 … HLK2) -HR -L2
+ /3 width=3 by inj, ex2_intro/
+| #L #L2 #_ #HL2 #IHL1 #K2 #s #e #HLK2 elim (HR … HL2 … HLK2) -HR -L2
+ #K #HLK #HK2 elim (IHL1 … HLK) -L
+ /3 width=5 by step, ex2_intro/
+]
+qed-.
+
+lemma l_deliftable_sn_llstar: ∀R. l_deliftable_sn R →
+ ∀l. l_deliftable_sn (llstar … R l).
+#R #HR #l #L #U1 #U2 #H @(lstar_ind_r … l U2 H) -l -U2
+[ /2 width=3 by lstar_O, ex2_intro/
+| #l #U #U2 #_ #HU2 #IHU1 #K #s #d #e #HLK #T1 #HTU1
+ elim (IHU1 … HLK … HTU1) -IHU1 -U1 #T #HTU #HT1
+ elim (HR … HU2 … HLK … HTU) -HR -L -U /3 width=5 by lstar_dx, ex2_intro/
+]
+qed-.
+
+(* Basic forvard lemmas *****************************************************)
+
+(* Basic_1: was: drop_S *)
+lemma drop_fwd_drop2: ∀L1,I2,K2,V2,s,e. ⇩[s, O, e] L1 ≡ K2. ⓑ{I2} V2 →
+ ⇩[s, O, e + 1] L1 ≡ K2.
+#L1 elim L1 -L1
+[ #I2 #K2 #V2 #s #e #H lapply (drop_inv_atom1 … H) -H * #H destruct
+| #K1 #I1 #V1 #IHL1 #I2 #K2 #V2 #s #e #H
+ elim (drop_inv_O1_pair1 … H) -H * #He #H
+ [ -IHL1 destruct /2 width=1 by drop_drop/
+ | @drop_drop >(plus_minus_m_m e 1) /2 width=3 by/
+ ]
+]
+qed-.
+
+lemma drop_fwd_length_ge: ∀L1,L2,d,e,s. ⇩[s, d, e] L1 ≡ L2 → |L1| ≤ d → |L2| = |L1|.
+#L1 #L2 #d #e #s #H elim H -L1 -L2 -d -e // normalize
+[ #I #L1 #L2 #V #e #_ #_ #H elim (le_plus_xSy_O_false … H)
+| /4 width=2 by le_plus_to_le_r, eq_f/
+]
+qed-.
+
+lemma drop_fwd_length_le_le: ∀L1,L2,d,e,s. ⇩[s, d, e] L1 ≡ L2 → d ≤ |L1| → e ≤ |L1| - d → |L2| = |L1| - e.
+#L1 #L2 #d #e #s #H elim H -L1 -L2 -d -e // normalize
+[ /3 width=2 by le_plus_to_le_r/
+| #I #L1 #L2 #V1 #V2 #d #e #_ #_ #IHL12 >minus_plus_plus_l
+ #Hd #He lapply (le_plus_to_le_r … Hd) -Hd
+ #Hd >IHL12 // -L2 >plus_minus /2 width=3 by transitive_le/
+]
+qed-.
+
+lemma drop_fwd_length_le_ge: ∀L1,L2,d,e,s. ⇩[s, d, e] L1 ≡ L2 → d ≤ |L1| → |L1| - d ≤ e → |L2| = d.
+#L1 #L2 #d #e #s #H elim H -L1 -L2 -d -e normalize
+[ /2 width=1 by le_n_O_to_eq/
+| #I #L #V #_ <minus_n_O #H elim (le_plus_xSy_O_false … H)
+| /3 width=2 by le_plus_to_le_r/
+| /4 width=2 by le_plus_to_le_r, eq_f/
+]
+qed-.
+
+lemma drop_fwd_length: ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ L2 → |L1| = |L2| + e.
+#L1 #L2 #d #e #H elim H -L1 -L2 -d -e // normalize /2 width=1 by/
+qed-.
+
+lemma drop_fwd_length_minus2: ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ L2 → |L2| = |L1| - e.
+#L1 #L2 #d #e #H lapply (drop_fwd_length … H) -H /2 width=1 by plus_minus, le_n/
+qed-.
+
+lemma drop_fwd_length_minus4: ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ L2 → e = |L1| - |L2|.
+#L1 #L2 #d #e #H lapply (drop_fwd_length … H) -H //
+qed-.
+
+lemma drop_fwd_length_le2: ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ L2 → e ≤ |L1|.
+#L1 #L2 #d #e #H lapply (drop_fwd_length … H) -H //
+qed-.
+
+lemma drop_fwd_length_le4: ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ L2 → |L2| ≤ |L1|.
+#L1 #L2 #d #e #H lapply (drop_fwd_length … H) -H //
+qed-.
+
+lemma drop_fwd_length_lt2: ∀L1,I2,K2,V2,d,e.
+ ⇩[Ⓕ, d, e] L1 ≡ K2. ⓑ{I2} V2 → e < |L1|.
+#L1 #I2 #K2 #V2 #d #e #H
+lapply (drop_fwd_length … H) normalize in ⊢ (%→?); -I2 -V2 //
+qed-.
+
+lemma drop_fwd_length_lt4: ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ L2 → 0 < e → |L2| < |L1|.
+#L1 #L2 #d #e #H lapply (drop_fwd_length … H) -H /2 width=1 by lt_minus_to_plus_r/
+qed-.
+
+lemma drop_fwd_length_eq1: ∀L1,L2,K1,K2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 →
+ |L1| = |L2| → |K1| = |K2|.
+#L1 #L2 #K1 #K2 #d #e #HLK1 #HLK2 #HL12
+lapply (drop_fwd_length … HLK1) -HLK1
+lapply (drop_fwd_length … HLK2) -HLK2
+/2 width=2 by injective_plus_r/
+qed-.
+
+lemma drop_fwd_length_eq2: ∀L1,L2,K1,K2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 →
+ |K1| = |K2| → |L1| = |L2|.
+#L1 #L2 #K1 #K2 #d #e #HLK1 #HLK2 #HL12
+lapply (drop_fwd_length … HLK1) -HLK1
+lapply (drop_fwd_length … HLK2) -HLK2 //
+qed-.
+
+lemma drop_fwd_lw: ∀L1,L2,s,d,e. ⇩[s, d, e] L1 ≡ L2 → ♯{L2} ≤ ♯{L1}.
+#L1 #L2 #s #d #e #H elim H -L1 -L2 -d -e // normalize
+[ /2 width=3 by transitive_le/
+| #I #L1 #L2 #V1 #V2 #d #e #_ #HV21 #IHL12
+ >(lift_fwd_tw … HV21) -HV21 /2 width=1 by monotonic_le_plus_l/
+]
+qed-.
+
+lemma drop_fwd_lw_lt: ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ L2 → 0 < e → ♯{L2} < ♯{L1}.
+#L1 #L2 #d #e #H elim H -L1 -L2 -d -e
+[ #d #e #H >H -H //
+| #I #L #V #H elim (lt_refl_false … H)
+| #I #L1 #L2 #V #e #HL12 #_ #_
+ lapply (drop_fwd_lw … HL12) -HL12 #HL12
+ @(le_to_lt_to_lt … HL12) -HL12 //
+| #I #L1 #L2 #V1 #V2 #d #e #_ #HV21 #IHL12 #H normalize in ⊢ (?%%); -I
+ >(lift_fwd_tw … HV21) -V2 /3 by lt_minus_to_plus/
+]
+qed-.
+
+lemma drop_fwd_rfw: ∀I,L,K,V,i. ⇩[i] L ≡ K.ⓑ{I}V → ∀T. ♯{K, V} < ♯{L, T}.
+#I #L #K #V #i #HLK lapply (drop_fwd_lw … HLK) -HLK
+normalize in ⊢ (%→?→?%%); /3 width=3 by le_to_lt_to_lt/
+qed-.
+
+(* Advanced inversion lemmas ************************************************)
+
+fact drop_inv_O2_aux: ∀L1,L2,s,d,e. ⇩[s, d, e] L1 ≡ L2 → e = 0 → L1 = L2.
+#L1 #L2 #s #d #e #H elim H -L1 -L2 -d -e
+[ //
+| //
+| #I #L1 #L2 #V #e #_ #_ >commutative_plus normalize #H destruct
+| #I #L1 #L2 #V1 #V2 #d #e #_ #HV21 #IHL12 #H
+ >(IHL12 H) -L1 >(lift_inv_O2_aux … HV21 … H) -V2 -d -e //
+]
+qed-.
+
+(* Basic_1: was: drop_gen_refl *)
+lemma drop_inv_O2: ∀L1,L2,s,d. ⇩[s, d, 0] L1 ≡ L2 → L1 = L2.
+/2 width=5 by drop_inv_O2_aux/ qed-.
+
+lemma drop_inv_length_eq: ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ L2 → |L1| = |L2| → e = 0.
+#L1 #L2 #d #e #H #HL12 lapply (drop_fwd_length_minus4 … H) //
+qed-.
+
+lemma drop_inv_refl: ∀L,d,e. ⇩[Ⓕ, d, e] L ≡ L → e = 0.
+/2 width=5 by drop_inv_length_eq/ qed-.
+
+fact drop_inv_FT_aux: ∀L1,L2,s,d,e. ⇩[s, d, e] L1 ≡ L2 →
+ ∀I,K,V. L2 = K.ⓑ{I}V → s = Ⓣ → d = 0 →
+ ⇩[Ⓕ, d, e] L1 ≡ K.ⓑ{I}V.
+#L1 #L2 #s #d #e #H elim H -L1 -L2 -d -e
+[ #d #e #_ #J #K #W #H destruct
+| #I #L #V #J #K #W #H destruct //
+| #I #L1 #L2 #V #e #_ #IHL12 #J #K #W #H1 #H2 destruct
+ /3 width=1 by drop_drop/
+| #I #L1 #L2 #V1 #V2 #d #e #_ #_ #_ #J #K #W #_ #_
+ <plus_n_Sm #H destruct
+]
+qed-.
+
+lemma drop_inv_FT: ∀I,L,K,V,e. ⇩[Ⓣ, 0, e] L ≡ K.ⓑ{I}V → ⇩[e] L ≡ K.ⓑ{I}V.
+/2 width=5 by drop_inv_FT_aux/ qed.
+
+lemma drop_inv_gen: ∀I,L,K,V,s,e. ⇩[s, 0, e] L ≡ K.ⓑ{I}V → ⇩[e] L ≡ K.ⓑ{I}V.
+#I #L #K #V * /2 width=1 by drop_inv_FT/
+qed-.
+
+lemma drop_inv_T: ∀I,L,K,V,s,e. ⇩[Ⓣ, 0, e] L ≡ K.ⓑ{I}V → ⇩[s, 0, e] L ≡ K.ⓑ{I}V.
+#I #L #K #V * /2 width=1 by drop_inv_FT/
+qed-.
+
+(* Basic_1: removed theorems 50:
+ drop_ctail drop_skip_flat
+ cimp_flat_sx cimp_flat_dx cimp_bind cimp_getl_conf
+ drop_clear drop_clear_O drop_clear_S
+ clear_gen_sort clear_gen_bind clear_gen_flat clear_gen_flat_r
+ clear_gen_all clear_clear clear_mono clear_trans clear_ctail clear_cle
+ getl_ctail_clen getl_gen_tail clear_getl_trans getl_clear_trans
+ getl_clear_bind getl_clear_conf getl_dec getl_drop getl_drop_conf_lt
+ getl_drop_conf_ge getl_conf_ge_drop getl_drop_conf_rev
+ drop_getl_trans_lt drop_getl_trans_le drop_getl_trans_ge
+ getl_drop_trans getl_flt getl_gen_all getl_gen_sort getl_gen_O
+ getl_gen_S getl_gen_2 getl_gen_flat getl_gen_bind getl_conf_le
+ getl_trans getl_refl getl_head getl_flat getl_ctail getl_mono
+*)
--- /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/grammar/lenv_append.ma".
+include "basic_2/substitution/drop.ma".
+
+(* DROPPING *****************************************************************)
+
+(* Properties on append for local environments ******************************)
+
+fact drop_O1_append_sn_le_aux: ∀L1,L2,s,d,e. ⇩[s, d, e] L1 ≡ L2 →
+ d = 0 → e ≤ |L1| →
+ ∀L. ⇩[s, 0, e] L @@ L1 ≡ L @@ L2.
+#L1 #L2 #s #d #e #H elim H -L1 -L2 -d -e normalize
+[2,3,4: /4 width=1 by drop_skip_lt, drop_drop, arith_b1, lt_minus_to_plus_r, monotonic_pred/ ]
+#d #e #_ #_ #H <(le_n_O_to_eq … H) -H //
+qed-.
+
+lemma drop_O1_append_sn_le: ∀L1,L2,s,e. ⇩[s, 0, e] L1 ≡ L2 → e ≤ |L1| →
+ ∀L. ⇩[s, 0, e] L @@ L1 ≡ L @@ L2.
+/2 width=3 by drop_O1_append_sn_le_aux/ qed.
+
+(* Inversion lemmas on append for local environments ************************)
+
+lemma drop_O1_inv_append1_ge: ∀K,L1,L2,s,e. ⇩[s, 0, e] L1 @@ L2 ≡ K →
+ |L2| ≤ e → ⇩[s, 0, e - |L2|] L1 ≡ K.
+#K #L1 #L2 elim L2 -L2 normalize //
+#L2 #I #V #IHL2 #s #e #H #H1e
+elim (drop_inv_O1_pair1 … H) -H * #H2e #HL12 destruct
+[ lapply (le_n_O_to_eq … H1e) -H1e -IHL2
+ >commutative_plus normalize #H destruct
+| <minus_plus >minus_minus_comm /3 width=1 by monotonic_pred/
+]
+qed-.
+
+lemma drop_O1_inv_append1_le: ∀K,L1,L2,s,e. ⇩[s, 0, e] L1 @@ L2 ≡ K → e ≤ |L2| →
+ ∀K2. ⇩[s, 0, e] L2 ≡ K2 → K = L1 @@ K2.
+#K #L1 #L2 elim L2 -L2 normalize
+[ #s #e #H1 #H2 #K2 #H3 lapply (le_n_O_to_eq … H2) -H2
+ #H2 elim (drop_inv_atom1 … H3) -H3 #H3 #_ destruct
+ >(drop_inv_O2 … H1) -H1 //
+| #L2 #I #V #IHL2 #s #e @(nat_ind_plus … e) -e [ -IHL2 ]
+ [ #H1 #_ #K2 #H2
+ lapply (drop_inv_O2 … H1) -H1 #H1
+ lapply (drop_inv_O2 … H2) -H2 #H2 destruct //
+ | /4 width=7 by drop_inv_drop1, le_plus_to_le_r/
+ ]
+]
+qed-.
--- /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/substitution/lift_lift.ma".
+include "basic_2/substitution/drop.ma".
+
+(* BASIC SLICING FOR LOCAL ENVIRONMENTS *************************************)
+
+(* Main properties **********************************************************)
+
+(* Basic_1: was: drop_mono *)
+theorem drop_mono: ∀L,L1,s1,d,e. ⇩[s1, d, e] L ≡ L1 →
+ ∀L2,s2. ⇩[s2, d, e] L ≡ L2 → L1 = L2.
+#L #L1 #s1 #d #e #H elim H -L -L1 -d -e
+[ #d #e #He #L2 #s2 #H elim (drop_inv_atom1 … H) -H //
+| #I #K #V #L2 #s2 #HL12 <(drop_inv_O2 … HL12) -L2 //
+| #I #L #K #V #e #_ #IHLK #L2 #s2 #H
+ lapply (drop_inv_drop1 … H) -H /2 width=2 by/
+| #I #L #K1 #T #V1 #d #e #_ #HVT1 #IHLK1 #X #s2 #H
+ elim (drop_inv_skip1 … H) -H // <minus_plus_m_m #K2 #V2 #HLK2 #HVT2 #H destruct
+ >(lift_inj … HVT1 … HVT2) -HVT1 -HVT2
+ >(IHLK1 … HLK2) -IHLK1 -HLK2 //
+]
+qed-.
+
+(* Basic_1: was: drop_conf_ge *)
+theorem drop_conf_ge: ∀L,L1,s1,d1,e1. ⇩[s1, d1, e1] L ≡ L1 →
+ ∀L2,s2,e2. ⇩[s2, 0, e2] L ≡ L2 → d1 + e1 ≤ e2 →
+ ⇩[s2, 0, e2 - e1] L1 ≡ L2.
+#L #L1 #s1 #d1 #e1 #H elim H -L -L1 -d1 -e1 //
+[ #d #e #_ #L2 #s2 #e2 #H #_ elim (drop_inv_atom1 … H) -H
+ #H #He destruct
+ @drop_atom #H >He // (**) (* explicit constructor *)
+| #I #L #K #V #e #_ #IHLK #L2 #s2 #e2 #H #He2
+ lapply (drop_inv_drop1_lt … H ?) -H /2 width=2 by ltn_to_ltO/ #HL2
+ <minus_plus >minus_minus_comm /3 width=1 by monotonic_pred/
+| #I #L #K #V1 #V2 #d #e #_ #_ #IHLK #L2 #s2 #e2 #H #Hdee2
+ lapply (transitive_le 1 … Hdee2) // #He2
+ lapply (drop_inv_drop1_lt … H ?) -H // -He2 #HL2
+ lapply (transitive_le (1+e) … Hdee2) // #Hee2
+ @drop_drop_lt >minus_minus_comm /3 width=1 by lt_minus_to_plus_r, monotonic_le_minus_r, monotonic_pred/ (**) (* explicit constructor *)
+]
+qed.
+
+(* Note: apparently this was missing in basic_1 *)
+theorem drop_conf_be: ∀L0,L1,s1,d1,e1. ⇩[s1, d1, e1] L0 ≡ L1 →
+ ∀L2,e2. ⇩[e2] L0 ≡ L2 → d1 ≤ e2 → e2 ≤ d1 + e1 →
+ ∃∃L. ⇩[s1, 0, d1 + e1 - e2] L2 ≡ L & ⇩[d1] L1 ≡ L.
+#L0 #L1 #s1 #d1 #e1 #H elim H -L0 -L1 -d1 -e1
+[ #d1 #e1 #He1 #L2 #e2 #H #Hd1 #_ elim (drop_inv_atom1 … H) -H #H #He2 destruct
+ >(He2 ?) in Hd1; // -He2 #Hd1 <(le_n_O_to_eq … Hd1) -d1
+ /4 width=3 by drop_atom, ex2_intro/
+| normalize #I #L #V #L2 #e2 #HL2 #_ #He2
+ lapply (le_n_O_to_eq … He2) -He2 #H destruct
+ lapply (drop_inv_O2 … HL2) -HL2 #H destruct /2 width=3 by drop_pair, ex2_intro/
+| normalize #I #L0 #K0 #V1 #e1 #HLK0 #IHLK0 #L2 #e2 #H #_ #He21
+ lapply (drop_inv_O1_pair1 … H) -H * * #He2 #HL20
+ [ -IHLK0 -He21 destruct <minus_n_O /3 width=3 by drop_drop, ex2_intro/
+ | -HLK0 <minus_le_minus_minus_comm //
+ elim (IHLK0 … HL20) -L0 /2 width=3 by monotonic_pred, ex2_intro/
+ ]
+| #I #L0 #K0 #V0 #V1 #d1 #e1 >plus_plus_comm_23 #_ #_ #IHLK0 #L2 #e2 #H #Hd1e2 #He2de1
+ elim (le_inv_plus_l … Hd1e2) #_ #He2
+ <minus_le_minus_minus_comm //
+ lapply (drop_inv_drop1_lt … H ?) -H // #HL02
+ elim (IHLK0 … HL02) -L0 /3 width=3 by drop_drop, monotonic_pred, ex2_intro/
+]
+qed-.
+
+(* Note: apparently this was missing in basic_1 *)
+theorem drop_conf_le: ∀L0,L1,s1,d1,e1. ⇩[s1, d1, e1] L0 ≡ L1 →
+ ∀L2,s2,e2. ⇩[s2, 0, e2] L0 ≡ L2 → e2 ≤ d1 →
+ ∃∃L. ⇩[s2, 0, e2] L1 ≡ L & ⇩[s1, d1 - e2, e1] L2 ≡ L.
+#L0 #L1 #s1 #d1 #e1 #H elim H -L0 -L1 -d1 -e1
+[ #d1 #e1 #He1 #L2 #s2 #e2 #H elim (drop_inv_atom1 … H) -H
+ #H #He2 #_ destruct /4 width=3 by drop_atom, ex2_intro/
+| #I #K0 #V0 #L2 #s2 #e2 #H1 #H2
+ lapply (le_n_O_to_eq … H2) -H2 #H destruct
+ lapply (drop_inv_pair1 … H1) -H1 #H destruct /2 width=3 by drop_pair, ex2_intro/
+| #I #K0 #K1 #V0 #e1 #HK01 #_ #L2 #s2 #e2 #H1 #H2
+ lapply (le_n_O_to_eq … H2) -H2 #H destruct
+ lapply (drop_inv_pair1 … H1) -H1 #H destruct /3 width=3 by drop_drop, ex2_intro/
+| #I #K0 #K1 #V0 #V1 #d1 #e1 #HK01 #HV10 #IHK01 #L2 #s2 #e2 #H #He2d1
+ elim (drop_inv_O1_pair1 … H) -H *
+ [ -IHK01 -He2d1 #H1 #H2 destruct /3 width=5 by drop_pair, drop_skip, ex2_intro/
+ | -HK01 -HV10 #He2 #HK0L2
+ elim (IHK01 … HK0L2) -IHK01 -HK0L2 /2 width=1 by monotonic_pred/
+ >minus_le_minus_minus_comm /3 width=3 by drop_drop_lt, ex2_intro/
+ ]
+]
+qed-.
+
+(* Note: with "s2", the conclusion parameter is "s1 ∨ s2" *)
+(* Basic_1: was: drop_trans_ge *)
+theorem drop_trans_ge: ∀L1,L,s1,d1,e1. ⇩[s1, d1, e1] L1 ≡ L →
+ ∀L2,e2. ⇩[e2] L ≡ L2 → d1 ≤ e2 → ⇩[s1, 0, e1 + e2] L1 ≡ L2.
+#L1 #L #s1 #d1 #e1 #H elim H -L1 -L -d1 -e1
+[ #d1 #e1 #He1 #L2 #e2 #H #_ elim (drop_inv_atom1 … H) -H
+ #H #He2 destruct /4 width=1 by drop_atom, eq_f2/
+| /2 width=1 by drop_gen/
+| /3 width=1 by drop_drop/
+| #I #L1 #L2 #V1 #V2 #d #e #_ #_ #IHL12 #L #e2 #H #Hde2
+ lapply (lt_to_le_to_lt 0 … Hde2) // #He2
+ lapply (lt_to_le_to_lt … (e + e2) He2 ?) // #Hee2
+ lapply (drop_inv_drop1_lt … H ?) -H // #HL2
+ @drop_drop_lt // >le_plus_minus /3 width=1 by monotonic_pred/
+]
+qed.
+
+(* Basic_1: was: drop_trans_le *)
+theorem drop_trans_le: ∀L1,L,s1,d1,e1. ⇩[s1, d1, e1] L1 ≡ L →
+ ∀L2,s2,e2. ⇩[s2, 0, e2] L ≡ L2 → e2 ≤ d1 →
+ ∃∃L0. ⇩[s2, 0, e2] L1 ≡ L0 & ⇩[s1, d1 - e2, e1] L0 ≡ L2.
+#L1 #L #s1 #d1 #e1 #H elim H -L1 -L -d1 -e1
+[ #d1 #e1 #He1 #L2 #s2 #e2 #H #_ elim (drop_inv_atom1 … H) -H
+ #H #He2 destruct /4 width=3 by drop_atom, ex2_intro/
+| #I #K #V #L2 #s2 #e2 #HL2 #H lapply (le_n_O_to_eq … H) -H
+ #H destruct /2 width=3 by drop_pair, ex2_intro/
+| #I #L1 #L2 #V #e #_ #IHL12 #L #s2 #e2 #HL2 #H lapply (le_n_O_to_eq … H) -H
+ #H destruct elim (IHL12 … HL2) -IHL12 -HL2 //
+ #L0 #H #HL0 lapply (drop_inv_O2 … H) -H #H destruct
+ /3 width=5 by drop_pair, drop_drop, ex2_intro/
+| #I #L1 #L2 #V1 #V2 #d #e #HL12 #HV12 #IHL12 #L #s2 #e2 #H #He2d
+ elim (drop_inv_O1_pair1 … H) -H *
+ [ -He2d -IHL12 #H1 #H2 destruct /3 width=5 by drop_pair, drop_skip, ex2_intro/
+ | -HL12 -HV12 #He2 #HL2
+ elim (IHL12 … HL2) -L2 [ >minus_le_minus_minus_comm // /3 width=3 by drop_drop_lt, ex2_intro/ | /2 width=1 by monotonic_pred/ ]
+ ]
+]
+qed-.
+
+(* Advanced properties ******************************************************)
+
+lemma l_liftable_llstar: ∀R. l_liftable R → ∀l. l_liftable (llstar … R l).
+#R #HR #l #K #T1 #T2 #H @(lstar_ind_r … l T2 H) -l -T2
+[ #L #s #d #e #_ #U1 #HTU1 #U2 #HTU2 -HR -K
+ >(lift_mono … HTU2 … HTU1) -T1 -U2 -d -e //
+| #l #T #T2 #_ #HT2 #IHT1 #L #s #d #e #HLK #U1 #HTU1 #U2 #HTU2
+ elim (lift_total T d e) /3 width=12 by lstar_dx/
+]
+qed-.
+
+(* Basic_1: was: drop_conf_lt *)
+lemma drop_conf_lt: ∀L,L1,s1,d1,e1. ⇩[s1, d1, e1] L ≡ L1 →
+ ∀I,K2,V2,s2,e2. ⇩[s2, 0, e2] L ≡ K2.ⓑ{I}V2 →
+ e2 < d1 → let d ≝ d1 - e2 - 1 in
+ ∃∃K1,V1. ⇩[s2, 0, e2] L1 ≡ K1.ⓑ{I}V1 &
+ ⇩[s1, d, e1] K2 ≡ K1 & ⇧[d, e1] V1 ≡ V2.
+#L #L1 #s1 #d1 #e1 #H1 #I #K2 #V2 #s2 #e2 #H2 #He2d1
+elim (drop_conf_le … H1 … H2) -L /2 width=2 by lt_to_le/ #K #HL1K #HK2
+elim (drop_inv_skip1 … HK2) -HK2 /2 width=1 by lt_plus_to_minus_r/
+#K1 #V1 #HK21 #HV12 #H destruct /2 width=5 by ex3_2_intro/
+qed-.
+
+(* Note: apparently this was missing in basic_1 *)
+lemma drop_trans_lt: ∀L1,L,s1,d1,e1. ⇩[s1, d1, e1] L1 ≡ L →
+ ∀I,L2,V2,s2,e2. ⇩[s2, 0, e2] L ≡ L2.ⓑ{I}V2 →
+ e2 < d1 → let d ≝ d1 - e2 - 1 in
+ ∃∃L0,V0. ⇩[s2, 0, e2] L1 ≡ L0.ⓑ{I}V0 &
+ ⇩[s1, d, e1] L0 ≡ L2 & ⇧[d, e1] V2 ≡ V0.
+#L1 #L #s1 #d1 #e1 #HL1 #I #L2 #V2 #s2 #e2 #HL2 #Hd21
+elim (drop_trans_le … HL1 … HL2) -L /2 width=1 by lt_to_le/ #L0 #HL10 #HL02
+elim (drop_inv_skip2 … HL02) -HL02 /2 width=1 by lt_plus_to_minus_r/ #L #V1 #HL2 #HV21 #H destruct /2 width=5 by ex3_2_intro/
+qed-.
+
+lemma drop_trans_ge_comm: ∀L1,L,L2,s1,d1,e1,e2.
+ ⇩[s1, d1, e1] L1 ≡ L → ⇩[e2] L ≡ L2 → d1 ≤ e2 →
+ ⇩[s1, 0, e2 + e1] L1 ≡ L2.
+#L1 #L #L2 #s1 #d1 #e1 #e2
+>commutative_plus /2 width=5 by drop_trans_ge/
+qed.
+
+lemma drop_conf_div: ∀I1,L,K,V1,e1. ⇩[e1] L ≡ K.ⓑ{I1}V1 →
+ ∀I2,V2,e2. ⇩[e2] L ≡ K.ⓑ{I2}V2 →
+ ∧∧ e1 = e2 & I1 = I2 & V1 = V2.
+#I1 #L #K #V1 #e1 #HLK1 #I2 #V2 #e2 #HLK2
+elim (le_or_ge e1 e2) #He
+[ lapply (drop_conf_ge … HLK1 … HLK2 ?)
+| lapply (drop_conf_ge … HLK2 … HLK1 ?)
+] -HLK1 -HLK2 // #HK
+lapply (drop_fwd_length_minus2 … HK) #H
+elim (discr_minus_x_xy … H) -H
+[1,3: normalize <plus_n_Sm #H destruct ]
+#H >H in HK; #HK
+lapply (drop_inv_O2 … HK) -HK #H destruct
+lapply (inv_eq_minus_O … H) -H /3 width=1 by le_to_le_to_eq, and3_intro/
+qed-.
+
+(* Advanced forward lemmas **************************************************)
+
+lemma drop_fwd_be: ∀L,K,s,d,e,i. ⇩[s, d, e] L ≡ K → |K| ≤ i → i < d → |L| ≤ i.
+#L #K #s #d #e #i #HLK #HK #Hd elim (lt_or_ge i (|L|)) //
+#HL elim (drop_O1_lt (Ⓕ) … HL) #I #K0 #V #HLK0 -HL
+elim (drop_conf_lt … HLK … HLK0) // -HLK -HLK0 -Hd
+#K1 #V1 #HK1 #_ #_ lapply (drop_fwd_length_lt2 … HK1) -I -K1 -V1
+#H elim (lt_refl_false i) /2 width=3 by lt_to_le_to_lt/
+qed-.
--- /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/grammar/leq_leq.ma".
+include "basic_2/substitution/drop.ma".
+
+(* BASIC SLICING FOR LOCAL ENVIRONMENTS *************************************)
+
+definition dedropable_sn: predicate (relation lenv) ≝
+ λR. ∀L1,K1,s,d,e. ⇩[s, d, e] L1 ≡ K1 → ∀K2. R K1 K2 →
+ ∃∃L2. R L1 L2 & ⇩[s, d, e] L2 ≡ K2 & L1 ⩬[d, e] L2.
+
+(* Properties on equivalence ************************************************)
+
+lemma leq_drop_trans_be: ∀L1,L2,d,e. L1 ⩬[d, e] L2 →
+ ∀I,K2,W,s,i. ⇩[s, 0, i] L2 ≡ K2.ⓑ{I}W →
+ d ≤ i → i < d + e →
+ ∃∃K1. K1 ⩬[0, ⫰(d+e-i)] K2 & ⇩[s, 0, i] L1 ≡ K1.ⓑ{I}W.
+#L1 #L2 #d #e #H elim H -L1 -L2 -d -e
+[ #d #e #J #K2 #W #s #i #H
+ elim (drop_inv_atom1 … H) -H #H destruct
+| #I1 #I2 #L1 #L2 #V1 #V2 #_ #_ #J #K2 #W #s #i #_ #_ #H
+ elim (ylt_yle_false … H) //
+| #I #L1 #L2 #V #e #HL12 #IHL12 #J #K2 #W #s #i #H #_ >yplus_O1
+ elim (drop_inv_O1_pair1 … H) -H * #Hi #HLK1 [ -IHL12 | -HL12 ]
+ [ #_ destruct >ypred_succ
+ /2 width=3 by drop_pair, ex2_intro/
+ | lapply (ylt_inv_O1 i ?) /2 width=1 by ylt_inj/
+ #H <H -H #H lapply (ylt_inv_succ … H) -H
+ #Hie elim (IHL12 … HLK1) -IHL12 -HLK1 // -Hie
+ >yminus_succ <yminus_inj /3 width=3 by drop_drop_lt, ex2_intro/
+ ]
+| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #IHL12 #J #K2 #W #s #i #HLK2 #Hdi
+ elim (yle_inv_succ1 … Hdi) -Hdi
+ #Hdi #Hi <Hi >yplus_succ1 #H lapply (ylt_inv_succ … H) -H
+ #Hide lapply (drop_inv_drop1_lt … HLK2 ?) -HLK2 /2 width=1 by ylt_O/
+ #HLK1 elim (IHL12 … HLK1) -IHL12 -HLK1 <yminus_inj >yminus_SO2
+ /4 width=3 by ylt_O, drop_drop_lt, ex2_intro/
+]
+qed-.
+
+lemma leq_drop_conf_be: ∀L1,L2,d,e. L1 ⩬[d, e] L2 →
+ ∀I,K1,W,s,i. ⇩[s, 0, i] L1 ≡ K1.ⓑ{I}W →
+ d ≤ i → i < d + e →
+ ∃∃K2. K1 ⩬[0, ⫰(d+e-i)] K2 & ⇩[s, 0, i] L2 ≡ K2.ⓑ{I}W.
+#L1 #L2 #d #e #HL12 #I #K1 #W #s #i #HLK1 #Hdi #Hide
+elim (leq_drop_trans_be … (leq_sym … HL12) … HLK1) // -L1 -Hdi -Hide
+/3 width=3 by leq_sym, ex2_intro/
+qed-.
+
+lemma drop_O1_ex: ∀K2,i,L1. |L1| = |K2| + i →
+ ∃∃L2. L1 ⩬[0, i] L2 & ⇩[i] L2 ≡ K2.
+#K2 #i @(nat_ind_plus … i) -i
+[ /3 width=3 by leq_O2, ex2_intro/
+| #i #IHi #Y #Hi elim (drop_O1_lt (Ⓕ) Y 0) //
+ #I #L1 #V #H lapply (drop_inv_O2 … H) -H #H destruct
+ normalize in Hi; elim (IHi L1) -IHi
+ /3 width=5 by drop_drop, leq_pair, injective_plus_l, ex2_intro/
+]
+qed-.
+
+lemma dedropable_sn_TC: ∀R. dedropable_sn R → dedropable_sn (TC … R).
+#R #HR #L1 #K1 #s #d #e #HLK1 #K2 #H elim H -K2
+[ #K2 #HK12 elim (HR … HLK1 … HK12) -HR -K1
+ /3 width=4 by inj, ex3_intro/
+| #K #K2 #_ #HK2 * #L #H1L1 #HLK #H2L1 elim (HR … HLK … HK2) -HR -K
+ /3 width=6 by leq_trans, step, ex3_intro/
+]
+qed-.
+
+(* Inversion lemmas on equivalence ******************************************)
+
+lemma drop_O1_inj: ∀i,L1,L2,K. ⇩[i] L1 ≡ K → ⇩[i] L2 ≡ K → L1 ⩬[i, ∞] L2.
+#i @(nat_ind_plus … i) -i
+[ #L1 #L2 #K #H <(drop_inv_O2 … H) -K #H <(drop_inv_O2 … H) -L1 //
+| #i #IHi * [2: #L1 #I1 #V1 ] * [2,4: #L2 #I2 #V2 ] #K #HLK1 #HLK2 //
+ lapply (drop_fwd_length … HLK1)
+ <(drop_fwd_length … HLK2) [ /4 width=5 by drop_inv_drop1, leq_succ/ ]
+ normalize <plus_n_Sm #H destruct
+]
+qed-.
include "basic_2/notation/relations/supterm_6.ma".
include "basic_2/grammar/cl_weight.ma".
-include "basic_2/substitution/ldrop.ma".
+include "basic_2/substitution/drop.ma".
(* SUPCLOSURE ***************************************************************)
qed.
lemma fqu_lref_S_lt: ∀I,G,L,V,i. 0 < i → ⦃G, L.ⓑ{I}V, #i⦄ ⊐ ⦃G, L, #(i-1)⦄.
-/3 width=3 by fqu_drop, ldrop_drop, lift_lref_ge_minus/
+/3 width=3 by fqu_drop, drop_drop, lift_lref_ge_minus/
qed.
(* Basic forward lemmas *****************************************************)
lemma fqu_fwd_fw: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → ♯{G2, L2, T2} < ♯{G1, L1, T1}.
#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 //
#G #L #K #T #U #e #HLK #HTU
-lapply (ldrop_fwd_lw_lt … HLK ?) -HLK // #HKL
+lapply (drop_fwd_lw_lt … HLK ?) -HLK // #HKL
lapply (lift_fwd_tw … HTU) -e #H
normalize in ⊢ (?%%); /2 width=1 by lt_minus_to_plus/
qed-.
#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
[1: normalize //
|3: #a
-|5: /2 width=4 by ldrop_fwd_length_lt4/
+|5: /2 width=4 by drop_fwd_length_lt4/
] #I #G #L #V #T #j #H destruct
qed-.
lemma fquq_fwd_fw: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → ♯{G2, L2, T2} ≤ ♯{G1, L1, T1}.
#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 /2 width=1 by lt_to_le/
#G1 #L1 #K1 #T1 #U1 #e #HLK1 #HTU1
-lapply (ldrop_fwd_lw … HLK1) -HLK1
+lapply (drop_fwd_lw … HLK1) -HLK1
lapply (lift_fwd_tw … HTU1) -HTU1
/2 width=1 by le_plus, le_n/
qed-.
#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 //
[ #a #I #G #L #V #T #j #H destruct
| #G1 #L1 #K1 #T1 #U1 #e #HLK1 #HTU1 #i #H destruct
- /2 width=3 by ldrop_fwd_length_le4/
+ /2 width=3 by drop_fwd_length_le4/
]
qed-.
⇩[e] L ≡ K → ⇧[0, e] T ≡ U → ⦃G, L, U⦄ ⊐⊐⸮ ⦃G, K, T⦄.
#G #L #K #T #U #e #HLK #HTU elim (eq_or_gt e)
/3 width=5 by fqu_drop_lt, or_introl/ #H destruct
->(ldrop_inv_O2 … HLK) -L >(lift_inv_O2 … HTU) -T //
+>(drop_inv_O2 … HLK) -L >(lift_inv_O2 … HTU) -T //
qed.
(* Main properties **********************************************************)
+++ /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 "ground_2/lib/bool.ma".
-include "ground_2/lib/lstar.ma".
-include "basic_2/notation/relations/rdrop_5.ma".
-include "basic_2/notation/relations/rdrop_4.ma".
-include "basic_2/notation/relations/rdrop_3.ma".
-include "basic_2/grammar/lenv_length.ma".
-include "basic_2/grammar/cl_restricted_weight.ma".
-include "basic_2/substitution/lift.ma".
-
-(* BASIC SLICING FOR LOCAL ENVIRONMENTS *************************************)
-
-(* Basic_1: includes: drop_skip_bind *)
-inductive ldrop (s:bool): relation4 nat nat lenv lenv ≝
-| ldrop_atom: ∀d,e. (s = Ⓕ → e = 0) → ldrop s d e (⋆) (⋆)
-| ldrop_pair: ∀I,L,V. ldrop s 0 0 (L.ⓑ{I}V) (L.ⓑ{I}V)
-| ldrop_drop: ∀I,L1,L2,V,e. ldrop s 0 e L1 L2 → ldrop s 0 (e+1) (L1.ⓑ{I}V) L2
-| ldrop_skip: ∀I,L1,L2,V1,V2,d,e.
- ldrop s d e L1 L2 → ⇧[d, e] V2 ≡ V1 →
- ldrop s (d+1) e (L1.ⓑ{I}V1) (L2.ⓑ{I}V2)
-.
-
-interpretation
- "basic slicing (local environment) abstract"
- 'RDrop s d e L1 L2 = (ldrop s d e L1 L2).
-(*
-interpretation
- "basic slicing (local environment) general"
- 'RDrop d e L1 L2 = (ldrop true d e L1 L2).
-*)
-interpretation
- "basic slicing (local environment) lget"
- 'RDrop e L1 L2 = (ldrop false O e L1 L2).
-
-definition l_liftable: predicate (lenv → relation term) ≝
- λR. ∀K,T1,T2. R K T1 T2 → ∀L,s,d,e. ⇩[s, d, e] L ≡ K →
- ∀U1. ⇧[d, e] T1 ≡ U1 → ∀U2. ⇧[d, e] T2 ≡ U2 → R L U1 U2.
-
-definition l_deliftable_sn: predicate (lenv → relation term) ≝
- λR. ∀L,U1,U2. R L U1 U2 → ∀K,s,d,e. ⇩[s, d, e] L ≡ K →
- ∀T1. ⇧[d, e] T1 ≡ U1 →
- ∃∃T2. ⇧[d, e] T2 ≡ U2 & R K T1 T2.
-
-definition dropable_sn: predicate (relation lenv) ≝
- λR. ∀L1,K1,s,d,e. ⇩[s, d, e] L1 ≡ K1 → ∀L2. R L1 L2 →
- ∃∃K2. R K1 K2 & ⇩[s, d, e] L2 ≡ K2.
-
-definition dropable_dx: predicate (relation lenv) ≝
- λR. ∀L1,L2. R L1 L2 → ∀K2,s,e. ⇩[s, 0, e] L2 ≡ K2 →
- ∃∃K1. ⇩[s, 0, e] L1 ≡ K1 & R K1 K2.
-
-(* Basic inversion lemmas ***************************************************)
-
-fact ldrop_inv_atom1_aux: ∀L1,L2,s,d,e. ⇩[s, d, e] L1 ≡ L2 → L1 = ⋆ →
- L2 = ⋆ ∧ (s = Ⓕ → e = 0).
-#L1 #L2 #s #d #e * -L1 -L2 -d -e
-[ /3 width=1 by conj/
-| #I #L #V #H destruct
-| #I #L1 #L2 #V #e #_ #H destruct
-| #I #L1 #L2 #V1 #V2 #d #e #_ #_ #H destruct
-]
-qed-.
-
-(* Basic_1: was: drop_gen_sort *)
-lemma ldrop_inv_atom1: ∀L2,s,d,e. ⇩[s, d, e] ⋆ ≡ L2 → L2 = ⋆ ∧ (s = Ⓕ → e = 0).
-/2 width=4 by ldrop_inv_atom1_aux/ qed-.
-
-fact ldrop_inv_O1_pair1_aux: ∀L1,L2,s,d,e. ⇩[s, d, e] L1 ≡ L2 → d = 0 →
- ∀K,I,V. L1 = K.ⓑ{I}V →
- (e = 0 ∧ L2 = K.ⓑ{I}V) ∨
- (0 < e ∧ ⇩[s, d, e-1] K ≡ L2).
-#L1 #L2 #s #d #e * -L1 -L2 -d -e
-[ #d #e #_ #_ #K #J #W #H destruct
-| #I #L #V #_ #K #J #W #HX destruct /3 width=1 by or_introl, conj/
-| #I #L1 #L2 #V #e #HL12 #_ #K #J #W #H destruct /3 width=1 by or_intror, conj/
-| #I #L1 #L2 #V1 #V2 #d #e #_ #_ >commutative_plus normalize #H destruct
-]
-qed-.
-
-lemma ldrop_inv_O1_pair1: ∀I,K,L2,V,s,e. ⇩[s, 0, e] K. ⓑ{I} V ≡ L2 →
- (e = 0 ∧ L2 = K.ⓑ{I}V) ∨
- (0 < e ∧ ⇩[s, 0, e-1] K ≡ L2).
-/2 width=3 by ldrop_inv_O1_pair1_aux/ qed-.
-
-lemma ldrop_inv_pair1: ∀I,K,L2,V,s. ⇩[s, 0, 0] K.ⓑ{I}V ≡ L2 → L2 = K.ⓑ{I}V.
-#I #K #L2 #V #s #H
-elim (ldrop_inv_O1_pair1 … H) -H * // #H destruct
-elim (lt_refl_false … H)
-qed-.
-
-(* Basic_1: was: drop_gen_drop *)
-lemma ldrop_inv_drop1_lt: ∀I,K,L2,V,s,e.
- ⇩[s, 0, e] K.ⓑ{I}V ≡ L2 → 0 < e → ⇩[s, 0, e-1] K ≡ L2.
-#I #K #L2 #V #s #e #H #He
-elim (ldrop_inv_O1_pair1 … H) -H * // #H destruct
-elim (lt_refl_false … He)
-qed-.
-
-lemma ldrop_inv_drop1: ∀I,K,L2,V,s,e.
- ⇩[s, 0, e+1] K.ⓑ{I}V ≡ L2 → ⇩[s, 0, e] K ≡ L2.
-#I #K #L2 #V #s #e #H lapply (ldrop_inv_drop1_lt … H ?) -H //
-qed-.
-
-fact ldrop_inv_skip1_aux: ∀L1,L2,s,d,e. ⇩[s, d, e] L1 ≡ L2 → 0 < d →
- ∀I,K1,V1. L1 = K1.ⓑ{I}V1 →
- ∃∃K2,V2. ⇩[s, d-1, e] K1 ≡ K2 &
- ⇧[d-1, e] V2 ≡ V1 &
- L2 = K2.ⓑ{I}V2.
-#L1 #L2 #s #d #e * -L1 -L2 -d -e
-[ #d #e #_ #_ #J #K1 #W1 #H destruct
-| #I #L #V #H elim (lt_refl_false … H)
-| #I #L1 #L2 #V #e #_ #H elim (lt_refl_false … H)
-| #I #L1 #L2 #V1 #V2 #d #e #HL12 #HV21 #_ #J #K1 #W1 #H destruct /2 width=5 by ex3_2_intro/
-]
-qed-.
-
-(* Basic_1: was: drop_gen_skip_l *)
-lemma ldrop_inv_skip1: ∀I,K1,V1,L2,s,d,e. ⇩[s, d, e] K1.ⓑ{I}V1 ≡ L2 → 0 < d →
- ∃∃K2,V2. ⇩[s, d-1, e] K1 ≡ K2 &
- ⇧[d-1, e] V2 ≡ V1 &
- L2 = K2.ⓑ{I}V2.
-/2 width=3 by ldrop_inv_skip1_aux/ qed-.
-
-lemma ldrop_inv_O1_pair2: ∀I,K,V,s,e,L1. ⇩[s, 0, e] L1 ≡ K.ⓑ{I}V →
- (e = 0 ∧ L1 = K.ⓑ{I}V) ∨
- ∃∃I1,K1,V1. ⇩[s, 0, e-1] K1 ≡ K.ⓑ{I}V & L1 = K1.ⓑ{I1}V1 & 0 < e.
-#I #K #V #s #e *
-[ #H elim (ldrop_inv_atom1 … H) -H #H destruct
-| #L1 #I1 #V1 #H
- elim (ldrop_inv_O1_pair1 … H) -H *
- [ #H1 #H2 destruct /3 width=1 by or_introl, conj/
- | /3 width=5 by ex3_3_intro, or_intror/
- ]
-]
-qed-.
-
-fact ldrop_inv_skip2_aux: ∀L1,L2,s,d,e. ⇩[s, d, e] L1 ≡ L2 → 0 < d →
- ∀I,K2,V2. L2 = K2.ⓑ{I}V2 →
- ∃∃K1,V1. ⇩[s, d-1, e] K1 ≡ K2 &
- ⇧[d-1, e] V2 ≡ V1 &
- L1 = K1.ⓑ{I}V1.
-#L1 #L2 #s #d #e * -L1 -L2 -d -e
-[ #d #e #_ #_ #J #K2 #W2 #H destruct
-| #I #L #V #H elim (lt_refl_false … H)
-| #I #L1 #L2 #V #e #_ #H elim (lt_refl_false … H)
-| #I #L1 #L2 #V1 #V2 #d #e #HL12 #HV21 #_ #J #K2 #W2 #H destruct /2 width=5 by ex3_2_intro/
-]
-qed-.
-
-(* Basic_1: was: drop_gen_skip_r *)
-lemma ldrop_inv_skip2: ∀I,L1,K2,V2,s,d,e. ⇩[s, d, e] L1 ≡ K2.ⓑ{I}V2 → 0 < d →
- ∃∃K1,V1. ⇩[s, d-1, e] K1 ≡ K2 & ⇧[d-1, e] V2 ≡ V1 &
- L1 = K1.ⓑ{I}V1.
-/2 width=3 by ldrop_inv_skip2_aux/ qed-.
-
-lemma ldrop_inv_O1_gt: ∀L,K,e,s. ⇩[s, 0, e] L ≡ K → |L| < e →
- s = Ⓣ ∧ K = ⋆.
-#L elim L -L [| #L #Z #X #IHL ] #K #e #s #H normalize in ⊢ (?%?→?); #H1e
-[ elim (ldrop_inv_atom1 … H) -H elim s -s /2 width=1 by conj/
- #_ #Hs lapply (Hs ?) // -Hs #H destruct elim (lt_zero_false … H1e)
-| elim (ldrop_inv_O1_pair1 … H) -H * #H2e #HLK destruct
- [ elim (lt_zero_false … H1e)
- | elim (IHL … HLK) -IHL -HLK /2 width=1 by lt_plus_to_minus_r, conj/
- ]
-]
-qed-.
-
-(* Basic properties *********************************************************)
-
-lemma ldrop_refl_atom_O2: ∀s,d. ⇩[s, d, O] ⋆ ≡ ⋆.
-/2 width=1 by ldrop_atom/ qed.
-
-(* Basic_1: was by definition: drop_refl *)
-lemma ldrop_refl: ∀L,d,s. ⇩[s, d, 0] L ≡ L.
-#L elim L -L //
-#L #I #V #IHL #d #s @(nat_ind_plus … d) -d /2 width=1 by ldrop_pair, ldrop_skip/
-qed.
-
-lemma ldrop_drop_lt: ∀I,L1,L2,V,s,e.
- ⇩[s, 0, e-1] L1 ≡ L2 → 0 < e → ⇩[s, 0, e] L1.ⓑ{I}V ≡ L2.
-#I #L1 #L2 #V #s #e #HL12 #He >(plus_minus_m_m e 1) /2 width=1 by ldrop_drop/
-qed.
-
-lemma ldrop_skip_lt: ∀I,L1,L2,V1,V2,s,d,e.
- ⇩[s, d-1, e] L1 ≡ L2 → ⇧[d-1, e] V2 ≡ V1 → 0 < d →
- ⇩[s, d, e] L1. ⓑ{I} V1 ≡ L2.ⓑ{I}V2.
-#I #L1 #L2 #V1 #V2 #s #d #e #HL12 #HV21 #Hd >(plus_minus_m_m d 1) /2 width=1 by ldrop_skip/
-qed.
-
-lemma ldrop_O1_le: ∀s,e,L. e ≤ |L| → ∃K. ⇩[s, 0, e] L ≡ K.
-#s #e @(nat_ind_plus … e) -e /2 width=2 by ex_intro/
-#e #IHe *
-[ #H elim (le_plus_xSy_O_false … H)
-| #L #I #V normalize #H elim (IHe L) -IHe /3 width=2 by ldrop_drop, monotonic_pred, ex_intro/
-]
-qed-.
-
-lemma ldrop_O1_lt: ∀s,L,e. e < |L| → ∃∃I,K,V. ⇩[s, 0, e] L ≡ K.ⓑ{I}V.
-#s #L elim L -L
-[ #e #H elim (lt_zero_false … H)
-| #L #I #V #IHL #e @(nat_ind_plus … e) -e /2 width=4 by ldrop_pair, ex1_3_intro/
- #e #_ normalize #H elim (IHL e) -IHL /3 width=4 by ldrop_drop, lt_plus_to_minus_r, lt_plus_to_lt_l, ex1_3_intro/
-]
-qed-.
-
-lemma ldrop_O1_pair: ∀L,K,e,s. ⇩[s, 0, e] L ≡ K → e ≤ |L| → ∀I,V.
- ∃∃J,W. ⇩[s, 0, e] L.ⓑ{I}V ≡ K.ⓑ{J}W.
-#L elim L -L [| #L #Z #X #IHL ] #K #e #s #H normalize #He #I #V
-[ elim (ldrop_inv_atom1 … H) -H #H <(le_n_O_to_eq … He) -e
- #Hs destruct /2 width=3 by ex1_2_intro/
-| elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK destruct /2 width=3 by ex1_2_intro/
- elim (IHL … HLK … Z X) -IHL -HLK
- /3 width=3 by ldrop_drop_lt, le_plus_to_minus, ex1_2_intro/
-]
-qed-.
-
-lemma ldrop_O1_ge: ∀L,e. |L| ≤ e → ⇩[Ⓣ, 0, e] L ≡ ⋆.
-#L elim L -L [ #e #_ @ldrop_atom #H destruct ]
-#L #I #V #IHL #e @(nat_ind_plus … e) -e [ #H elim (le_plus_xSy_O_false … H) ]
-normalize /4 width=1 by ldrop_drop, monotonic_pred/
-qed.
-
-lemma ldrop_O1_eq: ∀L,s. ⇩[s, 0, |L|] L ≡ ⋆.
-#L elim L -L /2 width=1 by ldrop_drop, ldrop_atom/
-qed.
-
-lemma ldrop_split: ∀L1,L2,d,e2,s. ⇩[s, d, e2] L1 ≡ L2 → ∀e1. e1 ≤ e2 →
- ∃∃L. ⇩[s, d, e2 - e1] L1 ≡ L & ⇩[s, d, e1] L ≡ L2.
-#L1 #L2 #d #e2 #s #H elim H -L1 -L2 -d -e2
-[ #d #e2 #Hs #e1 #He12 @(ex2_intro … (⋆))
- @ldrop_atom #H lapply (Hs H) -s #H destruct /2 width=1 by le_n_O_to_eq/
-| #I #L1 #V #e1 #He1 lapply (le_n_O_to_eq … He1) -He1
- #H destruct /2 width=3 by ex2_intro/
-| #I #L1 #L2 #V #e2 #HL12 #IHL12 #e1 @(nat_ind_plus … e1) -e1
- [ /3 width=3 by ldrop_drop, ex2_intro/
- | -HL12 #e1 #_ #He12 lapply (le_plus_to_le_r … He12) -He12
- #He12 elim (IHL12 … He12) -IHL12 >minus_plus_plus_l
- #L #HL1 #HL2 elim (lt_or_ge (|L1|) (e2-e1)) #H0
- [ elim (ldrop_inv_O1_gt … HL1 H0) -HL1 #H1 #H2 destruct
- elim (ldrop_inv_atom1 … HL2) -HL2 #H #_ destruct
- @(ex2_intro … (⋆)) [ @ldrop_O1_ge normalize // ]
- @ldrop_atom #H destruct
- | elim (ldrop_O1_pair … HL1 H0 I V) -HL1 -H0 /3 width=5 by ldrop_drop, ex2_intro/
- ]
- ]
-| #I #L1 #L2 #V1 #V2 #d #e2 #_ #HV21 #IHL12 #e1 #He12 elim (IHL12 … He12) -IHL12
- #L #HL1 #HL2 elim (lift_split … HV21 d e1) -HV21 /3 width=5 by ldrop_skip, ex2_intro/
-]
-qed-.
-
-lemma ldrop_FT: ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ L2 → ⇩[Ⓣ, d, e] L1 ≡ L2.
-#L1 #L2 #d #e #H elim H -L1 -L2 -d -e
-/3 width=1 by ldrop_atom, ldrop_drop, ldrop_skip/
-qed.
-
-lemma ldrop_gen: ∀L1,L2,s,d,e. ⇩[Ⓕ, d, e] L1 ≡ L2 → ⇩[s, d, e] L1 ≡ L2.
-#L1 #L2 * /2 width=1 by ldrop_FT/
-qed-.
-
-lemma ldrop_T: ∀L1,L2,s,d,e. ⇩[s, d, e] L1 ≡ L2 → ⇩[Ⓣ, d, e] L1 ≡ L2.
-#L1 #L2 * /2 width=1 by ldrop_FT/
-qed-.
-
-lemma l_liftable_LTC: ∀R. l_liftable R → l_liftable (LTC … R).
-#R #HR #K #T1 #T2 #H elim H -T2
-[ /3 width=10 by inj/
-| #T #T2 #_ #HT2 #IHT1 #L #s #d #e #HLK #U1 #HTU1 #U2 #HTU2
- elim (lift_total T d e) /4 width=12 by step/
-]
-qed-.
-
-lemma l_deliftable_sn_LTC: ∀R. l_deliftable_sn R → l_deliftable_sn (LTC … R).
-#R #HR #L #U1 #U2 #H elim H -U2
-[ #U2 #HU12 #K #s #d #e #HLK #T1 #HTU1
- elim (HR … HU12 … HLK … HTU1) -HR -L -U1 /3 width=3 by inj, ex2_intro/
-| #U #U2 #_ #HU2 #IHU1 #K #s #d #e #HLK #T1 #HTU1
- elim (IHU1 … HLK … HTU1) -IHU1 -U1 #T #HTU #HT1
- elim (HR … HU2 … HLK … HTU) -HR -L -U /3 width=5 by step, ex2_intro/
-]
-qed-.
-
-lemma dropable_sn_TC: ∀R. dropable_sn R → dropable_sn (TC … R).
-#R #HR #L1 #K1 #s #d #e #HLK1 #L2 #H elim H -L2
-[ #L2 #HL12 elim (HR … HLK1 … HL12) -HR -L1
- /3 width=3 by inj, ex2_intro/
-| #L #L2 #_ #HL2 * #K #HK1 #HLK elim (HR … HLK … HL2) -HR -L
- /3 width=3 by step, ex2_intro/
-]
-qed-.
-
-lemma dropable_dx_TC: ∀R. dropable_dx R → dropable_dx (TC … R).
-#R #HR #L1 #L2 #H elim H -L2
-[ #L2 #HL12 #K2 #s #e #HLK2 elim (HR … HL12 … HLK2) -HR -L2
- /3 width=3 by inj, ex2_intro/
-| #L #L2 #_ #HL2 #IHL1 #K2 #s #e #HLK2 elim (HR … HL2 … HLK2) -HR -L2
- #K #HLK #HK2 elim (IHL1 … HLK) -L
- /3 width=5 by step, ex2_intro/
-]
-qed-.
-
-lemma l_deliftable_sn_llstar: ∀R. l_deliftable_sn R →
- ∀l. l_deliftable_sn (llstar … R l).
-#R #HR #l #L #U1 #U2 #H @(lstar_ind_r … l U2 H) -l -U2
-[ /2 width=3 by lstar_O, ex2_intro/
-| #l #U #U2 #_ #HU2 #IHU1 #K #s #d #e #HLK #T1 #HTU1
- elim (IHU1 … HLK … HTU1) -IHU1 -U1 #T #HTU #HT1
- elim (HR … HU2 … HLK … HTU) -HR -L -U /3 width=5 by lstar_dx, ex2_intro/
-]
-qed-.
-
-(* Basic forvard lemmas *****************************************************)
-
-(* Basic_1: was: drop_S *)
-lemma ldrop_fwd_drop2: ∀L1,I2,K2,V2,s,e. ⇩[s, O, e] L1 ≡ K2. ⓑ{I2} V2 →
- ⇩[s, O, e + 1] L1 ≡ K2.
-#L1 elim L1 -L1
-[ #I2 #K2 #V2 #s #e #H lapply (ldrop_inv_atom1 … H) -H * #H destruct
-| #K1 #I1 #V1 #IHL1 #I2 #K2 #V2 #s #e #H
- elim (ldrop_inv_O1_pair1 … H) -H * #He #H
- [ -IHL1 destruct /2 width=1 by ldrop_drop/
- | @ldrop_drop >(plus_minus_m_m e 1) /2 width=3 by/
- ]
-]
-qed-.
-
-lemma ldrop_fwd_length_ge: ∀L1,L2,d,e,s. ⇩[s, d, e] L1 ≡ L2 → |L1| ≤ d → |L2| = |L1|.
-#L1 #L2 #d #e #s #H elim H -L1 -L2 -d -e // normalize
-[ #I #L1 #L2 #V #e #_ #_ #H elim (le_plus_xSy_O_false … H)
-| /4 width=2 by le_plus_to_le_r, eq_f/
-]
-qed-.
-
-lemma ldrop_fwd_length_le_le: ∀L1,L2,d,e,s. ⇩[s, d, e] L1 ≡ L2 → d ≤ |L1| → e ≤ |L1| - d → |L2| = |L1| - e.
-#L1 #L2 #d #e #s #H elim H -L1 -L2 -d -e // normalize
-[ /3 width=2 by le_plus_to_le_r/
-| #I #L1 #L2 #V1 #V2 #d #e #_ #_ #IHL12 >minus_plus_plus_l
- #Hd #He lapply (le_plus_to_le_r … Hd) -Hd
- #Hd >IHL12 // -L2 >plus_minus /2 width=3 by transitive_le/
-]
-qed-.
-
-lemma ldrop_fwd_length_le_ge: ∀L1,L2,d,e,s. ⇩[s, d, e] L1 ≡ L2 → d ≤ |L1| → |L1| - d ≤ e → |L2| = d.
-#L1 #L2 #d #e #s #H elim H -L1 -L2 -d -e normalize
-[ /2 width=1 by le_n_O_to_eq/
-| #I #L #V #_ <minus_n_O #H elim (le_plus_xSy_O_false … H)
-| /3 width=2 by le_plus_to_le_r/
-| /4 width=2 by le_plus_to_le_r, eq_f/
-]
-qed-.
-
-lemma ldrop_fwd_length: ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ L2 → |L1| = |L2| + e.
-#L1 #L2 #d #e #H elim H -L1 -L2 -d -e // normalize /2 width=1 by/
-qed-.
-
-lemma ldrop_fwd_length_minus2: ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ L2 → |L2| = |L1| - e.
-#L1 #L2 #d #e #H lapply (ldrop_fwd_length … H) -H /2 width=1 by plus_minus, le_n/
-qed-.
-
-lemma ldrop_fwd_length_minus4: ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ L2 → e = |L1| - |L2|.
-#L1 #L2 #d #e #H lapply (ldrop_fwd_length … H) -H //
-qed-.
-
-lemma ldrop_fwd_length_le2: ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ L2 → e ≤ |L1|.
-#L1 #L2 #d #e #H lapply (ldrop_fwd_length … H) -H //
-qed-.
-
-lemma ldrop_fwd_length_le4: ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ L2 → |L2| ≤ |L1|.
-#L1 #L2 #d #e #H lapply (ldrop_fwd_length … H) -H //
-qed-.
-
-lemma ldrop_fwd_length_lt2: ∀L1,I2,K2,V2,d,e.
- ⇩[Ⓕ, d, e] L1 ≡ K2. ⓑ{I2} V2 → e < |L1|.
-#L1 #I2 #K2 #V2 #d #e #H
-lapply (ldrop_fwd_length … H) normalize in ⊢ (%→?); -I2 -V2 //
-qed-.
-
-lemma ldrop_fwd_length_lt4: ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ L2 → 0 < e → |L2| < |L1|.
-#L1 #L2 #d #e #H lapply (ldrop_fwd_length … H) -H /2 width=1 by lt_minus_to_plus_r/
-qed-.
-
-lemma ldrop_fwd_length_eq1: ∀L1,L2,K1,K2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 →
- |L1| = |L2| → |K1| = |K2|.
-#L1 #L2 #K1 #K2 #d #e #HLK1 #HLK2 #HL12
-lapply (ldrop_fwd_length … HLK1) -HLK1
-lapply (ldrop_fwd_length … HLK2) -HLK2
-/2 width=2 by injective_plus_r/
-qed-.
-
-lemma ldrop_fwd_length_eq2: ∀L1,L2,K1,K2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 →
- |K1| = |K2| → |L1| = |L2|.
-#L1 #L2 #K1 #K2 #d #e #HLK1 #HLK2 #HL12
-lapply (ldrop_fwd_length … HLK1) -HLK1
-lapply (ldrop_fwd_length … HLK2) -HLK2 //
-qed-.
-
-lemma ldrop_fwd_lw: ∀L1,L2,s,d,e. ⇩[s, d, e] L1 ≡ L2 → ♯{L2} ≤ ♯{L1}.
-#L1 #L2 #s #d #e #H elim H -L1 -L2 -d -e // normalize
-[ /2 width=3 by transitive_le/
-| #I #L1 #L2 #V1 #V2 #d #e #_ #HV21 #IHL12
- >(lift_fwd_tw … HV21) -HV21 /2 width=1 by monotonic_le_plus_l/
-]
-qed-.
-
-lemma ldrop_fwd_lw_lt: ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ L2 → 0 < e → ♯{L2} < ♯{L1}.
-#L1 #L2 #d #e #H elim H -L1 -L2 -d -e
-[ #d #e #H >H -H //
-| #I #L #V #H elim (lt_refl_false … H)
-| #I #L1 #L2 #V #e #HL12 #_ #_
- lapply (ldrop_fwd_lw … HL12) -HL12 #HL12
- @(le_to_lt_to_lt … HL12) -HL12 //
-| #I #L1 #L2 #V1 #V2 #d #e #_ #HV21 #IHL12 #H normalize in ⊢ (?%%); -I
- >(lift_fwd_tw … HV21) -V2 /3 by lt_minus_to_plus/
-]
-qed-.
-
-lemma ldrop_fwd_rfw: ∀I,L,K,V,i. ⇩[i] L ≡ K.ⓑ{I}V → ∀T. ♯{K, V} < ♯{L, T}.
-#I #L #K #V #i #HLK lapply (ldrop_fwd_lw … HLK) -HLK
-normalize in ⊢ (%→?→?%%); /3 width=3 by le_to_lt_to_lt/
-qed-.
-
-(* Advanced inversion lemmas ************************************************)
-
-fact ldrop_inv_O2_aux: ∀L1,L2,s,d,e. ⇩[s, d, e] L1 ≡ L2 → e = 0 → L1 = L2.
-#L1 #L2 #s #d #e #H elim H -L1 -L2 -d -e
-[ //
-| //
-| #I #L1 #L2 #V #e #_ #_ >commutative_plus normalize #H destruct
-| #I #L1 #L2 #V1 #V2 #d #e #_ #HV21 #IHL12 #H
- >(IHL12 H) -L1 >(lift_inv_O2_aux … HV21 … H) -V2 -d -e //
-]
-qed-.
-
-(* Basic_1: was: drop_gen_refl *)
-lemma ldrop_inv_O2: ∀L1,L2,s,d. ⇩[s, d, 0] L1 ≡ L2 → L1 = L2.
-/2 width=5 by ldrop_inv_O2_aux/ qed-.
-
-lemma ldrop_inv_length_eq: ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ L2 → |L1| = |L2| → e = 0.
-#L1 #L2 #d #e #H #HL12 lapply (ldrop_fwd_length_minus4 … H) //
-qed-.
-
-lemma ldrop_inv_refl: ∀L,d,e. ⇩[Ⓕ, d, e] L ≡ L → e = 0.
-/2 width=5 by ldrop_inv_length_eq/ qed-.
-
-fact ldrop_inv_FT_aux: ∀L1,L2,s,d,e. ⇩[s, d, e] L1 ≡ L2 →
- ∀I,K,V. L2 = K.ⓑ{I}V → s = Ⓣ → d = 0 →
- ⇩[Ⓕ, d, e] L1 ≡ K.ⓑ{I}V.
-#L1 #L2 #s #d #e #H elim H -L1 -L2 -d -e
-[ #d #e #_ #J #K #W #H destruct
-| #I #L #V #J #K #W #H destruct //
-| #I #L1 #L2 #V #e #_ #IHL12 #J #K #W #H1 #H2 destruct
- /3 width=1 by ldrop_drop/
-| #I #L1 #L2 #V1 #V2 #d #e #_ #_ #_ #J #K #W #_ #_
- <plus_n_Sm #H destruct
-]
-qed-.
-
-lemma ldrop_inv_FT: ∀I,L,K,V,e. ⇩[Ⓣ, 0, e] L ≡ K.ⓑ{I}V → ⇩[e] L ≡ K.ⓑ{I}V.
-/2 width=5 by ldrop_inv_FT_aux/ qed.
-
-lemma ldrop_inv_gen: ∀I,L,K,V,s,e. ⇩[s, 0, e] L ≡ K.ⓑ{I}V → ⇩[e] L ≡ K.ⓑ{I}V.
-#I #L #K #V * /2 width=1 by ldrop_inv_FT/
-qed-.
-
-lemma ldrop_inv_T: ∀I,L,K,V,s,e. ⇩[Ⓣ, 0, e] L ≡ K.ⓑ{I}V → ⇩[s, 0, e] L ≡ K.ⓑ{I}V.
-#I #L #K #V * /2 width=1 by ldrop_inv_FT/
-qed-.
-
-(* Basic_1: removed theorems 50:
- drop_ctail drop_skip_flat
- cimp_flat_sx cimp_flat_dx cimp_bind cimp_getl_conf
- drop_clear drop_clear_O drop_clear_S
- clear_gen_sort clear_gen_bind clear_gen_flat clear_gen_flat_r
- clear_gen_all clear_clear clear_mono clear_trans clear_ctail clear_cle
- getl_ctail_clen getl_gen_tail clear_getl_trans getl_clear_trans
- getl_clear_bind getl_clear_conf getl_dec getl_drop getl_drop_conf_lt
- getl_drop_conf_ge getl_conf_ge_drop getl_drop_conf_rev
- drop_getl_trans_lt drop_getl_trans_le drop_getl_trans_ge
- getl_drop_trans getl_flt getl_gen_all getl_gen_sort getl_gen_O
- getl_gen_S getl_gen_2 getl_gen_flat getl_gen_bind getl_conf_le
- getl_trans getl_refl getl_head getl_flat getl_ctail getl_mono
-*)
+++ /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/grammar/lenv_append.ma".
-include "basic_2/substitution/ldrop.ma".
-
-(* DROPPING *****************************************************************)
-
-(* Properties on append for local environments ******************************)
-
-fact ldrop_O1_append_sn_le_aux: ∀L1,L2,s,d,e. ⇩[s, d, e] L1 ≡ L2 →
- d = 0 → e ≤ |L1| →
- ∀L. ⇩[s, 0, e] L @@ L1 ≡ L @@ L2.
-#L1 #L2 #s #d #e #H elim H -L1 -L2 -d -e normalize
-[2,3,4: /4 width=1 by ldrop_skip_lt, ldrop_drop, arith_b1, lt_minus_to_plus_r, monotonic_pred/ ]
-#d #e #_ #_ #H <(le_n_O_to_eq … H) -H //
-qed-.
-
-lemma ldrop_O1_append_sn_le: ∀L1,L2,s,e. ⇩[s, 0, e] L1 ≡ L2 → e ≤ |L1| →
- ∀L. ⇩[s, 0, e] L @@ L1 ≡ L @@ L2.
-/2 width=3 by ldrop_O1_append_sn_le_aux/ qed.
-
-(* Inversion lemmas on append for local environments ************************)
-
-lemma ldrop_O1_inv_append1_ge: ∀K,L1,L2,s,e. ⇩[s, 0, e] L1 @@ L2 ≡ K →
- |L2| ≤ e → ⇩[s, 0, e - |L2|] L1 ≡ K.
-#K #L1 #L2 elim L2 -L2 normalize //
-#L2 #I #V #IHL2 #s #e #H #H1e
-elim (ldrop_inv_O1_pair1 … H) -H * #H2e #HL12 destruct
-[ lapply (le_n_O_to_eq … H1e) -H1e -IHL2
- >commutative_plus normalize #H destruct
-| <minus_plus >minus_minus_comm /3 width=1 by monotonic_pred/
-]
-qed-.
-
-lemma ldrop_O1_inv_append1_le: ∀K,L1,L2,s,e. ⇩[s, 0, e] L1 @@ L2 ≡ K → e ≤ |L2| →
- ∀K2. ⇩[s, 0, e] L2 ≡ K2 → K = L1 @@ K2.
-#K #L1 #L2 elim L2 -L2 normalize
-[ #s #e #H1 #H2 #K2 #H3 lapply (le_n_O_to_eq … H2) -H2
- #H2 elim (ldrop_inv_atom1 … H3) -H3 #H3 #_ destruct
- >(ldrop_inv_O2 … H1) -H1 //
-| #L2 #I #V #IHL2 #s #e @(nat_ind_plus … e) -e [ -IHL2 ]
- [ #H1 #_ #K2 #H2
- lapply (ldrop_inv_O2 … H1) -H1 #H1
- lapply (ldrop_inv_O2 … H2) -H2 #H2 destruct //
- | /4 width=7 by ldrop_inv_drop1, le_plus_to_le_r/
- ]
-]
-qed-.
+++ /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/substitution/lift_lift.ma".
-include "basic_2/substitution/ldrop.ma".
-
-(* BASIC SLICING FOR LOCAL ENVIRONMENTS *************************************)
-
-(* Main properties **********************************************************)
-
-(* Basic_1: was: drop_mono *)
-theorem ldrop_mono: ∀L,L1,s1,d,e. ⇩[s1, d, e] L ≡ L1 →
- ∀L2,s2. ⇩[s2, d, e] L ≡ L2 → L1 = L2.
-#L #L1 #s1 #d #e #H elim H -L -L1 -d -e
-[ #d #e #He #L2 #s2 #H elim (ldrop_inv_atom1 … H) -H //
-| #I #K #V #L2 #s2 #HL12 <(ldrop_inv_O2 … HL12) -L2 //
-| #I #L #K #V #e #_ #IHLK #L2 #s2 #H
- lapply (ldrop_inv_drop1 … H) -H /2 width=2 by/
-| #I #L #K1 #T #V1 #d #e #_ #HVT1 #IHLK1 #X #s2 #H
- elim (ldrop_inv_skip1 … H) -H // <minus_plus_m_m #K2 #V2 #HLK2 #HVT2 #H destruct
- >(lift_inj … HVT1 … HVT2) -HVT1 -HVT2
- >(IHLK1 … HLK2) -IHLK1 -HLK2 //
-]
-qed-.
-
-(* Basic_1: was: drop_conf_ge *)
-theorem ldrop_conf_ge: ∀L,L1,s1,d1,e1. ⇩[s1, d1, e1] L ≡ L1 →
- ∀L2,s2,e2. ⇩[s2, 0, e2] L ≡ L2 → d1 + e1 ≤ e2 →
- ⇩[s2, 0, e2 - e1] L1 ≡ L2.
-#L #L1 #s1 #d1 #e1 #H elim H -L -L1 -d1 -e1 //
-[ #d #e #_ #L2 #s2 #e2 #H #_ elim (ldrop_inv_atom1 … H) -H
- #H #He destruct
- @ldrop_atom #H >He // (**) (* explicit constructor *)
-| #I #L #K #V #e #_ #IHLK #L2 #s2 #e2 #H #He2
- lapply (ldrop_inv_drop1_lt … H ?) -H /2 width=2 by ltn_to_ltO/ #HL2
- <minus_plus >minus_minus_comm /3 width=1 by monotonic_pred/
-| #I #L #K #V1 #V2 #d #e #_ #_ #IHLK #L2 #s2 #e2 #H #Hdee2
- lapply (transitive_le 1 … Hdee2) // #He2
- lapply (ldrop_inv_drop1_lt … H ?) -H // -He2 #HL2
- lapply (transitive_le (1+e) … Hdee2) // #Hee2
- @ldrop_drop_lt >minus_minus_comm /3 width=1 by lt_minus_to_plus_r, monotonic_le_minus_r, monotonic_pred/ (**) (* explicit constructor *)
-]
-qed.
-
-(* Note: apparently this was missing in basic_1 *)
-theorem ldrop_conf_be: ∀L0,L1,s1,d1,e1. ⇩[s1, d1, e1] L0 ≡ L1 →
- ∀L2,e2. ⇩[e2] L0 ≡ L2 → d1 ≤ e2 → e2 ≤ d1 + e1 →
- ∃∃L. ⇩[s1, 0, d1 + e1 - e2] L2 ≡ L & ⇩[d1] L1 ≡ L.
-#L0 #L1 #s1 #d1 #e1 #H elim H -L0 -L1 -d1 -e1
-[ #d1 #e1 #He1 #L2 #e2 #H #Hd1 #_ elim (ldrop_inv_atom1 … H) -H #H #He2 destruct
- >(He2 ?) in Hd1; // -He2 #Hd1 <(le_n_O_to_eq … Hd1) -d1
- /4 width=3 by ldrop_atom, ex2_intro/
-| normalize #I #L #V #L2 #e2 #HL2 #_ #He2
- lapply (le_n_O_to_eq … He2) -He2 #H destruct
- lapply (ldrop_inv_O2 … HL2) -HL2 #H destruct /2 width=3 by ldrop_pair, ex2_intro/
-| normalize #I #L0 #K0 #V1 #e1 #HLK0 #IHLK0 #L2 #e2 #H #_ #He21
- lapply (ldrop_inv_O1_pair1 … H) -H * * #He2 #HL20
- [ -IHLK0 -He21 destruct <minus_n_O /3 width=3 by ldrop_drop, ex2_intro/
- | -HLK0 <minus_le_minus_minus_comm //
- elim (IHLK0 … HL20) -L0 /2 width=3 by monotonic_pred, ex2_intro/
- ]
-| #I #L0 #K0 #V0 #V1 #d1 #e1 >plus_plus_comm_23 #_ #_ #IHLK0 #L2 #e2 #H #Hd1e2 #He2de1
- elim (le_inv_plus_l … Hd1e2) #_ #He2
- <minus_le_minus_minus_comm //
- lapply (ldrop_inv_drop1_lt … H ?) -H // #HL02
- elim (IHLK0 … HL02) -L0 /3 width=3 by ldrop_drop, monotonic_pred, ex2_intro/
-]
-qed-.
-
-(* Note: apparently this was missing in basic_1 *)
-theorem ldrop_conf_le: ∀L0,L1,s1,d1,e1. ⇩[s1, d1, e1] L0 ≡ L1 →
- ∀L2,s2,e2. ⇩[s2, 0, e2] L0 ≡ L2 → e2 ≤ d1 →
- ∃∃L. ⇩[s2, 0, e2] L1 ≡ L & ⇩[s1, d1 - e2, e1] L2 ≡ L.
-#L0 #L1 #s1 #d1 #e1 #H elim H -L0 -L1 -d1 -e1
-[ #d1 #e1 #He1 #L2 #s2 #e2 #H elim (ldrop_inv_atom1 … H) -H
- #H #He2 #_ destruct /4 width=3 by ldrop_atom, ex2_intro/
-| #I #K0 #V0 #L2 #s2 #e2 #H1 #H2
- lapply (le_n_O_to_eq … H2) -H2 #H destruct
- lapply (ldrop_inv_pair1 … H1) -H1 #H destruct /2 width=3 by ldrop_pair, ex2_intro/
-| #I #K0 #K1 #V0 #e1 #HK01 #_ #L2 #s2 #e2 #H1 #H2
- lapply (le_n_O_to_eq … H2) -H2 #H destruct
- lapply (ldrop_inv_pair1 … H1) -H1 #H destruct /3 width=3 by ldrop_drop, ex2_intro/
-| #I #K0 #K1 #V0 #V1 #d1 #e1 #HK01 #HV10 #IHK01 #L2 #s2 #e2 #H #He2d1
- elim (ldrop_inv_O1_pair1 … H) -H *
- [ -IHK01 -He2d1 #H1 #H2 destruct /3 width=5 by ldrop_pair, ldrop_skip, ex2_intro/
- | -HK01 -HV10 #He2 #HK0L2
- elim (IHK01 … HK0L2) -IHK01 -HK0L2 /2 width=1 by monotonic_pred/
- >minus_le_minus_minus_comm /3 width=3 by ldrop_drop_lt, ex2_intro/
- ]
-]
-qed-.
-
-(* Note: with "s2", the conclusion parameter is "s1 ∨ s2" *)
-(* Basic_1: was: drop_trans_ge *)
-theorem ldrop_trans_ge: ∀L1,L,s1,d1,e1. ⇩[s1, d1, e1] L1 ≡ L →
- ∀L2,e2. ⇩[e2] L ≡ L2 → d1 ≤ e2 → ⇩[s1, 0, e1 + e2] L1 ≡ L2.
-#L1 #L #s1 #d1 #e1 #H elim H -L1 -L -d1 -e1
-[ #d1 #e1 #He1 #L2 #e2 #H #_ elim (ldrop_inv_atom1 … H) -H
- #H #He2 destruct /4 width=1 by ldrop_atom, eq_f2/
-| /2 width=1 by ldrop_gen/
-| /3 width=1 by ldrop_drop/
-| #I #L1 #L2 #V1 #V2 #d #e #_ #_ #IHL12 #L #e2 #H #Hde2
- lapply (lt_to_le_to_lt 0 … Hde2) // #He2
- lapply (lt_to_le_to_lt … (e + e2) He2 ?) // #Hee2
- lapply (ldrop_inv_drop1_lt … H ?) -H // #HL2
- @ldrop_drop_lt // >le_plus_minus /3 width=1 by monotonic_pred/
-]
-qed.
-
-(* Basic_1: was: drop_trans_le *)
-theorem ldrop_trans_le: ∀L1,L,s1,d1,e1. ⇩[s1, d1, e1] L1 ≡ L →
- ∀L2,s2,e2. ⇩[s2, 0, e2] L ≡ L2 → e2 ≤ d1 →
- ∃∃L0. ⇩[s2, 0, e2] L1 ≡ L0 & ⇩[s1, d1 - e2, e1] L0 ≡ L2.
-#L1 #L #s1 #d1 #e1 #H elim H -L1 -L -d1 -e1
-[ #d1 #e1 #He1 #L2 #s2 #e2 #H #_ elim (ldrop_inv_atom1 … H) -H
- #H #He2 destruct /4 width=3 by ldrop_atom, ex2_intro/
-| #I #K #V #L2 #s2 #e2 #HL2 #H lapply (le_n_O_to_eq … H) -H
- #H destruct /2 width=3 by ldrop_pair, ex2_intro/
-| #I #L1 #L2 #V #e #_ #IHL12 #L #s2 #e2 #HL2 #H lapply (le_n_O_to_eq … H) -H
- #H destruct elim (IHL12 … HL2) -IHL12 -HL2 //
- #L0 #H #HL0 lapply (ldrop_inv_O2 … H) -H #H destruct
- /3 width=5 by ldrop_pair, ldrop_drop, ex2_intro/
-| #I #L1 #L2 #V1 #V2 #d #e #HL12 #HV12 #IHL12 #L #s2 #e2 #H #He2d
- elim (ldrop_inv_O1_pair1 … H) -H *
- [ -He2d -IHL12 #H1 #H2 destruct /3 width=5 by ldrop_pair, ldrop_skip, ex2_intro/
- | -HL12 -HV12 #He2 #HL2
- elim (IHL12 … HL2) -L2 [ >minus_le_minus_minus_comm // /3 width=3 by ldrop_drop_lt, ex2_intro/ | /2 width=1 by monotonic_pred/ ]
- ]
-]
-qed-.
-
-(* Advanced properties ******************************************************)
-
-lemma l_liftable_llstar: ∀R. l_liftable R → ∀l. l_liftable (llstar … R l).
-#R #HR #l #K #T1 #T2 #H @(lstar_ind_r … l T2 H) -l -T2
-[ #L #s #d #e #_ #U1 #HTU1 #U2 #HTU2 -HR -K
- >(lift_mono … HTU2 … HTU1) -T1 -U2 -d -e //
-| #l #T #T2 #_ #HT2 #IHT1 #L #s #d #e #HLK #U1 #HTU1 #U2 #HTU2
- elim (lift_total T d e) /3 width=12 by lstar_dx/
-]
-qed-.
-
-(* Basic_1: was: drop_conf_lt *)
-lemma ldrop_conf_lt: ∀L,L1,s1,d1,e1. ⇩[s1, d1, e1] L ≡ L1 →
- ∀I,K2,V2,s2,e2. ⇩[s2, 0, e2] L ≡ K2.ⓑ{I}V2 →
- e2 < d1 → let d ≝ d1 - e2 - 1 in
- ∃∃K1,V1. ⇩[s2, 0, e2] L1 ≡ K1.ⓑ{I}V1 &
- ⇩[s1, d, e1] K2 ≡ K1 & ⇧[d, e1] V1 ≡ V2.
-#L #L1 #s1 #d1 #e1 #H1 #I #K2 #V2 #s2 #e2 #H2 #He2d1
-elim (ldrop_conf_le … H1 … H2) -L /2 width=2 by lt_to_le/ #K #HL1K #HK2
-elim (ldrop_inv_skip1 … HK2) -HK2 /2 width=1 by lt_plus_to_minus_r/
-#K1 #V1 #HK21 #HV12 #H destruct /2 width=5 by ex3_2_intro/
-qed-.
-
-(* Note: apparently this was missing in basic_1 *)
-lemma ldrop_trans_lt: ∀L1,L,s1,d1,e1. ⇩[s1, d1, e1] L1 ≡ L →
- ∀I,L2,V2,s2,e2. ⇩[s2, 0, e2] L ≡ L2.ⓑ{I}V2 →
- e2 < d1 → let d ≝ d1 - e2 - 1 in
- ∃∃L0,V0. ⇩[s2, 0, e2] L1 ≡ L0.ⓑ{I}V0 &
- ⇩[s1, d, e1] L0 ≡ L2 & ⇧[d, e1] V2 ≡ V0.
-#L1 #L #s1 #d1 #e1 #HL1 #I #L2 #V2 #s2 #e2 #HL2 #Hd21
-elim (ldrop_trans_le … HL1 … HL2) -L /2 width=1 by lt_to_le/ #L0 #HL10 #HL02
-elim (ldrop_inv_skip2 … HL02) -HL02 /2 width=1 by lt_plus_to_minus_r/ #L #V1 #HL2 #HV21 #H destruct /2 width=5 by ex3_2_intro/
-qed-.
-
-lemma ldrop_trans_ge_comm: ∀L1,L,L2,s1,d1,e1,e2.
- ⇩[s1, d1, e1] L1 ≡ L → ⇩[e2] L ≡ L2 → d1 ≤ e2 →
- ⇩[s1, 0, e2 + e1] L1 ≡ L2.
-#L1 #L #L2 #s1 #d1 #e1 #e2
->commutative_plus /2 width=5 by ldrop_trans_ge/
-qed.
-
-lemma ldrop_conf_div: ∀I1,L,K,V1,e1. ⇩[e1] L ≡ K.ⓑ{I1}V1 →
- ∀I2,V2,e2. ⇩[e2] L ≡ K.ⓑ{I2}V2 →
- ∧∧ e1 = e2 & I1 = I2 & V1 = V2.
-#I1 #L #K #V1 #e1 #HLK1 #I2 #V2 #e2 #HLK2
-elim (le_or_ge e1 e2) #He
-[ lapply (ldrop_conf_ge … HLK1 … HLK2 ?)
-| lapply (ldrop_conf_ge … HLK2 … HLK1 ?)
-] -HLK1 -HLK2 // #HK
-lapply (ldrop_fwd_length_minus2 … HK) #H
-elim (discr_minus_x_xy … H) -H
-[1,3: normalize <plus_n_Sm #H destruct ]
-#H >H in HK; #HK
-lapply (ldrop_inv_O2 … HK) -HK #H destruct
-lapply (inv_eq_minus_O … H) -H /3 width=1 by le_to_le_to_eq, and3_intro/
-qed-.
-
-(* Advanced forward lemmas **************************************************)
-
-lemma ldrop_fwd_be: ∀L,K,s,d,e,i. ⇩[s, d, e] L ≡ K → |K| ≤ i → i < d → |L| ≤ i.
-#L #K #s #d #e #i #HLK #HK #Hd elim (lt_or_ge i (|L|)) //
-#HL elim (ldrop_O1_lt (Ⓕ) … HL) #I #K0 #V #HLK0 -HL
-elim (ldrop_conf_lt … HLK … HLK0) // -HLK -HLK0 -Hd
-#K1 #V1 #HK1 #_ #_ lapply (ldrop_fwd_length_lt2 … HK1) -I -K1 -V1
-#H elim (lt_refl_false i) /2 width=3 by lt_to_le_to_lt/
-qed-.
+++ /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/grammar/leq_leq.ma".
-include "basic_2/substitution/ldrop.ma".
-
-(* BASIC SLICING FOR LOCAL ENVIRONMENTS *************************************)
-
-definition dedropable_sn: predicate (relation lenv) ≝
- λR. ∀L1,K1,s,d,e. ⇩[s, d, e] L1 ≡ K1 → ∀K2. R K1 K2 →
- ∃∃L2. R L1 L2 & ⇩[s, d, e] L2 ≡ K2 & L1 ≃[d, e] L2.
-
-(* Properties on equivalence ************************************************)
-
-lemma leq_ldrop_trans_be: ∀L1,L2,d,e. L1 ≃[d, e] L2 →
- ∀I,K2,W,s,i. ⇩[s, 0, i] L2 ≡ K2.ⓑ{I}W →
- d ≤ i → i < d + e →
- ∃∃K1. K1 ≃[0, ⫰(d+e-i)] K2 & ⇩[s, 0, i] L1 ≡ K1.ⓑ{I}W.
-#L1 #L2 #d #e #H elim H -L1 -L2 -d -e
-[ #d #e #J #K2 #W #s #i #H
- elim (ldrop_inv_atom1 … H) -H #H destruct
-| #I1 #I2 #L1 #L2 #V1 #V2 #_ #_ #J #K2 #W #s #i #_ #_ #H
- elim (ylt_yle_false … H) //
-| #I #L1 #L2 #V #e #HL12 #IHL12 #J #K2 #W #s #i #H #_ >yplus_O1
- elim (ldrop_inv_O1_pair1 … H) -H * #Hi #HLK1 [ -IHL12 | -HL12 ]
- [ #_ destruct >ypred_succ
- /2 width=3 by ldrop_pair, ex2_intro/
- | lapply (ylt_inv_O1 i ?) /2 width=1 by ylt_inj/
- #H <H -H #H lapply (ylt_inv_succ … H) -H
- #Hie elim (IHL12 … HLK1) -IHL12 -HLK1 // -Hie
- >yminus_succ <yminus_inj /3 width=3 by ldrop_drop_lt, ex2_intro/
- ]
-| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #IHL12 #J #K2 #W #s #i #HLK2 #Hdi
- elim (yle_inv_succ1 … Hdi) -Hdi
- #Hdi #Hi <Hi >yplus_succ1 #H lapply (ylt_inv_succ … H) -H
- #Hide lapply (ldrop_inv_drop1_lt … HLK2 ?) -HLK2 /2 width=1 by ylt_O/
- #HLK1 elim (IHL12 … HLK1) -IHL12 -HLK1 <yminus_inj >yminus_SO2
- /4 width=3 by ylt_O, ldrop_drop_lt, ex2_intro/
-]
-qed-.
-
-lemma leq_ldrop_conf_be: ∀L1,L2,d,e. L1 ≃[d, e] L2 →
- ∀I,K1,W,s,i. ⇩[s, 0, i] L1 ≡ K1.ⓑ{I}W →
- d ≤ i → i < d + e →
- ∃∃K2. K1 ≃[0, ⫰(d+e-i)] K2 & ⇩[s, 0, i] L2 ≡ K2.ⓑ{I}W.
-#L1 #L2 #d #e #HL12 #I #K1 #W #s #i #HLK1 #Hdi #Hide
-elim (leq_ldrop_trans_be … (leq_sym … HL12) … HLK1) // -L1 -Hdi -Hide
-/3 width=3 by leq_sym, ex2_intro/
-qed-.
-
-lemma ldrop_O1_ex: ∀K2,i,L1. |L1| = |K2| + i →
- ∃∃L2. L1 ≃[0, i] L2 & ⇩[i] L2 ≡ K2.
-#K2 #i @(nat_ind_plus … i) -i
-[ /3 width=3 by leq_O2, ex2_intro/
-| #i #IHi #Y #Hi elim (ldrop_O1_lt (Ⓕ) Y 0) //
- #I #L1 #V #H lapply (ldrop_inv_O2 … H) -H #H destruct
- normalize in Hi; elim (IHi L1) -IHi
- /3 width=5 by ldrop_drop, leq_pair, injective_plus_l, ex2_intro/
-]
-qed-.
-
-lemma dedropable_sn_TC: ∀R. dedropable_sn R → dedropable_sn (TC … R).
-#R #HR #L1 #K1 #s #d #e #HLK1 #K2 #H elim H -K2
-[ #K2 #HK12 elim (HR … HLK1 … HK12) -HR -K1
- /3 width=4 by inj, ex3_intro/
-| #K #K2 #_ #HK2 * #L #H1L1 #HLK #H2L1 elim (HR … HLK … HK2) -HR -K
- /3 width=6 by leq_trans, step, ex3_intro/
-]
-qed-.
-
-(* Inversion lemmas on equivalence ******************************************)
-
-lemma ldrop_O1_inj: ∀i,L1,L2,K. ⇩[i] L1 ≡ K → ⇩[i] L2 ≡ K → L1 ≃[i, ∞] L2.
-#i @(nat_ind_plus … i) -i
-[ #L1 #L2 #K #H <(ldrop_inv_O2 … H) -K #H <(ldrop_inv_O2 … H) -L1 //
-| #i #IHi * [2: #L1 #I1 #V1 ] * [2,4: #L2 #I2 #V2 ] #K #HLK1 #HLK2 //
- lapply (ldrop_fwd_length … HLK1)
- <(ldrop_fwd_length … HLK2) [ /4 width=5 by ldrop_inv_drop1, leq_succ/ ]
- normalize <plus_n_Sm #H destruct
-]
-qed-.
(* *)
(**************************************************************************)
-include "basic_2/substitution/ldrop.ma".
+include "basic_2/substitution/drop.ma".
include "basic_2/substitution/lpx_sn.ma".
(* SN POINTWISE EXTENSION OF A CONTEXT-SENSITIVE REALTION FOR TERMS *********)
#H #HV12 destruct @(ex3_2_intro … K2 V2) // -HV12
@conj // -HK12
#J1 #J2 #L1 #L2 #W1 #W2 #i #HKL1 #HKL2 elim (IH J1 J2 L1 L2 W1 W2 (i+1)) -IH
-/2 width=1 by ldrop_drop, conj/
+/2 width=1 by drop_drop, conj/
qed-.
lemma lpx_sn_alt_inv_atom2: ∀R,L1. lpx_sn_alt R L1 (⋆) → L1 = ⋆.
#H #HV12 destruct @(ex3_2_intro … K1 V1) // -HV12
@conj // -HK12
#J1 #J2 #L1 #L2 #W1 #W2 #i #HKL1 #HKL2 elim (IH J1 J2 L1 L2 W1 W2 (i+1)) -IH
-/2 width=1 by ldrop_drop, conj/
+/2 width=1 by drop_drop, conj/
qed-.
(* Basic properties *********************************************************)
lemma lpx_sn_alt_atom: ∀R. lpx_sn_alt R (⋆) (⋆).
#R @conj //
-#I1 #I2 #K1 #K2 #V1 #V2 #i #HLK1 elim (ldrop_inv_atom1 … HLK1) -HLK1
+#I1 #I2 #K1 #K2 #V1 #V2 #i #HLK1 elim (drop_inv_atom1 … HLK1) -HLK1
#H destruct
qed.
#HL12 #IH @conj normalize //
#I1 #I2 #K1 #K2 #W1 #W2 #i @(nat_ind_plus … i) -i
[ #HLK1 #HLK2
- lapply (ldrop_inv_O2 … HLK1) -HLK1 #H destruct
- lapply (ldrop_inv_O2 … HLK2) -HLK2 #H destruct
+ lapply (drop_inv_O2 … HLK1) -HLK1 #H destruct
+ lapply (drop_inv_O2 … HLK2) -HLK2 #H destruct
/2 width=1 by conj/
-| -HL12 -HV12 /3 width=6 by ldrop_inv_drop1/
+| -HL12 -HV12 /3 width=6 by drop_inv_drop1/
]
qed.
--- /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/substitution/drop_leq.ma".
+include "basic_2/substitution/lpx_sn.ma".
+
+(* SN POINTWISE EXTENSION OF A CONTEXT-SENSITIVE REALTION FOR TERMS *********)
+
+(* Properies on dropping ****************************************************)
+
+lemma lpx_sn_drop_conf: ∀R,L1,L2. lpx_sn R L1 L2 →
+ ∀I,K1,V1,i. ⇩[i] L1 ≡ K1.ⓑ{I}V1 →
+ ∃∃K2,V2. ⇩[i] L2 ≡ K2.ⓑ{I}V2 & lpx_sn R K1 K2 & R K1 V1 V2.
+#R #L1 #L2 #H elim H -L1 -L2
+[ #I0 #K0 #V0 #i #H elim (drop_inv_atom1 … H) -H #H destruct
+| #I #K1 #K2 #V1 #V2 #HK12 #HV12 #IHK12 #I0 #K0 #V0 #i #H elim (drop_inv_O1_pair1 … H) * -H
+ [ -IHK12 #H1 #H2 destruct /3 width=5 by drop_pair, ex3_2_intro/
+ | -HK12 -HV12 #Hi #HK10 elim (IHK12 … HK10) -IHK12 -HK10
+ /3 width=5 by drop_drop_lt, ex3_2_intro/
+ ]
+]
+qed-.
+
+lemma lpx_sn_drop_trans: ∀R,L1,L2. lpx_sn R L1 L2 →
+ ∀I,K2,V2,i. ⇩[i] L2 ≡ K2.ⓑ{I}V2 →
+ ∃∃K1,V1. ⇩[i] L1 ≡ K1.ⓑ{I}V1 & lpx_sn R K1 K2 & R K1 V1 V2.
+#R #L1 #L2 #H elim H -L1 -L2
+[ #I0 #K0 #V0 #i #H elim (drop_inv_atom1 … H) -H #H destruct
+| #I #K1 #K2 #V1 #V2 #HK12 #HV12 #IHK12 #I0 #K0 #V0 #i #H elim (drop_inv_O1_pair1 … H) * -H
+ [ -IHK12 #H1 #H2 destruct /3 width=5 by drop_pair, ex3_2_intro/
+ | -HK12 -HV12 #Hi #HK10 elim (IHK12 … HK10) -IHK12 -HK10
+ /3 width=5 by drop_drop_lt, ex3_2_intro/
+ ]
+]
+qed-.
+
+lemma lpx_sn_deliftable_dropable: ∀R. l_deliftable_sn R → dropable_sn (lpx_sn R).
+#R #HR #L1 #K1 #s #d #e #H elim H -L1 -K1 -d -e
+[ #d #e #He #X #H >(lpx_sn_inv_atom1 … H) -H
+ /4 width=3 by drop_atom, lpx_sn_atom, ex2_intro/
+| #I #K1 #V1 #X #H elim (lpx_sn_inv_pair1 … H) -H
+ #L2 #V2 #HL12 #HV12 #H destruct
+ /3 width=5 by drop_pair, lpx_sn_pair, ex2_intro/
+| #I #L1 #K1 #V1 #e #_ #IHLK1 #X #H elim (lpx_sn_inv_pair1 … H) -H
+ #L2 #V2 #HL12 #HV12 #H destruct
+ elim (IHLK1 … HL12) -L1 /3 width=3 by drop_drop, ex2_intro/
+| #I #L1 #K1 #V1 #W1 #d #e #HLK1 #HWV1 #IHLK1 #X #H
+ elim (lpx_sn_inv_pair1 … H) -H #L2 #V2 #HL12 #HV12 #H destruct
+ elim (HR … HV12 … HLK1 … HWV1) -V1
+ elim (IHLK1 … HL12) -L1 /3 width=5 by drop_skip, lpx_sn_pair, ex2_intro/
+]
+qed-.
+
+lemma lpx_sn_liftable_dedropable: ∀R. (∀L. reflexive ? (R L)) →
+ l_liftable R → dedropable_sn (lpx_sn R).
+#R #H1R #H2R #L1 #K1 #s #d #e #H elim H -L1 -K1 -d -e
+[ #d #e #He #X #H >(lpx_sn_inv_atom1 … H) -H
+ /4 width=4 by drop_atom, lpx_sn_atom, ex3_intro/
+| #I #K1 #V1 #X #H elim (lpx_sn_inv_pair1 … H) -H
+ #K2 #V2 #HK12 #HV12 #H destruct
+ lapply (lpx_sn_fwd_length … HK12)
+ #H @(ex3_intro … (K2.ⓑ{I}V2)) (**) (* explicit constructor *)
+ /3 width=1 by lpx_sn_pair, monotonic_le_plus_l/
+ @leq_O2 normalize //
+| #I #L1 #K1 #V1 #e #_ #IHLK1 #K2 #HK12 elim (IHLK1 … HK12) -K1
+ /3 width=5 by drop_drop, leq_pair, lpx_sn_pair, ex3_intro/
+| #I #L1 #K1 #V1 #W1 #d #e #HLK1 #HWV1 #IHLK1 #X #H
+ elim (lpx_sn_inv_pair1 … H) -H #K2 #W2 #HK12 #HW12 #H destruct
+ elim (lift_total W2 d e) #V2 #HWV2
+ lapply (H2R … HW12 … HLK1 … HWV1 … HWV2) -W1
+ elim (IHLK1 … HK12) -K1
+ /3 width=6 by drop_skip, leq_succ, lpx_sn_pair, ex3_intro/
+]
+qed-.
+
+fact lpx_sn_dropable_aux: ∀R,L2,K2,s,d,e. ⇩[s, d, e] L2 ≡ K2 → ∀L1. lpx_sn R L1 L2 →
+ d = 0 → ∃∃K1. ⇩[s, 0, e] L1 ≡ K1 & lpx_sn R K1 K2.
+#R #L2 #K2 #s #d #e #H elim H -L2 -K2 -d -e
+[ #d #e #He #X #H >(lpx_sn_inv_atom2 … H) -H
+ /4 width=3 by drop_atom, lpx_sn_atom, ex2_intro/
+| #I #K2 #V2 #X #H elim (lpx_sn_inv_pair2 … H) -H
+ #K1 #V1 #HK12 #HV12 #H destruct
+ /3 width=5 by drop_pair, lpx_sn_pair, ex2_intro/
+| #I #L2 #K2 #V2 #e #_ #IHLK2 #X #H #_ elim (lpx_sn_inv_pair2 … H) -H
+ #L1 #V1 #HL12 #HV12 #H destruct
+ elim (IHLK2 … HL12) -L2 /3 width=3 by drop_drop, ex2_intro/
+| #I #L2 #K2 #V2 #W2 #d #e #_ #_ #_ #L1 #_
+ <plus_n_Sm #H destruct
+]
+qed-.
+
+lemma lpx_sn_dropable: ∀R. dropable_dx (lpx_sn R).
+/2 width=5 by lpx_sn_dropable_aux/ qed-.
+++ /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/substitution/ldrop_leq.ma".
-include "basic_2/substitution/lpx_sn.ma".
-
-(* SN POINTWISE EXTENSION OF A CONTEXT-SENSITIVE REALTION FOR TERMS *********)
-
-(* Properies on dropping ****************************************************)
-
-lemma lpx_sn_ldrop_conf: ∀R,L1,L2. lpx_sn R L1 L2 →
- ∀I,K1,V1,i. ⇩[i] L1 ≡ K1.ⓑ{I}V1 →
- ∃∃K2,V2. ⇩[i] L2 ≡ K2.ⓑ{I}V2 & lpx_sn R K1 K2 & R K1 V1 V2.
-#R #L1 #L2 #H elim H -L1 -L2
-[ #I0 #K0 #V0 #i #H elim (ldrop_inv_atom1 … H) -H #H destruct
-| #I #K1 #K2 #V1 #V2 #HK12 #HV12 #IHK12 #I0 #K0 #V0 #i #H elim (ldrop_inv_O1_pair1 … H) * -H
- [ -IHK12 #H1 #H2 destruct /3 width=5 by ldrop_pair, ex3_2_intro/
- | -HK12 -HV12 #Hi #HK10 elim (IHK12 … HK10) -IHK12 -HK10
- /3 width=5 by ldrop_drop_lt, ex3_2_intro/
- ]
-]
-qed-.
-
-lemma lpx_sn_ldrop_trans: ∀R,L1,L2. lpx_sn R L1 L2 →
- ∀I,K2,V2,i. ⇩[i] L2 ≡ K2.ⓑ{I}V2 →
- ∃∃K1,V1. ⇩[i] L1 ≡ K1.ⓑ{I}V1 & lpx_sn R K1 K2 & R K1 V1 V2.
-#R #L1 #L2 #H elim H -L1 -L2
-[ #I0 #K0 #V0 #i #H elim (ldrop_inv_atom1 … H) -H #H destruct
-| #I #K1 #K2 #V1 #V2 #HK12 #HV12 #IHK12 #I0 #K0 #V0 #i #H elim (ldrop_inv_O1_pair1 … H) * -H
- [ -IHK12 #H1 #H2 destruct /3 width=5 by ldrop_pair, ex3_2_intro/
- | -HK12 -HV12 #Hi #HK10 elim (IHK12 … HK10) -IHK12 -HK10
- /3 width=5 by ldrop_drop_lt, ex3_2_intro/
- ]
-]
-qed-.
-
-lemma lpx_sn_deliftable_dropable: ∀R. l_deliftable_sn R → dropable_sn (lpx_sn R).
-#R #HR #L1 #K1 #s #d #e #H elim H -L1 -K1 -d -e
-[ #d #e #He #X #H >(lpx_sn_inv_atom1 … H) -H
- /4 width=3 by ldrop_atom, lpx_sn_atom, ex2_intro/
-| #I #K1 #V1 #X #H elim (lpx_sn_inv_pair1 … H) -H
- #L2 #V2 #HL12 #HV12 #H destruct
- /3 width=5 by ldrop_pair, lpx_sn_pair, ex2_intro/
-| #I #L1 #K1 #V1 #e #_ #IHLK1 #X #H elim (lpx_sn_inv_pair1 … H) -H
- #L2 #V2 #HL12 #HV12 #H destruct
- elim (IHLK1 … HL12) -L1 /3 width=3 by ldrop_drop, ex2_intro/
-| #I #L1 #K1 #V1 #W1 #d #e #HLK1 #HWV1 #IHLK1 #X #H
- elim (lpx_sn_inv_pair1 … H) -H #L2 #V2 #HL12 #HV12 #H destruct
- elim (HR … HV12 … HLK1 … HWV1) -V1
- elim (IHLK1 … HL12) -L1 /3 width=5 by ldrop_skip, lpx_sn_pair, ex2_intro/
-]
-qed-.
-
-lemma lpx_sn_liftable_dedropable: ∀R. (∀L. reflexive ? (R L)) →
- l_liftable R → dedropable_sn (lpx_sn R).
-#R #H1R #H2R #L1 #K1 #s #d #e #H elim H -L1 -K1 -d -e
-[ #d #e #He #X #H >(lpx_sn_inv_atom1 … H) -H
- /4 width=4 by ldrop_atom, lpx_sn_atom, ex3_intro/
-| #I #K1 #V1 #X #H elim (lpx_sn_inv_pair1 … H) -H
- #K2 #V2 #HK12 #HV12 #H destruct
- lapply (lpx_sn_fwd_length … HK12)
- #H @(ex3_intro … (K2.ⓑ{I}V2)) (**) (* explicit constructor *)
- /3 width=1 by lpx_sn_pair, monotonic_le_plus_l/
- @leq_O2 normalize //
-| #I #L1 #K1 #V1 #e #_ #IHLK1 #K2 #HK12 elim (IHLK1 … HK12) -K1
- /3 width=5 by ldrop_drop, leq_pair, lpx_sn_pair, ex3_intro/
-| #I #L1 #K1 #V1 #W1 #d #e #HLK1 #HWV1 #IHLK1 #X #H
- elim (lpx_sn_inv_pair1 … H) -H #K2 #W2 #HK12 #HW12 #H destruct
- elim (lift_total W2 d e) #V2 #HWV2
- lapply (H2R … HW12 … HLK1 … HWV1 … HWV2) -W1
- elim (IHLK1 … HK12) -K1
- /3 width=6 by ldrop_skip, leq_succ, lpx_sn_pair, ex3_intro/
-]
-qed-.
-
-fact lpx_sn_dropable_aux: ∀R,L2,K2,s,d,e. ⇩[s, d, e] L2 ≡ K2 → ∀L1. lpx_sn R L1 L2 →
- d = 0 → ∃∃K1. ⇩[s, 0, e] L1 ≡ K1 & lpx_sn R K1 K2.
-#R #L2 #K2 #s #d #e #H elim H -L2 -K2 -d -e
-[ #d #e #He #X #H >(lpx_sn_inv_atom2 … H) -H
- /4 width=3 by ldrop_atom, lpx_sn_atom, ex2_intro/
-| #I #K2 #V2 #X #H elim (lpx_sn_inv_pair2 … H) -H
- #K1 #V1 #HK12 #HV12 #H destruct
- /3 width=5 by ldrop_pair, lpx_sn_pair, ex2_intro/
-| #I #L2 #K2 #V2 #e #_ #IHLK2 #X #H #_ elim (lpx_sn_inv_pair2 … H) -H
- #L1 #V1 #HL12 #HV12 #H destruct
- elim (IHLK2 … HL12) -L2 /3 width=3 by ldrop_drop, ex2_intro/
-| #I #L2 #K2 #V2 #W2 #d #e #_ #_ #_ #L1 #_
- <plus_n_Sm #H destruct
-]
-qed-.
-
-lemma lpx_sn_dropable: ∀R. dropable_dx (lpx_sn R).
-/2 width=5 by lpx_sn_dropable_aux/ qed-.
include "ground_2/ynat/ynat_plus.ma".
include "basic_2/notation/relations/lrsubeq_4.ma".
-include "basic_2/substitution/ldrop.ma".
+include "basic_2/substitution/drop.ma".
(* LOCAL ENVIRONMENT REFINEMENT FOR EXTENDED SUBSTITUTION *******************)
(* Properties on basic slicing **********************************************)
-lemma lsuby_ldrop_trans_be: ∀L1,L2,d,e. L1 ⊆[d, e] L2 →
- ∀I2,K2,W,s,i. ⇩[s, 0, i] L2 ≡ K2.ⓑ{I2}W →
- d ≤ i → i < d + e →
- ∃∃I1,K1. K1 ⊆[0, ⫰(d+e-i)] K2 & ⇩[s, 0, i] L1 ≡ K1.ⓑ{I1}W.
+lemma lsuby_drop_trans_be: ∀L1,L2,d,e. L1 ⊆[d, e] L2 →
+ ∀I2,K2,W,s,i. ⇩[s, 0, i] L2 ≡ K2.ⓑ{I2}W →
+ d ≤ i → i < d + e →
+ ∃∃I1,K1. K1 ⊆[0, ⫰(d+e-i)] K2 & ⇩[s, 0, i] L1 ≡ K1.ⓑ{I1}W.
#L1 #L2 #d #e #H elim H -L1 -L2 -d -e
[ #L1 #d #e #J2 #K2 #W #s #i #H
- elim (ldrop_inv_atom1 … H) -H #H destruct
+ elim (drop_inv_atom1 … H) -H #H destruct
| #I1 #I2 #L1 #L2 #V1 #V2 #_ #_ #J2 #K2 #W #s #i #_ #_ #H
elim (ylt_yle_false … H) //
| #I1 #I2 #L1 #L2 #V #e #HL12 #IHL12 #J2 #K2 #W #s #i #H #_ >yplus_O1
- elim (ldrop_inv_O1_pair1 … H) -H * #Hi #HLK1 [ -IHL12 | -HL12 ]
+ elim (drop_inv_O1_pair1 … H) -H * #Hi #HLK1 [ -IHL12 | -HL12 ]
[ #_ destruct -I2 >ypred_succ
- /2 width=4 by ldrop_pair, ex2_2_intro/
+ /2 width=4 by drop_pair, ex2_2_intro/
| lapply (ylt_inv_O1 i ?) /2 width=1 by ylt_inj/
#H <H -H #H lapply (ylt_inv_succ … H) -H
#Hie elim (IHL12 … HLK1) -IHL12 -HLK1 // -Hie
- >yminus_succ <yminus_inj /3 width=4 by ldrop_drop_lt, ex2_2_intro/
+ >yminus_succ <yminus_inj /3 width=4 by drop_drop_lt, ex2_2_intro/
]
| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #IHL12 #J2 #K2 #W #s #i #HLK2 #Hdi
elim (yle_inv_succ1 … Hdi) -Hdi
#Hdi #Hi <Hi >yplus_succ1 #H lapply (ylt_inv_succ … H) -H
- #Hide lapply (ldrop_inv_drop1_lt … HLK2 ?) -HLK2 /2 width=1 by ylt_O/
+ #Hide lapply (drop_inv_drop1_lt … HLK2 ?) -HLK2 /2 width=1 by ylt_O/
#HLK1 elim (IHL12 … HLK1) -IHL12 -HLK1 <yminus_inj >yminus_SO2
- /4 width=4 by ylt_O, ldrop_drop_lt, ex2_2_intro/
+ /4 width=4 by ylt_O, drop_drop_lt, ex2_2_intro/
]
qed-.
[ /2 width=1 by sta_lstasa/
| #G #L #l #k #X #H >(sta_inv_sort1 … H) -X >commutative_plus //
| #G #L #K #V #W #U #i #l #HLK #_ #HWU #IHVW #U2 #HU2
- lapply (ldrop_fwd_drop2 … HLK) #H
+ lapply (drop_fwd_drop2 … HLK) #H
elim (sta_inv_lift1 … HU2 … H … HWU) -H -U /3 width=6 by lstasa_ldef/
| #G #L #K #W #V #V0 #U #i #l #HLK #HWl0 #_ #HVU #IHWV #U2 #HU2
- lapply (ldrop_fwd_drop2 … HLK) #H
+ lapply (drop_fwd_drop2 … HLK) #H
elim (sta_inv_lift1 … HU2 … H … HVU) -H -U /3 width=8 by lstasa_ldec/
| #a #I #G #L #V #T1 #U1 #l #_ #IHTU1 #X #H
elim (sta_inv_bind1 … H) -H #U #HU1 #H destruct /3 width=1 by lstasa_bind/
#h #G #L #U #i #l #H elim (lstas_inv_step_sn … H) -H
#X #H #HXU elim (sta_inv_lref1 … H) -H
* #K #V #W #HLK #HVW #HWX
-lapply (ldrop_fwd_drop2 … HLK) #H0LK
+lapply (drop_fwd_drop2 … HLK) #H0LK
elim (lstas_inv_lift1 … HXU … H0LK … HWX) -H0LK -X
/4 width=8 by lstas_step_sn, ex4_4_intro, ex3_3_intro, or_introl, or_intror/
qed-.
∀W,l. ⦃G, K⦄ ⊢ V •*[h, l+1] W →
∀U. ⇧[0, i+1] W ≡ U → ⦃G, L⦄ ⊢ #i •*[h, l+1] U.
#h #G #L #K #V #i #HLK #W #l #HVW #U #HWU
-lapply (ldrop_fwd_drop2 … HLK)
+lapply (drop_fwd_drop2 … HLK)
elim (lstas_inv_step_sn … HVW) -HVW #W0
elim (lift_total W0 0 (i+1)) /3 width=12 by lstas_step_sn, sta_ldef, lstas_lift/
qed.
∀V,l. ⦃G, K⦄ ⊢ W •*[h, l] V →
∀U. ⇧[0, i+1] V ≡ U → ⦃G, L⦄ ⊢ #i •*[h, l+1] U.
#h #G #L #K #W #i #HLK #V0 #HWV0 #V #l #HWV #U #HVU
-lapply (ldrop_fwd_drop2 … HLK) #H
+lapply (drop_fwd_drop2 … HLK) #H
elim (lift_total W 0 (i+1)) /3 width=12 by lstas_step_sn, sta_ldec, lstas_lift/
qed.
include "basic_2/notation/relations/unfold_4.ma".
include "basic_2/grammar/lenv_append.ma".
include "basic_2/grammar/genv.ma".
-include "basic_2/substitution/ldrop.ma".
+include "basic_2/substitution/drop.ma".
(* CONTEXT-SENSITIVE UNFOLD FOR TERMS ***************************************)
[ { "dynamic typing" * } {
(*
[ { "local env. ref. for native type assignment" * } {
- [ "lsubn ( ? ⊢ ? :⫃ ? )" "lsubn_ldrop" "lsubn_cpcs" "lsubn_nta" * ]
+ [ "lsubn ( ? ⊢ ? :⫃ ? )" "lsubn_drop" "lsubn_cpcs" "lsubn_nta" * ]
}
]
[ { "native type assignment" * } {
]
[ { "strongly normalizing extended computation" * } {
[ "lcosx ( ? ⊢ ~⬊*[?,?,?] ? )" "lcosx_cpx" * ]
- [ "lsx ( ? ⊢ ⬊*[?,?,?,?] ? )" "lsx_alt ( ? ⊢ ⬊⬊*[?,?,?,?] ? )" "lsx_ldrop" + "lsx_lpx" + "lsx_lpxs" + "llsx_csx" * ]
+ [ "lsx ( ? ⊢ ⬊*[?,?,?,?] ? )" "lsx_alt ( ? ⊢ ⬊⬊*[?,?,?,?] ? )" "lsx_drop" + "lsx_lpx" + "lsx_lpxs" + "llsx_csx" * ]
[ "csx_vector ( ⦃?,?⦄ ⊢ ⬊*[?,?] ? )" "csx_tstc_vector" + "csx_aaa" * ]
[ "csx ( ⦃?,?⦄ ⊢ ⬊*[?,?] ? )" "csx_alt ( ⦃?,?⦄ ⊢ ⬊⬊*[?,?] ? )" "csx_lift" + "csx_lleq" + "csx_lpx" + "csx_lpxs" + "csx_fpbs" * ]
}
}
]
[ { "context-sensitive extended computation" * } {
- [ "lpxs ( ⦃?,?⦄ ⊢ ➡*[?,?] ? )" "lpxs_ldrop" + "lpxs_lleq" + "lpxs_aaa" + "lpxs_cpxs" + "lpxs_lpxs" * ]
+ [ "lpxs ( ⦃?,?⦄ ⊢ ➡*[?,?] ? )" "lpxs_drop" + "lpxs_lleq" + "lpxs_aaa" + "lpxs_cpxs" + "lpxs_lpxs" * ]
[ "cpxs ( ⦃?,?⦄ ⊢ ? ➡*[?,?] ? )" "cpxs_tstc" + "cpxs_tstc_vector" + "cpxs_leq" + "cpxs_lift" + "cpxs_lleq" + "cpxs_aaa" + "cpxs_cpxs" * ]
}
]
[ { "context-sensitive computation" * } {
- [ "lprs ( ⦃?,?⦄ ⊢ ➡* ? )" "lprs_ldrop" + "lprs_cprs" + "lprs_lprs" * ]
+ [ "lprs ( ⦃?,?⦄ ⊢ ➡* ? )" "lprs_drop" + "lprs_cprs" + "lprs_lprs" * ]
[ "cprs ( ⦃?,?⦄ ⊢ ? ➡* ?)" "cprs_lift" + "cprs_cprs" * ]
}
]
[ { "local env. ref. for abstract candidates of reducibility" * } {
- [ "lsubc ( ? ⊢ ? ⫃[?] ? )" "lsubc_ldrop" + "lsubc_ldrops" + "lsubc_lsuba" * ]
+ [ "lsubc ( ? ⊢ ? ⫃[?] ? )" "lsubc_drop" + "lsubc_drops" + "lsubc_lsuba" * ]
}
]
[ { "support for abstract computation properties" * } {
}
]
[ { "context-sensitive extended reduction" * } {
- [ "lpx ( ⦃?,?⦄ ⊢ ➡[?,?] ? )" "lpx_ldrop" + "lpx_frees" "lpx_lleq" + "lpx_aaa" * ]
+ [ "lpx ( ⦃?,?⦄ ⊢ ➡[?,?] ? )" "lpx_drop" + "lpx_frees" "lpx_lleq" + "lpx_aaa" * ]
[ "cpx ( ⦃?,?⦄ ⊢ ? ➡[?,?] ? )" "cpx_leq" + "cpx_lift" + "cpx_llpx_sn" + "cpx_lleq" + "cpx_cix" * ]
}
]
}
]
[ { "context-sensitive reduction" * } {
- [ "lpr ( ⦃?,?⦄ ⊢ ➡ ? )" "lpr_ldrop" + "lpr_lpr" * ]
+ [ "lpr ( ⦃?,?⦄ ⊢ ➡ ? )" "lpr_drop" + "lpr_lpr" * ]
[ "cpr ( ⦃?,?⦄ ⊢ ? ➡ ? )" "cpr_lift" + "cpr_llpx_sn" + "cpr_cir" * ]
}
]
[ { "multiple substitution" * } {
[ { "lazy equivalence" * } {
[ "fleq ( ⦃?,?,?⦄ ≡[?] ⦃?,?,?⦄ )" "fleq_fleq" * ]
- [ "lleq ( ? ≡[?,?] ? )" "lleq_alt" + "lleq_alt_rec" + "lleq_leq" + "lleq_ldrop" + "lleq_fqus" + "lleq_llor" + "lleq_lleq" * ]
+ [ "lleq ( ? ≡[?,?] ? )" "lleq_alt" + "lleq_alt_rec" + "lleq_leq" + "lleq_drop" + "lleq_fqus" + "lleq_llor" + "lleq_lleq" * ]
}
]
[ { "lazy pointwise extension of a relation" * } {
- [ "llpx_sn" "llpx_sn_alt" + "llpx_sn_alt_rec" + "llpx_sn_tc" + "llpx_sn_leq" + "llpx_sn_ldrop" + "llpx_sn_lpx_sn" + "llpx_sn_frees" + "llpx_sn_llor" * ]
+ [ "llpx_sn" "llpx_sn_alt" + "llpx_sn_alt_rec" + "llpx_sn_tc" + "llpx_sn_leq" + "llpx_sn_drop" + "llpx_sn_lpx_sn" + "llpx_sn_frees" + "llpx_sn_llor" * ]
}
]
[ { "pointwise union for local environments" * } {
- [ "llor ( ? ⩖[?,?] ? ≡ ? )" "llor_alt" + "llor_ldrop" * ]
+ [ "llor ( ? ⩖[?,?] ? ≡ ? )" "llor_alt" + "llor_drop" * ]
}
]
[ { "context-sensitive exclusion from free variables" * } {
}
]
[ { "iterated local env. slicing" * } {
- [ "ldrops ( ⇩*[?,?] ? ≡ ? )" "ldrops_ldrop" + "ldrops_ldrops" * ]
+ [ "drops ( ⇩*[?,?] ? ≡ ? )" "drops_drop" + "drops_drops" * ]
}
]
[ { "generic term relocation" * } {
}
]
[ { "pointwise extension of a relation" * } {
- [ "lpx_sn" "lpx_sn_alt" + "lpx_sn_tc" + "lpx_sn_ldrop" + "lpx_sn_lpx_sn" * ]
+ [ "lpx_sn" "lpx_sn_alt" + "lpx_sn_tc" + "lpx_sn_drop" + "lpx_sn_lpx_sn" * ]
}
]
[ { "basic local env. slicing" * } {
- [ "ldrop ( ⇩[?,?,?] ? ≡ ? )" "ldrop_append" + "ldrop_leq" + "ldrop_ldrop" * ]
+ [ "drop ( ⇩[?,?,?] ? ≡ ? )" "drop_append" + "drop_leq" + "drop_drop" * ]
}
]
[ { "basic term relocation" * } {
class "red"
[ { "grammar" * } {
[ { "equivalence for local environments" * } {
- [ "leq ( ? â\89\83[?,?] ? )" "leq_leq" * ]
+ [ "leq ( ? ⩬[?,?] ? )" "leq_leq" * ]
}
]
[ { "same top term constructor" * } {
- [ "tstc ( ? â\89\83 ? )" "tstc_tstc" + "tstc_vector" * ]
+ [ "tstc ( ? â\89\82 ? )" "tstc_tstc" + "tstc_vector" * ]
}
]
[ { "closures" * } {