applications of λδ version 2.
In particular it contains the components below.
</body>
- <news date="MLTT1.">
- Martin-Löf's Type Theory with one universe
- using λδ as theory of expressions.
- </news>
- <news date="Functional.">
- The validation algorithm for λδ as implemented in
- <rlink to="implementation.html#helena">Helena 0.8</rlink>.
- </news>
+ <topitem name="MLTT1">
+ <notice class="alpha" notice="MLTT1."/>
+ Martin-Löf's Type Theory with one universe
+ using λδ as theory of expressions.
+ </topitem>
+ <topitem name="functional">
+ <notice class="alpha" notice="Functional."/>
+ The validation algorithm for λδ as implemented in
+ <rlink to="implementation.html#helena">Helena 0.8</rlink>.
+ </topitem>
<section>Summary of the Specification</section>
<body>Here is a numerical acount of the specification's contents
and its timeline.
- Nodes are counted according to the "intrinsinc complexity measure"
- [F. Guidi: "Procedural Representation of CIC Proof Terms"
- Journal of Automated Reasoning 44(1-2), Springer (February 2010),
- pp. 53-78].
</body>
<table name="apps_2_sum"/>
- <news date="2012 February 24.">
+ <news class="alpha" date="2012 February 24.">
The Applications directory is started.
</news>
- <news date="2011 December 20.">
+ <news class="alpha" date="2011 December 20.">
The Functional component is started
inside the specification of λδ version 2.
</news>
- <news date="2011 December 12.">
+ <news class="alpha" date="2011 December 12.">
The MLTT1 component is started.
</news>
</body>
<table name="apps_2_src"/>
- <section>Physical Structure of the Specification</section>
- <body>The source files are grouped in directories,
- one for each component.
- </body>
<footer/>
</page>
(* Relocation properties ****************************************************)
(* Basic_1: was: pr3_lift *)
-lemma cprs_lift: ∀G. l_liftable (cprs G).
-/3 width=10 by l_liftable_LTC, cpr_lift/ qed.
+lemma cprs_lift: ∀G. d_liftable (cprs G).
+/3 width=10 by d_liftable_LTC, cpr_lift/ qed.
(* Basic_1: was: pr3_gen_lift *)
-lemma cprs_inv_lift1: ∀G. l_deliftable_sn (cprs G).
-/3 width=6 by l_deliftable_sn_LTC, cpr_inv_lift1/
+lemma cprs_inv_lift1: ∀G. d_deliftable_sn (cprs G).
+/3 width=6 by d_deliftable_sn_LTC, cpr_inv_lift1/
qed-.
#h #g #G #L #T1 #T2 #H @(cprs_ind … H) -T2 /3 width=3 by cpxs_strap1, cpr_cpx/
qed.
-lemma cpxs_sort: ∀h,g,G,L,k,l1. deg h g k l1 →
- ∀l2. l2 ≤ l1 → ⦃G, L⦄ ⊢ ⋆k ➡*[h, g] ⋆((next h)^l2 k).
-#h #g #G #L #k #l1 #Hkl1 #l2 @(nat_ind_plus … l2) -l2 /2 width=1 by cpx_cpxs/
-#l2 #IHl2 #Hl21 >iter_SO
-@(cpxs_strap1 … (⋆(iter l2 ℕ (next h) k)))
+lemma cpxs_sort: ∀h,g,G,L,k,d1. deg h g k d1 →
+ ∀d2. d2 ≤ d1 → ⦃G, L⦄ ⊢ ⋆k ➡*[h, g] ⋆((next h)^d2 k).
+#h #g #G #L #k #d1 #Hkd1 #d2 @(nat_ind_plus … d2) -d2 /2 width=1 by cpx_cpxs/
+#d2 #IHd2 #Hd21 >iter_SO
+@(cpxs_strap1 … (⋆(iter d2 ℕ (next h) k)))
[ /3 width=3 by lt_to_le/
-| @(cpx_st … (l1-l2-1)) <plus_minus_m_m
+| @(cpx_st … (d1-d2-1)) <plus_minus_m_m
/2 width=1 by deg_iter, monotonic_le_minus_r/
]
qed.
(* Basic inversion lemmas ***************************************************)
lemma cpxs_inv_sort1: ∀h,g,G,L,U2,k. ⦃G, L⦄ ⊢ ⋆k ➡*[h, g] U2 →
- ∃∃n,l. deg h g k (n+l) & U2 = ⋆((next h)^n k).
+ ∃∃n,d. deg h g k (n+d) & U2 = ⋆((next h)^n k).
#h #g #G #L #U2 #k #H @(cpxs_ind … H) -U2
-[ elim (deg_total h g k) #l #Hkl
- @(ex2_2_intro … 0 … Hkl) -Hkl //
-| #U #U2 #_ #HU2 * #n #l #Hknl #H destruct
+[ elim (deg_total h g k) #d #Hkd
+ @(ex2_2_intro … 0 … Hkd) -Hkd //
+| #U #U2 #_ #HU2 * #n #d #Hknd #H destruct
elim (cpx_inv_sort1 … HU2) -HU2
[ #H destruct /2 width=4 by ex2_2_intro/
- | * #l0 #Hkl0 #H destruct -l
- @(ex2_2_intro … (n+1) l0) /2 width=1 by deg_inv_prec/ >iter_SO //
+ | * #d0 #Hkd0 #H destruct -d
+ @(ex2_2_intro … (n+1) d0) /2 width=1 by deg_inv_prec/ >iter_SO //
]
]
qed-.
[1,3: /2 width=4 by fqu_flat_dx, cpxs_flat/
| #H0 destruct /2 width=1 by/
]
-| #G #L #K #T1 #U1 #e #HLK #HTU1 #T2 #HT12 #H elim (lift_total T2 0 (e+1))
+| #G #L #K #T1 #U1 #m #HLK #HTU1 #T2 #HT12 #H elim (lift_total T2 0 (m+1))
#U2 #HTU2 @(ex3_intro … U2)
[1,3: /2 width=10 by cpxs_lift, fqu_drop/
| #H0 destruct /3 width=5 by lift_inj/
+++ /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/cpx_leq.ma".
-include "basic_2/computation/cpxs.ma".
-
-(* CONTEXT-SENSITIVE EXTENDED PARALLEL COMPUTATION ON TERMS *****************)
-
-(* Properties on equivalence for local environments *************************)
-
-lemma leq_cpxs_trans: ∀h,g,G. lsub_trans … (cpxs h g G) (leq 0 (∞)).
-/3 width=5 by leq_cpx_trans, LTC_lsub_trans/
-qed-.
]
qed.
-lemma lstas_cpxs: ∀h,g,G,L,T1,T2,l2. ⦃G, L⦄ ⊢ T1 •*[h, l2] T2 →
- ∀l1. ⦃G, L⦄ ⊢ T1 ▪[h, g] l1 → l2 ≤ l1 → ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2.
-#h #g #G #L #T1 #T2 #l2 #H elim H -G -L -T1 -T2 -l2 //
+lemma lstas_cpxs: ∀h,g,G,L,T1,T2,d2. ⦃G, L⦄ ⊢ T1 •*[h, d2] T2 →
+ ∀d1. ⦃G, L⦄ ⊢ T1 ▪[h, g] d1 → d2 ≤ d1 → ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2.
+#h #g #G #L #T1 #T2 #d2 #H elim H -G -L -T1 -T2 -d2 //
[ /3 width=3 by cpxs_sort, da_inv_sort/
-| #G #L #K #V1 #V2 #W2 #i #l2 #HLK #_ #HVW2 #IHV12 #l1 #H #Hl21
- elim (da_inv_lref … H) -H * #K0 #V0 [| #l0 ] #HLK0
+| #G #L #K #V1 #V2 #W2 #i #d2 #HLK #_ #HVW2 #IHV12 #d1 #H #Hd21
+ elim (da_inv_lref … H) -H * #K0 #V0 [| #d0 ] #HLK0
lapply (drop_mono … HLK0 … HLK) -HLK0 #H destruct /3 width=7 by cpxs_delta/
-| #G #L #K #V1 #V2 #W2 #i #l2 #HLK #_ #HVW2 #IHV12 #l1 #H #Hl21
- elim (da_inv_lref … H) -H * #K0 #V0 [| #l0 ] #HLK0
+| #G #L #K #V1 #V2 #W2 #i #d2 #HLK #_ #HVW2 #IHV12 #d1 #H #Hd21
+ elim (da_inv_lref … H) -H * #K0 #V0 [| #d0 ] #HLK0
lapply (drop_mono … HLK0 … HLK) -HLK0 #H destruct
- #HV1 #H destruct lapply (le_plus_to_le_r … Hl21) -Hl21
+ #HV1 #H destruct lapply (le_plus_to_le_r … Hd21) -Hd21
/3 width=7 by cpxs_delta/
| /4 width=3 by cpxs_bind_dx, da_inv_bind/
| /4 width=3 by cpxs_flat_dx, da_inv_flat/
(* Relocation properties ****************************************************)
-lemma cpxs_lift: ∀h,g,G. l_liftable (cpxs h g G).
-/3 width=10 by cpx_lift, cpxs_strap1, l_liftable_LTC/ qed.
+lemma cpxs_lift: ∀h,g,G. d_liftable (cpxs h g G).
+/3 width=10 by cpx_lift, cpxs_strap1, d_liftable_LTC/ qed.
-lemma cpxs_inv_lift1: ∀h,g,G. l_deliftable_sn (cpxs h g G).
-/3 width=6 by l_deliftable_sn_LTC, cpx_inv_lift1/
+lemma cpxs_inv_lift1: ∀h,g,G. d_deliftable_sn (cpxs h g G).
+/3 width=6 by d_deliftable_sn_LTC, cpx_inv_lift1/
qed-.
(* Properties on supclosure *************************************************)
qed-.
lemma fquq_lstas_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ∀U2,l1. ⦃G2, L2⦄ ⊢ T2 •*[h, l1] U2 →
- ∀l2. ⦃G2, L2⦄ ⊢ T2 ▪[h, g] l2 → l1 ≤ l2 →
+ ∀U2,d1. ⦃G2, L2⦄ ⊢ T2 •*[h, d1] U2 →
+ ∀d2. ⦃G2, L2⦄ ⊢ T2 ▪[h, g] d2 → d1 ≤ d2 →
∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, g] U1 & ⦃G1, L1, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
/3 width=5 by fquq_cpxs_trans, lstas_cpxs/ qed-.
qed-.
lemma fqus_lstas_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
- ∀U2,l1. ⦃G2, L2⦄ ⊢ T2 •*[h, l1] U2 →
- ∀l2. ⦃G2, L2⦄ ⊢ T2 ▪[h, g] l2 → l1 ≤ l2 →
+ ∀U2,d1. ⦃G2, L2⦄ ⊢ T2 •*[h, d1] U2 →
+ ∀d2. ⦃G2, L2⦄ ⊢ T2 ▪[h, g] d2 → d1 ≤ d2 →
∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, g] U1 & ⦃G1, L1, U1⦄ ⊐* ⦃G2, L2, U2⦄.
/3 width=6 by fqus_cpxs_trans, lstas_cpxs/ 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/cpx_lreq.ma".
+include "basic_2/computation/cpxs.ma".
+
+(* CONTEXT-SENSITIVE EXTENDED PARALLEL COMPUTATION ON TERMS *****************)
+
+(* Properties on equivalence for local environments *************************)
+
+lemma lreq_cpxs_trans: ∀h,g,G. lsub_trans … (cpxs h g G) (lreq 0 (∞)).
+/3 width=5 by lreq_cpx_trans, LTC_lsub_trans/
+qed-.
lemma cpxs_fwd_sort: ∀h,g,G,L,U,k. ⦃G, L⦄ ⊢ ⋆k ➡*[h, g] U →
⋆k ≂ 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/
-| #n #IHn #k >plus_plus_comm_23 #Hnl #H destruct
- lapply (deg_next_SO … Hnl) -Hnl #Hnl
- elim (IHn … Hnl) -IHn
+elim (cpxs_inv_sort1 … H) -H #n #d generalize in match k; -k @(nat_ind_plus … n) -n
+[ #k #_ #H -d destruct /2 width=1 by or_introl/
+| #n #IHn #k >plus_plus_comm_23 #Hnd #H destruct
+ lapply (deg_next_SO … Hnd) -Hnd #Hnd
+ elim (IHn … Hnd) -IHn
[ #H lapply (tsts_inv_atom1 … H) -H #H >H -H /2 width=1 by or_intror/
- | generalize in match Hnl; -Hnl @(nat_ind_plus … n) -n
+ | generalize in match Hnd; -Hnd @(nat_ind_plus … n) -n
/4 width=3 by cpxs_strap2, cpx_st, or_intror/
| >iter_SO >iter_n_Sm //
]
lemma csx_sort: ∀h,g,G,L,k. ⦃G, L⦄ ⊢ ⬊*[h, g] ⋆k.
#h #g #G #L #k elim (deg_total h g k)
-#l generalize in match k; -k @(nat_ind_plus … l) -l /3 width=6 by cnx_csx, cnx_sort/
-#l #IHl #k #Hkl lapply (deg_next_SO … Hkl) -Hkl
-#Hkl @csx_intro #X #H #HX elim (cpx_inv_sort1 … H) -H
+#d generalize in match k; -k @(nat_ind_plus … d) -d /3 width=6 by cnx_csx, cnx_sort/
+#d #IHd #k #Hkd lapply (deg_next_SO … Hkd) -Hkd
+#Hkd @csx_intro #X #H #HX elim (cpx_inv_sort1 … H) -H
[ #H destruct elim HX //
-| -HX * #l0 #_ #H destruct -l0 /2 width=1 by/
+| -HX * #d0 #_ #H destruct -d0 /2 width=1 by/
]
qed.
(* Relocation properties ****************************************************)
(* Basic_1: was just: sn3_lift *)
-lemma csx_lift: ∀h,g,G,L2,L1,T1,s,d,e. ⦃G, L1⦄ ⊢ ⬊*[h, g] T1 →
- ∀T2. ⬇[s, d, e] L2 ≡ L1 → ⬆[d, e] T1 ≡ T2 → ⦃G, L2⦄ ⊢ ⬊*[h, g] T2.
-#h #g #G #L2 #L1 #T1 #s #d #e #H elim H -T1 #T1 #_ #IHT1 #T2 #HL21 #HT12
+lemma csx_lift: ∀h,g,G,L2,L1,T1,s,l,m. ⦃G, L1⦄ ⊢ ⬊*[h, g] T1 →
+ ∀T2. ⬇[s, l, m] L2 ≡ L1 → ⬆[l, m] T1 ≡ T2 → ⦃G, L2⦄ ⊢ ⬊*[h, g] T2.
+#h #g #G #L2 #L1 #T1 #s #l #m #H elim H -T1 #T1 #_ #IHT1 #T2 #HL21 #HT12
@csx_intro #T #HLT2 #HT2
elim (cpx_inv_lift1 … HLT2 … HL21 … HT12) -HLT2 #T0 #HT0 #HLT10
@(IHT1 … HLT10) // -L1 -L2 #H destruct
qed.
(* Basic_1: was just: sn3_gen_lift *)
-lemma csx_inv_lift: ∀h,g,G,L2,L1,T1,s,d,e. ⦃G, L1⦄ ⊢ ⬊*[h, g] T1 →
- ∀T2. ⬇[s, d, e] L1 ≡ L2 → ⬆[d, e] T2 ≡ T1 → ⦃G, L2⦄ ⊢ ⬊*[h, g] T2.
-#h #g #G #L2 #L1 #T1 #s #d #e #H elim H -T1 #T1 #_ #IHT1 #T2 #HL12 #HT21
+lemma csx_inv_lift: ∀h,g,G,L2,L1,T1,s,l,m. ⦃G, L1⦄ ⊢ ⬊*[h, g] T1 →
+ ∀T2. ⬇[s, l, m] L1 ≡ L2 → ⬆[l, m] T2 ≡ T1 → ⦃G, L2⦄ ⊢ ⬊*[h, g] T2.
+#h #g #G #L2 #L1 #T1 #s #l #m #H elim H -T1 #T1 #_ #IHT1 #T2 #HL12 #HT21
@csx_intro #T #HLT2 #HT2
-elim (lift_total T d e) #T0 #HT0
+elim (lift_total T l m) #T0 #HT0
lapply (cpx_lift … HLT2 … HL12 … HT21 … HT0) -HLT2 #HLT10
@(IHT1 … HLT10) // -L1 -L2 #H destruct
>(lift_inj … HT0 … HT21) in HT2; -T1 /2 width=1 by/
lemma csx_applv_sort: ∀h,g,G,L,k,Vs. ⦃G, L⦄ ⊢ ⬊*[h, g] Vs → ⦃G, L⦄ ⊢ ⬊*[h, g] ⒶVs.⋆k.
#h #g #G #L #k elim (deg_total h g k)
-#l generalize in match k; -k @(nat_ind_plus … l) -l [ /3 width=6 by csx_applv_cnx, cnx_sort, simple_atom/ ]
-#l #IHl #k #Hkl lapply (deg_next_SO … Hkl) -Hkl
-#Hkl #Vs elim Vs -Vs /2 width=1 by/
+#d generalize in match k; -k @(nat_ind_plus … d) -d [ /3 width=6 by csx_applv_cnx, cnx_sort, simple_atom/ ]
+#d #IHd #k #Hkd lapply (deg_next_SO … Hkd) -Hkd
+#Hkd #Vs elim Vs -Vs /2 width=1 by/
#V #Vs #IHVs #HVVs
elim (csxv_inv_cons … HVVs) #HV #HVs
@csx_appl_simple_tsts /2 width=1 by applv_simple, simple_atom/ -IHVs -HV -HVs
/4 width=5 by cpxs_fpbs, fpb_cpx, ex2_3_intro/
qed.
-lemma lstas_fpbg: ∀h,g,G,L,T1,T2,l2. ⦃G, L⦄ ⊢ T1 •*[h, l2] T2 → (T1 = T2 → ⊥) →
- ∀l1. l2 ≤ l1 → ⦃G, L⦄ ⊢ T1 ▪[h, g] l1 → ⦃G, L, T1⦄ >≡[h, g] ⦃G, L, T2⦄.
+lemma lstas_fpbg: ∀h,g,G,L,T1,T2,d2. ⦃G, L⦄ ⊢ T1 •*[h, d2] T2 → (T1 = T2 → ⊥) →
+ ∀d1. d2 ≤ d1 → ⦃G, L⦄ ⊢ T1 ▪[h, g] d1 → ⦃G, L, T1⦄ >≡[h, g] ⦃G, L, T2⦄.
/3 width=5 by lstas_cpxs, cpxs_fpbg/ qed.
lemma lpxs_fpbg: ∀h,g,G,L1,L2,T. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 →
(* Advanced properties ******************************************************)
-lemma sta_fpbg: ∀h,g,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 ▪[h, g] l+1 →
+lemma sta_fpbg: ∀h,g,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 ▪[h, g] d+1 →
⦃G, L⦄ ⊢ T1 •*[h, 1] T2 → ⦃G, L, T1⦄ >≡[h, g] ⦃G, L, T2⦄.
/4 width=2 by fpb_fpbg, sta_fpb/ qed.
(* Advanced properties ******************************************************)
-lemma lstas_fpbs: ∀h,g,G,L,T1,T2,l2. ⦃G, L⦄ ⊢ T1 •*[h, l2] T2 →
- ∀l1. l2 ≤ l1 → ⦃G, L⦄ ⊢ T1 ▪[h, g] l1 → ⦃G, L, T1⦄ ≥[h, g] ⦃G, L, T2⦄.
+lemma lstas_fpbs: ∀h,g,G,L,T1,T2,d2. ⦃G, L⦄ ⊢ T1 •*[h, d2] T2 →
+ ∀d1. d2 ≤ d1 → ⦃G, L⦄ ⊢ T1 ▪[h, g] d1 → ⦃G, L, T1⦄ ≥[h, g] ⦃G, L, T2⦄.
/3 width=5 by cpxs_fpbs, lstas_cpxs/ qed.
-lemma sta_fpbs: ∀h,g,G,L,T,U,l.
- ⦃G, L⦄ ⊢ T ▪[h, g] l+1 → ⦃G, L⦄ ⊢ T •*[h, 1] U →
+lemma sta_fpbs: ∀h,g,G,L,T,U,d.
+ ⦃G, L⦄ ⊢ T ▪[h, g] d+1 → ⦃G, L⦄ ⊢ T •*[h, 1] U →
⦃G, L, T⦄ ≥[h, g] ⦃G, L, U⦄.
/2 width=5 by lstas_fpbs/ qed.
(* Note: this is used in the closure proof *)
-lemma cpr_lpr_sta_fpbs: ∀h,g,G,L1,L2,T1,T2,U2,l2.
+lemma cpr_lpr_sta_fpbs: ∀h,g,G,L1,L2,T1,T2,U2,d2.
⦃G, L1⦄ ⊢ T1 ➡ T2 → ⦃G, L1⦄ ⊢ ➡ L2 →
- ⦃G, L2⦄ ⊢ T2 ▪[h, g] l2+1 → ⦃G, L2⦄ ⊢ T2 •*[h, 1] U2 →
+ ⦃G, L2⦄ ⊢ T2 ▪[h, g] d2+1 → ⦃G, L2⦄ ⊢ T2 •*[h, 1] U2 →
⦃G, L1, T1⦄ ≥[h, g] ⦃G, L2, U2⦄.
/4 width=5 by fpbs_strap1, cpr_lpr_fpbs, sta_cpx, fpbq_cpx/ qed.
#G0 #L0 #T0 #IHu #H10 lapply (csx_fpbs_conf … H10) // -HT1
#HT0 generalize in match IHu; -IHu generalize in match H10; -H10
@(lsx_ind … (csx_lsx … HT0 0)) -L0
-#L0 #_ #IHl #H10 #IHu @fsb_intro
-#G2 #L2 #T2 * -G2 -L2 -T2 [ -IHl -IHc | -IHu -IHl | ]
+#L0 #_ #IHd #H10 #IHu @fsb_intro
+#G2 #L2 #T2 * -G2 -L2 -T2 [ -IHd -IHc | -IHu -IHd | ]
[ /4 width=5 by fpbs_fqup_trans, fqu_fqup/
| #T2 #HT02 #HnT02 elim (fpbs_cpx_trans_neq … H10 … HT02 HnT02) -T0
/3 width=4 by/
-| #L2 #HL02 #HnL02 @(IHl … HL02 HnL02) -IHl -HnL02 [ -IHu -IHc | ]
+| #L2 #HL02 #HnL02 @(IHd … HL02 HnL02) -IHd -HnL02 [ -IHu -IHc | ]
[ /3 width=3 by fpbs_lpxs_trans, lpx_lpxs/
| #G3 #L3 #T3 #H03 #_ elim (lpx_fqup_trans … H03 … HL02) -L2
#L4 #T4 elim (eq_term_dec T0 T4) [ -IHc | -IHu ]
definition candidate: Type[0] ≝ relation3 genv lenv term.
definition CP0 ≝ λRR:relation4 genv lenv term term. λRS:relation term.
- ∀G. l_liftable1 (nf RR RS G) (Ⓕ).
+ ∀G. d_liftable1 (nf RR RS G) (Ⓕ).
definition CP1 ≝ λRR:relation4 genv lenv term term. λRS:relation term.
∀G,L. ∃k. NF … (RR G L) RS (⋆k).
-definition CP2 ≝ λRP:candidate. ∀G. l_liftable1 (RP G) (Ⓕ).
+definition CP2 ≝ λRP:candidate. ∀G. d_liftable1 (RP G) (Ⓕ).
definition CP3 ≝ λRP:candidate.
∀G,L,T,k. RP G L (ⓐ⋆k.T) → RP G L T.
(* Basic properties *********************************************************)
(* Basic_1: was: nf2_lift1 *)
-lemma gcp0_lifts: ∀RR,RS,RP. gcp RR RS RP → ∀G. l_liftables1 (nf RR RS G) (Ⓕ).
-#RR #RS #RP #H #G @l1_liftable_liftables @(cp0 … H)
+lemma gcp0_lifts: ∀RR,RS,RP. gcp RR RS RP → ∀G. d_liftables1 (nf RR RS G) (Ⓕ).
+#RR #RS #RP #H #G @d1_liftable_liftables @(cp0 … H)
qed.
-lemma gcp2_lifts: ∀RR,RS,RP. gcp RR RS RP → ∀G. l_liftables1 (RP G) (Ⓕ).
-#RR #RS #RP #H #G @l1_liftable_liftables @(cp2 … H)
+lemma gcp2_lifts: ∀RR,RS,RP. gcp RR RS RP → ∀G. d_liftables1 (RP G) (Ⓕ).
+#RR #RS #RP #H #G @d1_liftable_liftables @(cp2 … H)
qed.
(* Basic_1: was only: sns3_lifts1 *)
-lemma gcp2_lifts_all: ∀RR,RS,RP. gcp RR RS RP → ∀G. l_liftables1_all (RP G) (Ⓕ).
-#RR #RS #RP #H #G @l1_liftables_liftables_all /2 width=7 by gcp2_lifts/
+lemma gcp2_lifts_all: ∀RR,RS,RP. gcp RR RS RP → ∀G. d_liftables1_all (RP G) (Ⓕ).
+#RR #RS #RP #H #G @d1_liftables_liftables_all /2 width=7 by gcp2_lifts/
qed.
lapply (acr_gcr … H1RP H2RP B) #HB
elim (lifts_inv_lref1 … H) -H #i1 #Hi1 #H destruct
lapply (drop_fwd_drop2 … HLK1) #HK1b
- elim (drops_drop_trans … HL01 … HLK1) #X #des1 #i0 #HL0 #H #Hi0 #Hdes1
+ elim (drops_drop_trans … HL01 … HLK1) #X #des1 #i0 #HL0 #H #Hi0 #Hcs1
>(at_mono … Hi1 … Hi0) -i1
- elim (drops_inv_skip2 … Hdes1 … H) -des1 #K0 #V0 #des0 #Hdes0 #HK01 #HV10 #H destruct
+ elim (drops_inv_skip2 … Hcs1 … H) -des1 #K0 #V0 #des0 #Hcs0 #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
+ elim (lifts_lift_trans … Hi0 … Hcs0 … HV10 … HV0) -HV10 #V2 #HV12 #HV2
lapply (s5 … HB ? G ? ? (◊) … HV0 HLK2) /3 width=7 by drops_cons, lifts_cons/ (* Note: uses IHB HL20 V2 HV0 *)
- | -HLK1 -IHB -HL01 -HL20 -HK1b -Hi0 -Hdes0
+ | -HLK1 -IHB -HL01 -HL20 -HK1b -Hi0 -Hcs0
#K2 #V2 #A2 #HKV2A #H1KV0A #H2KV0A #_ #H1 #H2 destruct
lapply (drop_fwd_drop2 … HLK2) #HLK2b
lapply (aaa_lifts … HK01 … HV10 HKV1B) -HKV1B -HK01 -HV10 #HKV0B
(* Basic properties *********************************************************)
(* Basic 1: was: sc3_lift *)
-lemma gcr_lift: ∀RR,RS,RP. gcp RR RS RP → ∀A,G. l_liftable1 (acr RP A G) (Ⓕ).
+lemma gcr_lift: ∀RR,RS,RP. gcp RR RS RP → ∀A,G. d_liftable1 (acr RP A G) (Ⓕ).
#RR #RS #RP #H #A elim A -A
/3 width=8 by cp2, drops_cons, lifts_cons/
qed.
(* Basic_1: was: sc3_lift1 *)
-lemma gcr_lifts: ∀RR,RS,RP. gcp RR RS RP → ∀A,G. l_liftables1 (acr RP A G) (Ⓕ).
-#RR #RS #RP #H #A #G @l1_liftable_liftables /2 width=7 by gcr_lift/
+lemma gcr_lifts: ∀RR,RS,RP. gcp RR RS RP → ∀A,G. d_liftables1 (acr RP A G) (Ⓕ).
+#RR #RS #RP #H #A #G @d1_liftable_liftables /2 width=7 by gcr_lift/
qed.
(* Basic_1: was:
| #I #G #L #K #Vs #V1 #V2 #i #HA #HV12 #HLK #L0 #V0 #X #des #HL0 #H #HB
elim (lifts_inv_applv1 … H) -H #V0s #Y #HV0s #HY #H destruct
elim (lifts_inv_lref1 … HY) -HY #i0 #Hi0 #H destruct
- elim (drops_drop_trans … HL0 … HLK) #X #des0 #i1 #HL02 #H #Hi1 #Hdes0
+ elim (drops_drop_trans … HL0 … HLK) #X #des0 #i1 #HL02 #H #Hi1 #Hcs0
>(at_mono … Hi1 … Hi0) in HL02; -i1 #HL02
- elim (drops_inv_skip2 … Hdes0 … H) -H -des0 #L2 #W1 #des0 #Hdes0 #HLK #HVW1 #H destruct
+ elim (drops_inv_skip2 … Hcs0 … H) -H -des0 #L2 #W1 #des0 #Hcs0 #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
+ elim (lifts_lift_trans … Hcs0 … HVW1 … HW12) // -Hcs0 -Hi0 #V3 #HV13 #HVW2
>(lift_mono … HV13 … HV12) in HVW2; -V3 #HVW2
@(s5 … IHA … (V0 @ V0s) … HW12 HL02) /3 width=5 by lifts_applv/
| #G #L #V1s #V2s #HV12s #a #V #T #HA #HV #L0 #V10 #X #des #HL0 #H #HB
(* SN EXTENDED STRONGLY CONORMALIZING LOCAL ENVIRONMENTS ********************)
inductive lcosx (h) (g) (G): relation2 ynat lenv ≝
-| lcosx_sort: ∀d. lcosx h g G d (⋆)
+| lcosx_sort: ∀l. lcosx h g G l (⋆)
| lcosx_skip: ∀I,L,T. lcosx h g G 0 L → lcosx h g G 0 (L.ⓑ{I}T)
-| lcosx_pair: ∀I,L,T,d. G ⊢ ⬊*[h, g, T, d] L →
- lcosx h g G d L → lcosx h g G (⫯d) (L.ⓑ{I}T)
+| lcosx_pair: ∀I,L,T,l. G ⊢ ⬊*[h, g, T, l] L →
+ lcosx h g G l L → lcosx h g G (⫯l) (L.ⓑ{I}T)
.
interpretation
"sn extended strong conormalization (local environment)"
- 'CoSN h g d G L = (lcosx h g G d L).
+ 'CoSN h g l G L = (lcosx h g G l L).
(* Basic properties *********************************************************)
#h #g #G #L elim L /2 width=1 by lcosx_skip/
qed.
-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 (drop_inv_atom1 … H) -H #H destruct
+lemma lcosx_drop_trans_lt: ∀h,g,G,L,l. G ⊢ ~⬊*[h, g, l] L →
+ ∀I,K,V,i. ⬇[i] L ≡ K.ⓑ{I}V → i < l →
+ G ⊢ ~⬊*[h, g, ⫰(l-i)] K ∧ G ⊢ ⬊*[h, g, V, ⫰(l-i)] K.
+#h #g #G #L #l #H elim H -L -l
+[ #l #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
+| #I #L #T #l #HT #HL #IHL #J #K #V #i #H #Hil
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
+ | lapply (ylt_pred … Hil ?) -Hil /2 width=1 by ylt_inj/ >ypred_succ #Hil
elim (IHL … HLK ?) -IHL -HLK <yminus_inj >yminus_SO2 //
- <(ypred_succ d) in ⊢ (%→%→?); >yminus_pred /2 width=1 by ylt_inj, conj/
+ <(ypred_succ l) in ⊢ (%→%→?); >yminus_pred /2 width=1 by ylt_inj, conj/
]
]
qed-.
(* Basic inversion lemmas ***************************************************)
-fact lcosx_inv_succ_aux: ∀h,g,G,L,x. G ⊢ ~⬊*[h, g, x] L → ∀d. x = ⫯d →
+fact lcosx_inv_succ_aux: ∀h,g,G,L,x. G ⊢ ~⬊*[h, g, x] L → ∀l. x = ⫯l →
L = ⋆ ∨
- ∃∃I,K,V. L = K.ⓑ{I}V & G ⊢ ~⬊*[h, g, d] K &
- G ⊢ ⬊*[h, g, V, d] K.
-#h #g #G #L #d * -L -d /2 width=1 by or_introl/
+ ∃∃I,K,V. L = K.ⓑ{I}V & G ⊢ ~⬊*[h, g, l] K &
+ G ⊢ ⬊*[h, g, V, l] K.
+#h #g #G #L #l * -L -l /2 width=1 by or_introl/
[ #I #L #T #_ #x #H elim (ysucc_inv_O_sn … H)
-| #I #L #T #d #HT #HL #x #H <(ysucc_inj … H) -x
+| #I #L #T #l #HT #HL #x #H <(ysucc_inj … H) -x
/3 width=6 by ex3_3_intro, or_intror/
]
qed-.
-lemma lcosx_inv_succ: ∀h,g,G,L,d. G ⊢ ~⬊*[h, g, ⫯d] L → L = ⋆ ∨
- ∃∃I,K,V. L = K.ⓑ{I}V & G ⊢ ~⬊*[h, g, d] K &
- G ⊢ ⬊*[h, g, V, d] K.
+lemma lcosx_inv_succ: ∀h,g,G,L,l. G ⊢ ~⬊*[h, g, ⫯l] L → L = ⋆ ∨
+ ∃∃I,K,V. L = K.ⓑ{I}V & G ⊢ ~⬊*[h, g, l] K &
+ G ⊢ ⬊*[h, g, V, l] K.
/2 width=3 by lcosx_inv_succ_aux/ qed-.
-lemma lcosx_inv_pair: ∀h,g,I,G,L,T,d. G ⊢ ~⬊*[h, g, ⫯d] L.ⓑ{I}T →
- G ⊢ ~⬊*[h, g, d] L ∧ G ⊢ ⬊*[h, g, T, d] L.
-#h #g #I #G #L #T #d #H elim (lcosx_inv_succ … H) -H
+lemma lcosx_inv_pair: ∀h,g,I,G,L,T,l. G ⊢ ~⬊*[h, g, ⫯l] L.ⓑ{I}T →
+ G ⊢ ~⬊*[h, g, l] L ∧ G ⊢ ⬊*[h, g, T, l] L.
+#h #g #I #G #L #T #l #H elim (lcosx_inv_succ … H) -H
[ #H destruct
| * #Z #Y #X #H destruct /2 width=1 by conj/
]
(* Properties on extended context-sensitive parallel reduction for term *****)
lemma lsx_cpx_trans_lcosx: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 →
- ∀d. G ⊢ ~⬊*[h, g, d] L →
- G ⊢ ⬊*[h, g, T1, d] L → G ⊢ ⬊*[h, g, T2, d] L.
+ ∀l. G ⊢ ~⬊*[h, g, l] L →
+ G ⊢ ⬊*[h, g, T1, l] L → G ⊢ ⬊*[h, g, T2, l] L.
#h #g #G #L #T1 #T2 #H elim H -G -L -T1 -T2 //
-[ #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_drop_trans_lt … HL … HLK) // -HL -Hdi
+[ #I #G #L #K #V1 #V2 #W2 #i #HLK #_ #HVW2 #IHV12 #l #HL #H
+ elim (ylt_split i l) #Hli [ -H | -HL ]
+ [ <(ymax_pre_sn l (⫯i)) /2 width=1 by ylt_fwd_le_succ/
+ elim (lcosx_drop_trans_lt … HL … HLK) // -HL -Hli
lapply (drop_fwd_drop2 … HLK) -HLK /3 width=7 by lsx_lift_ge/
- | lapply (lsx_fwd_lref_be … H … HLK) // -H -Hdi
+ | lapply (lsx_fwd_lref_be … H … HLK) // -H -Hli
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
+| #a #I #G #L #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #l #HL #H
elim (lsx_inv_bind … H) -H #HV1 #HT1
@lsx_bind /2 width=2 by/ (**) (* explicit constructor *)
- @(lsx_leq_conf … (L.ⓑ{I}V1)) /3 width=1 by lcosx_pair, leq_succ/
-| #I #G #L #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #d #HL #H
+ @(lsx_lreq_conf … (L.ⓑ{I}V1)) /3 width=1 by lcosx_pair, lreq_succ/
+| #I #G #L #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #l #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
+| #G #L #V #U1 #U2 #T2 #_ #HTU2 #IHU12 #l #HL #H
elim (lsx_inv_bind … H) -H
/4 width=9 by lcosx_pair, lsx_inv_lift_ge, drop_drop/
-| #G #L #V #T1 #T2 #_ #IHT12 #d #HL #H
+| #G #L #V #T1 #T2 #_ #IHT12 #l #HL #H
elim (lsx_inv_flat … H) -H /2 width=1 by/
-| #G #L #V1 #V2 #T #_ #IHV12 #d #HL #H
+| #G #L #V1 #V2 #T #_ #IHV12 #l #HL #H
elim (lsx_inv_flat … H) -H /2 width=1 by/
-| #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #d #HL #H
+| #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #l #HL #H
elim (lsx_inv_flat … H) -H #HV1 #H
elim (lsx_inv_bind … H) -H #HW1 #HT1
@lsx_bind /3 width=1 by lsx_flat/ (**) (* explicit constructor *)
- @(lsx_leq_conf … (L.ⓛW1)) /3 width=1 by lcosx_pair, leq_succ/
-| #a #G #L #V1 #V2 #U2 #W1 #W2 #T1 #T2 #_ #HVU2 #_ #_ #IHV12 #IHW12 #IHT12 #d #HL #H
+ @(lsx_lreq_conf … (L.ⓛW1)) /3 width=1 by lcosx_pair, lreq_succ/
+| #a #G #L #V1 #V2 #U2 #W1 #W2 #T1 #T2 #_ #HVU2 #_ #_ #IHV12 #IHW12 #IHT12 #l #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, drop_drop/ ]
- @(lsx_leq_conf … (L.ⓓW1)) /3 width=1 by lcosx_pair, leq_succ/
+ @(lsx_lreq_conf … (L.ⓓW1)) /3 width=1 by lcosx_pair, lreq_succ/
]
qed-.
include "basic_2/multiple/lleq_lleq.ma".
include "basic_2/reduction/lpx_lleq.ma".
-include "basic_2/computation/cpxs_leq.ma".
+include "basic_2/computation/cpxs_lreq.ma".
include "basic_2/computation/lpxs_drop.ma".
include "basic_2/computation/lpxs_cpxs.ma".
(* Properties on lazy equivalence for local environments ********************)
lemma lleq_lpxs_trans: ∀h,g,G,L2,K2. ⦃G, L2⦄ ⊢ ➡*[h, g] K2 →
- ∀L1,T,d. L1 ≡[T, d] L2 →
- ∃∃K1. ⦃G, L1⦄ ⊢ ➡*[h, g] K1 & K1 ≡[T, d] K2.
+ ∀L1,T,l. L1 ≡[T, l] L2 →
+ ∃∃K1. ⦃G, L1⦄ ⊢ ➡*[h, g] K1 & K1 ≡[T, l] K2.
#h #g #G #L2 #K2 #H @(lpxs_ind … H) -K2 /2 width=3 by ex2_intro/
-#K #K2 #_ #HK2 #IH #L1 #T #d #HT elim (IH … HT) -L2
+#K #K2 #_ #HK2 #IH #L1 #T #l #HT elim (IH … HT) -L2
#L #HL1 #HT elim (lleq_lpx_trans … HK2 … HT) -K
/3 width=3 by lpxs_strap1, ex2_intro/
qed-.
-lemma lpxs_nlleq_inv_step_sn: ∀h,g,G,L1,L2,T,d. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → (L1 ≡[T, d] L2 → ⊥) →
- ∃∃L,L0. ⦃G, L1⦄ ⊢ ➡[h, g] L & L1 ≡[T, d] L → ⊥ & ⦃G, L⦄ ⊢ ➡*[h, g] L0 & L0 ≡[T, d] L2.
-#h #g #G #L1 #L2 #T #d #H @(lpxs_ind_dx … H) -L1
+lemma lpxs_nlleq_inv_step_sn: ∀h,g,G,L1,L2,T,l. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → (L1 ≡[T, l] L2 → ⊥) →
+ ∃∃L,L0. ⦃G, L1⦄ ⊢ ➡[h, g] L & L1 ≡[T, l] L → ⊥ & ⦃G, L⦄ ⊢ ➡*[h, g] L0 & L0 ≡[T, l] L2.
+#h #g #G #L1 #L2 #T #l #H @(lpxs_ind_dx … H) -L1
[ #H elim H -H //
-| #L1 #L #H1 #H2 #IH2 #H12 elim (lleq_dec T L1 L d) #H
+| #L1 #L #H1 #H2 #IH2 #H12 elim (lleq_dec T L1 L l) #H
[ -H1 -H2 elim IH2 -IH2 /3 width=3 by lleq_trans/ -H12
#L0 #L3 #H1 #H2 #H3 #H4 lapply (lleq_nlleq_trans … H … H2) -H2
#H2 elim (lleq_lpx_trans … H1 … H) -L
/3 width=4 by lpxs_pair, fqu_bind_dx, ex3_intro/
| #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 (drop_O1_le (Ⓕ) (e+1) K1)
+| #G1 #L1 #L #T1 #U1 #m #HL1 #HTU1 #K1 #H1KL1 #H2KL1
+ elim (drop_O1_le (Ⓕ) (m+1) K1)
[ #K #HK1 lapply (lleq_inv_lift_le … H2KL1 … HK1 HL1 … HTU1 ?) -H2KL1 //
#H2KL elim (lpxs_drop_trans_O1 … H1KL1 … HL1) -L1
#K0 #HK10 #H1KL lapply (drop_mono … HK10 … HK1) -HK10 #H destruct
]
qed-.
-fact leq_lpxs_trans_lleq_aux: ∀h,g,G,L1,L0,d,e. L1 ⩬[d, e] L0 → e = ∞ →
- ∀L2. ⦃G, L0⦄ ⊢ ➡*[h, g] L2 →
- ∃∃L. 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
+fact lreq_lpxs_trans_lleq_aux: ∀h,g,G,L1,L0,l,m. L1 ⩬[l, m] L0 → m = ∞ →
+ ∀L2. ⦃G, L0⦄ ⊢ ➡*[h, g] L2 →
+ ∃∃L. L ⩬[l, m] L2 & ⦃G, L1⦄ ⊢ ➡*[h, g] L &
+ (∀T. L0 ≡[T, l] L2 ↔ L1 ≡[T, l] L).
+#h #g #G #L1 #L0 #l #m #H elim H -L1 -L0 -l -m
+[ #l #m #_ #L2 #H >(lpxs_inv_atom1 … H) -H
/3 width=5 by ex3_intro, conj/
-| #I1 #I0 #L1 #L0 #V1 #V0 #_ #_ #He destruct
-| #I #L1 #L0 #V1 #e #HL10 #IHL10 #He #Y #H
+| #I1 #I0 #L1 #L0 #V1 #V0 #_ #_ #Hm destruct
+| #I #L1 #L0 #V1 #m #HL10 #IHL10 #Hm #Y #H
elim (lpxs_inv_pair1 … H) -H #L2 #V2 #HL02 #HV02 #H destruct
- lapply (ysucc_inv_Y_dx … He) -He #He
+ lapply (ysucc_inv_Y_dx … Hm) -Hm #Hm
elim (IHL10 … HL02) // -IHL10 -HL02 #L #HL2 #HL1 #IH
- @(ex3_intro … (L.ⓑ{I}V2)) /3 width=3 by lpxs_pair, leq_cpxs_trans, leq_pair/
+ @(ex3_intro … (L.ⓑ{I}V2)) /3 width=3 by lpxs_pair, lreq_cpxs_trans, lreq_pair/
#T elim (IH T) #HL0dx #HL0sn
- @conj #H @(lleq_leq_repl … H) -H /3 width=1 by leq_sym, leq_pair_O_Y/
-| #I1 #I0 #L1 #L0 #V1 #V0 #d #e #HL10 #IHL10 #He #Y #H
+ @conj #H @(lleq_lreq_repl … H) -H /3 width=1 by lreq_sym, lreq_pair_O_Y/
+| #I1 #I0 #L1 #L0 #V1 #V0 #l #m #HL10 #IHL10 #Hm #Y #H
elim (lpxs_inv_pair1 … H) -H #L2 #V2 #HL02 #HV02 #H destruct
elim (IHL10 … HL02) // -IHL10 -HL02 #L #HL2 #HL1 #IH
- @(ex3_intro … (L.ⓑ{I1}V1)) /3 width=1 by lpxs_pair, leq_succ/
+ @(ex3_intro … (L.ⓑ{I1}V1)) /3 width=1 by lpxs_pair, lreq_succ/
#T elim (IH T) #HL0dx #HL0sn
- @conj #H @(lleq_leq_repl … H) -H /3 width=1 by leq_sym, leq_succ/
+ @conj #H @(lleq_lreq_repl … H) -H /3 width=1 by lreq_sym, lreq_succ/
]
qed-.
-lemma leq_lpxs_trans_lleq: ∀h,g,G,L1,L0,d. L1 ⩬[d, ∞] L0 →
- ∀L2. ⦃G, L0⦄ ⊢ ➡*[h, g] L2 →
- ∃∃L. 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-.
+lemma lreq_lpxs_trans_lleq: ∀h,g,G,L1,L0,l. L1 ⩬[l, ∞] L0 →
+ ∀L2. ⦃G, L0⦄ ⊢ ➡*[h, g] L2 →
+ ∃∃L. L ⩬[l, ∞] L2 & ⦃G, L1⦄ ⊢ ➡*[h, g] L &
+ (∀T. L0 ≡[T, l] L2 ↔ L1 ≡[T, l] L).
+/2 width=1 by lreq_lpxs_trans_lleq_aux/ qed-.
(* 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.
+lemma lsubc_drop_O1_trans: ∀RP,G,L1,L2. G ⊢ L1 ⫃[RP] L2 → ∀K2,s,m. ⬇[s, 0, m] L2 ≡ K2 →
+ ∃∃K1. ⬇[s, 0, m] 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
+[ #X #s #m #H elim (drop_inv_atom1 … H) -H /4 width=3 by drop_atom, ex2_intro/
+| #I #L1 #L2 #V #_ #IHL12 #X #s #m #H
+ elim (drop_inv_O1_pair1 … H) -H * #Hm #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
+| #L1 #L2 #V #W #A #HV #H1W #H2W #_ #IHL12 #X #s #m #H
+ elim (drop_inv_O1_pair1 … H) -H * #Hm #H destruct
[ elim (IHL12 L2 s 0) -IHL12 // #X #H <(drop_inv_O2 … H) -H
/3 width=8 by lsubc_beta, drop_pair, ex2_intro/
| elim (IHL12 … H) -L2 /3 width=3 by drop_drop_lt, ex2_intro/
(* Basic_1: was: csubc_drop_conf_rev *)
lemma drop_lsubc_trans: ∀RR,RS,RP. gcp RR RS RP →
- ∀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 #Hgcp #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/
+ ∀G,L1,K1,l,m. ⬇[Ⓕ, l, m] L1 ≡ K1 → ∀K2. G ⊢ K1 ⫃[RP] K2 →
+ ∃∃L2. G ⊢ L1 ⫃[RP] L2 & ⬇[Ⓕ, l, m] L2 ≡ K2.
+#RR #RS #RP #Hgcp #G #L1 #K1 #l #m #H elim H -L1 -K1 -l -m
+[ #l #m #Hm #X #H elim (lsubc_inv_atom1 … H) -H
+ >Hm /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_beta, drop_pair, ex2_intro/
]
-| #I #L1 #K1 #V1 #e #_ #IHLK1 #K2 #HK12
+| #I #L1 #K1 #V1 #m #_ #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
+| #I #L1 #K1 #V1 #V2 #l #m #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/
∃∃L2. G ⊢ L1 ⫃[RP] L2 & ⬇*[Ⓕ, des] L2 ≡ K2.
#RR #RS #RP #Hgcp #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
+| #L1 #L #K1 #des #l #m #_ #HLK1 #IHL #K2 #HK12
elim (drop_lsubc_trans … Hgcp … HLK1 … HK12) -Hgcp -K1 #K #HLK #HK2
elim (IHL … HLK) -IHL -HLK /3 width=5 by drops_cons, ex2_intro/
]
(* SN EXTENDED STRONGLY NORMALIZING LOCAL ENVIRONMENTS **********************)
definition lsx: ∀h. sd h → relation4 ynat term genv lenv ≝
- λh,g,d,T,G. SN … (lpx h g G) (lleq d T).
+ λh,g,l,T,G. SN … (lpx h g G) (lleq l T).
interpretation
"extended strong normalization (local environment)"
- 'SN h g d T G L = (lsx h g T d G L).
+ 'SN h g l T G L = (lsx h g T l G L).
(* Basic eliminators ********************************************************)
-lemma lsx_ind: ∀h,g,G,T,d. ∀R:predicate lenv.
- (∀L1. G ⊢ ⬊*[h, g, T, d] L1 →
- (∀L2. ⦃G, L1⦄ ⊢ ➡[h, g] L2 → (L1 ≡[T, d] L2 → ⊥) → R L2) →
+lemma lsx_ind: ∀h,g,G,T,l. ∀R:predicate lenv.
+ (∀L1. G ⊢ ⬊*[h, g, T, l] L1 →
+ (∀L2. ⦃G, L1⦄ ⊢ ➡[h, g] L2 → (L1 ≡[T, l] L2 → ⊥) → R L2) →
R L1
) →
- ∀L. G ⊢ ⬊*[h, g, T, d] L → R L.
-#h #g #G #T #d #R #H0 #L1 #H elim H -L1
+ ∀L. G ⊢ ⬊*[h, g, T, l] L → R L.
+#h #g #G #T #l #R #H0 #L1 #H elim H -L1
/5 width=1 by lleq_sym, SN_intro/
qed-.
(* Basic properties *********************************************************)
-lemma lsx_intro: ∀h,g,G,L1,T,d.
- (∀L2. ⦃G, L1⦄ ⊢ ➡[h, g] L2 → (L1 ≡[T, d] L2 → ⊥) → G ⊢ ⬊*[h, g, T, d] L2) →
- G ⊢ ⬊*[h, g, T, d] L1.
+lemma lsx_intro: ∀h,g,G,L1,T,l.
+ (∀L2. ⦃G, L1⦄ ⊢ ➡[h, g] L2 → (L1 ≡[T, l] L2 → ⊥) → G ⊢ ⬊*[h, g, T, l] L2) →
+ G ⊢ ⬊*[h, g, T, l] L1.
/5 width=1 by lleq_sym, SN_intro/ qed.
-lemma lsx_atom: ∀h,g,G,T,d. G ⊢ ⬊*[h, g, T, d] ⋆.
-#h #g #G #T #d @lsx_intro
+lemma lsx_atom: ∀h,g,G,T,l. G ⊢ ⬊*[h, g, T, l] ⋆.
+#h #g #G #T #l @lsx_intro
#X #H #HT lapply (lpx_inv_atom1 … H) -H
#H destruct elim HT -HT //
qed.
-lemma lsx_sort: ∀h,g,G,L,d,k. G ⊢ ⬊*[h, g, ⋆k, d] L.
-#h #g #G #L1 #d #k @lsx_intro
+lemma lsx_sort: ∀h,g,G,L,l,k. G ⊢ ⬊*[h, g, ⋆k, l] L.
+#h #g #G #L1 #l #k @lsx_intro
#L2 #HL12 #H elim H -H
/3 width=4 by lpx_fwd_length, lleq_sort/
qed.
-lemma lsx_gref: ∀h,g,G,L,d,p. G ⊢ ⬊*[h, g, §p, d] L.
-#h #g #G #L1 #d #p @lsx_intro
+lemma lsx_gref: ∀h,g,G,L,l,p. G ⊢ ⬊*[h, g, §p, l] L.
+#h #g #G #L1 #l #p @lsx_intro
#L2 #HL12 #H elim H -H
/3 width=4 by lpx_fwd_length, lleq_gref/
qed.
-lemma lsx_ge_up: ∀h,g,G,L,T,U,dt,d,e. dt ≤ yinj d + yinj e →
- ⬆[d, e] T ≡ U → G ⊢ ⬊*[h, g, U, dt] L → G ⊢ ⬊*[h, g, U, d] L.
-#h #g #G #L #T #U #dt #d #e #Hdtde #HTU #H @(lsx_ind … H) -L
+lemma lsx_ge_up: ∀h,g,G,L,T,U,lt,l,m. lt ≤ yinj l + yinj m →
+ ⬆[l, m] T ≡ U → G ⊢ ⬊*[h, g, U, lt] L → G ⊢ ⬊*[h, g, U, l] L.
+#h #g #G #L #T #U #lt #l #m #Hltlm #HTU #H @(lsx_ind … H) -L
/5 width=7 by lsx_intro, lleq_ge_up/
qed-.
-lemma lsx_ge: ∀h,g,G,L,T,d1,d2. d1 ≤ d2 →
- G ⊢ ⬊*[h, g, T, d1] L → G ⊢ ⬊*[h, g, T, d2] L.
-#h #g #G #L #T #d1 #d2 #Hd12 #H @(lsx_ind … H) -L
+lemma lsx_ge: ∀h,g,G,L,T,l1,l2. l1 ≤ l2 →
+ G ⊢ ⬊*[h, g, T, l1] L → G ⊢ ⬊*[h, g, T, l2] L.
+#h #g #G #L #T #l1 #l2 #Hl12 #H @(lsx_ind … H) -L
/5 width=7 by lsx_intro, lleq_ge/
qed-.
(* Basic forward lemmas *****************************************************)
-lemma lsx_fwd_bind_sn: ∀h,g,a,I,G,L,V,T,d. G ⊢ ⬊*[h, g, ⓑ{a,I}V.T, d] L →
- G ⊢ ⬊*[h, g, V, d] L.
-#h #g #a #I #G #L #V #T #d #H @(lsx_ind … H) -L
+lemma lsx_fwd_bind_sn: ∀h,g,a,I,G,L,V,T,l. G ⊢ ⬊*[h, g, ⓑ{a,I}V.T, l] L →
+ G ⊢ ⬊*[h, g, V, l] L.
+#h #g #a #I #G #L #V #T #l #H @(lsx_ind … H) -L
#L1 #_ #IHL1 @lsx_intro
#L2 #HL12 #HV @IHL1 /3 width=4 by lleq_fwd_bind_sn/
qed-.
-lemma lsx_fwd_flat_sn: ∀h,g,I,G,L,V,T,d. G ⊢ ⬊*[h, g, ⓕ{I}V.T, d] L →
- G ⊢ ⬊*[h, g, V, d] L.
-#h #g #I #G #L #V #T #d #H @(lsx_ind … H) -L
+lemma lsx_fwd_flat_sn: ∀h,g,I,G,L,V,T,l. G ⊢ ⬊*[h, g, ⓕ{I}V.T, l] L →
+ G ⊢ ⬊*[h, g, V, l] L.
+#h #g #I #G #L #V #T #l #H @(lsx_ind … H) -L
#L1 #_ #IHL1 @lsx_intro
#L2 #HL12 #HV @IHL1 /3 width=3 by lleq_fwd_flat_sn/
qed-.
-lemma lsx_fwd_flat_dx: ∀h,g,I,G,L,V,T,d. G ⊢ ⬊*[h, g, ⓕ{I}V.T, d] L →
- G ⊢ ⬊*[h, g, T, d] L.
-#h #g #I #G #L #V #T #d #H @(lsx_ind … H) -L
+lemma lsx_fwd_flat_dx: ∀h,g,I,G,L,V,T,l. G ⊢ ⬊*[h, g, ⓕ{I}V.T, l] L →
+ G ⊢ ⬊*[h, g, T, l] L.
+#h #g #I #G #L #V #T #l #H @(lsx_ind … H) -L
#L1 #_ #IHL1 @lsx_intro
#L2 #HL12 #HV @IHL1 /3 width=3 by lleq_fwd_flat_dx/
qed-.
-lemma lsx_fwd_pair_sn: ∀h,g,I,G,L,V,T,d. G ⊢ ⬊*[h, g, ②{I}V.T, d] L →
- G ⊢ ⬊*[h, g, V, d] L.
+lemma lsx_fwd_pair_sn: ∀h,g,I,G,L,V,T,l. G ⊢ ⬊*[h, g, ②{I}V.T, l] L →
+ G ⊢ ⬊*[h, g, V, l] L.
#h #g * /2 width=4 by lsx_fwd_bind_sn, lsx_fwd_flat_sn/
qed-.
(* Basic inversion lemmas ***************************************************)
-lemma lsx_inv_flat: ∀h,g,I,G,L,V,T,d. G ⊢ ⬊*[h, g, ⓕ{I}V.T, d] L →
- G ⊢ ⬊*[h, g, V, d] L ∧ G ⊢ ⬊*[h, g, T, d] L.
+lemma lsx_inv_flat: ∀h,g,I,G,L,V,T,l. G ⊢ ⬊*[h, g, ⓕ{I}V.T, l] L →
+ G ⊢ ⬊*[h, g, V, l] L ∧ G ⊢ ⬊*[h, g, T, l] L.
/3 width=3 by lsx_fwd_flat_sn, lsx_fwd_flat_dx, conj/ qed-.
(* alternative definition of lsx *)
definition lsxa: ∀h. sd h → relation4 ynat term genv lenv ≝
- λh,g,d,T,G. SN … (lpxs h g G) (lleq d T).
+ λh,g,l,T,G. SN … (lpxs h g G) (lleq l T).
interpretation
"extended strong normalization (local environment) alternative"
- 'SNAlt h g d T G L = (lsxa h g T d G L).
+ 'SNAlt h g l T G L = (lsxa h g T l G L).
(* Basic eliminators ********************************************************)
-lemma lsxa_ind: ∀h,g,G,T,d. ∀R:predicate lenv.
- (∀L1. G ⊢ ⬊⬊*[h, g, T, d] L1 →
- (∀L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → (L1 ≡[T, d] L2 → ⊥) → R L2) →
+lemma lsxa_ind: ∀h,g,G,T,l. ∀R:predicate lenv.
+ (∀L1. G ⊢ ⬊⬊*[h, g, T, l] L1 →
+ (∀L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → (L1 ≡[T, l] L2 → ⊥) → R L2) →
R L1
) →
- ∀L. G ⊢ ⬊⬊*[h, g, T, d] L → R L.
-#h #g #G #T #d #R #H0 #L1 #H elim H -L1
+ ∀L. G ⊢ ⬊⬊*[h, g, T, l] L → R L.
+#h #g #G #T #l #R #H0 #L1 #H elim H -L1
/5 width=1 by lleq_sym, SN_intro/
qed-.
(* Basic properties *********************************************************)
-lemma lsxa_intro: ∀h,g,G,L1,T,d.
- (∀L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → (L1 ≡[T, d] L2 → ⊥) → G ⊢ ⬊⬊*[h, g, T, d] L2) →
- G ⊢ ⬊⬊*[h, g, T, d] L1.
+lemma lsxa_intro: ∀h,g,G,L1,T,l.
+ (∀L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → (L1 ≡[T, l] L2 → ⊥) → G ⊢ ⬊⬊*[h, g, T, l] L2) →
+ G ⊢ ⬊⬊*[h, g, T, l] L1.
/5 width=1 by lleq_sym, SN_intro/ qed.
-fact lsxa_intro_aux: ∀h,g,G,L1,T,d.
- (∀L,L2. ⦃G, L⦄ ⊢ ➡*[h, g] L2 → L1 ≡[T, d] L → (L1 ≡[T, d] L2 → ⊥) → G ⊢ ⬊⬊*[h, g, T, d] L2) →
- G ⊢ ⬊⬊*[h, g, T, d] L1.
+fact lsxa_intro_aux: ∀h,g,G,L1,T,l.
+ (∀L,L2. ⦃G, L⦄ ⊢ ➡*[h, g] L2 → L1 ≡[T, l] L → (L1 ≡[T, l] L2 → ⊥) → G ⊢ ⬊⬊*[h, g, T, l] L2) →
+ G ⊢ ⬊⬊*[h, g, T, l] L1.
/4 width=3 by lsxa_intro/ qed-.
-lemma lsxa_lleq_trans: ∀h,g,T,G,L1,d. G ⊢ ⬊⬊*[h, g, T, d] L1 →
- ∀L2. L1 ≡[T, d] L2 → G ⊢ ⬊⬊*[h, g, T, d] L2.
-#h #g #T #G #L1 #d #H @(lsxa_ind … H) -L1
+lemma lsxa_lleq_trans: ∀h,g,T,G,L1,l. G ⊢ ⬊⬊*[h, g, T, l] L1 →
+ ∀L2. L1 ≡[T, l] L2 → G ⊢ ⬊⬊*[h, g, T, l] L2.
+#h #g #T #G #L1 #l #H @(lsxa_ind … H) -L1
#L1 #_ #IHL1 #L2 #HL12 @lsxa_intro
#K2 #HLK2 #HnLK2 elim (lleq_lpxs_trans … HLK2 … HL12) -HLK2
/5 width=4 by lleq_canc_sn, lleq_trans/
qed-.
-lemma lsxa_lpxs_trans: ∀h,g,T,G,L1,d. G ⊢ ⬊⬊*[h, g, T, d] L1 →
- ∀L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → G ⊢ ⬊⬊*[h, g, T, d] L2.
-#h #g #T #G #L1 #d #H @(lsxa_ind … H) -L1 #L1 #HL1 #IHL1 #L2 #HL12
-elim (lleq_dec T L1 L2 d) /3 width=4 by lsxa_lleq_trans/
+lemma lsxa_lpxs_trans: ∀h,g,T,G,L1,l. G ⊢ ⬊⬊*[h, g, T, l] L1 →
+ ∀L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → G ⊢ ⬊⬊*[h, g, T, l] L2.
+#h #g #T #G #L1 #l #H @(lsxa_ind … H) -L1 #L1 #HL1 #IHL1 #L2 #HL12
+elim (lleq_dec T L1 L2 l) /3 width=4 by lsxa_lleq_trans/
qed-.
-lemma lsxa_intro_lpx: ∀h,g,G,L1,T,d.
- (∀L2. ⦃G, L1⦄ ⊢ ➡[h, g] L2 → (L1 ≡[T, d] L2 → ⊥) → G ⊢ ⬊⬊*[h, g, T, d] L2) →
- G ⊢ ⬊⬊*[h, g, T, d] L1.
-#h #g #G #L1 #T #d #IH @lsxa_intro_aux
+lemma lsxa_intro_lpx: ∀h,g,G,L1,T,l.
+ (∀L2. ⦃G, L1⦄ ⊢ ➡[h, g] L2 → (L1 ≡[T, l] L2 → ⊥) → G ⊢ ⬊⬊*[h, g, T, l] L2) →
+ G ⊢ ⬊⬊*[h, g, T, l] L1.
+#h #g #G #L1 #T #l #IH @lsxa_intro_aux
#L #L2 #H @(lpxs_ind_dx … H) -L
[ #H destruct #H elim H //
-| #L0 #L elim (lleq_dec T L1 L d) /3 width=1 by/
+| #L0 #L elim (lleq_dec T L1 L l) /3 width=1 by/
#HnT #HL0 #HL2 #_ #HT #_ elim (lleq_lpx_trans … HL0 … HT) -L0
#L0 #HL10 #HL0 @(lsxa_lpxs_trans … HL2) -HL2
/5 width=3 by lsxa_lleq_trans, lleq_trans/
(* Main properties **********************************************************)
-theorem lsx_lsxa: ∀h,g,G,L,T,d. G ⊢ ⬊*[h, g, T, d] L → G ⊢ ⬊⬊*[h, g, T, d] L.
-#h #g #G #L #T #d #H @(lsx_ind … H) -L
+theorem lsx_lsxa: ∀h,g,G,L,T,l. G ⊢ ⬊*[h, g, T, l] L → G ⊢ ⬊⬊*[h, g, T, l] L.
+#h #g #G #L #T #l #H @(lsx_ind … H) -L
/4 width=1 by lsxa_intro_lpx/
qed.
(* Main inversion lemmas ****************************************************)
-theorem lsxa_inv_lsx: ∀h,g,G,L,T,d. G ⊢ ⬊⬊*[h, g, T, d] L → G ⊢ ⬊*[h, g, T, d] L.
-#h #g #G #L #T #d #H @(lsxa_ind … H) -L
+theorem lsxa_inv_lsx: ∀h,g,G,L,T,l. G ⊢ ⬊⬊*[h, g, T, l] L → G ⊢ ⬊*[h, g, T, l] L.
+#h #g #G #L #T #l #H @(lsxa_ind … H) -L
/4 width=1 by lsx_intro, lpx_lpxs/
qed-.
(* Advanced properties ******************************************************)
-lemma lsx_intro_alt: ∀h,g,G,L1,T,d.
- (∀L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → (L1 ≡[T, d] L2 → ⊥) → G ⊢ ⬊*[h, g, T, d] L2) →
- G ⊢ ⬊*[h, g, T, d] L1.
+lemma lsx_intro_alt: ∀h,g,G,L1,T,l.
+ (∀L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → (L1 ≡[T, l] L2 → ⊥) → G ⊢ ⬊*[h, g, T, l] L2) →
+ G ⊢ ⬊*[h, g, T, l] L1.
/6 width=1 by lsxa_inv_lsx, lsx_lsxa, lsxa_intro/ qed.
-lemma lsx_lpxs_trans: ∀h,g,G,L1,T,d. G ⊢ ⬊*[h, g, T, d] L1 →
- ∀L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → G ⊢ ⬊*[h, g, T, d] L2.
+lemma lsx_lpxs_trans: ∀h,g,G,L1,T,l. G ⊢ ⬊*[h, g, T, l] L1 →
+ ∀L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → G ⊢ ⬊*[h, g, T, l] L2.
/4 width=3 by lsxa_inv_lsx, lsx_lsxa, lsxa_lpxs_trans/ qed-.
(* Advanced eliminators *****************************************************)
-lemma lsx_ind_alt: ∀h,g,G,T,d. ∀R:predicate lenv.
- (∀L1. G ⊢ ⬊*[h, g, T, d] L1 →
- (∀L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → (L1 ≡[T, d] L2 → ⊥) → R L2) →
+lemma lsx_ind_alt: ∀h,g,G,T,l. ∀R:predicate lenv.
+ (∀L1. G ⊢ ⬊*[h, g, T, l] L1 →
+ (∀L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → (L1 ≡[T, l] L2 → ⊥) → R L2) →
R L1
) →
- ∀L. G ⊢ ⬊*[h, g, T, d] L → R L.
-#h #g #G #T #d #R #IH #L #H @(lsxa_ind h g G T d … L)
+ ∀L. G ⊢ ⬊*[h, g, T, l] L → R L.
+#h #g #G #T #l #R #IH #L #H @(lsxa_ind h g G T l … L)
/4 width=1 by lsxa_inv_lsx, lsx_lsxa/
qed-.
(* Advanced properties ******************************************************)
-lemma lsx_lref_be_lpxs: ∀h,g,I,G,K1,V,i,d. d ≤ yinj i → ⦃G, K1⦄ ⊢ ⬊*[h, g] V →
+lemma lsx_lref_be_lpxs: ∀h,g,I,G,K1,V,i,l. l ≤ yinj i → ⦃G, K1⦄ ⊢ ⬊*[h, g] V →
∀K2. G ⊢ ⬊*[h, g, V, 0] K2 → ⦃G, K1⦄ ⊢ ➡*[h, g] K2 →
- ∀L2. ⬇[i] L2 ≡ K2.ⓑ{I}V → G ⊢ ⬊*[h, g, #i, d] L2.
-#h #g #I #G #K1 #V #i #d #Hdi #H @(csx_ind_alt … H) -V
+ ∀L2. ⬇[i] L2 ≡ K2.ⓑ{I}V → G ⊢ ⬊*[h, g, #i, l] L2.
+#h #g #I #G #K1 #V #i #l #Hli #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_drop_conf … HLK0 … HL02) -HL02
]
qed.
-lemma lsx_lref_be: ∀h,g,I,G,K,V,i,d. d ≤ yinj i → ⦃G, K⦄ ⊢ ⬊*[h, g] V →
+lemma lsx_lref_be: ∀h,g,I,G,K,V,i,l. l ≤ yinj i → ⦃G, K⦄ ⊢ ⬊*[h, g] V →
G ⊢ ⬊*[h, g, V, 0] K →
- ∀L. ⬇[i] L ≡ K.ⓑ{I}V → G ⊢ ⬊*[h, g, #i, d] L.
+ ∀L. ⬇[i] L ≡ K.ⓑ{I}V → G ⊢ ⬊*[h, g, #i, l] L.
/2 width=8 by lsx_lref_be_lpxs/ qed.
(* Main properties **********************************************************)
-theorem csx_lsx: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → ∀d. G ⊢ ⬊*[h, g, T, d] L.
+theorem csx_lsx: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → ∀l. G ⊢ ⬊*[h, g, T, l] L.
#h #g #G #L #T @(fqup_wf_ind_eq … G L T) -G -L -T
#Z #Y #X #IH #G #L * * //
-[ #i #HG #HL #HT #H #d destruct
+[ #i #HG #HL #HT #H #l 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 (drop_O1_lt (Ⓕ) … Hi) -Hi
+ elim (ylt_split i l) /2 width=1 by lsx_lref_skip/
+ #Hli #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
+| #a #I #V #T #HG #HL #HT #H #l destruct
elim (csx_fwd_bind … H) -H /3 width=1 by lsx_bind/
-| #I #V #T #HG #HL #HT #H #d destruct
+| #I #V #T #HG #HL #HT #H #l destruct
elim (csx_fwd_flat … H) -H /3 width=1 by lsx_flat/
]
qed.
(* 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
+lemma lsx_lref_free: ∀h,g,G,L,l,i. |L| ≤ i → G ⊢ ⬊*[h, g, #i, l] L.
+#h #g #G #L1 #l #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
+lemma lsx_lref_skip: ∀h,g,G,L,l,i. yinj i < l → G ⊢ ⬊*[h, g, #i, l] L.
+#h #g #G #L1 #l #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 →
+lemma lsx_fwd_lref_be: ∀h,g,I,G,L,l,i. l ≤ yinj i → G ⊢ ⬊*[h, g, #i, l] 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
+#h #g #I #G #L #l #i #Hli #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/
+#L2 #HL12 #H2LK2 #H elim (lreq_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/
(* 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
+lemma lsx_lift_le: ∀h,g,G,K,T,U,lt,l,m. lt ≤ yinj l →
+ ⬆[l, m] T ≡ U → G ⊢ ⬊*[h, g, T, lt] K →
+ ∀L. ⬇[Ⓕ, l, m] L ≡ K → G ⊢ ⬊*[h, g, U, lt] L.
+#h #g #G #K #T #U #lt #l #m #Hltl #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
+lemma lsx_lift_ge: ∀h,g,G,K,T,U,lt,l,m. yinj l ≤ lt →
+ ⬆[l, m] T ≡ U → G ⊢ ⬊*[h, g, T, lt] K →
+ ∀L. ⬇[Ⓕ, l, m] L ≡ K → G ⊢ ⬊*[h, g, U, lt + m] L.
+#h #g #G #K #T #U #lt #l #m #Hllt #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/
(* 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
+lemma lsx_inv_lift_le: ∀h,g,G,L,T,U,lt,l,m. lt ≤ yinj l →
+ ⬆[l, m] T ≡ U → G ⊢ ⬊*[h, g, U, lt] L →
+ ∀K. ⬇[Ⓕ, l, m] L ≡ K → G ⊢ ⬊*[h, g, T, lt] K.
+#h #g #G #L #T #U #lt #l #m #Hltl #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
+lemma lsx_inv_lift_be: ∀h,g,G,L,T,U,lt,l,m. yinj l ≤ lt → lt ≤ l + m →
+ ⬆[l, m] T ≡ U → G ⊢ ⬊*[h, g, U, lt] L →
+ ∀K. ⬇[Ⓕ, l, m] L ≡ K → G ⊢ ⬊*[h, g, T, l] K.
+#h #g #G #L #T #U #lt #l #m #Hllt #Hltlm #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
+lemma lsx_inv_lift_ge: ∀h,g,G,L,T,U,lt,l,m. yinj l + yinj m ≤ lt →
+ ⬆[l, m] T ≡ U → G ⊢ ⬊*[h, g, U, lt] L →
+ ∀K. ⬇[Ⓕ, l, m] L ≡ K → G ⊢ ⬊*[h, g, T, lt-m] K.
+#h #g #G #L #T #U #lt #l #m #Hlmlt #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/
(* Advanced properties ******************************************************)
-lemma lsx_lleq_trans: ∀h,g,T,G,L1,d. G ⊢ ⬊*[h, g, T, d] L1 →
- ∀L2. L1 ≡[T, d] L2 → G ⊢ ⬊*[h, g, T, d] L2.
-#h #g #T #G #L1 #d #H @(lsx_ind … H) -L1
+lemma lsx_lleq_trans: ∀h,g,T,G,L1,l. G ⊢ ⬊*[h, g, T, l] L1 →
+ ∀L2. L1 ≡[T, l] L2 → G ⊢ ⬊*[h, g, T, l] L2.
+#h #g #T #G #L1 #l #H @(lsx_ind … H) -L1
#L1 #_ #IHL1 #L2 #HL12 @lsx_intro
#K2 #HLK2 #HnLK2 elim (lleq_lpx_trans … HLK2 … HL12) -HLK2
/5 width=4 by lleq_canc_sn, lleq_trans/
qed-.
-lemma lsx_lpx_trans: ∀h,g,T,G,L1,d. G ⊢ ⬊*[h, g, T, d] L1 →
- ∀L2. ⦃G, L1⦄ ⊢ ➡[h, g] L2 → G ⊢ ⬊*[h, g, T, d] L2.
-#h #g #T #G #L1 #d #H @(lsx_ind … H) -L1 #L1 #HL1 #IHL1 #L2 #HL12
-elim (lleq_dec T L1 L2 d) /3 width=4 by lsx_lleq_trans/
+lemma lsx_lpx_trans: ∀h,g,T,G,L1,l. G ⊢ ⬊*[h, g, T, l] L1 →
+ ∀L2. ⦃G, L1⦄ ⊢ ➡[h, g] L2 → G ⊢ ⬊*[h, g, T, l] L2.
+#h #g #T #G #L1 #l #H @(lsx_ind … H) -L1 #L1 #HL1 #IHL1 #L2 #HL12
+elim (lleq_dec T L1 L2 l) /3 width=4 by lsx_lleq_trans/
qed-.
-lemma lsx_leq_conf: ∀h,g,G,L1,T,d. G ⊢ ⬊*[h, g, T, d] L1 →
- ∀L2. L1 ⩬[d, ∞] L2 → G ⊢ ⬊*[h, g, T, d] L2.
-#h #g #G #L1 #T #d #H @(lsx_ind … H) -L1
+lemma lsx_lreq_conf: ∀h,g,G,L1,T,l. G ⊢ ⬊*[h, g, T, l] L1 →
+ ∀L2. L1 ⩬[l, ∞] L2 → G ⊢ ⬊*[h, g, T, l] L2.
+#h #g #G #L1 #T #l #H @(lsx_ind … H) -L1
#L1 #_ #IHL1 #L2 #HL12 @lsx_intro
-#L3 #HL23 #HnL23 elim (leq_lpx_trans_lleq … HL12 … HL23) -HL12 -HL23
+#L3 #HL23 #HnL23 elim (lreq_lpx_trans_lleq … HL12 … HL23) -HL12 -HL23
#L0 #HL03 #HL10 #H elim (H T) -H /4 width=4 by/
qed-.
(* Advanced forward lemmas **************************************************)
-lemma lsx_fwd_bind_dx: ∀h,g,a,I,G,L,V,T,d. G ⊢ ⬊*[h, g, ⓑ{a,I}V.T, d] L →
- G ⊢ ⬊*[h, g, T, ⫯d] L.ⓑ{I}V.
-#h #g #a #I #G #L #V1 #T #d #H @(lsx_ind … H) -L
+lemma lsx_fwd_bind_dx: ∀h,g,a,I,G,L,V,T,l. G ⊢ ⬊*[h, g, ⓑ{a,I}V.T, l] L →
+ G ⊢ ⬊*[h, g, T, ⫯l] L.ⓑ{I}V.
+#h #g #a #I #G #L #V1 #T #l #H @(lsx_ind … H) -L
#L1 #_ #IHL1 @lsx_intro
#Y #H #HT elim (lpx_inv_pair1 … H) -H
#L2 #V2 #HL12 #_ #H destruct
-@(lsx_leq_conf … (L2.ⓑ{I}V1)) /2 width=1 by leq_succ/
+@(lsx_lreq_conf … (L2.ⓑ{I}V1)) /2 width=1 by lreq_succ/
@IHL1 // #H @HT -IHL1 -HL12 -HT
-@(lleq_leq_trans … (L2.ⓑ{I}V1))
-/2 width=2 by lleq_fwd_bind_dx, leq_succ/
+@(lleq_lreq_trans … (L2.ⓑ{I}V1))
+/2 width=2 by lleq_fwd_bind_dx, lreq_succ/
qed-.
(* Advanced inversion lemmas ************************************************)
-lemma lsx_inv_bind: ∀h,g,a,I,G,L,V,T,d. G ⊢ ⬊*[h, g, ⓑ{a, I}V.T, d] L →
- G ⊢ ⬊*[h, g, V, d] L ∧ G ⊢ ⬊*[h, g, T, ⫯d] L.ⓑ{I}V.
+lemma lsx_inv_bind: ∀h,g,a,I,G,L,V,T,l. G ⊢ ⬊*[h, g, ⓑ{a, I}V.T, l] L →
+ G ⊢ ⬊*[h, g, V, l] L ∧ G ⊢ ⬊*[h, g, T, ⫯l] L.ⓑ{I}V.
/3 width=4 by lsx_fwd_bind_sn, lsx_fwd_bind_dx, conj/ qed-.
(* Advanced properties ******************************************************)
-fact lsx_bind_lpxs_aux: ∀h,g,a,I,G,L1,V,d. G ⊢ ⬊*[h, g, V, d] L1 →
- ∀Y,T. G ⊢ ⬊*[h, g, T, ⫯d] Y →
+fact lsx_bind_lpxs_aux: ∀h,g,a,I,G,L1,V,l. G ⊢ ⬊*[h, g, V, l] L1 →
+ ∀Y,T. G ⊢ ⬊*[h, g, T, ⫯l] Y →
∀L2. Y = L2.ⓑ{I}V → ⦃G, L1⦄ ⊢ ➡*[h, g] L2 →
- G ⊢ ⬊*[h, g, ⓑ{a,I}V.T, d] L2.
-#h #g #a #I #G #L1 #V #d #H @(lsx_ind_alt … H) -L1
+ G ⊢ ⬊*[h, g, ⓑ{a,I}V.T, l] L2.
+#h #g #a #I #G #L1 #V #l #H @(lsx_ind_alt … H) -L1
#L1 #HL1 #IHL1 #Y #T #H @(lsx_ind_alt … H) -Y
#Y #HY #IHY #L2 #H #HL12 destruct @lsx_intro_alt
#L0 #HL20 lapply (lpxs_trans … HL12 … HL20)
#HL10 #H elim (nlleq_inv_bind … H) -H [ -HL1 -IHY | -HY -IHL1 ]
-[ #HnV elim (lleq_dec V L1 L2 d)
+[ #HnV elim (lleq_dec V L1 L2 l)
[ #HV @(IHL1 … L0) /3 width=5 by lsx_lpxs_trans, lpxs_pair, lleq_canc_sn/ (**) (* full auto too slow *)
| -HnV -HL10 /4 width=5 by lsx_lpxs_trans, lpxs_pair/
]
]
qed-.
-lemma lsx_bind: ∀h,g,a,I,G,L,V,d. G ⊢ ⬊*[h, g, V, d] L →
- ∀T. G ⊢ ⬊*[h, g, T, ⫯d] L.ⓑ{I}V →
- G ⊢ ⬊*[h, g, ⓑ{a,I}V.T, d] L.
+lemma lsx_bind: ∀h,g,a,I,G,L,V,l. G ⊢ ⬊*[h, g, V, l] L →
+ ∀T. G ⊢ ⬊*[h, g, T, ⫯l] L.ⓑ{I}V →
+ G ⊢ ⬊*[h, g, ⓑ{a,I}V.T, l] L.
/2 width=3 by lsx_bind_lpxs_aux/ qed.
-lemma lsx_flat_lpxs: ∀h,g,I,G,L1,V,d. G ⊢ ⬊*[h, g, V, d] L1 →
- ∀L2,T. G ⊢ ⬊*[h, g, T, d] L2 → ⦃G, L1⦄ ⊢ ➡*[h, g] L2 →
- G ⊢ ⬊*[h, g, ⓕ{I}V.T, d] L2.
-#h #g #I #G #L1 #V #d #H @(lsx_ind_alt … H) -L1
+lemma lsx_flat_lpxs: ∀h,g,I,G,L1,V,l. G ⊢ ⬊*[h, g, V, l] L1 →
+ ∀L2,T. G ⊢ ⬊*[h, g, T, l] L2 → ⦃G, L1⦄ ⊢ ➡*[h, g] L2 →
+ G ⊢ ⬊*[h, g, ⓕ{I}V.T, l] L2.
+#h #g #I #G #L1 #V #l #H @(lsx_ind_alt … H) -L1
#L1 #HL1 #IHL1 #L2 #T #H @(lsx_ind_alt … H) -L2
#L2 #HL2 #IHL2 #HL12 @lsx_intro_alt
#L0 #HL20 lapply (lpxs_trans … HL12 … HL20)
#HL10 #H elim (nlleq_inv_flat … H) -H [ -HL1 -IHL2 | -HL2 -IHL1 ]
-[ #HnV elim (lleq_dec V L1 L2 d)
+[ #HnV elim (lleq_dec V L1 L2 l)
[ #HV @(IHL1 … L0) /3 width=3 by lsx_lpxs_trans, lleq_canc_sn/ (**) (* full auto too slow: 47s *)
| -HnV -HL10 /3 width=4 by lsx_lpxs_trans/
]
]
qed-.
-lemma lsx_flat: ∀h,g,I,G,L,V,d. G ⊢ ⬊*[h, g, V, d] L →
- ∀T. G ⊢ ⬊*[h, g, T, d] L → G ⊢ ⬊*[h, g, ⓕ{I}V.T, d] L.
+lemma lsx_flat: ∀h,g,I,G,L,V,l. G ⊢ ⬊*[h, g, V, l] L →
+ ∀T. G ⊢ ⬊*[h, g, T, l] L → G ⊢ ⬊*[h, g, ⓕ{I}V.T, l] L.
/2 width=3 by lsx_flat_lpxs/ qed.
(* STRATIFIED DECOMPOSED PARALLEL COMPUTATION ON TERMS **********************)
definition scpds: ∀h. sd h → nat → relation4 genv lenv term term ≝
- λh,g,l2,G,L,T1,T2.
- ∃∃T,l1. l2 ≤ l1 & ⦃G, L⦄ ⊢ T1 ▪[h, g] l1 & ⦃G, L⦄ ⊢ T1 •*[h, l2] T & ⦃G, L⦄ ⊢ T ➡* T2.
+ λh,g,d2,G,L,T1,T2.
+ ∃∃T,d1. d2 ≤ d1 & ⦃G, L⦄ ⊢ T1 ▪[h, g] d1 & ⦃G, L⦄ ⊢ T1 •*[h, d2] T & ⦃G, L⦄ ⊢ T ➡* T2.
interpretation "stratified decomposed parallel computation (term)"
- 'DPRedStar h g l G L T1 T2 = (scpds h g l G L T1 T2).
+ 'DPRedStar h g d G L T1 T2 = (scpds h g d G L T1 T2).
(* Basic properties *********************************************************)
-lemma sta_cprs_scpds: ∀h,g,G,L,T1,T,T2,l. ⦃G, L⦄ ⊢ T1 ▪[h, g] l+1 → ⦃G, L⦄ ⊢ T1 •*[h, 1] T →
+lemma sta_cprs_scpds: ∀h,g,G,L,T1,T,T2,d. ⦃G, L⦄ ⊢ T1 ▪[h, g] d+1 → ⦃G, L⦄ ⊢ T1 •*[h, 1] T →
⦃G, L⦄ ⊢ T ➡* T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g, 1] T2.
/2 width=6 by ex4_2_intro/ qed.
-lemma lstas_scpds: ∀h,g,G,L,T1,T2,l1. ⦃G, L⦄ ⊢ T1 ▪[h, g] l1 →
- ∀l2. l2 ≤ l1 → ⦃G, L⦄ ⊢ T1 •*[h, l2] T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g, l2] T2.
+lemma lstas_scpds: ∀h,g,G,L,T1,T2,d1. ⦃G, L⦄ ⊢ T1 ▪[h, g] d1 →
+ ∀d2. d2 ≤ d1 → ⦃G, L⦄ ⊢ T1 •*[h, d2] T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g, d2] T2.
/2 width=6 by ex4_2_intro/ qed.
-lemma scpds_strap1: ∀h,g,G,L,T1,T,T2,l.
- ⦃G, L⦄ ⊢ T1 •*➡*[h, g, l] T → ⦃G, L⦄ ⊢ T ➡ T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g, l] T2.
-#h #g #G #L #T1 #T #T2 #l * /3 width=8 by cprs_strap1, ex4_2_intro/
+lemma scpds_strap1: ∀h,g,G,L,T1,T,T2,d.
+ ⦃G, L⦄ ⊢ T1 •*➡*[h, g, d] T → ⦃G, L⦄ ⊢ T ➡ T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g, d] T2.
+#h #g #G #L #T1 #T #T2 #d * /3 width=8 by cprs_strap1, ex4_2_intro/
qed.
(* Basic forward lemmas *****************************************************)
(* Properties on atomic arity assignment for terms **************************)
-lemma scpds_aaa_conf: ∀h,g,G,L,l. Conf3 … (aaa G L) (scpds h g l G L).
-#h #g #G #L #l #A #T #HT #U * /3 width=6 by lstas_aaa_conf, cprs_aaa_conf/
+lemma scpds_aaa_conf: ∀h,g,G,L,d. Conf3 … (aaa G L) (scpds h g d G L).
+#h #g #G #L #d #A #T #HT #U * /3 width=6 by lstas_aaa_conf, cprs_aaa_conf/
qed.
(* Relocation properties ****************************************************)
-lemma scpds_lift: ∀h,g,G,l. l_liftable (scpds h g l G).
-#h #g #G #l2 #K #T1 #T2 * #T #l1 #Hl21 #Hl1 #HT1 #HT2 #L #s #d #e
-elim (lift_total T d e)
+lemma scpds_lift: ∀h,g,G,d. d_liftable (scpds h g d G).
+#h #g #G #d2 #K #T1 #T2 * #T #d1 #Hd21 #Hd1 #HT1 #HT2 #L #s #l #m
+elim (lift_total T l m)
/3 width=15 by cprs_lift, da_lift, lstas_lift, ex4_2_intro/
qed.
-lemma scpds_inv_lift1: ∀h,g,G,l. l_deliftable_sn (scpds h g l G).
-#h #g #G #l2 #L #U1 #U2 * #U #l1 #Hl21 #Hl1 #HU1 #HU2 #K #s #d #e #HLK #T1 #HTU1
-lapply (da_inv_lift … Hl1 … HLK … HTU1) -Hl1 #Hl1
+lemma scpds_inv_lift1: ∀h,g,G,d. d_deliftable_sn (scpds h g d G).
+#h #g #G #d2 #L #U1 #U2 * #U #d1 #Hd21 #Hd1 #HU1 #HU2 #K #s #l #m #HLK #T1 #HTU1
+lapply (da_inv_lift … Hd1 … HLK … HTU1) -Hd1 #Hd1
elim (lstas_inv_lift1 … HU1 … HLK … HTU1) -U1 #T #HTU #HT1
elim (cprs_inv_lift1 … HU2 … HLK … HTU) -U -L
/3 width=8 by ex4_2_intro, ex2_intro/
(* Advanced properties ******************************************************)
-lemma scpds_strap2: ∀h,g,G,L,T1,T,T2,l1,l. ⦃G, L⦄ ⊢ T1 ▪[h, g] l1+1 →
- ⦃G, L⦄ ⊢ T1 •*[h, 1] T → ⦃G, L⦄ ⊢ T •*➡*[h, g, l] T2 →
- ⦃G, L⦄ ⊢ T1 •*➡*[h, g, l+1] T2.
-#h #g #G #L #T1 #T #T2 #l1 #l #Hl1 #HT1 *
-#T0 #l0 #Hl0 #HTl0 #HT0 #HT02
-lapply (lstas_da_conf … HT1 … Hl1) <minus_plus_m_m #HTl1
-lapply (da_mono … HTl0 … HTl1) -HTl0 -HTl1 #H destruct
+lemma scpds_strap2: ∀h,g,G,L,T1,T,T2,d1,d. ⦃G, L⦄ ⊢ T1 ▪[h, g] d1+1 →
+ ⦃G, L⦄ ⊢ T1 •*[h, 1] T → ⦃G, L⦄ ⊢ T •*➡*[h, g, d] T2 →
+ ⦃G, L⦄ ⊢ T1 •*➡*[h, g, d+1] T2.
+#h #g #G #L #T1 #T #T2 #d1 #d #Hd1 #HT1 *
+#T0 #d0 #Hd0 #HTd0 #HT0 #HT02
+lapply (lstas_da_conf … HT1 … Hd1) <minus_plus_m_m #HTd1
+lapply (da_mono … HTd0 … HTd1) -HTd0 -HTd1 #H destruct
lapply (lstas_trans … HT1 … HT0) -T >commutative_plus
/3 width=6 by le_S_S, ex4_2_intro/
qed.
-lemma scpds_cprs_trans: ∀h,g,G,L,T1,T,T2,l.
- ⦃G, L⦄ ⊢ T1 •*➡*[h, g, l] T → ⦃G, L⦄ ⊢ T ➡* T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g, l] T2.
-#h #g #G #L #T1 #T #T2 #l * /3 width=8 by cprs_trans, ex4_2_intro/
+lemma scpds_cprs_trans: ∀h,g,G,L,T1,T,T2,d.
+ ⦃G, L⦄ ⊢ T1 •*➡*[h, g, d] T → ⦃G, L⦄ ⊢ T ➡* T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g, d] T2.
+#h #g #G #L #T1 #T #T2 #d * /3 width=8 by cprs_trans, ex4_2_intro/
qed-.
-lemma lstas_scpds_trans: ∀h,g,G,L,T1,T,T2,l1,l2,l.
- l2 ≤ l1 → ⦃G, L⦄ ⊢ T1 ▪[h, g] l1 →
- ⦃G, L⦄ ⊢ T1 •*[h, l2] T → ⦃G, L⦄ ⊢ T •*➡*[h, g, l] T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g, l2+l] T2.
-#h #g #G #L #T1 #T #T2 #l1 #l2 #l #Hl21 #HTl1 #HT1 * #T0 #l0 #Hl0 #HTl0 #HT0 #HT02
-lapply (lstas_da_conf … HT1 … HTl1) #HTl12
-lapply (da_mono … HTl12 … HTl0) -HTl12 -HTl0 #H destruct
-lapply (le_minus_to_plus_r … Hl21 Hl0) -Hl21 -Hl0
+lemma lstas_scpds_trans: ∀h,g,G,L,T1,T,T2,d1,d2,d.
+ d2 ≤ d1 → ⦃G, L⦄ ⊢ T1 ▪[h, g] d1 →
+ ⦃G, L⦄ ⊢ T1 •*[h, d2] T → ⦃G, L⦄ ⊢ T •*➡*[h, g, d] T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g, d2+d] T2.
+#h #g #G #L #T1 #T #T2 #d1 #d2 #d #Hd21 #HTd1 #HT1 * #T0 #d0 #Hd0 #HTd0 #HT0 #HT02
+lapply (lstas_da_conf … HT1 … HTd1) #HTd12
+lapply (da_mono … HTd12 … HTd0) -HTd12 -HTd0 #H destruct
+lapply (le_minus_to_plus_r … Hd21 Hd0) -Hd21 -Hd0
/3 width=7 by lstas_trans, ex4_2_intro/
qed-.
(* Advanced inversion lemmas ************************************************)
-lemma scpds_inv_abst1: ∀h,g,a,G,L,V1,T1,U2,l. ⦃G, L⦄ ⊢ ⓛ{a}V1.T1 •*➡*[h, g, l] U2 →
- ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡* V2 & ⦃G, L.ⓛV1⦄ ⊢ T1 •*➡*[h, g, l] T2 &
+lemma scpds_inv_abst1: ∀h,g,a,G,L,V1,T1,U2,d. ⦃G, L⦄ ⊢ ⓛ{a}V1.T1 •*➡*[h, g, d] U2 →
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡* V2 & ⦃G, L.ⓛV1⦄ ⊢ T1 •*➡*[h, g, d] T2 &
U2 = ⓛ{a}V2.T2.
-#h #g #a #G #L #V1 #T1 #U2 #l2 * #X #l1 #Hl21 #Hl1 #H1 #H2
-lapply (da_inv_bind … Hl1) -Hl1 #Hl1
+#h #g #a #G #L #V1 #T1 #U2 #d2 * #X #d1 #Hd21 #Hd1 #H1 #H2
+lapply (da_inv_bind … Hd1) -Hd1 #Hd1
elim (lstas_inv_bind1 … H1) -H1 #U #HTU1 #H destruct
elim (cprs_inv_abst1 … H2) -H2 #V2 #T2 #HV12 #HUT2 #H destruct
/3 width=6 by ex4_2_intro, ex3_2_intro/
qed-.
-lemma scpds_inv_abbr_abst: ∀h,g,a1,a2,G,L,V1,W2,T1,T2,l. ⦃G, L⦄ ⊢ ⓓ{a1}V1.T1 •*➡*[h, g, l] ⓛ{a2}W2.T2 →
- ∃∃T. ⦃G, L.ⓓV1⦄ ⊢ T1 •*➡*[h, g, l] T & ⬆[0, 1] ⓛ{a2}W2.T2 ≡ T & a1 = true.
-#h #g #a1 #a2 #G #L #V1 #W2 #T1 #T2 #l2 * #X #l1 #Hl21 #Hl1 #H1 #H2
-lapply (da_inv_bind … Hl1) -Hl1 #Hl1
+lemma scpds_inv_abbr_abst: ∀h,g,a1,a2,G,L,V1,W2,T1,T2,d. ⦃G, L⦄ ⊢ ⓓ{a1}V1.T1 •*➡*[h, g, d] ⓛ{a2}W2.T2 →
+ ∃∃T. ⦃G, L.ⓓV1⦄ ⊢ T1 •*➡*[h, g, d] T & ⬆[0, 1] ⓛ{a2}W2.T2 ≡ T & a1 = true.
+#h #g #a1 #a2 #G #L #V1 #W2 #T1 #T2 #d2 * #X #d1 #Hd21 #Hd1 #H1 #H2
+lapply (da_inv_bind … Hd1) -Hd1 #Hd1
elim (lstas_inv_bind1 … H1) -H1 #U1 #HTU1 #H destruct
elim (cprs_inv_abbr1 … H2) -H2 *
[ #V2 #U2 #HV12 #HU12 #H destruct
]
qed-.
-lemma scpds_inv_lstas_eq: ∀h,g,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 •*➡*[h, g, l] T2 →
- ∀T. ⦃G, L⦄ ⊢ T1 •*[h, l] T → ⦃G, L⦄ ⊢ T ➡* T2.
-#h #g #G #L #T1 #T2 #l2 *
-#T0 #l1 #_ #_ #HT10 #HT02 #T #HT1
+lemma scpds_inv_lstas_eq: ∀h,g,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 •*➡*[h, g, d] T2 →
+ ∀T. ⦃G, L⦄ ⊢ T1 •*[h, d] T → ⦃G, L⦄ ⊢ T ➡* T2.
+#h #g #G #L #T1 #T2 #d2 *
+#T0 #d1 #_ #_ #HT10 #HT02 #T #HT1
lapply (lstas_mono … HT10 … HT1) #H destruct //
qed-.
(* Advanced forward lemmas **************************************************)
-lemma scpds_fwd_cpxs: ∀h,g,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 •*➡*[h, g, l] T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2.
-#h #g #G #L #T1 #T2 #l * /3 width=5 by cpxs_trans, lstas_cpxs, cprs_cpxs/
+lemma scpds_fwd_cpxs: ∀h,g,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 •*➡*[h, g, d] T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2.
+#h #g #G #L #T1 #T2 #d * /3 width=5 by cpxs_trans, lstas_cpxs, cprs_cpxs/
qed-.
(* Main properties **********************************************************)
-theorem scpds_conf_eq: ∀h,g,G,L,T0,T1,l. ⦃G, L⦄ ⊢ T0 •*➡*[h, g, l] T1 →
- ∀T2. ⦃G, L⦄ ⊢ T0 •*➡*[h, g, l] T2 →
+theorem scpds_conf_eq: ∀h,g,G,L,T0,T1,d. ⦃G, L⦄ ⊢ T0 •*➡*[h, g, d] T1 →
+ ∀T2. ⦃G, L⦄ ⊢ T0 •*➡*[h, g, d] T2 →
∃∃T. ⦃G, L⦄ ⊢ T1 ➡* T & ⦃G, L⦄ ⊢ T2 ➡* T.
-#h #g #G #L #T0 #T1 #l0 * #U1 #l1 #_ #_ #H1 #HUT1 #T2 * #U2 #l2 #_ #_ #H2 #HUT2 -l1 -l2
-lapply (lstas_mono … H1 … H2) #H destruct -h -l0 /2 width=3 by cprs_conf/
+#h #g #G #L #T0 #T1 #d0 * #U1 #d1 #_ #_ #H1 #HUT1 #T2 * #U2 #d2 #_ #_ #H2 #HUT2 -d1 -d2
+lapply (lstas_mono … H1 … H2) #H destruct -h -d0 /2 width=3 by cprs_conf/
qed-.
| lsubsv_atom: lsubsv h g G (⋆) (⋆)
| lsubsv_pair: ∀I,L1,L2,V. lsubsv h g G L1 L2 →
lsubsv h g G (L1.ⓑ{I}V) (L2.ⓑ{I}V)
-| lsubsv_beta: ∀L1,L2,W,V,l1. ⦃G, L1⦄ ⊢ ⓝW.V ¡[h, g, l1] → ⦃G, L2⦄ ⊢ W ¡[h, g] →
- ⦃G, L1⦄ ⊢ V ▪[h, g] l1+1 → ⦃G, L2⦄ ⊢ W ▪[h, g] l1 →
+| lsubsv_beta: ∀L1,L2,W,V,d1. ⦃G, L1⦄ ⊢ ⓝW.V ¡[h, g, d1] → ⦃G, L2⦄ ⊢ W ¡[h, g] →
+ ⦃G, L1⦄ ⊢ V ▪[h, g] d1+1 → ⦃G, L2⦄ ⊢ W ▪[h, g] d1 →
lsubsv h g G L1 L2 → lsubsv h g G (L1.ⓓⓝW.V) (L2.ⓛW)
.
#h #g #G #L1 #L2 * -L1 -L2
[ //
| #I #L1 #L2 #V #_ #H destruct
-| #L1 #L2 #W #V #l1 #_ #_ #_ #_ #_ #H destruct
+| #L1 #L2 #W #V #d1 #_ #_ #_ #_ #_ #H destruct
]
qed-.
fact lsubsv_inv_pair1_aux: ∀h,g,G,L1,L2. G ⊢ L1 ⫃¡[h, g] L2 →
∀I,K1,X. L1 = K1.ⓑ{I}X →
(∃∃K2. G ⊢ K1 ⫃¡[h, g] K2 & L2 = K2.ⓑ{I}X) ∨
- ∃∃K2,W,V,l1. ⦃G, K1⦄ ⊢ ⓝW.V ¡[h, g, l1] & ⦃G, K2⦄ ⊢ W ¡[h, g] &
- ⦃G, K1⦄ ⊢ V ▪[h, g] l1+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] l1 &
+ ∃∃K2,W,V,d1. ⦃G, K1⦄ ⊢ ⓝW.V ¡[h, g, d1] & ⦃G, K2⦄ ⊢ W ¡[h, g] &
+ ⦃G, K1⦄ ⊢ V ▪[h, g] d1+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] d1 &
G ⊢ K1 ⫃¡[h, g] K2 &
I = Abbr & L2 = K2.ⓛW & X = ⓝW.V.
#h #g #G #L1 #L2 * -L1 -L2
[ #J #K1 #X #H destruct
| #I #L1 #L2 #V #HL12 #J #K1 #X #H destruct /3 width=3 by ex2_intro, or_introl/
-| #L1 #L2 #W #V #l1 #HWV #HW #HVl1 #HWl1 #HL12 #J #K1 #X #H destruct /3 width=11 by or_intror, ex8_4_intro/
+| #L1 #L2 #W #V #d1 #HWV #HW #HVd1 #HWd1 #HL12 #J #K1 #X #H destruct /3 width=11 by or_intror, ex8_4_intro/
]
qed-.
lemma lsubsv_inv_pair1: ∀h,g,I,G,K1,L2,X. G ⊢ K1.ⓑ{I}X ⫃¡[h, g] L2 →
(∃∃K2. G ⊢ K1 ⫃¡[h, g] K2 & L2 = K2.ⓑ{I}X) ∨
- ∃∃K2,W,V,l1. ⦃G, K1⦄ ⊢ ⓝW.V ¡[h, g, l1] & ⦃G, K2⦄ ⊢ W ¡[h, g] &
- ⦃G, K1⦄ ⊢ V ▪[h, g] l1+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] l1 &
+ ∃∃K2,W,V,d1. ⦃G, K1⦄ ⊢ ⓝW.V ¡[h, g, d1] & ⦃G, K2⦄ ⊢ W ¡[h, g] &
+ ⦃G, K1⦄ ⊢ V ▪[h, g] d1+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] d1 &
G ⊢ K1 ⫃¡[h, g] K2 &
I = Abbr & L2 = K2.ⓛW & X = ⓝW.V.
/2 width=3 by lsubsv_inv_pair1_aux/ qed-.
#h #g #G #L1 #L2 * -L1 -L2
[ //
| #I #L1 #L2 #V #_ #H destruct
-| #L1 #L2 #W #V #l1 #_ #_ #_ #_ #_ #H destruct
+| #L1 #L2 #W #V #d1 #_ #_ #_ #_ #_ #H destruct
]
qed-.
fact lsubsv_inv_pair2_aux: ∀h,g,G,L1,L2. G ⊢ L1 ⫃¡[h, g] L2 →
∀I,K2,W. L2 = K2.ⓑ{I}W →
(∃∃K1. G ⊢ K1 ⫃¡[h, g] K2 & L1 = K1.ⓑ{I}W) ∨
- ∃∃K1,V,l1. ⦃G, K1⦄ ⊢ ⓝW.V ¡[h, g, l1] & ⦃G, K2⦄ ⊢ W ¡[h, g] &
- ⦃G, K1⦄ ⊢ V ▪[h, g] l1+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] l1 &
+ ∃∃K1,V,d1. ⦃G, K1⦄ ⊢ ⓝW.V ¡[h, g, d1] & ⦃G, K2⦄ ⊢ W ¡[h, g] &
+ ⦃G, K1⦄ ⊢ V ▪[h, g] d1+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] d1 &
G ⊢ K1 ⫃¡[h, g] K2 & I = Abst & L1 = K1.ⓓⓝW.V.
#h #g #G #L1 #L2 * -L1 -L2
[ #J #K2 #U #H destruct
| #I #L1 #L2 #V #HL12 #J #K2 #U #H destruct /3 width=3 by ex2_intro, or_introl/
-| #L1 #L2 #W #V #l1 #HWV #HW #HVl1 #HWl1 #HL12 #J #K2 #U #H destruct /3 width=8 by or_intror, ex7_3_intro/
+| #L1 #L2 #W #V #d1 #HWV #HW #HVd1 #HWd1 #HL12 #J #K2 #U #H destruct /3 width=8 by or_intror, ex7_3_intro/
]
qed-.
lemma lsubsv_inv_pair2: ∀h,g,I,G,L1,K2,W. G ⊢ L1 ⫃¡[h, g] K2.ⓑ{I}W →
(∃∃K1. G ⊢ K1 ⫃¡[h, g] K2 & L1 = K1.ⓑ{I}W) ∨
- ∃∃K1,V,l1. ⦃G, K1⦄ ⊢ ⓝW.V ¡[h, g, l1] & ⦃G, K2⦄ ⊢ W ¡[h, g] &
- ⦃G, K1⦄ ⊢ V ▪[h, g] l1+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] l1 &
+ ∃∃K1,V,d1. ⦃G, K1⦄ ⊢ ⓝW.V ¡[h, g, d1] & ⦃G, K2⦄ ⊢ W ¡[h, g] &
+ ⦃G, K1⦄ ⊢ V ▪[h, g] d1+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] d1 &
G ⊢ K1 ⫃¡[h, g] K2 & I = Abst & L1 = K1.ⓓⓝW.V.
/2 width=3 by lsubsv_inv_pair2_aux/ qed-.
(* Note: the constant 0 cannot be generalized *)
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.
+ ∀K1,s,m. ⬇[s, 0, m] L1 ≡ K1 →
+ ∃∃K2. G ⊢ K1 ⫃¡[h, g] K2 & ⬇[s, 0, m] 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 (drop_inv_O1_pair1 … H) -H * #He #HLK1
+| #I #L1 #L2 #V #_ #IHL12 #K1 #s #m #H
+ elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK1
[ destruct
elim (IHL12 L1 s 0) -IHL12 // #X #HL12 #H
<(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 #l1 #HWV #HW #HVl1 #HWl1 #_ #IHL12 #K1 #s #e #H
- elim (drop_inv_O1_pair1 … H) -H * #He #HLK1
+| #L1 #L2 #W #V #d1 #HWV #HW #HVd1 #HWd1 #_ #IHL12 #K1 #s #m #H
+ elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK1
[ destruct
elim (IHL12 L1 s 0) -IHL12 // #X #HL12 #H
<(drop_inv_O2 … H) in HL12; -H /3 width=4 by lsubsv_beta, drop_pair, ex2_intro/
(* Note: the constant 0 cannot be generalized *)
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.
+ ∀K2,s, m. ⬇[s, 0, m] L2 ≡ K2 →
+ ∃∃K1. G ⊢ K1 ⫃¡[h, g] K2 & ⬇[s, 0, m] 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 (drop_inv_O1_pair1 … H) -H * #He #HLK2
+| #I #L1 #L2 #V #_ #IHL12 #K2 #s #m #H
+ elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK2
[ destruct
elim (IHL12 L2 s 0) -IHL12 // #X #HL12 #H
<(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 #l1 #HWV #HW #HVl1 #HWl1 #_ #IHL12 #K2 #s #e #H
- elim (drop_inv_O1_pair1 … H) -H * #He #HLK2
+| #L1 #L2 #W #V #d1 #HWV #HW #HVd1 #HWd1 #_ #IHL12 #K2 #s #m #H
+ elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK2
[ destruct
elim (IHL12 L2 s 0) -IHL12 // #X #HL12 #H
<(drop_inv_O2 … H) in HL12; -H /3 width=4 by lsubsv_beta, drop_pair, ex2_intro/
(* Properties on nat-iterated static type assignment ************************)
-lemma lsubsv_lstas_trans: ∀h,g,G,L2,T,U2,l2. ⦃G, L2⦄ ⊢ T •*[h, l2] U2 →
- ∀l1. l2 ≤ l1 → ⦃G, L2⦄ ⊢ T ▪[h, g] l1 →
+lemma lsubsv_lstas_trans: ∀h,g,G,L2,T,U2,d2. ⦃G, L2⦄ ⊢ T •*[h, d2] U2 →
+ ∀d1. d2 ≤ d1 → ⦃G, L2⦄ ⊢ T ▪[h, g] d1 →
∀L1. G ⊢ L1 ⫃¡[h, g] L2 →
- ∃∃U1. ⦃G, L1⦄ ⊢ T •*[h, l2] U1 & ⦃G, L1⦄ ⊢ U1 ⬌* U2.
-#h #g #G #L2 #T #U #l2 #H elim H -G -L2 -T -U -l2
+ ∃∃U1. ⦃G, L1⦄ ⊢ T •*[h, d2] U1 & ⦃G, L1⦄ ⊢ U1 ⬌* U2.
+#h #g #G #L2 #T #U #d2 #H elim H -G -L2 -T -U -d2
[ /2 width=3 by ex2_intro/
-| #G #L2 #K2 #V #W #U #i #l2 #HLK2 #_ #HWU #IHVW #l1 #Hl21 #Hl1 #L1 #HL12
- elim (da_inv_lref … Hl1) -Hl1 * #K0 #V0 [| #l0 ] #HK0 #HV0
+| #G #L2 #K2 #V #W #U #i #d2 #HLK2 #_ #HWU #IHVW #d1 #Hd21 #Hd1 #L1 #HL12
+ elim (da_inv_lref … Hd1) -Hd1 * #K0 #V0 [| #d0 ] #HK0 #HV0
lapply (drop_mono … HK0 … HLK2) -HK0 #H destruct
elim (lsubsv_drop_O1_trans … HL12 … HLK2) -L2 #Y #H #HLK1
elim (lsubsv_inv_pair2 … H) -H * #K1 [ | -HWU -IHVW -HLK1 ]
[ #HK12 #H destruct
- elim (IHVW … Hl21 HV0 … HK12) -K2 -l1 #T #HVT #HTW
+ elim (IHVW … Hd21 HV0 … HK12) -K2 -d1 #T #HVT #HTW
lapply (drop_fwd_drop2 … HLK1) #H
elim (lift_total T 0 (i+1))
/3 width=12 by lstas_ldef, cpcs_lift, ex2_intro/
- | #V0 #l0 #_ #_ #_ #_ #_ #H destruct
+ | #V0 #d0 #_ #_ #_ #_ #_ #H destruct
]
-| #G #L2 #K2 #V #W #i #HLK2 #_ #IHVW #l1 #_ #Hl1 #L1 #HL12
- elim (da_inv_lref … Hl1) -Hl1 * #K0 #V0 [| #l0 ] #HK0 #HV0 [| #H1 ]
+| #G #L2 #K2 #V #W #i #HLK2 #_ #IHVW #d1 #_ #Hd1 #L1 #HL12
+ elim (da_inv_lref … Hd1) -Hd1 * #K0 #V0 [| #d0 ] #HK0 #HV0 [| #H1 ]
lapply (drop_mono … HK0 … HLK2) -HK0 #H2 destruct
elim (lsubsv_drop_O1_trans … HL12 … HLK2) -L2 #Y #H #HLK1
elim (lsubsv_inv_pair2 … H) -H * #K1
[ #HK12 #H destruct
elim (IHVW … HV0 … HK12) -K2 /3 width=5 by lstas_zero, ex2_intro/
- | #V1 #l1 #_ #_ #HV1 #HV #HK12 #_ #H destruct
+ | #V1 #d1 #_ #_ #HV1 #HV #HK12 #_ #H destruct
lapply (da_mono … HV0 … HV) -HV #H destruct
elim (da_lstas … HV1 0) -HV1 #W1 #HVW1 #_
elim (lift_total W1 0 (i+1)) #U1 #HWU1
@cpcs_cprs_sn @(cprs_delta … HLK1 … HWU1)
/4 width=2 by cprs_strap1, cpr_cprs, lstas_cpr, cpr_eps/
]
-| #G #L2 #K2 #V #W #U #i #l2 #HLK2 #_ #HWU #IHVW #l1 #Hl21 #Hl1 #L1 #HL12
- elim (da_inv_lref … Hl1) -Hl1 * #K0 #V0 [| #l0 ] #HK0 #HV0 [| #H1 ]
+| #G #L2 #K2 #V #W #U #i #d2 #HLK2 #_ #HWU #IHVW #d1 #Hd21 #Hd1 #L1 #HL12
+ elim (da_inv_lref … Hd1) -Hd1 * #K0 #V0 [| #d0 ] #HK0 #HV0 [| #H1 ]
lapply (drop_mono … HK0 … HLK2) -HK0 #H2 destruct
- lapply (le_plus_to_le_r … Hl21) -Hl21 #Hl21
+ lapply (le_plus_to_le_r … Hd21) -Hd21 #Hd21
elim (lsubsv_drop_O1_trans … HL12 … HLK2) -L2 #Y #H #HLK1
elim (lsubsv_inv_pair2 … H) -H * #K1
[ #HK12 #H destruct
- elim (IHVW … Hl21 HV0 … HK12) -K2 -Hl21 #X
+ elim (IHVW … Hd21 HV0 … HK12) -K2 -Hd21 #X
lapply (drop_fwd_drop2 … HLK1)
elim (lift_total X 0 (i+1))
/3 width=12 by lstas_succ, cpcs_lift, ex2_intro/
- | #V1 #l1 #H0 #_ #HV1 #HV #HK12 #_ #H destruct
+ | #V1 #d1 #H0 #_ #HV1 #HV #HK12 #_ #H destruct
lapply (da_mono … HV0 … HV) -HV #H destruct
elim (shnv_inv_cast … H0) -H0 #_ #_ #H
- lapply (H … Hl21) -H #HVV1
- elim (IHVW … Hl21 HV0 … HK12) -K2 -Hl21 #X #HVX #HXW
- elim (da_lstas … HV1 (l2+1)) -HV1 #X1 #HVX1 #_
+ lapply (H … Hd21) -H #HVV1
+ elim (IHVW … Hd21 HV0 … HK12) -K2 -Hd21 #X #HVX #HXW
+ elim (da_lstas … HV1 (d2+1)) -HV1 #X1 #HVX1 #_
lapply (scpes_inv_lstas_eq … HVV1 … HVX … HVX1) -HVV1 -HVX #HXX1
lapply (cpcs_canc_sn … HXX1 … HXW) -X
elim (lift_total X1 0 (i+1))
lapply (drop_fwd_drop2 … HLK1)
/4 width=12 by cpcs_lift, lstas_cast, lstas_ldef, ex2_intro/
]
-| #a #I #G #L2 #V2 #T2 #U2 #l1 #_ #IHTU2 #l2 #Hl12 #Hl2 #L1 #HL12
- lapply (da_inv_bind … Hl2) -Hl2 #Hl2
- elim (IHTU2 … Hl2 (L1.ⓑ{I}V2) …)
+| #a #I #G #L2 #V2 #T2 #U2 #d1 #_ #IHTU2 #d2 #Hd12 #Hd2 #L1 #HL12
+ lapply (da_inv_bind … Hd2) -Hd2 #Hd2
+ elim (IHTU2 … Hd2 (L1.ⓑ{I}V2) …)
/3 width=3 by lsubsv_pair, lstas_bind, cpcs_bind_dx, ex2_intro/
-| #G #L2 #V2 #T2 #U2 #l1 #_ #IHTU2 #l2 #Hl12 #Hl2 #L1 #HL12
- lapply (da_inv_flat … Hl2) -Hl2 #Hl2
- elim (IHTU2 … Hl2 … HL12) -L2
+| #G #L2 #V2 #T2 #U2 #d1 #_ #IHTU2 #d2 #Hd12 #Hd2 #L1 #HL12
+ lapply (da_inv_flat … Hd2) -Hd2 #Hd2
+ elim (IHTU2 … Hd2 … HL12) -L2
/3 width=5 by lstas_appl, cpcs_flat, ex2_intro/
-| #G #L2 #W2 #T2 #U2 #l1 #_ #IHTU2 #l2 #Hl12 #Hl2 #L1 #HL12
- lapply (da_inv_flat … Hl2) -Hl2 #Hl2
- elim (IHTU2 … Hl2 … HL12) -L2
+| #G #L2 #W2 #T2 #U2 #d1 #_ #IHTU2 #d2 #Hd12 #Hd2 #L1 #HL12
+ lapply (da_inv_flat … Hd2) -Hd2 #Hd2
+ elim (IHTU2 … Hd2 … HL12) -L2
/3 width=3 by lstas_cast, ex2_intro/
]
qed-.
lemma lsubsv_sta_trans: ∀h,g,G,L2,T,U2. ⦃G, L2⦄ ⊢ T •*[h, 1] U2 →
- ∀l. ⦃G, L2⦄ ⊢ T ▪[h, g] l+1 →
+ ∀d. ⦃G, L2⦄ ⊢ T ▪[h, g] d+1 →
∀L1. G ⊢ L1 ⫃¡[h, g] L2 →
∃∃U1. ⦃G, L1⦄ ⊢ T •*[h, 1] U1 & ⦃G, L1⦄ ⊢ U1 ⬌* U2.
/2 width=7 by lsubsv_lstas_trans/ qed-.
lemma lsubsv_fwd_lsuba: ∀h,g,G,L1,L2. G ⊢ L1 ⫃¡[h, g] L2 → G ⊢ L1 ⫃⁝ L2.
#h #g #G #L1 #L2 #H elim H -L1 -L2 /2 width=1 by lsuba_pair/
-#L1 #L2 #W #V #l1 #H #_ #_ #_ #_ #IHL12
+#L1 #L2 #W #V #d1 #H #_ #_ #_ #_ #IHL12
elim (shnv_inv_cast … H) -H #HW #HV #H
-lapply (H 0 ?) // -l1 #HWV
+lapply (H 0 ?) // -d1 #HWV
elim (snv_fwd_aaa … HW) -HW #B #HW
elim (snv_fwd_aaa … HV) -HV #A #HV
lapply (scpes_aaa_mono … HWV … HW … HV) #H destruct
(* Properties on decomposed extended parallel computation on terms **********)
-lemma lsubsv_scpds_trans: ∀h,g,G,L2,T1,T2,l. ⦃G, L2⦄ ⊢ T1 •*➡*[h, g, l] T2 →
+lemma lsubsv_scpds_trans: ∀h,g,G,L2,T1,T2,d. ⦃G, L2⦄ ⊢ T1 •*➡*[h, g, d] T2 →
∀L1. G ⊢ L1 ⫃¡[h, g] L2 →
- ∃∃T. ⦃G, L1⦄ ⊢ T1 •*➡*[h, g, l] T & ⦃G, L1⦄ ⊢ T2 ➡* T.
-#h #g #G #L2 #T1 #T2 #l2 * #T #l1 #Hl21 #Hl1 #HT1 #HT2 #L1 #HL12
+ ∃∃T. ⦃G, L1⦄ ⊢ T1 •*➡*[h, g, d] T & ⦃G, L1⦄ ⊢ T2 ➡* T.
+#h #g #G #L2 #T1 #T2 #d2 * #T #d1 #Hd21 #Hd1 #HT1 #HT2 #L1 #HL12
lapply (lsubsv_cprs_trans … HL12 … HT2) -HT2 #HT2
-elim (lsubsv_lstas_trans … HT1 … Hl1 … HL12) // #T0 #HT10 #HT0
+elim (lsubsv_lstas_trans … HT1 … Hd1 … HL12) // #T0 #HT10 #HT0
lapply (cpcs_cprs_strap1 … HT0 … HT2) -T #HT02
elim (cpcs_inv_cprs … HT02) -HT02
/5 width=5 by lsubsv_fwd_lsubd, lsubd_da_trans, ex4_2_intro, ex2_intro/
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 #HVW #_ #_ #_ #_ #H1 #H2 destruct -IHV
+ | #W #d #HVW #_ #_ #_ #_ #H1 #H2 destruct -IHV
/3 width=6 by shnv_inv_snv, snv_lref/
]
| #a #I #G #L2 #V #T #_ #_ #IHV #IHT #L1 #HL12 destruct
/4 width=1 by snv_bind, lsubsv_pair/
-| #a #G #L2 #V #W0 #T #U0 #l #_ #_ #HVW0 #HTU0 #IHV #IHT #L1 #HL12
+| #a #G #L2 #V #W0 #T #U0 #d #_ #_ #HVW0 #HTU0 #IHV #IHT #L1 #HL12
elim (lsubsv_scpds_trans … HVW0 … HL12) -HVW0 #V0 #HV0 #HWV0
elim (lsubsv_scpds_trans … HTU0 … HL12) -HTU0 #X #HT0 #H
elim (cprs_inv_abst1 … H) -H #W #T0 #HW0 #_ #H destruct
(* STRATIFIED HIGHER NATIVE VALIDITY FOR TERMS ******************************)
-inductive shnv (h) (g) (l1) (G) (L): predicate term ≝
+inductive shnv (h) (g) (d1) (G) (L): predicate term ≝
| shnv_cast: ∀U,T. ⦃G, L⦄ ⊢ U ¡[h, g] → ⦃G, L⦄ ⊢ T ¡[h, g] →
- (∀l2. l2 ≤ l1 → ⦃G, L⦄ ⊢ U •*⬌*[h, g, l2, l2+1] T) →
- shnv h g l1 G L (ⓝU.T)
+ (∀d2. d2 ≤ d1 → ⦃G, L⦄ ⊢ U •*⬌*[h, g, d2, d2+1] T) →
+ shnv h g d1 G L (ⓝU.T)
.
interpretation "stratified higher native validity (term)"
- 'NativeValid h g l G L T = (shnv h g l G L T).
+ 'NativeValid h g d G L T = (shnv h g d G L T).
(* Basic inversion lemmas ***************************************************)
-fact shnv_inv_cast_aux: ∀h,g,G,L,X,l1. ⦃G, L⦄ ⊢ X ¡[h, g, l1] → ∀U,T. X = ⓝU.T →
+fact shnv_inv_cast_aux: ∀h,g,G,L,X,d1. ⦃G, L⦄ ⊢ X ¡[h, g, d1] → ∀U,T. X = ⓝU.T →
∧∧ ⦃G, L⦄ ⊢ U ¡[h, g] & ⦃G, L⦄ ⊢ T ¡[h, g]
- & (∀l2. l2 ≤ l1 → ⦃G, L⦄ ⊢ U •*⬌*[h, g, l2, l2+1] T).
-#h #g #G #L #X #l1 * -X
+ & (∀d2. d2 ≤ d1 → ⦃G, L⦄ ⊢ U •*⬌*[h, g, d2, d2+1] T).
+#h #g #G #L #X #d1 * -X
#U #T #HU #HT #HUT #U1 #T1 #H destruct /3 width=1 by and3_intro/
qed-.
-lemma shnv_inv_cast: ∀h,g,G,L,U,T,l1. ⦃G, L⦄ ⊢ ⓝU.T ¡[h, g, l1] →
+lemma shnv_inv_cast: ∀h,g,G,L,U,T,d1. ⦃G, L⦄ ⊢ ⓝU.T ¡[h, g, d1] →
∧∧ ⦃G, L⦄ ⊢ U ¡[h, g] & ⦃G, L⦄ ⊢ T ¡[h, g]
- & (∀l2. l2 ≤ l1 → ⦃G, L⦄ ⊢ U •*⬌*[h, g, l2, l2+1] T).
+ & (∀d2. d2 ≤ d1 → ⦃G, L⦄ ⊢ U •*⬌*[h, g, d2, d2+1] T).
/2 width=3 by shnv_inv_cast_aux/ qed-.
-lemma shnv_inv_snv: ∀h,g,G,L,T,l. ⦃G, L⦄ ⊢ T ¡[h, g, l] → ⦃G, L⦄ ⊢ T ¡[h, g].
-#h #g #G #L #T #l * -T
+lemma shnv_inv_snv: ∀h,g,G,L,T,d. ⦃G, L⦄ ⊢ T ¡[h, g, d] → ⦃G, L⦄ ⊢ T ¡[h, g].
+#h #g #G #L #T #d * -T
#U #T #HU #HT #HUT elim (HUT 0) -HUT /2 width=3 by snv_cast/
qed-.
lemma snv_shnv_cast: ∀h,g,G,L,U,T. ⦃G, L⦄ ⊢ ⓝU.T ¡[h, g] → ⦃G, L⦄ ⊢ ⓝU.T ¡[h, g, 0].
#h #g #G #L #U #T #H elim (snv_inv_cast … H) -H
#U0 #HU #HT #HU0 #HTU0 @shnv_cast // -HU -HT
-#l #H <(le_n_O_to_eq … H) -l /2 width=3 by scpds_div/
+#d #H <(le_n_O_to_eq … H) -d /2 width=3 by scpds_div/
qed.
| snv_sort: ∀G,L,k. snv h g G L (⋆k)
| snv_lref: ∀I,G,L,K,V,i. ⬇[i] L ≡ K.ⓑ{I}V → snv h g G K V → snv h g G L (#i)
| snv_bind: ∀a,I,G,L,V,T. snv h g G L V → snv h g G (L.ⓑ{I}V) T → snv h g G L (ⓑ{a,I}V.T)
-| snv_appl: ∀a,G,L,V,W0,T,U0,l. snv h g G L V → snv h g G L T →
- ⦃G, L⦄ ⊢ V •*➡*[h, g, 1] W0 → ⦃G, L⦄ ⊢ T •*➡*[h, g, l] ⓛ{a}W0.U0 → snv h g G L (ⓐV.T)
+| snv_appl: ∀a,G,L,V,W0,T,U0,d. snv h g G L V → snv h g G L T →
+ ⦃G, L⦄ ⊢ V •*➡*[h, g, 1] W0 → ⦃G, L⦄ ⊢ T •*➡*[h, g, d] ⓛ{a}W0.U0 → snv h g G L (ⓐV.T)
| snv_cast: ∀G,L,U,T,U0. snv h g G L U → snv h g G L T →
⦃G, L⦄ ⊢ U •*➡*[h, g, 0] U0 → ⦃G, L⦄ ⊢ T •*➡*[h, g, 1] U0 → snv h g G L (ⓝU.T)
.
[ #G #L #k #i #H destruct
| #I #G #L #K #V #i0 #HLK #HV #i #H destruct /2 width=5 by ex2_3_intro/
| #a #I #G #L #V #T #_ #_ #i #H destruct
-| #a #G #L #V #W0 #T #U0 #l #_ #_ #_ #_ #i #H destruct
+| #a #G #L #V #W0 #T #U0 #d #_ #_ #_ #_ #i #H destruct
| #G #L #U #T #U0 #_ #_ #_ #_ #i #H destruct
]
qed-.
[ #G #L #k #p #H destruct
| #I #G #L #K #V #i #_ #_ #p #H destruct
| #a #I #G #L #V #T #_ #_ #p #H destruct
-| #a #G #L #V #W0 #T #U0 #l #_ #_ #_ #_ #p #H destruct
+| #a #G #L #V #W0 #T #U0 #d #_ #_ #_ #_ #p #H destruct
| #G #L #U #T #U0 #_ #_ #_ #_ #p #H destruct
]
qed-.
[ #G #L #k #b #Z #X1 #X2 #H destruct
| #I #G #L #K #V #i #_ #_ #b #Z #X1 #X2 #H destruct
| #a #I #G #L #V #T #HV #HT #b #Z #X1 #X2 #H destruct /2 width=1 by conj/
-| #a #G #L #V #W0 #T #U0 #l #_ #_ #_ #_ #b #Z #X1 #X2 #H destruct
+| #a #G #L #V #W0 #T #U0 #d #_ #_ #_ #_ #b #Z #X1 #X2 #H destruct
| #G #L #U #T #U0 #_ #_ #_ #_ #b #Z #X1 #X2 #H destruct
]
qed-.
/2 width=4 by snv_inv_bind_aux/ qed-.
fact snv_inv_appl_aux: ∀h,g,G,L,X. ⦃G, L⦄ ⊢ X ¡[h, g] → ∀V,T. X = ⓐV.T →
- ∃∃a,W0,U0,l. ⦃G, L⦄ ⊢ V ¡[h, g] & ⦃G, L⦄ ⊢ T ¡[h, g] &
- ⦃G, L⦄ ⊢ V •*➡*[h, g, 1] W0 & ⦃G, L⦄ ⊢ T •*➡*[h, g, l] ⓛ{a}W0.U0.
+ ∃∃a,W0,U0,d. ⦃G, L⦄ ⊢ V ¡[h, g] & ⦃G, L⦄ ⊢ T ¡[h, g] &
+ ⦃G, L⦄ ⊢ V •*➡*[h, g, 1] W0 & ⦃G, L⦄ ⊢ T •*➡*[h, g, d] ⓛ{a}W0.U0.
#h #g #G #L #X * -L -X
[ #G #L #k #X1 #X2 #H destruct
| #I #G #L #K #V #i #_ #_ #X1 #X2 #H destruct
| #a #I #G #L #V #T #_ #_ #X1 #X2 #H destruct
-| #a #G #L #V #W0 #T #U0 #l #HV #HT #HVW0 #HTU0 #X1 #X2 #H destruct /2 width=6 by ex4_4_intro/
+| #a #G #L #V #W0 #T #U0 #d #HV #HT #HVW0 #HTU0 #X1 #X2 #H destruct /2 width=6 by ex4_4_intro/
| #G #L #U #T #U0 #_ #_ #_ #_ #X1 #X2 #H destruct
]
qed-.
lemma snv_inv_appl: ∀h,g,G,L,V,T. ⦃G, L⦄ ⊢ ⓐV.T ¡[h, g] →
- ∃∃a,W0,U0,l. ⦃G, L⦄ ⊢ V ¡[h, g] & ⦃G, L⦄ ⊢ T ¡[h, g] &
- ⦃G, L⦄ ⊢ V •*➡*[h, g, 1] W0 & ⦃G, L⦄ ⊢ T •*➡*[h, g, l] ⓛ{a}W0.U0.
+ ∃∃a,W0,U0,d. ⦃G, L⦄ ⊢ V ¡[h, g] & ⦃G, L⦄ ⊢ T ¡[h, g] &
+ ⦃G, L⦄ ⊢ V •*➡*[h, g, 1] W0 & ⦃G, L⦄ ⊢ T •*➡*[h, g, d] ⓛ{a}W0.U0.
/2 width=3 by snv_inv_appl_aux/ qed-.
fact snv_inv_cast_aux: ∀h,g,G,L,X. ⦃G, L⦄ ⊢ X ¡[h, g] → ∀U,T. X = ⓝU.T →
[ #G #L #k #X1 #X2 #H destruct
| #I #G #L #K #V #i #_ #_ #X1 #X2 #H destruct
| #a #I #G #L #V #T #_ #_ #X1 #X2 #H destruct
-| #a #G #L #V #W0 #T #U0 #l #_ #_ #_ #_ #X1 #X2 #H destruct
+| #a #G #L #V #W0 #T #U0 #d #_ #_ #_ #_ #X1 #X2 #H destruct
| #G #L #U #T #U0 #HV #HT #HU0 #HTU0 #X1 #X2 #H destruct /2 width=3 by ex4_intro/
]
qed-.
[ /2 width=2 by aaa_sort, ex_intro/
| #I #G #L #K #V #i #HLK #_ * /3 width=6 by aaa_lref, ex_intro/
| #a * #G #L #V #T #_ #_ * #B #HV * #A #HA /3 width=2 by aaa_abbr, aaa_abst, ex_intro/
-| #a #G #L #V #W0 #T #U0 #l #_ #_ #HVW0 #HTU0 * #B #HV * #X #HT
+| #a #G #L #V #W0 #T #U0 #d #_ #_ #HVW0 #HTU0 * #B #HV * #X #HT
lapply (scpds_aaa_conf … HV … HVW0) -HVW0 #HW0
lapply (scpds_aaa_conf … HT … HTU0) -HTU0 #H
elim (aaa_inv_abst … H) -H #B0 #A #H1 #HU #H2 destruct
(* Advanced forward lemmas **************************************************)
-lemma snv_fwd_da: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ T ¡[h, g] → ∃l. ⦃G, L⦄ ⊢ T ▪[h, g] l.
+lemma snv_fwd_da: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ T ¡[h, g] → ∃d. ⦃G, L⦄ ⊢ T ▪[h, g] d.
#h #g #G #L #T #H elim (snv_fwd_aaa … H) -H /2 width=2 by aaa_da/
qed-.
lemma snv_fwd_lstas: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ T ¡[h, g] →
- ∀l. ∃U. ⦃G, L⦄ ⊢ T •*[h, l] U.
-#h #g #G #L #T #H #l elim (snv_fwd_aaa … H) -H
-#A #HT elim (aaa_lstas h … HT l) -HT /2 width=2 by ex_intro/
+ ∀d. ∃U. ⦃G, L⦄ ⊢ T •*[h, d] U.
+#h #g #G #L #T #H #d elim (snv_fwd_aaa … H) -H
+#A #HT elim (aaa_lstas h … HT d) -HT /2 width=2 by ex_intro/
qed-.
(∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G1, L1, T1⦄ → IH_da_cpr_lpr h g G1 L1 T1) →
∀G1,L1,T1. G0 = G1 → L0 = L1 → T0 = T1 → IH_da_cpr_lpr h g G1 L1 T1.
#h #g #G0 #L0 #T0 #IH3 #IH2 #IH1 #G1 #L1 * * [|||| * ]
-[ #k #_ #_ #_ #_ #l #H2 #X3 #H3 #L2 #_ -IH3 -IH2 -IH1
+[ #k #_ #_ #_ #_ #d #H2 #X3 #H3 #L2 #_ -IH3 -IH2 -IH1
lapply (da_inv_sort … H2) -H2
lapply (cpr_inv_sort1 … H3) -H3 #H destruct /2 width=1 by da_sort/
-| #i #HG0 #HL0 #HT0 #H1 #l #H2 #X3 #H3 #L2 #HL12 destruct -IH3 -IH2
+| #i #HG0 #HL0 #HT0 #H1 #d #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
+ elim (da_inv_lref … H2) -H2 * #K1 [ #V1 | #W1 #d1 ] #HLK1 [ #HV1 | #HW1 #H ] destruct
lapply (drop_mono … H … HLK1) -H #H destruct
elim (cpr_inv_lref1 … H3) -H3
[1,3: #H destruct
]
| #p #_ #_ #HT0 #H1 destruct -IH3 -IH2 -IH1
elim (snv_inv_gref … H1)
-| #a #I #V1 #T1 #HG0 #HL0 #HT0 #H1 #l #H2 #X3 #H3 #L2 #HL12 destruct -IH2
+| #a #I #V1 #T1 #HG0 #HL0 #HT0 #H1 #d #H2 #X3 #H3 #L2 #HL12 destruct -IH2
elim (snv_inv_bind … H1) -H1 #_ #HT1
lapply (da_inv_bind … H2) -H2
elim (cpr_inv_bind1 … H3) -H3 *
| #T2 #HT12 #HT2 #H1 #H2 destruct
/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 #b1 #W1 #U1 #l1 #HV1 #HT1 #HVW1 #HTU1
- lapply (da_inv_flat … H2) -H2 #Hl
+| #V1 #T1 #HG0 #HL0 #HT0 #H1 #d #H2 #X3 #H3 #L2 #HL12 destruct
+ elim (snv_inv_appl … H1) -H1 #b1 #W1 #U1 #d1 #HV1 #HT1 #HVW1 #HTU1
+ lapply (da_inv_flat … H2) -H2 #Hd
elim (cpr_inv_appl1 … H3) -H3 *
[ #V2 #T2 #HV12 #HT12 #H destruct -IH3 -IH2 /4 width=7 by da_flat, fqup_fpbg/
| #b #V2 #W #W2 #U #U2 #HV12 #HW2 #HU2 #H1 #H2 destruct
elim (snv_inv_bind … HT1) -HT1 #HW #HU
- lapply (da_inv_bind … Hl) -Hl #Hl
- elim (scpds_inv_abst1 … HTU1) -HTU1 #W3 #U3 #HW3 #_ #H destruct -U3 -l1
- elim (snv_fwd_da … HV1) #l1 #Hl1
- elim (snv_fwd_da … HW) #l0 #Hl0
- lapply (cprs_scpds_div … HW3 … Hl0 … 1 HVW1) -W3 #H
- elim (da_scpes_aux … IH3 IH2 IH1 … Hl0 … Hl1 … H) -IH3 -IH2 -H /2 width=1 by fqup_fpbg/ #_ #H1
- <minus_n_O #H destruct >(plus_minus_m_m l1 1) in Hl1; // -H1 #Hl1
- lapply (IH1 … HV1 … Hl1 … HV12 … HL12) -HV1 -Hl1 -HV12 [ /2 by fqup_fpbg/ ]
- lapply (IH1 … Hl0 … HW2 … HL12) -Hl0 /2 width=1 by fqup_fpbg/ -HW
- lapply (IH1 … HU … Hl … HU2 (L2.ⓛW2) ?) -IH1 -HU -Hl -HU2 [1,2: /2 by fqup_fpbg, lpr_pair/ ] -HL12 -HW2
+ lapply (da_inv_bind … Hd) -Hd #Hd
+ elim (scpds_inv_abst1 … HTU1) -HTU1 #W3 #U3 #HW3 #_ #H destruct -U3 -d1
+ elim (snv_fwd_da … HV1) #d1 #Hd1
+ elim (snv_fwd_da … HW) #d0 #Hd0
+ lapply (cprs_scpds_div … HW3 … Hd0 … 1 HVW1) -W3 #H
+ elim (da_scpes_aux … IH3 IH2 IH1 … Hd0 … Hd1 … H) -IH3 -IH2 -H /2 width=1 by fqup_fpbg/ #_ #H1
+ <minus_n_O #H destruct >(plus_minus_m_m d1 1) in Hd1; // -H1 #Hd1
+ lapply (IH1 … HV1 … Hd1 … HV12 … HL12) -HV1 -Hd1 -HV12 [ /2 by fqup_fpbg/ ]
+ lapply (IH1 … Hd0 … HW2 … HL12) -Hd0 /2 width=1 by fqup_fpbg/ -HW
+ lapply (IH1 … HU … Hd … HU2 (L2.ⓛW2) ?) -IH1 -HU -Hd -HU2 [1,2: /2 by fqup_fpbg, lpr_pair/ ] -HL12 -HW2
/4 width=6 by da_bind, lsubd_da_trans, lsubd_beta/
- | #b #V0 #V2 #W #W2 #U #U2 #HV10 #HV02 #HW2 #HU2 #H1 #H2 destruct -IH3 -IH2 -b1 -V0 -W1 -U1 -l1 -HV1
+ | #b #V0 #V2 #W #W2 #U #U2 #HV10 #HV02 #HW2 #HU2 #H1 #H2 destruct -IH3 -IH2 -b1 -V0 -W1 -U1 -d1 -HV1
elim (snv_inv_bind … HT1) -HT1 #_
- lapply (da_inv_bind … Hl) -Hl
+ lapply (da_inv_bind … Hd) -Hd
/5 width=9 by da_bind, da_flat, fqup_fpbg, lpr_pair/
]
-| #W1 #T1 #HG0 #HL0 #HT0 #H1 #l #H2 #X3 #H3 #L2 #HL12 destruct -IH3 -IH2
+| #W1 #T1 #HG0 #HL0 #HT0 #H1 #d #H2 #X3 #H3 #L2 #HL12 destruct -IH3 -IH2
elim (snv_inv_cast … H1) -H1 #U1 #HW1 #HT1 #HWU1 #HTU1
- lapply (da_inv_flat … H2) -H2 #Hl
+ lapply (da_inv_flat … H2) -H2 #Hd
elim (cpr_inv_cast1 … H3) -H3
[ * #W2 #T2 #HW12 #HT12 #H destruct /4 width=7 by da_flat, fqup_fpbg/
| /3 width=7 by fqup_fpbg/
(* Relocation properties ****************************************************)
-lemma snv_lift: ∀h,g,G,K,T. ⦃G, K⦄ ⊢ T ¡[h, g] → ∀L,s,d,e. ⬇[s, d, e] L ≡ K →
- ∀U. ⬆[d, e] T ≡ U → ⦃G, L⦄ ⊢ U ¡[h, g].
+lemma snv_lift: ∀h,g,G,K,T. ⦃G, K⦄ ⊢ T ¡[h, g] → ∀L,s,l,m. ⬇[s, l, m] L ≡ K →
+ ∀U. ⬆[l, m] T ≡ U → ⦃G, L⦄ ⊢ U ¡[h, g].
#h #g #G #K #T #H elim H -G -K -T
-[ #G #K #k #L #s #d #e #_ #X #H
- >(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
+[ #G #K #k #L #s #l #m #_ #X #H
+ >(lift_inv_sort1 … H) -X -K -l -m //
+| #I #G #K #K0 #V #i #HK0 #_ #IHV #L #s #l #m #HLK #X #H
+ elim (lift_inv_lref1 … H) * #Hil #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
+ elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hil #L0 #W #HLK0 #HVW #H destruct
/3 width=9 by snv_lref/
| 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
+| #a #I #G #K #V #T #_ #_ #IHV #IHT #L #s #l #m #HLK #X #H
elim (lift_inv_bind1 … H) -H #W #U #HVW #HTU #H destruct
/4 width=5 by snv_bind, drop_skip/
-| #a #G #K #V #W0 #T #U0 #l #_ #_ #HVW0 #HTU0 #IHV #IHT #L #s #d #e #HLK #X #H
+| #a #G #K #V #W0 #T #U0 #d #_ #_ #HVW0 #HTU0 #IHV #IHT #L #s #l #m #HLK #X #H
elim (lift_inv_flat1 … H) -H #W #U #HVW #HTU #H destruct
- elim (lift_total W0 d e)
- elim (lift_total U0 (d+1) e)
+ elim (lift_total W0 l m)
+ elim (lift_total U0 (l+1) m)
/4 width=17 by snv_appl, scpds_lift, lift_bind/
-| #G #K #V #T #U0 #_ #_ #HVU0 #HTU0 #IHV #IHT #L #s #d #e #HLK #X #H
+| #G #K #V #T #U0 #_ #_ #HVU0 #HTU0 #IHV #IHT #L #s #l #m #HLK #X #H
elim (lift_inv_flat1 … H) -H #W #U #HVW #HTU #H destruct
- elim (lift_total U0 d e)
+ elim (lift_total U0 l m)
/3 width=12 by snv_cast, scpds_lift/
]
qed.
-lemma snv_inv_lift: ∀h,g,G,L,U. ⦃G, L⦄ ⊢ U ¡[h, g] → ∀K,s,d,e. ⬇[s, d, e] L ≡ K →
- ∀T. ⬆[d, e] T ≡ U → ⦃G, K⦄ ⊢ T ¡[h, g].
+lemma snv_inv_lift: ∀h,g,G,L,U. ⦃G, L⦄ ⊢ U ¡[h, g] → ∀K,s,l,m. ⬇[s, l, m] L ≡ K →
+ ∀T. ⬆[l, m] T ≡ U → ⦃G, K⦄ ⊢ T ¡[h, g].
#h #g #G #L #U #H elim H -G -L -U
-[ #G #L #k #K #s #d #e #_ #X #H
- >(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
+[ #G #L #k #K #s #l #m #_ #X #H
+ >(lift_inv_sort2 … H) -X -L -l -m //
+| #I #G #L #L0 #W #i #HL0 #_ #IHW #K #s #l #m #HLK #X #H
+ elim (lift_inv_lref2 … H) * #Hil #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
+ elim (drop_inv_skip1 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hil #K0 #V #HLK0 #HVW #H destruct
/3 width=12 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
+| #a #I #G #L #W #U #_ #_ #IHW #IHU #K #s #l #m #HLK #X #H
elim (lift_inv_bind2 … H) -H #V #T #HVW #HTU #H destruct
/4 width=5 by snv_bind, drop_skip/
-| #a #G #L #W #W1 #U #U1 #l #_ #_ #HW1 #HU1 #IHW #IHU #K #s #d #e #HLK #X #H
+| #a #G #L #W #W1 #U #U1 #d #_ #_ #HW1 #HU1 #IHW #IHU #K #s #l #m #HLK #X #H
elim (lift_inv_flat2 … H) -H #V #T #HVW #HTU #H destruct
elim (scpds_inv_lift1 … HW1 … HLK … HVW) -HW1 #W0 #HW01 #HVW0
elim (scpds_inv_lift1 … HU1 … HLK … HTU) -HU1 #X #H #HTU0
elim (lift_inv_bind2 … H) -H #Y #U0 #HY #HU01 #H destruct
lapply (lift_inj … HY … HW01) -HY #H destruct
/3 width=6 by snv_appl/
-| #G #L #W #U #U1 #_ #_ #HWU1 #HU1 #IHW #IHU #K #s #d #e #HLK #X #H
+| #G #L #W #U #U1 #_ #_ #HWU1 #HU1 #IHW #IHU #K #s #l #m #HLK #X #H
elim (lift_inv_flat2 … H) -H #V #T #HVW #HTU #H destruct
elim (scpds_inv_lift1 … HWU1 … HLK … HVW) -HWU1 #U0 #HU01 #HVU0
elim (scpds_inv_lift1 … HU1 … HLK … HTU) -HU1 #X #HX #HTU0
/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 #W1 #U1 #l0 #HV1 #HT1 #HVW1 #HTU1
+ elim (snv_inv_appl … H1) -H1 #a #W1 #U1 #d0 #HV1 #HT1 #HVW1 #HTU1
elim (cpr_inv_appl1 … H2) -H2 *
[ #V2 #T2 #HV12 #HT12 #H destruct -IH4
lapply (IH1 … HV12 … HL12) /2 width=1 by fqup_fpbg/ #HV2
/2 width=7 by snv_appl, cprs_bind/
| #b #V2 #W10 #W20 #T10 #T20 #HV12 #HW120 #HT120 #H1 #H2 destruct
elim (snv_inv_bind … HT1) -HT1 #HW10 #HT10
- elim (scpds_inv_abst1 … HTU1) -HTU1 #W30 #T30 #HW130 #_ #H destruct -T30 -l0
- elim (snv_fwd_da … HV1) #l #HV1l
- elim (snv_fwd_da … HW10) #l0 #HW10l
- lapply (cprs_scpds_div … HW130 … HW10l … 1 HVW1) -W30 #HVW10
- elim (da_scpes_aux … IH4 IH1 IH2 … HW10l … HV1l … HVW10) /2 width=1 by fqup_fpbg/
- #_ #Hl <minus_n_O #H destruct >(plus_minus_m_m l 1) in HV1l; // -Hl #HV1l
+ elim (scpds_inv_abst1 … HTU1) -HTU1 #W30 #T30 #HW130 #_ #H destruct -T30 -d0
+ elim (snv_fwd_da … HV1) #d #HV1d
+ elim (snv_fwd_da … HW10) #d0 #HW10d
+ lapply (cprs_scpds_div … HW130 … HW10d … 1 HVW1) -W30 #HVW10
+ elim (da_scpes_aux … IH4 IH1 IH2 … HW10d … HV1d … HVW10) /2 width=1 by fqup_fpbg/
+ #_ #Hd <minus_n_O #H destruct >(plus_minus_m_m d 1) in HV1d; // -Hd #HV1d
lapply (scpes_cpr_lpr_aux … IH2 IH3 … HVW10 … HW120 … HV12 … HL12) /2 width=1 by fqup_fpbg/ -HVW10 #HVW20
- lapply (IH2 … HV1l … HV12 … HL12) /2 width=1 by fqup_fpbg/ -HV1l #HV2l
- lapply (IH2 … HW10l … HW120 … HL12) /2 width=1 by fqup_fpbg/ -HW10l #HW20l
+ lapply (IH2 … HV1d … HV12 … HL12) /2 width=1 by fqup_fpbg/ -HV1d #HV2d
+ lapply (IH2 … HW10d … HW120 … HL12) /2 width=1 by fqup_fpbg/ -HW10d #HW20d
lapply (IH1 … HV12 … HL12) /2 width=1 by fqup_fpbg/ #HV2
lapply (IH1 … HW120 … HL12) /2 width=1 by fqup_fpbg/ -HW10 #HW20
lapply (IH1 … HT10 … HT120 … (L2.ⓛW20) ?) /2 width=1 by fqup_fpbg, lpr_pair/ -HT10 #HT20
@snv_bind /2 width=1 by snv_cast_scpes/
@(lsubsv_snv_trans … HT20) -HT20
- @(lsubsv_beta … (l-1)) //
- @shnv_cast [1,2: // ] #l0 #Hl0
- lapply (scpes_le_aux … IH4 IH1 IH2 IH3 … HW20l … HV2l … l0 … HVW20) -IH4 -IH3 -IH2 -IH1 -HW20l -HV2l -HVW20
+ @(lsubsv_beta … (d-1)) //
+ @shnv_cast [1,2: // ] #d0 #Hd0
+ lapply (scpes_le_aux … IH4 IH1 IH2 IH3 … HW20d … HV2d … d0 … HVW20) -IH4 -IH3 -IH2 -IH1 -HW20d -HV2d -HVW20
/3 width=5 by fpbg_fpbs_trans, fqup_fpbg, cpr_lpr_fpbs, le_S_S/
| #b #V0 #V2 #W0 #W2 #T0 #T2 #HV10 #HV02 #HW02 #HT02 #H1 #H2 destruct -IH4
elim (snv_inv_bind … HT1) -HT1 #HW0 #HT0
elim (snv_inv_cast … H1) -H1 #U1 #HW1 #HT1 #HWU1 #HTU1
elim (cpr_inv_cast1 … H2) -H2
[ * #W2 #T2 #HW12 #HT12 #H destruct
- elim (snv_fwd_da … HW1) #l #HW1l
- lapply (scpds_div … HWU1 … HTU1) -U1 -l #HWT1
+ elim (snv_fwd_da … HW1) #d #HW1d
+ lapply (scpds_div … HWU1 … HTU1) -U1 -d #HWT1
lapply (scpes_cpr_lpr_aux … IH2 IH3 … HWT1 … HW12 … HT12 … HL12) /2 width=1 by fqup_fpbg/
lapply (IH1 … HW12 … HL12) /2 width=1 by fqup_fpbg/
lapply (IH1 … HT12 … HL12) /2 width=1 by fqup_fpbg/ -L1
(∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G1, L1, T1⦄ → IH_snv_lstas h g G1 L1 T1) →
∀G1,L1,T1. G0 = G1 → L0 = L1 → T0 = T1 → IH_snv_lstas h g G1 L1 T1.
#h #g #G0 #L0 #T0 #IH4 #IH3 #IH2 #IH1 #G1 #L1 * * [|||| * ]
-[ #k #HG0 #HL0 #HT0 #_ #l1 #l2 #Hl21 #Hl1 #X #H2 destruct -IH4 -IH3 -IH2 -IH1
+[ #k #HG0 #HL0 #HT0 #_ #d1 #d2 #Hd21 #Hd1 #X #H2 destruct -IH4 -IH3 -IH2 -IH1
>(lstas_inv_sort1 … H2) -X //
-| #i #HG0 #HL0 #HT0 #H1 #l1 #l2 #Hl21 #Hl1 #T #H2 destruct -IH4 -IH3 -IH2
+| #i #HG0 #HL0 #HT0 #H1 #d1 #d2 #Hd21 #Hd1 #T #H2 destruct -IH4 -IH3 -IH2
elim (snv_inv_lref … H1) -H1 #I0 #K0 #X0 #HLK0 #HX0
- elim (da_inv_lref … Hl1) -Hl1 * #K1 [ #V1 | #W1 #l0 ] #HLK1 [ #Hl1 | #Hl0 #H ]
+ elim (da_inv_lref … Hd1) -Hd1 * #K1 [ #V1 | #W1 #d0 ] #HLK1 [ #Hd1 | #Hd0 #H ]
lapply (drop_mono … HLK0 … HLK1) -HLK0 #H0 destruct
- elim (lstas_inv_lref1 … H2) -H2 * #K #Y #X [3,6: #l ] #HLK #HYX [1,2: #HXT #H0 |3,5: #HXT |4,6: #H1 #H2 ]
+ elim (lstas_inv_lref1 … H2) -H2 * #K #Y #X [3,6: #d ] #HLK #HYX [1,2: #HXT #H0 |3,5: #HXT |4,6: #H1 #H2 ]
lapply (drop_mono … HLK … HLK1) -HLK #H destruct
- [ lapply (le_plus_to_le_r … Hl21) -Hl21 #Hl21 |3: -Hl21 ]
+ [ lapply (le_plus_to_le_r … Hd21) -Hd21 #Hd21 |3: -Hd21 ]
lapply (fqup_lref … G1 … HLK1) #H
lapply (drop_fwd_drop2 … HLK1) /4 width=8 by snv_lift, snv_lref, fqup_fpbg/
-| #p #HG0 #HL0 #HT0 #H1 #l1 #l2 #Hl21 #Hl1 #X #H2 destruct -IH4 -IH3 -IH2 -IH1
+| #p #HG0 #HL0 #HT0 #H1 #d1 #d2 #Hd21 #Hd1 #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
+| #a #I #V1 #T1 #HG0 #HL0 #HT0 #H1 #d1 #d2 #Hd21 #Hd1 #X #H2 destruct -IH4 -IH3 -IH2
elim (snv_inv_bind … H1) -H1 #HV1 #HT1
- lapply (da_inv_bind … Hl1) -Hl1 #Hl1
+ lapply (da_inv_bind … Hd1) -Hd1 #Hd1
elim (lstas_inv_bind1 … H2) -H2 #U1 #HTU1 #H destruct /4 width=8 by fqup_fpbg, snv_bind/
-| #V1 #T1 #HG0 #HL0 #HT0 #H1 #l1 #l2 #Hl21 #Hl1 #X #H2 destruct
- elim (snv_inv_appl … H1) -H1 #a #W1 #U1 #l0 #HV1 #HT1 #HVW1 #HTU1
- lapply (da_inv_flat … Hl1) -Hl1 #Hl1
+| #V1 #T1 #HG0 #HL0 #HT0 #H1 #d1 #d2 #Hd21 #Hd1 #X #H2 destruct
+ elim (snv_inv_appl … H1) -H1 #a #W1 #U1 #d0 #HV1 #HT1 #HVW1 #HTU1
+ lapply (da_inv_flat … Hd1) -Hd1 #Hd1
elim (lstas_inv_appl1 … H2) -H2 #T0 #HT10 #H destruct
- lapply (IH1 … HT1 … Hl1 … HT10) /2 width=1 by fqup_fpbg/ #HT0
- lapply (lstas_scpds_aux … IH1 IH4 IH3 IH2 … Hl1 … HT10 … HTU1) -IH4 -IH3 -IH2 -IH1 /2 width=1 by fqup_fpbg/ -T1 -l1 #H
+ lapply (IH1 … HT1 … Hd1 … HT10) /2 width=1 by fqup_fpbg/ #HT0
+ lapply (lstas_scpds_aux … IH1 IH4 IH3 IH2 … Hd1 … HT10 … HTU1) -IH4 -IH3 -IH2 -IH1 /2 width=1 by fqup_fpbg/ -T1 -d1 #H
elim (scpes_inv_abst2 … H) -H /3 width=6 by snv_appl, scpds_cprs_trans/
-| #U1 #T1 #HG0 #HL0 #HT0 #H1 #l1 #l2 #Hl21 #Hl1 #X #H2 destruct -IH4 -IH3 -IH2
+| #U1 #T1 #HG0 #HL0 #HT0 #H1 #d1 #d2 #Hd21 #Hd1 #X #H2 destruct -IH4 -IH3 -IH2
elim (snv_inv_cast … H1) -H1
- lapply (da_inv_flat … Hl1) -Hl1
+ lapply (da_inv_flat … Hd1) -Hd1
lapply (lstas_inv_cast1 … H2) -H2 /3 width=8 by fqup_fpbg/
]
qed-.
(∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G1, L1, T1⦄ → IH_lstas_cpr_lpr h g G1 L1 T1) →
∀G1,L1,T1. G0 = G1 → L0 = L1 → T0 = T1 → IH_lstas_cpr_lpr h g G1 L1 T1.
#h #g #G0 #L0 #T0 #IH4 #IH3 #IH2 #IH1 #G1 #L1 * * [|||| * ]
-[ #k #_ #_ #_ #_ #l1 #l2 #_ #_ #X2 #H2 #X3 #H3 #L2 #_ -IH4 -IH3 -IH2 -IH1
+[ #k #_ #_ #_ #_ #d1 #d2 #_ #_ #X2 #H2 #X3 #H3 #L2 #_ -IH4 -IH3 -IH2 -IH1
>(lstas_inv_sort1 … H2) -X2
>(cpr_inv_sort1 … H3) -X3 /2 width=3 by ex2_intro/
-| #i #HG0 #HL0 #HT0 #H1 #l1 #l2 #Hl21 #Hl1 #X2 #H2 #X3 #H3 #L2 #HL12 destruct -IH4 -IH3
+| #i #HG0 #HL0 #HT0 #H1 #d1 #d2 #Hd21 #Hd1 #X2 #H2 #X3 #H3 #L2 #HL12 destruct -IH4 -IH3
elim (snv_inv_lref … H1) -H1 #I0 #K0 #X0 #HK0 #HX0
- elim (da_inv_lref … Hl1) -Hl1 * #K1 [ #V1 | #W1 #l ] #HLK1 [ #HVl1 | #HWl1 #H destruct ]
+ elim (da_inv_lref … Hd1) -Hd1 * #K1 [ #V1 | #W1 #d ] #HLK1 [ #HVd1 | #HWd1 #H destruct ]
lapply (drop_mono … HK0 … HLK1) -HK0 #H destruct
- elim (lstas_inv_lref1 … H2) -H2 * #K0 #V0 #X0 [3,6: #l0 ] #HK0 #HVX0 [1,2: #HX02 #H |3,5: #HX02 |4,6: #H1 #H2 ] destruct
+ elim (lstas_inv_lref1 … H2) -H2 * #K0 #V0 #X0 [3,6: #d0 ] #HK0 #HVX0 [1,2: #HX02 #H |3,5: #HX02 |4,6: #H1 #H2 ] destruct
lapply (drop_mono … HK0 … HLK1) -HK0 #H destruct
- [ lapply (le_plus_to_le_r … Hl21) -Hl21 #Hl21 |3: -Hl21 ]
+ [ lapply (le_plus_to_le_r … Hd21) -Hd21 #Hd21 |3: -Hd21 ]
lapply (fqup_lref … G1 … HLK1) #HKV1
elim (lpr_drop_conf … HLK1 … HL12) -HL12 #X #H #HLK2
elim (lpr_inv_pair1 … H) -H #K2 [ #W2 | #W2 | #V2 ] #HK12 [ #HW12 | #HW12 | #HV12 ] #H destruct
|2,4,6: * #K #V #X #H #HVX #HX3
lapply (drop_mono … H … HLK1) -H -HLK1 #H destruct
]
- [ lapply (IH2 … HWl1 … HW12 … HK12) /2 width=1 by fqup_fpbg/ -IH2 #H
- elim (da_lstas … H l0) -H
- elim (IH1 … HWl1 … HVX0 … HW12 … HK12) -IH1 -HVX0 /2 width=1 by fqup_fpbg/ #V2 #HWV2 #HV2
+ [ lapply (IH2 … HWd1 … HW12 … HK12) /2 width=1 by fqup_fpbg/ -IH2 #H
+ elim (da_lstas … H d0) -H
+ elim (IH1 … HWd1 … HVX0 … HW12 … HK12) -IH1 -HVX0 /2 width=1 by fqup_fpbg/ #V2 #HWV2 #HV2
elim (lift_total V2 0 (i+1))
/3 width=12 by cpcs_lift, lstas_succ, ex2_intro/
- | elim (IH1 … HWl1 … HVX0 … HW12 … HK12) -IH1 -HVX0
+ | elim (IH1 … HWd1 … HVX0 … HW12 … HK12) -IH1 -HVX0
/3 width=5 by fqup_fpbg, lstas_zero, ex2_intro/
- | elim (IH1 … HVl1 … HVX0 … HV12 … HK12) -IH1 -HVl1 -HVX0 -HV12 -HK12 -IH2 /2 width=1 by fqup_fpbg/ #W2 #HVW2 #HW02
+ | elim (IH1 … HVd1 … HVX0 … HV12 … HK12) -IH1 -HVd1 -HVX0 -HV12 -HK12 -IH2 /2 width=1 by fqup_fpbg/ #W2 #HVW2 #HW02
elim (lift_total W2 0 (i+1))
/4 width=12 by cpcs_lift, lstas_ldef, ex2_intro/
- | elim (IH1 … HVl1 … HVX0 … HVX … HK12) -IH1 -HVl1 -HVX0 -HVX -HK12 -IH2 -V2 /2 width=1 by fqup_fpbg/ -l1 #W2 #HXW2 #HW02
+ | elim (IH1 … HVd1 … HVX0 … HVX … HK12) -IH1 -HVd1 -HVX0 -HVX -HK12 -IH2 -V2 /2 width=1 by fqup_fpbg/ -d1 #W2 #HXW2 #HW02
elim (lift_total W2 0 (i+1))
/3 width=12 by cpcs_lift, lstas_lift, ex2_intro/
]
| #p #_ #_ #HT0 #H1 destruct -IH4 -IH3 -IH2 -IH1
elim (snv_inv_gref … H1)
-| #a #I #V1 #T1 #HG0 #HL0 #HT0 #H1 #l1 #l2 #Hl21 #Hl1 #X2 #H2 #X3 #H3 #L2 #HL12 destruct -IH4 -IH3 -IH2
+| #a #I #V1 #T1 #HG0 #HL0 #HT0 #H1 #d1 #d2 #Hd21 #Hd1 #X2 #H2 #X3 #H3 #L2 #HL12 destruct -IH4 -IH3 -IH2
elim (snv_inv_bind … H1) -H1 #_ #HT1
- lapply (da_inv_bind … Hl1) -Hl1 #Hl1
+ lapply (da_inv_bind … Hd1) -Hd1 #Hd1
elim (lstas_inv_bind1 … H2) -H2 #U1 #HTU1 #H destruct
elim (cpr_inv_bind1 … H3) -H3 *
[ #V2 #T2 #HV12 #HT12 #H destruct
- elim (IH1 … Hl1 … HTU1 … HT12 (L2.ⓑ{I}V2)) -IH1 -Hl1 -HTU1 -HT12 /2 width=1 by fqup_fpbg, lpr_pair/ -T1
+ elim (IH1 … Hd1 … HTU1 … HT12 (L2.ⓑ{I}V2)) -IH1 -Hd1 -HTU1 -HT12 /2 width=1 by fqup_fpbg, lpr_pair/ -T1
/4 width=5 by cpcs_bind2, lpr_cpr_conf, lstas_bind, ex2_intro/
| #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 (IH1 … Hd1 … HTU1 … HT13 (L2.ⓓV1)) -IH1 -Hd1 -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, 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 #U1 #l0 #HV1 #HT1 #HVW1 #HTU1
- lapply (da_inv_flat … Hl1) -Hl1 #Hl1
+| #V1 #T1 #HG0 #HL0 #HT0 #H1 #d1 #d2 #Hd21 #Hd1 #X2 #H2 #X3 #H3 #L2 #HL12 destruct
+ elim (snv_inv_appl … H1) -H1 #a #W1 #U1 #d0 #HV1 #HT1 #HVW1 #HTU1
+ lapply (da_inv_flat … Hd1) -Hd1 #Hd1
elim (lstas_inv_appl1 … H2) -H2 #X #HT1U #H destruct
elim (cpr_inv_appl1 … H3) -H3 *
- [ #V2 #T2 #HV12 #HT12 #H destruct -a -l0 -W1 -U1 -HV1 -IH4 -IH3 -IH2
- elim (IH1 … Hl1 … HT1U … HT12 … HL12) -IH1 -Hl1 -HT1U
+ [ #V2 #T2 #HV12 #HT12 #H destruct -a -d0 -W1 -U1 -HV1 -IH4 -IH3 -IH2
+ elim (IH1 … Hd1 … HT1U … HT12 … HL12) -IH1 -Hd1 -HT1U
/4 width=5 by fqup_fpbg, cpcs_flat, lpr_cpr_conf, lstas_appl, ex2_intro/
| #b #V2 #W2 #W3 #T2 #T3 #HV12 #HW23 #HT23 #H1 #H2 destruct
elim (snv_inv_bind … HT1) -HT1 #HW2 #HT2
- lapply (da_inv_bind … Hl1) -Hl1 #Hl1
+ lapply (da_inv_bind … Hd1) -Hd1 #Hd1
elim (lstas_inv_bind1 … HT1U) -HT1U #U #HT2U #H destruct
- elim (scpds_inv_abst1 … HTU1) -HTU1 #W0 #U0 #HW20 #_ #H destruct -U0 -l0
- elim (snv_fwd_da … HW2) #l0 #HW2l
- lapply (cprs_scpds_div … HW20 … HW2l … HVW1) -W0 #H21
- elim (snv_fwd_da … HV1) #l #HV1l
- elim (da_scpes_aux … IH4 IH3 IH2 … HW2l … HV1l … H21) /2 width=1 by fqup_fpbg/ #_ #H
- <minus_n_O #H0 destruct >(plus_minus_m_m l 1) in HV1l; // -H #HV1l
+ elim (scpds_inv_abst1 … HTU1) -HTU1 #W0 #U0 #HW20 #_ #H destruct -U0 -d0
+ elim (snv_fwd_da … HW2) #d0 #HW2d
+ lapply (cprs_scpds_div … HW20 … HW2d … HVW1) -W0 #H21
+ elim (snv_fwd_da … HV1) #d #HV1d
+ elim (da_scpes_aux … IH4 IH3 IH2 … HW2d … HV1d … H21) /2 width=1 by fqup_fpbg/ #_ #H
+ <minus_n_O #H0 destruct >(plus_minus_m_m d 1) in HV1d; // -H #HV1d
lapply (scpes_cpr_lpr_aux … IH2 IH1 … H21 … HW23 … HV12 … HL12) -H21 /2 width=1 by fqup_fpbg/ #H32
lapply (IH3 … HW23 … HL12) /2 width=1 by fqup_fpbg/ #HW3
lapply (IH3 … HV12 … HL12) /2 width=1 by fqup_fpbg/ #HV2
- lapply (IH2 … HW2l … HW23 … HL12) /2 width=1 by fqup_fpbg/ -HW2 -HW2l #HW3l
- lapply (IH2 … HV1l … HV12 … HL12) /2 width=1 by fqup_fpbg/ -HV1 -HV1l #HV2l
- elim (IH1 … Hl1 … HT2U … HT23 (L2.ⓛW3)) -HT2U /2 width=1 by fqup_fpbg, lpr_pair/ #U3 #HTU3 #HU23
- elim (lsubsv_lstas_trans … g … HTU3 … Hl21 … (L2.ⓓⓝW3.V2)) -HTU3
- [ #U4 #HT3U4 #HU43 -IH1 -IH2 -IH3 -IH4 -l -l1 -HW3 -HV2 -HT2
+ lapply (IH2 … HW2d … HW23 … HL12) /2 width=1 by fqup_fpbg/ -HW2 -HW2d #HW3d
+ lapply (IH2 … HV1d … HV12 … HL12) /2 width=1 by fqup_fpbg/ -HV1 -HV1d #HV2d
+ elim (IH1 … Hd1 … HT2U … HT23 (L2.ⓛW3)) -HT2U /2 width=1 by fqup_fpbg, lpr_pair/ #U3 #HTU3 #HU23
+ elim (lsubsv_lstas_trans … g … HTU3 … Hd21 … (L2.ⓓⓝW3.V2)) -HTU3
+ [ #U4 #HT3U4 #HU43 -IH1 -IH2 -IH3 -IH4 -d -d1 -HW3 -HV2 -HT2
@(ex2_intro … (ⓓ{b}ⓝW3.V2.U4)) /2 width=1 by lstas_bind/ -HT3U4
@(cpcs_canc_dx … (ⓓ{b}ⓝW3.V2.U3)) /2 width=1 by cpcs_bind_dx/ -HU43
@(cpcs_cpr_strap1 … (ⓐV2.ⓛ{b}W3.U3)) /2 width=1 by cpr_beta/
/4 width=3 by cpcs_flat, cpcs_bind2, lpr_cpr_conf/
| -U3
- @(lsubsv_beta … (l-1)) /3 width=7 by fqup_fpbg/
- @shnv_cast [1,2: // ] #l0 #Hl0
- lapply (scpes_le_aux … IH4 IH3 IH2 IH1 … HW3l … HV2l … l0 … H32) -IH4 -IH3 -IH2 -IH1 -HW3l -HV2l -H32
+ @(lsubsv_beta … (d-1)) /3 width=7 by fqup_fpbg/
+ @shnv_cast [1,2: // ] #d0 #Hd0
+ lapply (scpes_le_aux … IH4 IH3 IH2 IH1 … HW3d … HV2d … d0 … H32) -IH4 -IH3 -IH2 -IH1 -HW3d -HV2d -H32
/3 width=5 by fpbg_fpbs_trans, fqup_fpbg, cpr_lpr_fpbs, le_S_S/
| -IH1 -IH3 -IH4 /3 width=9 by fqup_fpbg, lpr_pair/
]
- | #b #V0 #V2 #W0 #W2 #T0 #T2 #HV10 #HV02 #HW02 #HT02 #H1 #H2 destruct -a -l0 -W1 -HV1 -IH4 -IH3 -IH2
+ | #b #V0 #V2 #W0 #W2 #T0 #T2 #HV10 #HV02 #HW02 #HT02 #H1 #H2 destruct -a -d0 -W1 -HV1 -IH4 -IH3 -IH2
elim (snv_inv_bind … HT1) -HT1 #_ #HT0
- lapply (da_inv_bind … Hl1) -Hl1 #Hl1
+ lapply (da_inv_bind … Hd1) -Hd1 #Hd1
elim (lstas_inv_bind1 … HT1U) -HT1U #U0 #HTU0 #H destruct
- elim (IH1 … Hl1 … HTU0 … HT02 (L2.ⓓW2)) -IH1 -Hl1 -HTU0 /2 width=1 by fqup_fpbg, lpr_pair/ -T0 #U2 #HTU2 #HU02
+ elim (IH1 … Hd1 … HTU0 … HT02 (L2.ⓓW2)) -IH1 -Hd1 -HTU0 /2 width=1 by fqup_fpbg, lpr_pair/ -T0 #U2 #HTU2 #HU02
lapply (lpr_cpr_conf … HL12 … HV10) -HV10 #HV10
lapply (lpr_cpr_conf … HL12 … HW02) -L1 #HW02
lapply (cpcs_bind2 b … HW02 … HU02) -HW02 -HU02 #HU02
lapply (cpcs_cpr_strap1 … HU02 (ⓓ{b}W2.ⓐV2.U2) ?)
/4 width=3 by lstas_appl, lstas_bind, cpr_theta, ex2_intro/
]
-| #W1 #T1 #HG0 #HL0 #HT0 #H1 #l1 #l2 #Hl21 #Hl1 #X2 #H2 #X3 #H3 #L2 #HL12 destruct -IH4 -IH3 -IH2
+| #W1 #T1 #HG0 #HL0 #HT0 #H1 #d1 #d2 #Hd21 #Hd1 #X2 #H2 #X3 #H3 #L2 #HL12 destruct -IH4 -IH3 -IH2
elim (snv_inv_cast … H1) -H1 #U1 #_ #HT1 #_ #_ -U1
- lapply (da_inv_flat … Hl1) -Hl1 #Hl1
+ lapply (da_inv_flat … Hd1) -Hd1 #Hd1
lapply (lstas_inv_cast1 … H2) -H2 #HTU1
elim (cpr_inv_cast1 … H3) -H3
[ * #U2 #T2 #_ #HT12 #H destruct
- elim (IH1 … Hl1 … HTU1 … HT12 … HL12) -IH1 -Hl1 -HTU1 -HL12
+ elim (IH1 … Hd1 … HTU1 … HT12 … HL12) -IH1 -Hd1 -HTU1 -HL12
/3 width=3 by fqup_fpbg, lstas_cast, ex2_intro/
- | #HT1X3 elim (IH1 … Hl1 … HTU1 … HT1X3 … HL12) -IH1 -Hl1 -HTU1 -HL12
+ | #HT1X3 elim (IH1 … Hd1 … HTU1 … HT1X3 … HL12) -IH1 -Hd1 -HTU1 -HL12
/2 width=3 by fqup_fpbg, ex2_intro/
]
]
qed-.
lemma da_cprs_lpr: ∀h,g,G,L1,T1. ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
- ∀l. ⦃G, L1⦄ ⊢ T1 ▪[h, g] l →
- ∀T2. ⦃G, L1⦄ ⊢ T1 ➡* T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L2⦄ ⊢ T2 ▪[h, g] l.
-#h #g #G #L1 #T1 #HT1 #l #Hl #T2 #H
+ ∀d. ⦃G, L1⦄ ⊢ T1 ▪[h, g] d →
+ ∀T2. ⦃G, L1⦄ ⊢ T1 ➡* T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L2⦄ ⊢ T2 ▪[h, g] d.
+#h #g #G #L1 #T1 #HT1 #d #Hd #T2 #H
@(cprs_ind … H) -T2 /3 width=6 by snv_cprs_lpr, da_cpr_lpr/
qed-.
lemma lstas_cprs_lpr: ∀h,g,G,L1,T1. ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
- ∀l1,l2. l2 ≤ l1 → ⦃G, L1⦄ ⊢ T1 ▪[h, g] l1 →
- ∀U1. ⦃G, L1⦄ ⊢ T1 •*[h, l2] U1 →
+ ∀d1,d2. d2 ≤ d1 → ⦃G, L1⦄ ⊢ T1 ▪[h, g] d1 →
+ ∀U1. ⦃G, L1⦄ ⊢ T1 •*[h, d2] U1 →
∀T2. ⦃G, L1⦄ ⊢ T1 ➡* T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
- ∃∃U2. ⦃G, L2⦄ ⊢ T2 •*[h, l2] U2 & ⦃G, L2⦄ ⊢ U1 ⬌* U2.
-#h #g #G #L1 #T1 #HT1 #l1 #l2 #Hl21 #Hl1 #U1 #HTU1 #T2 #H
+ ∃∃U2. ⦃G, L2⦄ ⊢ T2 •*[h, d2] U2 & ⦃G, L2⦄ ⊢ U1 ⬌* U2.
+#h #g #G #L1 #T1 #HT1 #d1 #d2 #Hd21 #Hd1 #U1 #HTU1 #T2 #H
@(cprs_ind … H) -T2 [ /2 width=10 by lstas_cpr_lpr/ ]
#T #T2 #HT1T #HTT2 #IHT1 #L2 #HL12
elim (IHT1 L1) // -IHT1 #U #HTU #HU1
-elim (lstas_cpr_lpr … g … Hl21 … HTU … HTT2 … HL12) -HTU -HTT2
-[2,3: /2 width=7 by snv_cprs_lpr, da_cprs_lpr/ ] -T1 -T -l1
+elim (lstas_cpr_lpr … g … Hd21 … HTU … HTT2 … HL12) -HTU -HTT2
+[2,3: /2 width=7 by snv_cprs_lpr, da_cprs_lpr/ ] -T1 -T -d1
/4 width=5 by lpr_cpcs_conf, cpcs_trans, ex2_intro/
qed-.
lemma lstas_cpcs_lpr: ∀h,g,G,L1,T1. ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
- ∀l,l1. l ≤ l1 → ⦃G, L1⦄ ⊢ T1 ▪[h, g] l1 → ∀U1. ⦃G, L1⦄ ⊢ T1 •*[h, l] U1 →
+ ∀d,d1. d ≤ d1 → ⦃G, L1⦄ ⊢ T1 ▪[h, g] d1 → ∀U1. ⦃G, L1⦄ ⊢ T1 •*[h, d] U1 →
∀T2. ⦃G, L1⦄ ⊢ T2 ¡[h, g] →
- ∀l2. l ≤ l2 → ⦃G, L1⦄ ⊢ T2 ▪[h, g] l2 → ∀U2. ⦃G, L1⦄ ⊢ T2 •*[h, l] U2 →
+ ∀d2. d ≤ d2 → ⦃G, L1⦄ ⊢ T2 ▪[h, g] d2 → ∀U2. ⦃G, L1⦄ ⊢ T2 •*[h, d] U2 →
⦃G, L1⦄ ⊢ T1 ⬌* T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L2⦄ ⊢ U1 ⬌* U2.
-#h #g #G #L1 #T1 #HT1 #l #l1 #Hl1 #HTl1 #U1 #HTU1 #T2 #HT2 #l2 #Hl2 #HTl2 #U2 #HTU2 #H #L2 #HL12
+#h #g #G #L1 #T1 #HT1 #d #d1 #Hd1 #HTd1 #U1 #HTU1 #T2 #HT2 #d2 #Hd2 #HTd2 #U2 #HTU2 #H #L2 #HL12
elim (cpcs_inv_cprs … H) -H #T #H1 #H2
-elim (lstas_cprs_lpr … HT1 … Hl1 HTl1 … HTU1 … H1 … HL12) -T1 #W1 #H1 #HUW1
-elim (lstas_cprs_lpr … HT2 … Hl2 HTl2 … HTU2 … H2 … HL12) -T2 #W2 #H2 #HUW2
-lapply (lstas_mono … H1 … H2) -h -T -l #H destruct /2 width=3 by cpcs_canc_dx/
+elim (lstas_cprs_lpr … HT1 … Hd1 HTd1 … HTU1 … H1 … HL12) -T1 #W1 #H1 #HUW1
+elim (lstas_cprs_lpr … HT2 … Hd2 HTd2 … HTU2 … H2 … HL12) -T2 #W2 #H2 #HUW2
+lapply (lstas_mono … H1 … H2) -h -T -d #H destruct /2 width=3 by cpcs_canc_dx/
qed-.
definition IH_da_cpr_lpr: ∀h:sh. sd h → relation3 genv lenv term ≝
λh,g,G,L1,T1. ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
- ∀l. ⦃G, L1⦄ ⊢ T1 ▪[h, g] l →
+ ∀d. ⦃G, L1⦄ ⊢ T1 ▪[h, g] d →
∀T2. ⦃G, L1⦄ ⊢ T1 ➡ T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
- ⦃G, L2⦄ ⊢ T2 ▪[h, g] l.
+ ⦃G, L2⦄ ⊢ T2 ▪[h, g] d.
definition IH_lstas_cpr_lpr: ∀h:sh. sd h → relation3 genv lenv term ≝
λh,g,G,L1,T1. ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
- ∀l1,l2. l2 ≤ l1 → ⦃G, L1⦄ ⊢ T1 ▪[h, g] l1 →
- ∀U1. ⦃G, L1⦄ ⊢ T1 •*[h, l2] U1 →
+ ∀d1,d2. d2 ≤ d1 → ⦃G, L1⦄ ⊢ T1 ▪[h, g] d1 →
+ ∀U1. ⦃G, L1⦄ ⊢ T1 •*[h, d2] U1 →
∀T2. ⦃G, L1⦄ ⊢ T1 ➡ T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
- ∃∃U2. ⦃G, L2⦄ ⊢ T2 •*[h, l2] U2 & ⦃G, L2⦄ ⊢ U1 ⬌* U2.
+ ∃∃U2. ⦃G, L2⦄ ⊢ T2 •*[h, d2] U2 & ⦃G, L2⦄ ⊢ U1 ⬌* U2.
definition IH_snv_lstas: ∀h:sh. sd h → relation3 genv lenv term ≝
λh,g,G,L,T. ⦃G, L⦄ ⊢ T ¡[h, g] →
- ∀l1,l2. l2 ≤ l1 → ⦃G, L⦄ ⊢ T ▪[h, g] l1 →
- ∀U. ⦃G, L⦄ ⊢ T •*[h, l2] U → ⦃G, L⦄ ⊢ U ¡[h, g].
+ ∀d1,d2. d2 ≤ d1 → ⦃G, L⦄ ⊢ T ▪[h, g] d1 →
+ ∀U. ⦃G, L⦄ ⊢ T •*[h, d2] U → ⦃G, L⦄ ⊢ U ¡[h, g].
(* Properties for the preservation results **********************************)
(∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h g G1 L1 T1) →
(∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G1, L1, T1⦄ → IH_da_cpr_lpr h g G1 L1 T1) →
∀G,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G, L1, T1⦄ → ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
- ∀l. ⦃G, L1⦄ ⊢ T1 ▪[h, g] l →
- ∀T2. ⦃G, L1⦄ ⊢ T1 ➡* T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L2⦄ ⊢ T2 ▪[h, g] l.
-#h #g #G0 #L0 #T0 #IH2 #IH1 #G #L1 #T1 #HLT0 #HT1 #l #Hl #T2 #H
+ ∀d. ⦃G, L1⦄ ⊢ T1 ▪[h, g] d →
+ ∀T2. ⦃G, L1⦄ ⊢ T1 ➡* T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L2⦄ ⊢ T2 ▪[h, g] d.
+#h #g #G0 #L0 #T0 #IH2 #IH1 #G #L1 #T1 #HLT0 #HT1 #d #Hd #T2 #H
@(cprs_ind … H) -T2 /4 width=10 by snv_cprs_lpr_aux, fpbg_fpbs_trans, cprs_fpbs/
qed-.
(∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h g G1 L1 T1) →
(∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G1, L1, T1⦄ → IH_da_cpr_lpr h g G1 L1 T1) →
∀G,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G, L1, T1⦄ → ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
- ∀l1. ⦃G, L1⦄ ⊢ T1 ▪[h, g] l1 →
- ∀T2,l2. ⦃G, L1⦄ ⊢ T1 •*➡*[h, g, l2] T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
- l2 ≤ l1 ∧ ⦃G, L2⦄ ⊢ T2 ▪[h, g] l1-l2.
-#h #g #G0 #L0 #T0 #IH3 #IH2 #IH1 #G #L1 #T1 #HLT0 #HT1 #l1 #Hl1 #T2 #l2 * #T #l0 #Hl20 #H #HT1 #HT2 #L2 #HL12
-lapply (da_mono … H … Hl1) -H #H destruct
-lapply (lstas_da_conf … HT1 … Hl1) #Hl12
-lapply (da_cprs_lpr_aux … IH2 IH1 … Hl12 … HT2 … HL12) -IH2 -IH1 -HT2 -HL12
+ ∀d1. ⦃G, L1⦄ ⊢ T1 ▪[h, g] d1 →
+ ∀T2,d2. ⦃G, L1⦄ ⊢ T1 •*➡*[h, g, d2] T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
+ d2 ≤ d1 ∧ ⦃G, L2⦄ ⊢ T2 ▪[h, g] d1-d2.
+#h #g #G0 #L0 #T0 #IH3 #IH2 #IH1 #G #L1 #T1 #HLT0 #HT1 #d1 #Hd1 #T2 #d2 * #T #d0 #Hd20 #H #HT1 #HT2 #L2 #HL12
+lapply (da_mono … H … Hd1) -H #H destruct
+lapply (lstas_da_conf … HT1 … Hd1) #Hd12
+lapply (da_cprs_lpr_aux … IH2 IH1 … Hd12 … HT2 … HL12) -IH2 -IH1 -HT2 -HL12
/3 width=8 by fpbg_fpbs_trans, lstas_fpbs, conj/
qed-.
(∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G1, L1, T1⦄ → IH_da_cpr_lpr h g G1 L1 T1) →
∀G,L,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G, L, T1⦄ → ⦃G, L⦄ ⊢ T1 ¡[h, g] →
∀T2. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G, L, T2⦄ → ⦃G, L⦄ ⊢ T2 ¡[h, g] →
- ∀l11. ⦃G, L⦄ ⊢ T1 ▪[h, g] l11 → ∀l12. ⦃G, L⦄ ⊢ T2 ▪[h, g] l12 →
- ∀l21,l22. ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, l21, l22] T2 →
- ∧∧ l21 ≤ l11 & l22 ≤ l12 & l11 - l21 = l12 - l22.
-#h #g #G0 #L0 #T0 #IH3 #IH2 #IH1 #G #L #T1 #HLT01 #HT1 #T2 #HLT02 #HT2 #l11 #Hl11 #l12 #Hl12 #l21 #l22 * #T #HT1 #HT2
-elim (da_scpds_lpr_aux … IH3 IH2 IH1 … Hl11 … HT1 … L) -Hl11 -HT1 //
-elim (da_scpds_lpr_aux … IH3 IH2 IH1 … Hl12 … HT2 … L) -Hl12 -HT2 //
+ ∀d11. ⦃G, L⦄ ⊢ T1 ▪[h, g] d11 → ∀d12. ⦃G, L⦄ ⊢ T2 ▪[h, g] d12 →
+ ∀d21,d22. ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, d21, d22] T2 →
+ ∧∧ d21 ≤ d11 & d22 ≤ d12 & d11 - d21 = d12 - d22.
+#h #g #G0 #L0 #T0 #IH3 #IH2 #IH1 #G #L #T1 #HLT01 #HT1 #T2 #HLT02 #HT2 #d11 #Hd11 #d12 #Hd12 #d21 #d22 * #T #HT1 #HT2
+elim (da_scpds_lpr_aux … IH3 IH2 IH1 … Hd11 … HT1 … L) -Hd11 -HT1 //
+elim (da_scpds_lpr_aux … IH3 IH2 IH1 … Hd12 … HT2 … L) -Hd12 -HT2 //
/3 width=7 by da_mono, and3_intro/
qed-.
(∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G1, L1, T1⦄ → IH_da_cpr_lpr h g G1 L1 T1) →
(∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G1, L1, T1⦄ → IH_lstas_cpr_lpr h g G1 L1 T1) →
∀G,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G, L1, T1⦄ → ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
- ∀l1,l2. l2 ≤ l1 → ⦃G, L1⦄ ⊢ T1 ▪[h, g] l1 →
- ∀U1. ⦃G, L1⦄ ⊢ T1 •*[h, l2] U1 →
+ ∀d1,d2. d2 ≤ d1 → ⦃G, L1⦄ ⊢ T1 ▪[h, g] d1 →
+ ∀U1. ⦃G, L1⦄ ⊢ T1 •*[h, d2] U1 →
∀T2. ⦃G, L1⦄ ⊢ T1 ➡* T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
- ∃∃U2. ⦃G, L2⦄ ⊢ T2 •*[h, l2] U2 & ⦃G, L2⦄ ⊢ U1 ⬌* U2.
-#h #g #G0 #L0 #T0 #IH3 #IH2 #IH1 #G #L1 #T1 #H01 #HT1 #l1 #l2 #Hl21 #Hl1 #U1 #HTU1 #T2 #H
+ ∃∃U2. ⦃G, L2⦄ ⊢ T2 •*[h, d2] U2 & ⦃G, L2⦄ ⊢ U1 ⬌* U2.
+#h #g #G0 #L0 #T0 #IH3 #IH2 #IH1 #G #L1 #T1 #H01 #HT1 #d1 #d2 #Hd21 #Hd1 #U1 #HTU1 #T2 #H
@(cprs_ind … H) -T2 [ /2 width=10 by/ ]
#T #T2 #HT1T #HTT2 #IHT1 #L2 #HL12
elim (IHT1 L1) // -IHT1 #U #HTU #HU1
-elim (IH1 … Hl21 … HTU … HTT2 … HL12) -IH1 -HTU -HTT2
+elim (IH1 … Hd21 … HTU … HTT2 … HL12) -IH1 -HTU -HTT2
[2: /3 width=12 by da_cprs_lpr_aux/
|3: /3 width=10 by snv_cprs_lpr_aux/
|4: /3 width=5 by fpbg_fpbs_trans, cprs_fpbs/
-] -G0 -L0 -T0 -T1 -T -l1
+] -G0 -L0 -T0 -T1 -T -d1
/4 width=5 by lpr_cpcs_conf, cpcs_trans, ex2_intro/
qed-.
(∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G1, L1, T1⦄ → IH_da_cpr_lpr h g G1 L1 T1) →
(∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G1, L1, T1⦄ → IH_lstas_cpr_lpr h g G1 L1 T1) →
∀G,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G, L1, T1⦄ → ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
- ∀U1,l. ⦃G, L1⦄ ⊢ T1 •*➡*[h, g, l] U1 →
+ ∀U1,d. ⦃G, L1⦄ ⊢ T1 •*➡*[h, g, d] U1 →
∀T2. ⦃G, L1⦄ ⊢ T1 ➡ T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
- ∃∃U2. ⦃G, L2⦄ ⊢ T2 •*➡*[h, g, l] U2 & ⦃G, L2⦄ ⊢ U1 ➡* U2.
-#h #g #G0 #L0 #T0 #IH2 #IH1 #G #L1 #T1 #H01 #HT1 #U1 #l2 * #W1 #l1 #Hl21 #HTl1 #HTW1 #HWU1 #T2 #HT12 #L2 #HL12
+ ∃∃U2. ⦃G, L2⦄ ⊢ T2 •*➡*[h, g, d] U2 & ⦃G, L2⦄ ⊢ U1 ➡* U2.
+#h #g #G0 #L0 #T0 #IH2 #IH1 #G #L1 #T1 #H01 #HT1 #U1 #d2 * #W1 #d1 #Hd21 #HTd1 #HTW1 #HWU1 #T2 #HT12 #L2 #HL12
elim (IH1 … H01 … HTW1 … HT12 … HL12) -IH1 // #W2 #HTW2 #HW12
-lapply (IH2 … H01 … HTl1 … HT12 … HL12) -L0 -T0 // -T1
+lapply (IH2 … H01 … HTd1 … HT12 … HL12) -L0 -T0 // -T1
lapply (lpr_cprs_conf … HL12 … HWU1) -L1 #HWU1
lapply (cpcs_canc_sn … HW12 HWU1) -W1 #H
elim (cpcs_inv_cprs … H) -H /3 width=6 by ex4_2_intro, ex2_intro/
(∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G1, L1, T1⦄ → IH_lstas_cpr_lpr h g G1 L1 T1) →
∀G,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G, L1, T1⦄ → ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
∀T2. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G, L1, T2⦄ → ⦃G, L1⦄ ⊢ T2 ¡[h, g] →
- ∀l1,l2. ⦃G, L1⦄ ⊢ T1 •*⬌*[h, g, l1, l2] T2 →
+ ∀d1,d2. ⦃G, L1⦄ ⊢ T1 •*⬌*[h, g, d1, d2] T2 →
∀U1. ⦃G, L1⦄ ⊢ T1 ➡ U1 → ∀U2. ⦃G, L1⦄ ⊢ T2 ➡ U2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
- ⦃G, L2⦄ ⊢ U1 •*⬌*[h, g, l1, l2] U2.
-#h #g #G0 #L0 #T0 #IH2 #IH1 #G #L1 #T1 #H01 #HT1 #T2 #HT02 #HT2 #l1 #l2 * #T0 #HT10 #HT20 #U1 #HTU1 #U2 #HTU2 #L2 #HL12
+ ⦃G, L2⦄ ⊢ U1 •*⬌*[h, g, d1, d2] U2.
+#h #g #G0 #L0 #T0 #IH2 #IH1 #G #L1 #T1 #H01 #HT1 #T2 #HT02 #HT2 #d1 #d2 * #T0 #HT10 #HT20 #U1 #HTU1 #U2 #HTU2 #L2 #HL12
elim (scpds_cpr_lpr_aux … IH2 IH1 … HT10 … HTU1 … HL12) -HT10 -HTU1 // #X1 #HUX1 #H1
elim (scpds_cpr_lpr_aux … IH2 IH1 … HT20 … HTU2 … HL12) -HT20 -HTU2 // #X2 #HUX2 #H2
elim (cprs_conf … H1 … H2) -T0
(∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G1, L1, T1⦄ → IH_da_cpr_lpr h g G1 L1 T1) →
(∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G1, L1, T1⦄ → IH_lstas_cpr_lpr h g G1 L1 T1) →
∀G,L,T. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G, L, T⦄ → ⦃G, L⦄ ⊢ T ¡[h, g] →
- ∀l,l1. l1 ≤ l → ⦃G, L⦄ ⊢ T ▪[h, g] l → ∀T1. ⦃G, L⦄ ⊢ T •*[h, l1] T1 →
- ∀T2,l2. ⦃G, L⦄ ⊢ T •*➡*[h, g, l2] T2 → ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, l2-l1, l1-l2] T2.
-#h #g #G0 #L0 #T0 #IH4 #IH3 #IH2 #IH1 #G #L #T #H0 #HT #l #l1 #Hl1 #HTl #T1 #HT1 #T2 #l2 * #X #l0 #Hl20 #H #HTX #HXT2
-lapply (da_mono … H … HTl) -H #H destruct
-lapply (lstas_da_conf … HT1 … HTl) #HTl1
-lapply (lstas_da_conf … HTX … HTl) #HXl
-lapply (da_cprs_lpr_aux … IH3 IH2 … HXl … HXT2 L ?)
-[1,2,3: /3 width=8 by fpbg_fpbs_trans, lstas_fpbs/ ] #HTl2
-elim (le_or_ge l1 l2) #Hl12 >(eq_minus_O … Hl12)
-[ elim (da_lstas … HTl2 0) #X2 #HTX2 #_ -IH4 -IH3 -IH2 -IH1 -H0 -HT -HTl -HXl
+ ∀d,d1. d1 ≤ d → ⦃G, L⦄ ⊢ T ▪[h, g] d → ∀T1. ⦃G, L⦄ ⊢ T •*[h, d1] T1 →
+ ∀T2,d2. ⦃G, L⦄ ⊢ T •*➡*[h, g, d2] T2 → ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, d2-d1, d1-d2] T2.
+#h #g #G0 #L0 #T0 #IH4 #IH3 #IH2 #IH1 #G #L #T #H0 #HT #d #d1 #Hd1 #HTd #T1 #HT1 #T2 #d2 * #X #d0 #Hd20 #H #HTX #HXT2
+lapply (da_mono … H … HTd) -H #H destruct
+lapply (lstas_da_conf … HT1 … HTd) #HTd1
+lapply (lstas_da_conf … HTX … HTd) #HXd
+lapply (da_cprs_lpr_aux … IH3 IH2 … HXd … HXT2 L ?)
+[1,2,3: /3 width=8 by fpbg_fpbs_trans, lstas_fpbs/ ] #HTd2
+elim (le_or_ge d1 d2) #Hd12 >(eq_minus_O … Hd12)
+[ elim (da_lstas … HTd2 0) #X2 #HTX2 #_ -IH4 -IH3 -IH2 -IH1 -H0 -HT -HTd -HXd
/5 width=6 by lstas_scpds, scpds_div, cprs_strap1, lstas_cpr, lstas_conf_le, monotonic_le_minus_l, ex4_2_intro/
-| elim (da_lstas … HTl1 0) #X1 #HTX1 #_
+| elim (da_lstas … HTd1 0) #X1 #HTX1 #_
lapply (lstas_conf_le … HTX … HT1) // #HXT1 -HT1
- elim (lstas_cprs_lpr_aux … IH3 IH2 IH1 … HXl … HXT1 … HXT2 L) -IH3 -IH2 -IH1 -HXl -HXT1 -HXT2
+ elim (lstas_cprs_lpr_aux … IH3 IH2 IH1 … HXd … HXT1 … HXT2 L) -IH3 -IH2 -IH1 -HXd -HXT1 -HXT2
/4 width=8 by cpcs_scpes, cpcs_cpr_conf, fpbg_fpbs_trans, lstas_fpbs, lstas_cpr, monotonic_le_minus_l/
]
qed-.
(∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G1, L1, T1⦄ → IH_lstas_cpr_lpr h g G1 L1 T1) →
∀G,L,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G, L, T1⦄ → ⦃G, L⦄ ⊢ T1 ¡[h, g] →
∀T2. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G, L, T2⦄ → ⦃G, L⦄ ⊢ T2 ¡[h, g] →
- ∀l11. ⦃G, L⦄ ⊢ T1 ▪[h, g] l11 → ∀l12. ⦃G, L⦄ ⊢ T2 ▪[h, g] l12 →
- ∀l21,l22,l. l21 + l ≤ l11 → l22 + l ≤ l12 →
- ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, l21, l22] T2 → ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, l21+l, l22+l] T2.
-#h #g #G0 #L0 #T0 #IH4 #IH3 #IH2 #IH1 #G #L #T1 #H01 #HT1 #T2 #H02 #HT2 #l11 #Hl11 #Hl12 #Hl12 #l21 #l22 #l #H1 #H2 * #T0 #HT10 #HT20
-elim (da_lstas … Hl11 (l21+l)) #X1 #HTX1 #_
-elim (da_lstas … Hl12 (l22+l)) #X2 #HTX2 #_
-lapply (lstas_scpds_aux … IH4 IH3 IH2 IH1 … Hl11 … HTX1 … HT10) -HT10
+ ∀d11. ⦃G, L⦄ ⊢ T1 ▪[h, g] d11 → ∀d12. ⦃G, L⦄ ⊢ T2 ▪[h, g] d12 →
+ ∀d21,d22,d. d21 + d ≤ d11 → d22 + d ≤ d12 →
+ ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, d21, d22] T2 → ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, d21+d, d22+d] T2.
+#h #g #G0 #L0 #T0 #IH4 #IH3 #IH2 #IH1 #G #L #T1 #H01 #HT1 #T2 #H02 #HT2 #d11 #Hd11 #Hd12 #Hd12 #d21 #d22 #d #H1 #H2 * #T0 #HT10 #HT20
+elim (da_lstas … Hd11 (d21+d)) #X1 #HTX1 #_
+elim (da_lstas … Hd12 (d22+d)) #X2 #HTX2 #_
+lapply (lstas_scpds_aux … IH4 IH3 IH2 IH1 … Hd11 … HTX1 … HT10) -HT10
[1,2,3: // | >eq_minus_O [2: // ] <minus_plus_m_m_commutative #HX1 ]
-lapply (lstas_scpds_aux … IH4 IH3 IH2 IH1 … Hl12 … HTX2 … HT20) -HT20
+lapply (lstas_scpds_aux … IH4 IH3 IH2 IH1 … Hd12 … HTX2 … HT20) -HT20
[1,2,3: // | >eq_minus_O [2: // ] <minus_plus_m_m_commutative #HX2 ]
-lapply (lstas_scpes_trans … Hl11 … HTX1 … HX1) [ // ] -Hl11 -HTX1 -HX1 -H1 #H1
-lapply (lstas_scpes_trans … Hl12 … HTX2 … HX2) [ // ] -Hl12 -HTX2 -HX2 -H2 #H2
+lapply (lstas_scpes_trans … Hd11 … HTX1 … HX1) [ // ] -Hd11 -HTX1 -HX1 -H1 #H1
+lapply (lstas_scpes_trans … Hd12 … HTX2 … HX2) [ // ] -Hd12 -HTX2 -HX2 -H2 #H2
/2 width=4 by scpes_canc_dx/ (**) (* full auto fails *)
qed-.
qed-.
(* Basic_1: was: pc3_gen_lift *)
-lemma cpcs_inv_lift: ∀G,L,K,s,d,e. ⬇[s, d, e] L ≡ K →
- ∀T1,U1. ⬆[d, e] T1 ≡ U1 → ∀T2,U2. ⬆[d, e] T2 ≡ U2 →
+lemma cpcs_inv_lift: ∀G,L,K,s,l,m. ⬇[s, l, m] L ≡ K →
+ ∀T1,U1. ⬆[l, m] T1 ≡ U1 → ∀T2,U2. ⬆[l, m] T2 ≡ U2 →
⦃G, L⦄ ⊢ U1 ⬌* U2 → ⦃G, K⦄ ⊢ T1 ⬌* T2.
-#G #L #K #s #d #e #HLK #T1 #U1 #HTU1 #T2 #U2 #HTU2 #HU12
+#G #L #K #s #l #m #HLK #T1 #U1 #HTU1 #T2 #U2 #HTU2 #HU12
elim (cpcs_inv_cprs … HU12) -HU12 #U #HU1 #HU2
elim (cprs_inv_lift1 … HU1 … HLK … HTU1) -U1 #T #HTU #HT1
elim (cprs_inv_lift1 … HU2 … HLK … HTU2) -L -U2 #X #HXU
->(lift_inj … HXU … HTU) -X -U -d -e /2 width=3 by cprs_div/
+>(lift_inj … HXU … HTU) -X -U -l -m /2 width=3 by cprs_div/
qed-.
(* Advanced properties ******************************************************)
qed-.
(* Basic_1: was: pc3_lift *)
-lemma cpcs_lift: ∀G,L,K,s,d,e. ⬇[s, d, e] L ≡ K →
- ∀T1,U1. ⬆[d, e] T1 ≡ U1 → ∀T2,U2. ⬆[d, e] T2 ≡ U2 →
+lemma cpcs_lift: ∀G,L,K,s,l,m. ⬇[s, l, m] L ≡ K →
+ ∀T1,U1. ⬆[l, m] T1 ≡ U1 → ∀T2,U2. ⬆[l, m] T2 ≡ U2 →
⦃G, K⦄ ⊢ T1 ⬌* T2 → ⦃G, L⦄ ⊢ U1 ⬌* U2.
-#G #L #K #s #d #e #HLK #T1 #U1 #HTU1 #T2 #U2 #HTU2 #HT12
+#G #L #K #s #l #m #HLK #T1 #U1 #HTU1 #T2 #U2 #HTU2 #HT12
elim (cpcs_inv_cprs … HT12) -HT12 #T #HT1 #HT2
-elim (lift_total T d e) /3 width=12 by cprs_div, cprs_lift/
+elim (lift_total T l m) /3 width=12 by cprs_div, cprs_lift/
qed.
lemma cpcs_strip: ∀G,L,T1,T. ⦃G, L⦄ ⊢ T ⬌* T1 → ∀T2. ⦃G, L⦄ ⊢ T ⬌ T2 →
(* STRATIFIED DECOMPOSED PARALLEL EQUIVALENCE FOR TERMS *********************)
definition scpes: ∀h. sd h → nat → nat → relation4 genv lenv term term ≝
- λh,g,l1,l2,G,L,T1,T2.
- ∃∃T. ⦃G, L⦄ ⊢ T1 •*➡*[h, g, l1] T & ⦃G, L⦄ ⊢ T2 •*➡*[h, g, l2] T.
+ λh,g,d1,d2,G,L,T1,T2.
+ ∃∃T. ⦃G, L⦄ ⊢ T1 •*➡*[h, g, d1] T & ⦃G, L⦄ ⊢ T2 •*➡*[h, g, d2] T.
interpretation "stratified decomposed parallel equivalence (term)"
- 'DPConvStar h g l1 l2 G L T1 T2 = (scpes h g l1 l2 G L T1 T2).
+ 'DPConvStar h g d1 d2 G L T1 T2 = (scpes h g d1 d2 G L T1 T2).
(* Basic properties *********************************************************)
-lemma scpds_div: ∀h,g,G,L,T1,T2,T,l1,l2.
- ⦃G, L⦄ ⊢ T1 •*➡*[h, g, l1] T → ⦃G, L⦄ ⊢ T2 •*➡*[h, g, l2] T →
- ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, l1, l2] T2.
+lemma scpds_div: ∀h,g,G,L,T1,T2,T,d1,d2.
+ ⦃G, L⦄ ⊢ T1 •*➡*[h, g, d1] T → ⦃G, L⦄ ⊢ T2 •*➡*[h, g, d2] T →
+ ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, d1, d2] T2.
/2 width=3 by ex2_intro/ qed.
-lemma scpes_sym: ∀h,g,G,L,T1,T2,l1,l2. ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, l1, l2] T2 →
- ⦃G, L⦄ ⊢ T2 •*⬌*[h, g, l2, l1] T1.
-#h #g #G #L #T1 #T2 #L1 #l2 * /2 width=3 by scpds_div/
+lemma scpes_sym: ∀h,g,G,L,T1,T2,d1,d2. ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, d1, d2] T2 →
+ ⦃G, L⦄ ⊢ T2 •*⬌*[h, g, d2, d1] T1.
+#h #g #G #L #T1 #T2 #L1 #d2 * /2 width=3 by scpds_div/
qed-.
(* Main inversion lemmas about atomic arity assignment on terms *************)
-theorem scpes_aaa_mono: ∀h,g,G,L,T1,T2,l1,l2. ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, l1, l2] T2 →
+theorem scpes_aaa_mono: ∀h,g,G,L,T1,T2,d1,d2. ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, d1, d2] T2 →
∀A1. ⦃G, L⦄ ⊢ T1 ⁝ A1 → ∀A2. ⦃G, L⦄ ⊢ T2 ⁝ A2 →
A1 = A2.
-#h #g #G #L #T1 #T2 #l1 #l2 * #T #HT1 #HT2 #A1 #HA1 #A2 #HA2
+#h #g #G #L #T1 #T2 #d1 #d2 * #T #HT1 #HT2 #A1 #HA1 #A2 #HA2
lapply (scpds_aaa_conf … HA1 … HT1) -T1 #HA1
lapply (scpds_aaa_conf … HA2 … HT2) -T2 #HA2
lapply (aaa_mono … HA1 … HA2) -L -T //
(* Inversion lemmas on parallel equivalence for terms ***********************)
-lemma scpes_inv_lstas_eq: ∀h,g,G,L,T1,T2,l1,l2. ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, l1, l2] T2 →
- ∀U1. ⦃G, L⦄ ⊢ T1 •*[h, l1] U1 →
- ∀U2. ⦃G, L⦄ ⊢ T2 •*[h, l2] U2 → ⦃G, L⦄ ⊢ U1 ⬌* U2.
-#h #g #G #L #T1 #T2 #l1 #l2 * #T #HT1 #HT2 #U1 #HTU1 #U2 #HTU2
+lemma scpes_inv_lstas_eq: ∀h,g,G,L,T1,T2,d1,d2. ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, d1, d2] T2 →
+ ∀U1. ⦃G, L⦄ ⊢ T1 •*[h, d1] U1 →
+ ∀U2. ⦃G, L⦄ ⊢ T2 •*[h, d2] U2 → ⦃G, L⦄ ⊢ U1 ⬌* U2.
+#h #g #G #L #T1 #T2 #d1 #d2 * #T #HT1 #HT2 #U1 #HTU1 #U2 #HTU2
/3 width=8 by scpds_inv_lstas_eq, cprs_div/
qed-.
(* Properties on parallel equivalence for terms *****************************)
-lemma cpcs_scpes: ∀h,g,G,L,T1,l11. ⦃G, L⦄ ⊢ T1 ▪[h, g] l11 →
- ∀U1,l12. l12 ≤ l11 → ⦃G, L⦄ ⊢ T1 •*[h, l12] U1 →
- ∀T2,l21. ⦃G, L⦄ ⊢ T2 ▪[h, g] l21 →
- ∀U2,l22. l22 ≤ l21 → ⦃G, L⦄ ⊢ T2 •*[h, l22] U2 →
- ⦃G, L⦄ ⊢ U1 ⬌* U2 → ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, l12, l22] T2.
-#h #g #G #L #T1 #l11 #HT1 #U1 #l12 #Hl121 #HTU1 #T2 #l21 #HT2 #U2 #l22 #Hl221 #HTU2 #HU12
+lemma cpcs_scpes: ∀h,g,G,L,T1,d11. ⦃G, L⦄ ⊢ T1 ▪[h, g] d11 →
+ ∀U1,d12. d12 ≤ d11 → ⦃G, L⦄ ⊢ T1 •*[h, d12] U1 →
+ ∀T2,d21. ⦃G, L⦄ ⊢ T2 ▪[h, g] d21 →
+ ∀U2,d22. d22 ≤ d21 → ⦃G, L⦄ ⊢ T2 •*[h, d22] U2 →
+ ⦃G, L⦄ ⊢ U1 ⬌* U2 → ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, d12, d22] T2.
+#h #g #G #L #T1 #d11 #HT1 #U1 #d12 #Hd121 #HTU1 #T2 #d21 #HT2 #U2 #d22 #Hd221 #HTU2 #HU12
elim (cpcs_inv_cprs … HU12) -HU12 /3 width=6 by scpds_div, ex4_2_intro/
qed.
(* Advanced inversion lemmas ************************************************)
-lemma scpes_inv_abst2: ∀h,g,a,G,L,T1,T2,W2,l1,l2. ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, l1, l2] ⓛ{a}W2.T2 →
- ∃∃W,T. ⦃G, L⦄ ⊢ T1 •*➡*[h, g, l1] ⓛ{a}W.T & ⦃G, L⦄ ⊢ W2 ➡* W &
- ⦃G, L.ⓛW2⦄ ⊢ T2 •*➡*[h, g, l2] T.
-#h #g #a #G #L #T1 #T2 #W2 #l1 #l2 * #T0 #HT10 #H
+lemma scpes_inv_abst2: ∀h,g,a,G,L,T1,T2,W2,d1,d2. ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, d1, d2] ⓛ{a}W2.T2 →
+ ∃∃W,T. ⦃G, L⦄ ⊢ T1 •*➡*[h, g, d1] ⓛ{a}W.T & ⦃G, L⦄ ⊢ W2 ➡* W &
+ ⦃G, L.ⓛW2⦄ ⊢ T2 •*➡*[h, g, d2] T.
+#h #g #a #G #L #T1 #T2 #W2 #d1 #d2 * #T0 #HT10 #H
elim (scpds_inv_abst1 … H) -H #W #T #HW2 #HT2 #H destruct /2 width=5 by ex3_2_intro/
qed-.
(* Advanced properties ******************************************************)
-lemma scpes_refl: ∀h,g,G,L,T,l1,l2. l2 ≤ l1 → ⦃G, L⦄ ⊢ T ▪[h, g] l1 →
- ⦃G, L⦄ ⊢ T •*⬌*[h, g, l2, l2] T.
-#h #g #G #L #T #l1 #l2 #Hl21 #Hl1
-elim (da_lstas … Hl1 … l2) #U #HTU #_
+lemma scpes_refl: ∀h,g,G,L,T,d1,d2. d2 ≤ d1 → ⦃G, L⦄ ⊢ T ▪[h, g] d1 →
+ ⦃G, L⦄ ⊢ T •*⬌*[h, g, d2, d2] T.
+#h #g #G #L #T #d1 #d2 #Hd21 #Hd1
+elim (da_lstas … Hd1 … d2) #U #HTU #_
/3 width=3 by scpds_div, lstas_scpds/
qed.
-lemma lstas_scpes_trans: ∀h,g,G,L,T1,l0,l1. ⦃G, L⦄ ⊢ T1 ▪[h, g] l0 → l1 ≤ l0 →
- ∀T. ⦃G, L⦄ ⊢ T1 •*[h, l1] T →
- ∀T2,l,l2. ⦃G, L⦄ ⊢ T •*⬌*[h,g,l,l2] T2 → ⦃G, L⦄ ⊢ T1 •*⬌*[h,g,l1+l,l2] T2.
-#h #g #G #L #T1 #l0 #l1 #Hl0 #Hl10 #T #HT1 #T2 #l #l2 *
+lemma lstas_scpes_trans: ∀h,g,G,L,T1,d0,d1. ⦃G, L⦄ ⊢ T1 ▪[h, g] d0 → d1 ≤ d0 →
+ ∀T. ⦃G, L⦄ ⊢ T1 •*[h, d1] T →
+ ∀T2,d,d2. ⦃G, L⦄ ⊢ T •*⬌*[h,g,d,d2] T2 → ⦃G, L⦄ ⊢ T1 •*⬌*[h,g,d1+d,d2] T2.
+#h #g #G #L #T1 #d0 #d1 #Hd0 #Hd10 #T #HT1 #T2 #d #d2 *
/3 width=3 by scpds_div, lstas_scpds_trans/ qed-.
(* Properties on parallel computation for terms *****************************)
lemma cprs_scpds_div: ∀h,g,G,L,T1,T. ⦃G, L⦄ ⊢ T1 ➡* T →
- ∀l. ⦃G, L⦄ ⊢ T1 ▪[h, g] l →
- ∀T2,l2. ⦃G, L⦄ ⊢ T2 •*➡*[h, g, l2] T →
- ⦃G, L⦄⊢ T1 •*⬌*[h, g, 0, l2] T2.
-#h #g #G #L #T1 #T #HT1 #l #Hl elim (da_lstas … Hl 0)
+ ∀d. ⦃G, L⦄ ⊢ T1 ▪[h, g] d →
+ ∀T2,d2. ⦃G, L⦄ ⊢ T2 •*➡*[h, g, d2] T →
+ ⦃G, L⦄⊢ T1 •*⬌*[h, g, 0, d2] T2.
+#h #g #G #L #T1 #T #HT1 #d #Hd elim (da_lstas … Hd 0)
#X1 #HTX1 #_ elim (cprs_strip … HT1 X1) -HT1
/3 width=5 by scpds_strap1, scpds_div, lstas_cpr, ex4_2_intro/
qed.
(* Main properties **********************************************************)
-theorem scpes_trans: ∀h,g,G,L,T1,T,l1,l. ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, l1, l] T →
- ∀T2,l2. ⦃G, L⦄ ⊢ T •*⬌*[h, g, l, l2] T2 → ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, l1, l2] T2.
-#h #g #G #L #T1 #T #l1 #l * #X1 #HT1X1 #HTX1 #T2 #l2 * #X2 #HTX2 #HT2X2
-elim (scpds_conf_eq … HTX1 … HTX2) -T -l /3 width=5 by scpds_cprs_trans, scpds_div/
+theorem scpes_trans: ∀h,g,G,L,T1,T,d1,d. ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, d1, d] T →
+ ∀T2,d2. ⦃G, L⦄ ⊢ T •*⬌*[h, g, d, d2] T2 → ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, d1, d2] T2.
+#h #g #G #L #T1 #T #d1 #d * #X1 #HT1X1 #HTX1 #T2 #d2 * #X2 #HTX2 #HT2X2
+elim (scpds_conf_eq … HTX1 … HTX2) -T -d /3 width=5 by scpds_cprs_trans, scpds_div/
qed-.
-theorem scpes_canc_sn: ∀h,g,G,L,T,T1,l,l1. ⦃G, L⦄ ⊢ T •*⬌*[h, g, l, l1] T1 →
- ∀T2,l2. ⦃G, L⦄ ⊢ T •*⬌*[h, g, l, l2] T2 → ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, l1, l2] T2.
+theorem scpes_canc_sn: ∀h,g,G,L,T,T1,d,d1. ⦃G, L⦄ ⊢ T •*⬌*[h, g, d, d1] T1 →
+ ∀T2,d2. ⦃G, L⦄ ⊢ T •*⬌*[h, g, d, d2] T2 → ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, d1, d2] T2.
/3 width=4 by scpes_trans, scpes_sym/ qed-.
-theorem scpes_canc_dx: ∀h,g,G,L,T1,T,l1,l. ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, l1, l] T →
- ∀T2,l2. ⦃G, L⦄ ⊢ T2 •*⬌*[h, g, l2, l] T → ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, l1, l2] T2.
+theorem scpes_canc_dx: ∀h,g,G,L,T1,T,d1,d. ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, d1, d] T →
+ ∀T2,d2. ⦃G, L⦄ ⊢ T2 •*⬌*[h, g, d2, d] T → ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, d1, d2] T2.
/3 width=4 by scpes_trans, scpes_sym/ qed-.
(* Basic properties *********************************************************)
-lemma Delta_lift: ∀W1,W2,d,e. ⬆[d, e] W1 ≡ W2 →
- ⬆[d, e] (Delta W1) ≡ (Delta W2).
+lemma Delta_lift: ∀W1,W2,l,m. ⬆[l, m] W1 ≡ W2 →
+ ⬆[l, m] (Delta W1) ≡ (Delta W2).
/4 width=1 by lift_flat, lift_bind, lift_lref_lt/ qed.
(* Main properties **********************************************************)
(* Basic properties *********************************************************)
-lemma ApplDelta_lift: ∀W1,W2,k,d,e. ⬆[d, e] W1 ≡ W2 →
- ⬆[d, e] (ApplDelta W1 k) ≡ (ApplDelta W2 k).
+lemma ApplDelta_lift: ∀W1,W2,k,l,m. ⬆[l, m] W1 ≡ W2 →
+ ⬆[l, m] (ApplDelta W1 k) ≡ (ApplDelta W2 k).
/5 width=1 by lift_flat, lift_bind, lift_lref_lt/ qed.
lemma cpr_ApplOmega_12: ∀G,L,W,k. ⦃G, L⦄ ⊢ ApplOmega1 W k ➡ ApplOmega2 W k.
(* extended validity of a closure, last arg of snv_appl > 1 *)
lemma snv_extended: ∀h,g,a,G,L,k. ⦃G, L.ⓛ⋆k.ⓛⓛ{a}⋆k.⋆k.ⓛ#0⦄ ⊢ ⓐ#2.#0 ¡[h, g].
#h #g #a #G #L #k elim (deg_total h g k)
-#l #Hl @(snv_appl … a … (⋆k) … (⋆k) (0+1+1))
+#d #Hd @(snv_appl … a … (⋆k) … (⋆k) (0+1+1))
[ /4 width=5 by snv_lref, drop_drop_lt/
| /4 width=13 by snv_bind, snv_lref/
| /5 width=6 by lstas_scpds, lstas_succ, da_ldec, da_sort, drop_drop_lt/
-| @(lstas_scpds … (l+1+1))
+| @(lstas_scpds … (d+1+1))
/5 width=11 by lstas_bind, lstas_succ, da_bind, da_ldec, da_sort, lift_bind/
]
qed.
(* restricted validity of the η-expanded closure, last arg of snv_appl = 1 **)
lemma snv_restricted: ∀h,g,a,G,L,k. ⦃G, L.ⓛ⋆k.ⓛⓛ{a}⋆k.⋆k.ⓛⓛ{a}⋆k.ⓐ#0.#1⦄ ⊢ ⓐ#2.#0 ¡[h, g].
#h #g #a #G #L #k elim (deg_total h g k)
-#l #Hl @(snv_appl … a … (⋆k) … (ⓐ#0.#2) (0+1))
+#d #Hd @(snv_appl … a … (⋆k) … (ⓐ#0.#2) (0+1))
[ /4 width=5 by snv_lref, drop_drop_lt/
| @snv_lref [4: // |1,2,3: skip ]
@snv_bind //
@(snv_appl … a … (⋆k) … (⋆k) (0+1))
[ @snv_lref [4: // |1,2,3: skip ] //
| @snv_lref [4: /2 width=1 by drop_drop_lt/ |1,2,3: skip ] @snv_bind //
- | @(lstas_scpds … (l+1)) /3 width=6 by da_sort, da_ldec, lstas_succ/
- | @(lstas_scpds … (l+1)) /3 width=8 by lstas_succ, lstas_bind, drop_drop, lift_bind/
+ | @(lstas_scpds … (d+1)) /3 width=6 by da_sort, da_ldec, lstas_succ/
+ | @(lstas_scpds … (d+1)) /3 width=8 by lstas_succ, lstas_bind, drop_drop, lift_bind/
@da_ldec [3: /2 width=1 by drop_drop_lt/ |1,2: skip ] /3 width=1 by da_sort, da_bind/
]
| /5 width=6 by lstas_scpds, lstas_succ, da_ldec, da_sort, drop_drop_lt/
-| @(lstas_scpds … (l+1+1)) //
+| @(lstas_scpds … (d+1+1)) //
[ @da_ldec [3: // |1,2: skip ]
@da_bind @da_flat @da_ldec [3: /2 width=1 by drop_drop_lt/ |1,2: skip ]
/3 width=1 by da_sort, da_bind/
(* Basic inversion lemmas ***************************************************)
+fact destruct_apair_apair_aux: ∀A1,A2,B1,B2. ②B1.A1 = ②B2.A2 → B1 = B2 ∧ A1 = A2.
+#A1 #A2 #B1 #B2 #H destruct /2 width=1 by conj/
+qed-.
+
lemma discr_apair_xy_x: ∀A,B. ②B. A = B → ⊥.
#A #B elim B -B
[ #H destruct
-| #Y #X #IHY #_ #H destruct
- -H >e0 in e1; normalize (**) (* destruct: one quality is not simplified, the destructed equality is not erased *)
- /2 width=1/
+| #Y #X #IHY #_ #H elim (destruct_apair_apair_aux … H) -H /2 width=1 by/ (**) (* destruct lemma needed *)
]
qed-.
lemma discr_tpair_xy_y: ∀B,A. ②B. A = A → ⊥.
#B #A elim A -A
[ #H destruct
-| #Y #X #_ #IHX #H destruct
- -H (**) (* destruct: the destructed equality is not erased *)
- /2 width=1/
+| #Y #X #_ #IHX #H elim (destruct_apair_apair_aux … H) -H /2 width=1 by/ (**) (* destruct lemma needed *)
]
qed-.
lemma eq_aarity_dec: ∀A1,A2:aarity. Decidable (A1 = A2).
#A1 elim A1 -A1
-[ #A2 elim A2 -A2 /2 width=1/
+[ #A2 elim A2 -A2 /2 width=1 by or_introl/
#B2 #A2 #_ #_ @or_intror #H destruct
| #B1 #A1 #IHB1 #IHA1 #A2 elim A2 -A2
[ -IHB1 -IHA1 @or_intror #H destruct
| #B2 #A2 #_ #_ elim (IHB1 B2) -IHB1
[ #H destruct elim (IHA1 A2) -IHA1
- [ #H destruct /2 width=1/
- | #HA12 @or_intror #H destruct /2 width=1/
+ [ #H destruct /2 width=1 by or_introl/
+ | #HA12 @or_intror #H destruct /2 width=1 by/
]
- | -IHA1 #HB12 @or_intror #H destruct /2 width=1/
+ | -IHA1 #HB12 @or_intror #H destruct /2 width=1 by/
]
]
]
| Flat2: flat2 → item2 (* non-binding item *)
.
+(* Basic inversion lemmas ***************************************************)
+
+fact destruct_sort_sort_aux: ∀k1,k2. Sort k1 = Sort k2 → k1 = k2.
+#k1 #k2 #H destruct //
+qed-.
+
(* Basic properties *********************************************************)
lemma eq_item0_dec: ∀I1,I2:item0. Decidable (I1 = I2).
(* Basic inversion lemmas ***************************************************)
-lemma destruct_lpair_lpair: ∀I1,I2,L1,L2,V1,V2. L1.ⓑ{I1}V1 = L2.ⓑ{I2}V2 →
- ∧∧L1 = L2 & I1 = I2 & V1 = V2.
+fact destruct_lpair_lpair_aux: ∀I1,I2,L1,L2,V1,V2. L1.ⓑ{I1}V1 = L2.ⓑ{I2}V2 →
+ ∧∧L1 = L2 & I1 = I2 & V1 = V2.
#I1 #I2 #L1 #L2 #V1 #V2 #H destruct /2 width=1 by and3_intro/
qed-.
#I #V #L elim L -L
[ #H destruct
| #L #J #W #IHL #H
- elim (destruct_lpair_lpair … H) -H #H1 #H2 #H3 destruct /2 width=1 by/ (**) (* destruct lemma needed *)
+ elim (destruct_lpair_lpair_aux … H) -H #H1 #H2 #H3 destruct /2 width=1 by/ (**) (* destruct lemma needed *)
]
qed-.
interpretation "tail abstraction (local environment)"
'SnAbst L T = (append (LPair LAtom Abst T) L).
-definition l_appendable_sn: predicate (lenv→relation term) ≝ λR.
+definition d_appendable_sn: predicate (lenv→relation term) ≝ λR.
∀K,T1,T2. R K T1 T2 → ∀L. R (L @@ K) T1 T2.
(* Basic properties *********************************************************)
| #K1 #I1 #V1 #IH * normalize
[ #L1 #L2 #_ <plus_n_Sm #H destruct
| #K2 #I2 #V2 #L1 #L2 #H1 #H2
- elim (destruct_lpair_lpair … H1) -H1 #H1 #H3 #H4 destruct (**) (* destruct lemma needed *)
+ elim (destruct_lpair_lpair_aux … H1) -H1 #H1 #H3 #H4 destruct (**) (* destruct lemma needed *)
elim (IH … H1) -IH -H1 /2 width=1 by conj/
]
]
normalize in H2; >append_length in H2; #H
elim (plus_xySz_x_false … (sym_eq … H))
| #K2 #I2 #V2 #L1 #L2 #H1 #H2
- elim (destruct_lpair_lpair … H1) -H1 #H1 #H3 #H4 destruct (**) (* destruct lemma needed *)
+ elim (destruct_lpair_lpair_aux … H1) -H1 #H1 #H3 #H4 destruct (**) (* destruct lemma needed *)
elim (IH … H1) -IH -H1 /2 width=1 by conj/
]
]
#I #L #K #V #H elim (append_inj_dx … (⋆.ⓑ{I}V) … H) //
qed-.
-lemma length_inv_pos_dx_ltail: ∀L,d. |L| = d + 1 →
- ∃∃I,K,V. |K| = d & L = ⓑ{I}V.K.
-#Y #d #H elim (length_inv_pos_dx … H) -H #I #L #V #Hd #HLK destruct
+lemma length_inv_pos_dx_ltail: ∀L,l. |L| = l + 1 →
+ ∃∃I,K,V. |K| = l & L = ⓑ{I}V.K.
+#Y #l #H elim (length_inv_pos_dx … H) -H #I #L #V #Hl #HLK destruct
elim (lpair_ltail L I V) /2 width=5 by ex2_3_intro/
qed-.
-lemma length_inv_pos_sn_ltail: ∀L,d. d + 1 = |L| →
- ∃∃I,K,V. d = |K| & L = ⓑ{I}V.K.
-#Y #d #H elim (length_inv_pos_sn … H) -H #I #L #V #Hd #HLK destruct
+lemma length_inv_pos_sn_ltail: ∀L,l. l + 1 = |L| →
+ ∃∃I,K,V. l = |K| & L = ⓑ{I}V.K.
+#Y #l #H elim (length_inv_pos_sn … H) -H #I #L #V #Hl #HLK destruct
elim (lpair_ltail L I V) /2 width=5 by ex2_3_intro/
qed-.
* // #L #I #V normalize <plus_n_Sm #H destruct
qed-.
-lemma length_inv_pos_dx: ∀d,L. |L| = d + 1 →
- ∃∃I,K,V. |K| = d & L = K. ⓑ{I}V.
-#d *
+lemma length_inv_pos_dx: ∀l,L. |L| = l + 1 →
+ ∃∃I,K,V. |K| = l & L = K. ⓑ{I}V.
+#l *
[ normalize <plus_n_Sm #H destruct
| #K #I #V #H
- lapply (injective_plus_l … H) -H /2 width=5/
+ lapply (injective_plus_l … H) -H /2 width=5 by ex2_3_intro/
]
qed-.
-lemma length_inv_pos_sn: ∀d,L. d + 1 = |L| →
- ∃∃I,K,V. d = |K| & L = K. ⓑ{I}V.
-#d *
+lemma length_inv_pos_sn: ∀l,L. l + 1 = |L| →
+ ∃∃I,K,V. l = |K| & L = K. ⓑ{I}V.
+#l *
[ normalize <plus_n_Sm #H destruct
| #K #I #V #H
- lapply (injective_plus_l … H) -H /2 width=5/
+ lapply (injective_plus_l … H) -H /2 width=5 by ex2_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 "ground_2/ynat/ynat_lt.ma".
-include "basic_2/notation/relations/midiso_4.ma".
-include "basic_2/grammar/lenv_length.ma".
-
-(* EQUIVALENCE FOR LOCAL ENVIRONMENTS ***************************************)
-
-inductive leq: relation4 ynat ynat lenv lenv ≝
-| leq_atom: ∀d,e. leq d e (⋆) (⋆)
-| leq_zero: ∀I1,I2,L1,L2,V1,V2.
- leq 0 0 L1 L2 → leq 0 0 (L1.ⓑ{I1}V1) (L2.ⓑ{I2}V2)
-| 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)
-.
-
-interpretation
- "equivalence (local environment)"
- 'MidIso d e L1 L2 = (leq d e L1 L2).
-
-(* Basic properties *********************************************************)
-
-lemma leq_pair_lt: ∀I,L1,L2,V,e. L1 ⩬[0, ⫰e] L2 → 0 < e →
- L1.ⓑ{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: ∀I1,I2,L1,L2,V1,V2,d,e. L1 ⩬[⫰d, e] L2 → 0 < d →
- L1.ⓑ{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: ∀L1,L2. L1 ⩬[0, ∞] L2 →
- ∀I,V. L1.ⓑ{I}V ⩬[0, ∞] L2.ⓑ{I}V.
-#L1 #L2 #HL12 #I #V lapply (leq_pair I … V … HL12) -HL12 //
-qed.
-
-lemma leq_refl: ∀L,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/
-#e elim (ynat_cases … e) [| * #x ]
-#He destruct /2 width=1 by leq_zero, leq_pair/
-qed.
-
-lemma leq_O2: ∀L1,L2,d. |L1| = |L2| → 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 ]
-#H lapply (injective_plus_l … H) -H #HL12
-elim (ynat_cases d) /3 width=1 by leq_zero/
-* /3 width=1 by leq_succ/
-qed.
-
-lemma leq_sym: ∀d,e. symmetric … (leq d e).
-#d #e #L1 #L2 #H elim H -L1 -L2 -d -e
-/2 width=1 by leq_zero, leq_pair, leq_succ/
-qed-.
-
-(* Basic inversion lemmas ***************************************************)
-
-fact leq_inv_atom1_aux: ∀L1,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
-| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #H destruct
-]
-qed-.
-
-lemma leq_inv_atom1: ∀L2,d,e. ⋆ ⩬[d, e] L2 → L2 = ⋆.
-/2 width=5 by leq_inv_atom1_aux/ qed-.
-
-fact leq_inv_zero1_aux: ∀L1,L2,d,e. L1 ⩬[d, e] L2 →
- ∀J1,K1,W1. L1 = K1.ⓑ{J1}W1 → d = 0 → e = 0 →
- ∃∃J2,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
- /2 width=5 by ex2_3_intro/
-| #I #L1 #L2 #V #e #_ #J1 #K1 #W1 #_ #_ #H
- elim (ysucc_inv_O_dx … H)
-| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #J1 #K1 #W1 #_ #H
- elim (ysucc_inv_O_dx … H)
-]
-qed-.
-
-lemma leq_inv_zero1: ∀I1,K1,L2,V1. K1.ⓑ{I1}V1 ⩬[0, 0] L2 →
- ∃∃I2,K2,V2. K1 ⩬[0, 0] K2 & L2 = K2.ⓑ{I2}V2.
-/2 width=9 by leq_inv_zero1_aux/ qed-.
-
-fact leq_inv_pair1_aux: ∀L1,L2,d,e. L1 ⩬[d, e] L2 →
- ∀J,K1,W. L1 = K1.ⓑ{J}W → d = 0 → 0 < e →
- ∃∃K2. 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
- elim (ylt_yle_false … H) //
-| #I #L1 #L2 #V #e #HL12 #J #K1 #W #H #_ #_ destruct
- /2 width=3 by ex2_intro/
-| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #J #K1 #W #_ #H
- elim (ysucc_inv_O_dx … H)
-]
-qed-.
-
-lemma leq_inv_pair1: ∀I,K1,L2,V,e. K1.ⓑ{I}V ⩬[0, e] L2 → 0 < e →
- ∃∃K2. K1 ⩬[0, ⫰e] K2 & L2 = K2.ⓑ{I}V.
-/2 width=6 by leq_inv_pair1_aux/ qed-.
-
-lemma leq_inv_pair: ∀I1,I2,L1,L2,V1,V2,e. L1.ⓑ{I1}V1 ⩬[0, e] L2.ⓑ{I2}V2 → 0 < e →
- ∧∧ 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: ∀L1,L2,d,e. L1 ⩬[d, e] L2 →
- ∀J1,K1,W1. L1 = K1.ⓑ{J1}W1 → 0 < d →
- ∃∃J2,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
- elim (ylt_yle_false … H) //
-| #I #L1 #L2 #V #e #_ #J1 #K1 #W1 #_ #H
- elim (ylt_yle_false … H) //
-| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #HL12 #J1 #K1 #W1 #H #_ destruct
- /2 width=5 by ex2_3_intro/
-]
-qed-.
-
-lemma leq_inv_succ1: ∀I1,K1,L2,V1,d,e. K1.ⓑ{I1}V1 ⩬[d, e] L2 → 0 < d →
- ∃∃I2,K2,V2. K1 ⩬[⫰d, e] K2 & L2 = K2.ⓑ{I2}V2.
-/2 width=5 by leq_inv_succ1_aux/ qed-.
-
-lemma leq_inv_atom2: ∀L1,d,e. L1 ⩬[d, e] ⋆ → L1 = ⋆.
-/3 width=3 by leq_inv_atom1, leq_sym/
-qed-.
-
-lemma leq_inv_succ: ∀I1,I2,L1,L2,V1,V2,d,e. L1.ⓑ{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: ∀I2,K2,L1,V2. L1 ⩬[0, 0] K2.ⓑ{I2}V2 →
- ∃∃I1,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: ∀I,K2,L1,V,e. L1 ⩬[0, e] K2.ⓑ{I}V → 0 < e →
- ∃∃K1. 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: ∀I2,K2,L1,V2,d,e. L1 ⩬[d, e] K2.ⓑ{I2}V2 → 0 < d →
- ∃∃I1,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: ∀L1,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: ∀L1,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/
-| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #_ #H elim (ysucc_inv_O_dx … H)
-]
-qed-.
-
-lemma leq_inv_O_Y: ∀L1,L2. L1 ⩬[0, ∞] L2 → L1 = L2.
-/2 width=5 by leq_inv_O_Y_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 "ground_2/ynat/ynat_plus.ma".
-include "basic_2/grammar/leq.ma".
-
-(* EQUIVALENCE FOR LOCAL ENVIRONMENTS ***************************************)
-
-(* Main properties **********************************************************)
-
-theorem leq_trans: ∀d,e. Transitive … (leq d e).
-#d #e #L1 #L2 #H elim H -L1 -L2 -d -e
-[ #d #e #X #H lapply (leq_inv_atom1 … H) -H
- #H destruct //
-| #I1 #I #L1 #L #V1 #V #_ #IHL1 #X #H elim (leq_inv_zero1 … H) -H
- #I2 #L2 #V2 #HL2 #H destruct /3 width=1 by leq_zero/
-| #I #L1 #L #V #e #_ #IHL1 #X #H elim (leq_inv_pair1 … H) -H //
- #L2 #HL2 #H destruct /3 width=1 by leq_pair/
-| #I1 #I #L1 #L #V1 #V #d #e #_ #IHL1 #X #H elim (leq_inv_succ1 … H) -H //
- #I2 #L2 #V2 #HL2 #H destruct /3 width=1 by leq_succ/
-]
-qed-.
-
-theorem leq_canc_sn: ∀d,e,L,L1,L2. L ⩬[d, e] L1 → L ⩬[d, e] L2 → L1 ⩬[d, e] L2.
-/3 width=3 by leq_trans, leq_sym/ qed-.
-
-theorem leq_canc_dx: ∀d,e,L,L1,L2. L1 ⩬[d, e] L → L2 ⩬[d, e] L → L1 ⩬[d, e] L2.
-/3 width=3 by leq_trans, leq_sym/ qed-.
-
-theorem leq_join: ∀L1,L2,d,i. L1 ⩬[d, i] L2 →
- ∀e. L1 ⩬[i+d, e] L2 → 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/
-| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #IHL12 #e #H
- lapply (leq_inv_succ … H ?) -H // >yplus_succ2 >ypred_succ /3 width=1 by leq_succ/
-]
-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 "ground_2/ynat/ynat_lt.ma".
+include "basic_2/notation/relations/midiso_4.ma".
+include "basic_2/grammar/lenv_length.ma".
+
+(* EQUIVALENCE FOR LOCAL ENVIRONMENTS ***************************************)
+
+inductive lreq: relation4 ynat ynat lenv lenv ≝
+| lreq_atom: ∀l,m. lreq l m (⋆) (⋆)
+| lreq_zero: ∀I1,I2,L1,L2,V1,V2.
+ lreq 0 0 L1 L2 → lreq 0 0 (L1.ⓑ{I1}V1) (L2.ⓑ{I2}V2)
+| lreq_pair: ∀I,L1,L2,V,m. lreq 0 m L1 L2 →
+ lreq 0 (⫯m) (L1.ⓑ{I}V) (L2.ⓑ{I}V)
+| lreq_succ: ∀I1,I2,L1,L2,V1,V2,l,m.
+ lreq l m L1 L2 → lreq (⫯l) m (L1.ⓑ{I1}V1) (L2.ⓑ{I2}V2)
+.
+
+interpretation
+ "equivalence (local environment)"
+ 'MidIso l m L1 L2 = (lreq l m L1 L2).
+
+(* Basic properties *********************************************************)
+
+lemma lreq_pair_lt: ∀I,L1,L2,V,m. L1 ⩬[0, ⫰m] L2 → 0 < m →
+ L1.ⓑ{I}V ⩬[0, m] L2.ⓑ{I}V.
+#I #L1 #L2 #V #m #HL12 #Hm <(ylt_inv_O1 … Hm) /2 width=1 by lreq_pair/
+qed.
+
+lemma lreq_succ_lt: ∀I1,I2,L1,L2,V1,V2,l,m. L1 ⩬[⫰l, m] L2 → 0 < l →
+ L1.ⓑ{I1}V1 ⩬[l, m] L2. ⓑ{I2}V2.
+#I1 #I2 #L1 #L2 #V1 #V2 #l #m #HL12 #Hl <(ylt_inv_O1 … Hl) /2 width=1 by lreq_succ/
+qed.
+
+lemma lreq_pair_O_Y: ∀L1,L2. L1 ⩬[0, ∞] L2 →
+ ∀I,V. L1.ⓑ{I}V ⩬[0, ∞] L2.ⓑ{I}V.
+#L1 #L2 #HL12 #I #V lapply (lreq_pair I … V … HL12) -HL12 //
+qed.
+
+lemma lreq_refl: ∀L,l,m. L ⩬[l, m] L.
+#L elim L -L //
+#L #I #V #IHL #l elim (ynat_cases … l) [| * #x ]
+#Hl destruct /2 width=1 by lreq_succ/
+#m elim (ynat_cases … m) [| * #x ]
+#Hm destruct /2 width=1 by lreq_zero, lreq_pair/
+qed.
+
+lemma lreq_O2: ∀L1,L2,l. |L1| = |L2| → L1 ⩬[l, yinj 0] L2.
+#L1 elim L1 -L1 [| #L1 #I1 #V1 #IHL1 ]
+* // [1,3: #L2 #I2 #V2 ] #l normalize
+[1,3: <plus_n_Sm #H destruct ]
+#H lapply (injective_plus_l … H) -H #HL12
+elim (ynat_cases l) /3 width=1 by lreq_zero/
+* /3 width=1 by lreq_succ/
+qed.
+
+lemma lreq_sym: ∀l,m. symmetric … (lreq l m).
+#l #m #L1 #L2 #H elim H -L1 -L2 -l -m
+/2 width=1 by lreq_zero, lreq_pair, lreq_succ/
+qed-.
+
+(* Basic inversion lemmas ***************************************************)
+
+fact lreq_inv_atom1_aux: ∀L1,L2,l,m. L1 ⩬[l, m] L2 → L1 = ⋆ → L2 = ⋆.
+#L1 #L2 #l #m * -L1 -L2 -l -m //
+[ #I1 #I2 #L1 #L2 #V1 #V2 #_ #H destruct
+| #I #L1 #L2 #V #m #_ #H destruct
+| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #_ #H destruct
+]
+qed-.
+
+lemma lreq_inv_atom1: ∀L2,l,m. ⋆ ⩬[l, m] L2 → L2 = ⋆.
+/2 width=5 by lreq_inv_atom1_aux/ qed-.
+
+fact lreq_inv_zero1_aux: ∀L1,L2,l,m. L1 ⩬[l, m] L2 →
+ ∀J1,K1,W1. L1 = K1.ⓑ{J1}W1 → l = 0 → m = 0 →
+ ∃∃J2,K2,W2. K1 ⩬[0, 0] K2 & L2 = K2.ⓑ{J2}W2.
+#L1 #L2 #l #m * -L1 -L2 -l -m
+[ #l #m #J1 #K1 #W1 #H destruct
+| #I1 #I2 #L1 #L2 #V1 #V2 #HL12 #J1 #K1 #W1 #H #_ #_ destruct
+ /2 width=5 by ex2_3_intro/
+| #I #L1 #L2 #V #m #_ #J1 #K1 #W1 #_ #_ #H
+ elim (ysucc_inv_O_dx … H)
+| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #_ #J1 #K1 #W1 #_ #H
+ elim (ysucc_inv_O_dx … H)
+]
+qed-.
+
+lemma lreq_inv_zero1: ∀I1,K1,L2,V1. K1.ⓑ{I1}V1 ⩬[0, 0] L2 →
+ ∃∃I2,K2,V2. K1 ⩬[0, 0] K2 & L2 = K2.ⓑ{I2}V2.
+/2 width=9 by lreq_inv_zero1_aux/ qed-.
+
+fact lreq_inv_pair1_aux: ∀L1,L2,l,m. L1 ⩬[l, m] L2 →
+ ∀J,K1,W. L1 = K1.ⓑ{J}W → l = 0 → 0 < m →
+ ∃∃K2. K1 ⩬[0, ⫰m] K2 & L2 = K2.ⓑ{J}W.
+#L1 #L2 #l #m * -L1 -L2 -l -m
+[ #l #m #J #K1 #W #H destruct
+| #I1 #I2 #L1 #L2 #V1 #V2 #_ #J #K1 #W #_ #_ #H
+ elim (ylt_yle_false … H) //
+| #I #L1 #L2 #V #m #HL12 #J #K1 #W #H #_ #_ destruct
+ /2 width=3 by ex2_intro/
+| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #_ #J #K1 #W #_ #H
+ elim (ysucc_inv_O_dx … H)
+]
+qed-.
+
+lemma lreq_inv_pair1: ∀I,K1,L2,V,m. K1.ⓑ{I}V ⩬[0, m] L2 → 0 < m →
+ ∃∃K2. K1 ⩬[0, ⫰m] K2 & L2 = K2.ⓑ{I}V.
+/2 width=6 by lreq_inv_pair1_aux/ qed-.
+
+lemma lreq_inv_pair: ∀I1,I2,L1,L2,V1,V2,m. L1.ⓑ{I1}V1 ⩬[0, m] L2.ⓑ{I2}V2 → 0 < m →
+ ∧∧ L1 ⩬[0, ⫰m] L2 & I1 = I2 & V1 = V2.
+#I1 #I2 #L1 #L2 #V1 #V2 #m #H #Hm elim (lreq_inv_pair1 … H) -H //
+#Y #HL12 #H destruct /2 width=1 by and3_intro/
+qed-.
+
+fact lreq_inv_succ1_aux: ∀L1,L2,l,m. L1 ⩬[l, m] L2 →
+ ∀J1,K1,W1. L1 = K1.ⓑ{J1}W1 → 0 < l →
+ ∃∃J2,K2,W2. K1 ⩬[⫰l, m] K2 & L2 = K2.ⓑ{J2}W2.
+#L1 #L2 #l #m * -L1 -L2 -l -m
+[ #l #m #J1 #K1 #W1 #H destruct
+| #I1 #I2 #L1 #L2 #V1 #V2 #_ #J1 #K1 #W1 #_ #H
+ elim (ylt_yle_false … H) //
+| #I #L1 #L2 #V #m #_ #J1 #K1 #W1 #_ #H
+ elim (ylt_yle_false … H) //
+| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #HL12 #J1 #K1 #W1 #H #_ destruct
+ /2 width=5 by ex2_3_intro/
+]
+qed-.
+
+lemma lreq_inv_succ1: ∀I1,K1,L2,V1,l,m. K1.ⓑ{I1}V1 ⩬[l, m] L2 → 0 < l →
+ ∃∃I2,K2,V2. K1 ⩬[⫰l, m] K2 & L2 = K2.ⓑ{I2}V2.
+/2 width=5 by lreq_inv_succ1_aux/ qed-.
+
+lemma lreq_inv_atom2: ∀L1,l,m. L1 ⩬[l, m] ⋆ → L1 = ⋆.
+/3 width=3 by lreq_inv_atom1, lreq_sym/
+qed-.
+
+lemma lreq_inv_succ: ∀I1,I2,L1,L2,V1,V2,l,m. L1.ⓑ{I1}V1 ⩬[l, m] L2.ⓑ{I2}V2 → 0 < l →
+ L1 ⩬[⫰l, m] L2.
+#I1 #I2 #L1 #L2 #V1 #V2 #l #m #H #Hl elim (lreq_inv_succ1 … H) -H //
+#Z #Y #X #HL12 #H destruct //
+qed-.
+
+lemma lreq_inv_zero2: ∀I2,K2,L1,V2. L1 ⩬[0, 0] K2.ⓑ{I2}V2 →
+ ∃∃I1,K1,V1. K1 ⩬[0, 0] K2 & L1 = K1.ⓑ{I1}V1.
+#I2 #K2 #L1 #V2 #H elim (lreq_inv_zero1 … (lreq_sym … H)) -H
+/3 width=5 by lreq_sym, ex2_3_intro/
+qed-.
+
+lemma lreq_inv_pair2: ∀I,K2,L1,V,m. L1 ⩬[0, m] K2.ⓑ{I}V → 0 < m →
+ ∃∃K1. K1 ⩬[0, ⫰m] K2 & L1 = K1.ⓑ{I}V.
+#I #K2 #L1 #V #m #H #Hm elim (lreq_inv_pair1 … (lreq_sym … H)) -H
+/3 width=3 by lreq_sym, ex2_intro/
+qed-.
+
+lemma lreq_inv_succ2: ∀I2,K2,L1,V2,l,m. L1 ⩬[l, m] K2.ⓑ{I2}V2 → 0 < l →
+ ∃∃I1,K1,V1. K1 ⩬[⫰l, m] K2 & L1 = K1.ⓑ{I1}V1.
+#I2 #K2 #L1 #V2 #l #m #H #Hl elim (lreq_inv_succ1 … (lreq_sym … H)) -H
+/3 width=5 by lreq_sym, ex2_3_intro/
+qed-.
+
+(* Basic forward lemmas *****************************************************)
+
+lemma lreq_fwd_length: ∀L1,L2,l,m. L1 ⩬[l, m] L2 → |L1| = |L2|.
+#L1 #L2 #l #m #H elim H -L1 -L2 -l -m normalize //
+qed-.
+
+(* Advanced inversion lemmas ************************************************)
+
+fact lreq_inv_O_Y_aux: ∀L1,L2,l,m. L1 ⩬[l, m] L2 → l = 0 → m = ∞ → L1 = L2.
+#L1 #L2 #l #m #H elim H -L1 -L2 -l -m //
+[ #I1 #I2 #L1 #L2 #V1 #V2 #_ #_ #_ #H destruct
+| /4 width=1 by eq_f3, ysucc_inv_Y_dx/
+| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #_ #_ #H elim (ysucc_inv_O_dx … H)
+]
+qed-.
+
+lemma lreq_inv_O_Y: ∀L1,L2. L1 ⩬[0, ∞] L2 → L1 = L2.
+/2 width=5 by lreq_inv_O_Y_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 "ground_2/ynat/ynat_plus.ma".
+include "basic_2/grammar/lreq.ma".
+
+(* EQUIVALENCE FOR LOCAL ENVIRONMENTS ***************************************)
+
+(* Main properties **********************************************************)
+
+theorem lreq_trans: ∀l,m. Transitive … (lreq l m).
+#l #m #L1 #L2 #H elim H -L1 -L2 -l -m
+[ #l #m #X #H lapply (lreq_inv_atom1 … H) -H
+ #H destruct //
+| #I1 #I #L1 #L #V1 #V #_ #IHL1 #X #H elim (lreq_inv_zero1 … H) -H
+ #I2 #L2 #V2 #HL2 #H destruct /3 width=1 by lreq_zero/
+| #I #L1 #L #V #m #_ #IHL1 #X #H elim (lreq_inv_pair1 … H) -H //
+ #L2 #HL2 #H destruct /3 width=1 by lreq_pair/
+| #I1 #I #L1 #L #V1 #V #l #m #_ #IHL1 #X #H elim (lreq_inv_succ1 … H) -H //
+ #I2 #L2 #V2 #HL2 #H destruct /3 width=1 by lreq_succ/
+]
+qed-.
+
+theorem lreq_canc_sn: ∀l,m,L,L1,L2. L ⩬[l, m] L1 → L ⩬[l, m] L2 → L1 ⩬[l, m] L2.
+/3 width=3 by lreq_trans, lreq_sym/ qed-.
+
+theorem lreq_canc_dx: ∀l,m,L,L1,L2. L1 ⩬[l, m] L → L2 ⩬[l, m] L → L1 ⩬[l, m] L2.
+/3 width=3 by lreq_trans, lreq_sym/ qed-.
+
+theorem lreq_join: ∀L1,L2,l,i. L1 ⩬[l, i] L2 →
+ ∀m. L1 ⩬[i+l, m] L2 → L1 ⩬[l, i+m] L2.
+#L1 #L2 #l #i #H elim H -L1 -L2 -l -i //
+[ #I #L1 #L2 #V #m #_ #IHL12 #m #H
+ lapply (lreq_inv_succ … H ?) -H // >ypred_succ /3 width=1 by lreq_pair/
+| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #_ #IHL12 #m #H
+ lapply (lreq_inv_succ … H ?) -H // >yplus_succ2 >ypred_succ /3 width=1 by lreq_succ/
+]
+qed-.
(* Basic inversion lemmas ***************************************************)
-lemma destruct_tpair_tpair: ∀I1,I2,T1,T2,V1,V2. ②{I1}T1.V1 = ②{I2}T2.V2 →
- ∧∧T1 = T2 & I1 = I2 & V1 = V2.
+fact destruct_tatom_tatom_aux: ∀I1,I2. ⓪{I1} = ⓪{I2} → I1 = I2.
+#I1 #I2 #H destruct //
+qed-.
+
+fact destruct_tpair_tpair_aux: ∀I1,I2,T1,T2,V1,V2. ②{I1}T1.V1 = ②{I2}T2.V2 →
+ ∧∧T1 = T2 & I1 = I2 & V1 = V2.
#I1 #I2 #T1 #T2 #V1 #V2 #H destruct /2 width=1 by and3_intro/
qed-.
lemma discr_tpair_xy_x: ∀I,T,V. ②{I} V. T = V → ⊥.
#I #T #V elim V -V
[ #J #H destruct
-| #J #W #U #IHW #_ #H elim (destruct_tpair_tpair … H) -H /2 width=1 by/ (**) (* destruct lemma needed *)
+| #J #W #U #IHW #_ #H elim (destruct_tpair_tpair_aux … H) -H /2 width=1 by/ (**) (* destruct lemma needed *)
]
qed-.
lemma discr_tpair_xy_y: ∀I,V,T. ②{I} V. T = T → ⊥.
#I #V #T elim T -T
[ #J #H destruct
-| #J #W #U #_ #IHU #H elim (destruct_tpair_tpair … H) -H /2 width=1 by/ (**) (* destruct lemma needed *)
+| #J #W #U #_ #IHU #H elim (destruct_tpair_tpair_aux … H) -H /2 width=1 by/ (**) (* destruct lemma needed *)
]
qed-.
(* CONTEXT-SENSITIVE EXTENDED MULTIPLE SUBSTITUTION FOR TERMS ***************)
definition cpys: ynat → ynat → relation4 genv lenv term term ≝
- λd,e,G. LTC … (cpy d e G).
+ λl,m,G. LTC … (cpy l m G).
interpretation "context-sensitive extended multiple substritution (term)"
- 'PSubstStar G L T1 d e T2 = (cpys d e G L T1 T2).
+ 'PSubstStar G L T1 l m T2 = (cpys l m G L T1 T2).
(* Basic eliminators ********************************************************)
-lemma cpys_ind: ∀G,L,T1,d,e. ∀R:predicate term. R T1 →
- (∀T,T2. ⦃G, L⦄ ⊢ T1 ▶*[d, e] T → ⦃G, L⦄ ⊢ T ▶[d, e] T2 → R T → R T2) →
- ∀T2. ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → R T2.
-#G #L #T1 #d #e #R #HT1 #IHT1 #T2 #HT12
+lemma cpys_ind: ∀G,L,T1,l,m. ∀R:predicate term. R T1 →
+ (∀T,T2. ⦃G, L⦄ ⊢ T1 ▶*[l, m] T → ⦃G, L⦄ ⊢ T ▶[l, m] T2 → R T → R T2) →
+ ∀T2. ⦃G, L⦄ ⊢ T1 ▶*[l, m] T2 → R T2.
+#G #L #T1 #l #m #R #HT1 #IHT1 #T2 #HT12
@(TC_star_ind … HT1 IHT1 … HT12) //
qed-.
-lemma cpys_ind_dx: ∀G,L,T2,d,e. ∀R:predicate term. R T2 →
- (∀T1,T. ⦃G, L⦄ ⊢ T1 ▶[d, e] T → ⦃G, L⦄ ⊢ T ▶*[d, e] T2 → R T → R T1) →
- ∀T1. ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → R T1.
-#G #L #T2 #d #e #R #HT2 #IHT2 #T1 #HT12
+lemma cpys_ind_dx: ∀G,L,T2,l,m. ∀R:predicate term. R T2 →
+ (∀T1,T. ⦃G, L⦄ ⊢ T1 ▶[l, m] T → ⦃G, L⦄ ⊢ T ▶*[l, m] T2 → R T → R T1) →
+ ∀T1. ⦃G, L⦄ ⊢ T1 ▶*[l, m] T2 → R T1.
+#G #L #T2 #l #m #R #HT2 #IHT2 #T1 #HT12
@(TC_star_ind_dx … HT2 IHT2 … HT12) //
qed-.
(* Basic properties *********************************************************)
-lemma cpy_cpys: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2.
+lemma cpy_cpys: ∀G,L,T1,T2,l,m. ⦃G, L⦄ ⊢ T1 ▶[l, m] T2 → ⦃G, L⦄ ⊢ T1 ▶*[l, m] T2.
/2 width=1 by inj/ qed.
-lemma cpys_strap1: ∀G,L,T1,T,T2,d,e.
- ⦃G, L⦄ ⊢ T1 ▶*[d, e] T → ⦃G, L⦄ ⊢ T ▶[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2.
+lemma cpys_strap1: ∀G,L,T1,T,T2,l,m.
+ ⦃G, L⦄ ⊢ T1 ▶*[l, m] T → ⦃G, L⦄ ⊢ T ▶[l, m] T2 → ⦃G, L⦄ ⊢ T1 ▶*[l, m] T2.
normalize /2 width=3 by step/ qed-.
-lemma cpys_strap2: ∀G,L,T1,T,T2,d,e.
- ⦃G, L⦄ ⊢ T1 ▶[d, e] T → ⦃G, L⦄ ⊢ T ▶*[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2.
+lemma cpys_strap2: ∀G,L,T1,T,T2,l,m.
+ ⦃G, L⦄ ⊢ T1 ▶[l, m] T → ⦃G, L⦄ ⊢ T ▶*[l, m] T2 → ⦃G, L⦄ ⊢ T1 ▶*[l, m] T2.
normalize /2 width=3 by TC_strap/ qed-.
-lemma lsuby_cpys_trans: ∀G,d,e. lsub_trans … (cpys d e G) (lsuby d e).
+lemma lsuby_cpys_trans: ∀G,l,m. lsub_trans … (cpys l m G) (lsuby l m).
/3 width=5 by lsuby_cpy_trans, LTC_lsub_trans/
qed-.
-lemma cpys_refl: ∀G,L,d,e. reflexive … (cpys d e G L).
+lemma cpys_refl: ∀G,L,l,m. reflexive … (cpys l m G L).
/2 width=1 by cpy_cpys/ qed.
-lemma cpys_bind: ∀G,L,V1,V2,d,e. ⦃G, L⦄ ⊢ V1 ▶*[d, e] V2 →
- ∀I,T1,T2. ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ▶*[⫯d, e] T2 →
- ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ▶*[d, e] ⓑ{a,I}V2.T2.
-#G #L #V1 #V2 #d #e #HV12 @(cpys_ind … HV12) -V2
+lemma cpys_bind: ∀G,L,V1,V2,l,m. ⦃G, L⦄ ⊢ V1 ▶*[l, m] V2 →
+ ∀I,T1,T2. ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ▶*[⫯l, m] T2 →
+ ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ▶*[l, m] ⓑ{a,I}V2.T2.
+#G #L #V1 #V2 #l #m #HV12 @(cpys_ind … HV12) -V2
[ #I #T1 #T2 #HT12 @(cpys_ind … HT12) -T2 /3 width=5 by cpys_strap1, cpy_bind/
| /3 width=5 by cpys_strap1, cpy_bind/
]
qed.
-lemma cpys_flat: ∀G,L,V1,V2,d,e. ⦃G, L⦄ ⊢ V1 ▶*[d, e] V2 →
- ∀T1,T2. ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 →
- ∀I. ⦃G, L⦄ ⊢ ⓕ{I}V1.T1 ▶*[d, e] ⓕ{I}V2.T2.
-#G #L #V1 #V2 #d #e #HV12 @(cpys_ind … HV12) -V2
+lemma cpys_flat: ∀G,L,V1,V2,l,m. ⦃G, L⦄ ⊢ V1 ▶*[l, m] V2 →
+ ∀T1,T2. ⦃G, L⦄ ⊢ T1 ▶*[l, m] T2 →
+ ∀I. ⦃G, L⦄ ⊢ ⓕ{I}V1.T1 ▶*[l, m] ⓕ{I}V2.T2.
+#G #L #V1 #V2 #l #m #HV12 @(cpys_ind … HV12) -V2
[ #T1 #T2 #HT12 @(cpys_ind … HT12) -T2 /3 width=5 by cpys_strap1, cpy_flat/
| /3 width=5 by cpys_strap1, cpy_flat/
qed.
-lemma cpys_weak: ∀G,L,T1,T2,d1,e1. ⦃G, L⦄ ⊢ T1 ▶*[d1, e1] T2 →
- ∀d2,e2. d2 ≤ d1 → d1 + e1 ≤ d2 + e2 →
- ⦃G, L⦄ ⊢ T1 ▶*[d2, e2] T2.
-#G #L #T1 #T2 #d1 #e1 #H #d1 #d2 #Hd21 #Hde12 @(cpys_ind … H) -T2
+lemma cpys_weak: ∀G,L,T1,T2,l1,m1. ⦃G, L⦄ ⊢ T1 ▶*[l1, m1] T2 →
+ ∀l2,m2. l2 ≤ l1 → l1 + m1 ≤ l2 + m2 →
+ ⦃G, L⦄ ⊢ T1 ▶*[l2, m2] T2.
+#G #L #T1 #T2 #l1 #m1 #H #l1 #l2 #Hl21 #Hlm12 @(cpys_ind … H) -T2
/3 width=7 by cpys_strap1, cpy_weak/
qed-.
-lemma cpys_weak_top: ∀G,L,T1,T2,d,e.
- ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶*[d, |L| - d] T2.
-#G #L #T1 #T2 #d #e #H @(cpys_ind … H) -T2
+lemma cpys_weak_top: ∀G,L,T1,T2,l,m.
+ ⦃G, L⦄ ⊢ T1 ▶*[l, m] T2 → ⦃G, L⦄ ⊢ T1 ▶*[l, |L| - l] T2.
+#G #L #T1 #T2 #l #m #H @(cpys_ind … H) -T2
/3 width=4 by cpys_strap1, cpy_weak_top/
qed-.
-lemma cpys_weak_full: ∀G,L,T1,T2,d,e.
- ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶*[0, |L|] T2.
-#G #L #T1 #T2 #d #e #H @(cpys_ind … H) -T2
+lemma cpys_weak_full: ∀G,L,T1,T2,l,m.
+ ⦃G, L⦄ ⊢ T1 ▶*[l, m] T2 → ⦃G, L⦄ ⊢ T1 ▶*[0, |L|] T2.
+#G #L #T1 #T2 #l #m #H @(cpys_ind … H) -T2
/3 width=5 by cpys_strap1, cpy_weak_full/
qed-.
(* Basic forward lemmas *****************************************************)
-lemma cpys_fwd_up: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶*[dt, et] U2 →
- ∀T1,d,e. ⬆[d, e] T1 ≡ U1 →
- d ≤ dt → d + e ≤ dt + et →
- ∃∃T2. ⦃G, L⦄ ⊢ U1 ▶*[d+e, dt+et-(d+e)] U2 & ⬆[d, e] T2 ≡ U2.
-#G #L #U1 #U2 #dt #et #H #T1 #d #e #HTU1 #Hddt #Hdedet @(cpys_ind … H) -U2
+lemma cpys_fwd_up: ∀G,L,U1,U2,lt,mt. ⦃G, L⦄ ⊢ U1 ▶*[lt, mt] U2 →
+ ∀T1,l,m. ⬆[l, m] T1 ≡ U1 →
+ l ≤ lt → l + m ≤ lt + mt →
+ ∃∃T2. ⦃G, L⦄ ⊢ U1 ▶*[l+m, lt+mt-(l+m)] U2 & ⬆[l, m] T2 ≡ U2.
+#G #L #U1 #U2 #lt #mt #H #T1 #l #m #HTU1 #Hllt #Hlmlmt @(cpys_ind … H) -U2
[ /2 width=3 by ex2_intro/
| -HTU1 #U #U2 #_ #HU2 * #T #HU1 #HTU
elim (cpy_fwd_up … HU2 … HTU) -HU2 -HTU /3 width=3 by cpys_strap1, ex2_intro/
]
qed-.
-lemma cpys_fwd_tw: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → ♯{T1} ≤ ♯{T2}.
-#G #L #T1 #T2 #d #e #H @(cpys_ind … H) -T2 //
+lemma cpys_fwd_tw: ∀G,L,T1,T2,l,m. ⦃G, L⦄ ⊢ T1 ▶*[l, m] T2 → ♯{T1} ≤ ♯{T2}.
+#G #L #T1 #T2 #l #m #H @(cpys_ind … H) -T2 //
#T #T2 #_ #HT2 #IHT1 lapply (cpy_fwd_tw … HT2) -HT2
/2 width=3 by transitive_le/
qed-.
(* Basic inversion lemmas ***************************************************)
(* Note: this can be derived from cpys_inv_atom1 *)
-lemma cpys_inv_sort1: ∀G,L,T2,k,d,e. ⦃G, L⦄ ⊢ ⋆k ▶*[d, e] T2 → T2 = ⋆k.
-#G #L #T2 #k #d #e #H @(cpys_ind … H) -T2 //
+lemma cpys_inv_sort1: ∀G,L,T2,k,l,m. ⦃G, L⦄ ⊢ ⋆k ▶*[l, m] T2 → T2 = ⋆k.
+#G #L #T2 #k #l #m #H @(cpys_ind … H) -T2 //
#T #T2 #_ #HT2 #IHT1 destruct
>(cpy_inv_sort1 … HT2) -HT2 //
qed-.
(* Note: this can be derived from cpys_inv_atom1 *)
-lemma cpys_inv_gref1: ∀G,L,T2,p,d,e. ⦃G, L⦄ ⊢ §p ▶*[d, e] T2 → T2 = §p.
-#G #L #T2 #p #d #e #H @(cpys_ind … H) -T2 //
+lemma cpys_inv_gref1: ∀G,L,T2,p,l,m. ⦃G, L⦄ ⊢ §p ▶*[l, m] T2 → T2 = §p.
+#G #L #T2 #p #l #m #H @(cpys_ind … H) -T2 //
#T #T2 #_ #HT2 #IHT1 destruct
>(cpy_inv_gref1 … HT2) -HT2 //
qed-.
-lemma cpys_inv_bind1: ∀a,I,G,L,V1,T1,U2,d,e. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ▶*[d, e] U2 →
- ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶*[d, e] V2 &
- ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ▶*[⫯d, e] T2 &
+lemma cpys_inv_bind1: ∀a,I,G,L,V1,T1,U2,l,m. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ▶*[l, m] U2 →
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶*[l, m] V2 &
+ ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ▶*[⫯l, m] T2 &
U2 = ⓑ{a,I}V2.T2.
-#a #I #G #L #V1 #T1 #U2 #d #e #H @(cpys_ind … H) -U2
+#a #I #G #L #V1 #T1 #U2 #l #m #H @(cpys_ind … H) -U2
[ /2 width=5 by ex3_2_intro/
| #U #U2 #_ #HU2 * #V #T #HV1 #HT1 #H destruct
elim (cpy_inv_bind1 … HU2) -HU2 #V2 #T2 #HV2 #HT2 #H
]
qed-.
-lemma cpys_inv_flat1: ∀I,G,L,V1,T1,U2,d,e. ⦃G, L⦄ ⊢ ⓕ{I}V1.T1 ▶*[d, e] U2 →
- ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶*[d, e] V2 & ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 &
+lemma cpys_inv_flat1: ∀I,G,L,V1,T1,U2,l,m. ⦃G, L⦄ ⊢ ⓕ{I}V1.T1 ▶*[l, m] U2 →
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶*[l, m] V2 & ⦃G, L⦄ ⊢ T1 ▶*[l, m] T2 &
U2 = ⓕ{I}V2.T2.
-#I #G #L #V1 #T1 #U2 #d #e #H @(cpys_ind … H) -U2
+#I #G #L #V1 #T1 #U2 #l #m #H @(cpys_ind … H) -U2
[ /2 width=5 by ex3_2_intro/
| #U #U2 #_ #HU2 * #V #T #HV1 #HT1 #H destruct
elim (cpy_inv_flat1 … HU2) -HU2
]
qed-.
-lemma cpys_inv_refl_O2: ∀G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 ▶*[d, 0] T2 → T1 = T2.
-#G #L #T1 #T2 #d #H @(cpys_ind … H) -T2 //
+lemma cpys_inv_refl_O2: ∀G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 ▶*[l, 0] T2 → T1 = T2.
+#G #L #T1 #T2 #l #H @(cpys_ind … H) -T2 //
#T #T2 #_ #HT2 #IHT1 <(cpy_inv_refl_O2 … HT2) -HT2 //
qed-.
-lemma cpys_inv_lift1_eq: ∀G,L,U1,U2. ∀d,e:nat.
- ⦃G, L⦄ ⊢ U1 ▶*[d, e] U2 → ∀T1. ⬆[d, e] T1 ≡ U1 → U1 = U2.
-#G #L #U1 #U2 #d #e #H #T1 #HTU1 @(cpys_ind … H) -U2
+lemma cpys_inv_lift1_eq: ∀G,L,U1,U2. ∀l,m:nat.
+ ⦃G, L⦄ ⊢ U1 ▶*[l, m] U2 → ∀T1. ⬆[l, m] T1 ≡ U1 → U1 = U2.
+#G #L #U1 #U2 #l #m #H #T1 #HTU1 @(cpys_ind … H) -U2
/2 width=7 by cpy_inv_lift1_eq/
qed-.
(* alternative definition of cpys *)
inductive cpysa: ynat → ynat → relation4 genv lenv term term ≝
-| cpysa_atom : ∀I,G,L,d,e. cpysa d e G L (⓪{I}) (⓪{I})
-| cpysa_subst: ∀I,G,L,K,V1,V2,W2,i,d,e. d ≤ yinj i → i < d+e →
- ⬇[i] L ≡ K.ⓑ{I}V1 → cpysa 0 (⫰(d+e-i)) G K V1 V2 →
- ⬆[0, i+1] V2 ≡ W2 → cpysa d e G L (#i) W2
-| cpysa_bind : ∀a,I,G,L,V1,V2,T1,T2,d,e.
- cpysa d e G L V1 V2 → cpysa (⫯d) e G (L.ⓑ{I}V1) T1 T2 →
- cpysa d e G L (ⓑ{a,I}V1.T1) (ⓑ{a,I}V2.T2)
-| cpysa_flat : ∀I,G,L,V1,V2,T1,T2,d,e.
- cpysa d e G L V1 V2 → cpysa d e G L T1 T2 →
- cpysa d e G L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2)
+| cpysa_atom : ∀I,G,L,l,m. cpysa l m G L (⓪{I}) (⓪{I})
+| cpysa_subst: ∀I,G,L,K,V1,V2,W2,i,l,m. l ≤ yinj i → i < l+m →
+ ⬇[i] L ≡ K.ⓑ{I}V1 → cpysa 0 (⫰(l+m-i)) G K V1 V2 →
+ ⬆[0, i+1] V2 ≡ W2 → cpysa l m G L (#i) W2
+| cpysa_bind : ∀a,I,G,L,V1,V2,T1,T2,l,m.
+ cpysa l m G L V1 V2 → cpysa (⫯l) m G (L.ⓑ{I}V1) T1 T2 →
+ cpysa l m G L (ⓑ{a,I}V1.T1) (ⓑ{a,I}V2.T2)
+| cpysa_flat : ∀I,G,L,V1,V2,T1,T2,l,m.
+ cpysa l m G L V1 V2 → cpysa l m G L T1 T2 →
+ cpysa l m G L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2)
.
interpretation
"context-sensitive extended multiple substritution (term) alternative"
- 'PSubstStarAlt G L T1 d e T2 = (cpysa d e G L T1 T2).
+ 'PSubstStarAlt G L T1 l m T2 = (cpysa l m G L T1 T2).
(* Basic properties *********************************************************)
-lemma lsuby_cpysa_trans: ∀G,d,e. lsub_trans … (cpysa d e G) (lsuby d e).
-#G #d #e #L1 #T1 #T2 #H elim H -G -L1 -T1 -T2 -d -e
+lemma lsuby_cpysa_trans: ∀G,l,m. lsub_trans … (cpysa l m G) (lsuby l m).
+#G #l #m #L1 #T1 #T2 #H elim H -G -L1 -T1 -T2 -l -m
[ //
-| #I #G #L1 #K1 #V1 #V2 #W2 #i #d #e #Hdi #Hide #HLK1 #_ #HVW2 #IHV12 #L2 #HL12
+| #I #G #L1 #K1 #V1 #V2 #W2 #i #l #m #Hli #Hilm #HLK1 #_ #HVW2 #IHV12 #L2 #HL12
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/
]
qed-.
-lemma cpysa_refl: ∀G,T,L,d,e. ⦃G, L⦄ ⊢ T ▶▶*[d, e] T.
+lemma cpysa_refl: ∀G,T,L,l,m. ⦃G, L⦄ ⊢ T ▶▶*[l, m] T.
#G #T elim T -T //
#I elim I -I /2 width=1 by cpysa_bind, cpysa_flat/
qed.
-lemma cpysa_cpy_trans: ∀G,L,T1,T,d,e. ⦃G, L⦄ ⊢ T1 ▶▶*[d, e] T →
- ∀T2. ⦃G, L⦄ ⊢ T ▶[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶▶*[d, e] T2.
-#G #L #T1 #T #d #e #H elim H -G -L -T1 -T -d -e
-[ #I #G #L #d #e #X #H
+lemma cpysa_cpy_trans: ∀G,L,T1,T,l,m. ⦃G, L⦄ ⊢ T1 ▶▶*[l, m] T →
+ ∀T2. ⦃G, L⦄ ⊢ T ▶[l, m] T2 → ⦃G, L⦄ ⊢ T1 ▶▶*[l, m] T2.
+#G #L #T1 #T #l #m #H elim H -G -L -T1 -T -l -m
+[ #I #G #L #l #m #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
+| #I #G #L #K #V1 #V2 #W2 #i #l #m #Hli #Hilm #HLK #_ #HVW2 #IHV12 #T2 #H
lapply (drop_fwd_drop2 … HLK) #H0LK
- lapply (cpy_weak … H 0 (d+e) ? ?) -H // #H
+ lapply (cpy_weak … H 0 (l+m) ? ?) -H // #H
elim (cpy_inv_lift1_be … H … H0LK … HVW2) -H -H0LK -HVW2
/3 width=7 by cpysa_subst, ylt_fwd_le_succ/
-| #a #I #G #L #V1 #V #T1 #T #d #e #_ #_ #IHV1 #IHT1 #X #H
+| #a #I #G #L #V1 #V #T1 #T #l #m #_ #_ #IHV1 #IHT1 #X #H
elim (cpy_inv_bind1 … H) -H #V2 #T2 #HV2 #HT2 #H destruct
/5 width=5 by cpysa_bind, lsuby_cpy_trans, lsuby_succ/
-| #I #G #L #V1 #V #T1 #T #d #e #_ #_ #IHV1 #IHT1 #X #H
+| #I #G #L #V1 #V #T1 #T #l #m #_ #_ #IHV1 #IHT1 #X #H
elim (cpy_inv_flat1 … H) -H #V2 #T2 #HV2 #HT2 #H destruct /3 width=1 by cpysa_flat/
]
qed-.
-lemma cpys_cpysa: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶▶*[d, e] T2.
+lemma cpys_cpysa: ∀G,L,T1,T2,l,m. ⦃G, L⦄ ⊢ T1 ▶*[l, m] T2 → ⦃G, L⦄ ⊢ T1 ▶▶*[l, m] T2.
/3 width=8 by cpysa_cpy_trans, cpys_ind/ qed.
(* Basic inversion lemmas ***************************************************)
-lemma cpysa_inv_cpys: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶▶*[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2.
-#G #L #T1 #T2 #d #e #H elim H -G -L -T1 -T2 -d -e
+lemma cpysa_inv_cpys: ∀G,L,T1,T2,l,m. ⦃G, L⦄ ⊢ T1 ▶▶*[l, m] T2 → ⦃G, L⦄ ⊢ T1 ▶*[l, m] T2.
+#G #L #T1 #T2 #l #m #H elim H -G -L -T1 -T2 -l -m
/2 width=7 by cpys_subst, cpys_flat, cpys_bind, cpy_cpys/
qed-.
(* Advanced eliminators *****************************************************)
lemma cpys_ind_alt: ∀R:ynat→ynat→relation4 genv lenv term term.
- (∀I,G,L,d,e. R d e G L (⓪{I}) (⓪{I})) →
- (∀I,G,L,K,V1,V2,W2,i,d,e. d ≤ yinj i → i < d + e →
- ⬇[i] L ≡ K.ⓑ{I}V1 → ⦃G, K⦄ ⊢ V1 ▶*[O, ⫰(d+e-i)] V2 →
- ⬆[O, i+1] V2 ≡ W2 → R O (⫰(d+e-i)) G K V1 V2 → R d e G L (#i) W2
+ (∀I,G,L,l,m. R l m G L (⓪{I}) (⓪{I})) →
+ (∀I,G,L,K,V1,V2,W2,i,l,m. l ≤ yinj i → i < l + m →
+ ⬇[i] L ≡ K.ⓑ{I}V1 → ⦃G, K⦄ ⊢ V1 ▶*[O, ⫰(l+m-i)] V2 →
+ ⬆[O, i+1] V2 ≡ W2 → R O (⫰(l+m-i)) G K V1 V2 → R l m G L (#i) W2
) →
- (∀a,I,G,L,V1,V2,T1,T2,d,e. ⦃G, L⦄ ⊢ V1 ▶*[d, e] V2 →
- ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ▶*[⫯d, e] T2 → R d e G L V1 V2 →
- R (⫯d) e G (L.ⓑ{I}V1) T1 T2 → R d e G L (ⓑ{a,I}V1.T1) (ⓑ{a,I}V2.T2)
+ (∀a,I,G,L,V1,V2,T1,T2,l,m. ⦃G, L⦄ ⊢ V1 ▶*[l, m] V2 →
+ ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ▶*[⫯l, m] T2 → R l m G L V1 V2 →
+ R (⫯l) m G (L.ⓑ{I}V1) T1 T2 → R l m G L (ⓑ{a,I}V1.T1) (ⓑ{a,I}V2.T2)
) →
- (∀I,G,L,V1,V2,T1,T2,d,e. ⦃G, L⦄ ⊢ V1 ▶*[d, e] V2 →
- ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → R d e G L V1 V2 →
- R d e G L T1 T2 → R d e G L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2)
+ (∀I,G,L,V1,V2,T1,T2,l,m. ⦃G, L⦄ ⊢ V1 ▶*[l, m] V2 →
+ ⦃G, L⦄ ⊢ T1 ▶*[l, m] T2 → R l m G L V1 V2 →
+ R l m G L T1 T2 → R l m G L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2)
) →
- ∀d,e,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → R d e G L T1 T2.
-#R #H1 #H2 #H3 #H4 #d #e #G #L #T1 #T2 #H elim (cpys_cpysa … H) -G -L -T1 -T2 -d -e
+ ∀l,m,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ▶*[l, m] T2 → R l m G L T1 T2.
+#R #H1 #H2 #H3 #H4 #l #m #G #L #T1 #T2 #H elim (cpys_cpysa … H) -G -L -T1 -T2 -l -m
/3 width=8 by cpysa_inv_cpys/
qed-.
(* Advanced inversion lemmas ************************************************)
-lemma cpys_inv_SO2: ∀G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 ▶*[d, 1] T2 → ⦃G, L⦄ ⊢ T1 ▶[d, 1] T2.
-#G #L #T1 #T2 #d #H @(cpys_ind … H) -T2 /2 width=3 by cpy_trans_ge/
+lemma cpys_inv_SO2: ∀G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 ▶*[l, 1] T2 → ⦃G, L⦄ ⊢ T1 ▶[l, 1] T2.
+#G #L #T1 #T2 #l #H @(cpys_ind … H) -T2 /2 width=3 by cpy_trans_ge/
qed-.
(* Advanced properties ******************************************************)
-lemma cpys_strip_eq: ∀G,L,T0,T1,d1,e1. ⦃G, L⦄ ⊢ T0 ▶*[d1, e1] T1 →
- ∀T2,d2,e2. ⦃G, L⦄ ⊢ T0 ▶[d2, e2] T2 →
- ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[d2, e2] T & ⦃G, L⦄ ⊢ T2 ▶*[d1, e1] T.
+lemma cpys_strip_eq: ∀G,L,T0,T1,l1,m1. ⦃G, L⦄ ⊢ T0 ▶*[l1, m1] T1 →
+ ∀T2,l2,m2. ⦃G, L⦄ ⊢ T0 ▶[l2, m2] T2 →
+ ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[l2, m2] T & ⦃G, L⦄ ⊢ T2 ▶*[l1, m1] T.
normalize /3 width=3 by cpy_conf_eq, TC_strip1/ qed-.
-lemma cpys_strip_neq: ∀G,L1,T0,T1,d1,e1. ⦃G, L1⦄ ⊢ T0 ▶*[d1, e1] T1 →
- ∀L2,T2,d2,e2. ⦃G, L2⦄ ⊢ T0 ▶[d2, e2] T2 →
- (d1 + e1 ≤ d2 ∨ d2 + e2 ≤ d1) →
- ∃∃T. ⦃G, L2⦄ ⊢ T1 ▶[d2, e2] T & ⦃G, L1⦄ ⊢ T2 ▶*[d1, e1] T.
+lemma cpys_strip_neq: ∀G,L1,T0,T1,l1,m1. ⦃G, L1⦄ ⊢ T0 ▶*[l1, m1] T1 →
+ ∀L2,T2,l2,m2. ⦃G, L2⦄ ⊢ T0 ▶[l2, m2] T2 →
+ (l1 + m1 ≤ l2 ∨ l2 + m2 ≤ l1) →
+ ∃∃T. ⦃G, L2⦄ ⊢ T1 ▶[l2, m2] T & ⦃G, L1⦄ ⊢ T2 ▶*[l1, m1] T.
normalize /3 width=3 by cpy_conf_neq, TC_strip1/ qed-.
-lemma cpys_strap1_down: ∀G,L,T1,T0,d1,e1. ⦃G, L⦄ ⊢ T1 ▶*[d1, e1] T0 →
- ∀T2,d2,e2. ⦃G, L⦄ ⊢ T0 ▶[d2, e2] T2 → d2 + e2 ≤ d1 →
- ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[d2, e2] T & ⦃G, L⦄ ⊢ T ▶*[d1, e1] T2.
+lemma cpys_strap1_down: ∀G,L,T1,T0,l1,m1. ⦃G, L⦄ ⊢ T1 ▶*[l1, m1] T0 →
+ ∀T2,l2,m2. ⦃G, L⦄ ⊢ T0 ▶[l2, m2] T2 → l2 + m2 ≤ l1 →
+ ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[l2, m2] T & ⦃G, L⦄ ⊢ T ▶*[l1, m1] T2.
normalize /3 width=3 by cpy_trans_down, TC_strap1/ qed.
-lemma cpys_strap2_down: ∀G,L,T1,T0,d1,e1. ⦃G, L⦄ ⊢ T1 ▶[d1, e1] T0 →
- ∀T2,d2,e2. ⦃G, L⦄ ⊢ T0 ▶*[d2, e2] T2 → d2 + e2 ≤ d1 →
- ∃∃T. ⦃G, L⦄ ⊢ T1 ▶*[d2, e2] T & ⦃G, L⦄ ⊢ T ▶[d1, e1] T2.
+lemma cpys_strap2_down: ∀G,L,T1,T0,l1,m1. ⦃G, L⦄ ⊢ T1 ▶[l1, m1] T0 →
+ ∀T2,l2,m2. ⦃G, L⦄ ⊢ T0 ▶*[l2, m2] T2 → l2 + m2 ≤ l1 →
+ ∃∃T. ⦃G, L⦄ ⊢ T1 ▶*[l2, m2] T & ⦃G, L⦄ ⊢ T ▶[l1, m1] T2.
normalize /3 width=3 by cpy_trans_down, TC_strap2/ qed-.
-lemma cpys_split_up: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 →
- ∀i. d ≤ i → i ≤ d + e →
- ∃∃T. ⦃G, L⦄ ⊢ T1 ▶*[d, i - d] T & ⦃G, L⦄ ⊢ T ▶*[i, d + e - i] T2.
-#G #L #T1 #T2 #d #e #H #i #Hdi #Hide @(cpys_ind … H) -T2
+lemma cpys_split_up: ∀G,L,T1,T2,l,m. ⦃G, L⦄ ⊢ T1 ▶*[l, m] T2 →
+ ∀i. l ≤ i → i ≤ l + m →
+ ∃∃T. ⦃G, L⦄ ⊢ T1 ▶*[l, i - l] T & ⦃G, L⦄ ⊢ T ▶*[i, l + m - i] T2.
+#G #L #T1 #T2 #l #m #H #i #Hli #Hilm @(cpys_ind … H) -T2
[ /2 width=3 by ex2_intro/
| #T #T2 #_ #HT12 * #T3 #HT13 #HT3
- elim (cpy_split_up … HT12 … Hide) -HT12 -Hide #T0 #HT0 #HT02
+ elim (cpy_split_up … HT12 … Hilm) -HT12 -Hilm #T0 #HT0 #HT02
elim (cpys_strap1_down … HT3 … HT0) -T /3 width=5 by cpys_strap1, ex2_intro/
>ymax_pre_sn_comm //
]
qed-.
-lemma cpys_inv_lift1_up: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶*[dt, et] U2 →
- ∀K,s,d,e. ⬇[s, d, e] L ≡ K → ∀T1. ⬆[d, e] T1 ≡ U1 →
- d ≤ dt → dt ≤ yinj d + e → yinj d + e ≤ dt + et →
- ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶*[d, dt + et - (yinj d + e)] T2 &
- ⬆[d, e] T2 ≡ U2.
-#G #L #U1 #U2 #dt #et #HU12 #K #s #d #e #HLK #T1 #HTU1 #Hddt #Hdtde #Hdedet
-elim (cpys_split_up … HU12 (d + e)) -HU12 // -Hdedet #U #HU1 #HU2
-lapply (cpys_weak … HU1 d e ? ?) -HU1 // [ >ymax_pre_sn_comm // ] -Hddt -Hdtde #HU1
+lemma cpys_inv_lift1_up: ∀G,L,U1,U2,lt,mt. ⦃G, L⦄ ⊢ U1 ▶*[lt, mt] U2 →
+ ∀K,s,l,m. ⬇[s, l, m] L ≡ K → ∀T1. ⬆[l, m] T1 ≡ U1 →
+ l ≤ lt → lt ≤ yinj l + m → yinj l + m ≤ lt + mt →
+ ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶*[l, lt + mt - (yinj l + m)] T2 &
+ ⬆[l, m] T2 ≡ U2.
+#G #L #U1 #U2 #lt #mt #HU12 #K #s #l #m #HLK #T1 #HTU1 #Hllt #Hltlm #Hlmlmt
+elim (cpys_split_up … HU12 (l + m)) -HU12 // -Hlmlmt #U #HU1 #HU2
+lapply (cpys_weak … HU1 l m ? ?) -HU1 // [ >ymax_pre_sn_comm // ] -Hllt -Hltlm #HU1
lapply (cpys_inv_lift1_eq … HU1 … HTU1) -HU1 #HU1 destruct
elim (cpys_inv_lift1_ge … HU2 … HLK … HTU1) -HU2 -HLK -HTU1 //
>yplus_minus_inj /2 width=3 by ex2_intro/
(* Main properties **********************************************************)
-theorem cpys_conf_eq: ∀G,L,T0,T1,d1,e1. ⦃G, L⦄ ⊢ T0 ▶*[d1, e1] T1 →
- ∀T2,d2,e2. ⦃G, L⦄ ⊢ T0 ▶*[d2, e2] T2 →
- ∃∃T. ⦃G, L⦄ ⊢ T1 ▶*[d2, e2] T & ⦃G, L⦄ ⊢ T2 ▶*[d1, e1] T.
+theorem cpys_conf_eq: ∀G,L,T0,T1,l1,m1. ⦃G, L⦄ ⊢ T0 ▶*[l1, m1] T1 →
+ ∀T2,l2,m2. ⦃G, L⦄ ⊢ T0 ▶*[l2, m2] T2 →
+ ∃∃T. ⦃G, L⦄ ⊢ T1 ▶*[l2, m2] T & ⦃G, L⦄ ⊢ T2 ▶*[l1, m1] T.
normalize /3 width=3 by cpy_conf_eq, TC_confluent2/ qed-.
-theorem cpys_conf_neq: ∀G,L1,T0,T1,d1,e1. ⦃G, L1⦄ ⊢ T0 ▶*[d1, e1] T1 →
- ∀L2,T2,d2,e2. ⦃G, L2⦄ ⊢ T0 ▶*[d2, e2] T2 →
- (d1 + e1 ≤ d2 ∨ d2 + e2 ≤ d1) →
- ∃∃T. ⦃G, L2⦄ ⊢ T1 ▶*[d2, e2] T & ⦃G, L1⦄ ⊢ T2 ▶*[d1, e1] T.
+theorem cpys_conf_neq: ∀G,L1,T0,T1,l1,m1. ⦃G, L1⦄ ⊢ T0 ▶*[l1, m1] T1 →
+ ∀L2,T2,l2,m2. ⦃G, L2⦄ ⊢ T0 ▶*[l2, m2] T2 →
+ (l1 + m1 ≤ l2 ∨ l2 + m2 ≤ l1) →
+ ∃∃T. ⦃G, L2⦄ ⊢ T1 ▶*[l2, m2] T & ⦃G, L1⦄ ⊢ T2 ▶*[l1, m1] T.
normalize /3 width=3 by cpy_conf_neq, TC_confluent2/ qed-.
-theorem cpys_trans_eq: ∀G,L,T1,T,T2,d,e.
- ⦃G, L⦄ ⊢ T1 ▶*[d, e] T → ⦃G, L⦄ ⊢ T ▶*[d, e] T2 →
- ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2.
+theorem cpys_trans_eq: ∀G,L,T1,T,T2,l,m.
+ ⦃G, L⦄ ⊢ T1 ▶*[l, m] T → ⦃G, L⦄ ⊢ T ▶*[l, m] T2 →
+ ⦃G, L⦄ ⊢ T1 ▶*[l, m] T2.
normalize /2 width=3 by trans_TC/ qed-.
-theorem cpys_trans_down: ∀G,L,T1,T0,d1,e1. ⦃G, L⦄ ⊢ T1 ▶*[d1, e1] T0 →
- ∀T2,d2,e2. ⦃G, L⦄ ⊢ T0 ▶*[d2, e2] T2 → d2 + e2 ≤ d1 →
- ∃∃T. ⦃G, L⦄ ⊢ T1 ▶*[d2, e2] T & ⦃G, L⦄ ⊢ T ▶*[d1, e1] T2.
+theorem cpys_trans_down: ∀G,L,T1,T0,l1,m1. ⦃G, L⦄ ⊢ T1 ▶*[l1, m1] T0 →
+ ∀T2,l2,m2. ⦃G, L⦄ ⊢ T0 ▶*[l2, m2] T2 → l2 + m2 ≤ l1 →
+ ∃∃T. ⦃G, L⦄ ⊢ T1 ▶*[l2, m2] T & ⦃G, L⦄ ⊢ T ▶*[l1, m1] T2.
normalize /3 width=3 by cpy_trans_down, TC_transitive2/ qed-.
-theorem cpys_antisym_eq: ∀G,L1,T1,T2,d,e. ⦃G, L1⦄ ⊢ T1 ▶*[d, e] T2 →
- ∀L2. ⦃G, L2⦄ ⊢ T2 ▶*[d, e] T1 → T1 = T2.
-#G #L1 #T1 #T2 #d #e #H @(cpys_ind_alt … H) -G -L1 -T1 -T2 //
-[ #I1 #G #L1 #K1 #V1 #V2 #W2 #i #d #e #Hdi #Hide #_ #_ #HVW2 #_ #L2 #HW2
- elim (lt_or_ge (|L2|) (i+1)) #Hi [ -Hdi -Hide | ]
+theorem cpys_antisym_eq: ∀G,L1,T1,T2,l,m. ⦃G, L1⦄ ⊢ T1 ▶*[l, m] T2 →
+ ∀L2. ⦃G, L2⦄ ⊢ T2 ▶*[l, m] T1 → T1 = T2.
+#G #L1 #T1 #T2 #l #m #H @(cpys_ind_alt … H) -G -L1 -T1 -T2 //
+[ #I1 #G #L1 #K1 #V1 #V2 #W2 #i #l #m #Hli #Hilm #_ #_ #HVW2 #_ #L2 #HW2
+ elim (lt_or_ge (|L2|) (i+1)) #Hi [ -Hli -Hilm | ]
[ lapply (cpys_weak_full … HW2) -HW2 #HW2
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 (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
+ /2 width=1 by ylt_fwd_le_succ, yle_succ_dx/ -Hli -Hilm
#X #_ #H elim (lift_inv_lref2_be … H) -H //
]
-| #a #I #G #L1 #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #L2 #H elim (cpys_inv_bind1 … H) -H
+| #a #I #G #L1 #V1 #V2 #T1 #T2 #l #m #_ #_ #IHV12 #IHT12 #L2 #H elim (cpys_inv_bind1 … H) -H
#V #T #HV2 #HT2 #H destruct
lapply (IHV12 … HV2) #H destruct -IHV12 -HV2 /3 width=2 by eq_f2/
-| #I #G #L1 #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #L2 #H elim (cpys_inv_flat1 … H) -H
+| #I #G #L1 #V1 #V2 #T1 #T2 #l #m #_ #_ #IHV12 #IHT12 #L2 #H elim (cpys_inv_flat1 … H) -H
#V #T #HV2 #HT2 #H destruct /3 width=2 by eq_f2/
]
qed-.
(* Advanced properties ******************************************************)
-lemma cpys_subst: ∀I,G,L,K,V,U1,i,d,e.
- d ≤ yinj i → i < d + e →
- ⬇[i] L ≡ K.ⓑ{I}V → ⦃G, K⦄ ⊢ V ▶*[0, ⫰(d+e-i)] U1 →
- ∀U2. ⬆[0, i+1] U1 ≡ U2 → ⦃G, L⦄ ⊢ #i ▶*[d, e] U2.
-#I #G #L #K #V #U1 #i #d #e #Hdi #Hide #HLK #H @(cpys_ind … H) -U1
+lemma cpys_subst: ∀I,G,L,K,V,U1,i,l,m.
+ l ≤ yinj i → i < l + m →
+ ⬇[i] L ≡ K.ⓑ{I}V → ⦃G, K⦄ ⊢ V ▶*[0, ⫰(l+m-i)] U1 →
+ ∀U2. ⬆[0, i+1] U1 ≡ U2 → ⦃G, L⦄ ⊢ #i ▶*[l, m] U2.
+#I #G #L #K #V #U1 #i #l #m #Hli #Hilm #HLK #H @(cpys_ind … H) -U1
[ /3 width=5 by cpy_cpys, cpy_subst/
| #U #U1 #_ #HU1 #IHU #U2 #HU12
elim (lift_total U 0 (i+1)) #U0 #HU0
lapply (IHU … HU0) -IHU #H
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
+ lapply (cpy_weak … HU02 l m ? ?) -HU02
[2,3: /2 width=3 by cpys_strap1, yle_succ_dx/ ]
>yplus_O1 <yplus_inj >ymax_pre_sn_comm /2 width=1 by ylt_fwd_le_succ/
]
qed.
-lemma cpys_subst_Y2: ∀I,G,L,K,V,U1,i,d.
- d ≤ yinj i →
+lemma cpys_subst_Y2: ∀I,G,L,K,V,U1,i,l.
+ l ≤ yinj i →
⬇[i] L ≡ K.ⓑ{I}V → ⦃G, K⦄ ⊢ V ▶*[0, ∞] U1 →
- ∀U2. ⬆[0, i+1] U1 ≡ U2 → ⦃G, L⦄ ⊢ #i ▶*[d, ∞] U2.
-#I #G #L #K #V #U1 #i #d #Hdi #HLK #HVU1 #U2 #HU12
+ ∀U2. ⬆[0, i+1] U1 ≡ U2 → ⦃G, L⦄ ⊢ #i ▶*[l, ∞] U2.
+#I #G #L #K #V #U1 #i #l #Hli #HLK #HVU1 #U2 #HU12
@(cpys_subst … HLK … HU12) >yminus_Y_inj //
qed.
(* Advanced inversion lemmas *************************************************)
-lemma cpys_inv_atom1: ∀I,G,L,T2,d,e. ⦃G, L⦄ ⊢ ⓪{I} ▶*[d, e] T2 →
+lemma cpys_inv_atom1: ∀I,G,L,T2,l,m. ⦃G, L⦄ ⊢ ⓪{I} ▶*[l, m] T2 →
T2 = ⓪{I} ∨
- ∃∃J,K,V1,V2,i. d ≤ yinj i & i < d + e &
+ ∃∃J,K,V1,V2,i. l ≤ yinj i & i < l + m &
⬇[i] L ≡ K.ⓑ{J}V1 &
- ⦃G, K⦄ ⊢ V1 ▶*[0, ⫰(d+e-i)] V2 &
+ ⦃G, K⦄ ⊢ V1 ▶*[0, ⫰(l+m-i)] V2 &
⬆[O, i+1] V2 ≡ T2 &
I = LRef i.
-#I #G #L #T2 #d #e #H @(cpys_ind … H) -T2
+#I #G #L #T2 #l #m #H @(cpys_ind … H) -T2
[ /2 width=1 by or_introl/
| #T #T2 #_ #HT2 *
[ #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
+ | * #J #K #V1 #V #i #Hli #Hilm #HLK #HV1 #HVT #HI
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/ ]
]
qed-.
-lemma cpys_inv_lref1: ∀G,L,T2,i,d,e. ⦃G, L⦄ ⊢ #i ▶*[d, e] T2 →
+lemma cpys_inv_lref1: ∀G,L,T2,i,l,m. ⦃G, L⦄ ⊢ #i ▶*[l, m] T2 →
T2 = #i ∨
- ∃∃I,K,V1,V2. d ≤ i & i < d + e &
+ ∃∃I,K,V1,V2. l ≤ i & i < l + m &
⬇[i] L ≡ K.ⓑ{I}V1 &
- ⦃G, K⦄ ⊢ V1 ▶*[0, ⫰(d+e-i)] V2 &
+ ⦃G, K⦄ ⊢ V1 ▶*[0, ⫰(l+m-i)] V2 &
⬆[O, i+1] V2 ≡ T2.
-#G #L #T2 #i #d #e #H elim (cpys_inv_atom1 … H) -H /2 width=1 by or_introl/
-* #I #K #V1 #V2 #j #Hdj #Hjde #HLK #HV12 #HVT2 #H destruct /3 width=7 by ex5_4_intro, or_intror/
+#G #L #T2 #i #l #m #H elim (cpys_inv_atom1 … H) -H /2 width=1 by or_introl/
+* #I #K #V1 #V2 #j #Hlj #Hjlm #HLK #HV12 #HVT2 #H destruct /3 width=7 by ex5_4_intro, or_intror/
qed-.
-lemma cpys_inv_lref1_Y2: ∀G,L,T2,i,d. ⦃G, L⦄ ⊢ #i ▶*[d, ∞] T2 →
+lemma cpys_inv_lref1_Y2: ∀G,L,T2,i,l. ⦃G, L⦄ ⊢ #i ▶*[l, ∞] T2 →
T2 = #i ∨
- ∃∃I,K,V1,V2. d ≤ i & ⬇[i] L ≡ K.ⓑ{I}V1 &
+ ∃∃I,K,V1,V2. l ≤ i & ⬇[i] L ≡ K.ⓑ{I}V1 &
⦃G, K⦄ ⊢ V1 ▶*[0, ∞] V2 & ⬆[O, i+1] V2 ≡ T2.
-#G #L #T2 #i #d #H elim (cpys_inv_lref1 … H) -H /2 width=1 by or_introl/
+#G #L #T2 #i #l #H elim (cpys_inv_lref1 … H) -H /2 width=1 by or_introl/
* >yminus_Y_inj /3 width=7 by or_intror, ex4_4_intro/
qed-.
-lemma cpys_inv_lref1_drop: ∀G,L,T2,i,d,e. ⦃G, L⦄ ⊢ #i ▶*[d, e] T2 →
+lemma cpys_inv_lref1_drop: ∀G,L,T2,i,l,m. ⦃G, L⦄ ⊢ #i ▶*[l, m] T2 →
∀I,K,V1. ⬇[i] L ≡ K.ⓑ{I}V1 →
∀V2. ⬆[O, i+1] V2 ≡ T2 →
- ∧∧ ⦃G, K⦄ ⊢ V1 ▶*[0, ⫰(d+e-i)] V2
- & d ≤ i
- & i < d + e.
-#G #L #T2 #i #d #e #H #I #K #V1 #HLK #V2 #HVT2 elim (cpys_inv_lref1 … H) -H
+ ∧∧ ⦃G, K⦄ ⊢ V1 ▶*[0, ⫰(l+m-i)] V2
+ & l ≤ i
+ & i < l + m.
+#G #L #T2 #i #l #m #H #I #K #V1 #HLK #V2 #HVT2 elim (cpys_inv_lref1 … H) -H
[ #H destruct elim (lift_inv_lref2_be … HVT2) -HVT2 -HLK //
-| * #Z #Y #X1 #X2 #Hdi #Hide #HLY #HX12 #HXT2
+| * #Z #Y #X1 #X2 #Hli #Hilm #HLY #HX12 #HXT2
lapply (lift_inj … HXT2 … HVT2) -T2 #H destruct
lapply (drop_mono … HLY … HLK) -L #H destruct
/2 width=1 by and3_intro/
(* Properties on relocation *************************************************)
-lemma cpys_lift_le: ∀G,K,T1,T2,dt,et. ⦃G, K⦄ ⊢ T1 ▶*[dt, et] T2 →
- ∀L,U1,s,d,e. dt + et ≤ yinj d → ⬇[s, d, e] L ≡ K →
- ⬆[d, e] T1 ≡ U1 → ∀U2. ⬆[d, e] T2 ≡ U2 →
- ⦃G, L⦄ ⊢ U1 ▶*[dt, et] U2.
-#G #K #T1 #T2 #dt #et #H #L #U1 #s #d #e #Hdetd #HLK #HTU1 @(cpys_ind … H) -T2
+lemma cpys_lift_le: ∀G,K,T1,T2,lt,mt. ⦃G, K⦄ ⊢ T1 ▶*[lt, mt] T2 →
+ ∀L,U1,s,l,m. lt + mt ≤ yinj l → ⬇[s, l, m] L ≡ K →
+ ⬆[l, m] T1 ≡ U1 → ∀U2. ⬆[l, m] T2 ≡ U2 →
+ ⦃G, L⦄ ⊢ U1 ▶*[lt, mt] U2.
+#G #K #T1 #T2 #lt #mt #H #L #U1 #s #l #m #Hlmtl #HLK #HTU1 @(cpys_ind … H) -T2
[ #U2 #H >(lift_mono … HTU1 … H) -H //
| -HTU1 #T #T2 #_ #HT2 #IHT #U2 #HTU2
- elim (lift_total T d e) #U #HTU
+ elim (lift_total T l m) #U #HTU
lapply (IHT … HTU) -IHT #HU1
lapply (cpy_lift_le … HT2 … HLK HTU HTU2 ?) -HT2 -HLK -HTU -HTU2 /2 width=3 by cpys_strap1/
]
qed-.
-lemma cpys_lift_be: ∀G,K,T1,T2,dt,et. ⦃G, K⦄ ⊢ T1 ▶*[dt, et] T2 →
- ∀L,U1,s,d,e. dt ≤ yinj d → d ≤ dt + et →
- ⬇[s, d, e] L ≡ K → ⬆[d, e] T1 ≡ U1 →
- ∀U2. ⬆[d, e] T2 ≡ U2 → ⦃G, L⦄ ⊢ U1 ▶*[dt, et + e] U2.
-#G #K #T1 #T2 #dt #et #H #L #U1 #s #d #e #Hdtd #Hddet #HLK #HTU1 @(cpys_ind … H) -T2
+lemma cpys_lift_be: ∀G,K,T1,T2,lt,mt. ⦃G, K⦄ ⊢ T1 ▶*[lt, mt] T2 →
+ ∀L,U1,s,l,m. lt ≤ yinj l → l ≤ lt + mt →
+ ⬇[s, l, m] L ≡ K → ⬆[l, m] T1 ≡ U1 →
+ ∀U2. ⬆[l, m] T2 ≡ U2 → ⦃G, L⦄ ⊢ U1 ▶*[lt, mt + m] U2.
+#G #K #T1 #T2 #lt #mt #H #L #U1 #s #l #m #Hltl #Hllmt #HLK #HTU1 @(cpys_ind … H) -T2
[ #U2 #H >(lift_mono … HTU1 … H) -H //
| -HTU1 #T #T2 #_ #HT2 #IHT #U2 #HTU2
- elim (lift_total T d e) #U #HTU
+ elim (lift_total T l m) #U #HTU
lapply (IHT … HTU) -IHT #HU1
lapply (cpy_lift_be … HT2 … HLK HTU HTU2 ? ?) -HT2 -HLK -HTU -HTU2 /2 width=3 by cpys_strap1/
]
qed-.
-lemma cpys_lift_ge: ∀G,K,T1,T2,dt,et. ⦃G, K⦄ ⊢ T1 ▶*[dt, et] T2 →
- ∀L,U1,s,d,e. yinj d ≤ dt → ⬇[s, d, e] L ≡ K →
- ⬆[d, e] T1 ≡ U1 → ∀U2. ⬆[d, e] T2 ≡ U2 →
- ⦃G, L⦄ ⊢ U1 ▶*[dt+e, et] U2.
-#G #K #T1 #T2 #dt #et #H #L #U1 #s #d #e #Hddt #HLK #HTU1 @(cpys_ind … H) -T2
+lemma cpys_lift_ge: ∀G,K,T1,T2,lt,mt. ⦃G, K⦄ ⊢ T1 ▶*[lt, mt] T2 →
+ ∀L,U1,s,l,m. yinj l ≤ lt → ⬇[s, l, m] L ≡ K →
+ ⬆[l, m] T1 ≡ U1 → ∀U2. ⬆[l, m] T2 ≡ U2 →
+ ⦃G, L⦄ ⊢ U1 ▶*[lt+m, mt] U2.
+#G #K #T1 #T2 #lt #mt #H #L #U1 #s #l #m #Hllt #HLK #HTU1 @(cpys_ind … H) -T2
[ #U2 #H >(lift_mono … HTU1 … H) -H //
| -HTU1 #T #T2 #_ #HT2 #IHT #U2 #HTU2
- elim (lift_total T d e) #U #HTU
+ elim (lift_total T l m) #U #HTU
lapply (IHT … HTU) -IHT #HU1
lapply (cpy_lift_ge … HT2 … HLK HTU HTU2 ?) -HT2 -HLK -HTU -HTU2 /2 width=3 by cpys_strap1/
]
(* Inversion lemmas for relocation ******************************************)
-lemma cpys_inv_lift1_le: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶*[dt, et] U2 →
- ∀K,s,d,e. ⬇[s, d, e] L ≡ K → ∀T1. ⬆[d, e] T1 ≡ U1 →
- dt + et ≤ d →
- ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶*[dt, et] T2 & ⬆[d, e] T2 ≡ U2.
-#G #L #U1 #U2 #dt #et #H #K #s #d #e #HLK #T1 #HTU1 #Hdetd @(cpys_ind … H) -U2
+lemma cpys_inv_lift1_le: ∀G,L,U1,U2,lt,mt. ⦃G, L⦄ ⊢ U1 ▶*[lt, mt] U2 →
+ ∀K,s,l,m. ⬇[s, l, m] L ≡ K → ∀T1. ⬆[l, m] T1 ≡ U1 →
+ lt + mt ≤ l →
+ ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶*[lt, mt] T2 & ⬆[l, m] T2 ≡ U2.
+#G #L #U1 #U2 #lt #mt #H #K #s #l #m #HLK #T1 #HTU1 #Hlmtl @(cpys_ind … H) -U2
[ /2 width=3 by ex2_intro/
| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU
elim (cpy_inv_lift1_le … HU2 … HLK … HTU) -HU2 -HLK -HTU /3 width=3 by cpys_strap1, ex2_intro/
]
qed-.
-lemma cpys_inv_lift1_be: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶*[dt, et] U2 →
- ∀K,s,d,e. ⬇[s, d, e] L ≡ K → ∀T1. ⬆[d, e] T1 ≡ U1 →
- dt ≤ d → yinj d + e ≤ dt + et →
- ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶*[dt, et - e] T2 & ⬆[d, e] T2 ≡ U2.
-#G #L #U1 #U2 #dt #et #H #K #s #d #e #HLK #T1 #HTU1 #Hdtd #Hdedet @(cpys_ind … H) -U2
+lemma cpys_inv_lift1_be: ∀G,L,U1,U2,lt,mt. ⦃G, L⦄ ⊢ U1 ▶*[lt, mt] U2 →
+ ∀K,s,l,m. ⬇[s, l, m] L ≡ K → ∀T1. ⬆[l, m] T1 ≡ U1 →
+ lt ≤ l → yinj l + m ≤ lt + mt →
+ ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶*[lt, mt - m] T2 & ⬆[l, m] T2 ≡ U2.
+#G #L #U1 #U2 #lt #mt #H #K #s #l #m #HLK #T1 #HTU1 #Hltl #Hlmlmt @(cpys_ind … H) -U2
[ /2 width=3 by ex2_intro/
| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU
elim (cpy_inv_lift1_be … HU2 … HLK … HTU) -HU2 -HLK -HTU /3 width=3 by cpys_strap1, ex2_intro/
]
qed-.
-lemma cpys_inv_lift1_ge: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶*[dt, et] U2 →
- ∀K,s,d,e. ⬇[s, d, e] L ≡ K → ∀T1. ⬆[d, e] T1 ≡ U1 →
- yinj d + e ≤ dt →
- ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶*[dt - e, et] T2 & ⬆[d, e] T2 ≡ U2.
-#G #L #U1 #U2 #dt #et #H #K #s #d #e #HLK #T1 #HTU1 #Hdedt @(cpys_ind … H) -U2
+lemma cpys_inv_lift1_ge: ∀G,L,U1,U2,lt,mt. ⦃G, L⦄ ⊢ U1 ▶*[lt, mt] U2 →
+ ∀K,s,l,m. ⬇[s, l, m] L ≡ K → ∀T1. ⬆[l, m] T1 ≡ U1 →
+ yinj l + m ≤ lt →
+ ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶*[lt - m, mt] T2 & ⬆[l, m] T2 ≡ U2.
+#G #L #U1 #U2 #lt #mt #H #K #s #l #m #HLK #T1 #HTU1 #Hlmlt @(cpys_ind … H) -U2
[ /2 width=3 by ex2_intro/
| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU
elim (cpy_inv_lift1_ge … HU2 … HLK … HTU) -HU2 -HLK -HTU /3 width=3 by cpys_strap1, ex2_intro/
(* Advanced inversion lemmas on relocation **********************************)
-lemma cpys_inv_lift1_ge_up: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶*[dt, et] U2 →
- ∀K,s,d,e. ⬇[s, d, e] L ≡ K → ∀T1. ⬆[d, e] T1 ≡ U1 →
- d ≤ dt → dt ≤ yinj d + e → yinj d + e ≤ dt + et →
- ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶*[d, dt + et - (yinj d + e)] T2 &
- ⬆[d, e] T2 ≡ U2.
-#G #L #U1 #U2 #dt #et #H #K #s #d #e #HLK #T1 #HTU1 #Hddt #Hdtde #Hdedet @(cpys_ind … H) -U2
+lemma cpys_inv_lift1_ge_up: ∀G,L,U1,U2,lt,mt. ⦃G, L⦄ ⊢ U1 ▶*[lt, mt] U2 →
+ ∀K,s,l,m. ⬇[s, l, m] L ≡ K → ∀T1. ⬆[l, m] T1 ≡ U1 →
+ l ≤ lt → lt ≤ yinj l + m → yinj l + m ≤ lt + mt →
+ ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶*[l, lt + mt - (yinj l + m)] T2 &
+ ⬆[l, m] T2 ≡ U2.
+#G #L #U1 #U2 #lt #mt #H #K #s #l #m #HLK #T1 #HTU1 #Hllt #Hltlm #Hlmlmt @(cpys_ind … H) -U2
[ /2 width=3 by ex2_intro/
| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU
elim (cpy_inv_lift1_ge_up … HU2 … HLK … HTU) -HU2 -HLK -HTU /3 width=3 by cpys_strap1, ex2_intro/
]
qed-.
-lemma cpys_inv_lift1_be_up: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶*[dt, et] U2 →
- ∀K,s,d,e. ⬇[s, d, e] L ≡ K → ∀T1. ⬆[d, e] T1 ≡ U1 →
- dt ≤ d → dt + et ≤ yinj d + e →
- ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶*[dt, d - dt] T2 & ⬆[d, e] T2 ≡ U2.
-#G #L #U1 #U2 #dt #et #H #K #s #d #e #HLK #T1 #HTU1 #Hdtd #Hdetde @(cpys_ind … H) -U2
+lemma cpys_inv_lift1_be_up: ∀G,L,U1,U2,lt,mt. ⦃G, L⦄ ⊢ U1 ▶*[lt, mt] U2 →
+ ∀K,s,l,m. ⬇[s, l, m] L ≡ K → ∀T1. ⬆[l, m] T1 ≡ U1 →
+ lt ≤ l → lt + mt ≤ yinj l + m →
+ ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶*[lt, l - lt] T2 & ⬆[l, m] T2 ≡ U2.
+#G #L #U1 #U2 #lt #mt #H #K #s #l #m #HLK #T1 #HTU1 #Hltl #Hlmtlm @(cpys_ind … H) -U2
[ /2 width=3 by ex2_intro/
| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU
elim (cpy_inv_lift1_be_up … HU2 … HLK … HTU) -HU2 -HLK -HTU /3 width=3 by cpys_strap1, ex2_intro/
]
qed-.
-lemma cpys_inv_lift1_le_up: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶*[dt, et] U2 →
- ∀K,s,d,e. ⬇[s, d, e] L ≡ K → ∀T1. ⬆[d, e] T1 ≡ U1 →
- dt ≤ d → d ≤ dt + et → dt + et ≤ yinj d + e →
- ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶*[dt, d - dt] T2 & ⬆[d, e] T2 ≡ U2.
-#G #L #U1 #U2 #dt #et #H #K #s #d #e #HLK #T1 #HTU1 #Hdtd #Hddet #Hdetde @(cpys_ind … H) -U2
+lemma cpys_inv_lift1_le_up: ∀G,L,U1,U2,lt,mt. ⦃G, L⦄ ⊢ U1 ▶*[lt, mt] U2 →
+ ∀K,s,l,m. ⬇[s, l, m] L ≡ K → ∀T1. ⬆[l, m] T1 ≡ U1 →
+ lt ≤ l → l ≤ lt + mt → lt + mt ≤ yinj l + m →
+ ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶*[lt, l - lt] T2 & ⬆[l, m] T2 ≡ U2.
+#G #L #U1 #U2 #lt #mt #H #K #s #l #m #HLK #T1 #HTU1 #Hltl #Hllmt #Hlmtlm @(cpys_ind … H) -U2
[ /2 width=3 by ex2_intro/
| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU
elim (cpy_inv_lift1_le_up … HU2 … HLK … HTU) -HU2 -HLK -HTU /3 width=3 by cpys_strap1, ex2_intro/
]
qed-.
-lemma cpys_inv_lift1_subst: ∀G,L,W1,W2,d,e. ⦃G, L⦄ ⊢ W1 ▶*[d, e] W2 →
+lemma cpys_inv_lift1_subst: ∀G,L,W1,W2,l,m. ⦃G, L⦄ ⊢ W1 ▶*[l, m] W2 →
∀K,V1,i. ⬇[i+1] L ≡ K → ⬆[O, i+1] V1 ≡ W1 →
- d ≤ yinj i → i < d + e →
- ∃∃V2. ⦃G, K⦄ ⊢ V1 ▶*[O, ⫰(d+e-i)] V2 & ⬆[O, i+1] V2 ≡ W2.
-#G #L #W1 #W2 #d #e #HW12 #K #V1 #i #HLK #HVW1 #Hdi #Hide
+ l ≤ yinj i → i < l + m →
+ ∃∃V2. ⦃G, K⦄ ⊢ V1 ▶*[O, ⫰(l+m-i)] V2 & ⬆[O, i+1] V2 ≡ W2.
+#G #L #W1 #W2 #l #m #HW12 #K #V1 #i #HLK #HVW1 #Hli #Hilm
elim (cpys_inv_lift1_ge_up … HW12 … HLK … HVW1 ? ? ?) //
>yplus_O1 <yplus_inj >yplus_SO2
[ >yminus_succ2 /2 width=3 by ex2_intro/
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
+| drops_cons: ∀L1,L,L2,des,l,m.
+ drops s des L1 L → ⬇[s, l, m] L ≡ L2 → drops s ({l, m} @ des) L1 L2
.
interpretation "iterated slicing (local environment) abstract"
'RDropStar des T1 T2 = (drops true des T1 T2).
*)
-definition l_liftable1: relation2 lenv term → predicate bool ≝
- λR,s. ∀K,T. R K T → ∀L,d,e. ⬇[s, d, e] L ≡ K →
- ∀U. ⬆[d, e] T ≡ U → R L U.
+definition d_liftable1: relation2 lenv term → predicate bool ≝
+ λR,s. ∀K,T. R K T → ∀L,l,m. ⬇[s, l, m] L ≡ K →
+ ∀U. ⬆[l, m] T ≡ U → R L U.
-definition l_liftables1: relation2 lenv term → predicate bool ≝
+definition d_liftables1: relation2 lenv term → predicate bool ≝
λR,s. ∀L,K,des. ⬇*[s, des] L ≡ K →
∀T,U. ⬆*[des] T ≡ U → R K T → R L U.
-definition l_liftables1_all: relation2 lenv term → predicate bool ≝
+definition d_liftables1_all: relation2 lenv term → predicate bool ≝
λR,s. ∀L,K,des. ⬇*[s, des] L ≡ K →
∀Ts,Us. ⬆*[des] Ts ≡ Us →
all … (R K) Ts → all … (R L) Us.
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
+#L1 #L #L2 #l #m #des #_ #_ #H destruct
qed-.
(* Basic_1: was: drop1_gen_pnil *)
/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.
+ ∀l,m,tl. des = {l, m} @ tl →
+ ∃∃L. ⬇*[s, tl] L1 ≡ L & ⬇[s, l, m] 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
+[ #L #l #m #tl #H destruct
+| #L1 #L #L2 #des #l #m #HT1 #HT2 #l0 #m0 #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.
+lemma drops_inv_cons: ∀L1,L2,s,l,m,des. ⬇*[s, {l, m} @ des] L1 ≡ L2 →
+ ∃∃L. ⬇*[s, des] L1 ≡ L & ⬇[s, l, m] L ≡ L2.
/2 width=3 by drops_inv_cons_aux/ qed-.
lemma drops_inv_skip2: ∀I,s,des,des2,i. des ▭ i ≡ des2 →
#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
+| #des #des2 #l #m #i #Hil #_ #IHcs2 #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
+ elim (IHcs2 … HL1) -IHcs2 -HL1 #K1 #V1 #des1 #Hcs1 #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
+| #des #des2 #l #m #i #Hil #_ #IHcs2 #L1 #K2 #V2 #H
+ elim (IHcs2 … H) -IHcs2 -H #K1 #V1 #des1 #Hcs1 #HK1 #HV1 #X destruct
/4 width=7 by minuss_ge, ex4_3_intro, le_S_S/
]
qed-.
#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
+| #L1 #L #L2 #des #l #m #_ #HL2 #IHL #V1 #V2 #H #I
elim (lifts_inv_cons … H) -H /3 width=5 by drop_skip, drops_cons/
].
qed.
-lemma l1_liftable_liftables: ∀R,s. l_liftable1 R s → l_liftables1 R s.
+lemma d1_liftable_liftables: ∀R,s. d_liftable1 R s → d_liftables1 R s.
#R #s #HR #L #K #des #H elim H -L -K -des
[ #L #T #U #H #HT <(lifts_inv_nil … H) -H //
-| #L1 #L #L2 #des #d #e #_ #HL2 #IHL #T2 #T1 #H #HLT2
+| #L1 #L #L2 #des #l #m #_ #HL2 #IHL #T2 #T1 #H #HLT2
elim (lifts_inv_cons … H) -H /3 width=10 by/
]
qed.
-lemma l1_liftables_liftables_all: ∀R,s. l_liftables1 R s → l_liftables1_all R s.
+lemma d1_liftables_liftables_all: ∀R,s. d_liftables1 R s → d_liftables1_all R s.
#R #s #HR #L #K #des #HLK #Ts #Us #H elim H -Ts -Us normalize //
#Ts #Us #T #U #HTU #_ #IHTUs * /3 width=7 by conj/
qed.
@⦃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
+| #L1 #L0 #L #des #l #m #_ #HL0 #IHL0 #L2 #i #HL2
+ elim (lt_or_ge i l) #Hil
[ 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
(* LAZY EQUIVALENCE FOR CLOSURES ********************************************)
-inductive fleq (d) (G) (L1) (T): relation3 genv lenv term ≝
-| fleq_intro: ∀L2. L1 ≡[T, d] L2 → fleq d G L1 T G L2 T
+inductive fleq (l) (G) (L1) (T): relation3 genv lenv term ≝
+| fleq_intro: ∀L2. L1 ≡[T, l] L2 → fleq l G L1 T G L2 T
.
interpretation
"lazy equivalence (closure)"
- 'LazyEq d G1 L1 T1 G2 L2 T2 = (fleq d G1 L1 T1 G2 L2 T2).
+ 'LazyEq l G1 L1 T1 G2 L2 T2 = (fleq l G1 L1 T1 G2 L2 T2).
(* Basic properties *********************************************************)
-lemma fleq_refl: ∀d. tri_reflexive … (fleq d).
+lemma fleq_refl: ∀l. tri_reflexive … (fleq l).
/2 width=1 by fleq_intro/ qed.
-lemma fleq_sym: ∀d. tri_symmetric … (fleq d).
-#d #G1 #L1 #T1 #G2 #L2 #T2 * /3 width=1 by fleq_intro, lleq_sym/
+lemma fleq_sym: ∀l. tri_symmetric … (fleq l).
+#l #G1 #L1 #T1 #G2 #L2 #T2 * /3 width=1 by fleq_intro, lleq_sym/
qed-.
(* Basic inversion lemmas ***************************************************)
-lemma fleq_inv_gen: ∀G1,G2,L1,L2,T1,T2,d. ⦃G1, L1, T1⦄ ≡[d] ⦃G2, L2, T2⦄ →
- ∧∧ G1 = G2 & L1 ≡[T1, d] L2 & T1 = T2.
-#G1 #G2 #L1 #L2 #T1 #T2 #d * -G2 -L2 -T2 /2 width=1 by and3_intro/
+lemma fleq_inv_gen: ∀G1,G2,L1,L2,T1,T2,l. ⦃G1, L1, T1⦄ ≡[l] ⦃G2, L2, T2⦄ →
+ ∧∧ G1 = G2 & L1 ≡[T1, l] L2 & T1 = T2.
+#G1 #G2 #L1 #L2 #T1 #T2 #l * -G2 -L2 -T2 /2 width=1 by and3_intro/
qed-.
(* Main properties **********************************************************)
-theorem fleq_trans: ∀d. tri_transitive … (fleq d).
-#d #G1 #G #L1 #L #T1 #T * -G -L -T
+theorem fleq_trans: ∀l. tri_transitive … (fleq l).
+#l #G1 #G #L1 #L #T1 #T * -G -L -T
#L #HT1 #G2 #L2 #T2 * -G2 -L2 -T2
/3 width=3 by lleq_trans, fleq_intro/
qed-.
-theorem fleq_canc_sn: ∀G,G1,G2,L,L1,L2,T,T1,T2,d.
- ⦃G, L, T⦄ ≡[d] ⦃G1, L1, T1⦄→ ⦃G, L, T⦄ ≡[d] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≡[d] ⦃G2, L2, T2⦄.
+theorem fleq_canc_sn: ∀G,G1,G2,L,L1,L2,T,T1,T2,l.
+ ⦃G, L, T⦄ ≡[l] ⦃G1, L1, T1⦄→ ⦃G, L, T⦄ ≡[l] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≡[l] ⦃G2, L2, T2⦄.
/3 width=5 by fleq_trans, fleq_sym/ qed-.
-theorem fleq_canc_dx: ∀G1,G2,G,L1,L2,L,T1,T2,T,d.
- ⦃G1, L1, T1⦄ ≡[d] ⦃G, L, T⦄ → ⦃G2, L2, T2⦄ ≡[d] ⦃G, L, T⦄ → ⦃G1, L1, T1⦄ ≡[d] ⦃G2, L2, T2⦄.
+theorem fleq_canc_dx: ∀G1,G2,G,L1,L2,L,T1,T2,T,l.
+ ⦃G1, L1, T1⦄ ≡[l] ⦃G, L, T⦄ → ⦃G2, L2, T2⦄ ≡[l] ⦃G, L, T⦄ → ⦃G1, L1, T1⦄ ≡[l] ⦃G2, L2, T2⦄.
/3 width=5 by fleq_trans, fleq_sym/ qed-.
⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄.
/2 width=5 by tri_TC_strap/ qed.
-lemma fqup_drop: ∀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,m. ⬇[m] L1 ≡ K1 → ⬆[0, m] 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
+#G1 #G2 #L1 #K1 #K2 #T1 #T2 #U1 #m #HLK1 #HTU1 #HT12 elim (eq_or_gt … m) #H destruct
[ >(drop_inv_O2 … HLK1) -L1 <(lift_inv_O2 … HTU1) -U1 //
| /3 width=5 by fqup_strap2, fqu_drop_lt/
]
/2 width=5 by tri_TC_strap/ qed-.
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 →
+ ∀L1,U1,m. ⬇[m] L1 ≡ K1 → ⬆[0, m] T1 ≡ U1 →
⦃G1, L1, U1⦄ ⊐* ⦃G2, K2, T2⦄.
#G1 #G2 #K1 #K2 #T1 #T2 #H @(fqus_ind … H) -G2 -K2 -T2
/3 width=5 by fqus_strap1, fquq_fqus, fquq_drop/
(* CONTEXT-SENSITIVE FREE VARIABLES *****************************************)
inductive frees: relation4 ynat lenv term nat ≝
-| frees_eq: ∀L,U,d,i. (∀T. ⬆[i, 1] T ≡ U → ⊥) → frees d L U i
-| frees_be: ∀I,L,K,U,W,d,i,j. d ≤ yinj j → j < i →
+| frees_eq: ∀L,U,l,i. (∀T. ⬆[i, 1] T ≡ U → ⊥) → frees l L U i
+| frees_be: ∀I,L,K,U,W,l,i,j. l ≤ yinj j → j < i →
(∀T. ⬆[j, 1] T ≡ U → ⊥) → ⬇[j]L ≡ K.ⓑ{I}W →
- frees 0 K W (i-j-1) → frees d L U i.
+ frees 0 K W (i-j-1) → frees l L U i.
interpretation
"context-sensitive free variables (term)"
- 'FreeStar L i d U = (frees d L U i).
+ 'FreeStar L i l U = (frees l L U i).
definition frees_trans: predicate (relation3 lenv term term) ≝
λR. ∀L,U1,U2,i. R L U1 U2 → L ⊢ i ϵ 𝐅*[0]⦃U2⦄ → L ⊢ i ϵ 𝐅*[0]⦃U1⦄.
(* Basic inversion lemmas ***************************************************)
-lemma frees_inv: ∀L,U,d,i. L ⊢ i ϵ 𝐅*[d]⦃U⦄ →
+lemma frees_inv: ∀L,U,l,i. L ⊢ i ϵ 𝐅*[l]⦃U⦄ →
(∀T. ⬆[i, 1] T ≡ U → ⊥) ∨
- ∃∃I,K,W,j. d ≤ yinj j & j < i & (∀T. ⬆[j, 1] T ≡ U → ⊥) &
+ ∃∃I,K,W,j. l ≤ yinj j & j < i & (∀T. ⬆[j, 1] T ≡ U → ⊥) &
⬇[j]L ≡ K.ⓑ{I}W & K ⊢ (i-j-1) ϵ 𝐅*[yinj 0]⦃W⦄.
-#L #U #d #i * -L -U -d -i /4 width=9 by ex5_4_intro, or_intror, or_introl/
+#L #U #l #i * -L -U -l -i /4 width=9 by ex5_4_intro, or_intror, or_introl/
qed-.
-lemma frees_inv_sort: ∀L,d,i,k. L ⊢ i ϵ 𝐅*[d]⦃⋆k⦄ → ⊥.
-#L #d #i #k #H elim (frees_inv … H) -H [|*] /2 width=2 by/
+lemma frees_inv_sort: ∀L,l,i,k. L ⊢ i ϵ 𝐅*[l]⦃⋆k⦄ → ⊥.
+#L #l #i #k #H elim (frees_inv … H) -H [|*] /2 width=2 by/
qed-.
-lemma frees_inv_gref: ∀L,d,i,p. L ⊢ i ϵ 𝐅*[d]⦃§p⦄ → ⊥.
-#L #d #i #p #H elim (frees_inv … H) -H [|*] /2 width=2 by/
+lemma frees_inv_gref: ∀L,l,i,p. L ⊢ i ϵ 𝐅*[l]⦃§p⦄ → ⊥.
+#L #l #i #p #H elim (frees_inv … H) -H [|*] /2 width=2 by/
qed-.
-lemma frees_inv_lref: ∀L,d,j,i. L ⊢ i ϵ 𝐅*[d]⦃#j⦄ →
+lemma frees_inv_lref: ∀L,l,j,i. L ⊢ i ϵ 𝐅*[l]⦃#j⦄ →
j = i ∨
- ∃∃I,K,W. d ≤ yinj j & j < i & ⬇[j] L ≡ K.ⓑ{I}W & K ⊢ (i-j-1) ϵ 𝐅*[yinj 0]⦃W⦄.
-#L #d #x #i #H elim (frees_inv … H) -H
+ ∃∃I,K,W. l ≤ yinj j & j < i & ⬇[j] L ≡ K.ⓑ{I}W & K ⊢ (i-j-1) ϵ 𝐅*[yinj 0]⦃W⦄.
+#L #l #x #i #H elim (frees_inv … H) -H
[ /4 width=2 by nlift_inv_lref_be_SO, or_introl/
-| * #I #K #W #j #Hdj #Hji #Hnx #HLK #HW
+| * #I #K #W #j #Hlj #Hji #Hnx #HLK #HW
>(nlift_inv_lref_be_SO … Hnx) -x /3 width=5 by ex4_3_intro, or_intror/
]
qed-.
-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 //
+lemma frees_inv_lref_free: ∀L,l,j,i. L ⊢ i ϵ 𝐅*[l]⦃#j⦄ → |L| ≤ j → j = i.
+#L #l #j #i #H #Hj elim (frees_inv_lref … H) -H //
* #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-.
-lemma frees_inv_lref_skip: ∀L,d,j,i. L ⊢ i ϵ 𝐅*[d]⦃#j⦄ → yinj j < d → j = i.
-#L #d #j #i #H #Hjd elim (frees_inv_lref … H) -H //
-* #I #K #W #Hdj elim (ylt_yle_false … Hdj) -Hdj //
+lemma frees_inv_lref_skip: ∀L,l,j,i. L ⊢ i ϵ 𝐅*[l]⦃#j⦄ → yinj j < l → j = i.
+#L #l #j #i #H #Hjl elim (frees_inv_lref … H) -H //
+* #I #K #W #Hlj elim (ylt_yle_false … Hlj) -Hlj //
qed-.
-lemma frees_inv_lref_ge: ∀L,d,j,i. L ⊢ i ϵ 𝐅*[d]⦃#j⦄ → i ≤ j → j = i.
-#L #d #j #i #H #Hij elim (frees_inv_lref … H) -H //
-* #I #K #W #_ #Hji elim (lt_refl_false j) -I -L -K -W -d /2 width=3 by lt_to_le_to_lt/
+lemma frees_inv_lref_ge: ∀L,l,j,i. L ⊢ i ϵ 𝐅*[l]⦃#j⦄ → i ≤ j → j = i.
+#L #l #j #i #H #Hij elim (frees_inv_lref … H) -H //
+* #I #K #W #_ #Hji elim (lt_refl_false j) -I -L -K -W -l /2 width=3 by lt_to_le_to_lt/
qed-.
-lemma frees_inv_lref_lt: ∀L,d,j,i.L ⊢ i ϵ 𝐅*[d]⦃#j⦄ → j < i →
- ∃∃I,K,W. d ≤ yinj j & ⬇[j] L ≡ K.ⓑ{I}W & K ⊢ (i-j-1) ϵ 𝐅*[yinj 0]⦃W⦄.
-#L #d #j #i #H #Hji elim (frees_inv_lref … H) -H
+lemma frees_inv_lref_lt: ∀L,l,j,i.L ⊢ i ϵ 𝐅*[l]⦃#j⦄ → j < i →
+ ∃∃I,K,W. l ≤ yinj j & ⬇[j] L ≡ K.ⓑ{I}W & K ⊢ (i-j-1) ϵ 𝐅*[yinj 0]⦃W⦄.
+#L #l #j #i #H #Hji elim (frees_inv_lref … H) -H
[ #H elim (lt_refl_false j) //
| * /2 width=5 by ex3_3_intro/
]
qed-.
-lemma frees_inv_bind: ∀a,I,L,W,U,d,i. L ⊢ i ϵ 𝐅*[d]⦃ⓑ{a,I}W.U⦄ →
- L ⊢ i ϵ 𝐅*[d]⦃W⦄ ∨ L.ⓑ{I}W ⊢ i+1 ϵ 𝐅*[⫯d]⦃U⦄ .
-#a #J #L #V #U #d #i #H elim (frees_inv … H) -H
+lemma frees_inv_bind: ∀a,I,L,W,U,l,i. L ⊢ i ϵ 𝐅*[l]⦃ⓑ{a,I}W.U⦄ →
+ L ⊢ i ϵ 𝐅*[l]⦃W⦄ ∨ L.ⓑ{I}W ⊢ i+1 ϵ 𝐅*[⫯l]⦃U⦄ .
+#a #J #L #V #U #l #i #H elim (frees_inv … H) -H
[ #HnX elim (nlift_inv_bind … HnX) -HnX
/4 width=2 by frees_eq, or_intror, or_introl/
-| * #I #K #W #j #Hdj #Hji #HnX #HLK #HW elim (nlift_inv_bind … HnX) -HnX
+| * #I #K #W #j #Hlj #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 drop_drop, yle_succ, lt_minus_to_plus/
]
qed-.
-lemma frees_inv_flat: ∀I,L,W,U,d,i. L ⊢ i ϵ 𝐅*[d]⦃ⓕ{I}W.U⦄ →
- L ⊢ i ϵ 𝐅*[d]⦃W⦄ ∨ L ⊢ i ϵ 𝐅*[d]⦃U⦄ .
-#J #L #V #U #d #i #H elim (frees_inv … H) -H
+lemma frees_inv_flat: ∀I,L,W,U,l,i. L ⊢ i ϵ 𝐅*[l]⦃ⓕ{I}W.U⦄ →
+ L ⊢ i ϵ 𝐅*[l]⦃W⦄ ∨ L ⊢ i ϵ 𝐅*[l]⦃U⦄ .
+#J #L #V #U #l #i #H elim (frees_inv … H) -H
[ #HnX elim (nlift_inv_flat … HnX) -HnX
/4 width=2 by frees_eq, or_intror, or_introl/
-| * #I #K #W #j #Hdj #Hji #HnX #HLK #HW elim (nlift_inv_flat … HnX) -HnX
+| * #I #K #W #j #Hlj #Hji #HnX #HLK #HW elim (nlift_inv_flat … HnX) -HnX
/4 width=9 by frees_be, or_intror, or_introl/
]
qed-.
(* Basic properties *********************************************************)
-lemma frees_lref_eq: ∀L,d,i. L ⊢ i ϵ 𝐅*[d]⦃#i⦄.
+lemma frees_lref_eq: ∀L,l,i. L ⊢ i ϵ 𝐅*[l]⦃#i⦄.
/3 width=7 by frees_eq, lift_inv_lref2_be/ qed.
-lemma frees_lref_be: ∀I,L,K,W,d,i,j. d ≤ yinj j → j < i → ⬇[j]L ≡ K.ⓑ{I}W →
- K ⊢ i-j-1 ϵ 𝐅*[0]⦃W⦄ → L ⊢ i ϵ 𝐅*[d]⦃#j⦄.
+lemma frees_lref_be: ∀I,L,K,W,l,i,j. l ≤ yinj j → j < i → ⬇[j]L ≡ K.ⓑ{I}W →
+ K ⊢ i-j-1 ϵ 𝐅*[0]⦃W⦄ → L ⊢ i ϵ 𝐅*[l]⦃#j⦄.
/3 width=9 by frees_be, lift_inv_lref2_be/ qed.
-lemma frees_bind_sn: ∀a,I,L,W,U,d,i. L ⊢ i ϵ 𝐅*[d]⦃W⦄ →
- L ⊢ i ϵ 𝐅*[d]⦃ⓑ{a,I}W.U⦄.
-#a #I #L #W #U #d #i #H elim (frees_inv … H) -H [|*]
+lemma frees_bind_sn: ∀a,I,L,W,U,l,i. L ⊢ i ϵ 𝐅*[l]⦃W⦄ →
+ L ⊢ i ϵ 𝐅*[l]⦃ⓑ{a,I}W.U⦄.
+#a #I #L #W #U #l #i #H elim (frees_inv … H) -H [|*]
/4 width=9 by frees_be, frees_eq, nlift_bind_sn/
qed.
-lemma frees_bind_dx: ∀a,I,L,W,U,d,i. L.ⓑ{I}W ⊢ i+1 ϵ 𝐅*[⫯d]⦃U⦄ →
- L ⊢ i ϵ 𝐅*[d]⦃ⓑ{a,I}W.U⦄.
-#a #J #L #V #U #d #i #H elim (frees_inv … H) -H
+lemma frees_bind_dx: ∀a,I,L,W,U,l,i. L.ⓑ{I}W ⊢ i+1 ϵ 𝐅*[⫯l]⦃U⦄ →
+ L ⊢ i ϵ 𝐅*[l]⦃ⓑ{a,I}W.U⦄.
+#a #J #L #V #U #l #i #H elim (frees_inv … H) -H
[ /4 width=9 by frees_eq, nlift_bind_dx/
-| * #I #K #W #j #Hdj #Hji #HnU #HLK #HW
- elim (yle_inv_succ1 … Hdj) -Hdj <yminus_SO2 #Hyj #H
+| * #I #K #W #j #Hlj #Hji #HnU #HLK #HW
+ elim (yle_inv_succ1 … Hlj) -Hlj <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, drop_inv_drop1_lt, lt_plus_to_minus/
]
qed.
-lemma frees_flat_sn: ∀I,L,W,U,d,i. L ⊢ i ϵ 𝐅*[d]⦃W⦄ →
- L ⊢ i ϵ 𝐅*[d]⦃ⓕ{I}W.U⦄.
-#I #L #W #U #d #i #H elim (frees_inv … H) -H [|*]
+lemma frees_flat_sn: ∀I,L,W,U,l,i. L ⊢ i ϵ 𝐅*[l]⦃W⦄ →
+ L ⊢ i ϵ 𝐅*[l]⦃ⓕ{I}W.U⦄.
+#I #L #W #U #l #i #H elim (frees_inv … H) -H [|*]
/4 width=9 by frees_be, frees_eq, nlift_flat_sn/
qed.
-lemma frees_flat_dx: ∀I,L,W,U,d,i. L ⊢ i ϵ 𝐅*[d]⦃U⦄ →
- L ⊢ i ϵ 𝐅*[d]⦃ⓕ{I}W.U⦄.
-#I #L #W #U #d #i #H elim (frees_inv … H) -H [|*]
+lemma frees_flat_dx: ∀I,L,W,U,l,i. L ⊢ i ϵ 𝐅*[l]⦃U⦄ →
+ L ⊢ i ϵ 𝐅*[l]⦃ⓕ{I}W.U⦄.
+#I #L #W #U #l #i #H elim (frees_inv … H) -H [|*]
/4 width=9 by frees_be, frees_eq, nlift_flat_dx/
qed.
-lemma frees_weak: ∀L,U,d1,i. L ⊢ i ϵ 𝐅*[d1]⦃U⦄ →
- ∀d2. d2 ≤ d1 → L ⊢ i ϵ 𝐅*[d2]⦃U⦄.
-#L #U #d1 #i #H elim H -L -U -d1 -i
+lemma frees_weak: ∀L,U,l1,i. L ⊢ i ϵ 𝐅*[l1]⦃U⦄ →
+ ∀l2. l2 ≤ l1 → L ⊢ i ϵ 𝐅*[l2]⦃U⦄.
+#L #U #l1 #i #H elim H -L -U -l1 -i
/3 width=9 by frees_be, frees_eq, yle_trans/
qed-.
(* Properties on append for local environments ******************************)
-lemma frees_append: ∀L2,U,d,i. L2 ⊢ i ϵ 𝐅*[d]⦃U⦄ → i ≤ |L2| →
- ∀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
+lemma frees_append: ∀L2,U,l,i. L2 ⊢ i ϵ 𝐅*[l]⦃U⦄ → i ≤ |L2| →
+ ∀L1. L1 @@ L2 ⊢ i ϵ 𝐅*[l]⦃U⦄.
+#L2 #U #l #i #H elim H -L2 -U -l -i /3 width=2 by frees_eq/
+#I #L2 #K2 #U #W #l #i #j #Hlj #Hji #HnU #HLK2 #_ #IHW #Hi #L1
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/
+[ -I -L1 -K2 -U -W -l /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/
]
(* Inversion lemmas on append for local environments ************************)
-fact frees_inv_append_aux: ∀L,U,d,i. L ⊢ i ϵ 𝐅*[d]⦃U⦄ → ∀L1,L2. L = L1 @@ 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 (drop_O1_lt (Ⓕ) L2 j) [2: -Z -Y -L1 -X -U -d /2 width=3 by lt_to_le_to_lt/ ]
+fact frees_inv_append_aux: ∀L,U,l,i. L ⊢ i ϵ 𝐅*[l]⦃U⦄ → ∀L1,L2. L = L1 @@ L2 →
+ i ≤ |L2| → L2 ⊢ i ϵ 𝐅*[l]⦃U⦄.
+#L #U #l #i #H elim H -L -U -l -i /3 width=2 by frees_eq/
+#Z #L #Y #U #X #l #i #j #Hlj #Hji #HnU #HLY #_ #IHW #L1 #L2 #H #Hi destruct
+elim (drop_O1_lt (Ⓕ) L2 j) [2: -Z -Y -L1 -X -U -l /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/
+[ -Z -I -Y -K2 -L1 -X -U -W -l /3 width=3 by lt_to_le, lt_to_le_to_lt/
| normalize #H destruct
@(frees_be … HnU HLK2) -HnU -HLK2 // @IHW -IHW //
>(minus_plus_m_m (|K2|) 1) >H0 -H0 /2 width=1 by monotonic_le_minus_l2/
]
qed-.
-lemma frees_inv_append: ∀L1,L2,U,d,i. L1 @@ L2 ⊢ i ϵ 𝐅*[d]⦃U⦄ →
- i ≤ |L2| → L2 ⊢ i ϵ 𝐅*[d]⦃U⦄.
+lemma frees_inv_append: ∀L1,L2,U,l,i. L1 @@ L2 ⊢ i ϵ 𝐅*[l]⦃U⦄ →
+ i ≤ |L2| → L2 ⊢ i ϵ 𝐅*[l]⦃U⦄.
/2 width=4 by frees_inv_append_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/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⦄ →
- ∀L1. 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_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⦄ →
- ∀L2. L1 ⩬[d, ∞] L2 → L2 ⊢ i ϵ 𝐅*[d]⦃U⦄.
-/3 width=3 by leq_sym, leq_frees_trans/ qed-.
(* Advanced properties ******************************************************)
-lemma frees_dec: ∀L,U,d,i. Decidable (frees d L U i).
+lemma frees_dec: ∀L,U,l,i. Decidable (frees l L U i).
#L #U @(f2_ind … rfw … L U) -L -U
#n #IH #L * *
[ -IH /3 width=5 by frees_inv_sort, or_intror/
-| #j #Hn #d #i elim (lt_or_eq_or_gt i j) #Hji
+| #j #Hn #l #i elim (lt_or_eq_or_gt i j) #Hji
[ -n @or_intror #H elim (lt_refl_false i)
- lapply (frees_inv_lref_ge … H ?) -L -d /2 width=1 by lt_to_le/
+ lapply (frees_inv_lref_ge … H ?) -L -l /2 width=1 by lt_to_le/
| -n /2 width=1 by or_introl/
- | elim (ylt_split j d) #Hdi
+ | elim (ylt_split j l) #Hli
[ -n @or_intror #H elim (lt_refl_false i)
lapply (frees_inv_lref_skip … H ?) -L //
| elim (lt_or_ge j (|L|)) #Hj
[ 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
+ @or_intror #H elim (frees_inv_lref_lt … H) // #Z #Y #X #_ #HLY -l
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 //
+ lapply (frees_inv_lref_free … H ?) -l //
]
]
]
| -IH /3 width=5 by frees_inv_gref, or_intror/
-| #a #I #W #U #Hn #d #i destruct
- elim (IH L W … d i) [1,3: /3 width=1 by frees_bind_sn, or_introl/ ] #HnW
- elim (IH (L.ⓑ{I}W) U … (⫯d) (i+1)) -IH [1,3: /3 width=1 by frees_bind_dx, or_introl/ ] #HnU
+| #a #I #W #U #Hn #l #i destruct
+ elim (IH L W … l i) [1,3: /3 width=1 by frees_bind_sn, or_introl/ ] #HnW
+ elim (IH (L.ⓑ{I}W) U … (⫯l) (i+1)) -IH [1,3: /3 width=1 by frees_bind_dx, or_introl/ ] #HnU
@or_intror #H elim (frees_inv_bind … H) -H /2 width=1 by/
-| #I #W #U #Hn #d #i destruct
- elim (IH L W … d i) [1,3: /3 width=1 by frees_flat_sn, or_introl/ ] #HnW
- elim (IH L U … d i) -IH [1,3: /3 width=1 by frees_flat_dx, or_introl/ ] #HnU
+| #I #W #U #Hn #l #i destruct
+ elim (IH L W … l i) [1,3: /3 width=1 by frees_flat_sn, or_introl/ ] #HnW
+ elim (IH L U … l i) -IH [1,3: /3 width=1 by frees_flat_dx, or_introl/ ] #HnU
@or_intror #H elim (frees_inv_flat … H) -H /2 width=1 by/
]
qed-.
-lemma frees_S: ∀L,U,d,i. L ⊢ i ϵ 𝐅*[yinj d]⦃U⦄ → ∀I,K,W. ⬇[d] L ≡ K.ⓑ{I}W →
- (K ⊢ i-d-1 ϵ 𝐅*[0]⦃W⦄ → ⊥) → L ⊢ i ϵ 𝐅*[⫯d]⦃U⦄.
-#L #U #d #i #H elim (frees_inv … H) -H /3 width=2 by frees_eq/
-* #I #K #W #j #Hdj #Hji #HnU #HLK #HW #I0 #K0 #W0 #HLK0 #HnW0
-lapply (yle_inv_inj … Hdj) -Hdj #Hdj
-elim (le_to_or_lt_eq … Hdj) -Hdj
+lemma frees_S: ∀L,U,l,i. L ⊢ i ϵ 𝐅*[yinj l]⦃U⦄ → ∀I,K,W. ⬇[l] L ≡ K.ⓑ{I}W →
+ (K ⊢ i-l-1 ϵ 𝐅*[0]⦃W⦄ → ⊥) → L ⊢ i ϵ 𝐅*[⫯l]⦃U⦄.
+#L #U #l #i #H elim (frees_inv … H) -H /3 width=2 by frees_eq/
+* #I #K #W #j #Hlj #Hji #HnU #HLK #HW #I0 #K0 #W0 #HLK0 #HnW0
+lapply (yle_inv_inj … Hlj) -Hlj #Hlj
+elim (le_to_or_lt_eq … Hlj) -Hlj
[ -I0 -K0 -W0 /3 width=9 by frees_be, yle_inj/
| -Hji -HnU #H destruct
lapply (drop_mono … HLK0 … HLK) #H destruct -I
(* Properties on relocation *************************************************)
-lemma frees_lift_ge: ∀K,T,d,i. K ⊢ i ϵ𝐅*[d]⦃T⦄ →
- ∀L,s,d0,e0. ⬇[s, d0, e0] L ≡ K →
- ∀U. ⬆[d0, e0] T ≡ U → d0 ≤ i →
- L ⊢ i+e0 ϵ 𝐅*[d]⦃U⦄.
-#K #T #d #i #H elim H -K -T -d -i
-[ #K #T #d #i #HnT #L #s #d0 #e0 #_ #U #HTU #Hd0i -K -s
+lemma frees_lift_ge: ∀K,T,l,i. K ⊢ i ϵ𝐅*[l]⦃T⦄ →
+ ∀L,s,l0,m0. ⬇[s, l0, m0] L ≡ K →
+ ∀U. ⬆[l0, m0] T ≡ U → l0 ≤ i →
+ L ⊢ i+m0 ϵ 𝐅*[l]⦃U⦄.
+#K #T #l #i #H elim H -K -T -l -i
+[ #K #T #l #i #HnT #L #s #l0 #m0 #_ #U #HTU #Hl0i -K -s
@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
+| #I #K #K0 #T #V #l #i #j #Hlj #Hji #HnT #HK0 #HV #IHV #L #s #l0 #m0 #HLK #U #HTU #Hl0i
+ elim (lt_or_ge j l0) #H1
[ 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
+ [ #X #HXU >(plus_minus_m_m l0 1) in HTU; /2 width=2 by ltn_to_ltO/ #HTU
elim (lift_div_le … HXU … HTU ?) -U /2 width=2 by monotonic_pred/
| >minus_plus <plus_minus // <minus_plus
/3 width=5 by monotonic_le_minus_l2/
(* Inversion lemmas on relocation *******************************************)
-lemma frees_inv_lift_be: ∀L,U,d,i. L ⊢ i ϵ 𝐅*[d]⦃U⦄ →
- ∀K,s,d0,e0. ⬇[s, d0, e0+1] L ≡ K →
- ∀T. ⬆[d0, e0+1] T ≡ U → d0 ≤ i → i ≤ d0 + e0 → ⊥.
-#L #U #d #i #H elim H -L -U -d -i
-[ #L #U #d #i #HnU #K #s #d0 #e0 #_ #T #HTU #Hd0i #Hide0
- 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
+lemma frees_inv_lift_be: ∀L,U,l,i. L ⊢ i ϵ 𝐅*[l]⦃U⦄ →
+ ∀K,s,l0,m0. ⬇[s, l0, m0+1] L ≡ K →
+ ∀T. ⬆[l0, m0+1] T ≡ U → l0 ≤ i → i ≤ l0 + m0 → ⊥.
+#L #U #l #i #H elim H -L -U -l -i
+[ #L #U #l #i #HnU #K #s #l0 #m0 #_ #T #HTU #Hl0i #Hilm0
+ elim (lift_split … HTU i m0) -HTU /2 width=2 by/
+| #I #L #K0 #U #W #l #i #j #Hli #Hij #HnU #HLK0 #_ #IHW #K #s #l0 #m0 #HLK #T #HTU #Hl0i #Hilm0
+ elim (lt_or_ge j l0) #H1
[ 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_split … HTU j e0) -HTU /3 width=3 by lt_to_le_to_lt, lt_to_le/
+ | elim (lift_split … HTU j m0) -HTU /3 width=3 by lt_to_le_to_lt, lt_to_le/
]
]
qed-.
-lemma frees_inv_lift_ge: ∀L,U,d,i. L ⊢ i ϵ 𝐅*[d]⦃U⦄ →
- ∀K,s,d0,e0. ⬇[s, d0, e0] L ≡ K →
- ∀T. ⬆[d0, e0] T ≡ U → d0 + e0 ≤ i →
- K ⊢ i-e0 ϵ𝐅*[d-yinj e0]⦃T⦄.
-#L #U #d #i #H elim H -L -U -d -i
-[ #L #U #d #i #HnU #K #s #d0 #e0 #HLK #T #HTU #Hde0i -L -s
- elim (le_inv_plus_l … Hde0i) -Hde0i #Hd0ie0 #He0i @frees_eq #X #HXT -K
+lemma frees_inv_lift_ge: ∀L,U,l,i. L ⊢ i ϵ 𝐅*[l]⦃U⦄ →
+ ∀K,s,l0,m0. ⬇[s, l0, m0] L ≡ K →
+ ∀T. ⬆[l0, m0] T ≡ U → l0 + m0 ≤ i →
+ K ⊢ i-m0 ϵ𝐅*[l-yinj m0]⦃T⦄.
+#L #U #l #i #H elim H -L -U -l -i
+[ #L #U #l #i #HnU #K #s #l0 #m0 #HLK #T #HTU #Hlm0i -L -s
+ elim (le_inv_plus_l … Hlm0i) -Hlm0i #Hl0im0 #Hm0i @frees_eq #X #HXT -K
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
+| #I #L #K0 #U #W #l #i #j #Hli #Hij #HnU #HLK0 #_ #IHW #K #s #l0 #m0 #HLK #T #HTU #Hlm0i
+ elim (lt_or_ge j l0) #H1
[ elim (drop_conf_lt … HLK … HLK0) -L // #L0 #V #H #HKL0 #HVW
- elim (le_inv_plus_l … Hde0i) #H0 #He0i
+ elim (le_inv_plus_l … Hlm0i) #H0 #Hm0i
@(frees_be … H) -H
[ /3 width=1 by yle_plus_dx1_trans, monotonic_yle_minus_dx/
| /2 width=3 by lt_to_le_to_lt/
| lapply (IHW … HKL0 … HVW ?) // -I -K -K0 -L0 -V -W -T -U -s
>minus_plus >minus_plus >plus_minus /2 width=1 by monotonic_le_minus_l/
]
- | elim (lt_or_ge j (d0+e0)) [ >commutative_plus |] #H2
+ | elim (lt_or_ge j (l0+m0)) [ >commutative_plus |] #H2
[ -L -I -W lapply (lt_plus_to_minus ???? H2) // -H2 #H2
- elim (lift_split … HTU j (e0-1)) -HTU //
+ elim (lift_split … HTU j (m0-1)) -HTU //
[ >minus_minus_associative /2 width=2 by ltn_to_ltO/ <minus_n_n
#X #_ #H elim (HnU … H)
| >commutative_plus /3 width=1 by le_minus_to_plus, monotonic_pred/
]
| lapply (drop_conf_ge … HLK … HLK0 ?) // -L #HK0
- elim (le_inv_plus_l … H2) -H2 #H2 #He0j
+ elim (le_inv_plus_l … H2) -H2 #H2 #Hm0j
@(frees_be … HK0)
[ /2 width=1 by monotonic_yle_minus_dx/
| /2 width=1 by monotonic_lt_minus_l/
--- /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_lreq.ma".
+include "basic_2/multiple/frees.ma".
+
+(* CONTEXT-SENSITIVE FREE VARIABLES *****************************************)
+
+(* Properties on equivalence for local environments *************************)
+
+lemma lreq_frees_trans: ∀L2,U,l,i. L2 ⊢ i ϵ 𝐅*[l]⦃U⦄ →
+ ∀L1. L1 ⩬[l, ∞] L2 → L1 ⊢ i ϵ 𝐅*[l]⦃U⦄.
+#L2 #U #l #i #H elim H -L2 -U -l -i /3 width=2 by frees_eq/
+#I2 #L2 #K2 #U #W2 #l #i #j #Hlj #Hji #HnU #HLK2 #_ #IHW2 #L1 #HL12
+elim (lreq_drop_trans_be … HL12 … HLK2) -L2 // >yminus_Y_inj #K1 #HK12 #HLK1
+lapply (lreq_inv_O_Y … HK12) -HK12 #H destruct /3 width=9 by frees_be/
+qed-.
+
+lemma frees_lreq_conf: ∀L1,U,l,i. L1 ⊢ i ϵ 𝐅*[l]⦃U⦄ →
+ ∀L2. L1 ⩬[l, ∞] L2 → L2 ⊢ i ϵ 𝐅*[l]⦃U⦄.
+/3 width=3 by lreq_sym, lreq_frees_trans/ qed-.
inductive lifts: list2 nat nat → relation term ≝
| lifts_nil : ∀T. lifts (◊) T T
-| lifts_cons: ∀T1,T,T2,des,d,e.
- ⬆[d,e] T1 ≡ T → lifts des T T2 → lifts ({d, e} @ des) T1 T2
+| lifts_cons: ∀T1,T,T2,des,l,m.
+ ⬆[l,m] T1 ≡ T → lifts des T T2 → lifts ({l, m} @ des) T1 T2
.
interpretation "generic relocation (term)"
fact lifts_inv_nil_aux: ∀T1,T2,des. ⬆*[des] T1 ≡ T2 → des = ◊ → T1 = T2.
#T1 #T2 #des * -T1 -T2 -des //
-#T1 #T #T2 #d #e #des #_ #_ #H destruct
+#T1 #T #T2 #l #m #des #_ #_ #H destruct
qed-.
lemma lifts_inv_nil: ∀T1,T2. ⬆*[◊] T1 ≡ T2 → T1 = T2.
/2 width=3 by lifts_inv_nil_aux/ qed-.
fact lifts_inv_cons_aux: ∀T1,T2,des. ⬆*[des] T1 ≡ T2 →
- ∀d,e,tl. des = {d, e} @ tl →
- ∃∃T. ⬆[d, e] T1 ≡ T & ⬆*[tl] T ≡ T2.
+ ∀l,m,tl. des = {l, m} @ tl →
+ ∃∃T. ⬆[l, m] T1 ≡ T & ⬆*[tl] T ≡ T2.
#T1 #T2 #des * -T1 -T2 -des
-[ #T #d #e #tl #H destruct
-| #T1 #T #T2 #des #d #e #HT1 #HT2 #hd #he #tl #H destruct
+[ #T #l #m #tl #H destruct
+| #T1 #T #T2 #des #l #m #HT1 #HT2 #l0 #m0 #tl #H destruct
/2 width=3 by ex2_intro/
qed-.
-lemma lifts_inv_cons: ∀T1,T2,d,e,des. ⬆*[{d, e} @ des] T1 ≡ T2 →
- ∃∃T. ⬆[d, e] T1 ≡ T & ⬆*[des] T ≡ T2.
+lemma lifts_inv_cons: ∀T1,T2,l,m,des. ⬆*[{l, m} @ des] T1 ≡ T2 →
+ ∃∃T. ⬆[l, m] T1 ≡ T & ⬆*[des] T ≡ T2.
/2 width=3 by lifts_inv_cons_aux/ qed-.
(* Basic_1: was: lift1_sort *)
lemma lifts_inv_sort1: ∀T2,k,des. ⬆*[des] ⋆k ≡ T2 → T2 = ⋆k.
#T2 #k #des elim des -des
[ #H <(lifts_inv_nil … H) -H //
-| #d #e #des #IH #H
+| #l #m #des #IH #H
elim (lifts_inv_cons … H) -H #X #H
>(lift_inv_sort1 … H) -H /2 width=1 by/
]
∃∃i2. @⦃i1, des⦄ ≡ i2 & T2 = #i2.
#T2 #des elim des -des
[ #i1 #H <(lifts_inv_nil … H) -H /2 width=3 by at_nil, ex2_intro/
-| #d #e #des #IH #i1 #H
+| #l #m #des #IH #i1 #H
elim (lifts_inv_cons … H) -H #X #H1 #H2
- elim (lift_inv_lref1 … H1) -H1 * #Hdi1 #H destruct
+ elim (lift_inv_lref1 … H1) -H1 * #Hli1 #H destruct
elim (IH … H2) -IH -H2 /3 width=3 by at_lt, at_ge, ex2_intro/
]
qed-.
lemma lifts_inv_gref1: ∀T2,p,des. ⬆*[des] §p ≡ T2 → T2 = §p.
#T2 #p #des elim des -des
[ #H <(lifts_inv_nil … H) -H //
-| #d #e #des #IH #H
+| #l #m #des #IH #H
elim (lifts_inv_cons … H) -H #X #H
>(lift_inv_gref1 … H) -H /2 width=1 by/
]
#a #I #T2 #des elim des -des
[ #V1 #U1 #H
<(lifts_inv_nil … H) -H /2 width=5 by ex3_2_intro, lifts_nil/
-| #d #e #des #IHdes #V1 #U1 #H
+| #l #m #des #IHcs #V1 #U1 #H
elim (lifts_inv_cons … H) -H #X #H #HT2
elim (lift_inv_bind1 … H) -H #V #U #HV1 #HU1 #H destruct
- elim (IHdes … HT2) -IHdes -HT2 #V2 #U2 #HV2 #HU2 #H destruct
+ elim (IHcs … HT2) -IHcs -HT2 #V2 #U2 #HV2 #HU2 #H destruct
/3 width=5 by ex3_2_intro, lifts_cons/
]
qed-.
#I #T2 #des elim des -des
[ #V1 #U1 #H
<(lifts_inv_nil … H) -H /2 width=5 by ex3_2_intro, lifts_nil/
-| #d #e #des #IHdes #V1 #U1 #H
+| #l #m #des #IHcs #V1 #U1 #H
elim (lifts_inv_cons … H) -H #X #H #HT2
elim (lift_inv_flat1 … H) -H #V #U #HV1 #HU1 #H destruct
- elim (IHdes … HT2) -IHdes -HT2 #V2 #U2 #HV2 #HU2 #H destruct
+ elim (IHcs … HT2) -IHcs -HT2 #V2 #U2 #HV2 #HU2 #H destruct
/3 width=5 by ex3_2_intro, lifts_cons/
]
qed-.
⬆*[des] ⓑ{a,I} V1. T1 ≡ ⓑ{a,I} V2. T2.
#a #I #T2 #V1 #V2 #des #H elim H -V1 -V2 -des
[ #V #T1 #H >(lifts_inv_nil … H) -H //
-| #V1 #V #V2 #des #d #e #HV1 #_ #IHV #T1 #H
+| #V1 #V #V2 #des #l #m #HV1 #_ #IHV #T1 #H
elim (lifts_inv_cons … H) -H /3 width=3 by lift_bind, lifts_cons/
]
qed.
⬆*[des] ⓕ{I} V1. T1 ≡ ⓕ{I} V2. T2.
#I #T2 #V1 #V2 #des #H elim H -V1 -V2 -des
[ #V #T1 #H >(lifts_inv_nil … H) -H //
-| #V1 #V #V2 #des #d #e #HV1 #_ #IHV #T1 #H
+| #V1 #V #V2 #des #l #m #HV1 #_ #IHV #T1 #H
elim (lifts_inv_cons … H) -H /3 width=3 by lift_flat, lifts_cons/
]
qed.
lemma lifts_total: ∀des,T1. ∃T2. ⬆*[des] T1 ≡ T2.
#des elim des -des /2 width=2 by lifts_nil, ex_intro/
-#d #e #des #IH #T1 elim (lift_total T1 d e)
+#l #m #des #IH #T1 elim (lift_total T1 l m)
#T #HT1 elim (IH T) -IH /3 width=4 by lifts_cons, ex_intro/
qed.
∃∃T0. ⬆[0, 1] T1 ≡ T0 & ⬆*[des + 1] T0 ≡ T2.
#T1 #T #des #H elim H -T1 -T -des
[ /2 width=3/
-| #T1 #T3 #T #des #d #e #HT13 #_ #IHT13 #T2 #HT2
+| #T1 #T3 #T #des #l #m #HT13 #_ #IHT13 #T2 #HT2
elim (IHT13 … HT2) -T #T #HT3 #HT2
elim (lift_trans_le … HT13 … HT3) -T3 // /3 width=5/
]
lapply (minuss_inv_nil1 … H2) -H2 #H
>(pluss_inv_nil2 … H) in HT10; -des0 #H
>(lifts_inv_nil … H) -T1 /2 width=3/
-| #d #e #des #IHdes #i #i0 #H1 #des0 #H2 #T1 #T0 #HT10 #T2 #HT02
- elim (at_inv_cons … H1) -H1 * #Hid #Hi0
- [ elim (minuss_inv_cons1_lt … H2) -H2 [2: /2 width=1/ ] #des1 #Hdes1 <minus_le_minus_minus_comm // <minus_plus_m_m #H
+| #l #m #des #IHcs #i #i0 #H1 #des0 #H2 #T1 #T0 #HT10 #T2 #HT02
+ elim (at_inv_cons … H1) -H1 * #Hil #Hi0
+ [ elim (minuss_inv_cons1_lt … H2) -H2 [2: /2 width=1/ ] #des1 #Hcs1 <minus_le_minus_minus_comm // <minus_plus_m_m #H
elim (pluss_inv_cons2 … H) -H #des2 #H1 #H2 destruct
elim (lifts_inv_cons … HT10) -HT10 #T >minus_plus #HT1 #HT0
- elim (IHdes … Hi0 … Hdes1 … HT0 … HT02) -IHdes -Hi0 -Hdes1 -T0 #T0 #HT0 #HT02
+ elim (IHcs … Hi0 … Hcs1 … HT0 … HT02) -IHcs -Hi0 -Hcs1 -T0 #T0 #HT0 #HT02
elim (lift_trans_le … HT1 … HT0) -T /2 width=1/ #T #HT1 <plus_minus_m_m /2 width=1/ /3 width=5/
| >commutative_plus in Hi0; #Hi0
- lapply (minuss_inv_cons1_ge … H2 ?) -H2 [ /2 width=1/ ] <associative_plus #Hdes0
- elim (IHdes … Hi0 … Hdes0 … HT10 … HT02) -IHdes -Hi0 -Hdes0 -T0 #T0 #HT0 #HT02
- elim (lift_split … HT0 d (i+1)) -HT0 /2 width=1/ /3 width=5/
+ lapply (minuss_inv_cons1_ge … H2 ?) -H2 [ /2 width=1/ ] <associative_plus #Hcs0
+ elim (IHcs … Hi0 … Hcs0 … HT10 … HT02) -IHcs -Hi0 -Hcs0 -T0 #T0 #HT0 #HT02
+ elim (lift_split … HT0 l (i+1)) -HT0 /2 width=1/ /3 width=5/
]
]
qed-.
interpretation
"lazy equivalence (local environment)"
- 'LazyEq T d L1 L2 = (lleq d T L1 L2).
+ 'LazyEq T l L1 L2 = (lleq l T L1 L2).
definition lleq_transitive: predicate (relation3 lenv term term) ≝
λR. ∀L2,T1,T2. R L2 T1 T2 → ∀L1. L1 ≡[T1, 0] L2 → R L1 T1 T2.
(* Basic inversion lemmas ***************************************************)
lemma lleq_ind: ∀R:relation4 ynat term lenv lenv. (
- ∀L1,L2,d,k. |L1| = |L2| → R d (⋆k) L1 L2
+ ∀L1,L2,l,k. |L1| = |L2| → R l (⋆k) L1 L2
) → (
- ∀L1,L2,d,i. |L1| = |L2| → yinj i < d → R d (#i) L1 L2
+ ∀L1,L2,l,i. |L1| = |L2| → yinj i < l → R l (#i) L1 L2
) → (
- ∀I,L1,L2,K1,K2,V,d,i. d ≤ yinj i →
+ ∀I,L1,L2,K1,K2,V,l,i. l ≤ yinj i →
⬇[i] L1 ≡ K1.ⓑ{I}V → ⬇[i] L2 ≡ K2.ⓑ{I}V →
- K1 ≡[V, yinj O] K2 → R (yinj O) V K1 K2 → R d (#i) L1 L2
+ K1 ≡[V, yinj O] K2 → R (yinj O) V K1 K2 → R l (#i) L1 L2
) → (
- ∀L1,L2,d,i. |L1| = |L2| → |L1| ≤ i → |L2| ≤ i → R d (#i) L1 L2
+ ∀L1,L2,l,i. |L1| = |L2| → |L1| ≤ i → |L2| ≤ i → R l (#i) L1 L2
) → (
- ∀L1,L2,d,p. |L1| = |L2| → R d (§p) L1 L2
+ ∀L1,L2,l,p. |L1| = |L2| → R l (§p) L1 L2
) → (
- ∀a,I,L1,L2,V,T,d.
- L1 ≡[V, d]L2 → L1.ⓑ{I}V ≡[T, ⫯d] L2.ⓑ{I}V →
- R d V L1 L2 → R (⫯d) T (L1.ⓑ{I}V) (L2.ⓑ{I}V) → R d (ⓑ{a,I}V.T) L1 L2
+ ∀a,I,L1,L2,V,T,l.
+ L1 ≡[V, l]L2 → L1.ⓑ{I}V ≡[T, ⫯l] L2.ⓑ{I}V →
+ R l V L1 L2 → R (⫯l) T (L1.ⓑ{I}V) (L2.ⓑ{I}V) → R l (ⓑ{a,I}V.T) L1 L2
) → (
- ∀I,L1,L2,V,T,d.
- L1 ≡[V, d]L2 → L1 ≡[T, d] L2 →
- R d V L1 L2 → R d T L1 L2 → R d (ⓕ{I}V.T) L1 L2
+ ∀I,L1,L2,V,T,l.
+ L1 ≡[V, l]L2 → L1 ≡[T, l] L2 →
+ R l V L1 L2 → R l T L1 L2 → R l (ⓕ{I}V.T) L1 L2
) →
- ∀d,T,L1,L2. L1 ≡[T, d] L2 → R d T L1 L2.
-#R #H1 #H2 #H3 #H4 #H5 #H6 #H7 #d #T #L1 #L2 #H elim H -L1 -L2 -T -d /2 width=8 by/
+ ∀l,T,L1,L2. L1 ≡[T, l] L2 → R l T L1 L2.
+#R #H1 #H2 #H3 #H4 #H5 #H6 #H7 #l #T #L1 #L2 #H elim H -L1 -L2 -T -l /2 width=8 by/
qed-.
-lemma lleq_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.
+lemma lleq_inv_bind: ∀a,I,L1,L2,V,T,l. L1 ≡[ⓑ{a,I}V.T, l] L2 →
+ L1 ≡[V, l] L2 ∧ L1.ⓑ{I}V ≡[T, ⫯l] L2.ⓑ{I}V.
/2 width=2 by llpx_sn_inv_bind/ qed-.
-lemma lleq_inv_flat: ∀I,L1,L2,V,T,d. L1 ≡[ⓕ{I}V.T, d] L2 →
- L1 ≡[V, d] L2 ∧ L1 ≡[T, d] L2.
+lemma lleq_inv_flat: ∀I,L1,L2,V,T,l. L1 ≡[ⓕ{I}V.T, l] L2 →
+ L1 ≡[V, l] L2 ∧ L1 ≡[T, l] L2.
/2 width=2 by llpx_sn_inv_flat/ qed-.
(* Basic forward lemmas *****************************************************)
-lemma lleq_fwd_length: ∀L1,L2,T,d. L1 ≡[T, d] L2 → |L1| = |L2|.
+lemma lleq_fwd_length: ∀L1,L2,T,l. L1 ≡[T, l] L2 → |L1| = |L2|.
/2 width=4 by llpx_sn_fwd_length/ qed-.
-lemma lleq_fwd_lref: ∀L1,L2,d,i. L1 ≡[#i, d] L2 →
+lemma lleq_fwd_lref: ∀L1,L2,l,i. L1 ≡[#i, l] L2 →
∨∨ |L1| ≤ i ∧ |L2| ≤ i
- | yinj i < d
+ | yinj i < l
| ∃∃I,K1,K2,V. ⬇[i] L1 ≡ K1.ⓑ{I}V &
⬇[i] L2 ≡ K2.ⓑ{I}V &
- K1 ≡[V, yinj 0] K2 & d ≤ yinj i.
-#L1 #L2 #d #i #H elim (llpx_sn_fwd_lref … H) /2 width=1/
+ K1 ≡[V, yinj 0] K2 & l ≤ yinj i.
+#L1 #L2 #l #i #H elim (llpx_sn_fwd_lref … H) /2 width=1/
* /3 width=7 by or3_intro2, ex4_4_intro/
qed-.
-lemma lleq_fwd_drop_sn: ∀L1,L2,T,d. L1 ≡[d, T] L2 → ∀K1,i. ⬇[i] L1 ≡ K1 →
+lemma lleq_fwd_drop_sn: ∀L1,L2,T,l. L1 ≡[l, T] L2 → ∀K1,i. ⬇[i] L1 ≡ K1 →
∃K2. ⬇[i] L2 ≡ K2.
/2 width=7 by llpx_sn_fwd_drop_sn/ qed-.
-lemma lleq_fwd_drop_dx: ∀L1,L2,T,d. L1 ≡[d, T] L2 → ∀K2,i. ⬇[i] L2 ≡ K2 →
+lemma lleq_fwd_drop_dx: ∀L1,L2,T,l. L1 ≡[l, T] L2 → ∀K2,i. ⬇[i] L2 ≡ K2 →
∃K1. ⬇[i] L1 ≡ K1.
/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.
+lemma lleq_fwd_bind_sn: ∀a,I,L1,L2,V,T,l.
+ L1 ≡[ⓑ{a,I}V.T, l] L2 → L1 ≡[V, l] L2.
/2 width=4 by llpx_sn_fwd_bind_sn/ qed-.
-lemma lleq_fwd_bind_dx: ∀a,I,L1,L2,V,T,d.
- L1 ≡[ⓑ{a,I}V.T, d] L2 → L1.ⓑ{I}V ≡[T, ⫯d] L2.ⓑ{I}V.
+lemma lleq_fwd_bind_dx: ∀a,I,L1,L2,V,T,l.
+ L1 ≡[ⓑ{a,I}V.T, l] L2 → L1.ⓑ{I}V ≡[T, ⫯l] L2.ⓑ{I}V.
/2 width=2 by llpx_sn_fwd_bind_dx/ qed-.
-lemma lleq_fwd_flat_sn: ∀I,L1,L2,V,T,d.
- L1 ≡[ⓕ{I}V.T, d] L2 → L1 ≡[V, d] L2.
+lemma lleq_fwd_flat_sn: ∀I,L1,L2,V,T,l.
+ L1 ≡[ⓕ{I}V.T, l] L2 → L1 ≡[V, l] L2.
/2 width=3 by llpx_sn_fwd_flat_sn/ qed-.
-lemma lleq_fwd_flat_dx: ∀I,L1,L2,V,T,d.
- L1 ≡[ⓕ{I}V.T, d] L2 → L1 ≡[T, d] L2.
+lemma lleq_fwd_flat_dx: ∀I,L1,L2,V,T,l.
+ L1 ≡[ⓕ{I}V.T, l] L2 → L1 ≡[T, l] L2.
/2 width=3 by llpx_sn_fwd_flat_dx/ qed-.
(* Basic properties *********************************************************)
-lemma lleq_sort: ∀L1,L2,d,k. |L1| = |L2| → L1 ≡[⋆k, d] L2.
+lemma lleq_sort: ∀L1,L2,l,k. |L1| = |L2| → L1 ≡[⋆k, l] L2.
/2 width=1 by llpx_sn_sort/ qed.
-lemma lleq_skip: ∀L1,L2,d,i. yinj i < d → |L1| = |L2| → L1 ≡[#i, d] L2.
+lemma lleq_skip: ∀L1,L2,l,i. yinj i < l → |L1| = |L2| → L1 ≡[#i, l] L2.
/2 width=1 by llpx_sn_skip/ qed.
-lemma lleq_lref: ∀I,L1,L2,K1,K2,V,d,i. d ≤ yinj i →
+lemma lleq_lref: ∀I,L1,L2,K1,K2,V,l,i. l ≤ yinj i →
⬇[i] L1 ≡ K1.ⓑ{I}V → ⬇[i] L2 ≡ K2.ⓑ{I}V →
- K1 ≡[V, 0] K2 → L1 ≡[#i, d] L2.
+ K1 ≡[V, 0] K2 → L1 ≡[#i, l] L2.
/2 width=9 by llpx_sn_lref/ qed.
-lemma lleq_free: ∀L1,L2,d,i. |L1| ≤ i → |L2| ≤ i → |L1| = |L2| → L1 ≡[#i, d] L2.
+lemma lleq_free: ∀L1,L2,l,i. |L1| ≤ i → |L2| ≤ i → |L1| = |L2| → L1 ≡[#i, l] L2.
/2 width=1 by llpx_sn_free/ qed.
-lemma lleq_gref: ∀L1,L2,d,p. |L1| = |L2| → L1 ≡[§p, d] L2.
+lemma lleq_gref: ∀L1,L2,l,p. |L1| = |L2| → L1 ≡[§p, l] L2.
/2 width=1 by llpx_sn_gref/ qed.
-lemma lleq_bind: ∀a,I,L1,L2,V,T,d.
- L1 ≡[V, d] L2 → L1.ⓑ{I}V ≡[T, ⫯d] L2.ⓑ{I}V →
- L1 ≡[ⓑ{a,I}V.T, d] L2.
+lemma lleq_bind: ∀a,I,L1,L2,V,T,l.
+ L1 ≡[V, l] L2 → L1.ⓑ{I}V ≡[T, ⫯l] L2.ⓑ{I}V →
+ L1 ≡[ⓑ{a,I}V.T, l] L2.
/2 width=1 by llpx_sn_bind/ qed.
-lemma lleq_flat: ∀I,L1,L2,V,T,d.
- L1 ≡[V, d] L2 → L1 ≡[T, d] L2 → L1 ≡[ⓕ{I}V.T, d] L2.
+lemma lleq_flat: ∀I,L1,L2,V,T,l.
+ L1 ≡[V, l] L2 → L1 ≡[T, l] L2 → L1 ≡[ⓕ{I}V.T, l] L2.
/2 width=1 by llpx_sn_flat/ qed.
-lemma lleq_refl: ∀d,T. reflexive … (lleq d T).
+lemma lleq_refl: ∀l,T. reflexive … (lleq l T).
/2 width=1 by llpx_sn_refl/ qed.
lemma lleq_Y: ∀L1,L2,T. |L1| = |L2| → L1 ≡[T, ∞] L2.
/2 width=1 by llpx_sn_Y/ qed.
-lemma lleq_sym: ∀d,T. symmetric … (lleq d T).
-#d #T #L1 #L2 #H @(lleq_ind … H) -d -T -L1 -L2
+lemma lleq_sym: ∀l,T. symmetric … (lleq l T).
+#l #T #L1 #L2 #H @(lleq_ind … H) -l -T -L1 -L2
/2 width=7 by lleq_sort, lleq_skip, lleq_lref, lleq_free, lleq_gref, lleq_bind, lleq_flat/
qed-.
-lemma lleq_ge_up: ∀L1,L2,U,dt. L1 ≡[U, dt] L2 →
- ∀T,d,e. ⬆[d, e] T ≡ U →
- dt ≤ d + e → L1 ≡[U, d] L2.
+lemma lleq_ge_up: ∀L1,L2,U,lt. L1 ≡[U, lt] L2 →
+ ∀T,l,m. ⬆[l, m] T ≡ U →
+ lt ≤ l + m → L1 ≡[U, l] L2.
/2 width=6 by llpx_sn_ge_up/ qed-.
-lemma lleq_ge: ∀L1,L2,T,d1. L1 ≡[T, d1] L2 → ∀d2. d1 ≤ d2 → L1 ≡[T, d2] L2.
+lemma lleq_ge: ∀L1,L2,T,l1. L1 ≡[T, l1] L2 → ∀l2. l1 ≤ l2 → L1 ≡[T, l2] L2.
/2 width=3 by llpx_sn_ge/ qed-.
lemma lleq_bind_O: ∀a,I,L1,L2,V,T. L1 ≡[V, 0] L2 → L1.ⓑ{I}V ≡[T, 0] L2.ⓑ{I}V →
(* Advanceded properties on lazy pointwise extensions ************************)
lemma llpx_sn_lrefl: ∀R. (∀L. reflexive … (R L)) →
- ∀L1,L2,T,d. L1 ≡[T, d] L2 → llpx_sn R d T L1 L2.
+ ∀L1,L2,T,l. L1 ≡[T, l] L2 → llpx_sn R l T L1 L2.
/2 width=3 by llpx_sn_co/ qed-.
(* Alternative definition (not recursive) ***********************************)
-theorem lleq_intro_alt: ∀L1,L2,T,d. |L1| = |L2| →
- (∀I1,I2,K1,K2,V1,V2,i. d ≤ yinj i → L1 ⊢ i ϵ 𝐅*[d]⦃T⦄ →
+theorem lleq_intro_alt: ∀L1,L2,T,l. |L1| = |L2| →
+ (∀I1,I2,K1,K2,V1,V2,i. l ≤ yinj i → L1 ⊢ i ϵ 𝐅*[l]⦃T⦄ →
⬇[i] L1 ≡ K1.ⓑ{I1}V1 → ⬇[i] L2 ≡ K2.ⓑ{I2}V2 →
I1 = I2 ∧ V1 = V2
- ) → L1 ≡[T, d] L2.
-#L1 #L2 #T #d #HL12 #IH @llpx_sn_alt_inv_llpx_sn @conj // -HL12
-#I1 #I2 #K1 #K2 #V1 #V2 #i #Hid #HnT #HLK1 #HLK2
+ ) → L1 ≡[T, l] L2.
+#L1 #L2 #T #l #HL12 #IH @llpx_sn_alt_inv_llpx_sn @conj // -HL12
+#I1 #I2 #K1 #K2 #V1 #V2 #i #Hil #HnT #HLK1 #HLK2
@(IH … HnT HLK1 HLK2) -IH -HnT -HLK1 -HLK2 //
qed.
-theorem lleq_inv_alt: ∀L1,L2,T,d. L1 ≡[T, d] L2 →
+theorem lleq_inv_alt: ∀L1,L2,T,l. L1 ≡[T, l] L2 →
|L1| = |L2| ∧
- ∀I1,I2,K1,K2,V1,V2,i. d ≤ yinj i → L1 ⊢ i ϵ 𝐅*[d]⦃T⦄ →
+ ∀I1,I2,K1,K2,V1,V2,i. l ≤ yinj i → L1 ⊢ i ϵ 𝐅*[l]⦃T⦄ →
⬇[i] L1 ≡ K1.ⓑ{I1}V1 → ⬇[i] L2 ≡ K2.ⓑ{I2}V2 →
I1 = I2 ∧ V1 = V2.
-#L1 #L2 #T #d #H elim (llpx_sn_llpx_sn_alt … H) -H
+#L1 #L2 #T #l #H elim (llpx_sn_llpx_sn_alt … H) -H
#HL12 #IH @conj //
-#I1 #I2 #K1 #K2 #V1 #V2 #i #Hid #HnT #HLK1 #HLK2
+#I1 #I2 #K1 #K2 #V1 #V2 #i #Hil #HnT #HLK1 #HLK2
@(IH … HnT HLK1 HLK2) -IH -HnT -HLK1 -HLK2 //
qed-.
(* Alternative definition (recursive) ***************************************)
-theorem lleq_intro_alt_r: ∀L1,L2,T,d. |L1| = |L2| →
- (∀I1,I2,K1,K2,V1,V2,i. d ≤ yinj i → (∀U. ⬆[i, 1] U ≡ T → ⊥) →
+theorem lleq_intro_alt_r: ∀L1,L2,T,l. |L1| = |L2| →
+ (∀I1,I2,K1,K2,V1,V2,i. l ≤ yinj i → (∀U. ⬆[i, 1] U ≡ T → ⊥) →
⬇[i] L1 ≡ K1.ⓑ{I1}V1 → ⬇[i] L2 ≡ K2.ⓑ{I2}V2 →
∧∧ I1 = I2 & V1 = V2 & K1 ≡[V1, 0] K2
- ) → L1 ≡[T, d] L2.
-#L1 #L2 #T #d #HL12 #IH @llpx_sn_intro_alt_r // -HL12
-#I1 #I2 #K1 #K2 #V1 #V2 #i #Hid #HnT #HLK1 #HLK2
+ ) → L1 ≡[T, l] L2.
+#L1 #L2 #T #l #HL12 #IH @llpx_sn_intro_alt_r // -HL12
+#I1 #I2 #K1 #K2 #V1 #V2 #i #Hil #HnT #HLK1 #HLK2
elim (IH … HnT HLK1 HLK2) -IH -HnT -HLK1 -HLK2 /2 width=1 by and3_intro/
qed.
theorem lleq_ind_alt_r: ∀S:relation4 ynat term lenv lenv.
- (∀L1,L2,T,d. |L1| = |L2| → (
- ∀I1,I2,K1,K2,V1,V2,i. d ≤ yinj i → (∀U. ⬆[i, 1] U ≡ T → ⊥) →
+ (∀L1,L2,T,l. |L1| = |L2| → (
+ ∀I1,I2,K1,K2,V1,V2,i. l ≤ yinj i → (∀U. ⬆[i, 1] U ≡ T → ⊥) →
⬇[i] L1 ≡ K1.ⓑ{I1}V1 → ⬇[i] L2 ≡ K2.ⓑ{I2}V2 →
∧∧ I1 = I2 & V1 = V2 & K1 ≡[V1, 0] K2 & S 0 V1 K1 K2
- ) → S d T L1 L2) →
- ∀L1,L2,T,d. L1 ≡[T, d] L2 → S d T L1 L2.
-#S #IH1 #L1 #L2 #T #d #H @(llpx_sn_ind_alt_r … H) -L1 -L2 -T -d
-#L1 #L2 #T #d #HL12 #IH2 @IH1 -IH1 // -HL12
-#I1 #I2 #K1 #K2 #V1 #V2 #i #Hid #HnT #HLK1 #HLK2
+ ) → S l T L1 L2) →
+ ∀L1,L2,T,l. L1 ≡[T, l] L2 → S l T L1 L2.
+#S #IH1 #L1 #L2 #T #l #H @(llpx_sn_ind_alt_r … H) -L1 -L2 -T -l
+#L1 #L2 #T #l #HL12 #IH2 @IH1 -IH1 // -HL12
+#I1 #I2 #K1 #K2 #V1 #V2 #i #Hil #HnT #HLK1 #HLK2
elim (IH2 … HnT HLK1 HLK2) -IH2 -HnT -HLK1 -HLK2 /2 width=1 by and4_intro/
qed-.
-theorem lleq_inv_alt_r: ∀L1,L2,T,d. L1 ≡[T, d] L2 →
+theorem lleq_inv_alt_r: ∀L1,L2,T,l. L1 ≡[T, l] L2 →
|L1| = |L2| ∧
- ∀I1,I2,K1,K2,V1,V2,i. d ≤ yinj i → (∀U. ⬆[i, 1] U ≡ T → ⊥) →
+ ∀I1,I2,K1,K2,V1,V2,i. l ≤ yinj i → (∀U. ⬆[i, 1] U ≡ T → ⊥) →
⬇[i] L1 ≡ K1.ⓑ{I1}V1 → ⬇[i] L2 ≡ K2.ⓑ{I2}V2 →
∧∧ I1 = I2 & V1 = V2 & K1 ≡[V1, 0] K2.
-#L1 #L2 #T #d #H elim (llpx_sn_inv_alt_r … H) -H
+#L1 #L2 #T #l #H elim (llpx_sn_inv_alt_r … H) -H
#HL12 #IH @conj //
-#I1 #I2 #K1 #K2 #V1 #V2 #i #Hid #HnT #HLK1 #HLK2
+#I1 #I2 #K1 #K2 #V1 #V2 #i #Hil #HnT #HLK1 #HLK2
elim (IH … HnT HLK1 HLK2) -IH -HnT -HLK1 -HLK2 /2 width=1 by and3_intro/
qed-.
∀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).
+lemma lleq_dec: ∀T,L1,L2,l. Decidable (L1 ≡[T, l] 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
+ ∀L1,L2,T,l. L1 ≡[T, l] L2 →
+ ∀L. llpx_sn R l T L2 L → llpx_sn R l T L1 L.
+#R #HR #L1 #L2 #T #l #H @(lleq_ind … H) -L1 -L2 -T -l
[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
+| #I #L1 #L2 #K1 #K2 #V #l #i #Hli #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
+| #a #I #L1 #L2 #V #T #l #_ #_ #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
+| #I #L1 #L2 #V #T #l #_ #_ #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.
+ ∀L1,L2,T,l. L1 ≡[T, l] L2 →
+ ∀L. llpx_sn R l T L1 L → llpx_sn R l 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 →
+lemma lleq_inv_lref_ge_dx: ∀L1,L2,l,i. L1 ≡[#i, l] L2 → l ≤ 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
+#L1 #L2 #l #i #H #Hli #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 →
+lemma lleq_inv_lref_ge_sn: ∀L1,L2,l,i. L1 ≡[#i, l] L2 → l ≤ 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
+#L1 #L2 #l #i #H #Hli #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 →
+lemma lleq_inv_lref_ge_bi: ∀L1,L2,l,i. L1 ≡[#i, l] L2 → l ≤ 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 →
+lemma lleq_inv_lref_ge: ∀L1,L2,l,i. L1 ≡[#i, l] L2 → l ≤ 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
+#L1 #L2 #l #i #HL12 #Hli #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.
+lemma lleq_inv_S: ∀L1,L2,T,l. L1 ≡[T, l+1] L2 →
+ ∀I,K1,K2,V. ⬇[l] L1 ≡ K1.ⓑ{I}V → ⬇[l] L2 ≡ K2.ⓑ{I}V →
+ K1 ≡[V, 0] K2 → L1 ≡[T, l] 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 →
(* Advanced forward lemmas **************************************************)
-lemma lleq_fwd_lref_dx: ∀L1,L2,d,i. L1 ≡[#i, d] L2 →
+lemma lleq_fwd_lref_dx: ∀L1,L2,l,i. L1 ≡[#i, l] 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
+ i < l ∨
+ ∃∃K1. ⬇[i] L1 ≡ K1.ⓑ{I}V & K1 ≡[V, 0] K2 & l ≤ i.
+#L1 #L2 #l #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 →
+lemma lleq_fwd_lref_sn: ∀L1,L2,l,i. L1 ≡[#i, l] 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
+ i < l ∨
+ ∃∃K2. ⬇[i] L2 ≡ K2.ⓑ{I}V & K1 ≡[V, 0] K2 & l ≤ i.
+#L1 #L2 #l #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-.
(* 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.
+lemma lleq_lift_le: ∀K1,K2,T,lt. K1 ≡[T, lt] K2 →
+ ∀L1,L2,l,m. ⬇[Ⓕ, l, m] L1 ≡ K1 → ⬇[Ⓕ, l, m] L2 ≡ K2 →
+ ∀U. ⬆[l, m] T ≡ U → lt ≤ l → L1 ≡[U, lt] 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.
+lemma lleq_lift_ge: ∀K1,K2,T,lt. K1 ≡[T, lt] K2 →
+ ∀L1,L2,l,m. ⬇[Ⓕ, l, m] L1 ≡ K1 → ⬇[Ⓕ, l, m] L2 ≡ K2 →
+ ∀U. ⬆[l, m] T ≡ U → l ≤ lt → L1 ≡[U, lt+m] 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.
+lemma lleq_inv_lift_le: ∀L1,L2,U,lt. L1 ≡[U, lt] L2 →
+ ∀K1,K2,l,m. ⬇[Ⓕ, l, m] L1 ≡ K1 → ⬇[Ⓕ, l, m] L2 ≡ K2 →
+ ∀T. ⬆[l, m] T ≡ U → lt ≤ l → K1 ≡[T, lt] 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.
+lemma lleq_inv_lift_be: ∀L1,L2,U,lt. L1 ≡[U, lt] L2 →
+ ∀K1,K2,l,m. ⬇[Ⓕ, l, m] L1 ≡ K1 → ⬇[Ⓕ, l, m] L2 ≡ K2 →
+ ∀T. ⬆[l, m] T ≡ U → l ≤ lt → lt ≤ yinj l + m → K1 ≡[T, l] 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.
+lemma lleq_inv_lift_ge: ∀L1,L2,U,lt. L1 ≡[U, lt] L2 →
+ ∀K1,K2,l,m. ⬇[Ⓕ, l, m] L1 ≡ K1 → ⬇[Ⓕ, l, m] L2 ≡ K2 →
+ ∀T. ⬆[l, m] T ≡ U → yinj l + m ≤ lt → K1 ≡[T, lt-m] 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 → ⊥).
+lemma nlleq_inv_bind: ∀a,I,L1,L2,V,T,l. (L1 ≡[ⓑ{a,I}V.T, l] L2 → ⊥) →
+ (L1 ≡[V, l] L2 → ⊥) ∨ (L1.ⓑ{I}V ≡[T, ⫯l] 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 → ⊥).
+lemma nlleq_inv_flat: ∀I,L1,L2,V,T,l. (L1 ≡[ⓕ{I}V.T, l] L2 → ⊥) →
+ (L1 ≡[V, l] L2 → ⊥) ∨ (L1 ≡[T, l] 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 → ⊥) →
#H3 #H4 /2 width=3 by fqu_bind_dx, ex2_intro/
| #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 (drop_O1_le (Ⓕ) (e+1) L1)
+| #G #L2 #K2 #T #U #m #HLK2 #HTU #L1 #HL12
+ elim (drop_O1_le (Ⓕ) (m+1) L1)
[ /3 width=12 by fqu_drop, lleq_inv_lift_le, ex2_intro/
| 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_leq.ma".
-include "basic_2/multiple/lleq.ma".
-
-(* LAZY EQUIVALENCE FOR LOCAL ENVIRONMENTS **********************************)
-
-(* Properties on equivalence for local environments *************************)
-
-lemma leq_lleq_trans: ∀L2,L,T,d. L2 ≡[T, d] L →
- ∀L1. 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 →
- ∀L2. L1 ⩬[d, ∞] L2 → L ≡[T, d] L2.
-/2 width=3 by llpx_sn_leq_trans/ qed-.
-
-lemma lleq_leq_repl: ∀L1,L2,T,d. L1 ≡[T, d] L2 → ∀K1. K1 ⩬[d, ∞] L1 →
- ∀K2. 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 →
- ∀J1,J2,W1,W2. L1.ⓑ{J1}W1 ≡[T, 1] L2.ⓑ{J2}W2.
-/2 width=5 by llpx_sn_bind_repl_SO/ qed-.
(* Main properties **********************************************************)
-theorem lleq_trans: ∀d,T. Transitive … (lleq d T).
+theorem lleq_trans: ∀l,T. Transitive … (lleq l T).
/2 width=3 by lleq_llpx_sn_trans/ qed-.
-theorem lleq_canc_sn: ∀L,L1,L2,T,d. L ≡[d, T] L1→ L ≡[d, T] L2 → L1 ≡[d, T] L2.
+theorem lleq_canc_sn: ∀L,L1,L2,T,l. L ≡[l, T] L1→ L ≡[l, T] L2 → L1 ≡[l, T] L2.
/3 width=3 by lleq_trans, lleq_sym/ qed-.
-theorem lleq_canc_dx: ∀L1,L2,L,T,d. L1 ≡[d, T] L → L2 ≡[d, T] L → L1 ≡[d, T] L2.
+theorem lleq_canc_dx: ∀L1,L2,L,T,l. L1 ≡[l, T] L → L2 ≡[l, T] L → L1 ≡[l, T] L2.
/3 width=3 by lleq_trans, lleq_sym/ qed-.
(* Advanced properies on negated lazy equivalence *****************************)
(* Note: for use in auto, works with /4 width=8/ so lleq_canc_sn is preferred *)
-lemma lleq_nlleq_trans: ∀d,T,L1,L. L1 ≡[T, d] L →
- ∀L2. (L ≡[T, d] L2 → ⊥) → (L1 ≡[T, d] L2 → ⊥).
+lemma lleq_nlleq_trans: ∀l,T,L1,L. L1 ≡[T, l] L →
+ ∀L2. (L ≡[T, l] L2 → ⊥) → (L1 ≡[T, l] L2 → ⊥).
/3 width=3 by lleq_canc_sn/ qed-.
-lemma nlleq_lleq_div: ∀d,T,L2,L. L2 ≡[T, d] L →
- ∀L1. (L1 ≡[T, d] L → ⊥) → (L1 ≡[T, d] L2 → ⊥).
+lemma nlleq_lleq_div: ∀l,T,L2,L. L2 ≡[T, l] L →
+ ∀L1. (L1 ≡[T, l] L → ⊥) → (L1 ≡[T, l] L2 → ⊥).
/3 width=3 by lleq_trans/ qed-.
(* Properties on pointwise union for local environments **********************)
lemma llpx_sn_llor_dx: ∀R. (s_r_confluent1 … R (llpx_sn R 0)) → (frees_trans R) →
- ∀L1,L2,T,d. llpx_sn R d T L1 L2 → ∀L. L1 ⋓[T, d] L2 ≡ L → L2 ≡[T, d] L.
-#R #H1R #H2R #L1 #L2 #T #d #H1 #L #H2
+ ∀L1,L2,T,l. llpx_sn R l T L1 L2 → ∀L. L1 ⋓[T, l] L2 ≡ L → L2 ≡[T, l] L.
+#R #H1R #H2R #L1 #L2 #T #l #H1 #L #H2
lapply (llpx_sn_frees_trans … H1R H2R … H1) -H1R -H2R #HR
elim (llpx_sn_llpx_sn_alt … H1) -H1 #HL12 #IH1
elim H2 -H2 #_ #HL1 #IH2
qed.
lemma llpx_sn_llor_dx_sym: ∀R. (s_r_confluent1 … R (llpx_sn R 0)) → (frees_trans R) →
- ∀L1,L2,T,d. llpx_sn R d T L1 L2 → ∀L. L1 ⋓[T, d] L2 ≡ L → L ≡[T, d] L2.
+ ∀L1,L2,T,l. llpx_sn R l T L1 L2 → ∀L. L1 ⋓[T, l] L2 ≡ L → L ≡[T, l] L2.
/3 width=6 by llpx_sn_llor_dx, lleq_sym/ 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/llpx_sn_lreq.ma".
+include "basic_2/multiple/lleq.ma".
+
+(* LAZY EQUIVALENCE FOR LOCAL ENVIRONMENTS **********************************)
+
+(* Properties on equivalence for local environments *************************)
+
+lemma lreq_lleq_trans: ∀L2,L,T,l. L2 ≡[T, l] L →
+ ∀L1. L1 ⩬[l, ∞] L2 → L1 ≡[T, l] L.
+/2 width=3 by lreq_llpx_sn_trans/ qed-.
+
+lemma lleq_lreq_trans: ∀L,L1,T,l. L ≡[T, l] L1 →
+ ∀L2. L1 ⩬[l, ∞] L2 → L ≡[T, l] L2.
+/2 width=3 by llpx_sn_lreq_trans/ qed-.
+
+lemma lleq_lreq_repl: ∀L1,L2,T,l. L1 ≡[T, l] L2 → ∀K1. K1 ⩬[l, ∞] L1 →
+ ∀K2. L2 ⩬[l, ∞] K2 → K1 ≡[T, l] K2.
+/2 width=5 by llpx_sn_lreq_repl/ qed-.
+
+lemma lleq_bind_repl_SO: ∀I1,I2,L1,L2,V1,V2,T. L1.ⓑ{I1}V1 ≡[T, 0] L2.ⓑ{I2}V2 →
+ ∀J1,J2,W1,W2. L1.ⓑ{J1}W1 ≡[T, 1] L2.ⓑ{J2}W2.
+/2 width=5 by llpx_sn_bind_repl_SO/ qed-.
(* POINTWISE UNION FOR LOCAL ENVIRONMENTS ***********************************)
-definition llor: ynat → relation4 term lenv lenv lenv ≝ λd,T,L2,L1,L.
+definition llor: ynat → relation4 term lenv lenv lenv ≝ λl,T,L2,L1,L.
∧∧ |L1| = |L2| & |L1| = |L|
& (∀I1,I2,I,K1,K2,K,V1,V2,V,i.
⬇[i] L1 ≡ K1.ⓑ{I1}V1 → ⬇[i] L2 ≡ K2.ⓑ{I2}V2 → ⬇[i] L ≡ K.ⓑ{I}V → ∨∨
- (∧∧ yinj i < d & I1 = I & V1 = V) |
- (∧∧ (L1 ⊢ i ϵ 𝐅*[d]⦃T⦄ → ⊥) & I1 = I & V1 = V) |
- (∧∧ d ≤ yinj i & L1 ⊢ i ϵ 𝐅*[d]⦃T⦄ & I2 = I & V2 = V)
+ (∧∧ yinj i < l & I1 = I & V1 = V) |
+ (∧∧ (L1 ⊢ i ϵ 𝐅*[l]⦃T⦄ → ⊥) & I1 = I & V1 = V) |
+ (∧∧ l ≤ yinj i & L1 ⊢ i ϵ 𝐅*[l]⦃T⦄ & I2 = I & V2 = V)
).
interpretation
"lazy union (local environment)"
- 'LazyOr L1 T d L2 L = (llor d T L2 L1 L).
+ 'LazyOr L1 T l L2 L = (llor l T L2 L1 L).
(* Basic properties *********************************************************)
(* Note: this can be proved by llor_skip *)
-lemma llor_atom: ∀T,d. ⋆ ⋓[T, d] ⋆ ≡ ⋆.
-#T #d @and3_intro //
+lemma llor_atom: ∀T,l. ⋆ ⋓[T, l] ⋆ ≡ ⋆.
+#T #l @and3_intro //
#I1 #I2 #I #K1 #K2 #K #V1 #V2 #V #i #HLK1
elim (drop_inv_atom1 … HLK1) -HLK1 #H destruct
qed.
(* Alternative definition ***************************************************)
-lemma llor_tail_frees: ∀L1,L2,L,U,d. L1 ⋓[U, d] L2 ≡ L → d ≤ yinj (|L1|) →
- ∀I1,W1. ⓑ{I1}W1.L1 ⊢ |L1| ϵ 𝐅*[d]⦃U⦄ →
- ∀I2,W2. ⓑ{I1}W1.L1 ⋓[U, d] ⓑ{I2}W2.L2 ≡ ⓑ{I2}W2.L.
-#L1 #L2 #L #U #d * #HL12 #HL1 #IH #Hd #I1 #W1 #HU #I2 #W2
+lemma llor_tail_frees: ∀L1,L2,L,U,l. L1 ⋓[U, l] L2 ≡ L → l ≤ yinj (|L1|) →
+ ∀I1,W1. ⓑ{I1}W1.L1 ⊢ |L1| ϵ 𝐅*[l]⦃U⦄ →
+ ∀I2,W2. ⓑ{I1}W1.L1 ⋓[U, l] ⓑ{I2}W2.L2 ≡ ⓑ{I2}W2.L.
+#L1 #L2 #L #U #l * #HL12 #HL1 #IH #Hl #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 (drop_fwd_length_lt2 … HLK1) >ltail_length #H
]
qed.
-lemma llor_tail_cofrees: ∀L1,L2,L,U,d. L1 ⋓[U, d] L2 ≡ L →
- ∀I1,W1. (ⓑ{I1}W1.L1 ⊢ |L1| ϵ 𝐅*[d]⦃U⦄ → ⊥) →
- ∀I2,W2. ⓑ{I1}W1.L1 ⋓[U, d] ⓑ{I2}W2.L2 ≡ ⓑ{I1}W1.L.
-#L1 #L2 #L #U #d * #HL12 #HL1 #IH #I1 #W1 #HU #I2 #W2
+lemma llor_tail_cofrees: ∀L1,L2,L,U,l. L1 ⋓[U, l] L2 ≡ L →
+ ∀I1,W1. (ⓑ{I1}W1.L1 ⊢ |L1| ϵ 𝐅*[l]⦃U⦄ → ⊥) →
+ ∀I2,W2. ⓑ{I1}W1.L1 ⋓[U, l] ⓑ{I2}W2.L2 ≡ ⓑ{I1}W1.L.
+#L1 #L2 #L #U #l * #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 (drop_fwd_length_lt2 … HLK1) >ltail_length #H
(* 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
+lemma llor_skip: ∀L1,L2,U,l. |L1| = |L2| → yinj (|L1|) ≤ l → L1 ⋓[U, l] L2 ≡ L1.
+#L1 #L2 #U #l #HL12 #Hl @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
qed.
(* Note: lemma 1400 concludes the "big tree" theorem *)
-lemma llor_total: ∀L1,L2,T,d. |L1| = |L2| → ∃L. L1 ⋓[T, d] L2 ≡ L.
+lemma llor_total: ∀L1,L2,T,l. |L1| = |L2| → ∃L. L1 ⋓[T, l] 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
+[ #L2 #T #l #H >(length_inv_zero_sn … H) -L2 /2 width=2 by ex_intro/
+| #I1 #L1 #V1 #IHL1 #Y #T #l >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
+ elim (ylt_split l (|ⓑ{I1}V1.L1|))
+ [ elim (frees_dec (ⓑ{I1}V1.L1) T l (|L1|)) #HnU
+ elim (IHL1 L2 T l) // -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/
]
(* LAZY SN POINTWISE EXTENSION OF A CONTEXT-SENSITIVE REALTION FOR TERMS ****)
inductive llpx_sn (R:relation3 lenv term term): relation4 ynat term lenv lenv ≝
-| llpx_sn_sort: ∀L1,L2,d,k. |L1| = |L2| → llpx_sn R d (⋆k) L1 L2
-| llpx_sn_skip: ∀L1,L2,d,i. |L1| = |L2| → yinj i < d → llpx_sn R d (#i) L1 L2
-| llpx_sn_lref: ∀I,L1,L2,K1,K2,V1,V2,d,i. d ≤ yinj i →
+| llpx_sn_sort: ∀L1,L2,l,k. |L1| = |L2| → llpx_sn R l (⋆k) L1 L2
+| llpx_sn_skip: ∀L1,L2,l,i. |L1| = |L2| → yinj i < l → llpx_sn R l (#i) L1 L2
+| llpx_sn_lref: ∀I,L1,L2,K1,K2,V1,V2,l,i. l ≤ yinj i →
⬇[i] L1 ≡ K1.ⓑ{I}V1 → ⬇[i] L2 ≡ K2.ⓑ{I}V2 →
- llpx_sn R (yinj 0) V1 K1 K2 → R K1 V1 V2 → llpx_sn R d (#i) L1 L2
-| llpx_sn_free: ∀L1,L2,d,i. |L1| ≤ i → |L2| ≤ i → |L1| = |L2| → llpx_sn R d (#i) L1 L2
-| llpx_sn_gref: ∀L1,L2,d,p. |L1| = |L2| → llpx_sn R d (§p) L1 L2
-| llpx_sn_bind: ∀a,I,L1,L2,V,T,d.
- llpx_sn R d V L1 L2 → llpx_sn R (⫯d) T (L1.ⓑ{I}V) (L2.ⓑ{I}V) →
- llpx_sn R d (ⓑ{a,I}V.T) L1 L2
-| llpx_sn_flat: ∀I,L1,L2,V,T,d.
- llpx_sn R d V L1 L2 → llpx_sn R d T L1 L2 → llpx_sn R d (ⓕ{I}V.T) L1 L2
+ llpx_sn R (yinj 0) V1 K1 K2 → R K1 V1 V2 → llpx_sn R l (#i) L1 L2
+| llpx_sn_free: ∀L1,L2,l,i. |L1| ≤ i → |L2| ≤ i → |L1| = |L2| → llpx_sn R l (#i) L1 L2
+| llpx_sn_gref: ∀L1,L2,l,p. |L1| = |L2| → llpx_sn R l (§p) L1 L2
+| llpx_sn_bind: ∀a,I,L1,L2,V,T,l.
+ llpx_sn R l V L1 L2 → llpx_sn R (⫯l) T (L1.ⓑ{I}V) (L2.ⓑ{I}V) →
+ llpx_sn R l (ⓑ{a,I}V.T) L1 L2
+| llpx_sn_flat: ∀I,L1,L2,V,T,l.
+ llpx_sn R l V L1 L2 → llpx_sn R l T L1 L2 → llpx_sn R l (ⓕ{I}V.T) L1 L2
.
(* Basic inversion lemmas ***************************************************)
-fact llpx_sn_inv_bind_aux: ∀R,L1,L2,X,d. llpx_sn R d X L1 L2 →
+fact llpx_sn_inv_bind_aux: ∀R,L1,L2,X,l. llpx_sn R l X L1 L2 →
∀a,I,V,T. X = ⓑ{a,I}V.T →
- llpx_sn R d V L1 L2 ∧ llpx_sn R (⫯d) T (L1.ⓑ{I}V) (L2.ⓑ{I}V).
-#R #L1 #L2 #X #d * -L1 -L2 -X -d
-[ #L1 #L2 #d #k #_ #b #J #W #U #H destruct
-| #L1 #L2 #d #i #_ #_ #b #J #W #U #H destruct
-| #I #L1 #L2 #K1 #K2 #V1 #V2 #d #i #_ #_ #_ #_ #_ #b #J #W #U #H destruct
-| #L1 #L2 #d #i #_ #_ #_ #b #J #W #U #H destruct
-| #L1 #L2 #d #p #_ #b #J #W #U #H destruct
-| #a #I #L1 #L2 #V #T #d #HV #HT #b #J #W #U #H destruct /2 width=1 by conj/
-| #I #L1 #L2 #V #T #d #_ #_ #b #J #W #U #H destruct
+ llpx_sn R l V L1 L2 ∧ llpx_sn R (⫯l) T (L1.ⓑ{I}V) (L2.ⓑ{I}V).
+#R #L1 #L2 #X #l * -L1 -L2 -X -l
+[ #L1 #L2 #l #k #_ #b #J #W #U #H destruct
+| #L1 #L2 #l #i #_ #_ #b #J #W #U #H destruct
+| #I #L1 #L2 #K1 #K2 #V1 #V2 #l #i #_ #_ #_ #_ #_ #b #J #W #U #H destruct
+| #L1 #L2 #l #i #_ #_ #_ #b #J #W #U #H destruct
+| #L1 #L2 #l #p #_ #b #J #W #U #H destruct
+| #a #I #L1 #L2 #V #T #l #HV #HT #b #J #W #U #H destruct /2 width=1 by conj/
+| #I #L1 #L2 #V #T #l #_ #_ #b #J #W #U #H destruct
]
qed-.
-lemma llpx_sn_inv_bind: ∀R,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).
+lemma llpx_sn_inv_bind: ∀R,a,I,L1,L2,V,T,l. llpx_sn R l (ⓑ{a,I}V.T) L1 L2 →
+ llpx_sn R l V L1 L2 ∧ llpx_sn R (⫯l) T (L1.ⓑ{I}V) (L2.ⓑ{I}V).
/2 width=4 by llpx_sn_inv_bind_aux/ qed-.
-fact llpx_sn_inv_flat_aux: ∀R,L1,L2,X,d. llpx_sn R d X L1 L2 →
+fact llpx_sn_inv_flat_aux: ∀R,L1,L2,X,l. llpx_sn R l X L1 L2 →
∀I,V,T. X = ⓕ{I}V.T →
- llpx_sn R d V L1 L2 ∧ llpx_sn R d T L1 L2.
-#R #L1 #L2 #X #d * -L1 -L2 -X -d
-[ #L1 #L2 #d #k #_ #J #W #U #H destruct
-| #L1 #L2 #d #i #_ #_ #J #W #U #H destruct
-| #I #L1 #L2 #K1 #K2 #V1 #V2 #d #i #_ #_ #_ #_ #_ #J #W #U #H destruct
-| #L1 #L2 #d #i #_ #_ #_ #J #W #U #H destruct
-| #L1 #L2 #d #p #_ #J #W #U #H destruct
-| #a #I #L1 #L2 #V #T #d #_ #_ #J #W #U #H destruct
-| #I #L1 #L2 #V #T #d #HV #HT #J #W #U #H destruct /2 width=1 by conj/
+ llpx_sn R l V L1 L2 ∧ llpx_sn R l T L1 L2.
+#R #L1 #L2 #X #l * -L1 -L2 -X -l
+[ #L1 #L2 #l #k #_ #J #W #U #H destruct
+| #L1 #L2 #l #i #_ #_ #J #W #U #H destruct
+| #I #L1 #L2 #K1 #K2 #V1 #V2 #l #i #_ #_ #_ #_ #_ #J #W #U #H destruct
+| #L1 #L2 #l #i #_ #_ #_ #J #W #U #H destruct
+| #L1 #L2 #l #p #_ #J #W #U #H destruct
+| #a #I #L1 #L2 #V #T #l #_ #_ #J #W #U #H destruct
+| #I #L1 #L2 #V #T #l #HV #HT #J #W #U #H destruct /2 width=1 by conj/
]
qed-.
-lemma llpx_sn_inv_flat: ∀R,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.
+lemma llpx_sn_inv_flat: ∀R,I,L1,L2,V,T,l. llpx_sn R l (ⓕ{I}V.T) L1 L2 →
+ llpx_sn R l V L1 L2 ∧ llpx_sn R l T L1 L2.
/2 width=4 by llpx_sn_inv_flat_aux/ qed-.
(* Basic forward lemmas *****************************************************)
-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
+lemma llpx_sn_fwd_length: ∀R,L1,L2,T,l. llpx_sn R l T L1 L2 → |L1| = |L2|.
+#R #L1 #L2 #T #l #H elim H -L1 -L2 -T -l //
+#I #L1 #L2 #K1 #K2 #V1 #V2 #l #i #_ #HLK1 #HLK2 #_ #_ #HK12
lapply (drop_fwd_length … HLK1) -HLK1
lapply (drop_fwd_length … HLK2) -HLK2
normalize //
qed-.
-lemma llpx_sn_fwd_drop_sn: ∀R,L1,L2,T,d. llpx_sn R d T L1 L2 →
+lemma llpx_sn_fwd_drop_sn: ∀R,L1,L2,T,l. llpx_sn R l 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
+#R #L1 #L2 #T #l #H #K1 #i #HLK1 lapply (llpx_sn_fwd_length … H) -H
#HL12 lapply (drop_fwd_length_le2 … HLK1) -HLK1 /2 width=1 by drop_O1_le/
qed-.
-lemma llpx_sn_fwd_drop_dx: ∀R,L1,L2,T,d. llpx_sn R d T L1 L2 →
+lemma llpx_sn_fwd_drop_dx: ∀R,L1,L2,T,l. llpx_sn R l 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
+#R #L1 #L2 #T #l #H #K2 #i #HLK2 lapply (llpx_sn_fwd_length … H) -H
#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 →
+fact llpx_sn_fwd_lref_aux: ∀R,L1,L2,X,l. llpx_sn R l X L1 L2 → ∀i. X = #i →
∨∨ |L1| ≤ i ∧ |L2| ≤ i
- | yinj i < d
+ | yinj i < l
| ∃∃I,K1,K2,V1,V2. ⬇[i] L1 ≡ K1.ⓑ{I}V1 &
⬇[i] L2 ≡ K2.ⓑ{I}V2 &
llpx_sn R (yinj 0) V1 K1 K2 &
- R K1 V1 V2 & d ≤ yinj i.
-#R #L1 #L2 #X #d * -L1 -L2 -X -d
-[ #L1 #L2 #d #k #_ #j #H destruct
-| #L1 #L2 #d #i #_ #Hid #j #H destruct /2 width=1 by or3_intro1/
-| #I #L1 #L2 #K1 #K2 #V1 #V2 #d #i #Hdi #HLK1 #HLK2 #HK12 #HV12 #j #H destruct
+ R K1 V1 V2 & l ≤ yinj i.
+#R #L1 #L2 #X #l * -L1 -L2 -X -l
+[ #L1 #L2 #l #k #_ #j #H destruct
+| #L1 #L2 #l #i #_ #Hil #j #H destruct /2 width=1 by or3_intro1/
+| #I #L1 #L2 #K1 #K2 #V1 #V2 #l #i #Hli #HLK1 #HLK2 #HK12 #HV12 #j #H destruct
/3 width=9 by or3_intro2, ex5_5_intro/
-| #L1 #L2 #d #i #HL1 #HL2 #_ #j #H destruct /3 width=1 by or3_intro0, conj/
-| #L1 #L2 #d #p #_ #j #H destruct
-| #a #I #L1 #L2 #V #T #d #_ #_ #j #H destruct
-| #I #L1 #L2 #V #T #d #_ #_ #j #H destruct
+| #L1 #L2 #l #i #HL1 #HL2 #_ #j #H destruct /3 width=1 by or3_intro0, conj/
+| #L1 #L2 #l #p #_ #j #H destruct
+| #a #I #L1 #L2 #V #T #l #_ #_ #j #H destruct
+| #I #L1 #L2 #V #T #l #_ #_ #j #H destruct
]
qed-.
-lemma llpx_sn_fwd_lref: ∀R,L1,L2,d,i. llpx_sn R d (#i) L1 L2 →
+lemma llpx_sn_fwd_lref: ∀R,L1,L2,l,i. llpx_sn R l (#i) L1 L2 →
∨∨ |L1| ≤ i ∧ |L2| ≤ i
- | yinj i < d
+ | yinj i < l
| ∃∃I,K1,K2,V1,V2. ⬇[i] L1 ≡ K1.ⓑ{I}V1 &
⬇[i] L2 ≡ K2.ⓑ{I}V2 &
llpx_sn R (yinj 0) V1 K1 K2 &
- R K1 V1 V2 & d ≤ yinj i.
+ R K1 V1 V2 & l ≤ yinj i.
/2 width=3 by llpx_sn_fwd_lref_aux/ qed-.
-lemma llpx_sn_fwd_bind_sn: ∀R,a,I,L1,L2,V,T,d. llpx_sn R d (ⓑ{a,I}V.T) L1 L2 →
- llpx_sn R d V L1 L2.
-#R #a #I #L1 #L2 #V #T #d #H elim (llpx_sn_inv_bind … H) -H //
+lemma llpx_sn_fwd_bind_sn: ∀R,a,I,L1,L2,V,T,l. llpx_sn R l (ⓑ{a,I}V.T) L1 L2 →
+ llpx_sn R l V L1 L2.
+#R #a #I #L1 #L2 #V #T #l #H elim (llpx_sn_inv_bind … H) -H //
qed-.
-lemma llpx_sn_fwd_bind_dx: ∀R,a,I,L1,L2,V,T,d. llpx_sn R d (ⓑ{a,I}V.T) L1 L2 →
- llpx_sn R (⫯d) T (L1.ⓑ{I}V) (L2.ⓑ{I}V).
-#R #a #I #L1 #L2 #V #T #d #H elim (llpx_sn_inv_bind … H) -H //
+lemma llpx_sn_fwd_bind_dx: ∀R,a,I,L1,L2,V,T,l. llpx_sn R l (ⓑ{a,I}V.T) L1 L2 →
+ llpx_sn R (⫯l) T (L1.ⓑ{I}V) (L2.ⓑ{I}V).
+#R #a #I #L1 #L2 #V #T #l #H elim (llpx_sn_inv_bind … H) -H //
qed-.
-lemma llpx_sn_fwd_flat_sn: ∀R,I,L1,L2,V,T,d. llpx_sn R d (ⓕ{I}V.T) L1 L2 →
- llpx_sn R d V L1 L2.
-#R #I #L1 #L2 #V #T #d #H elim (llpx_sn_inv_flat … H) -H //
+lemma llpx_sn_fwd_flat_sn: ∀R,I,L1,L2,V,T,l. llpx_sn R l (ⓕ{I}V.T) L1 L2 →
+ llpx_sn R l V L1 L2.
+#R #I #L1 #L2 #V #T #l #H elim (llpx_sn_inv_flat … H) -H //
qed-.
-lemma llpx_sn_fwd_flat_dx: ∀R,I,L1,L2,V,T,d. llpx_sn R d (ⓕ{I}V.T) L1 L2 →
- llpx_sn R d T L1 L2.
-#R #I #L1 #L2 #V #T #d #H elim (llpx_sn_inv_flat … H) -H //
+lemma llpx_sn_fwd_flat_dx: ∀R,I,L1,L2,V,T,l. llpx_sn R l (ⓕ{I}V.T) L1 L2 →
+ llpx_sn R l T L1 L2.
+#R #I #L1 #L2 #V #T #l #H elim (llpx_sn_inv_flat … H) -H //
qed-.
-lemma llpx_sn_fwd_pair_sn: ∀R,I,L1,L2,V,T,d. llpx_sn R d (②{I}V.T) L1 L2 →
- llpx_sn R d V L1 L2.
+lemma llpx_sn_fwd_pair_sn: ∀R,I,L1,L2,V,T,l. llpx_sn R l (②{I}V.T) L1 L2 →
+ llpx_sn R l V L1 L2.
#R * /2 width=4 by llpx_sn_fwd_flat_sn, llpx_sn_fwd_bind_sn/
qed-.
(* Basic properties *********************************************************)
-lemma llpx_sn_refl: ∀R. (∀L. reflexive … (R L)) → ∀T,L,d. llpx_sn R d T L L.
+lemma llpx_sn_refl: ∀R. (∀L. reflexive … (R L)) → ∀T,L,l. llpx_sn R l T L L.
#R #HR #T #L @(f2_ind … rfw … L T) -L -T
#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/
+#HiL #l elim (ylt_split i l) /2 width=1 by llpx_sn_skip/
elim (drop_O1_lt … HiL) -HiL destruct /4 width=9 by llpx_sn_lref, drop_fwd_rfw/
qed-.
@IH -IH // normalize /2 width=1 by eq_f2/
qed-.
-lemma llpx_sn_ge_up: ∀R,L1,L2,U,dt. llpx_sn R dt U L1 L2 → ∀T,d,e. ⬆[d, e] T ≡ U →
- dt ≤ d + e → llpx_sn R d U L1 L2.
-#R #L1 #L2 #U #dt #H elim H -L1 -L2 -U -dt
-[ #L1 #L2 #dt #k #HL12 #X #d #e #H #_ >(lift_inv_sort2 … H) -H /2 width=1 by llpx_sn_sort/
-| #L1 #L2 #dt #i #HL12 #Hidt #X #d #e #H #Hdtde
- elim (lift_inv_lref2 … H) -H * #Hid #H destruct /3 width=1 by llpx_sn_skip, ylt_inj/ -HL12
- elim (ylt_yle_false … Hidt) -Hidt
- @(yle_trans … Hdtde) /2 width=1 by yle_inj/ (**) (* full auto too slow 11s *)
-| #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
+lemma llpx_sn_ge_up: ∀R,L1,L2,U,lt. llpx_sn R lt U L1 L2 → ∀T,l,m. ⬆[l, m] T ≡ U →
+ lt ≤ l + m → llpx_sn R l U L1 L2.
+#R #L1 #L2 #U #lt #H elim H -L1 -L2 -U -lt
+[ #L1 #L2 #lt #k #HL12 #X #l #m #H #_ >(lift_inv_sort2 … H) -H /2 width=1 by llpx_sn_sort/
+| #L1 #L2 #lt #i #HL12 #Hilt #X #l #m #H #Hltlm
+ elim (lift_inv_lref2 … H) -H * #Hil #H destruct /3 width=1 by llpx_sn_skip, ylt_inj/ -HL12
+ elim (ylt_yle_false … Hilt) -Hilt
+ @(yle_trans … Hltlm) /2 width=1 by yle_inj/ (**) (* full auto too slow 11s *)
+| #I #L1 #L2 #K1 #K2 #W1 #W2 #lt #i #Hlti #HLK1 #HLK2 #HW1 #HW12 #_ #X #l #m #H #_
+ elim (lift_inv_lref2 … H) -H * #Hil #H destruct
[ lapply (llpx_sn_fwd_length … HW1) -HW1 #HK12
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/
+ normalize in ⊢ (%→%→?); -I -W1 -W2 -lt /3 width=1 by llpx_sn_skip, ylt_inj/
| /4 width=9 by llpx_sn_lref, yle_inj, le_plus_b/
]
| /2 width=1 by llpx_sn_free/
-| #L1 #L2 #dt #p #HL12 #X #d #e #H #_ >(lift_inv_gref2 … H) -H /2 width=1 by llpx_sn_gref/
-| #a #I #L1 #L2 #W #U #dt #_ #_ #IHV #IHT #X #d #e #H #Hdtde destruct
+| #L1 #L2 #lt #p #HL12 #X #l #m #H #_ >(lift_inv_gref2 … H) -H /2 width=1 by llpx_sn_gref/
+| #a #I #L1 #L2 #W #U #lt #_ #_ #IHV #IHT #X #l #m #H #Hltlm destruct
elim (lift_inv_bind2 … H) -H #V #T #HVW >commutative_plus #HTU #H destruct
@(llpx_sn_bind) /2 width=4 by/ (**) (* full auto fails *)
@(IHT … HTU) /2 width=1 by yle_succ/
-| #I #L1 #L2 #W #U #dt #_ #_ #IHV #IHT #X #d #e #H #Hdtde destruct
+| #I #L1 #L2 #W #U #lt #_ #_ #IHV #IHT #X #l #m #H #Hltlm destruct
elim (lift_inv_flat2 … H) -H #HVW #HTU #H destruct
/3 width=4 by llpx_sn_flat/
]
qed-.
(**) (* the minor premise comes first *)
-lemma llpx_sn_ge: ∀R,L1,L2,T,d1,d2. d1 ≤ d2 →
- llpx_sn R d1 T L1 L2 → llpx_sn R d2 T L1 L2.
-#R #L1 #L2 #T #d1 #d2 * -d1 -d2 (**) (* destructed yle *)
+lemma llpx_sn_ge: ∀R,L1,L2,T,l1,l2. l1 ≤ l2 →
+ llpx_sn R l1 T L1 L2 → llpx_sn R l2 T L1 L2.
+#R #L1 #L2 #T #l1 #l2 * -l1 -l2 (**) (* destructed yle *)
/3 width=6 by llpx_sn_ge_up, llpx_sn_Y, llpx_sn_fwd_length, yle_inj/
qed-.
/3 width=3 by llpx_sn_ge, llpx_sn_bind/ qed-.
lemma llpx_sn_co: ∀R1,R2. (∀L,T1,T2. R1 L T1 T2 → R2 L T1 T2) →
- ∀L1,L2,T,d. llpx_sn R1 d T L1 L2 → llpx_sn R2 d T L1 L2.
-#R1 #R2 #HR12 #L1 #L2 #T #d #H elim H -L1 -L2 -T -d
+ ∀L1,L2,T,l. llpx_sn R1 l T L1 L2 → llpx_sn R2 l T L1 L2.
+#R1 #R2 #HR12 #L1 #L2 #T #l #H elim H -L1 -L2 -T -l
/3 width=9 by llpx_sn_sort, llpx_sn_skip, llpx_sn_lref, llpx_sn_free, llpx_sn_gref, llpx_sn_bind, llpx_sn_flat/
qed-.
(* alternative definition of llpx_sn (not recursive) *)
definition llpx_sn_alt: relation3 lenv term term → relation4 ynat term lenv lenv ≝
- λR,d,T,L1,L2. |L1| = |L2| ∧
- (∀I1,I2,K1,K2,V1,V2,i. d ≤ yinj i → L1 ⊢ i ϵ 𝐅*[d]⦃T⦄ →
+ λR,l,T,L1,L2. |L1| = |L2| ∧
+ (∀I1,I2,K1,K2,V1,V2,i. l ≤ yinj i → L1 ⊢ i ϵ 𝐅*[l]⦃T⦄ →
⬇[i] L1 ≡ K1.ⓑ{I1}V1 → ⬇[i] L2 ≡ K2.ⓑ{I2}V2 →
I1 = I2 ∧ R K1 V1 V2
).
(* Main properties **********************************************************)
-theorem llpx_sn_llpx_sn_alt: ∀R,T,L1,L2,d. llpx_sn R d T L1 L2 → llpx_sn_alt R d T L1 L2.
+theorem llpx_sn_llpx_sn_alt: ∀R,T,L1,L2,l. llpx_sn R l T L1 L2 → llpx_sn_alt R l T L1 L2.
#R #U #L1 @(f2_ind … rfw … L1 U) -L1 -U
-#n #IHn #L1 #U #Hn #L2 #d #H elim (llpx_sn_inv_alt_r … H) -H
+#n #IHn #L1 #U #Hn #L2 #l #H elim (llpx_sn_inv_alt_r … H) -H
#HL12 #IHU @conj //
-#I1 #I2 #K1 #K2 #V1 #V2 #i #Hdi #H #HLK1 #HLK2 elim (frees_inv … H) -H
+#I1 #I2 #K1 #K2 #V1 #V2 #i #Hli #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
+| * #J1 #K10 #W10 #j #Hlj #Hji #HnU #HLK10 #HnW10 destruct
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
]
qed.
-theorem llpx_sn_alt_inv_llpx_sn: ∀R,T,L1,L2,d. llpx_sn_alt R d T L1 L2 → llpx_sn R d T L1 L2.
+theorem llpx_sn_alt_inv_llpx_sn: ∀R,T,L1,L2,l. llpx_sn_alt R l T L1 L2 → llpx_sn R l T L1 L2.
#R #U #L1 @(f2_ind … rfw … L1 U) -L1 -U
-#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
+#n #IHn #L1 #U #Hn #L2 #l * #HL12 #IHU @llpx_sn_intro_alt_r //
+#I1 #I2 #K1 #K2 #V1 #V2 #i #Hli #HnU #HLK1 #HLK2 destruct
elim (IHU … HLK1 HLK2) /3 width=2 by frees_eq/
#H #HV12 @and3_intro // @IHn -IHn /3 width=6 by drop_fwd_rfw/
lapply (drop_fwd_drop2 … HLK1) #H1
(* alternative definition of llpx_sn (recursive) *)
inductive llpx_sn_alt_r (R:relation3 lenv term term): relation4 ynat term lenv lenv ≝
-| llpx_sn_alt_r_intro: ∀L1,L2,T,d.
- (∀I1,I2,K1,K2,V1,V2,i. d ≤ yinj i → (∀U. ⬆[i, 1] U ≡ T → ⊥) →
+| llpx_sn_alt_r_intro: ∀L1,L2,T,l.
+ (∀I1,I2,K1,K2,V1,V2,i. l ≤ yinj i → (∀U. ⬆[i, 1] U ≡ T → ⊥) →
⬇[i] L1 ≡ K1.ⓑ{I1}V1 → ⬇[i] L2 ≡ K2.ⓑ{I2}V2 → I1 = I2 ∧ R K1 V1 V2
) →
- (∀I1,I2,K1,K2,V1,V2,i. d ≤ yinj i → (∀U. ⬆[i, 1] U ≡ T → ⊥) →
+ (∀I1,I2,K1,K2,V1,V2,i. l ≤ yinj i → (∀U. ⬆[i, 1] U ≡ T → ⊥) →
⬇[i] L1 ≡ K1.ⓑ{I1}V1 → ⬇[i] L2 ≡ K2.ⓑ{I2}V2 → llpx_sn_alt_r R 0 V1 K1 K2
- ) → |L1| = |L2| → llpx_sn_alt_r R d T L1 L2
+ ) → |L1| = |L2| → llpx_sn_alt_r R l T L1 L2
.
(* Compact definition of llpx_sn_alt_r **************************************)
-lemma llpx_sn_alt_r_intro_alt: ∀R,L1,L2,T,d. |L1| = |L2| →
- (∀I1,I2,K1,K2,V1,V2,i. d ≤ yinj i → (∀U. ⬆[i, 1] U ≡ T → ⊥) →
+lemma llpx_sn_alt_r_intro_alt: ∀R,L1,L2,T,l. |L1| = |L2| →
+ (∀I1,I2,K1,K2,V1,V2,i. l ≤ yinj i → (∀U. ⬆[i, 1] U ≡ T → ⊥) →
⬇[i] L1 ≡ K1.ⓑ{I1}V1 → ⬇[i] L2 ≡ K2.ⓑ{I2}V2 →
∧∧ I1 = I2 & R K1 V1 V2 & llpx_sn_alt_r R 0 V1 K1 K2
- ) → llpx_sn_alt_r R d T L1 L2.
-#R #L1 #L2 #T #d #HL12 #IH @llpx_sn_alt_r_intro // -HL12
-#I1 #I2 #K1 #K2 #V1 #V2 #i #Hid #HnT #HLK1 #HLK2
+ ) → llpx_sn_alt_r R l T L1 L2.
+#R #L1 #L2 #T #l #HL12 #IH @llpx_sn_alt_r_intro // -HL12
+#I1 #I2 #K1 #K2 #V1 #V2 #i #Hil #HnT #HLK1 #HLK2
elim (IH … HnT HLK1 HLK2) -IH -HnT -HLK1 -HLK2 /2 width=1 by conj/
qed.
lemma llpx_sn_alt_r_ind_alt: ∀R. ∀S:relation4 ynat term lenv lenv.
- (∀L1,L2,T,d. |L1| = |L2| → (
- ∀I1,I2,K1,K2,V1,V2,i. d ≤ yinj i → (∀U. ⬆[i, 1] U ≡ T → ⊥) →
+ (∀L1,L2,T,l. |L1| = |L2| → (
+ ∀I1,I2,K1,K2,V1,V2,i. l ≤ yinj i → (∀U. ⬆[i, 1] U ≡ T → ⊥) →
⬇[i] L1 ≡ K1.ⓑ{I1}V1 → ⬇[i] L2 ≡ K2.ⓑ{I2}V2 →
∧∧ I1 = I2 & R K1 V1 V2 & llpx_sn_alt_r R 0 V1 K1 K2 & S 0 V1 K1 K2
- ) → S d T L1 L2) →
- ∀L1,L2,T,d. llpx_sn_alt_r R d T L1 L2 → S d T L1 L2.
-#R #S #IH #L1 #L2 #T #d #H elim H -L1 -L2 -T -d
-#L1 #L2 #T #d #H1 #H2 #HL12 #IH2 @IH -IH // -HL12
-#I1 #I2 #K1 #K2 #V1 #V2 #i #Hid #HnT #HLK1 #HLK2
+ ) → S l T L1 L2) →
+ ∀L1,L2,T,l. llpx_sn_alt_r R l T L1 L2 → S l T L1 L2.
+#R #S #IH #L1 #L2 #T #l #H elim H -L1 -L2 -T -l
+#L1 #L2 #T #l #H1 #H2 #HL12 #IH2 @IH -IH // -HL12
+#I1 #I2 #K1 #K2 #V1 #V2 #i #Hil #HnT #HLK1 #HLK2
elim (H1 … HnT HLK1 HLK2) -H1 /4 width=8 by and4_intro/
qed-.
-lemma llpx_sn_alt_r_inv_alt: ∀R,L1,L2,T,d. llpx_sn_alt_r R d T L1 L2 →
+lemma llpx_sn_alt_r_inv_alt: ∀R,L1,L2,T,l. llpx_sn_alt_r R l T L1 L2 →
|L1| = |L2| ∧
- ∀I1,I2,K1,K2,V1,V2,i. d ≤ yinj i → (∀U. ⬆[i, 1] U ≡ T → ⊥) →
+ ∀I1,I2,K1,K2,V1,V2,i. l ≤ yinj i → (∀U. ⬆[i, 1] U ≡ T → ⊥) →
⬇[i] L1 ≡ K1.ⓑ{I1}V1 → ⬇[i] L2 ≡ K2.ⓑ{I2}V2 →
∧∧ I1 = I2 & R K1 V1 V2 & llpx_sn_alt_r R 0 V1 K1 K2.
-#R #L1 #L2 #T #d #H @(llpx_sn_alt_r_ind_alt … H) -L1 -L2 -T -d
-#L1 #L2 #T #d #HL12 #IH @conj // -HL12
-#I1 #I2 #K1 #K2 #V1 #V2 #i #Hid #HnT #HLK1 #HLK2
+#R #L1 #L2 #T #l #H @(llpx_sn_alt_r_ind_alt … H) -L1 -L2 -T -l
+#L1 #L2 #T #l #HL12 #IH @conj // -HL12
+#I1 #I2 #K1 #K2 #V1 #V2 #i #Hil #HnT #HLK1 #HLK2
elim (IH … HnT HLK1 HLK2) -IH -HnT -HLK1 -HLK2 /2 width=1 by and3_intro/
qed-.
(* Basic inversion lemmas ***************************************************)
-lemma llpx_sn_alt_r_inv_flat: ∀R,I,L1,L2,V,T,d. llpx_sn_alt_r R d (ⓕ{I}V.T) L1 L2 →
- llpx_sn_alt_r R d V L1 L2 ∧ llpx_sn_alt_r R d T L1 L2.
-#R #I #L1 #L2 #V #T #d #H elim (llpx_sn_alt_r_inv_alt … H) -H
+lemma llpx_sn_alt_r_inv_flat: ∀R,I,L1,L2,V,T,l. llpx_sn_alt_r R l (ⓕ{I}V.T) L1 L2 →
+ llpx_sn_alt_r R l V L1 L2 ∧ llpx_sn_alt_r R l T L1 L2.
+#R #I #L1 #L2 #V #T #l #H elim (llpx_sn_alt_r_inv_alt … H) -H
#HL12 #IH @conj @llpx_sn_alt_r_intro_alt // -HL12
-#I1 #I2 #K1 #K2 #V1 #V2 #i #Hdi #H #HLK1 #HLK2
+#I1 #I2 #K1 #K2 #V1 #V2 #i #Hli #H #HLK1 #HLK2
elim (IH … HLK1 HLK2) -IH -HLK1 -HLK2 //
/3 width=8 by nlift_flat_sn, nlift_flat_dx, and3_intro/
qed-.
-lemma llpx_sn_alt_r_inv_bind: ∀R,a,I,L1,L2,V,T,d. llpx_sn_alt_r R d (ⓑ{a,I}V.T) L1 L2 →
- llpx_sn_alt_r R d V L1 L2 ∧ llpx_sn_alt_r R (⫯d) T (L1.ⓑ{I}V) (L2.ⓑ{I}V).
-#R #a #I #L1 #L2 #V #T #d #H elim (llpx_sn_alt_r_inv_alt … H) -H
+lemma llpx_sn_alt_r_inv_bind: ∀R,a,I,L1,L2,V,T,l. llpx_sn_alt_r R l (ⓑ{a,I}V.T) L1 L2 →
+ llpx_sn_alt_r R l V L1 L2 ∧ llpx_sn_alt_r R (⫯l) T (L1.ⓑ{I}V) (L2.ⓑ{I}V).
+#R #a #I #L1 #L2 #V #T #l #H elim (llpx_sn_alt_r_inv_alt … H) -H
#HL12 #IH @conj @llpx_sn_alt_r_intro_alt [1,3: normalize // ] -HL12
-#I1 #I2 #K1 #K2 #V1 #V2 #i #Hdi #H #HLK1 #HLK2
+#I1 #I2 #K1 #K2 #V1 #V2 #i #Hli #H #HLK1 #HLK2
[ 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 (yle_inv_succ1 … Hli) -Hli * #Hli #Hi
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/
(* Basic forward lemmas *****************************************************)
-lemma llpx_sn_alt_r_fwd_length: ∀R,L1,L2,T,d. llpx_sn_alt_r R d T L1 L2 → |L1| = |L2|.
-#R #L1 #L2 #T #d #H elim (llpx_sn_alt_r_inv_alt … H) -H //
+lemma llpx_sn_alt_r_fwd_length: ∀R,L1,L2,T,l. llpx_sn_alt_r R l T L1 L2 → |L1| = |L2|.
+#R #L1 #L2 #T #l #H elim (llpx_sn_alt_r_inv_alt … H) -H //
qed-.
-lemma llpx_sn_alt_r_fwd_lref: ∀R,L1,L2,d,i. llpx_sn_alt_r R d (#i) L1 L2 →
+lemma llpx_sn_alt_r_fwd_lref: ∀R,L1,L2,l,i. llpx_sn_alt_r R l (#i) L1 L2 →
∨∨ |L1| ≤ i ∧ |L2| ≤ i
- | yinj i < d
+ | yinj i < l
| ∃∃I,K1,K2,V1,V2. ⬇[i] L1 ≡ K1.ⓑ{I}V1 &
⬇[i] L2 ≡ K2.ⓑ{I}V2 &
llpx_sn_alt_r R (yinj 0) V1 K1 K2 &
- R K1 V1 V2 & d ≤ yinj i.
-#R #L1 #L2 #d #i #H elim (llpx_sn_alt_r_inv_alt … H) -H
+ R K1 V1 V2 & l ≤ yinj i.
+#R #L1 #L2 #l #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 (drop_O1_lt (Ⓕ) … HL1)
+elim (ylt_split i l) /3 width=1 by or3_intro1/
+#Hli #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/
(* Basic properties *********************************************************)
-lemma llpx_sn_alt_r_sort: ∀R,L1,L2,d,k. |L1| = |L2| → llpx_sn_alt_r R d (⋆k) L1 L2.
-#R #L1 #L2 #d #k #HL12 @llpx_sn_alt_r_intro_alt // -HL12
+lemma llpx_sn_alt_r_sort: ∀R,L1,L2,l,k. |L1| = |L2| → llpx_sn_alt_r R l (⋆k) L1 L2.
+#R #L1 #L2 #l #k #HL12 @llpx_sn_alt_r_intro_alt // -HL12
#I1 #I2 #K1 #K2 #V1 #V2 #i #_ #H elim (H (⋆k)) //
qed.
-lemma llpx_sn_alt_r_gref: ∀R,L1,L2,d,p. |L1| = |L2| → llpx_sn_alt_r R d (§p) L1 L2.
-#R #L1 #L2 #d #p #HL12 @llpx_sn_alt_r_intro_alt // -HL12
+lemma llpx_sn_alt_r_gref: ∀R,L1,L2,l,p. |L1| = |L2| → llpx_sn_alt_r R l (§p) L1 L2.
+#R #L1 #L2 #l #p #HL12 @llpx_sn_alt_r_intro_alt // -HL12
#I1 #I2 #K1 #K2 #V1 #V2 #i #_ #H elim (H (§p)) //
qed.
-lemma llpx_sn_alt_r_skip: ∀R,L1,L2,d,i. |L1| = |L2| → yinj i < d → llpx_sn_alt_r R d (#i) L1 L2.
-#R #L1 #L2 #d #i #HL12 #Hid @llpx_sn_alt_r_intro_alt // -HL12
-#I1 #I2 #K1 #K2 #V1 #V2 #j #Hdj #H elim (H (#i)) -H
+lemma llpx_sn_alt_r_skip: ∀R,L1,L2,l,i. |L1| = |L2| → yinj i < l → llpx_sn_alt_r R l (#i) L1 L2.
+#R #L1 #L2 #l #i #HL12 #Hil @llpx_sn_alt_r_intro_alt // -HL12
+#I1 #I2 #K1 #K2 #V1 #V2 #j #Hlj #H elim (H (#i)) -H
/4 width=3 by lift_lref_lt, ylt_yle_trans, ylt_inv_inj/
qed.
-lemma llpx_sn_alt_r_free: ∀R,L1,L2,d,i. |L1| ≤ i → |L2| ≤ i → |L1| = |L2| →
- llpx_sn_alt_r R d (#i) L1 L2.
-#R #L1 #L2 #d #i #HL1 #_ #HL12 @llpx_sn_alt_r_intro_alt // -HL12
+lemma llpx_sn_alt_r_free: ∀R,L1,L2,l,i. |L1| ≤ i → |L2| ≤ i → |L1| = |L2| →
+ llpx_sn_alt_r R l (#i) L1 L2.
+#R #L1 #L2 #l #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 (drop_fwd_length_lt2 … HLK1) -HLK1
/3 width=3 by lift_lref_ge_minus, lt_to_le_to_lt/
qed.
-lemma llpx_sn_alt_r_lref: ∀R,I,L1,L2,K1,K2,V1,V2,d,i. d ≤ yinj i →
+lemma llpx_sn_alt_r_lref: ∀R,I,L1,L2,K1,K2,V1,V2,l,i. l ≤ yinj i →
⬇[i] L1 ≡ K1.ⓑ{I}V1 → ⬇[i] L2 ≡ K2.ⓑ{I}V2 →
llpx_sn_alt_r R 0 V1 K1 K2 → R K1 V1 V2 →
- 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
+ llpx_sn_alt_r R l (#i) L1 L2.
+#R #I #L1 #L2 #K1 #K2 #V1 #V2 #l #i #Hli #HLK1 #HLK2 #HK12 #HV12 @llpx_sn_alt_r_intro_alt
[ lapply (llpx_sn_alt_r_fwd_length … HK12) -HK12 #HK12
@(drop_fwd_length_eq2 … HLK1 HLK2) normalize //
-| #Z1 #Z2 #Y1 #Y2 #X1 #X2 #j #Hdj #H #HLY1 #HLY2
+| #Z1 #Z2 #Y1 #Y2 #X1 #X2 #j #Hlj #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 (drop_mono … HLY1 … HLK1) -HLY1 -HLK1 #H destruct
]
qed.
-lemma llpx_sn_alt_r_flat: ∀R,I,L1,L2,V,T,d.
- llpx_sn_alt_r R d V L1 L2 → llpx_sn_alt_r R d T L1 L2 →
- llpx_sn_alt_r R d (ⓕ{I}V.T) L1 L2.
-#R #I #L1 #L2 #V #T #d #HV #HT
+lemma llpx_sn_alt_r_flat: ∀R,I,L1,L2,V,T,l.
+ llpx_sn_alt_r R l V L1 L2 → llpx_sn_alt_r R l T L1 L2 →
+ llpx_sn_alt_r R l (ⓕ{I}V.T) L1 L2.
+#R #I #L1 #L2 #V #T #l #HV #HT
elim (llpx_sn_alt_r_inv_alt … HV) -HV #HL12 #IHV
elim (llpx_sn_alt_r_inv_alt … HT) -HT #_ #IHT
@llpx_sn_alt_r_intro_alt // -HL12
-#I1 #I2 #K1 #K2 #V1 #V2 #i #Hdi #HnVT #HLK1 #HLK2
+#I1 #I2 #K1 #K2 #V1 #V2 #i #Hli #HnVT #HLK1 #HLK2
elim (nlift_inv_flat … HnVT) -HnVT #H
[ elim (IHV … HLK1 … HLK2) -IHV /2 width=2 by and3_intro/
| elim (IHT … HLK1 … HLK2) -IHT /3 width=2 by and3_intro/
]
qed.
-lemma llpx_sn_alt_r_bind: ∀R,a,I,L1,L2,V,T,d.
- llpx_sn_alt_r R d V L1 L2 →
- llpx_sn_alt_r R (⫯d) T (L1.ⓑ{I}V) (L2.ⓑ{I}V) →
- llpx_sn_alt_r R d (ⓑ{a,I}V.T) L1 L2.
-#R #a #I #L1 #L2 #V #T #d #HV #HT
+lemma llpx_sn_alt_r_bind: ∀R,a,I,L1,L2,V,T,l.
+ llpx_sn_alt_r R l V L1 L2 →
+ llpx_sn_alt_r R (⫯l) T (L1.ⓑ{I}V) (L2.ⓑ{I}V) →
+ llpx_sn_alt_r R l (ⓑ{a,I}V.T) L1 L2.
+#R #a #I #L1 #L2 #V #T #l #HV #HT
elim (llpx_sn_alt_r_inv_alt … HV) -HV #HL12 #IHV
elim (llpx_sn_alt_r_inv_alt … HT) -HT #_ #IHT
@llpx_sn_alt_r_intro_alt // -HL12
-#I1 #I2 #K1 #K2 #V1 #V2 #i #Hdi #HnVT #HLK1 #HLK2
+#I1 #I2 #K1 #K2 #V1 #V2 #i #Hli #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 drop_drop, yle_succ, and3_intro/
(* Main properties **********************************************************)
-theorem llpx_sn_lpx_sn_alt_r: ∀R,L1,L2,T,d. llpx_sn R d T L1 L2 → llpx_sn_alt_r R d T L1 L2.
-#R #L1 #L2 #T #d #H elim H -L1 -L2 -T -d
+theorem llpx_sn_lpx_sn_alt_r: ∀R,L1,L2,T,l. llpx_sn R l T L1 L2 → llpx_sn_alt_r R l T L1 L2.
+#R #L1 #L2 #T #l #H elim H -L1 -L2 -T -l
/2 width=9 by llpx_sn_alt_r_sort, llpx_sn_alt_r_gref, llpx_sn_alt_r_skip, llpx_sn_alt_r_free, llpx_sn_alt_r_lref, llpx_sn_alt_r_flat, llpx_sn_alt_r_bind/
qed.
(* Main inversion lemmas ****************************************************)
-theorem llpx_sn_alt_r_inv_lpx_sn: ∀R,T,L1,L2,d. llpx_sn_alt_r R d T L1 L2 → llpx_sn R d T L1 L2.
+theorem llpx_sn_alt_r_inv_lpx_sn: ∀R,T,L1,L2,l. llpx_sn_alt_r R l T L1 L2 → llpx_sn R l T L1 L2.
#R #T #L1 @(f2_ind … rfw … L1 T) -L1 -T #n #IH #L1 * *
[1,3: /3 width=4 by llpx_sn_alt_r_fwd_length, llpx_sn_gref, llpx_sn_sort/
-| #i #Hn #L2 #d #H lapply (llpx_sn_alt_r_fwd_length … H)
+| #i #Hn #L2 #l #H lapply (llpx_sn_alt_r_fwd_length … H)
#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, drop_fwd_rfw/
]
-| #a #I #V #T #Hn #L2 #d #H elim (llpx_sn_alt_r_inv_bind … H) -H
+| #a #I #V #T #Hn #L2 #l #H elim (llpx_sn_alt_r_inv_bind … H) -H
/3 width=1 by llpx_sn_bind/
-| #I #V #T #Hn #L2 #d #H elim (llpx_sn_alt_r_inv_flat … H) -H
+| #I #V #T #Hn #L2 #l #H elim (llpx_sn_alt_r_inv_flat … H) -H
/3 width=1 by llpx_sn_flat/
]
qed-.
(* Alternative definition of llpx_sn (recursive) ****************************)
-lemma llpx_sn_intro_alt_r: ∀R,L1,L2,T,d. |L1| = |L2| →
- (∀I1,I2,K1,K2,V1,V2,i. d ≤ yinj i → (∀U. ⬆[i, 1] U ≡ T → ⊥) →
+lemma llpx_sn_intro_alt_r: ∀R,L1,L2,T,l. |L1| = |L2| →
+ (∀I1,I2,K1,K2,V1,V2,i. l ≤ yinj i → (∀U. ⬆[i, 1] U ≡ T → ⊥) →
⬇[i] L1 ≡ K1.ⓑ{I1}V1 → ⬇[i] L2 ≡ K2.ⓑ{I2}V2 →
∧∧ I1 = I2 & R K1 V1 V2 & llpx_sn R 0 V1 K1 K2
- ) → llpx_sn R d T L1 L2.
-#R #L1 #L2 #T #d #HL12 #IH @llpx_sn_alt_r_inv_lpx_sn
+ ) → llpx_sn R l T L1 L2.
+#R #L1 #L2 #T #l #HL12 #IH @llpx_sn_alt_r_inv_lpx_sn
@llpx_sn_alt_r_intro_alt // -HL12
-#I1 #I2 #K1 #K2 #V1 #V2 #i #Hid #HnT #HLK1 #HLK2
+#I1 #I2 #K1 #K2 #V1 #V2 #i #Hil #HnT #HLK1 #HLK2
elim (IH … HnT HLK1 HLK2) -IH -HnT -HLK1 -HLK2 /3 width=1 by llpx_sn_lpx_sn_alt_r, and3_intro/
qed.
lemma llpx_sn_ind_alt_r: ∀R. ∀S:relation4 ynat term lenv lenv.
- (∀L1,L2,T,d. |L1| = |L2| → (
- ∀I1,I2,K1,K2,V1,V2,i. d ≤ yinj i → (∀U. ⬆[i, 1] U ≡ T → ⊥) →
+ (∀L1,L2,T,l. |L1| = |L2| → (
+ ∀I1,I2,K1,K2,V1,V2,i. l ≤ yinj i → (∀U. ⬆[i, 1] U ≡ T → ⊥) →
⬇[i] L1 ≡ K1.ⓑ{I1}V1 → ⬇[i] L2 ≡ K2.ⓑ{I2}V2 →
∧∧ I1 = I2 & R K1 V1 V2 & llpx_sn R 0 V1 K1 K2 & S 0 V1 K1 K2
- ) → S d T L1 L2) →
- ∀L1,L2,T,d. llpx_sn R d T L1 L2 → S d T L1 L2.
-#R #S #IH1 #L1 #L2 #T #d #H lapply (llpx_sn_lpx_sn_alt_r … H) -H
-#H @(llpx_sn_alt_r_ind_alt … H) -L1 -L2 -T -d
-#L1 #L2 #T #d #HL12 #IH2 @IH1 -IH1 // -HL12
-#I1 #I2 #K1 #K2 #V1 #V2 #i #Hid #HnT #HLK1 #HLK2
+ ) → S l T L1 L2) →
+ ∀L1,L2,T,l. llpx_sn R l T L1 L2 → S l T L1 L2.
+#R #S #IH1 #L1 #L2 #T #l #H lapply (llpx_sn_lpx_sn_alt_r … H) -H
+#H @(llpx_sn_alt_r_ind_alt … H) -L1 -L2 -T -l
+#L1 #L2 #T #l #HL12 #IH2 @IH1 -IH1 // -HL12
+#I1 #I2 #K1 #K2 #V1 #V2 #i #Hil #HnT #HLK1 #HLK2
elim (IH2 … HnT HLK1 HLK2) -IH2 -HnT -HLK1 -HLK2 /3 width=1 by llpx_sn_alt_r_inv_lpx_sn, and4_intro/
qed-.
-lemma llpx_sn_inv_alt_r: ∀R,L1,L2,T,d. llpx_sn R d T L1 L2 →
+lemma llpx_sn_inv_alt_r: ∀R,L1,L2,T,l. llpx_sn R l T L1 L2 →
|L1| = |L2| ∧
- ∀I1,I2,K1,K2,V1,V2,i. d ≤ yinj i → (∀U. ⬆[i, 1] U ≡ T → ⊥) →
+ ∀I1,I2,K1,K2,V1,V2,i. l ≤ yinj i → (∀U. ⬆[i, 1] U ≡ T → ⊥) →
⬇[i] L1 ≡ K1.ⓑ{I1}V1 → ⬇[i] L2 ≡ K2.ⓑ{I2}V2 →
∧∧ I1 = I2 & R K1 V1 V2 & llpx_sn R 0 V1 K1 K2.
-#R #L1 #L2 #T #d #H lapply (llpx_sn_lpx_sn_alt_r … H) -H
+#R #L1 #L2 #T #l #H lapply (llpx_sn_lpx_sn_alt_r … H) -H
#H elim (llpx_sn_alt_r_inv_alt … H) -H
#HL12 #IH @conj //
-#I1 #I2 #K1 #K2 #V1 #V2 #i #Hid #HnT #HLK1 #HLK2
+#I1 #I2 #K1 #K2 #V1 #V2 #i #Hil #HnT #HLK1 #HLK2
elim (IH … HnT HLK1 HLK2) -IH -HnT -HLK1 -HLK2 /3 width=1 by llpx_sn_alt_r_inv_lpx_sn, and3_intro/
qed-.
(**************************************************************************)
include "basic_2/substitution/drop_drop.ma".
-include "basic_2/multiple/llpx_sn_leq.ma".
+include "basic_2/multiple/llpx_sn_lreq.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 →
+lemma llpx_sn_fwd_lref_dx: ∀R,L1,L2,l,i. llpx_sn R l (#i) L1 L2 →
∀I,K2,V2. ⬇[i] L2 ≡ K2.ⓑ{I}V2 →
- i < d ∨
+ i < l ∨
∃∃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 [ * || * ]
+ R K1 V1 V2 & l ≤ i.
+#R #L1 #L2 #l #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
+| #I #K11 #K22 #V11 #V22 #HLK11 #HLK22 #HK12 #HV12 #Hli
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 →
+lemma llpx_sn_fwd_lref_sn: ∀R,L1,L2,l,i. llpx_sn R l (#i) L1 L2 →
∀I,K1,V1. ⬇[i] L1 ≡ K1.ⓑ{I}V1 →
- i < d ∨
+ i < l ∨
∃∃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 [ * || * ]
+ R K1 V1 V2 & l ≤ i.
+#R #L1 #L2 #l #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
+| #I #K11 #K22 #V11 #V22 #HLK11 #HLK22 #HK12 #HV12 #Hli
lapply (drop_mono … HLK11 … HLK1) -L1 #H destruct
/3 width=5 by ex4_2_intro, or_intror/
]
(* Advanced inversion lemmas ************************************************)
-lemma llpx_sn_inv_lref_ge_dx: ∀R,L1,L2,d,i. llpx_sn R d (#i) L1 L2 → d ≤ i →
+lemma llpx_sn_inv_lref_ge_dx: ∀R,L1,L2,l,i. llpx_sn R l (#i) L1 L2 → l ≤ 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)
+#R #L1 #L2 #l #i #H #Hli #I #K2 #V2 #HLK2 elim (llpx_sn_fwd_lref_dx … H … HLK2) -L2
+[ #H elim (ylt_yle_false … H Hli)
| * /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 →
+lemma llpx_sn_inv_lref_ge_sn: ∀R,L1,L2,l,i. llpx_sn R l (#i) L1 L2 → l ≤ 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)
+#R #L1 #L2 #l #i #H #Hli #I #K1 #V1 #HLK1 elim (llpx_sn_fwd_lref_sn … H … HLK1) -L1
+[ #H elim (ylt_yle_false … H Hli)
| * /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 →
+lemma llpx_sn_inv_lref_ge_bi: ∀R,L1,L2,l,i. llpx_sn R l (#i) L1 L2 → l ≤ 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
+#R #L1 #L2 #l #i #HL12 #Hli #I1 #I2 #K1 #K2 #V1 #V2 #HLK1 #HLK2
+elim (llpx_sn_inv_lref_ge_sn … HL12 … HLK1) // -L1 -l
#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
+fact llpx_sn_inv_S_aux: ∀R,L1,L2,T,l0. llpx_sn R l0 T L1 L2 → ∀l. l0 = l + 1 →
+ ∀K1,K2,I,V1,V2. ⬇[l] L1 ≡ K1.ⓑ{I}V1 → ⬇[l] L2 ≡ K2.ⓑ{I}V2 →
+ llpx_sn R 0 V1 K1 K2 → R K1 V1 V2 → llpx_sn R l T L1 L2.
+#R #L1 #L2 #T #l0 #H elim H -L1 -L2 -T -l0
/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
+[ #L1 #L2 #l0 #i #HL12 #Hil #l #H #K1 #K2 #I #V1 #V2 #HLK1 #HLK2 #HK12 #HV12 destruct
+ elim (yle_split_eq i l) /2 width=1 by llpx_sn_skip, ylt_fwd_succ2/ -HL12 -Hil
#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
+| #I #L1 #L2 #K11 #K22 #V1 #V2 #l0 #i #Hl0i #HLK11 #HLK22 #HK12 #HV12 #_ #l #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
+| #a #I #L1 #L2 #V #T #l0 #_ #_ #IHV #IHT #l #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
+| #I #L1 #L2 #V #T #l0 #_ #_ #IHV #IHT #l #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.
+lemma llpx_sn_inv_S: ∀R,L1,L2,T,l. llpx_sn R (l + 1) T L1 L2 →
+ ∀K1,K2,I,V1,V2. ⬇[l] L1 ≡ K1.ⓑ{I}V1 → ⬇[l] L2 ≡ K2.ⓑ{I}V2 →
+ llpx_sn R 0 V1 K1 K2 → R K1 V1 V2 → llpx_sn R l T L1 L2.
/2 width=9 by llpx_sn_inv_S_aux/ qed-.
lemma llpx_sn_inv_bind_O: ∀R. (∀L. reflexive … (R L)) →
/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).
+ ∀T,L1,L2,l. Decidable (llpx_sn R l 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
+ [ #HL12 #l elim (ylt_split i l) /3 width=1 by llpx_sn_skip, or_introl/
+ #Hli 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
#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/
+| #a #I #V #T #Hn #L2 #l destruct
+ elim (IH L1 V … L2 l) /2 width=1 by/
+ elim (IH (L1.ⓑ{I}V) T … (L2.ⓑ{I}V) (⫯l)) -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/
+| #I #V #T #Hn #L2 #l destruct
+ elim (IH L1 V … L2 l) /2 width=1 by/
+ elim (IH L1 T … L2 l) -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/
]
(* 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
+lemma llpx_sn_lift_le: ∀R. d_liftable R →
+ ∀K1,K2,T,l0. llpx_sn R l0 T K1 K2 →
+ ∀L1,L2,l,m. ⬇[Ⓕ, l, m] L1 ≡ K1 → ⬇[Ⓕ, l, m] L2 ≡ K2 →
+ ∀U. ⬆[l, m] T ≡ U → l0 ≤ l → llpx_sn R l0 U L1 L2.
+#R #HR #K1 #K2 #T #l0 #H elim H -K1 -K2 -T -l0
+[ #K1 #K2 #l0 #k #HK12 #L1 #L2 #l #m #HLK1 #HLK2 #X #H #_ >(lift_inv_sort1 … H) -X
+ lapply (drop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 -l
/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
+| #K1 #K2 #l0 #i #HK12 #Hil0 #L1 #L2 #l #m #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_lref1 … H) -H
+ * #Hli #H destruct
+ [ lapply (drop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 -l
/2 width=1 by llpx_sn_skip/
- | elim (ylt_yle_false … Hid0) -L1 -L2 -K1 -K2 -e -Hid0
+ | elim (ylt_yle_false … Hil0) -L1 -L2 -K1 -K2 -m -Hil0
/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 ]
+| #I #K1 #K2 #K11 #K22 #V1 #V2 #l0 #i #Hil0 #HK11 #HK22 #HK12 #HV12 #IHK12 #L1 #L2 #l #m #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_lref1 … H) -H
+ * #Hli #H destruct [ -HK12 | -IHK12 ]
[ elim (drop_trans_lt … HLK1 … HK11) // -K1
- elim (drop_trans_lt … HLK2 … HK22) // -Hdi -K2
+ elim (drop_trans_lt … HLK2 … HK22) // -Hli -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
+ lapply (drop_trans_ge_comm … HLK2 … HK22 ?) // -Hli -Hl0 -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
+| #K1 #K2 #l0 #i #HK1 #HK2 #HK12 #L1 #L2 #l #m #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_lref1 … H) -H
+ * #Hil #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 *)
+ @llpx_sn_free [ >HLK1 | >HLK2 ] -Hil -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
+| #K1 #K2 #l0 #p #HK12 #L1 #L2 #l #m #HLK1 #HLK2 #X #H #_ >(lift_inv_gref1 … H) -X
+ lapply (drop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 -l -m
/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
+| #a #I #K1 #K2 #V #T #l0 #_ #_ #IHV #IHT #L1 #L2 #l #m #HLK1 #HLK2 #X #H #Hl0 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
+| #I #K1 #K2 #V #T #l0 #_ #_ #IHV #IHT #L1 #L2 #l #m #HLK1 #HLK2 #X #H #Hl0 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
+lemma llpx_sn_lift_ge: ∀R,K1,K2,T,l0. llpx_sn R l0 T K1 K2 →
+ ∀L1,L2,l,m. ⬇[Ⓕ, l, m] L1 ≡ K1 → ⬇[Ⓕ, l, m] L2 ≡ K2 →
+ ∀U. ⬆[l, m] T ≡ U → l ≤ l0 → llpx_sn R (l0+m) U L1 L2.
+#R #K1 #K2 #T #l0 #H elim H -K1 -K2 -T -l0
+[ #K1 #K2 #l0 #k #HK12 #L1 #L2 #l #m #HLK1 #HLK2 #X #H #_ >(lift_inv_sort1 … H) -X
+ lapply (drop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 -l
/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
+| #K1 #K2 #l0 #i #HK12 #Hil0 #L1 #L2 #l #m #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
+| #I #K1 #K2 #K11 #K22 #V1 #V2 #l0 #i #Hil0 #HK11 #HK22 #HK12 #HV12 #_ #L1 #L2 #l #m #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_lref1 … H) -H
+ * #Hil #H destruct
+ [ elim (ylt_yle_false … Hil0) -I -L1 -L2 -K1 -K2 -K11 -K22 -V1 -V2 -m -Hil0
/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
+ lapply (drop_trans_ge_comm … HLK2 … HK22 ?) // -Hil -Hl0 -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
+| #K1 #K2 #l0 #i #HK1 #HK2 #HK12 #L1 #L2 #l #m #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_lref1 … H) -H
+ * #Hil #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 *)
+ @llpx_sn_free [ >HLK1 | >HLK2 ] -Hil -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
+| #K1 #K2 #l0 #p #HK12 #L1 #L2 #l #m #HLK1 #HLK2 #X #H #_ >(lift_inv_gref1 … H) -X
+ lapply (drop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 -l
/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
+| #a #I #K1 #K2 #V #T #l0 #_ #_ #IHV #IHT #L1 #L2 #l #m #HLK1 #HLK2 #X #H #Hl0 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
+| #I #K1 #K2 #V #T #l0 #_ #_ #IHV #IHT #L1 #L2 #l #m #HLK1 #HLK2 #X #H #Hl0 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
+lemma llpx_sn_inv_lift_le: ∀R. d_deliftable_sn R →
+ ∀L1,L2,U,l0. llpx_sn R l0 U L1 L2 →
+ ∀K1,K2,l,m. ⬇[Ⓕ, l, m] L1 ≡ K1 → ⬇[Ⓕ, l, m] L2 ≡ K2 →
+ ∀T. ⬆[l, m] T ≡ U → l0 ≤ l → llpx_sn R l0 T K1 K2.
+#R #HR #L1 #L2 #U #l0 #H elim H -L1 -L2 -U -l0
+[ #L1 #L2 #l0 #k #HL12 #K1 #K2 #l #m #HLK1 #HLK2 #X #H #_ >(lift_inv_sort2 … H) -X
+ lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -l -m
/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
+| #L1 #L2 #l0 #i #HL12 #Hil0 #K1 #K2 #l #m #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 ]
+| #I #L1 #L2 #K11 #K22 #W1 #W2 #l0 #i #Hil0 #HLK11 #HLK22 #HK12 #HW12 #IHK12 #K1 #K2 #l #m #HLK1 #HLK2 #X #H #Hl0 elim (lift_inv_lref2 … H) -H
+ * #Hil #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 (drop_conf_lt … HLK2 … HLK22) // -Hil -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 *)
+ lapply (drop_conf_ge … HLK2 … HLK22 ?) // -L2 -Hil0
+ elim (le_inv_plus_l … Hil) -Hil /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
+| #L1 #L2 #l0 #i #HL1 #HL2 #HL12 #K1 #K2 #l #m #HLK1 #HLK2 #X #H #Hl0 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 … 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
+| #L1 #L2 #l0 #p #HL12 #K1 #K2 #l #m #HLK1 #HLK2 #X #H #_ >(lift_inv_gref2 … H) -X
+ lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -l -m
/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
+| #a #I #L1 #L2 #W #U #l0 #_ #_ #IHW #IHU #K1 #K2 #l #m #HLK1 #HLK2 #X #H #Hl0 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
+| #I #L1 #L2 #W #U #l0 #_ #_ #IHW #IHU #K1 #K2 #l #m #HLK1 #HLK2 #X #H #Hl0 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
+lemma llpx_sn_inv_lift_be: ∀R,L1,L2,U,l0. llpx_sn R l0 U L1 L2 →
+ ∀K1,K2,l,m. ⬇[Ⓕ, l, m] L1 ≡ K1 → ⬇[Ⓕ, l, m] L2 ≡ K2 →
+ ∀T. ⬆[l, m] T ≡ U → l ≤ l0 → l0 ≤ yinj l + m → llpx_sn R l T K1 K2.
+#R #L1 #L2 #U #l0 #H elim H -L1 -L2 -U -l0
+[ #L1 #L2 #l0 #k #HL12 #K1 #K2 #l #m #HLK1 #HLK2 #X #H #_ #_ >(lift_inv_sort2 … H) -X
+ lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -l0 -m
/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
+| #L1 #L2 #l0 #i #HL12 #Hil0 #K1 #K2 #l #m #HLK1 #HLK2 #X #H #Hl0 #Hl0m elim (lift_inv_lref2 … H) -H
+ * #Hil #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
+ -Hil0 /3 width=1 by llpx_sn_skip, ylt_inj/
+ | elim (ylt_yle_false … Hil0) -L1 -L2 -Hl0 -Hil0
/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
+| #I #L1 #L2 #K11 #K22 #W1 #W2 #l0 #i #Hil0 #HLK11 #HLK22 #HK12 #HW12 #_ #K1 #K2 #l #m #HLK1 #HLK2 #X #H #Hl0 #Hl0m elim (lift_inv_lref2 … H) -H
+ * #Hil #H destruct
+ [ elim (ylt_yle_false … Hil0) -I -L1 -L2 -K11 -K22 -W1 -W2 -Hl0m -Hil0
/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/
+ lapply (drop_conf_ge … HLK2 … HLK22 ?) // -L2 -Hil0 -Hl0 -Hl0m
+ elim (le_inv_plus_l … Hil) -Hil /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
+| #L1 #L2 #l0 #i #HL1 #HL2 #HL12 #K1 #K2 #l #m #HLK1 #HLK2 #X #H #Hl0 #Hl0m 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 … 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
+| #L1 #L2 #l0 #p #HL12 #K1 #K2 #l #m #HLK1 #HLK2 #X #H #_ #_ >(lift_inv_gref2 … H) -X
+ lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -l0 -m
/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
+| #a #I #L1 #L2 #W #U #l0 #_ #_ #IHW #IHU #K1 #K2 #l #m #HLK1 #HLK2 #X #H #Hl0 #Hl0m elim (lift_inv_bind2 … H) -H
>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
+| #I #L1 #L2 #W #U #l0 #_ #_ #IHW #IHU #K1 #K2 #l #m #HLK1 #HLK2 #X #H #Hl0 #Hl0m 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
+lemma llpx_sn_inv_lift_ge: ∀R,L1,L2,U,l0. llpx_sn R l0 U L1 L2 →
+ ∀K1,K2,l,m. ⬇[Ⓕ, l, m] L1 ≡ K1 → ⬇[Ⓕ, l, m] L2 ≡ K2 →
+ ∀T. ⬆[l, m] T ≡ U → yinj l + m ≤ l0 → llpx_sn R (l0-m) T K1 K2.
+#R #L1 #L2 #U #l0 #H elim H -L1 -L2 -U -l0
+[ #L1 #L2 #l0 #k #HL12 #K1 #K2 #l #m #HLK1 #HLK2 #X #H #_ >(lift_inv_sort2 … H) -X
+ lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -l
/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 ]
+| #L1 #L2 #l0 #i #HL12 #Hil0 #K1 #K2 #l #m #HLK1 #HLK2 #X #H #Hlml0 elim (lift_inv_lref2 … H) -H
+ * #Hil #H destruct [ -Hil0 | -Hlml0 ]
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 #_
+ | elim (le_inv_plus_l … Hil) -Hil #_
/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
+| #I #L1 #L2 #K11 #K22 #W1 #W2 #l0 #i #Hil0 #HLK11 #HLK22 #HK12 #HW12 #_ #K1 #K2 #l #m #HLK1 #HLK2 #X #H #Hlml0 elim (lift_inv_lref2 … H) -H
+ * #Hil #H destruct
+ [ elim (ylt_yle_false … Hil0) -I -L1 -L2 -K11 -K22 -W1 -W2 -Hil0
/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
+ lapply (drop_conf_ge … HLK2 … HLK22 ?) // -L2 -Hlml0 -Hil
/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
+| #L1 #L2 #l0 #i #HL1 #HL2 #HL12 #K1 #K2 #l #m #HLK1 #HLK2 #X #H #Hlml0 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 … 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
+| #L1 #L2 #l0 #p #HL12 #K1 #K2 #l #m #HLK1 #HLK2 #X #H #_ >(lift_inv_gref2 … H) -X
+ lapply (drop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -l
/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
+| #a #I #L1 #L2 #W #U #l0 #_ #_ #IHW #IHU #K1 #K2 #l #m #HLK1 #HLK2 #X #H #Hlml0 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
+| #I #L1 #L2 #W #U #l0 #_ #_ #IHW #IHU #K1 #K2 #l #m #HLK1 #HLK2 #X #H #Hlml0 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.
+ ∀K1,K2,m. ⬇[m] L1 ≡ K1 → ⬇[m] L2 ≡ K2 →
+ ∀T. ⬆[0, m] 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)
+ ∀K1,m. ⬇[m] L1 ≡ K1 → ∀T. ⬆[0, m] T ≡ U →
+ ∃∃K2. ⬇[m] L2 ≡ K2 & llpx_sn R 0 T K1 K2.
+#R #L1 #L2 #U #HU #K1 #m #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)
+ ∀K2,m. ⬇[m] L2 ≡ K2 → ∀T. ⬆[0, m] T ≡ U →
+ ∃∃K1. ⬇[m] L1 ≡ K1 & llpx_sn R 0 T K1 K2.
+#R #L1 #L2 #U #HU #K2 #m #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)
+ ∀a,I,L1,L2,V,T,l. (llpx_sn R l (ⓑ{a,I}V.T) L1 L2 → ⊥) →
+ (llpx_sn R l V L1 L2 → ⊥) ∨ (llpx_sn R (⫯l) T (L1.ⓑ{I}V) (L2.ⓑ{I}V) → ⊥).
+#R #HR #a #I #L1 #L2 #V #T #l #H elim (llpx_sn_dec … HR V L1 L2 l)
/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)
+ ∀I,L1,L2,V,T,l. (llpx_sn R l (ⓕ{I}V.T) L1 L2 → ⊥) →
+ (llpx_sn R l V L1 L2 → ⊥) ∨ (llpx_sn R l T L1 L2 → ⊥).
+#R #HR #I #L1 #L2 #V #T #l #H elim (llpx_sn_dec … HR V L1 L2 l)
/4 width=1 by llpx_sn_flat, or_intror, or_introl/
qed-.
(* Properties on context-sensitive free variables ***************************)
fact llpx_sn_frees_trans_aux: ∀R. (s_r_confluent1 … R (llpx_sn R 0)) → (frees_trans R) →
- ∀L2,U,d,i. L2 ⊢ i ϵ 𝐅*[d]⦃U⦄ →
- ∀L1. llpx_sn R d U L1 L2 → L1 ⊢ i ϵ 𝐅*[d]⦃U⦄.
-#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
+ ∀L2,U,l,i. L2 ⊢ i ϵ 𝐅*[l]⦃U⦄ →
+ ∀L1. llpx_sn R l U L1 L2 → L1 ⊢ i ϵ 𝐅*[l]⦃U⦄.
+#R #H1R #H2R #L2 #U #l #i #H elim H -L2 -U -l -i /3 width=2 by frees_eq/
+#I2 #L2 #K2 #U #W2 #l #i #j #Hlj #Hji #HnU #HLK2 #_ #IHW2 #L1 #HL12
elim (llpx_sn_inv_alt_r … HL12) -HL12 #HL12 #IH
lapply (drop_fwd_length_lt2 … HLK2) #Hj
elim (drop_O1_lt (Ⓕ) L1 j) // -Hj -HL12 #I1 #K1 #W1 #HLK1
qed-.
lemma llpx_sn_frees_trans: ∀R. (s_r_confluent1 … R (llpx_sn R 0)) → (frees_trans R) →
- ∀L1,L2,U,d. llpx_sn R d U L1 L2 →
- ∀i. L2 ⊢ i ϵ 𝐅*[d]⦃U⦄ → L1 ⊢ i ϵ 𝐅*[d]⦃U⦄.
+ ∀L1,L2,U,l. llpx_sn R l U L1 L2 →
+ ∀i. L2 ⊢ i ϵ 𝐅*[l]⦃U⦄ → L1 ⊢ i ϵ 𝐅*[l]⦃U⦄.
/2 width=6 by llpx_sn_frees_trans_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/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 →
- ∀L1. 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_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 →
- ∀L2. 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_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: ∀R,L1,L2,T,d. llpx_sn R d T L1 L2 → ∀K1. K1 ⩬[d, ∞] L1 →
- ∀K2. 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) →
- ∀J1,J2,W1,W2. llpx_sn R 1 T (L1.ⓑ{J1}W1) (L2.ⓑ{J2}W2).
-#R #I1 #I2 #L1 #L2 #V1 #V2 #T #HT #J1 #J2 #W1 #W2 lapply (llpx_sn_ge R … 1 … HT) -HT
-/3 width=7 by llpx_sn_leq_repl, leq_succ/
-qed-.
(* Inversion lemmas on pointwise union for local environments ****************)
lemma llpx_sn_llor_fwd_sn: ∀R. (∀L. reflexive … (R L)) →
- ∀L1,L2,T,d. llpx_sn R d T L1 L2 →
- ∀L. L1 ⋓[T, d] L2 ≡ L → lpx_sn R L1 L.
-#R #HR #L1 #L2 #T #d #H1 #L #H2
+ ∀L1,L2,T,l. llpx_sn R l T L1 L2 →
+ ∀L. L1 ⋓[T, l] L2 ≡ L → lpx_sn R L1 L.
+#R #HR #L1 #L2 #T #l #H1 #L #H2
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
(* Properties on pointwise extensions ***************************************)
lemma lpx_sn_llpx_sn: ∀R. (∀L. reflexive … (R L)) →
- ∀T,L1,L2,d. lpx_sn R L1 L2 → llpx_sn R d T L1 L2.
+ ∀T,L1,L2,l. lpx_sn R L1 L2 → llpx_sn R l T L1 L2.
#R #HR #T #L1 @(f2_ind … rfw … L1 T) -L1 -T
#n #IH #L1 * *
[ -HR -IH /4 width=2 by lpx_sn_fwd_length, llpx_sn_sort/
| -HR #i elim (lt_or_ge i (|L1|))
[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)
+ #Hi #Hn #L2 #l elim (ylt_split i l)
[ -n /3 width=2 by llpx_sn_skip, lpx_sn_fwd_length/ ]
- #Hdi #HL12 elim (drop_O1_lt (Ⓕ) L1 i) //
+ #Hli #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/
--- /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_lreq.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 lreq_llpx_sn_trans: ∀R,L2,L,T,l. llpx_sn R l T L2 L →
+ ∀L1. L1 ⩬[l, ∞] L2 → llpx_sn R l T L1 L.
+#R #L2 #L #T #l #H elim H -L2 -L -T -l
+/4 width=5 by llpx_sn_flat, llpx_sn_gref, llpx_sn_skip, llpx_sn_sort, lreq_fwd_length, trans_eq/
+[ #I #L2 #L #K2 #K #V2 #V #l #i #Hli #HLK2 #HLK #HK2 #HV2 #_ #L1 #HL12
+ elim (lreq_drop_trans_be … HL12 … HLK2) -L2 // >yminus_Y_inj #K1 #HK12 #HLK1
+ lapply (lreq_inv_O_Y … HK12) -HK12 #H destruct /2 width=9 by llpx_sn_lref/
+| /4 width=5 by llpx_sn_free, lreq_fwd_length, le_repl_sn_trans_aux, trans_eq/
+| /4 width=1 by llpx_sn_bind, lreq_succ/
+]
+qed-.
+
+lemma llpx_sn_lreq_trans: ∀R,L,L1,T,l. llpx_sn R l T L L1 →
+ ∀L2. L1 ⩬[l, ∞] L2 → llpx_sn R l T L L2.
+#R #L #L1 #T #l #H elim H -L -L1 -T -l
+/4 width=5 by llpx_sn_flat, llpx_sn_gref, llpx_sn_skip, llpx_sn_sort, lreq_fwd_length, trans_eq/
+[ #I #L #L1 #K #K1 #V #V1 #l #i #Hli #HLK #HLK1 #HK1 #HV1 #_ #L2 #HL12
+ elim (lreq_drop_conf_be … HL12 … HLK1) -L1 // >yminus_Y_inj #K2 #HK12 #HLK2
+ lapply (lreq_inv_O_Y … HK12) -HK12 #H destruct /2 width=9 by llpx_sn_lref/
+| /4 width=5 by llpx_sn_free, lreq_fwd_length, le_repl_sn_conf_aux, trans_eq/
+| /4 width=1 by llpx_sn_bind, lreq_succ/
+]
+qed-.
+
+lemma llpx_sn_lreq_repl: ∀R,L1,L2,T,l. llpx_sn R l T L1 L2 → ∀K1. K1 ⩬[l, ∞] L1 →
+ ∀K2. L2 ⩬[l, ∞] K2 → llpx_sn R l T K1 K2.
+/3 width=4 by llpx_sn_lreq_trans, lreq_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) →
+ ∀J1,J2,W1,W2. llpx_sn R 1 T (L1.ⓑ{J1}W1) (L2.ⓑ{J2}W2).
+#R #I1 #I2 #L1 #L2 #V1 #V2 #T #HT #J1 #J2 #W1 #W2 lapply (llpx_sn_ge R … 1 … HT) -HT
+/3 width=7 by llpx_sn_lreq_repl, lreq_succ/
+qed-.
inductive at: list2 nat nat → relation nat ≝
| at_nil: ∀i. at (◊) i i
-| at_lt : ∀des,d,e,i1,i2. i1 < d →
- at des i1 i2 → at ({d, e} @ des) i1 i2
-| at_ge : ∀des,d,e,i1,i2. d ≤ i1 →
- at des (i1 + e) i2 → at ({d, e} @ des) i1 i2
+| at_lt : ∀des,l,m,i1,i2. i1 < l →
+ at des i1 i2 → at ({l, m} @ des) i1 i2
+| at_ge : ∀des,l,m,i1,i2. l ≤ i1 →
+ at des (i1 + m) i2 → at ({l, m} @ des) i1 i2
.
interpretation "application (multiple relocation with pairs)"
fact at_inv_nil_aux: ∀des,i1,i2. @⦃i1, des⦄ ≡ i2 → des = ◊ → i1 = i2.
#des #i1 #i2 * -des -i1 -i2
[ //
-| #des #d #e #i1 #i2 #_ #_ #H destruct
-| #des #d #e #i1 #i2 #_ #_ #H destruct
+| #des #l #m #i1 #i2 #_ #_ #H destruct
+| #des #l #m #i1 #i2 #_ #_ #H destruct
]
qed-.
/2 width=3 by at_inv_nil_aux/ qed-.
fact at_inv_cons_aux: ∀des,i1,i2. @⦃i1, des⦄ ≡ i2 →
- ∀d,e,des0. des = {d, e} @ des0 →
- i1 < d ∧ @⦃i1, des0⦄ ≡ i2 ∨
- d ≤ i1 ∧ @⦃i1 + e, des0⦄ ≡ i2.
+ ∀l,m,des0. des = {l, m} @ des0 →
+ i1 < l ∧ @⦃i1, des0⦄ ≡ i2 ∨
+ l ≤ i1 ∧ @⦃i1 + m, des0⦄ ≡ i2.
#des #i1 #i2 * -des -i1 -i2
-[ #i #d #e #des #H destruct
-| #des1 #d1 #e1 #i1 #i2 #Hid1 #Hi12 #d2 #e2 #des2 #H destruct /3 width=1 by or_introl, conj/
-| #des1 #d1 #e1 #i1 #i2 #Hdi1 #Hi12 #d2 #e2 #des2 #H destruct /3 width=1 by or_intror, conj/
+[ #i #l #m #des #H destruct
+| #des1 #l1 #m1 #i1 #i2 #Hil1 #Hi12 #l2 #m2 #des2 #H destruct /3 width=1 by or_introl, conj/
+| #des1 #l1 #m1 #i1 #i2 #Hli1 #Hi12 #l2 #m2 #des2 #H destruct /3 width=1 by or_intror, conj/
]
qed-.
-lemma at_inv_cons: ∀des,d,e,i1,i2. @⦃i1, {d, e} @ des⦄ ≡ i2 →
- i1 < d ∧ @⦃i1, des⦄ ≡ i2 ∨
- d ≤ i1 ∧ @⦃i1 + e, des⦄ ≡ i2.
+lemma at_inv_cons: ∀des,l,m,i1,i2. @⦃i1, {l, m} @ des⦄ ≡ i2 →
+ i1 < l ∧ @⦃i1, des⦄ ≡ i2 ∨
+ l ≤ i1 ∧ @⦃i1 + m, des⦄ ≡ i2.
/2 width=3 by at_inv_cons_aux/ qed-.
-lemma at_inv_cons_lt: ∀des,d,e,i1,i2. @⦃i1, {d, e} @ des⦄ ≡ i2 →
- i1 < d → @⦃i1, des⦄ ≡ i2.
-#des #d #e #i1 #e2 #H
-elim (at_inv_cons … H) -H * // #Hdi1 #_ #Hi1d
-lapply (le_to_lt_to_lt … Hdi1 Hi1d) -Hdi1 -Hi1d #Hd
-elim (lt_refl_false … Hd)
+lemma at_inv_cons_lt: ∀des,l,m,i1,i2. @⦃i1, {l, m} @ des⦄ ≡ i2 →
+ i1 < l → @⦃i1, des⦄ ≡ i2.
+#des #l #m #i1 #m2 #H
+elim (at_inv_cons … H) -H * // #Hli1 #_ #Hi1l
+lapply (le_to_lt_to_lt … Hli1 Hi1l) -Hli1 -Hi1l #Hl
+elim (lt_refl_false … Hl)
qed-.
-lemma at_inv_cons_ge: ∀des,d,e,i1,i2. @⦃i1, {d, e} @ des⦄ ≡ i2 →
- d ≤ i1 → @⦃i1 + e, des⦄ ≡ i2.
-#des #d #e #i1 #e2 #H
-elim (at_inv_cons … H) -H * // #Hi1d #_ #Hdi1
-lapply (le_to_lt_to_lt … Hdi1 Hi1d) -Hdi1 -Hi1d #Hd
-elim (lt_refl_false … Hd)
+lemma at_inv_cons_ge: ∀des,l,m,i1,i2. @⦃i1, {l, m} @ des⦄ ≡ i2 →
+ l ≤ i1 → @⦃i1 + m, des⦄ ≡ i2.
+#des #l #m #i1 #m2 #H
+elim (at_inv_cons … H) -H * // #Hi1l #_ #Hli1
+lapply (le_to_lt_to_lt … Hli1 Hi1l) -Hli1 -Hi1l #Hl
+elim (lt_refl_false … Hl)
qed-.
inductive minuss: nat → relation (list2 nat nat) ≝
| minuss_nil: ∀i. minuss i (◊) (◊)
-| minuss_lt : ∀des1,des2,d,e,i. i < d → minuss i des1 des2 →
- minuss i ({d, e} @ des1) ({d - i, e} @ des2)
-| minuss_ge : ∀des1,des2,d,e,i. d ≤ i → minuss (e + i) des1 des2 →
- minuss i ({d, e} @ des1) des2
+| minuss_lt : ∀des1,des2,l,m,i. i < l → minuss i des1 des2 →
+ minuss i ({l, m} @ des1) ({l - i, m} @ des2)
+| minuss_ge : ∀des1,des2,l,m,i. l ≤ i → minuss (m + i) des1 des2 →
+ minuss i ({l, m} @ des1) des2
.
interpretation "minus (multiple relocation with pairs)"
fact minuss_inv_nil1_aux: ∀des1,des2,i. des1 ▭ i ≡ des2 → des1 = ◊ → des2 = ◊.
#des1 #des2 #i * -des1 -des2 -i
[ //
-| #des1 #des2 #d #e #i #_ #_ #H destruct
-| #des1 #des2 #d #e #i #_ #_ #H destruct
+| #des1 #des2 #l #m #i #_ #_ #H destruct
+| #des1 #des2 #l #m #i #_ #_ #H destruct
]
qed-.
/2 width=4 by minuss_inv_nil1_aux/ qed-.
fact minuss_inv_cons1_aux: ∀des1,des2,i. des1 ▭ i ≡ des2 →
- ∀d,e,des. des1 = {d, e} @ des →
- d ≤ i ∧ des ▭ e + i ≡ des2 ∨
- ∃∃des0. i < d & des ▭ i ≡ des0 &
- des2 = {d - i, e} @ des0.
+ ∀l,m,des. des1 = {l, m} @ des →
+ l ≤ i ∧ des ▭ m + i ≡ des2 ∨
+ ∃∃des0. i < l & des ▭ i ≡ des0 &
+ des2 = {l - i, m} @ des0.
#des1 #des2 #i * -des1 -des2 -i
-[ #i #d #e #des #H destruct
-| #des1 #des #d1 #e1 #i1 #Hid1 #Hdes #d2 #e2 #des2 #H destruct /3 width=3 by ex3_intro, or_intror/
-| #des1 #des #d1 #e1 #i1 #Hdi1 #Hdes #d2 #e2 #des2 #H destruct /3 width=1 by or_introl, conj/
+[ #i #l #m #des #H destruct
+| #des1 #des #l1 #m1 #i1 #Hil1 #Hcs #l2 #m2 #des2 #H destruct /3 width=3 by ex3_intro, or_intror/
+| #des1 #des #l1 #m1 #i1 #Hli1 #Hcs #l2 #m2 #des2 #H destruct /3 width=1 by or_introl, conj/
]
qed-.
-lemma minuss_inv_cons1: ∀des1,des2,d,e,i. {d, e} @ des1 ▭ i ≡ des2 →
- d ≤ i ∧ des1 ▭ e + i ≡ des2 ∨
- ∃∃des. i < d & des1 ▭ i ≡ des &
- des2 = {d - i, e} @ des.
+lemma minuss_inv_cons1: ∀des1,des2,l,m,i. {l, m} @ des1 ▭ i ≡ des2 →
+ l ≤ i ∧ des1 ▭ m + i ≡ des2 ∨
+ ∃∃des. i < l & des1 ▭ i ≡ des &
+ des2 = {l - i, m} @ des.
/2 width=3 by minuss_inv_cons1_aux/ qed-.
-lemma minuss_inv_cons1_ge: ∀des1,des2,d,e,i. {d, e} @ des1 ▭ i ≡ des2 →
- d ≤ i → des1 ▭ e + i ≡ des2.
-#des1 #des2 #d #e #i #H
-elim (minuss_inv_cons1 … H) -H * // #des #Hid #_ #_ #Hdi
-lapply (lt_to_le_to_lt … Hid Hdi) -Hid -Hdi #Hi
+lemma minuss_inv_cons1_ge: ∀des1,des2,l,m,i. {l, m} @ des1 ▭ i ≡ des2 →
+ l ≤ i → des1 ▭ m + i ≡ des2.
+#des1 #des2 #l #m #i #H
+elim (minuss_inv_cons1 … H) -H * // #des #Hil #_ #_ #Hli
+lapply (lt_to_le_to_lt … Hil Hli) -Hil -Hli #Hi
elim (lt_refl_false … Hi)
qed-.
-lemma minuss_inv_cons1_lt: ∀des1,des2,d,e,i. {d, e} @ des1 ▭ i ≡ des2 →
- i < d →
- ∃∃des. des1 ▭ i ≡ des & des2 = {d - i, e} @ des.
-#des1 #des2 #d #e #i #H elim (minuss_inv_cons1 … H) -H * /2 width=3 by ex2_intro/
-#Hdi #_ #Hid lapply (lt_to_le_to_lt … Hid Hdi) -Hid -Hdi
+lemma minuss_inv_cons1_lt: ∀des1,des2,l,m,i. {l, m} @ des1 ▭ i ≡ des2 →
+ i < l →
+ ∃∃des. des1 ▭ i ≡ des & des2 = {l - i, m} @ des.
+#des1 #des2 #l #m #i #H elim (minuss_inv_cons1 … H) -H * /2 width=3 by ex2_intro/
+#Hli #_ #Hil lapply (lt_to_le_to_lt … Hil Hli) -Hil -Hli
#Hi elim (lt_refl_false … Hi)
qed-.
theorem at_mono: ∀des,i,i1. @⦃i, des⦄ ≡ i1 → ∀i2. @⦃i, des⦄ ≡ i2 → i1 = i2.
#des #i #i1 #H elim H -des -i -i1
[ #i #x #H <(at_inv_nil … H) -x //
-| #des #d #e #i #i1 #Hid #_ #IHi1 #x #H
- lapply (at_inv_cons_lt … H Hid) -H -Hid /2 width=1 by/
-| #des #d #e #i #i1 #Hdi #_ #IHi1 #x #H
- lapply (at_inv_cons_ge … H Hdi) -H -Hdi /2 width=1 by/
+| #des #l #m #i #i1 #Hil #_ #IHi1 #x #H
+ lapply (at_inv_cons_lt … H Hil) -H -Hil /2 width=1 by/
+| #des #l #m #i #i1 #Hli #_ #IHi1 #x #H
+ lapply (at_inv_cons_ge … H Hli) -H -Hli /2 width=1 by/
]
qed-.
let rec pluss (des:list2 nat nat) (i:nat) on des ≝ match des with
[ nil2 ⇒ ◊
-| cons2 d e des ⇒ {d + i, e} @ pluss des i
+| cons2 l m des ⇒ {l + i, m} @ pluss des i
].
interpretation "plus (multiple relocation with pairs)"
lemma pluss_inv_nil2: ∀i,des. des + i = ◊ → des = ◊.
#i * // normalize
-#d #e #des #H destruct
+#l #m #des #H destruct
qed.
-lemma pluss_inv_cons2: ∀i,d,e,des2,des. des + i = {d, e} @ des2 →
- ∃∃des1. des1 + i = des2 & des = {d - i, e} @ des1.
-#i #d #e #des2 * normalize
+lemma pluss_inv_cons2: ∀i,l,m,des2,des. des + i = {l, m} @ des2 →
+ ∃∃des1. des1 + i = des2 & des = {l - i, m} @ des1.
+#i #l #m #des2 * normalize
[ #H destruct
-| #d1 #e1 #des1 #H destruct /2 width=3/
+| #l1 #m1 #des1 #H destruct /2 width=3/
]
qed-.
X,Y,Z : reserved: transient objet denoted by a capital letter
a,b : binder polarity
-c : reserved: future use (\lambda\delta 3)
-d : relocation depth
-e : relocation height
+c : relocation
+d : term degree
+e : reserved: future use (\lambda\delta 3)
f :
g : sort degree parameter
h : sort hierarchy parameter
i,j : local reference position index (de Bruijn's)
k : sort index
-l : term degree
-m,n : reserved: future use
+l : relocation depth
+m : relocation height
+n : type iterations
o :
p,q : global reference position
r : reduction kind parameter (true = ordinary, false = extended)
(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-notation "hvbox( G ⊢ ~ ⬊ * break [ term 46 h , break term 46 g , break term 46 d ] break term 46 L )"
+notation "hvbox( G ⊢ ~ ⬊ * break [ term 46 h , break term 46 g , break term 46 l ] break term 46 L )"
non associative with precedence 45
- for @{ 'CoSN $h $g $d $G $L }.
+ for @{ 'CoSN $h $g $l $G $L }.
(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T ▪ break [ term 46 h , break term 46 g ] break term 46 l )"
+notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T ▪ break [ term 46 h , break term 46 g ] break term 46 d )"
non associative with precedence 45
- for @{ 'Degree $h $g $G $L $T $l }.
+ for @{ 'Degree $h $g $G $L $T $d }.
(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 • * ⬌ * break [ term 46 h , break term 46 g , break term 46 l1 , break term 46 l2 ] break term 46 T2 )"
+notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 • * ⬌ * break [ term 46 h , break term 46 g , break term 46 n1 , break term 46 n2 ] break term 46 T2 )"
non associative with precedence 45
- for @{ 'DPConvStar $h $g $l1 $l2 $G $L $T1 $T2 }.
+ for @{ 'DPConvStar $h $g $n1 $n2 $G $L $T1 $T2 }.
(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 • * ➡ * break [ term 46 h , break term 46 g , break term 46 l ] break term 46 T2 )"
+notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 • * ➡ * break [ term 46 h , break term 46 g , break term 46 n ] break term 46 T2 )"
non associative with precedence 45
- for @{ 'DPRedStar $h $g $l $G $L $T1 $T2 }.
+ for @{ 'DPRedStar $h $g $n $G $L $T1 $T2 }.
(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-notation "hvbox( L ⊢ break term 46 i ϵ 𝐅 * [ break term 46 d ] ⦃ break term 46 T ⦄ )"
+notation "hvbox( L ⊢ break term 46 i ϵ 𝐅 * [ break term 46 l ] ⦃ break term 46 T ⦄ )"
non associative with precedence 45
- for @{ 'FreeStar $L $i $d $T }.
+ for @{ 'FreeStar $L $i $l $T }.
(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-notation "hvbox( L1 ≡ break [ term 46 T , break term 46 d ] break term 46 L2 )"
+notation "hvbox( L1 ≡ break [ term 46 T , break term 46 l ] break term 46 L2 )"
non associative with precedence 45
- for @{ 'LazyEq $T $d $L1 $L2 }.
+ for @{ 'LazyEq $T $l $L1 $L2 }.
(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-notation "hvbox( ⦃ term 46 G1, break term 46 L1, break term 46 T1 ⦄ ≡ break [ term 46 d ] break ⦃ term 46 G2, break term 46 L2 , break term 46 T2 ⦄ )"
+notation "hvbox( ⦃ term 46 G1, break term 46 L1, break term 46 T1 ⦄ ≡ break [ term 46 l ] break ⦃ term 46 G2, break term 46 L2 , break term 46 T2 ⦄ )"
non associative with precedence 45
- for @{ 'LazyEq $d $G1 $L1 $T1 $G2 $L2 $T2 }.
+ for @{ 'LazyEq $l $G1 $L1 $T1 $G2 $L2 $T2 }.
(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-notation "hvbox( L1 ⋓ break [ term 46 T , break term 46 d ] break term 46 L2 ≡ break term 46 L )"
+notation "hvbox( L1 ⋓ break [ term 46 T , break term 46 l ] break term 46 L2 ≡ break term 46 L )"
non associative with precedence 45
- for @{ 'LazyOr $L1 $T $d $L2 $L }.
+ for @{ 'LazyOr $L1 $T $l $L2 $L }.
(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-notation "hvbox( L1 break ⊆ [ term 46 d , break term 46 e ] break term 46 L2 )"
+notation "hvbox( L1 break ⊆ [ term 46 l , break term 46 m ] break term 46 L2 )"
non associative with precedence 45
- for @{ 'LRSubEq $L1 $d $e $L2 }.
+ for @{ 'LRSubEq $L1 $l $m $L2 }.
(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-notation "hvbox( L1 ⩬ break [ term 46 d , break term 46 e ] break term 46 L2 )"
+notation "hvbox( L1 ⩬ break [ term 46 l , break term 46 m ] break term 46 L2 )"
non associative with precedence 45
- for @{ 'MidIso $d $e $L1 $L2 }.
+ for @{ 'MidIso $l $m $L1 $L2 }.
(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T ¡ break [ term 46 h , break term 46 g , break term 46 l ] )"
+notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T ¡ break [ term 46 h , break term 46 g , break term 46 d ] )"
non associative with precedence 45
- for @{ 'NativeValid $h $g $l $G $L $T }.
+ for @{ 'NativeValid $h $g $d $G $L $T }.
(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 break ▶ [ term 46 d , break term 46 e ] break term 46 T2 )"
+notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 break ▶ [ term 46 l , break term 46 m ] break term 46 T2 )"
non associative with precedence 45
- for @{ 'PSubst $G $L $T1 $d $e $T2 }.
+ for @{ 'PSubst $G $L $T1 $l $m $T2 }.
(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 break ▶ * [ term 46 d , break term 46 e ] break term 46 T2 )"
+notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 break ▶ * [ term 46 l , break term 46 m ] break term 46 T2 )"
non associative with precedence 45
- for @{ 'PSubstStar $G $L $T1 $d $e $T2 }.
+ for @{ 'PSubstStar $G $L $T1 $l $m $T2 }.
(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 break ▶ ▶ * [ term 46 d , break term 46 e ] break term 46 T2 )"
+notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 break ▶ ▶ * [ term 46 l , break term 46 m ] break term 46 T2 )"
non associative with precedence 45
- for @{ 'PSubstStarAlt $G $L $T1 $d $e $T2 }.
+ for @{ 'PSubstStarAlt $G $L $T1 $l $m $T2 }.
(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-notation "hvbox( ⬇ [ term 46 e ] break term 46 L1 ≡ break term 46 L2 )"
+notation "hvbox( ⬇ [ term 46 m ] break term 46 L1 ≡ break term 46 L2 )"
non associative with precedence 45
- for @{ 'RDrop $e $L1 $L2 }.
+ for @{ 'RDrop $m $L1 $L2 }.
(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-notation "hvbox( ⬇ [ term 46 d , break term 46 e ] break term 46 L1 ≡ break term 46 L2 )"
+notation "hvbox( ⬇ [ term 46 l , break term 46 m ] break term 46 L1 ≡ break term 46 L2 )"
non associative with precedence 45
- for @{ 'RDrop $d $e $L1 $L2 }.
+ for @{ 'RDrop $l $m $L1 $L2 }.
(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-notation "hvbox( ⬇ [ term 46 s , break term 46 d , break term 46 e ] break term 46 L1 ≡ break term 46 L2 )"
+notation "hvbox( ⬇ [ term 46 s , break term 46 l , break term 46 m ] break term 46 L1 ≡ break term 46 L2 )"
non associative with precedence 45
- for @{ 'RDrop $s $d $e $L1 $L2 }.
+ for @{ 'RDrop $s $l $m $L1 $L2 }.
(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-notation "hvbox( ⬇ * [ term 46 e ] break term 46 L1 ≡ break term 46 L2 )"
+notation "hvbox( ⬇ * [ term 46 m ] break term 46 L1 ≡ break term 46 L2 )"
non associative with precedence 45
- for @{ 'RDropStar $e $L1 $L2 }.
+ for @{ 'RDropStar $m $L1 $L2 }.
(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-notation "hvbox( ⬇ * [ term 46 s , break term 46 e ] break term 46 L1 ≡ break term 46 L2 )"
+notation "hvbox( ⬇ * [ term 46 s , break term 46 m ] break term 46 L1 ≡ break term 46 L2 )"
non associative with precedence 45
- for @{ 'RDropStar $s $e $L1 $L2 }.
+ for @{ 'RDropStar $s $m $L1 $L2 }.
(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-notation "hvbox( ⬆ [ term 46 d , break term 46 e ] break term 46 T1 ≡ break term 46 T2 )"
+notation "hvbox( ⬆ [ term 46 l , break term 46 m ] break term 46 T1 ≡ break term 46 T2 )"
non associative with precedence 45
- for @{ 'RLift $d $e $T1 $T2 }.
+ for @{ 'RLift $l $m $T1 $T2 }.
(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-notation "hvbox( ⬆ * [ term 46 e ] break term 46 T1 ≡ break term 46 T2 )"
+notation "hvbox( ⬆ * [ term 46 m ] break term 46 T1 ≡ break term 46 T2 )"
non associative with precedence 45
- for @{ 'RLiftStar $e $T1 $T2 }.
+ for @{ 'RLiftStar $m $T1 $T2 }.
(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-notation "hvbox( G ⊢ ⬊ * break [ term 46 h , break term 46 g , break term 46 T , break term 46 d ] break term 46 L )"
+notation "hvbox( G ⊢ ⬊ * break [ term 46 h , break term 46 g , break term 46 T , break term 46 l ] break term 46 L )"
non associative with precedence 45
- for @{ 'SN $h $g $T $d $G $L }.
+ for @{ 'SN $h $g $T $l $G $L }.
(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-notation "hvbox( G ⊢ ⬊ ⬊ * break [ term 46 h , break term 46 g , break term 46 T , break term 46 d ] break term 46 L )"
+notation "hvbox( G ⊢ ⬊ ⬊ * break [ term 46 h , break term 46 g , break term 46 T , break term 46 l ] break term 46 L )"
non associative with precedence 45
- for @{ 'SNAlt $h $g $T $d $G $L }.
+ for @{ 'SNAlt $h $g $T $l $G $L }.
(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 •* break [ term 46 h , break term 46 l ] break term 46 T2 )"
+notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 •* break [ term 46 h , break term 46 n ] break term 46 T2 )"
non associative with precedence 45
- for @{ 'StaticTypeStar $h $G $L $l $T1 $T2 }.
+ for @{ 'StaticTypeStar $h $G $L $n $T1 $T2 }.
(* Properties on relocation *************************************************)
-lemma cir_lift: ∀G,K,T. ⦃G, K⦄ ⊢ ➡ 𝐈⦃T⦄ → ∀L,s,d,e. ⬇[s, d, e] L ≡ K →
- ∀U. ⬆[d, e] T ≡ U → ⦃G, L⦄ ⊢ ➡ 𝐈⦃U⦄.
+lemma cir_lift: ∀G,K,T. ⦃G, K⦄ ⊢ ➡ 𝐈⦃T⦄ → ∀L,s,l,m. ⬇[s, l, m] L ≡ K →
+ ∀U. ⬆[l, m] T ≡ U → ⦃G, L⦄ ⊢ ➡ 𝐈⦃U⦄.
/3 width=8 by crr_inv_lift/ qed.
-lemma cir_inv_lift: ∀G,L,U. ⦃G, L⦄ ⊢ ➡ 𝐈⦃U⦄ → ∀K,s,d,e. ⬇[s, d, e] L ≡ K →
- ∀T. ⬆[d, e] T ≡ U → ⦃G, K⦄ ⊢ ➡ 𝐈⦃T⦄.
+lemma cir_inv_lift: ∀G,L,U. ⦃G, L⦄ ⊢ ➡ 𝐈⦃U⦄ → ∀K,s,l,m. ⬇[s, l, m] L ≡ K →
+ ∀T. ⬆[l, m] T ≡ U → ⦃G, K⦄ ⊢ ➡ 𝐈⦃T⦄.
/3 width=8 by crr_lift/ qed-.
(* Basic inversion lemmas ***************************************************)
-lemma cix_inv_sort: ∀h,g,G,L,k,l. deg h g k (l+1) → ⦃G, L⦄ ⊢ ➡[h, g] 𝐈⦃⋆k⦄ → ⊥.
+lemma cix_inv_sort: ∀h,g,G,L,k,d. deg h g k (d+1) → ⦃G, L⦄ ⊢ ➡[h, g] 𝐈⦃⋆k⦄ → ⊥.
/3 width=2 by crx_sort/ qed-.
lemma cix_inv_delta: ∀h,g,I,G,L,K,V,i. ⬇[i] L ≡ K.ⓑ{I}V → ⦃G, L⦄ ⊢ ➡[h, g] 𝐈⦃#i⦄ → ⊥.
(* Basic properties *********************************************************)
lemma cix_sort: ∀h,g,G,L,k. deg h g k 0 → ⦃G, L⦄ ⊢ ➡[h, g] 𝐈⦃⋆k⦄.
-#h #g #G #L #k #Hk #H elim (crx_inv_sort … H) -L #l #Hkl
-lapply (deg_mono … Hk Hkl) -h -k <plus_n_Sm #H destruct
+#h #g #G #L #k #Hk #H elim (crx_inv_sort … H) -L #d #Hkd
+lapply (deg_mono … Hk Hkd) -h -k <plus_n_Sm #H destruct
qed.
lemma tix_lref: ∀h,g,G,i. ⦃G, ⋆⦄ ⊢ ➡[h, g] 𝐈⦃#i⦄.
-#h #g #G #i #H elim (trx_inv_atom … H) -H #k #l #_ #H destruct
+#h #g #G #i #H elim (trx_inv_atom … H) -H #k #d #_ #H destruct
qed.
lemma cix_gref: ∀h,g,G,L,p. ⦃G, L⦄ ⊢ ➡[h, g] 𝐈⦃§p⦄.
(* Properties on relocation *************************************************)
-lemma cix_lift: ∀h,g,G,K,T. ⦃G, K⦄ ⊢ ➡[h, g] 𝐈⦃T⦄ → ∀L,s,d,e. ⬇[s, d, e] L ≡ K →
- ∀U. ⬆[d, e] T ≡ U → ⦃G, L⦄ ⊢ ➡[h, g] 𝐈⦃U⦄.
+lemma cix_lift: ∀h,g,G,K,T. ⦃G, K⦄ ⊢ ➡[h, g] 𝐈⦃T⦄ → ∀L,s,l,m. ⬇[s, l, m] L ≡ K →
+ ∀U. ⬆[l, m] T ≡ U → ⦃G, L⦄ ⊢ ➡[h, g] 𝐈⦃U⦄.
/3 width=8 by crx_inv_lift/ qed.
-lemma cix_inv_lift: ∀h,g,G,L,U. ⦃G, L⦄ ⊢ ➡[h, g] 𝐈⦃U⦄ → ∀K,s,d,e. ⬇[s, d, e] L ≡ K →
- ∀T. ⬆[d, e] T ≡ U → ⦃G, K⦄ ⊢ ➡[h, g] 𝐈⦃T⦄.
+lemma cix_inv_lift: ∀h,g,G,L,U. ⦃G, L⦄ ⊢ ➡[h, g] 𝐈⦃U⦄ → ∀K,s,l,m. ⬇[s, l, m] L ≡ K →
+ ∀T. ⬆[l, m] T ≡ U → ⦃G, K⦄ ⊢ ➡[h, g] 𝐈⦃T⦄.
/3 width=8 by crx_lift/ qed-.
(* Relocation properties ****************************************************)
(* Basic_1: was: nf2_lift *)
-lemma cnr_lift: ∀G,L0,L,T,T0,s,d,e. ⦃G, L⦄ ⊢ ➡ 𝐍⦃T⦄ →
- ⬇[s, d, e] L0 ≡ L → ⬆[d, e] T ≡ T0 → ⦃G, L0⦄ ⊢ ➡ 𝐍⦃T0⦄.
-#G #L0 #L #T #T0 #s #d #e #HLT #HL0 #HT0 #X #H
+lemma cnr_lift: ∀G,L0,L,T,T0,s,l,m. ⦃G, L⦄ ⊢ ➡ 𝐍⦃T⦄ →
+ ⬇[s, l, m] L0 ≡ L → ⬆[l, m] T ≡ T0 → ⦃G, L0⦄ ⊢ ➡ 𝐍⦃T0⦄.
+#G #L0 #L #T #T0 #s #l #m #HLT #HL0 #HT0 #X #H
elim (cpr_inv_lift1 … H … HL0 … HT0) -L0 #T1 #HT10 #HT1
<(HLT … HT1) in HT0; -L #HT0
>(lift_mono … HT10 … HT0) -T1 -X //
qed.
(* Note: this was missing in basic_1 *)
-lemma cnr_inv_lift: ∀G,L0,L,T,T0,s,d,e. ⦃G, L0⦄ ⊢ ➡ 𝐍⦃T0⦄ →
- ⬇[s, d, e] L0 ≡ L → ⬆[d, e] T ≡ T0 → ⦃G, L⦄ ⊢ ➡ 𝐍⦃T⦄.
-#G #L0 #L #T #T0 #s #d #e #HLT0 #HL0 #HT0 #X #H
-elim (lift_total X d e) #X0 #HX0
+lemma cnr_inv_lift: ∀G,L0,L,T,T0,s,l,m. ⦃G, L0⦄ ⊢ ➡ 𝐍⦃T0⦄ →
+ ⬇[s, l, m] L0 ≡ L → ⬆[l, m] T ≡ T0 → ⦃G, L⦄ ⊢ ➡ 𝐍⦃T⦄.
+#G #L0 #L #T #T0 #s #l #m #HLT0 #HL0 #HT0 #X #H
+elim (lift_total X l m) #X0 #HX0
lapply (cpr_lift … H … HL0 … HT0 … HX0) -L #HTX0
>(HLT0 … HTX0) in HX0; -L0 -X0 #H
->(lift_inj … H … HT0) -T0 -X -s -d -e //
+>(lift_inj … H … HT0) -T0 -X -s -l -m //
qed-.
lemma cnx_inv_sort: ∀h,g,G,L,k. ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃⋆k⦄ → deg h g k 0.
#h #g #G #L #k #H elim (deg_total h g k)
-#l @(nat_ind_plus … l) -l // #l #_ #Hkl
-lapply (H (⋆(next h k)) ?) -H /2 width=2 by cpx_st/ -L -l #H destruct -H -e0 (**) (* destruct does not remove some premises *)
-lapply (next_lt h k) >e1 -e1 #H elim (lt_refl_false … H)
+#d @(nat_ind_plus … d) -d // #d #_ #Hkd
+lapply (H (⋆(next h k)) ?) -H /2 width=2 by cpx_st/ -L -d #H
+lapply (destruct_tatom_tatom_aux … H) -H #H (**) (* destruct lemma needed *)
+lapply (destruct_sort_sort_aux … H) -H #H (**) (* destruct lemma needed *)
+lapply (next_lt h k) >H -H #H elim (lt_refl_false … H)
qed-.
lemma cnx_inv_delta: ∀h,g,I,G,L,K,V,i. ⬇[i] L ≡ K.ⓑ{I}V → ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃#i⦄ → ⊥.
(* Basic properties *********************************************************)
lemma cnx_sort: ∀h,g,G,L,k. deg h g k 0 → ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃⋆k⦄.
-#h #g #G #L #k #Hk #X #H elim (cpx_inv_sort1 … H) -H // * #l #Hkl #_
-lapply (deg_mono … Hkl Hk) -h -L <plus_n_Sm #H destruct
+#h #g #G #L #k #Hk #X #H elim (cpx_inv_sort1 … H) -H // * #d #Hkd #_
+lapply (deg_mono … Hkd Hk) -h -L <plus_n_Sm #H destruct
qed.
-lemma cnx_sort_iter: ∀h,g,G,L,k,l. deg h g k l → ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃⋆((next h)^l k)⦄.
-#h #g #G #L #k #l #Hkl
-lapply (deg_iter … l Hkl) -Hkl <minus_n_n /2 width=6 by cnx_sort/
+lemma cnx_sort_iter: ∀h,g,G,L,k,d. deg h g k d → ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃⋆((next h)^d k)⦄.
+#h #g #G #L #k #d #Hkd
+lapply (deg_iter … d Hkd) -Hkd <minus_n_n /2 width=6 by cnx_sort/
qed.
lemma cnx_lref_free: ∀h,g,G,L,i. |L| ≤ i → ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃#i⦄.
(* Note: this property is unusual *)
lemma cnx_inv_crx: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ ➡[h, g] 𝐑⦃T⦄ → ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃T⦄ → ⊥.
#h #g #G #L #T #H elim H -L -T
-[ #L #k #l #Hkl #H
+[ #L #k #d #Hkd #H
lapply (cnx_inv_sort … H) -H #H
- lapply (deg_mono … H Hkl) -h -L -k <plus_n_Sm #H destruct
+ lapply (deg_mono … H Hkd) -h -L -k <plus_n_Sm #H destruct
| #I #L #K #V #i #HLK #H
elim (cnx_inv_delta … HLK H)
| #L #V #T #_ #IHV #H
(* Relocation properties ****************************************************)
-lemma cnx_lift: ∀h,g,G,L0,L,T,T0,s,d,e. ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃T⦄ → ⬇[s, d, e] L0 ≡ L →
- ⬆[d, e] T ≡ T0 → ⦃G, L0⦄ ⊢ ➡[h, g] 𝐍⦃T0⦄.
-#h #g #G #L0 #L #T #T0 #s #d #e #HLT #HL0 #HT0 #X #H
+lemma cnx_lift: ∀h,g,G,L0,L,T,T0,s,l,m. ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃T⦄ → ⬇[s, l, m] L0 ≡ L →
+ ⬆[l, m] T ≡ T0 → ⦃G, L0⦄ ⊢ ➡[h, g] 𝐍⦃T0⦄.
+#h #g #G #L0 #L #T #T0 #s #l #m #HLT #HL0 #HT0 #X #H
elim (cpx_inv_lift1 … H … HL0 … HT0) -L0 #T1 #HT10 #HT1
<(HLT … HT1) in HT0; -L #HT0
>(lift_mono … HT10 … HT0) -T1 -X //
qed.
-lemma cnx_inv_lift: ∀h,g,G,L0,L,T,T0,s,d,e. ⦃G, L0⦄ ⊢ ➡[h, g] 𝐍⦃T0⦄ → ⬇[s, d, e] L0 ≡ L →
- ⬆[d, e] T ≡ T0 → ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃T⦄.
-#h #g #G #L0 #L #T #T0 #s #d #e #HLT0 #HL0 #HT0 #X #H
-elim (lift_total X d e) #X0 #HX0
+lemma cnx_inv_lift: ∀h,g,G,L0,L,T,T0,s,l,m. ⦃G, L0⦄ ⊢ ➡[h, g] 𝐍⦃T0⦄ → ⬇[s, l, m] L0 ≡ L →
+ ⬆[l, m] T ≡ T0 → ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃T⦄.
+#h #g #G #L0 #L #T #T0 #s #l #m #HLT0 #HL0 #HT0 #X #H
+elim (lift_total X l m) #X0 #HX0
lapply (cpx_lift … H … HL0 … HT0 … HX0) -L #HTX0
>(HLT0 … HTX0) in HX0; -L0 -X0 #H
->(lift_inj … H … HT0) -T0 -X -d -e //
+>(lift_inj … H … HT0) -T0 -X -l -m //
qed-.
∀T. ⦃G, L⦄ ⊢ ②{I}V1.T ➡ ②{I}V2.T.
* /2 width=1 by cpr_bind, cpr_flat/ qed.
-lemma cpr_delift: ∀G,K,V,T1,L,d. ⬇[d] L ≡ (K.ⓓV) →
- ∃∃T2,T. ⦃G, L⦄ ⊢ T1 ➡ T2 & ⬆[d, 1] T ≡ T2.
+lemma cpr_delift: ∀G,K,V,T1,L,l. ⬇[l] L ≡ (K.ⓓV) →
+ ∃∃T2,T. ⦃G, L⦄ ⊢ T1 ➡ T2 & ⬆[l, 1] T ≡ T2.
#G #K #V #T1 elim T1 -T1
[ * /2 width=4 by cpr_atom, lift_sort, lift_gref, ex2_2_intro/
- #i #L #d #HLK elim (lt_or_eq_or_gt i d)
- #Hid [1,3: /3 width=4 by cpr_atom, lift_lref_ge_minus, lift_lref_lt, ex2_2_intro/ ]
+ #i #L #l #HLK elim (lt_or_eq_or_gt i l)
+ #Hil [1,3: /3 width=4 by cpr_atom, lift_lref_ge_minus, lift_lref_lt, ex2_2_intro/ ]
destruct
elim (lift_total V 0 (i+1)) #W #HVW
elim (lift_split … HVW i i) /3 width=6 by cpr_delta, ex2_2_intro/
-| * [ #a ] #I #W1 #U1 #IHW1 #IHU1 #L #d #HLK
+| * [ #a ] #I #W1 #U1 #IHW1 #IHU1 #L #l #HLK
elim (IHW1 … HLK) -IHW1 #W2 #W #HW12 #HW2
- [ elim (IHU1 (L. ⓑ{I}W1) (d+1)) -IHU1 /3 width=9 by drop_drop, cpr_bind, lift_bind, ex2_2_intro/
+ [ elim (IHU1 (L. ⓑ{I}W1) (l+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/
]
]
qed-.
-fact lstas_cpr_aux: ∀h,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 •*[h, l] T2 →
- l = 0 → ⦃G, L⦄ ⊢ T1 ➡ T2.
-#h #G #L #T1 #T2 #l #H elim H -G -L -T1 -T2 -l
+fact lstas_cpr_aux: ∀h,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 •*[h, d] T2 →
+ d = 0 → ⦃G, L⦄ ⊢ T1 ➡ T2.
+#h #G #L #T1 #T2 #d #H elim H -G -L -T1 -T2 -d
/3 width=1 by cpr_eps, cpr_flat, cpr_bind/
-[ #G #L #l #k #H0 destruct normalize //
-| #G #L #K #V1 #V2 #W2 #i #l #HLK #_ #HVW2 #IHV12 #H destruct
+[ #G #L #d #k #H0 destruct normalize //
+| #G #L #K #V1 #V2 #W2 #i #d #HLK #_ #HVW2 #IHV12 #H destruct
/3 width=6 by cpr_delta/
-| #G #L #K #V1 #V2 #W2 #i #l #_ #_ #_ #_ <plus_n_Sm #H destruct
+| #G #L #K #V1 #V2 #W2 #i #d #_ #_ #_ #_ <plus_n_Sm #H destruct
]
qed-.
(* Relocation properties ****************************************************)
(* Basic_1: includes: pr0_lift pr2_lift *)
-lemma cpr_lift: ∀G. l_liftable (cpr G).
+lemma cpr_lift: ∀G. d_liftable (cpr G).
#G #K #T1 #T2 #H elim H -G -K -T1 -T2
-[ #I #G #K #L #s #d #e #_ #U1 #H1 #U2 #H2
+[ #I #G #K #L #s #l #m #_ #U1 #H1 #U2 #H2
>(lift_mono … H1 … H2) -H1 -H2 //
-| #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
+| #G #K #KV #V #V2 #W2 #i #HKV #HV2 #HVW2 #IHV2 #L #s #l #m #HLK #U1 #H #U2 #HWU2
+ elim (lift_inv_lref1 … H) * #Hil #H destruct
[ elim (lift_trans_ge … HVW2 … HWU2) -W2 // <minus_plus #W2 #HVW2 #HWU2
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
+ elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hil
#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 (drop_trans_ge_comm … HLK … HKV ?) -K // -Hid /3 width=6 by cpr_delta, drop_inv_gen/
+ lapply (drop_trans_ge_comm … HLK … HKV ?) -K // -Hil /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
+| #a #I #G #K #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L #s #l #m #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, drop_skip/
-| #I #G #K #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L #s #d #e #HLK #U1 #H1 #U2 #H2
+| #I #G #K #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L #s #l #m #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
+| #G #K #V #T1 #T #T2 #_ #HT2 #IHT1 #L #s #l #m #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, drop_skip/
-| #G #K #V #T1 #T2 #_ #IHT12 #L #s #d #e #HLK #U1 #H #U2 #HTU2
+| #G #K #V #T1 #T2 #_ #IHT12 #L #s #l #m #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
+| #a #G #K #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #L #s #l #m #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, 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
+| #a #G #K #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV1 #IHW12 #IHT12 #L #s #l #m #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
qed.
(* Basic_1: includes: pr0_gen_lift pr2_gen_lift *)
-lemma cpr_inv_lift1: ∀G. l_deliftable_sn (cpr G).
+lemma cpr_inv_lift1: ∀G. d_deliftable_sn (cpr G).
#G #L #U1 #U2 #H elim H -G -L -U1 -U2
-[ * #i #G #L #K #s #d #e #_ #T1 #H
+[ * #i #G #L #K #s #l #m #_ #T1 #H
[ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3 by ex2_intro/
- | elim (lift_inv_lref2 … H) -H * #Hid #H destruct /3 width=3 by lift_lref_ge_minus, lift_lref_lt, ex2_intro/
+ | elim (lift_inv_lref2 … H) -H * #Hil #H destruct /3 width=3 by lift_lref_ge_minus, lift_lref_lt, ex2_intro/
| lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3 by ex2_intro/
]
-| #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
+| #G #L #LV #V #V2 #W2 #i #HLV #HV2 #HVW2 #IHV2 #K #s #l #m #HLK #T1 #H
+ elim (lift_inv_lref2 … H) -H * #Hil #H destruct
[ 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
+ | elim (le_inv_plus_l … Hil) #Hlim #Hmi
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
+ elim (lift_split … HVW2 l (i - m + 1)) -HVW2 [2,3,4: /2 width=1 by le_S_S, le_S/ ] -Hil -Hlim
#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
+| #a #I #G #L #V1 #V2 #U1 #U2 #_ #_ #IHV12 #IHU12 #K #s #l #m #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, drop_skip, lift_bind, ex2_intro/
-| #I #G #L #V1 #V2 #U1 #U2 #_ #_ #IHV12 #IHU12 #K #s #d #e #HLK #X #H
+| #I #G #L #V1 #V2 #U1 #U2 #_ #_ #IHV12 #IHU12 #K #s #l #m #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
+| #G #L #V #U1 #U #U2 #_ #HU2 #IHU1 #K #s #l #m #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 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
+| #G #L #V #U1 #U2 #_ #IHU12 #K #s #l #m #HLK #X #H
elim (lift_inv_flat2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct
elim (IHU12 … HLK … HTU1) -L -U1 /3 width=3 by cpr_eps, ex2_intro/
-| #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #K #s #d #e #HLK #X #HX
+| #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #K #s #l #m #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 (IHV12 … HLK … HV01) -V1
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
+| #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV1 #IHW12 #IHT12 #K #s #l #m #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
(* Properties on lazy sn pointwise extensions *******************************)
-lemma cpr_llpx_sn_conf: ∀R. (∀L. reflexive … (R L)) → l_liftable R → l_deliftable_sn R →
+lemma cpr_llpx_sn_conf: ∀R. (∀L. reflexive … (R L)) → d_liftable R → d_deliftable_sn R →
∀G. s_r_confluent1 … (cpr G) (llpx_sn R 0).
#R #H1R #H2R #H3R #G #Ls #T1 #T2 #H elim H -G -Ls -T1 -T2
[ //
| #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
+ #Kd #V1l #HLKd #HV1s #HV1sd
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 *)
(* avtivate genv *)
inductive cpx (h) (g): relation4 genv lenv term term ≝
| cpx_atom : ∀I,G,L. cpx h g G L (⓪{I}) (⓪{I})
-| cpx_st : ∀G,L,k,l. deg h g k (l+1) → cpx h g G L (⋆k) (⋆(next h k))
+| cpx_st : ∀G,L,k,d. deg h g k (d+1) → cpx h g G L (⋆k) (⋆(next h k))
| cpx_delta: ∀I,G,L,K,V,V2,W2,i.
⬇[i] L ≡ K.ⓑ{I}V → cpx h g G K V V2 →
⬆[0, i+1] V2 ≡ W2 → cpx h g G L (#i) W2
#h #g * /2 width=1 by cpx_bind, cpx_flat/
qed.
-lemma cpx_delift: ∀h,g,I,G,K,V,T1,L,d. ⬇[d] L ≡ (K.ⓑ{I}V) →
- ∃∃T2,T. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 & ⬆[d, 1] T ≡ T2.
+lemma cpx_delift: ∀h,g,I,G,K,V,T1,L,l. ⬇[l] L ≡ (K.ⓑ{I}V) →
+ ∃∃T2,T. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 & ⬆[l, 1] T ≡ T2.
#h #g #I #G #K #V #T1 elim T1 -T1
-[ * #i #L #d /2 width=4 by cpx_atom, lift_sort, lift_gref, ex2_2_intro/
- elim (lt_or_eq_or_gt i d) #Hid [1,3: /3 width=4 by cpx_atom, lift_lref_ge_minus, lift_lref_lt, ex2_2_intro/ ]
+[ * #i #L #l /2 width=4 by cpx_atom, lift_sort, lift_gref, ex2_2_intro/
+ elim (lt_or_eq_or_gt i l) #Hil [1,3: /3 width=4 by cpx_atom, lift_lref_ge_minus, lift_lref_lt, ex2_2_intro/ ]
destruct
elim (lift_total V 0 (i+1)) #W #HVW
elim (lift_split … HVW i i) /3 width=7 by cpx_delta, ex2_2_intro/
-| * [ #a ] #I #W1 #U1 #IHW1 #IHU1 #L #d #HLK
+| * [ #a ] #I #W1 #U1 #IHW1 #IHU1 #L #l #HLK
elim (IHW1 … HLK) -IHW1 #W2 #W #HW12 #HW2
- [ elim (IHU1 (L. ⓑ{I} W1) (d+1)) -IHU1 /3 width=9 by cpx_bind, drop_drop, lift_bind, ex2_2_intro/
+ [ elim (IHU1 (L. ⓑ{I} W1) (l+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/
]
]
fact cpx_inv_atom1_aux: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 → ∀J. T1 = ⓪{J} →
∨∨ T2 = ⓪{J}
- | ∃∃k,l. deg h g k (l+1) & T2 = ⋆(next h k) & J = Sort k
+ | ∃∃k,d. deg h g k (d+1) & T2 = ⋆(next h k) & J = Sort k
| ∃∃I,K,V,V2,i. ⬇[i] L ≡ K.ⓑ{I}V & ⦃G, K⦄ ⊢ V ➡[h, g] V2 &
⬆[O, i+1] V2 ≡ T2 & J = LRef i.
#G #h #g #L #T1 #T2 * -L -T1 -T2
[ #I #G #L #J #H destruct /2 width=1 by or3_intro0/
-| #G #L #k #l #Hkl #J #H destruct /3 width=5 by or3_intro1, ex3_2_intro/
+| #G #L #k #d #Hkd #J #H destruct /3 width=5 by or3_intro1, ex3_2_intro/
| #I #G #L #K #V #V2 #T2 #i #HLK #HV2 #HVT2 #J #H destruct /3 width=9 by or3_intro2, ex4_5_intro/
| #a #I #G #L #V1 #V2 #T1 #T2 #_ #_ #J #H destruct
| #I #G #L #V1 #V2 #T1 #T2 #_ #_ #J #H destruct
lemma cpx_inv_atom1: ∀h,g,J,G,L,T2. ⦃G, L⦄ ⊢ ⓪{J} ➡[h, g] T2 →
∨∨ T2 = ⓪{J}
- | ∃∃k,l. deg h g k (l+1) & T2 = ⋆(next h k) & J = Sort k
+ | ∃∃k,d. deg h g k (d+1) & T2 = ⋆(next h k) & J = Sort k
| ∃∃I,K,V,V2,i. ⬇[i] L ≡ K.ⓑ{I}V & ⦃G, K⦄ ⊢ V ➡[h, g] V2 &
⬆[O, i+1] V2 ≡ T2 & J = LRef i.
/2 width=3 by cpx_inv_atom1_aux/ qed-.
lemma cpx_inv_sort1: ∀h,g,G,L,T2,k. ⦃G, L⦄ ⊢ ⋆k ➡[h, g] T2 → T2 = ⋆k ∨
- ∃∃l. deg h g k (l+1) & T2 = ⋆(next h k).
+ ∃∃d. deg h g k (d+1) & T2 = ⋆(next h k).
#h #g #G #L #T2 #k #H
elim (cpx_inv_atom1 … H) -H /2 width=1 by or_introl/ *
-[ #k0 #l0 #Hkl0 #H1 #H2 destruct /3 width=4 by ex2_intro, or_intror/
+[ #k0 #d0 #Hkd0 #H1 #H2 destruct /3 width=4 by ex2_intro, or_intror/
| #I #K #V #V2 #i #_ #_ #_ #H destruct
]
qed-.
⬆[O, i+1] V2 ≡ T2.
#h #g #G #L #T2 #i #H
elim (cpx_inv_atom1 … H) -H /2 width=1 by or_introl/ *
-[ #k #l #_ #_ #H destruct
+[ #k #d #_ #_ #H destruct
| #I #K #V #V2 #j #HLK #HV2 #HVT2 #H destruct /3 width=7 by ex3_4_intro, or_intror/
]
qed-.
lemma cpx_inv_gref1: ∀h,g,G,L,T2,p. ⦃G, L⦄ ⊢ §p ➡[h, g] T2 → T2 = §p.
#h #g #G #L #T2 #p #H
elim (cpx_inv_atom1 … H) -H // *
-[ #k #l #_ #_ #H destruct
+[ #k #d #_ #_ #H destruct
| #I #K #V #V2 #i #_ #_ #_ #H destruct
]
qed-.
a = true & J = Abbr.
#h #g #G #L #U1 #U2 * -L -U1 -U2
[ #I #G #L #b #J #W #U1 #H destruct
-| #G #L #k #l #_ #b #J #W #U1 #H destruct
+| #G #L #k #d #_ #b #J #W #U1 #H destruct
| #I #G #L #K #V #V2 #W2 #i #_ #_ #_ #b #J #W #U1 #H destruct
| #a #I #G #L #V1 #V2 #T1 #T2 #HV12 #HT12 #b #J #W #U1 #H destruct /3 width=5 by ex3_2_intro, or_introl/
| #I #G #L #V1 #V2 #T1 #T2 #_ #_ #b #J #W #U1 #H destruct
U2 = ⓓ{a}W2.ⓐV2.T2 & J = Appl.
#h #g #G #L #U #U2 * -L -U -U2
[ #I #G #L #J #W #U1 #H destruct
-| #G #L #k #l #_ #J #W #U1 #H destruct
+| #G #L #k #d #_ #J #W #U1 #H destruct
| #I #G #L #K #V #V2 #W2 #i #_ #_ #_ #J #W #U1 #H destruct
| #a #I #G #L #V1 #V2 #T1 #T2 #_ #_ #J #W #U1 #H destruct
| #I #G #L #V1 #V2 #T1 #T2 #HV12 #HT12 #J #W #U1 #H destruct /3 width=5 by or5_intro0, ex3_2_intro/
lemma cpx_fwd_cix: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 → ⦃G, L⦄ ⊢ ➡[h, g] 𝐈⦃T1⦄ → T2 = T1.
#h #g #G #L #T1 #T2 #H elim H -G -L -T1 -T2
[ //
-| #G #L #k #l #Hkl #H elim (cix_inv_sort … Hkl H)
+| #G #L #k #d #Hkd #H elim (cix_inv_sort … Hkd H)
| #I #G #L #K #V1 #V2 #W2 #i #HLK #_ #HVW2 #IHV12 #H
elim (cix_inv_delta … HLK) //
| #a * #G #L #V1 #V2 #T1 #T2 #_ #_ #IHV1 #IHT1 #H
+++ /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/reduction/cpx.ma".
-
-(* CONTEXT-SENSITIVE EXTENDED PARALLEL REDUCTION FOR TERMS ******************)
-
-(* Properties on equivalence for local environments *************************)
-
-lemma leq_cpx_trans: ∀h,g,G. lsub_trans … (cpx h g G) (leq 0 (∞)).
-#h #g #G #L1 #T1 #T2 #H elim H -G -L1 -T1 -T2
-[ //
-| /2 width=2 by cpx_st/
-| #I #G #L1 #K1 #V1 #V2 #W2 #i #HLK1 #_ #HVW2 #IHV12 #L2 #HL12
- 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/
-]
-qed-.
(* Advanced properties ******************************************************)
-fact sta_cpx_aux: ∀h,g,G,L,T1,T2,l2,l1. ⦃G, L⦄ ⊢ T1 •*[h, l2] T2 → l2 = 1 →
- ⦃G, L⦄ ⊢ T1 ▪[h, g] l1+1 → ⦃G, L⦄ ⊢ T1 ➡[h, g] T2.
-#h #g #G #L #T1 #T2 #l2 #l1 #H elim H -G -L -T1 -T2 -l2
-[ #G #L #l2 #k #H0 destruct normalize
+fact sta_cpx_aux: ∀h,g,G,L,T1,T2,d2,d1. ⦃G, L⦄ ⊢ T1 •*[h, d2] T2 → d2 = 1 →
+ ⦃G, L⦄ ⊢ T1 ▪[h, g] d1+1 → ⦃G, L⦄ ⊢ T1 ➡[h, g] T2.
+#h #g #G #L #T1 #T2 #d2 #d1 #H elim H -G -L -T1 -T2 -d2
+[ #G #L #d2 #k #H0 destruct normalize
/3 width=4 by cpx_st, da_inv_sort/
-| #G #L #K #V1 #V2 #W2 #i #l2 #HLK #_ #HVW2 #IHV12 #H0 #H destruct
- elim (da_inv_lref … H) -H * #K0 #V0 [| #l0 ] #HLK0
+| #G #L #K #V1 #V2 #W2 #i #d2 #HLK #_ #HVW2 #IHV12 #H0 #H destruct
+ elim (da_inv_lref … H) -H * #K0 #V0 [| #d0 ] #HLK0
lapply (drop_mono … HLK0 … HLK) -HLK0 #H destruct /3 width=7 by cpx_delta/
| #G #L #K #V1 #V2 #i #_ #_ #_ #H destruct
-| #G #L #K #V1 #V2 #W2 #i #l2 #HLK #HV12 #HVW2 #_ #H0 #H
+| #G #L #K #V1 #V2 #W2 #i #d2 #HLK #HV12 #HVW2 #_ #H0 #H
lapply (discr_plus_xy_y … H0) -H0 #H0 destruct
- elim (da_inv_lref … H) -H * #K0 #V0 [| #l0 ] #HLK0
+ elim (da_inv_lref … H) -H * #K0 #V0 [| #d0 ] #HLK0
lapply (drop_mono … HLK0 … HLK) -HLK0 #H destruct
/4 width=7 by cpx_delta, cpr_cpx, lstas_cpr/
| /4 width=2 by cpx_bind, da_inv_bind/
]
qed-.
-lemma sta_cpx: ∀h,g,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 •*[h, 1] T2 →
- ⦃G, L⦄ ⊢ T1 ▪[h, g] l+1 → ⦃G, L⦄ ⊢ T1 ➡[h, g] T2.
+lemma sta_cpx: ∀h,g,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 •*[h, 1] T2 →
+ ⦃G, L⦄ ⊢ T1 ▪[h, g] d+1 → ⦃G, L⦄ ⊢ T1 ➡[h, g] T2.
/2 width=3 by sta_cpx_aux/ qed.
(* Relocation properties ****************************************************)
-lemma cpx_lift: ∀h,g,G. l_liftable (cpx h g G).
+lemma cpx_lift: ∀h,g,G. d_liftable (cpx h g G).
#h #g #G #K #T1 #T2 #H elim H -G -K -T1 -T2
-[ #I #G #K #L #s #d #e #_ #U1 #H1 #U2 #H2
+[ #I #G #K #L #s #l #m #_ #U1 #H1 #U2 #H2
>(lift_mono … H1 … H2) -H1 -H2 //
-| #G #K #k #l #Hkl #L #s #d #e #_ #U1 #H1 #U2 #H2
+| #G #K #k #d #Hkd #L #s #l #m #_ #U1 #H1 #U2 #H2
>(lift_inv_sort1 … H1) -U1
>(lift_inv_sort1 … H2) -U2 /2 width=2 by cpx_st/
-| #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
+| #I #G #K #KV #V #V2 #W2 #i #HKV #HV2 #HVW2 #IHV2 #L #s #l #m #HLK #U1 #H #U2 #HWU2
+ elim (lift_inv_lref1 … H) * #Hil #H destruct
[ elim (lift_trans_ge … HVW2 … HWU2) -W2 // <minus_plus #W2 #HVW2 #HWU2
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
+ elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hil
#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 (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
+| #a #I #G #K #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L #s #l #m #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, drop_skip/
-| #I #G #K #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L #s #d #e #HLK #U1 #H1 #U2 #H2
+| #I #G #K #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L #s #l #m #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
+| #G #K #V #T1 #T #T2 #_ #HT2 #IHT1 #L #s #l #m #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, drop_skip/
-| #G #K #V #T1 #T2 #_ #IHT12 #L #s #d #e #HLK #U1 #H #U2 #HTU2
+| #G #K #V #T1 #T2 #_ #IHT12 #L #s #l #m #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
+| #G #K #V1 #V2 #T #_ #IHV12 #L #s #l #m #HLK #U1 #H #U2 #HVU2
elim (lift_inv_flat1 … H) -H #VV1 #TT1 #HVV1 #HTT1 #H destruct /3 width=6 by cpx_ct/
-| #a #G #K #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #L #s #d #e #HLK #X1 #HX1 #X2 #HX2
+| #a #G #K #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #L #s #l #m #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 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
+| #a #G #K #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV1 #IHW12 #IHT12 #L #s #l #m #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
]
qed.
-lemma cpx_inv_lift1: ∀h,g,G. l_deliftable_sn (cpx h g G).
+lemma cpx_inv_lift1: ∀h,g,G. d_deliftable_sn (cpx h g G).
#h #g #G #L #U1 #U2 #H elim H -G -L -U1 -U2
-[ * #i #G #L #K #s #d #e #_ #T1 #H
+[ * #i #G #L #K #s #l #m #_ #T1 #H
[ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3 by cpx_atom, lift_sort, ex2_intro/
- | elim (lift_inv_lref2 … H) -H * #Hid #H destruct /3 width=3 by cpx_atom, lift_lref_ge_minus, lift_lref_lt, ex2_intro/
+ | elim (lift_inv_lref2 … H) -H * #Hil #H destruct /3 width=3 by cpx_atom, lift_lref_ge_minus, lift_lref_lt, ex2_intro/
| lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3 by cpx_atom, lift_gref, ex2_intro/
]
-| #G #L #k #l #Hkl #K #s #d #e #_ #T1 #H
+| #G #L #k #d #Hkd #K #s #l #m #_ #T1 #H
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
+| #I #G #L #LV #V #V2 #W2 #i #HLV #HV2 #HVW2 #IHV2 #K #s #l #m #HLK #T1 #H
+ elim (lift_inv_lref2 … H) -H * #Hil #H destruct
[ 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
+ | elim (le_inv_plus_l … Hil) #Hlim #Hmi
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
+ elim (lift_split … HVW2 l (i - m + 1)) -HVW2 /3 width=1 by le_S, le_S_S/ -Hil -Hlim
#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
+| #a #I #G #L #V1 #V2 #U1 #U2 #_ #_ #IHV12 #IHU12 #K #s #l #m #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, drop_skip, lift_bind, ex2_intro/
-| #I #G #L #V1 #V2 #U1 #U2 #_ #_ #IHV12 #IHU12 #K #s #d #e #HLK #X #H
+| #I #G #L #V1 #V2 #U1 #U2 #_ #_ #IHV12 #IHU12 #K #s #l #m #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=5 by cpx_flat, lift_flat, ex2_intro/
-| #G #L #V #U1 #U #U2 #_ #HU2 #IHU1 #K #s #d #e #HLK #X #H
+| #G #L #V #U1 #U #U2 #_ #HU2 #IHU1 #K #s #l #m #HLK #X #H
elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct
elim (IHU1 (K.ⓓW1) s … HTU1) /2 width=1/ -L -U1 #T #HTU #HT1
elim (lift_div_le … HU2 … HTU) -U /3 width=5 by cpx_zeta, ex2_intro/
-| #G #L #V #U1 #U2 #_ #IHU12 #K #s #d #e #HLK #X #H
+| #G #L #V #U1 #U2 #_ #IHU12 #K #s #l #m #HLK #X #H
elim (lift_inv_flat2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct
elim (IHU12 … HLK … HTU1) -L -U1 /3 width=3 by cpx_eps, ex2_intro/
-| #G #L #V1 #V2 #U1 #_ #IHV12 #K #s #d #e #HLK #X #H
+| #G #L #V1 #V2 #U1 #_ #IHV12 #K #s #l #m #HLK #X #H
elim (lift_inv_flat2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct
elim (IHV12 … HLK … HWV1) -L -V1 /3 width=3 by cpx_ct, ex2_intro/
-| #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #K #s #d #e #HLK #X #HX
+| #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #K #s #l #m #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 (IHV12 … HLK … HV01) -V1 #V3 #HV32 #HV03
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
+| #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV1 #IHW12 #IHT12 #K #s #l #m #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
[ #I #G #L #V2 #U2 #HVU2
elim (lift_total U2 0 1)
/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))
+| #G #L #K #T1 #U1 #m #HLK1 #HTU1 #T2 #HTU2
+ elim (lift_total T2 0 (m+1))
/3 width=11 by cpx_lift, fqu_drop, ex2_intro/
]
qed-.
lemma fqu_sta_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
∀U2. ⦃G2, L2⦄ ⊢ T2 •*[h, 1] U2 →
- ∀l. ⦃G2, L2⦄ ⊢ T2 ▪[h, g] l+1 →
+ ∀d. ⦃G2, L2⦄ ⊢ T2 ▪[h, g] d+1 →
∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, g] U1 & ⦃G1, L1, U1⦄ ⊐ ⦃G2, L2, U2⦄.
/3 width=5 by fqu_cpx_trans, sta_cpx/ qed-.
lemma fquq_sta_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
∀U2. ⦃G2, L2⦄ ⊢ T2 •*[h, 1] U2 →
- ∀l. ⦃G2, L2⦄ ⊢ T2 ▪[h, g] l+1 →
+ ∀d. ⦃G2, L2⦄ ⊢ T2 ▪[h, g] d+1 →
∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, g] U1 & ⦃G1, L1, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
/3 width=5 by fquq_cpx_trans, sta_cpx/ qed-.
[1,3: /2 width=4 by fqu_flat_dx, cpx_flat/
| #H0 destruct /2 width=1 by/
]
-| #G #L #K #T1 #U1 #e #HLK #HTU1 #T2 #HT12 #H elim (lift_total T2 0 (e+1))
+| #G #L #K #T1 #U1 #m #HLK #HTU1 #T2 #HT12 #H elim (lift_total T2 0 (m+1))
#U2 #HTU2 @(ex3_intro … U2)
[1,3: /2 width=10 by cpx_lift, fqu_drop/
| #H0 destruct /3 width=5 by lift_inj/
(* Properties on lazy sn pointwise extensions *******************************)
(* Note: lemma 1000 *)
-lemma cpx_llpx_sn_conf: ∀R. (∀L. reflexive … (R L)) → l_liftable R → l_deliftable_sn R →
+lemma cpx_llpx_sn_conf: ∀R. (∀L. reflexive … (R L)) → d_liftable R → d_deliftable_sn R →
∀h,g,G. s_r_confluent1 … (cpx h g G) (llpx_sn R 0).
#R #H1R #H2R #H3R #h #g #G #Ls #T1 #T2 #H elim H -G -Ls -T1 -T2
[ //
| /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
+ #Kd #V1l #HLKd #HV1s #HV1sd
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 *)
--- /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_lreq.ma".
+include "basic_2/reduction/cpx.ma".
+
+(* CONTEXT-SENSITIVE EXTENDED PARALLEL REDUCTION FOR TERMS ******************)
+
+(* Properties on equivalence for local environments *************************)
+
+lemma lreq_cpx_trans: ∀h,g,G. lsub_trans … (cpx h g G) (lreq 0 (∞)).
+#h #g #G #L1 #T1 #T2 #H elim H -G -L1 -T1 -T2
+[ //
+| /2 width=2 by cpx_st/
+| #I #G #L1 #K1 #V1 #V2 #W2 #i #HLK1 #_ #HVW2 #IHV12 #L2 #HL12
+ elim (lreq_drop_trans_be … HL12 … HLK1) // -HL12 -HLK1 /3 width=7 by cpx_delta/
+|4,9: /4 width=1 by cpx_bind, cpx_beta, lreq_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, lreq_pair_O_Y/
+]
+qed-.
(* Properties on relocation *************************************************)
-lemma crr_lift: ∀G,K,T. ⦃G, K⦄ ⊢ ➡ 𝐑⦃T⦄ → ∀L,s,d,e. ⬇[s, d, e] L ≡ K →
- ∀U. ⬆[d, e] T ≡ U → ⦃G, L⦄ ⊢ ➡ 𝐑⦃U⦄.
+lemma crr_lift: ∀G,K,T. ⦃G, K⦄ ⊢ ➡ 𝐑⦃T⦄ → ∀L,s,l,m. ⬇[s, l, m] L ≡ K →
+ ∀U. ⬆[l, m] T ≡ U → ⦃G, L⦄ ⊢ ➡ 𝐑⦃U⦄.
#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
+[ #K #K0 #V #i #HK0 #L #s #l #m #HLK #X #H
+ elim (lift_inv_lref1 … H) -H * #Hil #H destruct
[ 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
+| #K #V #T #_ #IHV #L #s #l #m #HLK #X #H
elim (lift_inv_flat1 … H) -H #W #U #HVW #_ #H destruct /3 width=5 by crr_appl_sn/
-| #K #V #T #_ #IHT #L #s #d #e #HLK #X #H
+| #K #V #T #_ #IHT #L #s #l #m #HLK #X #H
elim (lift_inv_flat1 … H) -H #W #U #_ #HTU #H destruct /3 width=5 by crr_appl_dx/
-| #I #K #V #T #HI #L #s #d #e #_ #X #H
+| #I #K #V #T #HI #L #s #l #m #_ #X #H
elim (lift_fwd_pair1 … H) -H #W #U #_ #H destruct /2 width=1 by crr_ri2/
-| #a #I #K #V #T #HI #_ #IHV #L #s #d #e #HLK #X #H
+| #a #I #K #V #T #HI #_ #IHV #L #s #l #m #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
+| #a #I #K #V #T #HI #_ #IHT #L #s #l #m #HLK #X #H
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
+| #a #K #V #V0 #T #L #s #l #m #_ #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/
-| #a #K #V #V0 #T #L #s #d #e #_ #X #H
+| #a #K #V #V0 #T #L #s #l #m #_ #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_theta/
]
qed.
-lemma crr_inv_lift: ∀G,L,U. ⦃G, L⦄ ⊢ ➡ 𝐑⦃U⦄ → ∀K,s,d,e. ⬇[s, d, e] L ≡ K →
- ∀T. ⬆[d, e] T ≡ U → ⦃G, K⦄ ⊢ ➡ 𝐑⦃T⦄.
+lemma crr_inv_lift: ∀G,L,U. ⦃G, L⦄ ⊢ ➡ 𝐑⦃U⦄ → ∀K,s,l,m. ⬇[s, l, m] L ≡ K →
+ ∀T. ⬆[l, m] T ≡ U → ⦃G, K⦄ ⊢ ➡ 𝐑⦃T⦄.
#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
+[ #L #L0 #W #i #HK0 #K #s #l #m #HLK #X #H
+ elim (lift_inv_lref2 … H) -H * #Hil #H destruct
[ 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
+| #L #W #U #_ #IHW #K #s #l #m #HLK #X #H
elim (lift_inv_flat2 … H) -H #V #T #HVW #_ #H destruct /3 width=5 by crr_appl_sn/
-| #L #W #U #_ #IHU #K #s #d #e #HLK #X #H
+| #L #W #U #_ #IHU #K #s #l #m #HLK #X #H
elim (lift_inv_flat2 … H) -H #V #T #_ #HTU #H destruct /3 width=5 by crr_appl_dx/
-| #I #L #W #U #HI #K #s #d #e #_ #X #H
+| #I #L #W #U #HI #K #s #l #m #_ #X #H
elim (lift_fwd_pair2 … H) -H #V #T #_ #H destruct /2 width=1 by crr_ri2/
-| #a #I #L #W #U #HI #_ #IHW #K #s #d #e #HLK #X #H
+| #a #I #L #W #U #HI #_ #IHW #K #s #l #m #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
+| #a #I #L #W #U #HI #_ #IHU #K #s #l #m #HLK #X #H
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
+| #a #L #W #W0 #U #K #s #l #m #_ #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/
-| #a #L #W #W0 #U #K #s #d #e #_ #X #H
+| #a #L #W #W0 #U #K #s #l #m #_ #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_theta/
]
(* activate genv *)
(* extended reducible terms *)
inductive crx (h) (g) (G:genv): relation2 lenv term ≝
-| crx_sort : ∀L,k,l. deg h g k (l+1) → crx h g G L (⋆k)
+| crx_sort : ∀L,k,d. deg h g k (d+1) → crx h g G L (⋆k)
| crx_delta : ∀I,L,K,V,i. ⬇[i] L ≡ K.ⓑ{I}V → crx h g G L (#i)
| crx_appl_sn: ∀L,V,T. crx h g G L V → crx h g G L (ⓐV.T)
| crx_appl_dx: ∀L,V,T. crx h g G L T → crx h g G L (ⓐV.T)
(* Basic inversion lemmas ***************************************************)
fact crx_inv_sort_aux: ∀h,g,G,L,T,k. ⦃G, L⦄ ⊢ ➡[h, g] 𝐑⦃T⦄ → T = ⋆k →
- ∃l. deg h g k (l+1).
+ ∃d. deg h g k (d+1).
#h #g #G #L #T #k0 * -L -T
-[ #L #k #l #Hkl #H destruct /2 width=2 by ex_intro/
+[ #L #k #d #Hkd #H destruct /2 width=2 by ex_intro/
| #I #L #K #V #i #HLK #H destruct
| #L #V #T #_ #H destruct
| #L #V #T #_ #H destruct
]
qed-.
-lemma crx_inv_sort: ∀h,g,G,L,k. ⦃G, L⦄ ⊢ ➡[h, g] 𝐑⦃⋆k⦄ → ∃l. deg h g k (l+1).
+lemma crx_inv_sort: ∀h,g,G,L,k. ⦃G, L⦄ ⊢ ➡[h, g] 𝐑⦃⋆k⦄ → ∃d. deg h g k (d+1).
/2 width=5 by crx_inv_sort_aux/ qed-.
fact crx_inv_lref_aux: ∀h,g,G,L,T,i. ⦃G, L⦄ ⊢ ➡[h, g] 𝐑⦃T⦄ → T = #i →
∃∃I,K,V. ⬇[i] L ≡ K.ⓑ{I}V.
#h #g #G #L #T #j * -L -T
-[ #L #k #l #_ #H destruct
+[ #L #k #d #_ #H destruct
| #I #L #K #V #i #HLK #H destruct /2 width=4 by ex1_3_intro/
| #L #V #T #_ #H destruct
| #L #V #T #_ #H destruct
fact crx_inv_gref_aux: ∀h,g,G,L,T,p. ⦃G, L⦄ ⊢ ➡[h, g] 𝐑⦃T⦄ → T = §p → ⊥.
#h #g #G #L #T #q * -L -T
-[ #L #k #l #_ #H destruct
+[ #L #k #d #_ #H destruct
| #I #L #K #V #i #HLK #H destruct
| #L #V #T #_ #H destruct
| #L #V #T #_ #H destruct
/2 width=8 by crx_inv_gref_aux/ qed-.
lemma trx_inv_atom: ∀h,g,I,G. ⦃G, ⋆⦄ ⊢ ➡[h, g] 𝐑⦃⓪{I}⦄ →
- ∃∃k,l. deg h g k (l+1) & I = Sort k.
+ ∃∃k,d. deg h g k (d+1) & I = Sort k.
#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
fact crx_inv_ib2_aux: ∀h,g,a,I,G,L,W,U,T. ib2 a I → ⦃G, L⦄ ⊢ ➡[h, g] 𝐑⦃T⦄ →
T = ⓑ{a,I}W.U → ⦃G, L⦄ ⊢ ➡[h, g] 𝐑⦃W⦄ ∨ ⦃G, L.ⓑ{I}W⦄ ⊢ ➡[h, g] 𝐑⦃U⦄.
#h #g #b #J #G #L #W0 #U #T #HI * -L -T
-[ #L #k #l #_ #H destruct
+[ #L #k #d #_ #H destruct
| #I #L #K #V #i #_ #H destruct
| #L #V #T #_ #H destruct
| #L #V #T #_ #H destruct
fact crx_inv_appl_aux: ∀h,g,G,L,W,U,T. ⦃G, L⦄ ⊢ ➡[h, g] 𝐑⦃T⦄ → T = ⓐW.U →
∨∨ ⦃G, L⦄ ⊢ ➡[h, g] 𝐑⦃W⦄ | ⦃G, L⦄ ⊢ ➡[h, g] 𝐑⦃U⦄ | (𝐒⦃U⦄ → ⊥).
#h #g #G #L #W0 #U #T * -L -T
-[ #L #k #l #_ #H destruct
+[ #L #k #d #_ #H destruct
| #I #L #K #V #i #_ #H destruct
| #L #V #T #HV #H destruct /2 width=1 by or3_intro0/
| #L #V #T #HT #H destruct /2 width=1 by or3_intro1/
(* Properties on relocation *************************************************)
-lemma crx_lift: ∀h,g,G,K,T. ⦃G, K⦄ ⊢ ➡[h, g] 𝐑⦃T⦄ → ∀L,s,d,e. ⬇[s, d, e] L ≡ K →
- ∀U. ⬆[d, e] T ≡ U → ⦃G, L⦄ ⊢ ➡[h, g] 𝐑⦃U⦄.
+lemma crx_lift: ∀h,g,G,K,T. ⦃G, K⦄ ⊢ ➡[h, g] 𝐑⦃T⦄ → ∀L,s,l,m. ⬇[s, l, m] L ≡ K →
+ ∀U. ⬆[l, m] T ≡ U → ⦃G, L⦄ ⊢ ➡[h, g] 𝐑⦃U⦄.
#h #g #G #K #T #H elim H -K -T
-[ #K #k #l #Hkl #L #s #d #e #_ #X #H
+[ #K #k #d #Hkd #L #s #l #m #_ #X #H
>(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
+| #I #K #K0 #V #i #HK0 #L #s #l #m #HLK #X #H
+ elim (lift_inv_lref1 … H) -H * #Hil #H destruct
[ 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
+| #K #V #T #_ #IHV #L #s #l #m #HLK #X #H
elim (lift_inv_flat1 … H) -H #W #U #HVW #_ #H destruct /3 width=5 by crx_appl_sn/
-| #K #V #T #_ #IHT #L #s #d #e #HLK #X #H
+| #K #V #T #_ #IHT #L #s #l #m #HLK #X #H
elim (lift_inv_flat1 … H) -H #W #U #_ #HTU #H destruct /3 width=5 by crx_appl_dx/
-| #I #K #V #T #HI #L #s #d #e #_ #X #H
+| #I #K #V #T #HI #L #s #l #m #_ #X #H
elim (lift_fwd_pair1 … H) -H #W #U #_ #H destruct /2 width=1 by crx_ri2/
-| #a #I #K #V #T #HI #_ #IHV #L #s #d #e #HLK #X #H
+| #a #I #K #V #T #HI #_ #IHV #L #s #l #m #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
+| #a #I #K #V #T #HI #_ #IHT #L #s #l #m #HLK #X #H
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
+| #a #K #V #V0 #T #L #s #l #m #_ #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/
-| #a #K #V #V0 #T #L #s #d #e #_ #X #H
+| #a #K #V #V0 #T #L #s #l #m #_ #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_theta/
]
qed.
-lemma crx_inv_lift: ∀h,g,G,L,U. ⦃G, L⦄ ⊢ ➡[h, g] 𝐑⦃U⦄ → ∀K,s,d,e. ⬇[s, d, e] L ≡ K →
- ∀T. ⬆[d, e] T ≡ U → ⦃G, K⦄ ⊢ ➡[h, g] 𝐑⦃T⦄.
+lemma crx_inv_lift: ∀h,g,G,L,U. ⦃G, L⦄ ⊢ ➡[h, g] 𝐑⦃U⦄ → ∀K,s,l,m. ⬇[s, l, m] L ≡ K →
+ ∀T. ⬆[l, m] T ≡ U → ⦃G, K⦄ ⊢ ➡[h, g] 𝐑⦃T⦄.
#h #g #G #L #U #H elim H -L -U
-[ #L #k #l #Hkl #K #s #d #e #_ #X #H
+[ #L #k #d #Hkd #K #s #l #m #_ #X #H
>(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
+| #I #L #L0 #W #i #HK0 #K #s #l #m #HLK #X #H
+ elim (lift_inv_lref2 … H) -H * #Hil #H destruct
[ 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
+| #L #W #U #_ #IHW #K #s #l #m #HLK #X #H
elim (lift_inv_flat2 … H) -H #V #T #HVW #_ #H destruct /3 width=5 by crx_appl_sn/
-| #L #W #U #_ #IHU #K #s #d #e #HLK #X #H
+| #L #W #U #_ #IHU #K #s #l #m #HLK #X #H
elim (lift_inv_flat2 … H) -H #V #T #_ #HTU #H destruct /3 width=5 by crx_appl_dx/
-| #I #L #W #U #HI #K #s #d #e #_ #X #H
+| #I #L #W #U #HI #K #s #l #m #_ #X #H
elim (lift_fwd_pair2 … H) -H #V #T #_ #H destruct /2 width=1 by crx_ri2/
-| #a #I #L #W #U #HI #_ #IHW #K #s #d #e #HLK #X #H
+| #a #I #L #W #U #HI #_ #IHW #K #s #l #m #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
+| #a #I #L #W #U #HI #_ #IHU #K #s #l #m #HLK #X #H
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
+| #a #L #W #W0 #U #K #s #l #m #_ #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/
-| #a #L #W #W0 #U #K #s #d #e #_ #X #H
+| #a #L #W #W0 #U #K #s #l #m #_ #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_theta/
]
(* Advanced properties ******************************************************)
-lemma sta_fpb: ∀h,g,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 ▪[h, g] l+1 →
+lemma sta_fpb: ∀h,g,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 ▪[h, g] d+1 →
⦃G, L⦄ ⊢ T1 •*[h, 1] T2 → ⦃G, L, T1⦄ ≻[h, g] ⦃G, L, T2⦄.
-#h #g #G #L #T1 #T2 #l #HT1 #HT12 elim (eq_term_dec T1 T2)
+#h #g #G #L #T1 #T2 #d #HT1 #HT12 elim (eq_term_dec T1 T2)
/3 width=2 by fpb_cpx, sta_cpx/ #H destruct
elim (lstas_inv_refl_pos h G L T2 0) //
qed.
(* Advanced properties ******************************************************)
-lemma sta_fpbq: ∀h,g,G,L,T1,T2,l.
- ⦃G, L⦄ ⊢ T1 ▪[h, g] l+1 → ⦃G, L⦄ ⊢ T1 •*[h, 1] T2 →
+lemma sta_fpbq: ∀h,g,G,L,T1,T2,d.
+ ⦃G, L⦄ ⊢ T1 ▪[h, g] d+1 → ⦃G, L⦄ ⊢ T1 •*[h, 1] T2 →
⦃G, L, T1⦄ ≽[h, g] ⦃G, L, T2⦄.
/3 width=4 by fpbq_cpx, sta_cpx/ qed.
∃∃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
+#G #L #K #U #T #m #HLK #HUT #U2 #HU2
+elim (lift_total U2 0 (m+1)) #T2 #HUT2
lapply (cpr_lift … HU2 … HLK … HUT … HUT2) -HU2 -HUT /3 width=9 by fqu_drop, ex3_2_intro/
qed-.
∃∃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
+#G #L #K #U #T #m #HLK #HUT #U2 #HU2
+elim (lift_total U2 0 (m+1)) #T2 #HUT2
lapply (cpr_lift … HU2 … HLK … HUT … HUT2) -HU2 -HUT /3 width=9 by fqu_drop, ex3_2_intro/
qed-.
[ #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
+| #G #L1 #K1 #T1 #U1 #m #HLK1 #HTU1 #K2 #HK12
elim (drop_lpr_trans … HLK1 … HK12) -HK12
/3 width=7 by fqu_drop, 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 cpx_pair_sn, fqu_bind_dx, ex3_2_intro/
-| #G #L1 #K1 #T1 #U1 #e #HLK1 #HTU1 #K2 #HK12
+| #G #L1 #K1 #T1 #U1 #m #HLK1 #HTU1 #K2 #HK12
elim (drop_lpx_trans … HLK1 … HK12) -HK12
/3 width=7 by fqu_drop, 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
+| #G #L1 #K1 #T1 #U1 #m #HLK1 #HTU1 #L0 #HL01
elim (lpx_drop_trans_O1 … HL01 … HLK1) -L1
/3 width=5 by fqu_drop, ex3_2_intro/
]
(* *)
(**************************************************************************)
-include "basic_2/multiple/frees_leq.ma".
+include "basic_2/multiple/frees_lreq.ma".
include "basic_2/multiple/frees_lift.ma".
include "basic_2/reduction/lpx_drop.ma".
#h #g #G #L1 #U1 @(fqup_wf_ind_eq … G L1 U1) -G -L1 -U1
#G0 #L0 #U0 #IH #G #L1 * *
[ -IH #k #HG #HL #HU #U2 #H1 #L2 #_ #i #H2 elim (cpx_inv_sort1 … H1) -H1
- [| * #l #_ ] #H destruct elim (frees_inv_sort … H2)
+ [| * #d #_ ] #H destruct elim (frees_inv_sort … H2)
| #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 (frees_inv_bind … Hi) -Hi #Hi
[ elim (frees_inv_flat … Hi) -Hi
/4 width=7 by frees_flat_dx, frees_flat_sn, frees_bind_sn/
- | lapply (leq_frees_trans … Hi (L2.ⓛV2) ?) /2 width=1 by leq_succ/ -Hi #HU2
+ | lapply (lreq_frees_trans … Hi (L2.ⓛV2) ?) /2 width=1 by lreq_succ/ -Hi #HU2
lapply (frees_weak … HU2 0 ?) -HU2
/5 width=7 by frees_bind_dx_O, frees_flat_dx, lpx_pair/
]
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".
+include "basic_2/multiple/lleq_lreq.ma".
include "basic_2/multiple/lleq_llor.ma".
-include "basic_2/reduction/cpx_leq.ma".
+include "basic_2/reduction/cpx_lreq.ma".
include "basic_2/reduction/cpx_lleq.ma".
include "basic_2/reduction/lpx_frees.ma".
(* Note: contains a proof of llpx_cpx_conf *)
lemma lleq_lpx_trans: ∀h,g,G,L2,K2. ⦃G, L2⦄ ⊢ ➡[h, g] K2 →
- ∀L1,T,d. L1 ≡[T, d] L2 →
- ∃∃K1. ⦃G, L1⦄ ⊢ ➡[h, g] K1 & K1 ≡[T, d] K2.
-#h #g #G #L2 #K2 #HLK2 #L1 #T #d #HL12
+ ∀L1,T,l. L1 ≡[T, l] L2 →
+ ∃∃K1. ⦃G, L1⦄ ⊢ ➡[h, g] K1 & K1 ≡[T, l] K2.
+#h #g #G #L2 #K2 #HLK2 #L1 #T #l #HL12
lapply (lpx_fwd_length … HLK2) #H1
lapply (lleq_fwd_length … HL12) #H2
-lapply (lpx_sn_llpx_sn … T … d HLK2) // -HLK2 #H
+lapply (lpx_sn_llpx_sn … T … l HLK2) // -HLK2 #H
lapply (lleq_llpx_sn_trans … HL12 … H) /2 width=3 by lleq_cpx_trans/ -HL12 -H #H
-elim (llor_total L1 K2 T d) // -H1 -H2 #K1 #HLK1
+elim (llor_total L1 K2 T l) // -H1 -H2 #K1 #HLK1
lapply (llpx_sn_llor_dx_sym … H … HLK1)
[ /2 width=6 by cpx_frees_trans/
| /3 width=10 by cpx_llpx_sn_conf, cpx_inv_lift1, cpx_lift/
/3 width=4 by lpx_pair, fqu_bind_dx, ex3_intro/
| #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 (drop_O1_le (Ⓕ) (e+1) K1)
+| #G1 #L1 #L #T1 #U1 #m #HL1 #HTU1 #K1 #H1KL1 #H2KL1
+ elim (drop_O1_le (Ⓕ) (m+1) K1)
[ #K #HK1 lapply (lleq_inv_lift_le … H2KL1 … HK1 HL1 … HTU1 ?) -H2KL1 //
#H2KL elim (lpx_drop_trans_O1 … H1KL1 … HL1) -L1
#K0 #HK10 #H1KL lapply (drop_mono … HK10 … HK1) -HK10 #H destruct
]
qed-.
-fact leq_lpx_trans_lleq_aux: ∀h,g,G,L1,L0,d,e. L1 ⩬[d, e] L0 → e = ∞ →
- ∀L2. ⦃G, L0⦄ ⊢ ➡[h, g] L2 →
- ∃∃L. 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
+fact lreq_lpx_trans_lleq_aux: ∀h,g,G,L1,L0,l,m. L1 ⩬[l, m] L0 → m = ∞ →
+ ∀L2. ⦃G, L0⦄ ⊢ ➡[h, g] L2 →
+ ∃∃L. L ⩬[l, m] L2 & ⦃G, L1⦄ ⊢ ➡[h, g] L &
+ (∀T. L0 ≡[T, l] L2 ↔ L1 ≡[T, l] L).
+#h #g #G #L1 #L0 #l #m #H elim H -L1 -L0 -l -m
+[ #l #m #_ #L2 #H >(lpx_inv_atom1 … H) -H
/3 width=5 by ex3_intro, conj/
-| #I1 #I0 #L1 #L0 #V1 #V0 #_ #_ #He destruct
-| #I #L1 #L0 #V1 #e #HL10 #IHL10 #He #Y #H
+| #I1 #I0 #L1 #L0 #V1 #V0 #_ #_ #Hm destruct
+| #I #L1 #L0 #V1 #m #HL10 #IHL10 #Hm #Y #H
elim (lpx_inv_pair1 … H) -H #L2 #V2 #HL02 #HV02 #H destruct
- lapply (ysucc_inv_Y_dx … He) -He #He
+ lapply (ysucc_inv_Y_dx … Hm) -Hm #Hm
elim (IHL10 … HL02) // -IHL10 -HL02 #L #HL2 #HL1 #IH
- @(ex3_intro … (L.ⓑ{I}V2)) /3 width=3 by lpx_pair, leq_cpx_trans, leq_pair/
+ @(ex3_intro … (L.ⓑ{I}V2)) /3 width=3 by lpx_pair, lreq_cpx_trans, lreq_pair/
#T elim (IH T) #HL0dx #HL0sn
- @conj #H @(lleq_leq_repl … H) -H /3 width=1 by leq_sym, leq_pair_O_Y/
-| #I1 #I0 #L1 #L0 #V1 #V0 #d #e #HL10 #IHL10 #He #Y #H
+ @conj #H @(lleq_lreq_repl … H) -H /3 width=1 by lreq_sym, lreq_pair_O_Y/
+| #I1 #I0 #L1 #L0 #V1 #V0 #l #m #HL10 #IHL10 #Hm #Y #H
elim (lpx_inv_pair1 … H) -H #L2 #V2 #HL02 #HV02 #H destruct
elim (IHL10 … HL02) // -IHL10 -HL02 #L #HL2 #HL1 #IH
- @(ex3_intro … (L.ⓑ{I1}V1)) /3 width=1 by lpx_pair, leq_succ/
+ @(ex3_intro … (L.ⓑ{I1}V1)) /3 width=1 by lpx_pair, lreq_succ/
#T elim (IH T) #HL0dx #HL0sn
- @conj #H @(lleq_leq_repl … H) -H /3 width=1 by leq_sym, leq_succ/
+ @conj #H @(lleq_lreq_repl … H) -H /3 width=1 by lreq_sym, lreq_succ/
]
qed-.
-lemma leq_lpx_trans_lleq: ∀h,g,G,L1,L0,d. L1 ⩬[d, ∞] L0 →
- ∀L2. ⦃G, L0⦄ ⊢ ➡[h, g] L2 →
- ∃∃L. 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-.
+lemma lreq_lpx_trans_lleq: ∀h,g,G,L1,L0,l. L1 ⩬[l, ∞] L0 →
+ ∀L2. ⦃G, L0⦄ ⊢ ➡[h, g] L2 →
+ ∃∃L. L ⩬[l, ∞] L2 & ⦃G, L1⦄ ⊢ ➡[h, g] L &
+ (∀T. L0 ≡[T, l] L2 ↔ L1 ≡[T, l] L).
+/2 width=1 by lreq_lpx_trans_lleq_aux/ qed-.
(* Properties on basic relocation *******************************************)
-lemma aaa_lift: ∀G,L1,T1,A. ⦃G, L1⦄ ⊢ T1 ⁝ A → ∀L2,s,d,e. ⬇[s, d, e] L2 ≡ L1 →
- ∀T2. ⬆[d, e] T1 ≡ T2 → ⦃G, L2⦄ ⊢ T2 ⁝ A.
+lemma aaa_lift: ∀G,L1,T1,A. ⦃G, L1⦄ ⊢ T1 ⁝ A → ∀L2,s,l,m. ⬇[s, l, m] L2 ≡ L1 →
+ ∀T2. ⬆[l, m] T1 ≡ T2 → ⦃G, L2⦄ ⊢ T2 ⁝ A.
#G #L1 #T1 #A #H elim H -G -L1 -T1 -A
-[ #G #L1 #k #L2 #s #d #e #_ #T2 #H
+[ #G #L1 #k #L2 #s #l #m #_ #T2 #H
>(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
+| #I #G #L1 #K1 #V1 #B #i #HLK1 #_ #IHB #L2 #s #l #m #HL21 #T2 #H
+ elim (lift_inv_lref1 … H) -H * #Hil #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
+ elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hil #K2 #V2 #HK21 #HV12 #H destruct
/3 width=9 by aaa_lref/
| 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
+| #a #G #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #s #l #m #HL21 #X #H
elim (lift_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct
/4 width=5 by aaa_abbr, drop_skip/
-| #a #G #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #s #d #e #HL21 #X #H
+| #a #G #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #s #l #m #HL21 #X #H
elim (lift_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct
/4 width=5 by aaa_abst, drop_skip/
-| #G #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #s #d #e #HL21 #X #H
+| #G #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #s #l #m #HL21 #X #H
elim (lift_inv_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct
/3 width=5 by aaa_appl/
-| #G #L1 #V1 #T1 #A #_ #_ #IH1 #IH2 #L2 #s #d #e #HL21 #X #H
+| #G #L1 #V1 #T1 #A #_ #_ #IH1 #IH2 #L2 #s #l #m #HL21 #X #H
elim (lift_inv_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct
/3 width=5 by aaa_cast/
]
qed.
-lemma aaa_inv_lift: ∀G,L2,T2,A. ⦃G, L2⦄ ⊢ T2 ⁝ A → ∀L1,s,d,e. ⬇[s, d, e] L2 ≡ L1 →
- ∀T1. ⬆[d, e] T1 ≡ T2 → ⦃G, L1⦄ ⊢ T1 ⁝ A.
+lemma aaa_inv_lift: ∀G,L2,T2,A. ⦃G, L2⦄ ⊢ T2 ⁝ A → ∀L1,s,l,m. ⬇[s, l, m] L2 ≡ L1 →
+ ∀T1. ⬆[l, m] T1 ≡ T2 → ⦃G, L1⦄ ⊢ T1 ⁝ A.
#G #L2 #T2 #A #H elim H -G -L2 -T2 -A
-[ #G #L2 #k #L1 #s #d #e #_ #T1 #H
+[ #G #L2 #k #L1 #s #l #m #_ #T1 #H
>(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
+| #I #G #L2 #K2 #V2 #B #i #HLK2 #_ #IHB #L1 #s #l #m #HL21 #T1 #H
+ elim (lift_inv_lref2 … H) -H * #Hil #H destruct
[ 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
+| #a #G #L2 #V2 #T2 #B #A #_ #_ #IHB #IHA #L1 #s #l #m #HL21 #X #H
elim (lift_inv_bind2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct
/4 width=5 by aaa_abbr, drop_skip/
-| #a #G #L2 #V2 #T2 #B #A #_ #_ #IHB #IHA #L1 #s #d #e #HL21 #X #H
+| #a #G #L2 #V2 #T2 #B #A #_ #_ #IHB #IHA #L1 #s #l #m #HL21 #X #H
elim (lift_inv_bind2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct
/4 width=5 by aaa_abst, drop_skip/
-| #G #L2 #V2 #T2 #B #A #_ #_ #IHB #IHA #L1 #s #d #e #HL21 #X #H
+| #G #L2 #V2 #T2 #B #A #_ #_ #IHB #IHA #L1 #s #l #m #HL21 #X #H
elim (lift_inv_flat2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct
/3 width=5 by aaa_appl/
-| #G #L2 #V2 #T2 #A #_ #_ #IH1 #IH2 #L1 #s #d #e #HL21 #X #H
+| #G #L2 #V2 #T2 #A #_ #_ #IH1 #IH2 #L1 #s #l #m #HL21 #X #H
elim (lift_inv_flat2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct
/3 width=5 by aaa_cast/
]
#G #L1 #L2 #T2 #A #s #des #H elim H -L1 -L2 -des
[ #L #T1 #H #HT1
<(lifts_inv_nil … H) -H //
-| #L1 #L #L2 #des #d #e #_ #HL2 #IHL1 #T1 #H #HT1
+| #L1 #L #L2 #des #l #m #_ #HL2 #IHL1 #T1 #H #HT1
elim (lifts_inv_cons … H) -H /3 width=10 by aaa_lift/
]
qed.
(* activate genv *)
inductive da (h:sh) (g:sd h): relation4 genv lenv term nat ≝
-| da_sort: ∀G,L,k,l. deg h g k l → da h g G L (⋆k) l
-| da_ldef: ∀G,L,K,V,i,l. ⬇[i] L ≡ K.ⓓV → da h g G K V l → da h g G L (#i) l
-| da_ldec: ∀G,L,K,W,i,l. ⬇[i] L ≡ K.ⓛW → da h g G K W l → da h g G L (#i) (l+1)
-| da_bind: ∀a,I,G,L,V,T,l. da h g G (L.ⓑ{I}V) T l → da h g G L (ⓑ{a,I}V.T) l
-| da_flat: ∀I,G,L,V,T,l. da h g G L T l → da h g G L (ⓕ{I}V.T) l
+| da_sort: ∀G,L,k,d. deg h g k d → da h g G L (⋆k) d
+| da_ldef: ∀G,L,K,V,i,d. ⬇[i] L ≡ K.ⓓV → da h g G K V d → da h g G L (#i) d
+| da_ldec: ∀G,L,K,W,i,d. ⬇[i] L ≡ K.ⓛW → da h g G K W d → da h g G L (#i) (d+1)
+| da_bind: ∀a,I,G,L,V,T,d. da h g G (L.ⓑ{I}V) T d → da h g G L (ⓑ{a,I}V.T) d
+| da_flat: ∀I,G,L,V,T,d. da h g G L T d → da h g G L (ⓕ{I}V.T) d
.
interpretation "degree assignment (term)"
- 'Degree h g G L T l = (da h g G L T l).
+ 'Degree h g G L T d = (da h g G L T d).
(* Basic inversion lemmas ***************************************************)
-fact da_inv_sort_aux: ∀h,g,G,L,T,l. ⦃G, L⦄ ⊢ T ▪[h, g] l →
- ∀k0. T = ⋆k0 → deg h g k0 l.
-#h #g #G #L #T #l * -G -L -T -l
-[ #G #L #k #l #Hkl #k0 #H destruct //
-| #G #L #K #V #i #l #_ #_ #k0 #H destruct
-| #G #L #K #W #i #l #_ #_ #k0 #H destruct
-| #a #I #G #L #V #T #l #_ #k0 #H destruct
-| #I #G #L #V #T #l #_ #k0 #H destruct
+fact da_inv_sort_aux: ∀h,g,G,L,T,d. ⦃G, L⦄ ⊢ T ▪[h, g] d →
+ ∀k0. T = ⋆k0 → deg h g k0 d.
+#h #g #G #L #T #d * -G -L -T -d
+[ #G #L #k #d #Hkd #k0 #H destruct //
+| #G #L #K #V #i #d #_ #_ #k0 #H destruct
+| #G #L #K #W #i #d #_ #_ #k0 #H destruct
+| #a #I #G #L #V #T #d #_ #k0 #H destruct
+| #I #G #L #V #T #d #_ #k0 #H destruct
]
qed-.
-lemma da_inv_sort: ∀h,g,G,L,k,l. ⦃G, L⦄ ⊢ ⋆k ▪[h, g] l → deg h g k l.
+lemma da_inv_sort: ∀h,g,G,L,k,d. ⦃G, L⦄ ⊢ ⋆k ▪[h, g] d → deg h g k d.
/2 width=5 by da_inv_sort_aux/ qed-.
-fact da_inv_lref_aux: ∀h,g,G,L,T,l. ⦃G, L⦄ ⊢ T ▪[h, g] l → ∀j. T = #j →
- (∃∃K,V. ⬇[j] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V ▪[h, g] l) ∨
- (∃∃K,W,l0. ⬇[j] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W ▪[h, g] l0 &
- l = l0 + 1
+fact da_inv_lref_aux: ∀h,g,G,L,T,d. ⦃G, L⦄ ⊢ T ▪[h, g] d → ∀j. T = #j →
+ (∃∃K,V. ⬇[j] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V ▪[h, g] d) ∨
+ (∃∃K,W,d0. ⬇[j] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W ▪[h, g] d0 &
+ d = d0 + 1
).
-#h #g #G #L #T #l * -G -L -T -l
-[ #G #L #k #l #_ #j #H destruct
-| #G #L #K #V #i #l #HLK #HV #j #H destruct /3 width=4 by ex2_2_intro, or_introl/
-| #G #L #K #W #i #l #HLK #HW #j #H destruct /3 width=6 by ex3_3_intro, or_intror/
-| #a #I #G #L #V #T #l #_ #j #H destruct
-| #I #G #L #V #T #l #_ #j #H destruct
+#h #g #G #L #T #d * -G -L -T -d
+[ #G #L #k #d #_ #j #H destruct
+| #G #L #K #V #i #d #HLK #HV #j #H destruct /3 width=4 by ex2_2_intro, or_introl/
+| #G #L #K #W #i #d #HLK #HW #j #H destruct /3 width=6 by ex3_3_intro, or_intror/
+| #a #I #G #L #V #T #d #_ #j #H destruct
+| #I #G #L #V #T #d #_ #j #H destruct
]
qed-.
-lemma da_inv_lref: ∀h,g,G,L,j,l. ⦃G, L⦄ ⊢ #j ▪[h, g] l →
- (∃∃K,V. ⬇[j] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V ▪[h, g] l) ∨
- (∃∃K,W,l0. ⬇[j] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W ▪[h, g] l0 & l = l0+1).
+lemma da_inv_lref: ∀h,g,G,L,j,d. ⦃G, L⦄ ⊢ #j ▪[h, g] d →
+ (∃∃K,V. ⬇[j] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V ▪[h, g] d) ∨
+ (∃∃K,W,d0. ⬇[j] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W ▪[h, g] d0 & d = d0+1).
/2 width=3 by da_inv_lref_aux/ qed-.
-fact da_inv_gref_aux: ∀h,g,G,L,T,l. ⦃G, L⦄ ⊢ T ▪[h, g] l → ∀p0. T = §p0 → ⊥.
-#h #g #G #L #T #l * -G -L -T -l
-[ #G #L #k #l #_ #p0 #H destruct
-| #G #L #K #V #i #l #_ #_ #p0 #H destruct
-| #G #L #K #W #i #l #_ #_ #p0 #H destruct
-| #a #I #G #L #V #T #l #_ #p0 #H destruct
-| #I #G #L #V #T #l #_ #p0 #H destruct
+fact da_inv_gref_aux: ∀h,g,G,L,T,d. ⦃G, L⦄ ⊢ T ▪[h, g] d → ∀p0. T = §p0 → ⊥.
+#h #g #G #L #T #d * -G -L -T -d
+[ #G #L #k #d #_ #p0 #H destruct
+| #G #L #K #V #i #d #_ #_ #p0 #H destruct
+| #G #L #K #W #i #d #_ #_ #p0 #H destruct
+| #a #I #G #L #V #T #d #_ #p0 #H destruct
+| #I #G #L #V #T #d #_ #p0 #H destruct
]
qed-.
-lemma da_inv_gref: ∀h,g,G,L,p,l. ⦃G, L⦄ ⊢ §p ▪[h, g] l → ⊥.
+lemma da_inv_gref: ∀h,g,G,L,p,d. ⦃G, L⦄ ⊢ §p ▪[h, g] d → ⊥.
/2 width=9 by da_inv_gref_aux/ qed-.
-fact da_inv_bind_aux: ∀h,g,G,L,T,l. ⦃G, L⦄ ⊢ T ▪[h, g] l →
- ∀b,J,X,Y. T = ⓑ{b,J}Y.X → ⦃G, L.ⓑ{J}Y⦄ ⊢ X ▪[h, g] l.
-#h #g #G #L #T #l * -G -L -T -l
-[ #G #L #k #l #_ #b #J #X #Y #H destruct
-| #G #L #K #V #i #l #_ #_ #b #J #X #Y #H destruct
-| #G #L #K #W #i #l #_ #_ #b #J #X #Y #H destruct
-| #a #I #G #L #V #T #l #HT #b #J #X #Y #H destruct //
-| #I #G #L #V #T #l #_ #b #J #X #Y #H destruct
+fact da_inv_bind_aux: ∀h,g,G,L,T,d. ⦃G, L⦄ ⊢ T ▪[h, g] d →
+ ∀b,J,X,Y. T = ⓑ{b,J}Y.X → ⦃G, L.ⓑ{J}Y⦄ ⊢ X ▪[h, g] d.
+#h #g #G #L #T #d * -G -L -T -d
+[ #G #L #k #d #_ #b #J #X #Y #H destruct
+| #G #L #K #V #i #d #_ #_ #b #J #X #Y #H destruct
+| #G #L #K #W #i #d #_ #_ #b #J #X #Y #H destruct
+| #a #I #G #L #V #T #d #HT #b #J #X #Y #H destruct //
+| #I #G #L #V #T #d #_ #b #J #X #Y #H destruct
]
qed-.
-lemma da_inv_bind: ∀h,g,b,J,G,L,Y,X,l. ⦃G, L⦄ ⊢ ⓑ{b,J}Y.X ▪[h, g] l → ⦃G, L.ⓑ{J}Y⦄ ⊢ X ▪[h, g] l.
+lemma da_inv_bind: ∀h,g,b,J,G,L,Y,X,d. ⦃G, L⦄ ⊢ ⓑ{b,J}Y.X ▪[h, g] d → ⦃G, L.ⓑ{J}Y⦄ ⊢ X ▪[h, g] d.
/2 width=4 by da_inv_bind_aux/ qed-.
-fact da_inv_flat_aux: ∀h,g,G,L,T,l. ⦃G, L⦄ ⊢ T ▪[h, g] l →
- ∀J,X,Y. T = ⓕ{J}Y.X → ⦃G, L⦄ ⊢ X ▪[h, g] l.
-#h #g #G #L #T #l * -G -L -T -l
-[ #G #L #k #l #_ #J #X #Y #H destruct
-| #G #L #K #V #i #l #_ #_ #J #X #Y #H destruct
-| #G #L #K #W #i #l #_ #_ #J #X #Y #H destruct
-| #a #I #G #L #V #T #l #_ #J #X #Y #H destruct
-| #I #G #L #V #T #l #HT #J #X #Y #H destruct //
+fact da_inv_flat_aux: ∀h,g,G,L,T,d. ⦃G, L⦄ ⊢ T ▪[h, g] d →
+ ∀J,X,Y. T = ⓕ{J}Y.X → ⦃G, L⦄ ⊢ X ▪[h, g] d.
+#h #g #G #L #T #d * -G -L -T -d
+[ #G #L #k #d #_ #J #X #Y #H destruct
+| #G #L #K #V #i #d #_ #_ #J #X #Y #H destruct
+| #G #L #K #W #i #d #_ #_ #J #X #Y #H destruct
+| #a #I #G #L #V #T #d #_ #J #X #Y #H destruct
+| #I #G #L #V #T #d #HT #J #X #Y #H destruct //
]
qed-.
-lemma da_inv_flat: ∀h,g,J,G,L,Y,X,l. ⦃G, L⦄ ⊢ ⓕ{J}Y.X ▪[h, g] l → ⦃G, L⦄ ⊢ X ▪[h, g] l.
+lemma da_inv_flat: ∀h,g,J,G,L,Y,X,d. ⦃G, L⦄ ⊢ ⓕ{J}Y.X ▪[h, g] d → ⦃G, L⦄ ⊢ X ▪[h, g] d.
/2 width=5 by da_inv_flat_aux/ qed-.
(* Properties on atomic arity assignment for terms **************************)
-lemma aaa_da: ∀h,g,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∃l. ⦃G, L⦄ ⊢ T ▪[h, g] l.
+lemma aaa_da: ∀h,g,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∃d. ⦃G, L⦄ ⊢ T ▪[h, g] d.
#h #g #G #L #T #A #H elim H -G -L -T -A
[ #G #L #k elim (deg_total h g k) /3 width=2 by da_sort, ex_intro/
| * #G #L #K #V #B #i #HLK #_ * /3 width=5 by da_ldef, da_ldec, ex_intro/
(* Main properties **********************************************************)
-theorem da_mono: ∀h,g,G,L,T,l1. ⦃G, L⦄ ⊢ T ▪[h, g] l1 →
- ∀l2. ⦃G, L⦄ ⊢ T ▪[h, g] l2 → l1 = l2.
-#h #g #G #L #T #l1 #H elim H -G -L -T -l1
-[ #G #L #k #l1 #Hkl1 #l2 #H
- lapply (da_inv_sort … H) -G -L #Hkl2
- >(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 ]
+theorem da_mono: ∀h,g,G,L,T,d1. ⦃G, L⦄ ⊢ T ▪[h, g] d1 →
+ ∀d2. ⦃G, L⦄ ⊢ T ▪[h, g] d2 → d1 = d2.
+#h #g #G #L #T #d1 #H elim H -G -L -T -d1
+[ #G #L #k #d1 #Hkd1 #d2 #H
+ lapply (da_inv_sort … H) -G -L #Hkd2
+ >(deg_mono … Hkd2 … Hkd1) -h -k -d2 //
+| #G #L #K #V #i #d1 #HLK #_ #IHV #d2 #H
+ elim (da_inv_lref … H) -H * #K0 #V0 [| #d0 ] #HLK0 #HV0 [| #Hd0 ]
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 ]
+| #G #L #K #W #i #d1 #HLK #_ #IHW #d2 #H
+ elim (da_inv_lref … H) -H * #K0 #W0 [| #d0 ] #HLK0 #HW0 [| #Hd0 ]
lapply (drop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct /3 width=1/
-| #a #I #G #L #V #T #l1 #_ #IHT #l2 #H
+| #a #I #G #L #V #T #d1 #_ #IHT #d2 #H
lapply (da_inv_bind … H) -H /2 width=1/
-| #I #G #L #V #T #l1 #_ #IHT #l2 #H
+| #I #G #L #V #T #d1 #_ #IHT #d2 #H
lapply (da_inv_flat … H) -H /2 width=1/
]
qed-.
(* Properties on relocation *************************************************)
-lemma da_lift: ∀h,g,G,L1,T1,l. ⦃G, L1⦄ ⊢ T1 ▪[h, g] l →
- ∀L2,s,d,e. ⬇[s, d, e] L2 ≡ L1 → ∀T2. ⬆[d, e] T1 ≡ T2 →
- ⦃G, L2⦄ ⊢ T2 ▪[h, g] l.
-#h #g #G #L1 #T1 #l #H elim H -G -L1 -T1 -l
-[ #G #L1 #k #l #Hkl #L2 #s #d #e #_ #X #H
+lemma da_lift: ∀h,g,G,L1,T1,d. ⦃G, L1⦄ ⊢ T1 ▪[h, g] d →
+ ∀L2,s,l,m. ⬇[s, l, m] L2 ≡ L1 → ∀T2. ⬆[l, m] T1 ≡ T2 →
+ ⦃G, L2⦄ ⊢ T2 ▪[h, g] d.
+#h #g #G #L1 #T1 #d #H elim H -G -L1 -T1 -d
+[ #G #L1 #k #d #Hkd #L2 #s #l #m #_ #X #H
>(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
+| #G #L1 #K1 #V1 #i #d #HLK1 #_ #IHV1 #L2 #s #l #m #HL21 #X #H
+ elim (lift_inv_lref1 … H) * #Hil #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
+ elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hil #K2 #V2 #HK21 #HV12 #H destruct
/3 width=9 by da_ldef/
| 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
+| #G #L1 #K1 #W1 #i #d #HLK1 #_ #IHW1 #L2 #s #l #m #HL21 #X #H
+ elim (lift_inv_lref1 … H) * #Hil #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
+ elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hil #K2 #W2 #HK21 #HW12 #H destruct
/3 width=8 by da_ldec/
| 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
+| #a #I #G #L1 #V1 #T1 #d #_ #IHT1 #L2 #s #l #m #HL21 #X #H
elim (lift_inv_bind1 … H) -H #V2 #T2 #HV12 #HU12 #H destruct
/4 width=5 by da_bind, drop_skip/
-| #I #G #L1 #V1 #T1 #l #_ #IHT1 #L2 #s #d #e #HL21 #X #H
+| #I #G #L1 #V1 #T1 #d #_ #IHT1 #L2 #s #l #m #HL21 #X #H
elim (lift_inv_flat1 … H) -H #V2 #T2 #HV12 #HU12 #H destruct
/3 width=5 by da_flat/
]
(* Inversion lemmas on relocation *******************************************)
-lemma da_inv_lift: ∀h,g,G,L2,T2,l. ⦃G, L2⦄ ⊢ T2 ▪[h, g] l →
- ∀L1,s,d,e. ⬇[s, d, e] L2 ≡ L1 → ∀T1. ⬆[d, e] T1 ≡ T2 →
- ⦃G, L1⦄ ⊢ T1 ▪[h, g] l.
-#h #g #G #L2 #T2 #l #H elim H -G -L2 -T2 -l
-[ #G #L2 #k #l #Hkl #L1 #s #d #e #_ #X #H
+lemma da_inv_lift: ∀h,g,G,L2,T2,d. ⦃G, L2⦄ ⊢ T2 ▪[h, g] d →
+ ∀L1,s,l,m. ⬇[s, l, m] L2 ≡ L1 → ∀T1. ⬆[l, m] T1 ≡ T2 →
+ ⦃G, L1⦄ ⊢ T1 ▪[h, g] d.
+#h #g #G #L2 #T2 #d #H elim H -G -L2 -T2 -d
+[ #G #L2 #k #d #Hkd #L1 #s #l #m #_ #X #H
>(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 ]
+| #G #L2 #K2 #V2 #i #d #HLK2 #HV2 #IHV2 #L1 #s #l #m #HL21 #X #H
+ elim (lift_inv_lref2 … H) * #Hil #H destruct [ -HV2 | -IHV2 ]
[ 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 ]
+| #G #L2 #K2 #W2 #i #d #HLK2 #HW2 #IHW2 #L1 #s #l #m #HL21 #X #H
+ elim (lift_inv_lref2 … H) * #Hil #H destruct [ -HW2 | -IHW2 ]
[ 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
+| #a #I #G #L2 #V2 #T2 #d #_ #IHT2 #L1 #s #l #m #HL21 #X #H
elim (lift_inv_bind2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct
/4 width=5 by da_bind, drop_skip/
-| #I #G #L2 #V2 #T2 #l #_ #IHT2 #L1 #s #d #e #HL21 #X #H
+| #I #G #L2 #V2 #T2 #d #_ #IHT2 #L1 #s #l #m #HL21 #X #H
elim (lift_inv_flat2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct
/3 width=5 by da_flat/
]
qed.
(* Note: the constant 0 cannot be generalized *)
-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.
+lemma lsuba_drop_O1_conf: ∀G,L1,L2. G ⊢ L1 ⫃⁝ L2 → ∀K1,s,m. ⬇[s, 0, m] L1 ≡ K1 →
+ ∃∃K2. G ⊢ K1 ⫃⁝ K2 & ⬇[s, 0, m] 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 (drop_inv_O1_pair1 … H) -H * #He #HLK1
+| #I #L1 #L2 #V #_ #IHL12 #K1 #s #m #H
+ elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK1
[ destruct
elim (IHL12 L1 s 0) -IHL12 // #X #HL12 #H
<(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 (drop_inv_O1_pair1 … H) -H * #He #HLK1
+| #L1 #L2 #W #V #A #HV #HW #_ #IHL12 #K1 #s #m #H
+ elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK1
[ destruct
elim (IHL12 L1 s 0) -IHL12 // #X #HL12 #H
<(drop_inv_O2 … H) in HL12; -H /3 width=3 by lsuba_beta, drop_pair, ex2_intro/
qed-.
(* Note: the constant 0 cannot be generalized *)
-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.
+lemma lsuba_drop_O1_trans: ∀G,L1,L2. G ⊢ L1 ⫃⁝ L2 → ∀K2,s,m. ⬇[s, 0, m] L2 ≡ K2 →
+ ∃∃K1. G ⊢ K1 ⫃⁝ K2 & ⬇[s, 0, m] 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 (drop_inv_O1_pair1 … H) -H * #He #HLK2
+| #I #L1 #L2 #V #_ #IHL12 #K2 #s #m #H
+ elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK2
[ destruct
elim (IHL12 L2 s 0) -IHL12 // #X #HL12 #H
<(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 (drop_inv_O1_pair1 … H) -H * #He #HLK2
+| #L1 #L2 #W #V #A #HV #HW #_ #IHL12 #K2 #s #m #H
+ elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK2
[ destruct
elim (IHL12 L2 s 0) -IHL12 // #X #HL12 #H
<(drop_inv_O2 … H) in HL12; -H /3 width=3 by lsuba_beta, drop_pair, ex2_intro/
| lsubd_atom: lsubd h g G (⋆) (⋆)
| lsubd_pair: ∀I,L1,L2,V. lsubd h g G L1 L2 →
lsubd h g G (L1.ⓑ{I}V) (L2.ⓑ{I}V)
-| lsubd_beta: ∀L1,L2,W,V,l. ⦃G, L1⦄ ⊢ V ▪[h, g] l+1 → ⦃G, L2⦄ ⊢ W ▪[h, g] l →
+| lsubd_beta: ∀L1,L2,W,V,d. ⦃G, L1⦄ ⊢ V ▪[h, g] d+1 → ⦃G, L2⦄ ⊢ W ▪[h, g] d →
lsubd h g G L1 L2 → lsubd h g G (L1.ⓓⓝW.V) (L2.ⓛW)
.
#h #g #G #L1 #L2 * -L1 -L2
[ //
| #I #L1 #L2 #V #_ #H destruct
-| #L1 #L2 #W #V #l #_ #_ #_ #H destruct
+| #L1 #L2 #W #V #d #_ #_ #_ #H destruct
]
qed-.
fact lsubd_inv_pair1_aux: ∀h,g,G,L1,L2. G ⊢ L1 ⫃▪[h, g] L2 →
∀I,K1,X. L1 = K1.ⓑ{I}X →
(∃∃K2. G ⊢ K1 ⫃▪[h, g] K2 & L2 = K2.ⓑ{I}X) ∨
- ∃∃K2,W,V,l. ⦃G, K1⦄ ⊢ V ▪[h, g] l+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] l &
+ ∃∃K2,W,V,d. ⦃G, K1⦄ ⊢ V ▪[h, g] d+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] d &
G ⊢ K1 ⫃▪[h, g] K2 &
I = Abbr & L2 = K2.ⓛW & X = ⓝW.V.
#h #g #G #L1 #L2 * -L1 -L2
[ #J #K1 #X #H destruct
| #I #L1 #L2 #V #HL12 #J #K1 #X #H destruct /3 width=3 by ex2_intro, or_introl/
-| #L1 #L2 #W #V #l #HV #HW #HL12 #J #K1 #X #H destruct /3 width=9 by ex6_4_intro, or_intror/
+| #L1 #L2 #W #V #d #HV #HW #HL12 #J #K1 #X #H destruct /3 width=9 by ex6_4_intro, or_intror/
]
qed-.
lemma lsubd_inv_pair1: ∀h,g,I,G,K1,L2,X. G ⊢ K1.ⓑ{I}X ⫃▪[h, g] L2 →
(∃∃K2. G ⊢ K1 ⫃▪[h, g] K2 & L2 = K2.ⓑ{I}X) ∨
- ∃∃K2,W,V,l. ⦃G, K1⦄ ⊢ V ▪[h, g] l+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] l &
+ ∃∃K2,W,V,d. ⦃G, K1⦄ ⊢ V ▪[h, g] d+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] d &
G ⊢ K1 ⫃▪[h, g] K2 &
I = Abbr & L2 = K2.ⓛW & X = ⓝW.V.
/2 width=3 by lsubd_inv_pair1_aux/ qed-.
#h #g #G #L1 #L2 * -L1 -L2
[ //
| #I #L1 #L2 #V #_ #H destruct
-| #L1 #L2 #W #V #l #_ #_ #_ #H destruct
+| #L1 #L2 #W #V #d #_ #_ #_ #H destruct
]
qed-.
fact lsubd_inv_pair2_aux: ∀h,g,G,L1,L2. G ⊢ L1 ⫃▪[h, g] L2 →
∀I,K2,W. L2 = K2.ⓑ{I}W →
(∃∃K1. G ⊢ K1 ⫃▪[h, g] K2 & L1 = K1.ⓑ{I}W) ∨
- ∃∃K1,V,l. ⦃G, K1⦄ ⊢ V ▪[h, g] l+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] l &
+ ∃∃K1,V,d. ⦃G, K1⦄ ⊢ V ▪[h, g] d+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] d &
G ⊢ K1 ⫃▪[h, g] K2 & I = Abst & L1 = K1. ⓓⓝW.V.
#h #g #G #L1 #L2 * -L1 -L2
[ #J #K2 #U #H destruct
| #I #L1 #L2 #V #HL12 #J #K2 #U #H destruct /3 width=3 by ex2_intro, or_introl/
-| #L1 #L2 #W #V #l #HV #HW #HL12 #J #K2 #U #H destruct /3 width=7 by ex5_3_intro, or_intror/
+| #L1 #L2 #W #V #d #HV #HW #HL12 #J #K2 #U #H destruct /3 width=7 by ex5_3_intro, or_intror/
]
qed-.
lemma lsubd_inv_pair2: ∀h,g,I,G,L1,K2,W. G ⊢ L1 ⫃▪[h, g] K2.ⓑ{I}W →
(∃∃K1. G ⊢ K1 ⫃▪[h, g] K2 & L1 = K1.ⓑ{I}W) ∨
- ∃∃K1,V,l. ⦃G, K1⦄ ⊢ V ▪[h, g] l+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] l &
+ ∃∃K1,V,d. ⦃G, K1⦄ ⊢ V ▪[h, g] d+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] d &
G ⊢ K1 ⫃▪[h, g] K2 & I = Abst & L1 = K1. ⓓⓝW.V.
/2 width=3 by lsubd_inv_pair2_aux/ qed-.
(* Note: the constant 0 cannot be generalized *)
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.
+ ∀K1,s,m. ⬇[s, 0, m] L1 ≡ K1 →
+ ∃∃K2. G ⊢ K1 ⫃▪[h, g] K2 & ⬇[s, 0, m] 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 (drop_inv_O1_pair1 … H) -H * #He #HLK1
+| #I #L1 #L2 #V #_ #IHL12 #K1 #s #m #H
+ elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK1
[ destruct
elim (IHL12 L1 s 0) -IHL12 // #X #HL12 #H
<(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 (drop_inv_O1_pair1 … H) -H * #He #HLK1
+| #L1 #L2 #W #V #d #HV #HW #_ #IHL12 #K1 #s #m #H
+ elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK1
[ destruct
elim (IHL12 L1 s 0) -IHL12 // #X #HL12 #H
<(drop_inv_O2 … H) in HL12; -H /3 width=3 by lsubd_beta, drop_pair, ex2_intro/
(* Note: the constant 0 cannot be generalized *)
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.
+ ∀K2,s,m. ⬇[s, 0, m] L2 ≡ K2 →
+ ∃∃K1. G ⊢ K1 ⫃▪[h, g] K2 & ⬇[s, 0, m] 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 (drop_inv_O1_pair1 … H) -H * #He #HLK2
+| #I #L1 #L2 #V #_ #IHL12 #K2 #s #m #H
+ elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK2
[ destruct
elim (IHL12 L2 s 0) -IHL12 // #X #HL12 #H
<(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 (drop_inv_O1_pair1 … H) -H * #He #HLK2
+| #L1 #L2 #W #V #d #HV #HW #_ #IHL12 #K2 #s #m #H
+ elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK2
[ destruct
elim (IHL12 L2 s 0) -IHL12 // #X #HL12 #H
<(drop_inv_O2 … H) in HL12; -H /3 width=3 by lsubd_beta, drop_pair, ex2_intro/
(* Properties on degree assignment ******************************************)
-lemma lsubd_da_trans: ∀h,g,G,L2,T,l. ⦃G, L2⦄ ⊢ T ▪[h, g] l →
- ∀L1. G ⊢ L1 ⫃▪[h, g] L2 → ⦃G, L1⦄ ⊢ T ▪[h, g] l.
-#h #g #G #L2 #T #l #H elim H -G -L2 -T -l
+lemma lsubd_da_trans: ∀h,g,G,L2,T,d. ⦃G, L2⦄ ⊢ T ▪[h, g] d →
+ ∀L1. G ⊢ L1 ⫃▪[h, g] L2 → ⦃G, L1⦄ ⊢ T ▪[h, g] d.
+#h #g #G #L2 #T #d #H elim H -G -L2 -T -d
[ /2 width=1/
-| #G #L2 #K2 #V #i #l #HLK2 #_ #IHV #L1 #HL12
+| #G #L2 #K2 #V #i #d #HLK2 #_ #IHV #L1 #HL12
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
+ | #W #d0 #_ #_ #_ #H destruct
]
-| #G #L2 #K2 #W #i #l #HLK2 #HW #IHW #L1 #HL12
+| #G #L2 #K2 #W #i #d #HLK2 #HW #IHW #L1 #HL12
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
+ | #V #d0 #HV #H0W #_ #_ #H destruct
lapply (da_mono … H0W … HW) -H0W -HW #H destruct /3 width=7/
]
| /4 width=1/
]
qed-.
-lemma lsubd_da_conf: ∀h,g,G,L1,T,l. ⦃G, L1⦄ ⊢ T ▪[h, g] l →
- ∀L2. G ⊢ L1 ⫃▪[h, g] L2 → ⦃G, L2⦄ ⊢ T ▪[h, g] l.
-#h #g #G #L1 #T #l #H elim H -G -L1 -T -l
+lemma lsubd_da_conf: ∀h,g,G,L1,T,d. ⦃G, L1⦄ ⊢ T ▪[h, g] d →
+ ∀L2. G ⊢ L1 ⫃▪[h, g] L2 → ⦃G, L2⦄ ⊢ T ▪[h, g] d.
+#h #g #G #L1 #T #d #H elim H -G -L1 -T -d
[ /2 width=1/
-| #G #L1 #K1 #V #i #l #HLK1 #HV #IHV #L2 #HL12
+| #G #L1 #K1 #V #i #d #HLK1 #HV #IHV #L2 #HL12
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
+ | #W0 #V0 #d0 #HV0 #HW0 #_ #_ #H1 #H2 destruct
lapply (da_inv_flat … HV) -HV #H0V0
lapply (da_mono … H0V0 … HV0) -H0V0 -HV0 #H destruct /2 width=4/
]
-| #G #L1 #K1 #W #i #l #HLK1 #HW #IHW #L2 #HL12
+| #G #L1 #K1 #W #i #d #HLK1 #HW #IHW #L2 #HL12
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
+ | #W0 #V0 #d0 #HV0 #HW0 #_ #H destruct
]
| /4 width=1/
| /3 width=1/
| #I #L1 #L #Y #HL1 #IHL1 #X #H
elim (lsubd_inv_pair1 … H) -H * #L2
[ #HL2 #H destruct /3 width=1/
- | #W #V #l #HV #HW #HL2 #H1 #H2 #H3 destruct
+ | #W #V #d #HV #HW #HL2 #H1 #H2 #H3 destruct
/3 width=3 by lsubd_beta, lsubd_da_trans/
]
-| #L1 #L #W #V #l #HV #HW #HL1 #IHL1 #X #H
+| #L1 #L #W #V #d #HV #HW #HL1 #IHL1 #X #H
elim (lsubd_inv_pair1 … H) -H * #L2
[ #HL2 #H destruct /3 width=5 by lsubd_beta, lsubd_da_conf/
- | #W0 #V0 #l0 #_ #_ #_ #H destruct
+ | #W0 #V0 #d0 #_ #_ #_ #H destruct
]
]
qed-.
(* sort degree specification *)
record sd (h:sh): Type[0] ≝ {
deg : relation nat; (* degree of the sort *)
- deg_total: ∀k. ∃l. deg k l; (* functional relation axioms *)
- deg_mono : ∀k,l1,l2. deg k l1 → deg k l2 → l1 = l2;
- deg_next : ∀k,l. deg k l → deg (next h k) (l - 1) (* compatibility condition *)
+ deg_total: ∀k. ∃d. deg k d; (* functional relation axioms *)
+ deg_mono : ∀k,d1,d2. deg k d1 → deg k d2 → d1 = d2;
+ deg_next : ∀k,d. deg k d → deg (next h k) (d - 1) (* compatibility condition *)
}.
(* Notable specifications ***************************************************)
-definition deg_O: relation nat ≝ λk,l. l = 0.
+definition deg_O: relation nat ≝ λk,d. d = 0.
definition sd_O: ∀h. sd h ≝ λh. mk_sd h deg_O ….
/2 width=2 by le_n_O_to_eq, le_n, ex_intro/ defined.
inductive deg_SO (h:sh) (k:nat) (k0:nat): predicate nat ≝
-| deg_SO_pos : ∀l0. (next h)^l0 k0 = k → deg_SO h k k0 (l0 + 1)
-| deg_SO_zero: ((∃l0. (next h)^l0 k0 = k) → ⊥) → deg_SO h k k0 0
+| deg_SO_pos : ∀d0. (next h)^d0 k0 = k → deg_SO h k k0 (d0 + 1)
+| deg_SO_zero: ((∃d0. (next h)^d0 k0 = k) → ⊥) → deg_SO h k k0 0
.
-fact deg_SO_inv_pos_aux: ∀h,k,k0,l0. deg_SO h k k0 l0 → ∀l. l0 = l + 1 →
- (next h)^l k0 = k.
-#h #k #k0 #l0 * -l0
-[ #l0 #Hl0 #l #H
+fact deg_SO_inv_pos_aux: ∀h,k,k0,d0. deg_SO h k k0 d0 → ∀d. d0 = d + 1 →
+ (next h)^d k0 = k.
+#h #k #k0 #d0 * -d0
+[ #d0 #Hd0 #d #H
lapply (injective_plus_l … H) -H #H destruct //
-| #_ #l0 <plus_n_Sm #H destruct
+| #_ #d0 <plus_n_Sm #H destruct
]
qed.
-lemma deg_SO_inv_pos: ∀h,k,k0,l0. deg_SO h k k0 (l0 + 1) → (next h)^l0 k0 = k.
+lemma deg_SO_inv_pos: ∀h,k,k0,d0. deg_SO h k k0 (d0 + 1) → (next h)^d0 k0 = k.
/2 width=3 by deg_SO_inv_pos_aux/ qed-.
lemma deg_SO_refl: ∀h,k. deg_SO h k k 1.
qed.
lemma deg_SO_gt: ∀h,k1,k2. k1 < k2 → deg_SO h k1 k2 0.
-#h #k1 #k2 #HK12 @deg_SO_zero * #l elim l -l normalize
+#h #k1 #k2 #HK12 @deg_SO_zero * #d elim d -d normalize
[ #H destruct
elim (lt_refl_false … HK12)
-| #l #_ #H
- lapply (next_lt h ((next h)^l k2)) >H -H #H
+| #d #_ #H
+ lapply (next_lt h ((next h)^d k2)) >H -H #H
lapply (transitive_lt … H HK12) -k1 #H1
- lapply (nexts_le h k2 l) #H2
- lapply (le_to_lt_to_lt … H2 H1) -h -l #H
+ lapply (nexts_le h k2 d) #H2
+ lapply (le_to_lt_to_lt … H2 H1) -h -d #H
elim (lt_refl_false … H)
]
qed.
[ #k0
lapply (nexts_dec h k0 k) *
[ * /3 width=2 by deg_SO_pos, ex_intro/ | /4 width=2 by deg_SO_zero, ex_intro/ ]
-| #K0 #l1 #l2 * [ #l01 ] #H1 * [1,3: #l02 ] #H2 //
+| #K0 #d1 #d2 * [ #d01 ] #H1 * [1,3: #d02 ] #H2 //
[ < H2 in H1; -H2 #H
lapply (nexts_inj … H) -H #H destruct //
| elim H1 /2 width=2 by ex_intro/
| elim H2 /2 width=2 by ex_intro/
]
-| #k0 #l0 *
- [ #l #H destruct elim l -l normalize
+| #k0 #d0 *
+ [ #d #H destruct elim d -d normalize
/2 width=1 by deg_SO_gt, deg_SO_pos, next_lt/
- | #H1 @deg_SO_zero * #l #H2 destruct
- @H1 -H1 @(ex_intro … (S l)) /2 width=1 by sym_eq/ (**) (* explicit constructor *)
+ | #H1 @deg_SO_zero * #d #H2 destruct
+ @H1 -H1 @(ex_intro … (S d)) /2 width=1 by sym_eq/ (**) (* explicit constructor *)
]
]
defined.
-let rec sd_l (h:sh) (k:nat) (l:nat) on l : sd h ≝
- match l with
+let rec sd_d (h:sh) (k:nat) (d:nat) on d : sd h ≝
+ match d with
[ O ⇒ sd_O h
- | S l ⇒ match l with
+ | S d ⇒ match d with
[ O ⇒ sd_SO h k
- | _ ⇒ sd_l h (next h k) l
+ | _ ⇒ sd_d h (next h k) d
]
].
(* Basic inversion lemmas ***************************************************)
-lemma deg_inv_pred: ∀h,g,k,l. deg h g (next h k) (l+1) → deg h g k (l+2).
-#h #g #k #l #H1
-elim (deg_total h g k) #l0 #H0
+lemma deg_inv_pred: ∀h,g,k,d. deg h g (next h k) (d+1) → deg h g k (d+2).
+#h #g #k #d #H1
+elim (deg_total h g k) #d0 #H0
lapply (deg_next … H0) #H2
lapply (deg_mono … H1 H2) -H1 -H2 #H
-<(associative_plus l 1 1) >H <plus_minus_m_m /2 width=3 by transitive_le/
+<(associative_plus d 1 1) >H <plus_minus_m_m /2 width=3 by transitive_le/
qed-.
-lemma deg_inv_prec: ∀h,g,k,l,l0. deg h g ((next h)^l k) (l0+1) → deg h g k (l+l0+1).
-#h #g #k #l @(nat_ind_plus … l) -l //
-#l #IHl #l0 >iter_SO #H
-lapply (deg_inv_pred … H) -H <(associative_plus l0 1 1) #H
-lapply (IHl … H) -IHl -H //
+lemma deg_inv_prec: ∀h,g,k,d,d0. deg h g ((next h)^d k) (d0+1) → deg h g k (d+d0+1).
+#h #g #k #d @(nat_ind_plus … d) -d //
+#d #IHd #d0 >iter_SO #H
+lapply (deg_inv_pred … H) -H <(associative_plus d0 1 1) #H
+lapply (IHd … H) -IHd -H //
qed-.
(* Basic properties *********************************************************)
-lemma deg_iter: ∀h,g,k,l1,l2. deg h g k l1 → deg h g ((next h)^l2 k) (l1-l2).
-#h #g #k #l1 #l2 @(nat_ind_plus … l2) -l2 [ <minus_n_O // ]
-#l2 #IHl2 #Hkl1 >iter_SO <minus_plus /3 width=1 by deg_next/
+lemma deg_iter: ∀h,g,k,d1,d2. deg h g k d1 → deg h g ((next h)^d2 k) (d1-d2).
+#h #g #k #d1 #d2 @(nat_ind_plus … d2) -d2 [ <minus_n_O // ]
+#d2 #IHd2 #Hkd1 >iter_SO <minus_plus /3 width=1 by deg_next/
qed.
-lemma deg_next_SO: ∀h,g,k,l. deg h g k (l+1) → deg h g (next h k) l.
-#h #g #k #l #Hkl
-lapply (deg_next … Hkl) -Hkl <minus_plus_m_m //
+lemma deg_next_SO: ∀h,g,k,d. deg h g k (d+1) → deg h g (next h k) d.
+#h #g #k #d #Hkd
+lapply (deg_next … Hkd) -Hkd <minus_plus_m_m //
qed-.
-lemma sd_l_SS: ∀h,k,l. sd_l h k (l + 2) = sd_l h (next h k) (l + 1).
-#h #k #l <plus_n_Sm <plus_n_Sm //
+lemma sd_d_SS: ∀h,k,d. sd_d h k (d + 2) = sd_d h (next h k) (d + 1).
+#h #k #d <plus_n_Sm <plus_n_Sm //
qed.
-lemma sd_l_correct: ∀h,l,k. deg h (sd_l h k l) k l.
-#h #l @(nat_ind_plus … l) -l // #l @(nat_ind_plus … l) -l /3 width=1 by deg_inv_pred/
+lemma sd_d_correct: ∀h,d,k. deg h (sd_d h k d) k d.
+#h #d @(nat_ind_plus … d) -d // #d @(nat_ind_plus … d) -d /3 width=1 by deg_inv_pred/
qed.
(* Basic properties *********************************************************)
-lemma nexts_le: ∀h,k,l. k ≤ (next h)^l k.
-#h #k #l elim l -l // normalize #l #IHl
-lapply (next_lt h ((next h)^l k)) #H
-lapply (le_to_lt_to_lt … IHl H) -IHl -H /2 width=2 by lt_to_le/
+lemma nexts_le: ∀h,k,d. k ≤ (next h)^d k.
+#h #k #d elim d -d // normalize #d #IHd
+lapply (next_lt h ((next h)^d k)) #H
+lapply (le_to_lt_to_lt … IHd H) -IHd -H /2 width=2 by lt_to_le/
qed.
-lemma nexts_lt: ∀h,k,l. k < (next h)^(l+1) k.
-#h #k #l >iter_SO
-lapply (nexts_le h k l) #H
+lemma nexts_lt: ∀h,k,d. k < (next h)^(d+1) k.
+#h #k #d >iter_SO
+lapply (nexts_le h k d) #H
@(le_to_lt_to_lt … H) //
qed.
-axiom nexts_dec: ∀h,k1,k2. Decidable (∃l. (next h)^l k1 = k2).
+axiom nexts_dec: ∀h,k1,k2. Decidable (∃d. (next h)^d k1 = k2).
-axiom nexts_inj: ∀h,k,l1,l2. (next h)^l1 k = (next h)^l2 k → l1 = l2.
+axiom nexts_inj: ∀h,k,d1,d2. (next h)^d1 k = (next h)^d2 k → d1 = d2.
(* activate genv *)
inductive cpy: ynat → ynat → relation4 genv lenv term term ≝
-| cpy_atom : ∀I,G,L,d,e. cpy d e G L (⓪{I}) (⓪{I})
-| cpy_subst: ∀I,G,L,K,V,W,i,d,e. d ≤ yinj i → i < d+e →
- ⬇[i] L ≡ K.ⓑ{I}V → ⬆[0, i+1] V ≡ W → cpy d e G L (#i) W
-| cpy_bind : ∀a,I,G,L,V1,V2,T1,T2,d,e.
- cpy d e G L V1 V2 → cpy (⫯d) e G (L.ⓑ{I}V1) T1 T2 →
- cpy d e G L (ⓑ{a,I}V1.T1) (ⓑ{a,I}V2.T2)
-| cpy_flat : ∀I,G,L,V1,V2,T1,T2,d,e.
- cpy d e G L V1 V2 → cpy d e G L T1 T2 →
- cpy d e G L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2)
+| cpy_atom : ∀I,G,L,l,m. cpy l m G L (⓪{I}) (⓪{I})
+| cpy_subst: ∀I,G,L,K,V,W,i,l,m. l ≤ yinj i → i < l+m →
+ ⬇[i] L ≡ K.ⓑ{I}V → ⬆[0, i+1] V ≡ W → cpy l m G L (#i) W
+| cpy_bind : ∀a,I,G,L,V1,V2,T1,T2,l,m.
+ cpy l m G L V1 V2 → cpy (⫯l) m G (L.ⓑ{I}V1) T1 T2 →
+ cpy l m G L (ⓑ{a,I}V1.T1) (ⓑ{a,I}V2.T2)
+| cpy_flat : ∀I,G,L,V1,V2,T1,T2,l,m.
+ cpy l m G L V1 V2 → cpy l m G L T1 T2 →
+ cpy l m G L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2)
.
interpretation "context-sensitive extended ordinary substritution (term)"
- 'PSubst G L T1 d e T2 = (cpy d e G L T1 T2).
+ 'PSubst G L T1 l m T2 = (cpy l m G L T1 T2).
(* Basic properties *********************************************************)
-lemma lsuby_cpy_trans: ∀G,d,e. lsub_trans … (cpy d e G) (lsuby d e).
-#G #d #e #L1 #T1 #T2 #H elim H -G -L1 -T1 -T2 -d -e
+lemma lsuby_cpy_trans: ∀G,l,m. lsub_trans … (cpy l m G) (lsuby l m).
+#G #l #m #L1 #T1 #T2 #H elim H -G -L1 -T1 -T2 -l -m
[ //
-| #I #G #L1 #K1 #V #W #i #d #e #Hdi #Hide #HLK1 #HVW #L2 #HL12
+| #I #G #L1 #K1 #V #W #i #l #m #Hli #Hilm #HLK1 #HVW #L2 #HL12
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/
]
qed-.
-lemma cpy_refl: ∀G,T,L,d,e. ⦃G, L⦄ ⊢ T ▶[d, e] T.
+lemma cpy_refl: ∀G,T,L,l,m. ⦃G, L⦄ ⊢ T ▶[l, m] T.
#G #T elim T -T // * /2 width=1 by cpy_bind, cpy_flat/
qed.
(* Basic_1: was: subst1_ex *)
-lemma cpy_full: ∀I,G,K,V,T1,L,d. ⬇[d] L ≡ K.ⓑ{I}V →
- ∃∃T2,T. ⦃G, L⦄ ⊢ T1 ▶[d, 1] T2 & ⬆[d, 1] T ≡ T2.
+lemma cpy_full: ∀I,G,K,V,T1,L,l. ⬇[l] L ≡ K.ⓑ{I}V →
+ ∃∃T2,T. ⦃G, L⦄ ⊢ T1 ▶[l, 1] T2 & ⬆[l, 1] T ≡ T2.
#I #G #K #V #T1 elim T1 -T1
-[ * #i #L #d #HLK
+[ * #i #L #l #HLK
/2 width=4 by lift_sort, lift_gref, ex2_2_intro/
- elim (lt_or_eq_or_gt i d) #Hid
+ elim (lt_or_eq_or_gt i l) #Hil
/3 width=4 by lift_lref_ge_minus, lift_lref_lt, ex2_2_intro/
destruct
elim (lift_total V 0 (i+1)) #W #HVW
elim (lift_split … HVW i i)
/4 width=5 by cpy_subst, ylt_inj, ex2_2_intro/
-| * [ #a ] #J #W1 #U1 #IHW1 #IHU1 #L #d #HLK
+| * [ #a ] #J #W1 #U1 #IHW1 #IHU1 #L #l #HLK
elim (IHW1 … HLK) -IHW1 #W2 #W #HW12 #HW2
- [ elim (IHU1 (L.ⓑ{J}W1) (d+1)) -IHU1
+ [ elim (IHU1 (L.ⓑ{J}W1) (l+1)) -IHU1
/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/
]
qed-.
-lemma cpy_weak: ∀G,L,T1,T2,d1,e1. ⦃G, L⦄ ⊢ T1 ▶[d1, e1] T2 →
- ∀d2,e2. d2 ≤ d1 → d1 + e1 ≤ d2 + e2 →
- ⦃G, L⦄ ⊢ T1 ▶[d2, e2] T2.
-#G #L #T1 #T2 #d1 #e1 #H elim H -G -L -T1 -T2 -d1 -e1 //
+lemma cpy_weak: ∀G,L,T1,T2,l1,m1. ⦃G, L⦄ ⊢ T1 ▶[l1, m1] T2 →
+ ∀l2,m2. l2 ≤ l1 → l1 + m1 ≤ l2 + m2 →
+ ⦃G, L⦄ ⊢ T1 ▶[l2, m2] T2.
+#G #L #T1 #T2 #l1 #m1 #H elim H -G -L -T1 -T2 -l1 -m1 //
[ /3 width=5 by cpy_subst, ylt_yle_trans, yle_trans/
| /4 width=3 by cpy_bind, ylt_yle_trans, yle_succ/
| /3 width=1 by cpy_flat/
]
qed-.
-lemma cpy_weak_top: ∀G,L,T1,T2,d,e.
- ⦃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
+lemma cpy_weak_top: ∀G,L,T1,T2,l,m.
+ ⦃G, L⦄ ⊢ T1 ▶[l, m] T2 → ⦃G, L⦄ ⊢ T1 ▶[l, |L| - l] T2.
+#G #L #T1 #T2 #l #m #H elim H -G -L -T1 -T2 -l -m //
+[ #I #G #L #K #V #W #i #l #m #Hli #_ #HLK #HVW
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 *)
]
qed-.
-lemma cpy_weak_full: ∀G,L,T1,T2,d,e.
- ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶[0, |L|] T2.
-#G #L #T1 #T2 #d #e #HT12
-lapply (cpy_weak … HT12 0 (d + e) ? ?) -HT12
+lemma cpy_weak_full: ∀G,L,T1,T2,l,m.
+ ⦃G, L⦄ ⊢ T1 ▶[l, m] T2 → ⦃G, L⦄ ⊢ T1 ▶[0, |L|] T2.
+#G #L #T1 #T2 #l #m #HT12
+lapply (cpy_weak … HT12 0 (l + m) ? ?) -HT12
/2 width=2 by cpy_weak_top/
qed-.
-lemma cpy_split_up: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 → ∀i. i ≤ d + e →
- ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[d, i-d] T & ⦃G, L⦄ ⊢ T ▶[i, d+e-i] T2.
-#G #L #T1 #T2 #d #e #H elim H -G -L -T1 -T2 -d -e
+lemma cpy_split_up: ∀G,L,T1,T2,l,m. ⦃G, L⦄ ⊢ T1 ▶[l, m] T2 → ∀i. i ≤ l + m →
+ ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[l, i-l] T & ⦃G, L⦄ ⊢ T ▶[i, l+m-i] T2.
+#G #L #T1 #T2 #l #m #H elim H -G -L -T1 -T2 -l -m
[ /2 width=3 by ex2_intro/
-| #I #G #L #K #V #W #i #d #e #Hdi #Hide #HLK #HVW #j #Hjde
- elim (ylt_split i j) [ -Hide -Hjde | -Hdi ]
+| #I #G #L #K #V #W #i #l #m #Hli #Hilm #HLK #HVW #j #Hjlm
+ elim (ylt_split i j) [ -Hilm -Hjlm | -Hli ]
/4 width=9 by cpy_subst, ylt_yle_trans, ex2_intro/
-| #a #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hide
+| #a #I #G #L #V1 #V2 #T1 #T2 #l #m #_ #_ #IHV12 #IHT12 #i #Hilm
elim (IHV12 i) -IHV12 // #V
- elim (IHT12 (i+1)) -IHT12 /2 width=1 by yle_succ/ -Hide
+ elim (IHT12 (i+1)) -IHT12 /2 width=1 by yle_succ/ -Hilm
>yplus_SO2 >yplus_succ1 #T #HT1 #HT2
lapply (lsuby_cpy_trans … HT2 (L.ⓑ{I}V) ?) -HT2
/3 width=5 by lsuby_succ, ex2_intro, cpy_bind/
-| #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hide
- elim (IHV12 i) -IHV12 // elim (IHT12 i) -IHT12 // -Hide
+| #I #G #L #V1 #V2 #T1 #T2 #l #m #_ #_ #IHV12 #IHT12 #i #Hilm
+ elim (IHV12 i) -IHV12 // elim (IHT12 i) -IHT12 // -Hilm
/3 width=5 by ex2_intro, cpy_flat/
]
qed-.
-lemma cpy_split_down: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 → ∀i. i ≤ d + e →
- ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[i, d+e-i] T & ⦃G, L⦄ ⊢ T ▶[d, i-d] T2.
-#G #L #T1 #T2 #d #e #H elim H -G -L -T1 -T2 -d -e
+lemma cpy_split_down: ∀G,L,T1,T2,l,m. ⦃G, L⦄ ⊢ T1 ▶[l, m] T2 → ∀i. i ≤ l + m →
+ ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[i, l+m-i] T & ⦃G, L⦄ ⊢ T ▶[l, i-l] T2.
+#G #L #T1 #T2 #l #m #H elim H -G -L -T1 -T2 -l -m
[ /2 width=3 by ex2_intro/
-| #I #G #L #K #V #W #i #d #e #Hdi #Hide #HLK #HVW #j #Hjde
- elim (ylt_split i j) [ -Hide -Hjde | -Hdi ]
+| #I #G #L #K #V #W #i #l #m #Hli #Hilm #HLK #HVW #j #Hjlm
+ elim (ylt_split i j) [ -Hilm -Hjlm | -Hli ]
/4 width=9 by cpy_subst, ylt_yle_trans, ex2_intro/
-| #a #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hide
+| #a #I #G #L #V1 #V2 #T1 #T2 #l #m #_ #_ #IHV12 #IHT12 #i #Hilm
elim (IHV12 i) -IHV12 // #V
- elim (IHT12 (i+1)) -IHT12 /2 width=1 by yle_succ/ -Hide
+ elim (IHT12 (i+1)) -IHT12 /2 width=1 by yle_succ/ -Hilm
>yplus_SO2 >yplus_succ1 #T #HT1 #HT2
lapply (lsuby_cpy_trans … HT2 (L.ⓑ{I}V) ?) -HT2
/3 width=5 by lsuby_succ, ex2_intro, cpy_bind/
-| #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hide
- elim (IHV12 i) -IHV12 // elim (IHT12 i) -IHT12 // -Hide
+| #I #G #L #V1 #V2 #T1 #T2 #l #m #_ #_ #IHV12 #IHT12 #i #Hilm
+ elim (IHV12 i) -IHV12 // elim (IHT12 i) -IHT12 // -Hilm
/3 width=5 by ex2_intro, cpy_flat/
]
qed-.
(* Basic forward lemmas *****************************************************)
-lemma cpy_fwd_up: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶[dt, et] U2 →
- ∀T1,d,e. ⬆[d, e] T1 ≡ U1 →
- d ≤ dt → d + e ≤ dt + et →
- ∃∃T2. ⦃G, L⦄ ⊢ U1 ▶[d+e, dt+et-(d+e)] U2 & ⬆[d, e] T2 ≡ U2.
-#G #L #U1 #U2 #dt #et #H elim H -G -L -U1 -U2 -dt -et
-[ * #i #G #L #dt #et #T1 #d #e #H #_
+lemma cpy_fwd_up: ∀G,L,U1,U2,lt,mt. ⦃G, L⦄ ⊢ U1 ▶[lt, mt] U2 →
+ ∀T1,l,m. ⬆[l, m] T1 ≡ U1 →
+ l ≤ lt → l + m ≤ lt + mt →
+ ∃∃T2. ⦃G, L⦄ ⊢ U1 ▶[l+m, lt+mt-(l+m)] U2 & ⬆[l, m] T2 ≡ U2.
+#G #L #U1 #U2 #lt #mt #H elim H -G -L -U1 -U2 -lt -mt
+[ * #i #G #L #lt #mt #T1 #l #m #H #_
[ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3 by ex2_intro/
- | elim (lift_inv_lref2 … H) -H * #Hid #H destruct /3 width=3 by lift_lref_ge_minus, lift_lref_lt, ex2_intro/
+ | elim (lift_inv_lref2 … H) -H * #Hil #H destruct /3 width=3 by lift_lref_ge_minus, lift_lref_lt, ex2_intro/
| lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3 by ex2_intro/
]
-| #I #G #L #K #V #W #i #dt #et #Hdti #Hidet #HLK #HVW #T1 #d #e #H #Hddt #Hdedet
- elim (lift_inv_lref2 … H) -H * #Hid #H destruct [ -V -Hidet -Hdedet | -Hdti -Hddt ]
- [ elim (ylt_yle_false … Hddt) -Hddt /3 width=3 by yle_ylt_trans, ylt_inj/
- | elim (le_inv_plus_l … Hid) #Hdie #Hei
- elim (lift_split … HVW d (i-e+1) ? ? ?) [2,3,4: /2 width=1 by le_S_S, le_S/ ] -Hdie
- #T2 #_ >plus_minus // <minus_minus /2 width=1 by le_S/ <minus_n_n <plus_n_O #H -Hei
+| #I #G #L #K #V #W #i #lt #mt #Hlti #Hilmt #HLK #HVW #T1 #l #m #H #Hllt #Hlmlmt
+ elim (lift_inv_lref2 … H) -H * #Hil #H destruct [ -V -Hilmt -Hlmlmt | -Hlti -Hllt ]
+ [ elim (ylt_yle_false … Hllt) -Hllt /3 width=3 by yle_ylt_trans, ylt_inj/
+ | elim (le_inv_plus_l … Hil) #Hlim #Hmi
+ elim (lift_split … HVW l (i-m+1) ? ? ?) [2,3,4: /2 width=1 by le_S_S, le_S/ ] -Hlim
+ #T2 #_ >plus_minus // <minus_minus /2 width=1 by le_S/ <minus_n_n <plus_n_O #H -Hmi
@(ex2_intro … H) -H @(cpy_subst … HLK HVW) /2 width=1 by yle_inj/ >ymax_pre_sn_comm // (**) (* explicit constructor *)
]
-| #a #I #G #L #W1 #W2 #U1 #U2 #dt #et #_ #_ #IHW12 #IHU12 #X #d #e #H #Hddt #Hdedet
+| #a #I #G #L #W1 #W2 #U1 #U2 #lt #mt #_ #_ #IHW12 #IHU12 #X #l #m #H #Hllt #Hlmlmt
elim (lift_inv_bind2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct
elim (IHW12 … HVW1) -V1 -IHW12 //
elim (IHU12 … HTU1) -T1 -IHU12 /2 width=1 by yle_succ/
<yplus_inj >yplus_SO2 >yplus_succ1 >yplus_succ1
/3 width=2 by cpy_bind, lift_bind, ex2_intro/
-| #I #G #L #W1 #W2 #U1 #U2 #dt #et #_ #_ #IHW12 #IHU12 #X #d #e #H #Hddt #Hdedet
+| #I #G #L #W1 #W2 #U1 #U2 #lt #mt #_ #_ #IHW12 #IHU12 #X #l #m #H #Hllt #Hlmlmt
elim (lift_inv_flat2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct
elim (IHW12 … HVW1) -V1 -IHW12 // elim (IHU12 … HTU1) -T1 -IHU12
/3 width=2 by cpy_flat, lift_flat, ex2_intro/
]
qed-.
-lemma cpy_fwd_tw: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 → ♯{T1} ≤ ♯{T2}.
-#G #L #T1 #T2 #d #e #H elim H -G -L -T1 -T2 -d -e normalize
+lemma cpy_fwd_tw: ∀G,L,T1,T2,l,m. ⦃G, L⦄ ⊢ T1 ▶[l, m] T2 → ♯{T1} ≤ ♯{T2}.
+#G #L #T1 #T2 #l #m #H elim H -G -L -T1 -T2 -l -m normalize
/3 width=1 by monotonic_le_plus_l, le_plus/
qed-.
(* Basic inversion lemmas ***************************************************)
-fact cpy_inv_atom1_aux: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 → ∀J. T1 = ⓪{J} →
+fact cpy_inv_atom1_aux: ∀G,L,T1,T2,l,m. ⦃G, L⦄ ⊢ T1 ▶[l, m] T2 → ∀J. T1 = ⓪{J} →
T2 = ⓪{J} ∨
- ∃∃I,K,V,i. d ≤ yinj i & i < d + e &
+ ∃∃I,K,V,i. l ≤ yinj i & i < l + m &
⬇[i] L ≡ K.ⓑ{I}V &
⬆[O, i+1] V ≡ T2 &
J = LRef i.
-#G #L #T1 #T2 #d #e * -G -L -T1 -T2 -d -e
-[ #I #G #L #d #e #J #H destruct /2 width=1 by or_introl/
-| #I #G #L #K #V #T2 #i #d #e #Hdi #Hide #HLK #HVT2 #J #H destruct /3 width=9 by ex5_4_intro, or_intror/
-| #a #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #J #H destruct
-| #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #J #H destruct
+#G #L #T1 #T2 #l #m * -G -L -T1 -T2 -l -m
+[ #I #G #L #l #m #J #H destruct /2 width=1 by or_introl/
+| #I #G #L #K #V #T2 #i #l #m #Hli #Hilm #HLK #HVT2 #J #H destruct /3 width=9 by ex5_4_intro, or_intror/
+| #a #I #G #L #V1 #V2 #T1 #T2 #l #m #_ #_ #J #H destruct
+| #I #G #L #V1 #V2 #T1 #T2 #l #m #_ #_ #J #H destruct
]
qed-.
-lemma cpy_inv_atom1: ∀I,G,L,T2,d,e. ⦃G, L⦄ ⊢ ⓪{I} ▶[d, e] T2 →
+lemma cpy_inv_atom1: ∀I,G,L,T2,l,m. ⦃G, L⦄ ⊢ ⓪{I} ▶[l, m] T2 →
T2 = ⓪{I} ∨
- ∃∃J,K,V,i. d ≤ yinj i & i < d + e &
+ ∃∃J,K,V,i. l ≤ yinj i & i < l + m &
⬇[i] L ≡ K.ⓑ{J}V &
⬆[O, i+1] V ≡ T2 &
I = LRef i.
/2 width=4 by cpy_inv_atom1_aux/ qed-.
(* Basic_1: was: subst1_gen_sort *)
-lemma cpy_inv_sort1: ∀G,L,T2,k,d,e. ⦃G, L⦄ ⊢ ⋆k ▶[d, e] T2 → T2 = ⋆k.
-#G #L #T2 #k #d #e #H
+lemma cpy_inv_sort1: ∀G,L,T2,k,l,m. ⦃G, L⦄ ⊢ ⋆k ▶[l, m] T2 → T2 = ⋆k.
+#G #L #T2 #k #l #m #H
elim (cpy_inv_atom1 … H) -H //
* #I #K #V #i #_ #_ #_ #_ #H destruct
qed-.
(* Basic_1: was: subst1_gen_lref *)
-lemma cpy_inv_lref1: ∀G,L,T2,i,d,e. ⦃G, L⦄ ⊢ #i ▶[d, e] T2 →
+lemma cpy_inv_lref1: ∀G,L,T2,i,l,m. ⦃G, L⦄ ⊢ #i ▶[l, m] T2 →
T2 = #i ∨
- ∃∃I,K,V. d ≤ i & i < d + e &
+ ∃∃I,K,V. l ≤ i & i < l + m &
⬇[i] L ≡ K.ⓑ{I}V &
⬆[O, i+1] V ≡ T2.
-#G #L #T2 #i #d #e #H
+#G #L #T2 #i #l #m #H
elim (cpy_inv_atom1 … H) -H /2 width=1 by or_introl/
-* #I #K #V #j #Hdj #Hjde #HLK #HVT2 #H destruct /3 width=5 by ex4_3_intro, or_intror/
+* #I #K #V #j #Hlj #Hjlm #HLK #HVT2 #H destruct /3 width=5 by ex4_3_intro, or_intror/
qed-.
-lemma cpy_inv_gref1: ∀G,L,T2,p,d,e. ⦃G, L⦄ ⊢ §p ▶[d, e] T2 → T2 = §p.
-#G #L #T2 #p #d #e #H
+lemma cpy_inv_gref1: ∀G,L,T2,p,l,m. ⦃G, L⦄ ⊢ §p ▶[l, m] T2 → T2 = §p.
+#G #L #T2 #p #l #m #H
elim (cpy_inv_atom1 … H) -H //
* #I #K #V #i #_ #_ #_ #_ #H destruct
qed-.
-fact cpy_inv_bind1_aux: ∀G,L,U1,U2,d,e. ⦃G, L⦄ ⊢ U1 ▶[d, e] U2 →
+fact cpy_inv_bind1_aux: ∀G,L,U1,U2,l,m. ⦃G, L⦄ ⊢ U1 ▶[l, m] U2 →
∀a,I,V1,T1. U1 = ⓑ{a,I}V1.T1 →
- ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶[d, e] V2 &
- ⦃G, L. ⓑ{I}V1⦄ ⊢ T1 ▶[⫯d, e] T2 &
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶[l, m] V2 &
+ ⦃G, L. ⓑ{I}V1⦄ ⊢ T1 ▶[⫯l, m] T2 &
U2 = ⓑ{a,I}V2.T2.
-#G #L #U1 #U2 #d #e * -G -L -U1 -U2 -d -e
-[ #I #G #L #d #e #b #J #W1 #U1 #H destruct
-| #I #G #L #K #V #W #i #d #e #_ #_ #_ #_ #b #J #W1 #U1 #H destruct
-| #a #I #G #L #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #b #J #W1 #U1 #H destruct /2 width=5 by ex3_2_intro/
-| #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #b #J #W1 #U1 #H destruct
+#G #L #U1 #U2 #l #m * -G -L -U1 -U2 -l -m
+[ #I #G #L #l #m #b #J #W1 #U1 #H destruct
+| #I #G #L #K #V #W #i #l #m #_ #_ #_ #_ #b #J #W1 #U1 #H destruct
+| #a #I #G #L #V1 #V2 #T1 #T2 #l #m #HV12 #HT12 #b #J #W1 #U1 #H destruct /2 width=5 by ex3_2_intro/
+| #I #G #L #V1 #V2 #T1 #T2 #l #m #_ #_ #b #J #W1 #U1 #H destruct
]
qed-.
-lemma cpy_inv_bind1: ∀a,I,G,L,V1,T1,U2,d,e. ⦃G, L⦄ ⊢ ⓑ{a,I} V1. T1 ▶[d, e] U2 →
- ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶[d, e] V2 &
- ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ▶[⫯d, e] T2 &
+lemma cpy_inv_bind1: ∀a,I,G,L,V1,T1,U2,l,m. ⦃G, L⦄ ⊢ ⓑ{a,I} V1. T1 ▶[l, m] U2 →
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶[l, m] V2 &
+ ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ▶[⫯l, m] T2 &
U2 = ⓑ{a,I}V2.T2.
/2 width=3 by cpy_inv_bind1_aux/ qed-.
-fact cpy_inv_flat1_aux: ∀G,L,U1,U2,d,e. ⦃G, L⦄ ⊢ U1 ▶[d, e] U2 →
+fact cpy_inv_flat1_aux: ∀G,L,U1,U2,l,m. ⦃G, L⦄ ⊢ U1 ▶[l, m] U2 →
∀I,V1,T1. U1 = ⓕ{I}V1.T1 →
- ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶[d, e] V2 &
- ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 &
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶[l, m] V2 &
+ ⦃G, L⦄ ⊢ T1 ▶[l, m] T2 &
U2 = ⓕ{I}V2.T2.
-#G #L #U1 #U2 #d #e * -G -L -U1 -U2 -d -e
-[ #I #G #L #d #e #J #W1 #U1 #H destruct
-| #I #G #L #K #V #W #i #d #e #_ #_ #_ #_ #J #W1 #U1 #H destruct
-| #a #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #J #W1 #U1 #H destruct
-| #I #G #L #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #J #W1 #U1 #H destruct /2 width=5 by ex3_2_intro/
+#G #L #U1 #U2 #l #m * -G -L -U1 -U2 -l -m
+[ #I #G #L #l #m #J #W1 #U1 #H destruct
+| #I #G #L #K #V #W #i #l #m #_ #_ #_ #_ #J #W1 #U1 #H destruct
+| #a #I #G #L #V1 #V2 #T1 #T2 #l #m #_ #_ #J #W1 #U1 #H destruct
+| #I #G #L #V1 #V2 #T1 #T2 #l #m #HV12 #HT12 #J #W1 #U1 #H destruct /2 width=5 by ex3_2_intro/
]
qed-.
-lemma cpy_inv_flat1: ∀I,G,L,V1,T1,U2,d,e. ⦃G, L⦄ ⊢ ⓕ{I} V1. T1 ▶[d, e] U2 →
- ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶[d, e] V2 &
- ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 &
+lemma cpy_inv_flat1: ∀I,G,L,V1,T1,U2,l,m. ⦃G, L⦄ ⊢ ⓕ{I} V1. T1 ▶[l, m] U2 →
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶[l, m] V2 &
+ ⦃G, L⦄ ⊢ T1 ▶[l, m] T2 &
U2 = ⓕ{I}V2.T2.
/2 width=3 by cpy_inv_flat1_aux/ qed-.
-fact cpy_inv_refl_O2_aux: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 → e = 0 → T1 = T2.
-#G #L #T1 #T2 #d #e #H elim H -G -L -T1 -T2 -d -e
+fact cpy_inv_refl_O2_aux: ∀G,L,T1,T2,l,m. ⦃G, L⦄ ⊢ T1 ▶[l, m] T2 → m = 0 → T1 = T2.
+#G #L #T1 #T2 #l #m #H elim H -G -L -T1 -T2 -l -m
[ //
-| #I #G #L #K #V #W #i #d #e #Hdi #Hide #_ #_ #H destruct
- elim (ylt_yle_false … Hdi) -Hdi //
+| #I #G #L #K #V #W #i #l #m #Hli #Hilm #_ #_ #H destruct
+ elim (ylt_yle_false … Hli) -Hli //
| /3 width=1 by eq_f2/
| /3 width=1 by eq_f2/
]
qed-.
-lemma cpy_inv_refl_O2: ∀G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 ▶[d, 0] T2 → T1 = T2.
+lemma cpy_inv_refl_O2: ∀G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 ▶[l, 0] T2 → T1 = T2.
/2 width=6 by cpy_inv_refl_O2_aux/ qed-.
(* Basic_1: was: subst1_gen_lift_eq *)
-lemma cpy_inv_lift1_eq: ∀G,T1,U1,d,e. ⬆[d, e] T1 ≡ U1 →
- ∀L,U2. ⦃G, L⦄ ⊢ U1 ▶[d, e] U2 → U1 = U2.
-#G #T1 #U1 #d #e #HTU1 #L #U2 #HU12 elim (cpy_fwd_up … HU12 … HTU1) -HU12 -HTU1
+lemma cpy_inv_lift1_eq: ∀G,T1,U1,l,m. ⬆[l, m] T1 ≡ U1 →
+ ∀L,U2. ⦃G, L⦄ ⊢ U1 ▶[l, m] U2 → U1 = U2.
+#G #T1 #U1 #l #m #HTU1 #L #U2 #HU12 elim (cpy_fwd_up … HU12 … HTU1) -HU12 -HTU1
/2 width=4 by cpy_inv_refl_O2/
qed-.
(* Main properties **********************************************************)
(* Basic_1: was: subst1_confluence_eq *)
-theorem cpy_conf_eq: ∀G,L,T0,T1,d1,e1. ⦃G, L⦄ ⊢ T0 ▶[d1, e1] T1 →
- ∀T2,d2,e2. ⦃G, L⦄ ⊢ T0 ▶[d2, e2] T2 →
- ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[d2, e2] T & ⦃G, L⦄ ⊢ T2 ▶[d1, e1] T.
-#G #L #T0 #T1 #d1 #e1 #H elim H -G -L -T0 -T1 -d1 -e1
+theorem cpy_conf_eq: ∀G,L,T0,T1,l1,m1. ⦃G, L⦄ ⊢ T0 ▶[l1, m1] T1 →
+ ∀T2,l2,m2. ⦃G, L⦄ ⊢ T0 ▶[l2, m2] T2 →
+ ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[l2, m2] T & ⦃G, L⦄ ⊢ T2 ▶[l1, m1] T.
+#G #L #T0 #T1 #l1 #m1 #H elim H -G -L -T0 -T1 -l1 -m1
[ /2 width=3 by ex2_intro/
-| #I1 #G #L #K1 #V1 #T1 #i0 #d1 #e1 #Hd1 #Hde1 #HLK1 #HVT1 #T2 #d2 #e2 #H
+| #I1 #G #L #K1 #V1 #T1 #i0 #l1 #m1 #Hl1 #Hlm1 #HLK1 #HVT1 #T2 #l2 #m2 #H
elim (cpy_inv_lref1 … H) -H
[ #HX destruct /3 width=7 by cpy_subst, ex2_intro/
- | -Hd1 -Hde1 * #I2 #K2 #V2 #_ #_ #HLK2 #HVT2
+ | -Hl1 -Hlm1 * #I2 #K2 #V2 #_ #_ #HLK2 #HVT2
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
+| #a #I #G #L #V0 #V1 #T0 #T1 #l1 #m1 #_ #_ #IHV01 #IHT01 #X #l2 #m2 #HX
elim (cpy_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct
elim (IHV01 … HV02) -IHV01 -HV02 #V #HV1 #HV2
elim (IHT01 … HT02) -T0 #T #HT1 #HT2
lapply (lsuby_cpy_trans … HT1 (L.ⓑ{I}V1) ?) -HT1 /2 width=1 by lsuby_succ/
lapply (lsuby_cpy_trans … HT2 (L.ⓑ{I}V2) ?) -HT2
/3 width=5 by cpy_bind, lsuby_succ, ex2_intro/
-| #I #G #L #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #X #d2 #e2 #HX
+| #I #G #L #V0 #V1 #T0 #T1 #l1 #m1 #_ #_ #IHV01 #IHT01 #X #l2 #m2 #HX
elim (cpy_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct
elim (IHV01 … HV02) -V0
elim (IHT01 … HT02) -T0 /3 width=5 by cpy_flat, ex2_intro/
qed-.
(* Basic_1: was: subst1_confluence_neq *)
-theorem cpy_conf_neq: ∀G,L1,T0,T1,d1,e1. ⦃G, L1⦄ ⊢ T0 ▶[d1, e1] T1 →
- ∀L2,T2,d2,e2. ⦃G, L2⦄ ⊢ T0 ▶[d2, e2] T2 →
- (d1 + e1 ≤ d2 ∨ d2 + e2 ≤ d1) →
- ∃∃T. ⦃G, L2⦄ ⊢ T1 ▶[d2, e2] T & ⦃G, L1⦄ ⊢ T2 ▶[d1, e1] T.
-#G #L1 #T0 #T1 #d1 #e1 #H elim H -G -L1 -T0 -T1 -d1 -e1
+theorem cpy_conf_neq: ∀G,L1,T0,T1,l1,m1. ⦃G, L1⦄ ⊢ T0 ▶[l1, m1] T1 →
+ ∀L2,T2,l2,m2. ⦃G, L2⦄ ⊢ T0 ▶[l2, m2] T2 →
+ (l1 + m1 ≤ l2 ∨ l2 + m2 ≤ l1) →
+ ∃∃T. ⦃G, L2⦄ ⊢ T1 ▶[l2, m2] T & ⦃G, L1⦄ ⊢ T2 ▶[l1, m1] T.
+#G #L1 #T0 #T1 #l1 #m1 #H elim H -G -L1 -T0 -T1 -l1 -m1
[ /2 width=3 by ex2_intro/
-| #I1 #G #L1 #K1 #V1 #T1 #i0 #d1 #e1 #Hd1 #Hde1 #HLK1 #HVT1 #L2 #T2 #d2 #e2 #H1 #H2
+| #I1 #G #L1 #K1 #V1 #T1 #i0 #l1 #m1 #Hl1 #Hlm1 #HLK1 #HVT1 #L2 #T2 #l2 #m2 #H1 #H2
elim (cpy_inv_lref1 … H1) -H1
[ #H destruct /3 width=7 by cpy_subst, ex2_intro/
- | -HLK1 -HVT1 * #I2 #K2 #V2 #Hd2 #Hde2 #_ #_ elim H2 -H2 #Hded [ -Hd1 -Hde2 | -Hd2 -Hde1 ]
- [ elim (ylt_yle_false … Hde1) -Hde1 /2 width=3 by yle_trans/
- | elim (ylt_yle_false … Hde2) -Hde2 /2 width=3 by yle_trans/
+ | -HLK1 -HVT1 * #I2 #K2 #V2 #Hl2 #Hlm2 #_ #_ elim H2 -H2 #Hlml [ -Hl1 -Hlm2 | -Hl2 -Hlm1 ]
+ [ elim (ylt_yle_false … Hlm1) -Hlm1 /2 width=3 by yle_trans/
+ | elim (ylt_yle_false … Hlm2) -Hlm2 /2 width=3 by yle_trans/
]
]
-| #a #I #G #L1 #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #L2 #X #d2 #e2 #HX #H
+| #a #I #G #L1 #V0 #V1 #T0 #T1 #l1 #m1 #_ #_ #IHV01 #IHT01 #L2 #X #l2 #m2 #HX #H
elim (cpy_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct
elim (IHV01 … HV02 H) -IHV01 -HV02 #V #HV1 #HV2
elim (IHT01 … HT02) -T0
lapply (lsuby_cpy_trans … HT2 (L1.ⓑ{I}V2) ?) -HT2 /3 width=5 by cpy_bind, lsuby_succ, ex2_intro/
| -HV1 -HV2 elim H -H /3 width=1 by yle_succ, or_introl, or_intror/
]
-| #I #G #L1 #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #L2 #X #d2 #e2 #HX #H
+| #I #G #L1 #V0 #V1 #T0 #T1 #l1 #m1 #_ #_ #IHV01 #IHT01 #L2 #X #l2 #m2 #HX #H
elim (cpy_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct
elim (IHV01 … HV02 H) -V0
elim (IHT01 … HT02 H) -T0 -H /3 width=5 by cpy_flat, ex2_intro/
(* Note: the constant 1 comes from cpy_subst *)
(* Basic_1: was: subst1_trans *)
-theorem cpy_trans_ge: ∀G,L,T1,T0,d,e. ⦃G, L⦄ ⊢ T1 ▶[d, e] T0 →
- ∀T2. ⦃G, L⦄ ⊢ T0 ▶[d, 1] T2 → 1 ≤ e → ⦃G, L⦄ ⊢ T1 ▶[d, e] T2.
-#G #L #T1 #T0 #d #e #H elim H -G -L -T1 -T0 -d -e
-[ #I #G #L #d #e #T2 #H #He
+theorem cpy_trans_ge: ∀G,L,T1,T0,l,m. ⦃G, L⦄ ⊢ T1 ▶[l, m] T0 →
+ ∀T2. ⦃G, L⦄ ⊢ T0 ▶[l, 1] T2 → 1 ≤ m → ⦃G, L⦄ ⊢ T1 ▶[l, m] T2.
+#G #L #T1 #T0 #l #m #H elim H -G -L -T1 -T0 -l -m
+[ #I #G #L #l #m #T2 #H #Hm
elim (cpy_inv_atom1 … H) -H
[ #H destruct //
- | * #J #K #V #i #Hd2i #Hide2 #HLK #HVT2 #H destruct
- lapply (ylt_yle_trans … (d+e) … Hide2) /2 width=5 by cpy_subst, monotonic_yle_plus_dx/
+ | * #J #K #V #i #Hl2i #Hilm2 #HLK #HVT2 #H destruct
+ lapply (ylt_yle_trans … (l+m) … Hilm2) /2 width=5 by cpy_subst, monotonic_yle_plus_dx/
]
-| #I #G #L #K #V #V2 #i #d #e #Hdi #Hide #HLK #HVW #T2 #HVT2 #He
+| #I #G #L #K #V #V2 #i #l #m #Hli #Hilm #HLK #HVW #T2 #HVT2 #Hm
lapply (cpy_weak … HVT2 0 (i+1) ? ?) -HVT2 /3 width=1 by yle_plus_dx2_trans, yle_succ/
>yplus_inj #HVT2 <(cpy_inv_lift1_eq … HVW … HVT2) -HVT2 /2 width=5 by cpy_subst/
-| #a #I #G #L #V1 #V0 #T1 #T0 #d #e #_ #_ #IHV10 #IHT10 #X #H #He
+| #a #I #G #L #V1 #V0 #T1 #T0 #l #m #_ #_ #IHV10 #IHT10 #X #H #Hm
elim (cpy_inv_bind1 … H) -H #V2 #T2 #HV02 #HT02 #H destruct
lapply (lsuby_cpy_trans … HT02 (L.ⓑ{I}V1) ?) -HT02 /2 width=1 by lsuby_succ/ #HT02
- lapply (IHT10 … HT02 He) -T0 /3 width=1 by cpy_bind/
-| #I #G #L #V1 #V0 #T1 #T0 #d #e #_ #_ #IHV10 #IHT10 #X #H #He
+ lapply (IHT10 … HT02 Hm) -T0 /3 width=1 by cpy_bind/
+| #I #G #L #V1 #V0 #T1 #T0 #l #m #_ #_ #IHV10 #IHT10 #X #H #Hm
elim (cpy_inv_flat1 … H) -H #V2 #T2 #HV02 #HT02 #H destruct /3 width=1 by cpy_flat/
]
qed-.
-theorem cpy_trans_down: ∀G,L,T1,T0,d1,e1. ⦃G, L⦄ ⊢ T1 ▶[d1, e1] T0 →
- ∀T2,d2,e2. ⦃G, L⦄ ⊢ T0 ▶[d2, e2] T2 → d2 + e2 ≤ d1 →
- ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[d2, e2] T & ⦃G, L⦄ ⊢ T ▶[d1, e1] T2.
-#G #L #T1 #T0 #d1 #e1 #H elim H -G -L -T1 -T0 -d1 -e1
+theorem cpy_trans_down: ∀G,L,T1,T0,l1,m1. ⦃G, L⦄ ⊢ T1 ▶[l1, m1] T0 →
+ ∀T2,l2,m2. ⦃G, L⦄ ⊢ T0 ▶[l2, m2] T2 → l2 + m2 ≤ l1 →
+ ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[l2, m2] T & ⦃G, L⦄ ⊢ T ▶[l1, m1] T2.
+#G #L #T1 #T0 #l1 #m1 #H elim H -G -L -T1 -T0 -l1 -m1
[ /2 width=3 by ex2_intro/
-| #I #G #L #K #V #W #i1 #d1 #e1 #Hdi1 #Hide1 #HLK #HVW #T2 #d2 #e2 #HWT2 #Hde2d1
- lapply (yle_trans … Hde2d1 … Hdi1) -Hde2d1 #Hde2i1
- lapply (cpy_weak … HWT2 0 (i1+1) ? ?) -HWT2 /3 width=1 by yle_succ, yle_pred_sn/ -Hde2i1
+| #I #G #L #K #V #W #i1 #l1 #m1 #Hli1 #Hilm1 #HLK #HVW #T2 #l2 #m2 #HWT2 #Hlm2l1
+ lapply (yle_trans … Hlm2l1 … Hli1) -Hlm2l1 #Hlm2i1
+ lapply (cpy_weak … HWT2 0 (i1+1) ? ?) -HWT2 /3 width=1 by yle_succ, yle_pred_sn/ -Hlm2i1
>yplus_inj #HWT2 <(cpy_inv_lift1_eq … HVW … HWT2) -HWT2 /3 width=9 by cpy_subst, ex2_intro/
-| #a #I #G #L #V1 #V0 #T1 #T0 #d1 #e1 #_ #_ #IHV10 #IHT10 #X #d2 #e2 #HX #de2d1
+| #a #I #G #L #V1 #V0 #T1 #T0 #l1 #m1 #_ #_ #IHV10 #IHT10 #X #l2 #m2 #HX #lm2l1
elim (cpy_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct
lapply (lsuby_cpy_trans … HT02 (L.ⓑ{I}V1) ?) -HT02 /2 width=1 by lsuby_succ/ #HT02
elim (IHV10 … HV02) -IHV10 -HV02 // #V
elim (IHT10 … HT02) -T0 /2 width=1 by yle_succ/ #T #HT1 #HT2
lapply (lsuby_cpy_trans … HT2 (L.ⓑ{I}V) ?) -HT2 /3 width=6 by cpy_bind, lsuby_succ, ex2_intro/
-| #I #G #L #V1 #V0 #T1 #T0 #d1 #e1 #_ #_ #IHV10 #IHT10 #X #d2 #e2 #HX #de2d1
+| #I #G #L #V1 #V0 #T1 #T0 #l1 #m1 #_ #_ #IHV10 #IHT10 #X #l2 #m2 #HX #lm2l1
elim (cpy_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct
elim (IHV10 … HV02) -V0 //
elim (IHT10 … HT02) -T0 /3 width=6 by cpy_flat, ex2_intro/
(* Properties on relocation *************************************************)
(* Basic_1: was: subst1_lift_lt *)
-lemma cpy_lift_le: ∀G,K,T1,T2,dt,et. ⦃G, K⦄ ⊢ T1 ▶[dt, et] T2 →
- ∀L,U1,U2,s,d,e. ⬇[s, d, e] L ≡ K →
- ⬆[d, e] T1 ≡ U1 → ⬆[d, e] T2 ≡ U2 →
- dt + et ≤ d → ⦃G, L⦄ ⊢ U1 ▶[dt, et] U2.
-#G #K #T1 #T2 #dt #et #H elim H -G -K -T1 -T2 -dt -et
-[ #I #G #K #dt #et #L #U1 #U2 #s #d #e #_ #H1 #H2 #_
+lemma cpy_lift_le: ∀G,K,T1,T2,lt,mt. ⦃G, K⦄ ⊢ T1 ▶[lt, mt] T2 →
+ ∀L,U1,U2,s,l,m. ⬇[s, l, m] L ≡ K →
+ ⬆[l, m] T1 ≡ U1 → ⬆[l, m] T2 ≡ U2 →
+ lt + mt ≤ l → ⦃G, L⦄ ⊢ U1 ▶[lt, mt] U2.
+#G #K #T1 #T2 #lt #mt #H elim H -G -K -T1 -T2 -lt -mt
+[ #I #G #K #lt #mt #L #U1 #U2 #s #l #m #_ #H1 #H2 #_
>(lift_mono … H1 … H2) -H1 -H2 //
-| #I #G #K #KV #V #W #i #dt #et #Hdti #Hidet #HKV #HVW #L #U1 #U2 #s #d #e #HLK #H #HWU2 #Hdetd
- lapply (ylt_yle_trans … Hdetd … Hidet) -Hdetd #Hid
- lapply (ylt_inv_inj … Hid) -Hid #Hid
- lapply (lift_inv_lref1_lt … H … Hid) -H #H destruct
+| #I #G #K #KV #V #W #i #lt #mt #Hlti #Hilmt #HKV #HVW #L #U1 #U2 #s #l #m #HLK #H #HWU2 #Hlmtl
+ lapply (ylt_yle_trans … Hlmtl … Hilmt) -Hlmtl #Hil
+ lapply (ylt_inv_inj … Hil) -Hil #Hil
+ lapply (lift_inv_lref1_lt … H … Hil) -H #H destruct
elim (lift_trans_ge … HVW … HWU2) -W // <minus_plus #W #HVW #HWU2
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
+ elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hil #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
+| #a #I #G #K #V1 #V2 #T1 #T2 #lt #mt #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #l #m #HLK #H1 #H2 #Hlmtl
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, 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
+| #G #I #K #V1 #V2 #T1 #T2 #lt #mt #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #l #m #HLK #H1 #H2 #Hlmtl
elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1
elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct
/3 width=7 by cpy_flat/
]
qed-.
-lemma cpy_lift_be: ∀G,K,T1,T2,dt,et. ⦃G, K⦄ ⊢ T1 ▶[dt, et] T2 →
- ∀L,U1,U2,s,d,e. ⬇[s, d, e] L ≡ K →
- ⬆[d, e] T1 ≡ U1 → ⬆[d, e] T2 ≡ U2 →
- dt ≤ d → d ≤ dt + et → ⦃G, L⦄ ⊢ U1 ▶[dt, et + e] U2.
-#G #K #T1 #T2 #dt #et #H elim H -G -K -T1 -T2 -dt -et
-[ #I #G #K #dt #et #L #U1 #U2 #s #d #e #_ #H1 #H2 #_ #_
+lemma cpy_lift_be: ∀G,K,T1,T2,lt,mt. ⦃G, K⦄ ⊢ T1 ▶[lt, mt] T2 →
+ ∀L,U1,U2,s,l,m. ⬇[s, l, m] L ≡ K →
+ ⬆[l, m] T1 ≡ U1 → ⬆[l, m] T2 ≡ U2 →
+ lt ≤ l → l ≤ lt + mt → ⦃G, L⦄ ⊢ U1 ▶[lt, mt + m] U2.
+#G #K #T1 #T2 #lt #mt #H elim H -G -K -T1 -T2 -lt -mt
+[ #I #G #K #lt #mt #L #U1 #U2 #s #l #m #_ #H1 #H2 #_ #_
>(lift_mono … H1 … H2) -H1 -H2 //
-| #I #G #K #KV #V #W #i #dt #et #Hdti #Hidet #HKV #HVW #L #U1 #U2 #s #d #e #HLK #H #HWU2 #Hdtd #_
- elim (lift_inv_lref1 … H) -H * #Hid #H destruct
- [ -Hdtd
- lapply (ylt_yle_trans … (dt+et+e) … Hidet) // -Hidet #Hidete
+| #I #G #K #KV #V #W #i #lt #mt #Hlti #Hilmt #HKV #HVW #L #U1 #U2 #s #l #m #HLK #H #HWU2 #Hltl #_
+ elim (lift_inv_lref1 … H) -H * #Hil #H destruct
+ [ -Hltl
+ lapply (ylt_yle_trans … (lt+mt+m) … Hilmt) // -Hilmt #Hilmtm
elim (lift_trans_ge … HVW … HWU2) -W // <minus_plus #W #HVW #HWU2
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
+ elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hil #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
+ | -Hlti
+ elim (yle_inv_inj2 … Hltl) -Hltl #ltt #Hltl #H destruct
+ lapply (transitive_le … Hltl Hil) -Hltl #Hlti
lapply (lift_trans_be … HVW … HWU2 ? ?) -W /2 width=1 by le_S/ >plus_plus_comm_23 #HVU2
- lapply (drop_trans_ge_comm … HLK … HKV ?) -K // -Hid
+ lapply (drop_trans_ge_comm … HLK … HKV ?) -K // -Hil
/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
+| #a #I #G #K #V1 #V2 #T1 #T2 #lt #mt #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #l #m #HLK #H1 #H2 #Hltl #Hllmt
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, 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
+| #I #G #K #V1 #V2 #T1 #T2 #lt #mt #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #l #m #HLK #H1 #H2 #Hlmtl
elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1
elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct
/3 width=7 by cpy_flat/
qed-.
(* Basic_1: was: subst1_lift_ge *)
-lemma cpy_lift_ge: ∀G,K,T1,T2,dt,et. ⦃G, K⦄ ⊢ T1 ▶[dt, et] T2 →
- ∀L,U1,U2,s,d,e. ⬇[s, d, e] L ≡ K →
- ⬆[d, e] T1 ≡ U1 → ⬆[d, e] T2 ≡ U2 →
- d ≤ dt → ⦃G, L⦄ ⊢ U1 ▶[dt+e, et] U2.
-#G #K #T1 #T2 #dt #et #H elim H -G -K -T1 -T2 -dt -et
-[ #I #G #K #dt #et #L #U1 #U2 #s #d #e #_ #H1 #H2 #_
+lemma cpy_lift_ge: ∀G,K,T1,T2,lt,mt. ⦃G, K⦄ ⊢ T1 ▶[lt, mt] T2 →
+ ∀L,U1,U2,s,l,m. ⬇[s, l, m] L ≡ K →
+ ⬆[l, m] T1 ≡ U1 → ⬆[l, m] T2 ≡ U2 →
+ l ≤ lt → ⦃G, L⦄ ⊢ U1 ▶[lt+m, mt] U2.
+#G #K #T1 #T2 #lt #mt #H elim H -G -K -T1 -T2 -lt -mt
+[ #I #G #K #lt #mt #L #U1 #U2 #s #l #m #_ #H1 #H2 #_
>(lift_mono … H1 … H2) -H1 -H2 //
-| #I #G #K #KV #V #W #i #dt #et #Hdti #Hidet #HKV #HVW #L #U1 #U2 #s #d #e #HLK #H #HWU2 #Hddt
- lapply (yle_trans … Hddt … Hdti) -Hddt #Hid
- elim (yle_inv_inj2 … Hid) -Hid #dd #Hddi #H0 destruct
- lapply (lift_inv_lref1_ge … H … Hddi) -H #H destruct
+| #I #G #K #KV #V #W #i #lt #mt #Hlti #Hilmt #HKV #HVW #L #U1 #U2 #s #l #m #HLK #H #HWU2 #Hllt
+ lapply (yle_trans … Hllt … Hlti) -Hllt #Hil
+ elim (yle_inv_inj2 … Hil) -Hil #ll #Hlli #H0 destruct
+ lapply (lift_inv_lref1_ge … H … Hlli) -H #H destruct
lapply (lift_trans_be … HVW … HWU2 ? ?) -W /2 width=1 by le_S/ >plus_plus_comm_23 #HVU2
- lapply (drop_trans_ge_comm … HLK … HKV ?) -K // -Hddi
+ lapply (drop_trans_ge_comm … HLK … HKV ?) -K // -Hlli
/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
+| #a #I #G #K #V1 #V2 #T1 #T2 #lt #mt #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #l #m #HLK #H1 #H2 #Hllt
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, 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
+| #I #G #K #V1 #V2 #T1 #T2 #lt #mt #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #l #m #HLK #H1 #H2 #Hllt
elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1
elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct
/3 width=6 by cpy_flat/
(* Inversion lemmas on relocation *******************************************)
(* Basic_1: was: subst1_gen_lift_lt *)
-lemma cpy_inv_lift1_le: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶[dt, et] U2 →
- ∀K,s,d,e. ⬇[s, d, e] L ≡ K → ∀T1. ⬆[d, e] T1 ≡ U1 →
- dt + et ≤ d →
- ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[dt, et] T2 & ⬆[d, e] T2 ≡ U2.
-#G #L #U1 #U2 #dt #et #H elim H -G -L -U1 -U2 -dt -et
-[ * #i #G #L #dt #et #K #s #d #e #_ #T1 #H #_
+lemma cpy_inv_lift1_le: ∀G,L,U1,U2,lt,mt. ⦃G, L⦄ ⊢ U1 ▶[lt, mt] U2 →
+ ∀K,s,l,m. ⬇[s, l, m] L ≡ K → ∀T1. ⬆[l, m] T1 ≡ U1 →
+ lt + mt ≤ l →
+ ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[lt, mt] T2 & ⬆[l, m] T2 ≡ U2.
+#G #L #U1 #U2 #lt #mt #H elim H -G -L -U1 -U2 -lt -mt
+[ * #i #G #L #lt #mt #K #s #l #m #_ #T1 #H #_
[ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3 by ex2_intro/
- | elim (lift_inv_lref2 … H) -H * #Hid #H destruct /3 width=3 by lift_lref_ge_minus, lift_lref_lt, ex2_intro/
+ | elim (lift_inv_lref2 … H) -H * #Hil #H destruct /3 width=3 by lift_lref_ge_minus, lift_lref_lt, ex2_intro/
| lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3 by ex2_intro/
]
-| #I #G #L #KV #V #W #i #dt #et #Hdti #Hidet #HLKV #HVW #K #s #d #e #HLK #T1 #H #Hdetd
- 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
+| #I #G #L #KV #V #W #i #lt #mt #Hlti #Hilmt #HLKV #HVW #K #s #l #m #HLK #T1 #H #Hlmtl
+ lapply (ylt_yle_trans … Hlmtl … Hilmt) -Hlmtl #Hil
+ lapply (ylt_inv_inj … Hil) -Hil #Hil
+ lapply (lift_inv_lref2_lt … H … Hil) -H #H destruct
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_trans_le … HUV … HVW) -V // >minus_plus <plus_minus_m_m // -Hil /3 width=5 by cpy_subst, ex2_intro/
+| #a #I #G #L #W1 #W2 #U1 #U2 #lt #mt #_ #_ #IHW12 #IHU12 #K #s #l #m #HLK #X #H #Hlmtl
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, 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
+| #I #G #L #W1 #W2 #U1 #U2 #lt #mt #_ #_ #IHW12 #IHU12 #K #s #l #m #HLK #X #H #Hlmtl
elim (lift_inv_flat2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct
elim (IHW12 … HLK … HVW1) -W1 //
elim (IHU12 … HLK … HTU1) -U1 -HLK
]
qed-.
-lemma cpy_inv_lift1_be: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶[dt, et] U2 →
- ∀K,s,d,e. ⬇[s, d, e] L ≡ K → ∀T1. ⬆[d, e] T1 ≡ U1 →
- dt ≤ d → yinj d + e ≤ dt + et →
- ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[dt, et-e] T2 & ⬆[d, e] T2 ≡ U2.
-#G #L #U1 #U2 #dt #et #H elim H -G -L -U1 -U2 -dt -et
-[ * #i #G #L #dt #et #K #s #d #e #_ #T1 #H #_ #_
+lemma cpy_inv_lift1_be: ∀G,L,U1,U2,lt,mt. ⦃G, L⦄ ⊢ U1 ▶[lt, mt] U2 →
+ ∀K,s,l,m. ⬇[s, l, m] L ≡ K → ∀T1. ⬆[l, m] T1 ≡ U1 →
+ lt ≤ l → yinj l + m ≤ lt + mt →
+ ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[lt, mt-m] T2 & ⬆[l, m] T2 ≡ U2.
+#G #L #U1 #U2 #lt #mt #H elim H -G -L -U1 -U2 -lt -mt
+[ * #i #G #L #lt #mt #K #s #l #m #_ #T1 #H #_ #_
[ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3 by ex2_intro/
- | elim (lift_inv_lref2 … H) -H * #Hid #H destruct /3 width=3 by lift_lref_ge_minus, lift_lref_lt, ex2_intro/
+ | elim (lift_inv_lref2 … H) -H * #Hil #H destruct /3 width=3 by lift_lref_ge_minus, lift_lref_lt, ex2_intro/
| lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3 by ex2_intro/
]
-| #I #G #L #KV #V #W #i #dt #et #Hdti #Hidet #HLKV #HVW #K #s #d #e #HLK #T1 #H #Hdtd #Hdedet
- lapply (yle_fwd_plus_ge_inj … Hdtd Hdedet) #Heet
- 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
+| #I #G #L #KV #V #W #i #lt #mt #Hlti #Hilmt #HLKV #HVW #K #s #l #m #HLK #T1 #H #Hltl #Hlmlmt
+ lapply (yle_fwd_plus_ge_inj … Hltl Hlmlmt) #Hmmt
+ elim (lift_inv_lref2 … H) -H * #Hil #H destruct [ -Hltl -Hilmt | -Hlti -Hlmlmt ]
+ [ lapply (ylt_yle_trans i l (lt+(mt-m)) ? ?) /2 width=1 by ylt_inj/
+ [ >yplus_minus_assoc_inj /2 width=1 by yle_plus1_to_minus_inj2/ ] -Hlmlmt #Hilmtm
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
+ elim (lift_trans_le … HUV … HVW) -V // >minus_plus <plus_minus_m_m // -Hil
/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
+ | elim (le_inv_plus_l … Hil) #Hlim #Hmi
+ lapply (yle_trans … Hltl (i-m) ?) /2 width=1 by yle_inj/ -Hltl #Hltim
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
+ elim (lift_split … HVW l (i-m+1)) -HVW [2,3,4: /2 width=1 by le_S_S, le_S/ ] -Hil -Hlim
#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 *)
>yplus_minus_assoc_inj /3 width=1 by monotonic_ylt_minus_dx, yle_inj/
]
-| #a #I #G #L #W1 #W2 #U1 #U2 #dt #et #_ #_ #IHW12 #IHU12 #K #s #d #e #HLK #X #H #Hdtd #Hdedet
+| #a #I #G #L #W1 #W2 #U1 #U2 #lt #mt #_ #_ #IHW12 #IHU12 #K #s #l #m #HLK #X #H #Hltl #Hlmlmt
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, 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
+| #I #G #L #W1 #W2 #U1 #U2 #lt #mt #_ #_ #IHW12 #IHU12 #K #s #l #m #HLK #X #H #Hltl #Hlmlmt
elim (lift_inv_flat2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct
elim (IHW12 … HLK … HVW1) -W1 //
elim (IHU12 … HLK … HTU1) -U1 -HLK //
qed-.
(* Basic_1: was: subst1_gen_lift_ge *)
-lemma cpy_inv_lift1_ge: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶[dt, et] U2 →
- ∀K,s,d,e. ⬇[s, d, e] L ≡ K → ∀T1. ⬆[d, e] T1 ≡ U1 →
- yinj d + e ≤ dt →
- ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[dt-e, et] T2 & ⬆[d, e] T2 ≡ U2.
-#G #L #U1 #U2 #dt #et #H elim H -G -L -U1 -U2 -dt -et
-[ * #i #G #L #dt #et #K #s #d #e #_ #T1 #H #_
+lemma cpy_inv_lift1_ge: ∀G,L,U1,U2,lt,mt. ⦃G, L⦄ ⊢ U1 ▶[lt, mt] U2 →
+ ∀K,s,l,m. ⬇[s, l, m] L ≡ K → ∀T1. ⬆[l, m] T1 ≡ U1 →
+ yinj l + m ≤ lt →
+ ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[lt-m, mt] T2 & ⬆[l, m] T2 ≡ U2.
+#G #L #U1 #U2 #lt #mt #H elim H -G -L -U1 -U2 -lt -mt
+[ * #i #G #L #lt #mt #K #s #l #m #_ #T1 #H #_
[ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3 by ex2_intro/
- | elim (lift_inv_lref2 … H) -H * #Hid #H destruct /3 width=3 by lift_lref_ge_minus, lift_lref_lt, ex2_intro/
+ | elim (lift_inv_lref2 … H) -H * #Hil #H destruct /3 width=3 by lift_lref_ge_minus, lift_lref_lt, ex2_intro/
| lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3 by ex2_intro/
]
-| #I #G #L #KV #V #W #i #dt #et #Hdti #Hidet #HLKV #HVW #K #s #d #e #HLK #T1 #H #Hdedt
- lapply (yle_trans … Hdedt … Hdti) #Hdei
- elim (yle_inv_plus_inj2 … Hdedt) -Hdedt #_ #Hedt
- elim (yle_inv_plus_inj2 … Hdei) #Hdie #Hei
+| #I #G #L #KV #V #W #i #lt #mt #Hlti #Hilmt #HLKV #HVW #K #s #l #m #HLK #T1 #H #Hlmlt
+ lapply (yle_trans … Hlmlt … Hlti) #Hlmi
+ elim (yle_inv_plus_inj2 … Hlmlt) -Hlmlt #_ #Hmlt
+ elim (yle_inv_plus_inj2 … Hlmi) #Hlim #Hmi
lapply (lift_inv_lref2_ge … H ?) -H /2 width=1 by yle_inv_inj/ #H destruct
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
+ elim (lift_split … HVW l (i-m+1)) -HVW [2,3,4: /3 width=1 by yle_inv_inj, le_S_S, le_S/ ] -Hlmi -Hlim
#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 *)
[ /2 width=1 by monotonic_yle_minus_dx/
| <yplus_minus_comm_inj /2 width=1 by monotonic_ylt_minus_dx/
]
-| #a #I #G #L #W1 #W2 #U1 #U2 #dt #et #_ #_ #IHW12 #IHU12 #K #s #d #e #HLK #X #H #Hdetd
+| #a #I #G #L #W1 #W2 #U1 #U2 #lt #mt #_ #_ #IHW12 #IHU12 #K #s #l #m #HLK #X #H #Hlmtl
elim (lift_inv_bind2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct
- elim (yle_inv_plus_inj2 … Hdetd) #_ #Hedt
+ elim (yle_inv_plus_inj2 … Hlmtl) #_ #Hmlt
elim (IHW12 … HLK … HVW1) -IHW12 // #V2 #HV12 #HVW2
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
+| #I #G #L #W1 #W2 #U1 #U2 #lt #mt #_ #_ #IHW12 #IHU12 #K #s #l #m #HLK #X #H #Hlmtl
elim (lift_inv_flat2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct
elim (IHW12 … HLK … HVW1) -W1 //
elim (IHU12 … HLK … HTU1) -U1 -HLK /3 width=5 by cpy_flat, lift_flat, ex2_intro/
(* Advanced inversion lemmas on relocation ***********************************)
-lemma cpy_inv_lift1_ge_up: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶[dt, et] U2 →
- ∀K,s,d,e. ⬇[s, d, e] L ≡ K → ∀T1. ⬆[d, e] T1 ≡ U1 →
- d ≤ dt → dt ≤ yinj d + e → yinj d + e ≤ dt + et →
- ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[d, dt + et - (yinj d + e)] T2 & ⬆[d, e] T2 ≡ U2.
-#G #L #U1 #U2 #dt #et #HU12 #K #s #d #e #HLK #T1 #HTU1 #Hddt #Hdtde #Hdedet
-elim (cpy_split_up … HU12 (d + e)) -HU12 // -Hdedet #U #HU1 #HU2
-lapply (cpy_weak … HU1 d e ? ?) -HU1 // [ >ymax_pre_sn_comm // ] -Hddt -Hdtde #HU1
+lemma cpy_inv_lift1_ge_up: ∀G,L,U1,U2,lt,mt. ⦃G, L⦄ ⊢ U1 ▶[lt, mt] U2 →
+ ∀K,s,l,m. ⬇[s, l, m] L ≡ K → ∀T1. ⬆[l, m] T1 ≡ U1 →
+ l ≤ lt → lt ≤ yinj l + m → yinj l + m ≤ lt + mt →
+ ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[l, lt + mt - (yinj l + m)] T2 & ⬆[l, m] T2 ≡ U2.
+#G #L #U1 #U2 #lt #mt #HU12 #K #s #l #m #HLK #T1 #HTU1 #Hllt #Hltlm #Hlmlmt
+elim (cpy_split_up … HU12 (l + m)) -HU12 // -Hlmlmt #U #HU1 #HU2
+lapply (cpy_weak … HU1 l m ? ?) -HU1 // [ >ymax_pre_sn_comm // ] -Hllt -Hltlm #HU1
lapply (cpy_inv_lift1_eq … HTU1 … HU1) -HU1 #HU1 destruct
elim (cpy_inv_lift1_ge … HU2 … HLK … HTU1) -U -L /2 width=3 by ex2_intro/
qed-.
-lemma cpy_inv_lift1_be_up: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶[dt, et] U2 →
- ∀K,s,d,e. ⬇[s, d, e] L ≡ K → ∀T1. ⬆[d, e] T1 ≡ U1 →
- dt ≤ d → dt + et ≤ yinj d + e →
- ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[dt, d-dt] T2 & ⬆[d, e] T2 ≡ U2.
-#G #L #U1 #U2 #dt #et #HU12 #K #s #d #e #HLK #T1 #HTU1 #Hdtd #Hdetde
-lapply (cpy_weak … HU12 dt (d+e-dt) ? ?) -HU12 //
-[ >ymax_pre_sn_comm /2 width=1 by yle_plus_dx1_trans/ ] -Hdetde #HU12
+lemma cpy_inv_lift1_be_up: ∀G,L,U1,U2,lt,mt. ⦃G, L⦄ ⊢ U1 ▶[lt, mt] U2 →
+ ∀K,s,l,m. ⬇[s, l, m] L ≡ K → ∀T1. ⬆[l, m] T1 ≡ U1 →
+ lt ≤ l → lt + mt ≤ yinj l + m →
+ ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[lt, l-lt] T2 & ⬆[l, m] T2 ≡ U2.
+#G #L #U1 #U2 #lt #mt #HU12 #K #s #l #m #HLK #T1 #HTU1 #Hltl #Hlmtlm
+lapply (cpy_weak … HU12 lt (l+m-lt) ? ?) -HU12 //
+[ >ymax_pre_sn_comm /2 width=1 by yle_plus_dx1_trans/ ] -Hlmtlm #HU12
elim (cpy_inv_lift1_be … HU12 … HLK … HTU1) -U1 -L /2 width=3 by ex2_intro/
qed-.
-lemma cpy_inv_lift1_le_up: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶[dt, et] U2 →
- ∀K,s,d,e. ⬇[s, d, e] L ≡ K → ∀T1. ⬆[d, e] T1 ≡ U1 →
- dt ≤ d → d ≤ dt + et → dt + et ≤ yinj d + e →
- ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[dt, d - dt] T2 & ⬆[d, e] T2 ≡ U2.
-#G #L #U1 #U2 #dt #et #HU12 #K #s #d #e #HLK #T1 #HTU1 #Hdtd #Hddet #Hdetde
-elim (cpy_split_up … HU12 d) -HU12 // #U #HU1 #HU2
+lemma cpy_inv_lift1_le_up: ∀G,L,U1,U2,lt,mt. ⦃G, L⦄ ⊢ U1 ▶[lt, mt] U2 →
+ ∀K,s,l,m. ⬇[s, l, m] L ≡ K → ∀T1. ⬆[l, m] T1 ≡ U1 →
+ lt ≤ l → l ≤ lt + mt → lt + mt ≤ yinj l + m →
+ ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[lt, l - lt] T2 & ⬆[l, m] T2 ≡ U2.
+#G #L #U1 #U2 #lt #mt #HU12 #K #s #l #m #HLK #T1 #HTU1 #Hltl #Hllmt #Hlmtlm
+elim (cpy_split_up … HU12 l) -HU12 // #U #HU1 #HU2
elim (cpy_inv_lift1_le … HU1 … HLK … HTU1) -U1
-[2: >ymax_pre_sn_comm // ] -Hdtd #T #HT1 #HTU
-lapply (cpy_weak … HU2 d e ? ?) -HU2 //
-[ >ymax_pre_sn_comm // ] -Hddet -Hdetde #HU2
+[2: >ymax_pre_sn_comm // ] -Hltl #T #HT1 #HTU
+lapply (cpy_weak … HU2 l m ? ?) -HU2 //
+[ >ymax_pre_sn_comm // ] -Hllmt -Hlmtlm #HU2
lapply (cpy_inv_lift1_eq … HTU … HU2) -L #H destruct /2 width=3 by ex2_intro/
qed-.
(* Inversion lemmas on negated relocation ***********************************)
-lemma cpy_fwd_nlift2_ge: ∀G,L,U1,U2,d,e. ⦃G, L⦄ ⊢ U1 ▶[d, e] U2 →
- ∀i. d ≤ yinj i → (∀T2. ⬆[i, 1] T2 ≡ U2 → ⊥) →
+lemma cpy_fwd_nlift2_ge: ∀G,L,U1,U2,l,m. ⦃G, L⦄ ⊢ U1 ▶[l, m] U2 →
+ ∀i. l ≤ yinj i → (∀T2. ⬆[i, 1] T2 ≡ U2 → ⊥) →
(∀T1. ⬆[i, 1] T1 ≡ U1 → ⊥) ∨
- ∃∃I,K,W,j. d ≤ yinj j & j < i & ⬇[j]L ≡ K.ⓑ{I}W &
+ ∃∃I,K,W,j. l ≤ yinj j & j < i & ⬇[j]L ≡ K.ⓑ{I}W &
(∀V. ⬆[i-j-1, 1] V ≡ W → ⊥) & (∀T1. ⬆[j, 1] T1 ≡ U1 → ⊥).
-#G #L #U1 #U2 #d #e #H elim H -G -L -U1 -U2 -d -e
+#G #L #U1 #U2 #l #m #H elim H -G -L -U1 -U2 -l -m
[ /3 width=2 by or_introl/
-| #I #G #L #K #V #W #j #d #e #Hdj #Hjde #HLK #HVW #i #Hdi #HnW
+| #I #G #L #K #V #W #j #l #m #Hlj #Hjlm #HLK #HVW #i #Hli #HnW
elim (lt_or_ge j i) #Hij
[ @or_intror @(ex5_4_intro … HLK) // -HLK
[ #X #HXV elim (lift_trans_le … HXV … HVW ?) -V //
| elim (lift_split … HVW i j) -HVW //
#X #_ #H elim HnW -HnW //
]
-| #a #I #G #L #W1 #W2 #U1 #U2 #d #e #_ #_ #IHW12 #IHU12 #i #Hdi #H elim (nlift_inv_bind … H) -H
+| #a #I #G #L #W1 #W2 #U1 #U2 #l #m #_ #_ #IHW12 #IHU12 #i #Hli #H elim (nlift_inv_bind … H) -H
[ #HnW2 elim (IHW12 … HnW2) -IHW12 -HnW2 -IHU12 //
[ /4 width=9 by nlift_bind_sn, or_introl/
| * /5 width=9 by nlift_bind_sn, ex5_4_intro, or_intror/
]
| #HnU2 elim (IHU12 … HnU2) -IHU12 -HnU2 -IHW12 /2 width=1 by yle_succ/
[ /4 width=9 by nlift_bind_dx, or_introl/
- | * #J #K #W #j #Hdj #Hji #HLK #HnW
- elim (yle_inv_succ1 … Hdj) -Hdj #Hdj #Hj
+ | * #J #K #W #j #Hlj #Hji #HLK #HnW
+ elim (yle_inv_succ1 … Hlj) -Hlj #Hlj #Hj
lapply (ylt_O … Hj) -Hj #Hj
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/ ]
/5 width=9 by nlift_bind_dx, monotonic_lt_pred, lt_plus_to_minus_r, ex5_4_intro, or_intror/
]
]
-| #I #G #L #W1 #W2 #U1 #U2 #d #e #_ #_ #IHW12 #IHU12 #i #Hdi #H elim (nlift_inv_flat … H) -H
+| #I #G #L #W1 #W2 #U1 #U2 #l #m #_ #_ #IHW12 #IHU12 #i #Hli #H elim (nlift_inv_flat … H) -H
[ #HnW2 elim (IHW12 … HnW2) -IHW12 -HnW2 -IHU12 //
[ /4 width=9 by nlift_flat_sn, or_introl/
| * /5 width=9 by nlift_flat_sn, ex5_4_intro, or_intror/
(* 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_atom: ∀l,m. (s = Ⓕ → m = 0) → drop s l m (⋆) (⋆)
| 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)
+| drop_drop: ∀I,L1,L2,V,m. drop s 0 m L1 L2 → drop s 0 (m+1) (L1.ⓑ{I}V) L2
+| drop_skip: ∀I,L1,L2,V1,V2,l,m.
+ drop s l m L1 L2 → ⬆[l, m] V2 ≡ V1 →
+ drop s (l+1) m (L1.ⓑ{I}V1) (L2.ⓑ{I}V2)
.
interpretation
"basic slicing (local environment) abstract"
- 'RDrop s d e L1 L2 = (drop s d e L1 L2).
+ 'RDrop s l m L1 L2 = (drop s l m L1 L2).
(*
interpretation
"basic slicing (local environment) general"
*)
interpretation
"basic slicing (local environment) lget"
- 'RDrop e L1 L2 = (drop false O e L1 L2).
+ 'RDrop m L1 L2 = (drop false O m 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 d_liftable: predicate (lenv → relation term) ≝
+ λR. ∀K,T1,T2. R K T1 T2 → ∀L,s,l,m. ⬇[s, l, m] L ≡ K →
+ ∀U1. ⬆[l, m] T1 ≡ U1 → ∀U2. ⬆[l, m] 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 d_deliftable_sn: predicate (lenv → relation term) ≝
+ λR. ∀L,U1,U2. R L U1 U2 → ∀K,s,l,m. ⬇[s, l, m] L ≡ K →
+ ∀T1. ⬆[l, m] T1 ≡ U1 →
+ ∃∃T2. ⬆[l, m] 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.
+ λR. ∀L1,K1,s,l,m. ⬇[s, l, m] L1 ≡ K1 → ∀L2. R L1 L2 →
+ ∃∃K2. R K1 K2 & ⬇[s, l, m] 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.
+ λR. ∀L1,L2. R L1 L2 → ∀K2,s,m. ⬇[s, 0, m] L2 ≡ K2 →
+ ∃∃K1. ⬇[s, 0, m] 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
+fact drop_inv_atom1_aux: ∀L1,L2,s,l,m. ⬇[s, l, m] L1 ≡ L2 → L1 = ⋆ →
+ L2 = ⋆ ∧ (s = Ⓕ → m = 0).
+#L1 #L2 #s #l #m * -L1 -L2 -l -m
[ /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
+| #I #L1 #L2 #V #m #_ #H destruct
+| #I #L1 #L2 #V1 #V2 #l #m #_ #_ #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).
+lemma drop_inv_atom1: ∀L2,s,l,m. ⬇[s, l, m] ⋆ ≡ L2 → L2 = ⋆ ∧ (s = Ⓕ → m = 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 →
+fact drop_inv_O1_pair1_aux: ∀L1,L2,s,l,m. ⬇[s, l, m] L1 ≡ L2 → l = 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
+ (m = 0 ∧ L2 = K.ⓑ{I}V) ∨
+ (0 < m ∧ ⬇[s, l, m-1] K ≡ L2).
+#L1 #L2 #s #l #m * -L1 -L2 -l -m
+[ #l #m #_ #_ #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
+| #I #L1 #L2 #V #m #HL12 #_ #K #J #W #H destruct /3 width=1 by or_intror, conj/
+| #I #L1 #L2 #V1 #V2 #l #m #_ #_ >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).
+lemma drop_inv_O1_pair1: ∀I,K,L2,V,s,m. ⬇[s, 0, m] K. ⓑ{I} V ≡ L2 →
+ (m = 0 ∧ L2 = K.ⓑ{I}V) ∨
+ (0 < m ∧ ⬇[s, 0, m-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.
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
+lemma drop_inv_drop1_lt: ∀I,K,L2,V,s,m.
+ ⬇[s, 0, m] K.ⓑ{I}V ≡ L2 → 0 < m → ⬇[s, 0, m-1] K ≡ L2.
+#I #K #L2 #V #s #m #H #Hm
elim (drop_inv_O1_pair1 … H) -H * // #H destruct
-elim (lt_refl_false … He)
+elim (lt_refl_false … Hm)
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 //
+lemma drop_inv_drop1: ∀I,K,L2,V,s,m.
+ ⬇[s, 0, m+1] K.ⓑ{I}V ≡ L2 → ⬇[s, 0, m] K ≡ L2.
+#I #K #L2 #V #s #m #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 →
+fact drop_inv_skip1_aux: ∀L1,L2,s,l,m. ⬇[s, l, m] L1 ≡ L2 → 0 < l →
∀I,K1,V1. L1 = K1.ⓑ{I}V1 →
- ∃∃K2,V2. ⬇[s, d-1, e] K1 ≡ K2 &
- ⬆[d-1, e] V2 ≡ V1 &
+ ∃∃K2,V2. ⬇[s, l-1, m] K1 ≡ K2 &
+ ⬆[l-1, m] V2 ≡ V1 &
L2 = K2.ⓑ{I}V2.
-#L1 #L2 #s #d #e * -L1 -L2 -d -e
-[ #d #e #_ #_ #J #K1 #W1 #H destruct
+#L1 #L2 #s #l #m * -L1 -L2 -l -m
+[ #l #m #_ #_ #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/
+| #I #L1 #L2 #V #m #_ #H elim (lt_refl_false … H)
+| #I #L1 #L2 #V1 #V2 #l #m #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 &
+lemma drop_inv_skip1: ∀I,K1,V1,L2,s,l,m. ⬇[s, l, m] K1.ⓑ{I}V1 ≡ L2 → 0 < l →
+ ∃∃K2,V2. ⬇[s, l-1, m] K1 ≡ K2 &
+ ⬆[l-1, m] 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 *
+lemma drop_inv_O1_pair2: ∀I,K,V,s,m,L1. ⬇[s, 0, m] L1 ≡ K.ⓑ{I}V →
+ (m = 0 ∧ L1 = K.ⓑ{I}V) ∨
+ ∃∃I1,K1,V1. ⬇[s, 0, m-1] K1 ≡ K.ⓑ{I}V & L1 = K1.ⓑ{I1}V1 & 0 < m.
+#I #K #V #s #m *
[ #H elim (drop_inv_atom1 … H) -H #H destruct
| #L1 #I1 #V1 #H
elim (drop_inv_O1_pair1 … H) -H *
]
qed-.
-fact drop_inv_skip2_aux: ∀L1,L2,s,d,e. ⬇[s, d, e] L1 ≡ L2 → 0 < d →
+fact drop_inv_skip2_aux: ∀L1,L2,s,l,m. ⬇[s, l, m] L1 ≡ L2 → 0 < l →
∀I,K2,V2. L2 = K2.ⓑ{I}V2 →
- ∃∃K1,V1. ⬇[s, d-1, e] K1 ≡ K2 &
- ⬆[d-1, e] V2 ≡ V1 &
+ ∃∃K1,V1. ⬇[s, l-1, m] K1 ≡ K2 &
+ ⬆[l-1, m] V2 ≡ V1 &
L1 = K1.ⓑ{I}V1.
-#L1 #L2 #s #d #e * -L1 -L2 -d -e
-[ #d #e #_ #_ #J #K2 #W2 #H destruct
+#L1 #L2 #s #l #m * -L1 -L2 -l -m
+[ #l #m #_ #_ #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/
+| #I #L1 #L2 #V #m #_ #H elim (lt_refl_false … H)
+| #I #L1 #L2 #V1 #V2 #l #m #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 &
+lemma drop_inv_skip2: ∀I,L1,K2,V2,s,l,m. ⬇[s, l, m] L1 ≡ K2.ⓑ{I}V2 → 0 < l →
+ ∃∃K1,V1. ⬇[s, l-1, m] K1 ≡ K2 & ⬆[l-1, m] 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 →
+lemma drop_inv_O1_gt: ∀L,K,m,s. ⬇[s, 0, m] L ≡ K → |L| < m →
s = Ⓣ ∧ K = ⋆.
-#L elim L -L [| #L #Z #X #IHL ] #K #e #s #H normalize in ⊢ (?%?→?); #H1e
+#L elim L -L [| #L #Z #X #IHL ] #K #m #s #H normalize in ⊢ (?%?→?); #H1m
[ 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)
+ #_ #Hs lapply (Hs ?) // -Hs #H destruct elim (lt_zero_false … H1m)
+| elim (drop_inv_O1_pair1 … H) -H * #H2m #HLK destruct
+ [ elim (lt_zero_false … H1m)
| elim (IHL … HLK) -IHL -HLK /2 width=1 by lt_plus_to_minus_r, conj/
]
]
(* Basic properties *********************************************************)
-lemma drop_refl_atom_O2: ∀s,d. ⬇[s, d, O] ⋆ ≡ ⋆.
+lemma drop_refl_atom_O2: ∀s,l. ⬇[s, l, 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.
+lemma drop_refl: ∀L,l,s. ⬇[s, l, 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/
+#L #I #V #IHL #l #s @(nat_ind_plus … l) -l /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/
+lemma drop_drop_lt: ∀I,L1,L2,V,s,m.
+ ⬇[s, 0, m-1] L1 ≡ L2 → 0 < m → ⬇[s, 0, m] L1.ⓑ{I}V ≡ L2.
+#I #L1 #L2 #V #s #m #HL12 #Hm >(plus_minus_m_m m 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/
+lemma drop_skip_lt: ∀I,L1,L2,V1,V2,s,l,m.
+ ⬇[s, l-1, m] L1 ≡ L2 → ⬆[l-1, m] V2 ≡ V1 → 0 < l →
+ ⬇[s, l, m] L1. ⓑ{I} V1 ≡ L2.ⓑ{I}V2.
+#I #L1 #L2 #V1 #V2 #s #l #m #HL12 #HV21 #Hl >(plus_minus_m_m l 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 *
+lemma drop_O1_le: ∀s,m,L. m ≤ |L| → ∃K. ⬇[s, 0, m] L ≡ K.
+#s #m @(nat_ind_plus … m) -m /2 width=2 by ex_intro/
+#m #IHm *
[ #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/
+| #L #I #V normalize #H elim (IHm L) -IHm /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.
+lemma drop_O1_lt: ∀s,L,m. m < |L| → ∃∃I,K,V. ⬇[s, 0, m] 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/
+[ #m #H elim (lt_zero_false … H)
+| #L #I #V #IHL #m @(nat_ind_plus … m) -m /2 width=4 by drop_pair, ex1_3_intro/
+ #m #_ normalize #H elim (IHL m) -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
+lemma drop_O1_pair: ∀L,K,m,s. ⬇[s, 0, m] L ≡ K → m ≤ |L| → ∀I,V.
+ ∃∃J,W. ⬇[s, 0, m] L.ⓑ{I}V ≡ K.ⓑ{J}W.
+#L elim L -L [| #L #Z #X #IHL ] #K #m #s #H normalize #Hm #I #V
+[ elim (drop_inv_atom1 … H) -H #H <(le_n_O_to_eq … Hm) -m
#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 (drop_inv_O1_pair1 … H) -H * #Hm #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) ]
+lemma drop_O1_ge: ∀L,m. |L| ≤ m → ⬇[Ⓣ, 0, m] L ≡ ⋆.
+#L elim L -L [ #m #_ @drop_atom #H destruct ]
+#L #I #V #IHL #m @(nat_ind_plus … m) -m [ #H elim (le_plus_xSy_O_false … H) ]
normalize /4 width=1 by drop_drop, monotonic_pred/
qed.
#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 … (⋆))
+lemma drop_split: ∀L1,L2,l,m2,s. ⬇[s, l, m2] L1 ≡ L2 → ∀m1. m1 ≤ m2 →
+ ∃∃L. ⬇[s, l, m2 - m1] L1 ≡ L & ⬇[s, l, m1] L ≡ L2.
+#L1 #L2 #l #m2 #s #H elim H -L1 -L2 -l -m2
+[ #l #m2 #Hs #m1 #Hm12 @(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
+| #I #L1 #V #m1 #Hm1 lapply (le_n_O_to_eq … Hm1) -Hm1
#H destruct /2 width=3 by ex2_intro/
-| #I #L1 #L2 #V #e2 #HL12 #IHL12 #e1 @(nat_ind_plus … e1) -e1
+| #I #L1 #L2 #V #m2 #HL12 #IHL12 #m1 @(nat_ind_plus … m1) -m1
[ /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
+ | -HL12 #m1 #_ #Hm12 lapply (le_plus_to_le_r … Hm12) -Hm12
+ #Hm12 elim (IHL12 … Hm12) -IHL12 >minus_plus_plus_l
+ #L #HL1 #HL2 elim (lt_or_ge (|L1|) (m2-m1)) #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 // ]
| 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/
+| #I #L1 #L2 #V1 #V2 #l #m2 #_ #HV21 #IHL12 #m1 #Hm12 elim (IHL12 … Hm12) -IHL12
+ #L #HL1 #HL2 elim (lift_split … HV21 l m1) -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
+lemma drop_FT: ∀L1,L2,l,m. ⬇[Ⓕ, l, m] L1 ≡ L2 → ⬇[Ⓣ, l, m] L1 ≡ L2.
+#L1 #L2 #l #m #H elim H -L1 -L2 -l -m
/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.
+lemma drop_gen: ∀L1,L2,s,l,m. ⬇[Ⓕ, l, m] L1 ≡ L2 → ⬇[s, l, m] 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.
+lemma drop_T: ∀L1,L2,s,l,m. ⬇[s, l, m] L1 ≡ L2 → ⬇[Ⓣ, l, m] L1 ≡ L2.
#L1 #L2 * /2 width=1 by drop_FT/
qed-.
-lemma l_liftable_LTC: ∀R. l_liftable R → l_liftable (LTC … R).
+lemma d_liftable_LTC: ∀R. d_liftable R → d_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/
+| #T #T2 #_ #HT2 #IHT1 #L #s #l #m #HLK #U1 #HTU1 #U2 #HTU2
+ elim (lift_total T l m) /4 width=12 by step/
]
qed-.
-lemma l_deliftable_sn_LTC: ∀R. l_deliftable_sn R → l_deliftable_sn (LTC … R).
+lemma d_deliftable_sn_LTC: ∀R. d_deliftable_sn R → d_deliftable_sn (LTC … R).
#R #HR #L #U1 #U2 #H elim H -U2
-[ #U2 #HU12 #K #s #d #e #HLK #T1 #HTU1
+[ #U2 #HU12 #K #s #l #m #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
+| #U #U2 #_ #HU2 #IHU1 #K #s #l #m #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
+#R #HR #L1 #K1 #s #l #m #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
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
+[ #L2 #HL12 #K2 #s #m #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
+| #L #L2 #_ #HL2 #IHL1 #K2 #s #m #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
+lemma d_deliftable_sn_llstar: ∀R. d_deliftable_sn R →
+ ∀d. d_deliftable_sn (llstar … R d).
+#R #HR #d #L #U1 #U2 #H @(lstar_ind_r … d U2 H) -d -U2
[ /2 width=3 by lstar_O, ex2_intro/
-| #l #U #U2 #_ #HU2 #IHU1 #K #s #d #e #HLK #T1 #HTU1
+| #d #U #U2 #_ #HU2 #IHU1 #K #s #l #m #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/
]
(* Basic forward 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.
+lemma drop_fwd_drop2: ∀L1,I2,K2,V2,s,m. ⬇[s, O, m] L1 ≡ K2. ⓑ{I2} V2 →
+ ⬇[s, O, m + 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
+[ #I2 #K2 #V2 #s #m #H lapply (drop_inv_atom1 … H) -H * #H destruct
+| #K1 #I1 #V1 #IHL1 #I2 #K2 #V2 #s #m #H
+ elim (drop_inv_O1_pair1 … H) -H * #Hm #H
[ -IHL1 destruct /2 width=1 by drop_drop/
- | @drop_drop >(plus_minus_m_m e 1) /2 width=3 by/
+ | @drop_drop >(plus_minus_m_m m 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)
+lemma drop_fwd_length_ge: ∀L1,L2,l,m,s. ⬇[s, l, m] L1 ≡ L2 → |L1| ≤ l → |L2| = |L1|.
+#L1 #L2 #l #m #s #H elim H -L1 -L2 -l -m // normalize
+[ #I #L1 #L2 #V #m #_ #_ #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
+lemma drop_fwd_length_le_le: ∀L1,L2,l,m,s. ⬇[s, l, m] L1 ≡ L2 → l ≤ |L1| → m ≤ |L1| - l → |L2| = |L1| - m.
+#L1 #L2 #l #m #s #H elim H -L1 -L2 -l -m // 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/
+| #I #L1 #L2 #V1 #V2 #l #m #_ #_ #IHL12 >minus_plus_plus_l
+ #Hl #Hm lapply (le_plus_to_le_r … Hl) -Hl
+ #Hl >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
+lemma drop_fwd_length_le_ge: ∀L1,L2,l,m,s. ⬇[s, l, m] L1 ≡ L2 → l ≤ |L1| → |L1| - l ≤ m → |L2| = l.
+#L1 #L2 #l #m #s #H elim H -L1 -L2 -l -m 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/
]
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/
+lemma drop_fwd_length: ∀L1,L2,l,m. ⬇[Ⓕ, l, m] L1 ≡ L2 → |L1| = |L2| + m.
+#L1 #L2 #l #m #H elim H -L1 -L2 -l -m // 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/
+lemma drop_fwd_length_minus2: ∀L1,L2,l,m. ⬇[Ⓕ, l, m] L1 ≡ L2 → |L2| = |L1| - m.
+#L1 #L2 #l #m #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 //
+lemma drop_fwd_length_minus4: ∀L1,L2,l,m. ⬇[Ⓕ, l, m] L1 ≡ L2 → m = |L1| - |L2|.
+#L1 #L2 #l #m #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 //
+lemma drop_fwd_length_le2: ∀L1,L2,l,m. ⬇[Ⓕ, l, m] L1 ≡ L2 → m ≤ |L1|.
+#L1 #L2 #l #m #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 //
+lemma drop_fwd_length_le4: ∀L1,L2,l,m. ⬇[Ⓕ, l, m] L1 ≡ L2 → |L2| ≤ |L1|.
+#L1 #L2 #l #m #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
+lemma drop_fwd_length_lt2: ∀L1,I2,K2,V2,l,m.
+ ⬇[Ⓕ, l, m] L1 ≡ K2. ⓑ{I2} V2 → m < |L1|.
+#L1 #I2 #K2 #V2 #l #m #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/
+lemma drop_fwd_length_lt4: ∀L1,L2,l,m. ⬇[Ⓕ, l, m] L1 ≡ L2 → 0 < m → |L2| < |L1|.
+#L1 #L2 #l #m #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 →
+lemma drop_fwd_length_eq1: ∀L1,L2,K1,K2,l,m. ⬇[Ⓕ, l, m] L1 ≡ K1 → ⬇[Ⓕ, l, m] L2 ≡ K2 →
|L1| = |L2| → |K1| = |K2|.
-#L1 #L2 #K1 #K2 #d #e #HLK1 #HLK2 #HL12
+#L1 #L2 #K1 #K2 #l #m #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 →
+lemma drop_fwd_length_eq2: ∀L1,L2,K1,K2,l,m. ⬇[Ⓕ, l, m] L1 ≡ K1 → ⬇[Ⓕ, l, m] L2 ≡ K2 →
|K1| = |K2| → |L1| = |L2|.
-#L1 #L2 #K1 #K2 #d #e #HLK1 #HLK2 #HL12
+#L1 #L2 #K1 #K2 #l #m #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
+lemma drop_fwd_lw: ∀L1,L2,s,l,m. ⬇[s, l, m] L1 ≡ L2 → ♯{L2} ≤ ♯{L1}.
+#L1 #L2 #s #l #m #H elim H -L1 -L2 -l -m // normalize
[ /2 width=3 by transitive_le/
-| #I #L1 #L2 #V1 #V2 #d #e #_ #HV21 #IHL12
+| #I #L1 #L2 #V1 #V2 #l #m #_ #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 //
+lemma drop_fwd_lw_lt: ∀L1,L2,l,m. ⬇[Ⓕ, l, m] L1 ≡ L2 → 0 < m → ♯{L2} < ♯{L1}.
+#L1 #L2 #l #m #H elim H -L1 -L2 -l -m
+[ #l #m #H >H -H //
| #I #L #V #H elim (lt_refl_false … H)
-| #I #L1 #L2 #V #e #HL12 #_ #_
+| #I #L1 #L2 #V #m #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
+| #I #L1 #L2 #V1 #V2 #l #m #_ #HV21 #IHL12 #H normalize in ⊢ (?%%); -I
>(lift_fwd_tw … HV21) -V2 /3 by lt_minus_to_plus/
]
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
+fact drop_inv_O2_aux: ∀L1,L2,s,l,m. ⬇[s, l, m] L1 ≡ L2 → m = 0 → L1 = L2.
+#L1 #L2 #s #l #m #H elim H -L1 -L2 -l -m
[ //
| //
-| #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 //
+| #I #L1 #L2 #V #m #_ #_ >commutative_plus normalize #H destruct
+| #I #L1 #L2 #V1 #V2 #l #m #_ #HV21 #IHL12 #H
+ >(IHL12 H) -L1 >(lift_inv_O2_aux … HV21 … H) -V2 -l -m //
]
qed-.
(* Basic_1: was: drop_gen_refl *)
-lemma drop_inv_O2: ∀L1,L2,s,d. ⬇[s, d, 0] L1 ≡ L2 → L1 = L2.
+lemma drop_inv_O2: ∀L1,L2,s,l. ⬇[s, l, 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) //
+lemma drop_inv_length_eq: ∀L1,L2,l,m. ⬇[Ⓕ, l, m] L1 ≡ L2 → |L1| = |L2| → m = 0.
+#L1 #L2 #l #m #H #HL12 lapply (drop_fwd_length_minus4 … H) //
qed-.
-lemma drop_inv_refl: ∀L,d,e. ⬇[Ⓕ, d, e] L ≡ L → e = 0.
+lemma drop_inv_refl: ∀L,l,m. ⬇[Ⓕ, l, m] L ≡ L → m = 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
+fact drop_inv_FT_aux: ∀L1,L2,s,l,m. ⬇[s, l, m] L1 ≡ L2 →
+ ∀I,K,V. L2 = K.ⓑ{I}V → s = Ⓣ → l = 0 →
+ ⬇[Ⓕ, l, m] L1 ≡ K.ⓑ{I}V.
+#L1 #L2 #s #l #m #H elim H -L1 -L2 -l -m
+[ #l #m #_ #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
+| #I #L1 #L2 #V #m #_ #IHL12 #J #K #W #H1 #H2 destruct
/3 width=1 by drop_drop/
-| #I #L1 #L2 #V1 #V2 #d #e #_ #_ #_ #J #K #W #_ #_
+| #I #L1 #L2 #V1 #V2 #l #m #_ #_ #_ #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.
+lemma drop_inv_FT: ∀I,L,K,V,m. ⬇[Ⓣ, 0, m] L ≡ K.ⓑ{I}V → ⬇[m] 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.
+lemma drop_inv_gen: ∀I,L,K,V,s,m. ⬇[s, 0, m] L ≡ K.ⓑ{I}V → ⬇[m] 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.
+lemma drop_inv_T: ∀I,L,K,V,s,m. ⬇[Ⓣ, 0, m] L ≡ K.ⓑ{I}V → ⬇[s, 0, m] L ≡ K.ⓑ{I}V.
#I #L #K #V * /2 width=1 by drop_inv_FT/
qed-.
(* 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
+fact drop_O1_append_sn_le_aux: ∀L1,L2,s,l,m. ⬇[s, l, m] L1 ≡ L2 →
+ l = 0 → m ≤ |L1| →
+ ∀L. ⬇[s, 0, m] L @@ L1 ≡ L @@ L2.
+#L1 #L2 #s #l #m #H elim H -L1 -L2 -l -m 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 //
+#l #m #_ #_ #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.
+lemma drop_O1_append_sn_le: ∀L1,L2,s,m. ⬇[s, 0, m] L1 ≡ L2 → m ≤ |L1| →
+ ∀L. ⬇[s, 0, m] 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.
+lemma drop_O1_inv_append1_ge: ∀K,L1,L2,s,m. ⬇[s, 0, m] L1 @@ L2 ≡ K →
+ |L2| ≤ m → ⬇[s, 0, m - |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
+#L2 #I #V #IHL2 #s #m #H #H1m
+elim (drop_inv_O1_pair1 … H) -H * #H2m #HL12 destruct
+[ lapply (le_n_O_to_eq … H1m) -H1m -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.
+lemma drop_O1_inv_append1_le: ∀K,L1,L2,s,m. ⬇[s, 0, m] L1 @@ L2 ≡ K → m ≤ |L2| →
+ ∀K2. ⬇[s, 0, m] 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
+[ #s #m #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 ]
+| #L2 #I #V #IHL2 #s #m @(nat_ind_plus … m) -m [ -IHL2 ]
[ #H1 #_ #K2 #H2
lapply (drop_inv_O2 … H1) -H1 #H1
lapply (drop_inv_O2 … H2) -H2 #H2 destruct //
(* 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 //
+theorem drop_mono: ∀L,L1,s1,l,m. ⬇[s1, l, m] L ≡ L1 →
+ ∀L2,s2. ⬇[s2, l, m] L ≡ L2 → L1 = L2.
+#L #L1 #s1 #l #m #H elim H -L -L1 -l -m
+[ #l #m #Hm #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
+| #I #L #K #V #m #_ #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
+| #I #L #K1 #T #V1 #l #m #_ #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
+theorem drop_conf_ge: ∀L,L1,s1,l1,m1. ⬇[s1, l1, m1] L ≡ L1 →
+ ∀L2,s2,m2. ⬇[s2, 0, m2] L ≡ L2 → l1 + m1 ≤ m2 →
+ ⬇[s2, 0, m2 - m1] L1 ≡ L2.
+#L #L1 #s1 #l1 #m1 #H elim H -L -L1 -l1 -m1 //
+[ #l #m #_ #L2 #s2 #m2 #H #_ elim (drop_inv_atom1 … H) -H
+ #H #Hm destruct
+ @drop_atom #H >Hm // (**) (* explicit constructor *)
+| #I #L #K #V #m #_ #IHLK #L2 #s2 #m2 #H #Hm2
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
+| #I #L #K #V1 #V2 #l #m #_ #_ #IHLK #L2 #s2 #m2 #H #Hlmm2
+ lapply (transitive_le 1 … Hlmm2) // #Hm2
+ lapply (drop_inv_drop1_lt … H ?) -H // -Hm2 #HL2
+ lapply (transitive_le (1+m) … Hlmm2) // #Hmm2
@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
+theorem drop_conf_be: ∀L0,L1,s1,l1,m1. ⬇[s1, l1, m1] L0 ≡ L1 →
+ ∀L2,m2. ⬇[m2] L0 ≡ L2 → l1 ≤ m2 → m2 ≤ l1 + m1 →
+ ∃∃L. ⬇[s1, 0, l1 + m1 - m2] L2 ≡ L & ⬇[l1] L1 ≡ L.
+#L0 #L1 #s1 #l1 #m1 #H elim H -L0 -L1 -l1 -m1
+[ #l1 #m1 #Hm1 #L2 #m2 #H #Hl1 #_ elim (drop_inv_atom1 … H) -H #H #Hm2 destruct
+ >(Hm2 ?) in Hl1; // -Hm2 #Hl1 <(le_n_O_to_eq … Hl1) -l1
/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
+| normalize #I #L #V #L2 #m2 #HL2 #_ #Hm2
+ lapply (le_n_O_to_eq … Hm2) -Hm2 #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/
+| normalize #I #L0 #K0 #V1 #m1 #HLK0 #IHLK0 #L2 #m2 #H #_ #Hm21
+ lapply (drop_inv_O1_pair1 … H) -H * * #Hm2 #HL20
+ [ -IHLK0 -Hm21 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
+| #I #L0 #K0 #V0 #V1 #l1 #m1 >plus_plus_comm_23 #_ #_ #IHLK0 #L2 #m2 #H #Hl1m2 #Hm2lm1
+ elim (le_inv_plus_l … Hl1m2) #_ #Hm2
<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
+theorem drop_conf_le: ∀L0,L1,s1,l1,m1. ⬇[s1, l1, m1] L0 ≡ L1 →
+ ∀L2,s2,m2. ⬇[s2, 0, m2] L0 ≡ L2 → m2 ≤ l1 →
+ ∃∃L. ⬇[s2, 0, m2] L1 ≡ L & ⬇[s1, l1 - m2, m1] L2 ≡ L.
+#L0 #L1 #s1 #l1 #m1 #H elim H -L0 -L1 -l1 -m1
+[ #l1 #m1 #Hm1 #L2 #s2 #m2 #H elim (drop_inv_atom1 … H) -H
+ #H #Hm2 #_ destruct /4 width=3 by drop_atom, ex2_intro/
+| #I #K0 #V0 #L2 #s2 #m2 #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
+| #I #K0 #K1 #V0 #m1 #HK01 #_ #L2 #s2 #m2 #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
+| #I #K0 #K1 #V0 #V1 #l1 #m1 #HK01 #HV10 #IHK01 #L2 #s2 #m2 #H #Hm2l1
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
+ [ -IHK01 -Hm2l1 #H1 #H2 destruct /3 width=5 by drop_pair, drop_skip, ex2_intro/
+ | -HK01 -HV10 #Hm2 #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/
]
(* 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/
+theorem drop_trans_ge: ∀L1,L,s1,l1,m1. ⬇[s1, l1, m1] L1 ≡ L →
+ ∀L2,m2. ⬇[m2] L ≡ L2 → l1 ≤ m2 → ⬇[s1, 0, m1 + m2] L1 ≡ L2.
+#L1 #L #s1 #l1 #m1 #H elim H -L1 -L -l1 -m1
+[ #l1 #m1 #Hm1 #L2 #m2 #H #_ elim (drop_inv_atom1 … H) -H
+ #H #Hm2 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
+| #I #L1 #L2 #V1 #V2 #l #m #_ #_ #IHL12 #L #m2 #H #Hlm2
+ lapply (lt_to_le_to_lt 0 … Hlm2) // #Hm2
+ lapply (lt_to_le_to_lt … (m + m2) Hm2 ?) // #Hmm2
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
+theorem drop_trans_le: ∀L1,L,s1,l1,m1. ⬇[s1, l1, m1] L1 ≡ L →
+ ∀L2,s2,m2. ⬇[s2, 0, m2] L ≡ L2 → m2 ≤ l1 →
+ ∃∃L0. ⬇[s2, 0, m2] L1 ≡ L0 & ⬇[s1, l1 - m2, m1] L0 ≡ L2.
+#L1 #L #s1 #l1 #m1 #H elim H -L1 -L -l1 -m1
+[ #l1 #m1 #Hm1 #L2 #s2 #m2 #H #_ elim (drop_inv_atom1 … H) -H
+ #H #Hm2 destruct /4 width=3 by drop_atom, ex2_intro/
+| #I #K #V #L2 #s2 #m2 #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
+| #I #L1 #L2 #V #m #_ #IHL12 #L #s2 #m2 #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
+| #I #L1 #L2 #V1 #V2 #l #m #HL12 #HV12 #IHL12 #L #s2 #m2 #H #Hm2l
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
+ [ -Hm2l -IHL12 #H1 #H2 destruct /3 width=5 by drop_pair, drop_skip, ex2_intro/
+ | -HL12 -HV12 #Hm2 #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/ ]
]
]
(* 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/
+lemma d_liftable_llstar: ∀R. d_liftable R → ∀d. d_liftable (llstar … R d).
+#R #HR #d #K #T1 #T2 #H @(lstar_ind_r … d T2 H) -d -T2
+[ #L #s #l #m #_ #U1 #HTU1 #U2 #HTU2 -HR -K
+ >(lift_mono … HTU2 … HTU1) -T1 -U2 -l -m //
+| #d #T #T2 #_ #HT2 #IHT1 #L #s #l #m #HLK #U1 #HTU1 #U2 #HTU2
+ elim (lift_total T l m) /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
+lemma drop_conf_lt: ∀L,L1,s1,l1,m1. ⬇[s1, l1, m1] L ≡ L1 →
+ ∀I,K2,V2,s2,m2. ⬇[s2, 0, m2] L ≡ K2.ⓑ{I}V2 →
+ m2 < l1 → let l ≝ l1 - m2 - 1 in
+ ∃∃K1,V1. ⬇[s2, 0, m2] L1 ≡ K1.ⓑ{I}V1 &
+ ⬇[s1, l, m1] K2 ≡ K1 & ⬆[l, m1] V1 ≡ V2.
+#L #L1 #s1 #l1 #m1 #H1 #I #K2 #V2 #s2 #m2 #H2 #Hm2l1
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
+lemma drop_trans_lt: ∀L1,L,s1,l1,m1. ⬇[s1, l1, m1] L1 ≡ L →
+ ∀I,L2,V2,s2,m2. ⬇[s2, 0, m2] L ≡ L2.ⓑ{I}V2 →
+ m2 < l1 → let l ≝ l1 - m2 - 1 in
+ ∃∃L0,V0. ⬇[s2, 0, m2] L1 ≡ L0.ⓑ{I}V0 &
+ ⬇[s1, l, m1] L0 ≡ L2 & ⬆[l, m1] V2 ≡ V0.
+#L1 #L #s1 #l1 #m1 #HL1 #I #L2 #V2 #s2 #m2 #HL2 #Hl21
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
+lemma drop_trans_ge_comm: ∀L1,L,L2,s1,l1,m1,m2.
+ ⬇[s1, l1, m1] L1 ≡ L → ⬇[m2] L ≡ L2 → l1 ≤ m2 →
+ ⬇[s1, 0, m2 + m1] L1 ≡ L2.
+#L1 #L #L2 #s1 #l1 #m1 #m2
>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
+lemma drop_conf_div: ∀I1,L,K,V1,m1. ⬇[m1] L ≡ K.ⓑ{I1}V1 →
+ ∀I2,V2,m2. ⬇[m2] L ≡ K.ⓑ{I2}V2 →
+ ∧∧ m1 = m2 & I1 = I2 & V1 = V2.
+#I1 #L #K #V1 #m1 #HLK1 #I2 #V2 #m2 #HLK2
+elim (le_or_ge m1 m2) #Hm
[ lapply (drop_conf_ge … HLK1 … HLK2 ?)
| lapply (drop_conf_ge … HLK2 … HLK1 ?)
] -HLK1 -HLK2 // #HK
(* 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|)) //
+lemma drop_fwd_be: ∀L,K,s,l,m,i. ⬇[s, l, m] L ≡ K → |K| ≤ i → i < l → |L| ≤ i.
+#L #K #s #l #m #i #HLK #HK #Hl 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
+elim (drop_conf_lt … HLK … HLK0) // -HLK -HLK0 -Hl
#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-.
--- /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/lreq_lreq.ma".
+include "basic_2/substitution/drop.ma".
+
+(* BASIC SLICING FOR LOCAL ENVIRONMENTS *************************************)
+
+definition dedropable_sn: predicate (relation lenv) ≝
+ λR. ∀L1,K1,s,l,m. ⬇[s, l, m] L1 ≡ K1 → ∀K2. R K1 K2 →
+ ∃∃L2. R L1 L2 & ⬇[s, l, m] L2 ≡ K2 & L1 ⩬[l, m] L2.
+
+(* Properties on equivalence ************************************************)
+
+lemma lreq_drop_trans_be: ∀L1,L2,l,m. L1 ⩬[l, m] L2 →
+ ∀I,K2,W,s,i. ⬇[s, 0, i] L2 ≡ K2.ⓑ{I}W →
+ l ≤ i → i < l + m →
+ ∃∃K1. K1 ⩬[0, ⫰(l+m-i)] K2 & ⬇[s, 0, i] L1 ≡ K1.ⓑ{I}W.
+#L1 #L2 #l #m #H elim H -L1 -L2 -l -m
+[ #l #m #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 #m #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
+ #Him elim (IHL12 … HLK1) -IHL12 -HLK1 // -Him
+ >yminus_succ <yminus_inj /3 width=3 by drop_drop_lt, ex2_intro/
+ ]
+| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #_ #IHL12 #J #K2 #W #s #i #HLK2 #Hli
+ elim (yle_inv_succ1 … Hli) -Hli
+ #Hli #Hi <Hi >yplus_succ1 #H lapply (ylt_inv_succ … H) -H
+ #Hilm 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 lreq_drop_conf_be: ∀L1,L2,l,m. L1 ⩬[l, m] L2 →
+ ∀I,K1,W,s,i. ⬇[s, 0, i] L1 ≡ K1.ⓑ{I}W →
+ l ≤ i → i < l + m →
+ ∃∃K2. K1 ⩬[0, ⫰(l+m-i)] K2 & ⬇[s, 0, i] L2 ≡ K2.ⓑ{I}W.
+#L1 #L2 #l #m #HL12 #I #K1 #W #s #i #HLK1 #Hli #Hilm
+elim (lreq_drop_trans_be … (lreq_sym … HL12) … HLK1) // -L1 -Hli -Hilm
+/3 width=3 by lreq_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 lreq_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, lreq_pair, injective_plus_l, ex2_intro/
+]
+qed-.
+
+lemma dedropable_sn_TC: ∀R. dedropable_sn R → dedropable_sn (TC … R).
+#R #HR #L1 #K1 #s #l #m #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 lreq_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, lreq_succ/ ]
+ normalize <plus_n_Sm #H destruct
+]
+qed-.
| fqu_pair_sn: ∀I,G,L,V,T. fqu G L (②{I}V.T) G L V
| fqu_bind_dx: ∀a,I,G,L,V,T. fqu G L (ⓑ{a,I}V.T) G (L.ⓑ{I}V) T
| fqu_flat_dx: ∀I,G,L,V,T. fqu G L (ⓕ{I}V.T) G L T
-| fqu_drop : ∀G,L,K,T,U,e.
- ⬇[e+1] L ≡ K → ⬆[0, e+1] T ≡ U → fqu G L U G K T
+| fqu_drop : ∀G,L,K,T,U,m.
+ ⬇[m+1] L ≡ K → ⬆[0, m+1] T ≡ U → fqu G L U G K T
.
interpretation
(* Basic properties *********************************************************)
-lemma fqu_drop_lt: ∀G,L,K,T,U,e. 0 < e →
- ⬇[e] L ≡ K → ⬆[0, e] T ≡ U → ⦃G, L, U⦄ ⊐ ⦃G, K, T⦄.
-#G #L #K #T #U #e #He >(plus_minus_m_m e 1) /2 width=3 by fqu_drop/
+lemma fqu_drop_lt: ∀G,L,K,T,U,m. 0 < m →
+ ⬇[m] L ≡ K → ⬆[0, m] T ≡ U → ⦃G, L, U⦄ ⊐ ⦃G, K, T⦄.
+#G #L #K #T #U #m #Hm >(plus_minus_m_m m 1) /2 width=3 by fqu_drop/
qed.
lemma fqu_lref_S_lt: ∀I,G,L,V,i. 0 < i → ⦃G, L.ⓑ{I}V, #i⦄ ⊐ ⦃G, L, #(i-1)⦄.
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
+#G #L #K #T #U #m #HLK #HTU
lapply (drop_fwd_lw_lt … HLK ?) -HLK // #HKL
-lapply (lift_fwd_tw … HTU) -e #H
+lapply (lift_fwd_tw … HTU) -m #H
normalize in ⊢ (?%%); /2 width=1 by lt_minus_to_plus/
qed-.
G1 = G2 → |L1| = |L2| → T1 = T2 → ⊥.
#G1 #G2 #L1 #L2 #T1 #T2 * -G1 -G2 -L1 -L2 -T1 -T2 normalize
/2 width=4 by discr_tpair_xy_y, discr_tpair_xy_x, plus_xSy_x_false/
-#G #L #K #T #U #e #HLK #_ #_ #H #_ -G -T -U >(drop_fwd_length … HLK) in H; -L
+#G #L #K #T #U #m #HLK #_ #_ #H #_ -G -T -U >(drop_fwd_length … HLK) in H; -L
/2 width=4 by plus_xySz_x_false/
qed-.
| fquq_pair_sn: ∀I,G,L,V,T. fquq G L (②{I}V.T) G L V
| fquq_bind_dx: ∀a,I,G,L,V,T. fquq G L (ⓑ{a,I}V.T) G (L.ⓑ{I}V) T
| fquq_flat_dx: ∀I,G, L,V,T. fquq G L (ⓕ{I}V.T) G L T
-| fquq_drop : ∀G,L,K,T,U,e.
- ⬇[e] L ≡ K → ⬆[0, e] T ≡ U → fquq G L U G K T
+| fquq_drop : ∀G,L,K,T,U,m.
+ ⬇[m] L ≡ K → ⬆[0, m] T ≡ U → fquq G L U G K T
.
interpretation
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
+#G1 #L1 #K1 #T1 #U1 #m #HLK1 #HTU1
lapply (drop_fwd_lw … HLK1) -HLK1
lapply (lift_fwd_tw … HTU1) -HTU1
/2 width=1 by le_plus, le_n/
∀i. T1 = #i → |L2| ≤ |L1|.
#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
+| #G1 #L1 #K1 #T1 #U1 #m #HLK1 #HTU1 #i #H destruct
/2 width=3 by drop_fwd_length_le4/
]
qed-.
lemma fquqa_refl: tri_reflexive … fquqa.
// qed.
-lemma fquqa_drop: ∀G,L,K,T,U,e.
- ⬇[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)
+lemma fquqa_drop: ∀G,L,K,T,U,m.
+ ⬇[m] L ≡ K → ⬆[0, m] T ≡ U → ⦃G, L, U⦄ ⊐⊐⸮ ⦃G, K, T⦄.
+#G #L #K #T #U #m #HLK #HTU elim (eq_or_gt m)
/3 width=5 by fqu_drop_lt, or_introl/ #H destruct
>(drop_inv_O2 … HLK) -L >(lift_inv_O2 … HTU) -T //
qed.
(* GLOBAL ENVIRONMENT READING ***********************************************)
-inductive gget (e:nat): relation genv ≝
-| gget_gt: ∀G. |G| ≤ e → gget e G (⋆)
-| gget_eq: ∀G. |G| = e + 1 → gget e G G
-| gget_lt: ∀I,G1,G2,V. e < |G1| → gget e G1 G2 → gget e (G1. ⓑ{I} V) G2
+inductive gget (m:nat): relation genv ≝
+| gget_gt: ∀G. |G| ≤ m → gget m G (⋆)
+| gget_eq: ∀G. |G| = m + 1 → gget m G G
+| gget_lt: ∀I,G1,G2,V. m < |G1| → gget m G1 G2 → gget m (G1. ⓑ{I} V) G2
.
interpretation "global reading"
- 'RDrop e G1 G2 = (gget e G1 G2).
+ 'RDrop m G1 G2 = (gget m G1 G2).
(* basic inversion lemmas ***************************************************)
-lemma gget_inv_gt: ∀G1,G2,e. ⬇[e] G1 ≡ G2 → |G1| ≤ e → G2 = ⋆.
-#G1 #G2 #e * -G1 -G2 //
+lemma gget_inv_gt: ∀G1,G2,m. ⬇[m] G1 ≡ G2 → |G1| ≤ m → G2 = ⋆.
+#G1 #G2 #m * -G1 -G2 //
[ #G #H >H -H >commutative_plus #H (**) (* lemma needed here *)
lapply (le_plus_to_le_r … 0 H) -H #H
lapply (le_n_O_to_eq … H) -H #H destruct
]
qed-.
-lemma gget_inv_eq: ∀G1,G2,e. ⬇[e] G1 ≡ G2 → |G1| = e + 1 → G1 = G2.
-#G1 #G2 #e * -G1 -G2 //
+lemma gget_inv_eq: ∀G1,G2,m. ⬇[m] G1 ≡ G2 → |G1| = m + 1 → G1 = G2.
+#G1 #G2 #m * -G1 -G2 //
[ #G #H1 #H2 >H2 in H1; -H2 >commutative_plus #H (**) (* lemma needed here *)
lapply (le_plus_to_le_r … 0 H) -H #H
lapply (le_n_O_to_eq … H) -H #H destruct
]
qed-.
-fact gget_inv_lt_aux: ∀I,G,G1,G2,V,e. ⬇[e] G ≡ G2 → G = G1. ⓑ{I} V →
- e < |G1| → ⬇[e] G1 ≡ G2.
-#I #G #G1 #G2 #V #e * -G -G2
+fact gget_inv_lt_aux: ∀I,G,G1,G2,V,m. ⬇[m] G ≡ G2 → G = G1. ⓑ{I} V →
+ m < |G1| → ⬇[m] G1 ≡ G2.
+#I #G #G1 #G2 #V #m * -G -G2
[ #G #H1 #H destruct #H2
lapply (le_to_lt_to_lt … H1 H2) -H1 -H2 normalize in ⊢ (? % ? → ?); >commutative_plus #H
lapply (lt_plus_to_lt_l … 0 H) -H #H
]
qed-.
-lemma gget_inv_lt: ∀I,G1,G2,V,e.
- ⬇[e] G1. ⓑ{I} V ≡ G2 → e < |G1| → ⬇[e] G1 ≡ G2.
+lemma gget_inv_lt: ∀I,G1,G2,V,m.
+ ⬇[m] G1. ⓑ{I} V ≡ G2 → m < |G1| → ⬇[m] G1 ≡ G2.
/2 width=5 by gget_inv_lt_aux/ qed-.
(* Basic properties *********************************************************)
-lemma gget_total: ∀e,G1. ∃G2. ⬇[e] G1 ≡ G2.
-#e #G1 elim G1 -G1 /3 width=2/
+lemma gget_total: ∀m,G1. ∃G2. ⬇[m] G1 ≡ G2.
+#m #G1 elim G1 -G1 /3 width=2 by gget_gt, ex_intro/
#I #V #G1 * #G2 #HG12
-elim (lt_or_eq_or_gt e (|G1|)) #He
+elim (lt_or_eq_or_gt m (|G1|)) #Hm
[ /3 width=2/
-| destruct /3 width=2/
-| @ex_intro [2: @gget_gt normalize /2 width=1/ | skip ] (**) (* explicit constructor *)
+| destruct /3 width=2 by gget_eq, ex_intro/
+| @ex_intro [2: @gget_gt normalize /2 width=1 by/ | skip ] (**) (* explicit constructor *)
]
qed-.
(* Main properties **********************************************************)
-theorem gget_mono: ∀G,G1,e. ⬇[e] G ≡ G1 → ∀G2. ⬇[e] G ≡ G2 → G1 = G2.
-#G #G1 #e #H elim H -G -G1
-[ #G #He #G2 #H
- >(gget_inv_gt … H He) -H -He //
-| #G #He #G2 #H
- >(gget_inv_eq … H He) -H -He //
-| #I #G #G1 #V #He #_ #IHG1 #G2 #H
- lapply (gget_inv_lt … H He) -H -He /2 width=1/
+theorem gget_mono: ∀G,G1,m. ⬇[m] G ≡ G1 → ∀G2. ⬇[m] G ≡ G2 → G1 = G2.
+#G #G1 #m #H elim H -G -G1
+[ #G #Hm #G2 #H
+ >(gget_inv_gt … H Hm) -H -Hm //
+| #G #Hm #G2 #H
+ >(gget_inv_eq … H Hm) -H -Hm //
+| #I #G #G1 #V #Hm #_ #IHG1 #G2 #H
+ lapply (gget_inv_lt … H Hm) -H -Hm /2 width=1/
]
qed-.
-lemma gget_dec: ∀G1,G2,e. Decidable (⬇[e] G1 ≡ G2).
-#G1 #G2 #e
-elim (gget_total e G1) #G #HG1
+lemma gget_dec: ∀G1,G2,m. Decidable (⬇[m] G1 ≡ G2).
+#G1 #G2 #m
+elim (gget_total m G1) #G #HG1
elim (eq_genv_dec G G2) #HG2
[ destruct /2 width=1/
| @or_intror #HG12
lift_sort lift_lref_lt lift_lref_ge lift_bind lift_flat
*)
inductive lift: relation4 nat nat term term ≝
-| lift_sort : ∀k,d,e. lift d e (⋆k) (⋆k)
-| lift_lref_lt: ∀i,d,e. i < d → lift d e (#i) (#i)
-| lift_lref_ge: ∀i,d,e. d ≤ i → lift d e (#i) (#(i + e))
-| lift_gref : ∀p,d,e. lift d e (§p) (§p)
-| lift_bind : ∀a,I,V1,V2,T1,T2,d,e.
- lift d e V1 V2 → lift (d + 1) e T1 T2 →
- lift d e (ⓑ{a,I} V1. T1) (ⓑ{a,I} V2. T2)
-| lift_flat : ∀I,V1,V2,T1,T2,d,e.
- lift d e V1 V2 → lift d e T1 T2 →
- lift d e (ⓕ{I} V1. T1) (ⓕ{I} V2. T2)
+| lift_sort : ∀k,l,m. lift l m (⋆k) (⋆k)
+| lift_lref_lt: ∀i,l,m. i < l → lift l m (#i) (#i)
+| lift_lref_ge: ∀i,l,m. l ≤ i → lift l m (#i) (#(i + m))
+| lift_gref : ∀p,l,m. lift l m (§p) (§p)
+| lift_bind : ∀a,I,V1,V2,T1,T2,l,m.
+ lift l m V1 V2 → lift (l + 1) m T1 T2 →
+ lift l m (ⓑ{a,I} V1. T1) (ⓑ{a,I} V2. T2)
+| lift_flat : ∀I,V1,V2,T1,T2,l,m.
+ lift l m V1 V2 → lift l m T1 T2 →
+ lift l m (ⓕ{I} V1. T1) (ⓕ{I} V2. T2)
.
-interpretation "relocation" 'RLift d e T1 T2 = (lift d e T1 T2).
+interpretation "relocation" 'RLift l m T1 T2 = (lift l m T1 T2).
(* Basic inversion lemmas ***************************************************)
-fact lift_inv_O2_aux: ∀d,e,T1,T2. ⬆[d, e] T1 ≡ T2 → e = 0 → T1 = T2.
-#d #e #T1 #T2 #H elim H -d -e -T1 -T2 /3 width=1 by eq_f2/
+fact lift_inv_O2_aux: ∀l,m,T1,T2. ⬆[l, m] T1 ≡ T2 → m = 0 → T1 = T2.
+#l #m #T1 #T2 #H elim H -l -m -T1 -T2 /3 width=1 by eq_f2/
qed-.
-lemma lift_inv_O2: ∀d,T1,T2. ⬆[d, 0] T1 ≡ T2 → T1 = T2.
+lemma lift_inv_O2: ∀l,T1,T2. ⬆[l, 0] T1 ≡ T2 → T1 = T2.
/2 width=4 by lift_inv_O2_aux/ qed-.
-fact lift_inv_sort1_aux: ∀d,e,T1,T2. ⬆[d,e] T1 ≡ T2 → ∀k. T1 = ⋆k → T2 = ⋆k.
-#d #e #T1 #T2 * -d -e -T1 -T2 //
-[ #i #d #e #_ #k #H destruct
-| #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct
-| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct
+fact lift_inv_sort1_aux: ∀l,m,T1,T2. ⬆[l,m] T1 ≡ T2 → ∀k. T1 = ⋆k → T2 = ⋆k.
+#l #m #T1 #T2 * -l -m -T1 -T2 //
+[ #i #l #m #_ #k #H destruct
+| #a #I #V1 #V2 #T1 #T2 #l #m #_ #_ #k #H destruct
+| #I #V1 #V2 #T1 #T2 #l #m #_ #_ #k #H destruct
]
qed-.
-lemma lift_inv_sort1: ∀d,e,T2,k. ⬆[d,e] ⋆k ≡ T2 → T2 = ⋆k.
+lemma lift_inv_sort1: ∀l,m,T2,k. ⬆[l,m] ⋆k ≡ T2 → T2 = ⋆k.
/2 width=5 by lift_inv_sort1_aux/ qed-.
-fact lift_inv_lref1_aux: ∀d,e,T1,T2. ⬆[d,e] T1 ≡ T2 → ∀i. T1 = #i →
- (i < d ∧ T2 = #i) ∨ (d ≤ i ∧ T2 = #(i + e)).
-#d #e #T1 #T2 * -d -e -T1 -T2
-[ #k #d #e #i #H destruct
-| #j #d #e #Hj #i #Hi destruct /3 width=1 by or_introl, conj/
-| #j #d #e #Hj #i #Hi destruct /3 width=1 by or_intror, conj/
-| #p #d #e #i #H destruct
-| #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct
-| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct
+fact lift_inv_lref1_aux: ∀l,m,T1,T2. ⬆[l,m] T1 ≡ T2 → ∀i. T1 = #i →
+ (i < l ∧ T2 = #i) ∨ (l ≤ i ∧ T2 = #(i + m)).
+#l #m #T1 #T2 * -l -m -T1 -T2
+[ #k #l #m #i #H destruct
+| #j #l #m #Hj #i #Hi destruct /3 width=1 by or_introl, conj/
+| #j #l #m #Hj #i #Hi destruct /3 width=1 by or_intror, conj/
+| #p #l #m #i #H destruct
+| #a #I #V1 #V2 #T1 #T2 #l #m #_ #_ #i #H destruct
+| #I #V1 #V2 #T1 #T2 #l #m #_ #_ #i #H destruct
]
qed-.
-lemma lift_inv_lref1: ∀d,e,T2,i. ⬆[d,e] #i ≡ T2 →
- (i < d ∧ T2 = #i) ∨ (d ≤ i ∧ T2 = #(i + e)).
+lemma lift_inv_lref1: ∀l,m,T2,i. ⬆[l,m] #i ≡ T2 →
+ (i < l ∧ T2 = #i) ∨ (l ≤ i ∧ T2 = #(i + m)).
/2 width=3 by lift_inv_lref1_aux/ qed-.
-lemma lift_inv_lref1_lt: ∀d,e,T2,i. ⬆[d,e] #i ≡ T2 → i < d → T2 = #i.
-#d #e #T2 #i #H elim (lift_inv_lref1 … H) -H * //
-#Hdi #_ #Hid lapply (le_to_lt_to_lt … Hdi Hid) -Hdi -Hid #Hdd
-elim (lt_refl_false … Hdd)
+lemma lift_inv_lref1_lt: ∀l,m,T2,i. ⬆[l,m] #i ≡ T2 → i < l → T2 = #i.
+#l #m #T2 #i #H elim (lift_inv_lref1 … H) -H * //
+#Hli #_ #Hil lapply (le_to_lt_to_lt … Hli Hil) -Hli -Hil #Hll
+elim (lt_refl_false … Hll)
qed-.
-lemma lift_inv_lref1_ge: ∀d,e,T2,i. ⬆[d,e] #i ≡ T2 → d ≤ i → T2 = #(i + e).
-#d #e #T2 #i #H elim (lift_inv_lref1 … H) -H * //
-#Hid #_ #Hdi lapply (le_to_lt_to_lt … Hdi Hid) -Hdi -Hid #Hdd
-elim (lt_refl_false … Hdd)
+lemma lift_inv_lref1_ge: ∀l,m,T2,i. ⬆[l,m] #i ≡ T2 → l ≤ i → T2 = #(i + m).
+#l #m #T2 #i #H elim (lift_inv_lref1 … H) -H * //
+#Hil #_ #Hli lapply (le_to_lt_to_lt … Hli Hil) -Hli -Hil #Hll
+elim (lt_refl_false … Hll)
qed-.
-fact lift_inv_gref1_aux: ∀d,e,T1,T2. ⬆[d,e] T1 ≡ T2 → ∀p. T1 = §p → T2 = §p.
-#d #e #T1 #T2 * -d -e -T1 -T2 //
-[ #i #d #e #_ #k #H destruct
-| #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct
-| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct
+fact lift_inv_gref1_aux: ∀l,m,T1,T2. ⬆[l,m] T1 ≡ T2 → ∀p. T1 = §p → T2 = §p.
+#l #m #T1 #T2 * -l -m -T1 -T2 //
+[ #i #l #m #_ #k #H destruct
+| #a #I #V1 #V2 #T1 #T2 #l #m #_ #_ #k #H destruct
+| #I #V1 #V2 #T1 #T2 #l #m #_ #_ #k #H destruct
]
qed-.
-lemma lift_inv_gref1: ∀d,e,T2,p. ⬆[d,e] §p ≡ T2 → T2 = §p.
+lemma lift_inv_gref1: ∀l,m,T2,p. ⬆[l,m] §p ≡ T2 → T2 = §p.
/2 width=5 by lift_inv_gref1_aux/ qed-.
-fact lift_inv_bind1_aux: ∀d,e,T1,T2. ⬆[d,e] T1 ≡ T2 →
+fact lift_inv_bind1_aux: ∀l,m,T1,T2. ⬆[l,m] T1 ≡ T2 →
∀a,I,V1,U1. T1 = ⓑ{a,I} V1.U1 →
- ∃∃V2,U2. ⬆[d,e] V1 ≡ V2 & ⬆[d+1,e] U1 ≡ U2 &
+ ∃∃V2,U2. ⬆[l,m] V1 ≡ V2 & ⬆[l+1,m] U1 ≡ U2 &
T2 = ⓑ{a,I} V2. U2.
-#d #e #T1 #T2 * -d -e -T1 -T2
-[ #k #d #e #a #I #V1 #U1 #H destruct
-| #i #d #e #_ #a #I #V1 #U1 #H destruct
-| #i #d #e #_ #a #I #V1 #U1 #H destruct
-| #p #d #e #a #I #V1 #U1 #H destruct
-| #b #J #W1 #W2 #T1 #T2 #d #e #HW #HT #a #I #V1 #U1 #H destruct /2 width=5 by ex3_2_intro/
-| #J #W1 #W2 #T1 #T2 #d #e #_ #HT #a #I #V1 #U1 #H destruct
+#l #m #T1 #T2 * -l -m -T1 -T2
+[ #k #l #m #a #I #V1 #U1 #H destruct
+| #i #l #m #_ #a #I #V1 #U1 #H destruct
+| #i #l #m #_ #a #I #V1 #U1 #H destruct
+| #p #l #m #a #I #V1 #U1 #H destruct
+| #b #J #W1 #W2 #T1 #T2 #l #m #HW #HT #a #I #V1 #U1 #H destruct /2 width=5 by ex3_2_intro/
+| #J #W1 #W2 #T1 #T2 #l #m #_ #HT #a #I #V1 #U1 #H destruct
]
qed-.
-lemma lift_inv_bind1: ∀d,e,T2,a,I,V1,U1. ⬆[d,e] ⓑ{a,I} V1. U1 ≡ T2 →
- ∃∃V2,U2. ⬆[d,e] V1 ≡ V2 & ⬆[d+1,e] U1 ≡ U2 &
+lemma lift_inv_bind1: ∀l,m,T2,a,I,V1,U1. ⬆[l,m] ⓑ{a,I} V1. U1 ≡ T2 →
+ ∃∃V2,U2. ⬆[l,m] V1 ≡ V2 & ⬆[l+1,m] U1 ≡ U2 &
T2 = ⓑ{a,I} V2. U2.
/2 width=3 by lift_inv_bind1_aux/ qed-.
-fact lift_inv_flat1_aux: ∀d,e,T1,T2. ⬆[d,e] T1 ≡ T2 →
+fact lift_inv_flat1_aux: ∀l,m,T1,T2. ⬆[l,m] T1 ≡ T2 →
∀I,V1,U1. T1 = ⓕ{I} V1.U1 →
- ∃∃V2,U2. ⬆[d,e] V1 ≡ V2 & ⬆[d,e] U1 ≡ U2 &
+ ∃∃V2,U2. ⬆[l,m] V1 ≡ V2 & ⬆[l,m] U1 ≡ U2 &
T2 = ⓕ{I} V2. U2.
-#d #e #T1 #T2 * -d -e -T1 -T2
-[ #k #d #e #I #V1 #U1 #H destruct
-| #i #d #e #_ #I #V1 #U1 #H destruct
-| #i #d #e #_ #I #V1 #U1 #H destruct
-| #p #d #e #I #V1 #U1 #H destruct
-| #a #J #W1 #W2 #T1 #T2 #d #e #_ #_ #I #V1 #U1 #H destruct
-| #J #W1 #W2 #T1 #T2 #d #e #HW #HT #I #V1 #U1 #H destruct /2 width=5 by ex3_2_intro/
+#l #m #T1 #T2 * -l -m -T1 -T2
+[ #k #l #m #I #V1 #U1 #H destruct
+| #i #l #m #_ #I #V1 #U1 #H destruct
+| #i #l #m #_ #I #V1 #U1 #H destruct
+| #p #l #m #I #V1 #U1 #H destruct
+| #a #J #W1 #W2 #T1 #T2 #l #m #_ #_ #I #V1 #U1 #H destruct
+| #J #W1 #W2 #T1 #T2 #l #m #HW #HT #I #V1 #U1 #H destruct /2 width=5 by ex3_2_intro/
]
qed-.
-lemma lift_inv_flat1: ∀d,e,T2,I,V1,U1. ⬆[d,e] ⓕ{I} V1. U1 ≡ T2 →
- ∃∃V2,U2. ⬆[d,e] V1 ≡ V2 & ⬆[d,e] U1 ≡ U2 &
+lemma lift_inv_flat1: ∀l,m,T2,I,V1,U1. ⬆[l,m] ⓕ{I} V1. U1 ≡ T2 →
+ ∃∃V2,U2. ⬆[l,m] V1 ≡ V2 & ⬆[l,m] U1 ≡ U2 &
T2 = ⓕ{I} V2. U2.
/2 width=3 by lift_inv_flat1_aux/ qed-.
-fact lift_inv_sort2_aux: ∀d,e,T1,T2. ⬆[d,e] T1 ≡ T2 → ∀k. T2 = ⋆k → T1 = ⋆k.
-#d #e #T1 #T2 * -d -e -T1 -T2 //
-[ #i #d #e #_ #k #H destruct
-| #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct
-| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct
+fact lift_inv_sort2_aux: ∀l,m,T1,T2. ⬆[l,m] T1 ≡ T2 → ∀k. T2 = ⋆k → T1 = ⋆k.
+#l #m #T1 #T2 * -l -m -T1 -T2 //
+[ #i #l #m #_ #k #H destruct
+| #a #I #V1 #V2 #T1 #T2 #l #m #_ #_ #k #H destruct
+| #I #V1 #V2 #T1 #T2 #l #m #_ #_ #k #H destruct
]
qed-.
(* Basic_1: was: lift_gen_sort *)
-lemma lift_inv_sort2: ∀d,e,T1,k. ⬆[d,e] T1 ≡ ⋆k → T1 = ⋆k.
+lemma lift_inv_sort2: ∀l,m,T1,k. ⬆[l,m] T1 ≡ ⋆k → T1 = ⋆k.
/2 width=5 by lift_inv_sort2_aux/ qed-.
-fact lift_inv_lref2_aux: ∀d,e,T1,T2. ⬆[d,e] T1 ≡ T2 → ∀i. T2 = #i →
- (i < d ∧ T1 = #i) ∨ (d + e ≤ i ∧ T1 = #(i - e)).
-#d #e #T1 #T2 * -d -e -T1 -T2
-[ #k #d #e #i #H destruct
-| #j #d #e #Hj #i #Hi destruct /3 width=1 by or_introl, conj/
-| #j #d #e #Hj #i #Hi destruct <minus_plus_m_m /4 width=1 by monotonic_le_plus_l, or_intror, conj/
-| #p #d #e #i #H destruct
-| #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct
-| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct
+fact lift_inv_lref2_aux: ∀l,m,T1,T2. ⬆[l,m] T1 ≡ T2 → ∀i. T2 = #i →
+ (i < l ∧ T1 = #i) ∨ (l + m ≤ i ∧ T1 = #(i - m)).
+#l #m #T1 #T2 * -l -m -T1 -T2
+[ #k #l #m #i #H destruct
+| #j #l #m #Hj #i #Hi destruct /3 width=1 by or_introl, conj/
+| #j #l #m #Hj #i #Hi destruct <minus_plus_m_m /4 width=1 by monotonic_le_plus_l, or_intror, conj/
+| #p #l #m #i #H destruct
+| #a #I #V1 #V2 #T1 #T2 #l #m #_ #_ #i #H destruct
+| #I #V1 #V2 #T1 #T2 #l #m #_ #_ #i #H destruct
]
qed-.
(* Basic_1: was: lift_gen_lref *)
-lemma lift_inv_lref2: ∀d,e,T1,i. ⬆[d,e] T1 ≡ #i →
- (i < d ∧ T1 = #i) ∨ (d + e ≤ i ∧ T1 = #(i - e)).
+lemma lift_inv_lref2: ∀l,m,T1,i. ⬆[l,m] T1 ≡ #i →
+ (i < l ∧ T1 = #i) ∨ (l + m ≤ i ∧ T1 = #(i - m)).
/2 width=3 by lift_inv_lref2_aux/ qed-.
(* Basic_1: was: lift_gen_lref_lt *)
-lemma lift_inv_lref2_lt: ∀d,e,T1,i. ⬆[d,e] T1 ≡ #i → i < d → T1 = #i.
-#d #e #T1 #i #H elim (lift_inv_lref2 … H) -H * //
-#Hdi #_ #Hid lapply (le_to_lt_to_lt … Hdi Hid) -Hdi -Hid #Hdd
-elim (lt_inv_plus_l … Hdd) -Hdd #Hdd
-elim (lt_refl_false … Hdd)
+lemma lift_inv_lref2_lt: ∀l,m,T1,i. ⬆[l,m] T1 ≡ #i → i < l → T1 = #i.
+#l #m #T1 #i #H elim (lift_inv_lref2 … H) -H * //
+#Hli #_ #Hil lapply (le_to_lt_to_lt … Hli Hil) -Hli -Hil #Hll
+elim (lt_inv_plus_l … Hll) -Hll #Hll
+elim (lt_refl_false … Hll)
qed-.
(* Basic_1: was: lift_gen_lref_false *)
-lemma lift_inv_lref2_be: ∀d,e,T1,i. ⬆[d,e] T1 ≡ #i →
- d ≤ i → i < d + e → ⊥.
-#d #e #T1 #i #H elim (lift_inv_lref2 … H) -H *
+lemma lift_inv_lref2_be: ∀l,m,T1,i. ⬆[l,m] T1 ≡ #i →
+ l ≤ i → i < l + m → ⊥.
+#l #m #T1 #i #H elim (lift_inv_lref2 … H) -H *
[ #H1 #_ #H2 #_ | #H2 #_ #_ #H1 ]
lapply (le_to_lt_to_lt … H2 H1) -H2 -H1 #H
elim (lt_refl_false … H)
qed-.
(* Basic_1: was: lift_gen_lref_ge *)
-lemma lift_inv_lref2_ge: ∀d,e,T1,i. ⬆[d,e] T1 ≡ #i → d + e ≤ i → T1 = #(i - e).
-#d #e #T1 #i #H elim (lift_inv_lref2 … H) -H * //
-#Hid #_ #Hdi lapply (le_to_lt_to_lt … Hdi Hid) -Hdi -Hid #Hdd
-elim (lt_inv_plus_l … Hdd) -Hdd #Hdd
-elim (lt_refl_false … Hdd)
+lemma lift_inv_lref2_ge: ∀l,m,T1,i. ⬆[l,m] T1 ≡ #i → l + m ≤ i → T1 = #(i - m).
+#l #m #T1 #i #H elim (lift_inv_lref2 … H) -H * //
+#Hil #_ #Hli lapply (le_to_lt_to_lt … Hli Hil) -Hli -Hil #Hll
+elim (lt_inv_plus_l … Hll) -Hll #Hll
+elim (lt_refl_false … Hll)
qed-.
-fact lift_inv_gref2_aux: ∀d,e,T1,T2. ⬆[d,e] T1 ≡ T2 → ∀p. T2 = §p → T1 = §p.
-#d #e #T1 #T2 * -d -e -T1 -T2 //
-[ #i #d #e #_ #k #H destruct
-| #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct
-| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct
+fact lift_inv_gref2_aux: ∀l,m,T1,T2. ⬆[l,m] T1 ≡ T2 → ∀p. T2 = §p → T1 = §p.
+#l #m #T1 #T2 * -l -m -T1 -T2 //
+[ #i #l #m #_ #k #H destruct
+| #a #I #V1 #V2 #T1 #T2 #l #m #_ #_ #k #H destruct
+| #I #V1 #V2 #T1 #T2 #l #m #_ #_ #k #H destruct
]
qed-.
-lemma lift_inv_gref2: ∀d,e,T1,p. ⬆[d,e] T1 ≡ §p → T1 = §p.
+lemma lift_inv_gref2: ∀l,m,T1,p. ⬆[l,m] T1 ≡ §p → T1 = §p.
/2 width=5 by lift_inv_gref2_aux/ qed-.
-fact lift_inv_bind2_aux: ∀d,e,T1,T2. ⬆[d,e] T1 ≡ T2 →
+fact lift_inv_bind2_aux: ∀l,m,T1,T2. ⬆[l,m] T1 ≡ T2 →
∀a,I,V2,U2. T2 = ⓑ{a,I} V2.U2 →
- ∃∃V1,U1. ⬆[d,e] V1 ≡ V2 & ⬆[d+1,e] U1 ≡ U2 &
+ ∃∃V1,U1. ⬆[l,m] V1 ≡ V2 & ⬆[l+1,m] U1 ≡ U2 &
T1 = ⓑ{a,I} V1. U1.
-#d #e #T1 #T2 * -d -e -T1 -T2
-[ #k #d #e #a #I #V2 #U2 #H destruct
-| #i #d #e #_ #a #I #V2 #U2 #H destruct
-| #i #d #e #_ #a #I #V2 #U2 #H destruct
-| #p #d #e #a #I #V2 #U2 #H destruct
-| #b #J #W1 #W2 #T1 #T2 #d #e #HW #HT #a #I #V2 #U2 #H destruct /2 width=5 by ex3_2_intro/
-| #J #W1 #W2 #T1 #T2 #d #e #_ #_ #a #I #V2 #U2 #H destruct
+#l #m #T1 #T2 * -l -m -T1 -T2
+[ #k #l #m #a #I #V2 #U2 #H destruct
+| #i #l #m #_ #a #I #V2 #U2 #H destruct
+| #i #l #m #_ #a #I #V2 #U2 #H destruct
+| #p #l #m #a #I #V2 #U2 #H destruct
+| #b #J #W1 #W2 #T1 #T2 #l #m #HW #HT #a #I #V2 #U2 #H destruct /2 width=5 by ex3_2_intro/
+| #J #W1 #W2 #T1 #T2 #l #m #_ #_ #a #I #V2 #U2 #H destruct
]
qed-.
(* Basic_1: was: lift_gen_bind *)
-lemma lift_inv_bind2: ∀d,e,T1,a,I,V2,U2. ⬆[d,e] T1 ≡ ⓑ{a,I} V2. U2 →
- ∃∃V1,U1. ⬆[d,e] V1 ≡ V2 & ⬆[d+1,e] U1 ≡ U2 &
+lemma lift_inv_bind2: ∀l,m,T1,a,I,V2,U2. ⬆[l,m] T1 ≡ ⓑ{a,I} V2. U2 →
+ ∃∃V1,U1. ⬆[l,m] V1 ≡ V2 & ⬆[l+1,m] U1 ≡ U2 &
T1 = ⓑ{a,I} V1. U1.
/2 width=3 by lift_inv_bind2_aux/ qed-.
-fact lift_inv_flat2_aux: ∀d,e,T1,T2. ⬆[d,e] T1 ≡ T2 →
+fact lift_inv_flat2_aux: ∀l,m,T1,T2. ⬆[l,m] T1 ≡ T2 →
∀I,V2,U2. T2 = ⓕ{I} V2.U2 →
- ∃∃V1,U1. ⬆[d,e] V1 ≡ V2 & ⬆[d,e] U1 ≡ U2 &
+ ∃∃V1,U1. ⬆[l,m] V1 ≡ V2 & ⬆[l,m] U1 ≡ U2 &
T1 = ⓕ{I} V1. U1.
-#d #e #T1 #T2 * -d -e -T1 -T2
-[ #k #d #e #I #V2 #U2 #H destruct
-| #i #d #e #_ #I #V2 #U2 #H destruct
-| #i #d #e #_ #I #V2 #U2 #H destruct
-| #p #d #e #I #V2 #U2 #H destruct
-| #a #J #W1 #W2 #T1 #T2 #d #e #_ #_ #I #V2 #U2 #H destruct
-| #J #W1 #W2 #T1 #T2 #d #e #HW #HT #I #V2 #U2 #H destruct /2 width=5 by ex3_2_intro/
+#l #m #T1 #T2 * -l -m -T1 -T2
+[ #k #l #m #I #V2 #U2 #H destruct
+| #i #l #m #_ #I #V2 #U2 #H destruct
+| #i #l #m #_ #I #V2 #U2 #H destruct
+| #p #l #m #I #V2 #U2 #H destruct
+| #a #J #W1 #W2 #T1 #T2 #l #m #_ #_ #I #V2 #U2 #H destruct
+| #J #W1 #W2 #T1 #T2 #l #m #HW #HT #I #V2 #U2 #H destruct /2 width=5 by ex3_2_intro/
]
qed-.
(* Basic_1: was: lift_gen_flat *)
-lemma lift_inv_flat2: ∀d,e,T1,I,V2,U2. ⬆[d,e] T1 ≡ ⓕ{I} V2. U2 →
- ∃∃V1,U1. ⬆[d,e] V1 ≡ V2 & ⬆[d,e] U1 ≡ U2 &
+lemma lift_inv_flat2: ∀l,m,T1,I,V2,U2. ⬆[l,m] T1 ≡ ⓕ{I} V2. U2 →
+ ∃∃V1,U1. ⬆[l,m] V1 ≡ V2 & ⬆[l,m] U1 ≡ U2 &
T1 = ⓕ{I} V1. U1.
/2 width=3 by lift_inv_flat2_aux/ qed-.
-lemma lift_inv_pair_xy_x: ∀d,e,I,V,T. ⬆[d, e] ②{I} V. T ≡ V → ⊥.
-#d #e #J #V elim V -V
+lemma lift_inv_pair_xy_x: ∀l,m,I,V,T. ⬆[l, m] ②{I} V. T ≡ V → ⊥.
+#l #m #J #V elim V -V
[ * #i #T #H
[ lapply (lift_inv_sort2 … H) -H #H destruct
| elim (lift_inv_lref2 … H) -H * #_ #H destruct
qed-.
(* Basic_1: was: thead_x_lift_y_y *)
-lemma lift_inv_pair_xy_y: ∀I,T,V,d,e. ⬆[d, e] ②{I} V. T ≡ T → ⊥.
+lemma lift_inv_pair_xy_y: ∀I,T,V,l,m. ⬆[l, m] ②{I} V. T ≡ T → ⊥.
#J #T elim T -T
-[ * #i #V #d #e #H
+[ * #i #V #l #m #H
[ lapply (lift_inv_sort2 … H) -H #H destruct
| elim (lift_inv_lref2 … H) -H * #_ #H destruct
| lapply (lift_inv_gref2 … H) -H #H destruct
]
-| * [ #a ] #I #W2 #U2 #_ #IHU2 #V #d #e #H
+| * [ #a ] #I #W2 #U2 #_ #IHU2 #V #l #m #H
[ elim (lift_inv_bind2 … H) -H #W1 #U1 #_ #HU12 #H destruct /2 width=4 by/
| elim (lift_inv_flat2 … H) -H #W1 #U1 #_ #HU12 #H destruct /2 width=4 by/
]
(* Basic forward lemmas *****************************************************)
-lemma lift_fwd_pair1: ∀I,T2,V1,U1,d,e. ⬆[d,e] ②{I}V1.U1 ≡ T2 →
- ∃∃V2,U2. ⬆[d,e] V1 ≡ V2 & T2 = ②{I}V2.U2.
-* [ #a ] #I #T2 #V1 #U1 #d #e #H
+lemma lift_fwd_pair1: ∀I,T2,V1,U1,l,m. ⬆[l,m] ②{I}V1.U1 ≡ T2 →
+ ∃∃V2,U2. ⬆[l,m] V1 ≡ V2 & T2 = ②{I}V2.U2.
+* [ #a ] #I #T2 #V1 #U1 #l #m #H
[ elim (lift_inv_bind1 … H) -H /2 width=4 by ex2_2_intro/
| elim (lift_inv_flat1 … H) -H /2 width=4 by ex2_2_intro/
]
qed-.
-lemma lift_fwd_pair2: ∀I,T1,V2,U2,d,e. ⬆[d,e] T1 ≡ ②{I}V2.U2 →
- ∃∃V1,U1. ⬆[d,e] V1 ≡ V2 & T1 = ②{I}V1.U1.
-* [ #a ] #I #T1 #V2 #U2 #d #e #H
+lemma lift_fwd_pair2: ∀I,T1,V2,U2,l,m. ⬆[l,m] T1 ≡ ②{I}V2.U2 →
+ ∃∃V1,U1. ⬆[l,m] V1 ≡ V2 & T1 = ②{I}V1.U1.
+* [ #a ] #I #T1 #V2 #U2 #l #m #H
[ elim (lift_inv_bind2 … H) -H /2 width=4 by ex2_2_intro/
| elim (lift_inv_flat2 … H) -H /2 width=4 by ex2_2_intro/
]
qed-.
-lemma lift_fwd_tw: ∀d,e,T1,T2. ⬆[d, e] T1 ≡ T2 → ♯{T1} = ♯{T2}.
-#d #e #T1 #T2 #H elim H -d -e -T1 -T2 normalize //
+lemma lift_fwd_tw: ∀l,m,T1,T2. ⬆[l, m] T1 ≡ T2 → ♯{T1} = ♯{T2}.
+#l #m #T1 #T2 #H elim H -l -m -T1 -T2 normalize //
qed-.
-lemma lift_simple_dx: ∀d,e,T1,T2. ⬆[d, e] T1 ≡ T2 → 𝐒⦃T1⦄ → 𝐒⦃T2⦄.
-#d #e #T1 #T2 #H elim H -d -e -T1 -T2 //
-#a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #_ #_ #H
+lemma lift_simple_dx: ∀l,m,T1,T2. ⬆[l, m] T1 ≡ T2 → 𝐒⦃T1⦄ → 𝐒⦃T2⦄.
+#l #m #T1 #T2 #H elim H -l -m -T1 -T2 //
+#a #I #V1 #V2 #T1 #T2 #l #m #_ #_ #_ #_ #H
elim (simple_inv_bind … H)
qed-.
-lemma lift_simple_sn: ∀d,e,T1,T2. ⬆[d, e] T1 ≡ T2 → 𝐒⦃T2⦄ → 𝐒⦃T1⦄.
-#d #e #T1 #T2 #H elim H -d -e -T1 -T2 //
-#a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #_ #_ #H
+lemma lift_simple_sn: ∀l,m,T1,T2. ⬆[l, m] T1 ≡ T2 → 𝐒⦃T2⦄ → 𝐒⦃T1⦄.
+#l #m #T1 #T2 #H elim H -l -m -T1 -T2 //
+#a #I #V1 #V2 #T1 #T2 #l #m #_ #_ #_ #_ #H
elim (simple_inv_bind … H)
qed-.
(* Basic properties *********************************************************)
(* Basic_1: was: lift_lref_gt *)
-lemma lift_lref_ge_minus: ∀d,e,i. d + e ≤ i → ⬆[d, e] #(i - e) ≡ #i.
-#d #e #i #H >(plus_minus_m_m i e) in ⊢ (? ? ? ? %); /3 width=2 by lift_lref_ge, le_plus_to_minus_r, le_plus_b/
+lemma lift_lref_ge_minus: ∀l,m,i. l + m ≤ i → ⬆[l, m] #(i - m) ≡ #i.
+#l #m #i #H >(plus_minus_m_m i m) in ⊢ (? ? ? ? %); /3 width=2 by lift_lref_ge, le_plus_to_minus_r, le_plus_b/
qed.
-lemma lift_lref_ge_minus_eq: ∀d,e,i,j. d + e ≤ i → j = i - e → ⬆[d, e] #j ≡ #i.
+lemma lift_lref_ge_minus_eq: ∀l,m,i,j. l + m ≤ i → j = i - m → ⬆[l, m] #j ≡ #i.
/2 width=1/ qed-.
(* Basic_1: was: lift_r *)
-lemma lift_refl: ∀T,d. ⬆[d, 0] T ≡ T.
+lemma lift_refl: ∀T,l. ⬆[l, 0] T ≡ T.
#T elim T -T
-[ * #i // #d elim (lt_or_ge i d) /2 width=1 by lift_lref_lt, lift_lref_ge/
+[ * #i // #l elim (lt_or_ge i l) /2 width=1 by lift_lref_lt, lift_lref_ge/
| * /2 width=1 by lift_bind, lift_flat/
]
qed.
-lemma lift_total: ∀T1,d,e. ∃T2. ⬆[d,e] T1 ≡ T2.
+lemma lift_total: ∀T1,l,m. ∃T2. ⬆[l,m] T1 ≡ T2.
#T1 elim T1 -T1
-[ * #i /2 width=2/ #d #e elim (lt_or_ge i d) /3 width=2 by lift_lref_lt, lift_lref_ge, ex_intro/
-| * [ #a ] #I #V1 #T1 #IHV1 #IHT1 #d #e
- elim (IHV1 d e) -IHV1 #V2 #HV12
- [ elim (IHT1 (d+1) e) -IHT1 /3 width=2 by lift_bind, ex_intro/
- | elim (IHT1 d e) -IHT1 /3 width=2 by lift_flat, ex_intro/
+[ * #i /2 width=2/ #l #m elim (lt_or_ge i l) /3 width=2 by lift_lref_lt, lift_lref_ge, ex_intro/
+| * [ #a ] #I #V1 #T1 #IHV1 #IHT1 #l #m
+ elim (IHV1 l m) -IHV1 #V2 #HV12
+ [ elim (IHT1 (l+1) m) -IHT1 /3 width=2 by lift_bind, ex_intro/
+ | elim (IHT1 l m) -IHT1 /3 width=2 by lift_flat, ex_intro/
]
]
qed.
(* Basic_1: was: lift_free (right to left) *)
-lemma lift_split: ∀d1,e2,T1,T2. ⬆[d1, e2] T1 ≡ T2 →
- ∀d2,e1. d1 ≤ d2 → d2 ≤ d1 + e1 → e1 ≤ e2 →
- ∃∃T. ⬆[d1, e1] T1 ≡ T & ⬆[d2, e2 - e1] T ≡ T2.
-#d1 #e2 #T1 #T2 #H elim H -d1 -e2 -T1 -T2
+lemma lift_split: ∀l1,m2,T1,T2. ⬆[l1, m2] T1 ≡ T2 →
+ ∀l2,m1. l1 ≤ l2 → l2 ≤ l1 + m1 → m1 ≤ m2 →
+ ∃∃T. ⬆[l1, m1] T1 ≡ T & ⬆[l2, m2 - m1] T ≡ T2.
+#l1 #m2 #T1 #T2 #H elim H -l1 -m2 -T1 -T2
[ /3 width=3/
-| #i #d1 #e2 #Hid1 #d2 #e1 #Hd12 #_ #_
- lapply (lt_to_le_to_lt … Hid1 Hd12) -Hd12 #Hid2 /4 width=3 by lift_lref_lt, ex2_intro/
-| #i #d1 #e2 #Hid1 #d2 #e1 #_ #Hd21 #He12
- lapply (transitive_le … (i+e1) Hd21 ?) /2 width=1 by monotonic_le_plus_l/ -Hd21 #Hd21
- >(plus_minus_m_m e2 e1 ?) /3 width=3 by lift_lref_ge, ex2_intro/
+| #i #l1 #m2 #Hil1 #l2 #m1 #Hl12 #_ #_
+ lapply (lt_to_le_to_lt … Hil1 Hl12) -Hl12 #Hil2 /4 width=3 by lift_lref_lt, ex2_intro/
+| #i #l1 #m2 #Hil1 #l2 #m1 #_ #Hl21 #Hm12
+ lapply (transitive_le … (i+m1) Hl21 ?) /2 width=1 by monotonic_le_plus_l/ -Hl21 #Hl21
+ >(plus_minus_m_m m2 m1 ?) /3 width=3 by lift_lref_ge, ex2_intro/
| /3 width=3/
-| #a #I #V1 #V2 #T1 #T2 #d1 #e2 #_ #_ #IHV #IHT #d2 #e1 #Hd12 #Hd21 #He12
- elim (IHV … Hd12 Hd21 He12) -IHV #V0 #HV0a #HV0b
- elim (IHT (d2+1) … ? ? He12) /3 width=5 by lift_bind, le_S_S, ex2_intro/
-| #I #V1 #V2 #T1 #T2 #d1 #e2 #_ #_ #IHV #IHT #d2 #e1 #Hd12 #Hd21 #He12
- elim (IHV … Hd12 Hd21 He12) -IHV #V0 #HV0a #HV0b
- elim (IHT d2 … ? ? He12) /3 width=5 by lift_flat, ex2_intro/
+| #a #I #V1 #V2 #T1 #T2 #l1 #m2 #_ #_ #IHV #IHT #l2 #m1 #Hl12 #Hl21 #Hm12
+ elim (IHV … Hl12 Hl21 Hm12) -IHV #V0 #HV0a #HV0b
+ elim (IHT (l2+1) … ? ? Hm12) /3 width=5 by lift_bind, le_S_S, ex2_intro/
+| #I #V1 #V2 #T1 #T2 #l1 #m2 #_ #_ #IHV #IHT #l2 #m1 #Hl12 #Hl21 #Hm12
+ elim (IHV … Hl12 Hl21 Hm12) -IHV #V0 #HV0a #HV0b
+ elim (IHT l2 … ? ? Hm12) /3 width=5 by lift_flat, ex2_intro/
]
qed.
(* Basic_1: was only: dnf_dec2 dnf_dec *)
-lemma is_lift_dec: ∀T2,d,e. Decidable (∃T1. ⬆[d,e] T1 ≡ T2).
+lemma is_lift_dec: ∀T2,l,m. Decidable (∃T1. ⬆[l,m] T1 ≡ T2).
#T1 elim T1 -T1
-[ * [1,3: /3 width=2 by lift_sort, lift_gref, ex_intro, or_introl/ ] #i #d #e
- elim (lt_or_ge i d) #Hdi
+[ * [1,3: /3 width=2 by lift_sort, lift_gref, ex_intro, or_introl/ ] #i #l #m
+ elim (lt_or_ge i l) #Hli
[ /4 width=3 by lift_lref_lt, ex_intro, or_introl/
- | elim (lt_or_ge i (d + e)) #Hide
- [ @or_intror * #T1 #H elim (lift_inv_lref2_be … H Hdi Hide)
- | -Hdi /4 width=2 by lift_lref_ge_minus, ex_intro, or_introl/
+ | elim (lt_or_ge i (l + m)) #Hilm
+ [ @or_intror * #T1 #H elim (lift_inv_lref2_be … H Hli Hilm)
+ | -Hli /4 width=2 by lift_lref_ge_minus, ex_intro, or_introl/
]
]
-| * [ #a ] #I #V2 #T2 #IHV2 #IHT2 #d #e
- [ elim (IHV2 d e) -IHV2
- [ * #V1 #HV12 elim (IHT2 (d+1) e) -IHT2
+| * [ #a ] #I #V2 #T2 #IHV2 #IHT2 #l #m
+ [ elim (IHV2 l m) -IHV2
+ [ * #V1 #HV12 elim (IHT2 (l+1) m) -IHT2
[ * #T1 #HT12 @or_introl /3 width=2 by lift_bind, ex_intro/
| -V1 #HT2 @or_intror * #X #H
elim (lift_inv_bind2 … H) -H /3 width=2 by ex_intro/
| -IHT2 #HV2 @or_intror * #X #H
elim (lift_inv_bind2 … H) -H /3 width=2 by ex_intro/
]
- | elim (IHV2 d e) -IHV2
- [ * #V1 #HV12 elim (IHT2 d e) -IHT2
+ | elim (IHV2 l m) -IHV2
+ [ * #V1 #HV12 elim (IHT2 l m) -IHT2
[ * #T1 #HT12 /4 width=2/
| -V1 #HT2 @or_intror * #X #H
elim (lift_inv_flat2 … H) -H /3 width=2 by ex_intro/
(* Main properties ***********************************************************)
(* Basic_1: was: lift_inj *)
-theorem lift_inj: ∀d,e,T1,U. ⬆[d,e] T1 ≡ U → ∀T2. ⬆[d,e] T2 ≡ U → T1 = T2.
-#d #e #T1 #U #H elim H -d -e -T1 -U
-[ #k #d #e #X #HX
+theorem lift_inj: ∀l,m,T1,U. ⬆[l,m] T1 ≡ U → ∀T2. ⬆[l,m] T2 ≡ U → T1 = T2.
+#l #m #T1 #U #H elim H -l -m -T1 -U
+[ #k #l #m #X #HX
lapply (lift_inv_sort2 … HX) -HX //
-| #i #d #e #Hid #X #HX
+| #i #l #m #Hil #X #HX
lapply (lift_inv_lref2_lt … HX ?) -HX //
-| #i #d #e #Hdi #X #HX
+| #i #l #m #Hli #X #HX
lapply (lift_inv_lref2_ge … HX ?) -HX /2 width=1 by monotonic_le_plus_l/
-| #p #d #e #X #HX
+| #p #l #m #X #HX
lapply (lift_inv_gref2 … HX) -HX //
-| #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
+| #a #I #V1 #V2 #T1 #T2 #l #m #_ #_ #IHV12 #IHT12 #X #HX
elim (lift_inv_bind2 … HX) -HX #V #T #HV1 #HT1 #HX destruct /3 width=1 by eq_f2/
-| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
+| #I #V1 #V2 #T1 #T2 #l #m #_ #_ #IHV12 #IHT12 #X #HX
elim (lift_inv_flat2 … HX) -HX #V #T #HV1 #HT1 #HX destruct /3 width=1 by eq_f2/
]
qed-.
(* Basic_1: was: lift_gen_lift *)
-theorem lift_div_le: ∀d1,e1,T1,T. ⬆[d1, e1] T1 ≡ T →
- ∀d2,e2,T2. ⬆[d2 + e1, e2] T2 ≡ T →
- d1 ≤ d2 →
- ∃∃T0. ⬆[d1, e1] T0 ≡ T2 & ⬆[d2, e2] T0 ≡ T1.
-#d1 #e1 #T1 #T #H elim H -d1 -e1 -T1 -T
-[ #k #d1 #e1 #d2 #e2 #T2 #Hk #Hd12
+theorem lift_div_le: ∀l1,m1,T1,T. ⬆[l1, m1] T1 ≡ T →
+ ∀l2,m2,T2. ⬆[l2 + m1, m2] T2 ≡ T →
+ l1 ≤ l2 →
+ ∃∃T0. ⬆[l1, m1] T0 ≡ T2 & ⬆[l2, m2] T0 ≡ T1.
+#l1 #m1 #T1 #T #H elim H -l1 -m1 -T1 -T
+[ #k #l1 #m1 #l2 #m2 #T2 #Hk #Hl12
lapply (lift_inv_sort2 … Hk) -Hk #Hk destruct /3 width=3 by lift_sort, ex2_intro/
-| #i #d1 #e1 #Hid1 #d2 #e2 #T2 #Hi #Hd12
- lapply (lt_to_le_to_lt … Hid1 Hd12) -Hd12 #Hid2
+| #i #l1 #m1 #Hil1 #l2 #m2 #T2 #Hi #Hl12
+ lapply (lt_to_le_to_lt … Hil1 Hl12) -Hl12 #Hil2
lapply (lift_inv_lref2_lt … Hi ?) -Hi /3 width=3 by lift_lref_lt, lt_plus_to_minus_r, lt_to_le_to_lt, ex2_intro/
-| #i #d1 #e1 #Hid1 #d2 #e2 #T2 #Hi #Hd12
- elim (lift_inv_lref2 … Hi) -Hi * #Hid2 #H destruct
- [ -Hd12 lapply (lt_plus_to_lt_l … Hid2) -Hid2 #Hid2 /3 width=3 by lift_lref_lt, lift_lref_ge, ex2_intro/
- | -Hid1 >plus_plus_comm_23 in Hid2; #H lapply (le_plus_to_le_r … H) -H #H
- elim (le_inv_plus_l … H) -H #Hide2 #He2i
- lapply (transitive_le … Hd12 Hide2) -Hd12 #Hd12
- >le_plus_minus_comm // >(plus_minus_m_m i e2) in ⊢ (? ? ? %);
+| #i #l1 #m1 #Hil1 #l2 #m2 #T2 #Hi #Hl12
+ elim (lift_inv_lref2 … Hi) -Hi * #Hil2 #H destruct
+ [ -Hl12 lapply (lt_plus_to_lt_l … Hil2) -Hil2 #Hil2 /3 width=3 by lift_lref_lt, lift_lref_ge, ex2_intro/
+ | -Hil1 >plus_plus_comm_23 in Hil2; #H lapply (le_plus_to_le_r … H) -H #H
+ elim (le_inv_plus_l … H) -H #Hilm2 #Hm2i
+ lapply (transitive_le … Hl12 Hilm2) -Hl12 #Hl12
+ >le_plus_minus_comm // >(plus_minus_m_m i m2) in ⊢ (? ? ? %);
/4 width=3 by lift_lref_ge, ex2_intro/
]
-| #p #d1 #e1 #d2 #e2 #T2 #Hk #Hd12
+| #p #l1 #m1 #l2 #m2 #T2 #Hk #Hl12
lapply (lift_inv_gref2 … Hk) -Hk #Hk destruct /3 width=3 by lift_gref, ex2_intro/
-| #a #I #W1 #W #U1 #U #d1 #e1 #_ #_ #IHW #IHU #d2 #e2 #T2 #H #Hd12
+| #a #I #W1 #W #U1 #U #l1 #m1 #_ #_ #IHW #IHU #l2 #m2 #T2 #H #Hl12
lapply (lift_inv_bind2 … H) -H * #W2 #U2 #HW2 #HU2 #H destruct
elim (IHW … HW2) // -IHW -HW2 #W0 #HW2 #HW1
>plus_plus_comm_23 in HU2; #HU2 elim (IHU … HU2) /3 width=5 by lift_bind, le_S_S, ex2_intro/
-| #I #W1 #W #U1 #U #d1 #e1 #_ #_ #IHW #IHU #d2 #e2 #T2 #H #Hd12
+| #I #W1 #W #U1 #U #l1 #m1 #_ #_ #IHW #IHU #l2 #m2 #T2 #H #Hl12
lapply (lift_inv_flat2 … H) -H * #W2 #U2 #HW2 #HU2 #H destruct
elim (IHW … HW2) // -IHW -HW2 #W0 #HW2 #HW1
elim (IHU … HU2) /3 width=5 by lift_flat, ex2_intro/
qed.
(* Note: apparently this was missing in basic_1 *)
-theorem lift_div_be: ∀d1,e1,T1,T. ⬆[d1, e1] T1 ≡ T →
- ∀e,e2,T2. ⬆[d1 + e, e2] T2 ≡ T →
- e ≤ e1 → e1 ≤ e + e2 →
- ∃∃T0. ⬆[d1, e] T0 ≡ T2 & ⬆[d1, e + e2 - e1] T0 ≡ T1.
-#d1 #e1 #T1 #T #H elim H -d1 -e1 -T1 -T
-[ #k #d1 #e1 #e #e2 #T2 #H >(lift_inv_sort2 … H) -H /2 width=3 by lift_sort, ex2_intro/
-| #i #d1 #e1 #Hid1 #e #e2 #T2 #H #He1 #He1e2
+theorem lift_div_be: ∀l1,m1,T1,T. ⬆[l1, m1] T1 ≡ T →
+ ∀m,m2,T2. ⬆[l1 + m, m2] T2 ≡ T →
+ m ≤ m1 → m1 ≤ m + m2 →
+ ∃∃T0. ⬆[l1, m] T0 ≡ T2 & ⬆[l1, m + m2 - m1] T0 ≡ T1.
+#l1 #m1 #T1 #T #H elim H -l1 -m1 -T1 -T
+[ #k #l1 #m1 #m #m2 #T2 #H >(lift_inv_sort2 … H) -H /2 width=3 by lift_sort, ex2_intro/
+| #i #l1 #m1 #Hil1 #m #m2 #T2 #H #Hm1 #Hm1m2
>(lift_inv_lref2_lt … H) -H /3 width=3 by lift_lref_lt, lt_plus_to_minus_r, lt_to_le_to_lt, ex2_intro/
-| #i #d1 #e1 #Hid1 #e #e2 #T2 #H #He1 #He1e2
- elim (lt_or_ge (i+e1) (d1+e+e2)) #Hie1d1e2
+| #i #l1 #m1 #Hil1 #m #m2 #T2 #H #Hm1 #Hm1m2
+ elim (lt_or_ge (i+m1) (l1+m+m2)) #Him1l1m2
[ elim (lift_inv_lref2_be … H) -H /2 width=1 by le_plus/
| >(lift_inv_lref2_ge … H ?) -H //
- lapply (le_plus_to_minus … Hie1d1e2) #Hd1e21i
- elim (le_inv_plus_l … Hie1d1e2) -Hie1d1e2 #Hd1e12 #He2ie1
- @ex2_intro [2: /2 width=1/ | skip ] -Hd1e12
+ lapply (le_plus_to_minus … Him1l1m2) #Hl1m21i
+ elim (le_inv_plus_l … Him1l1m2) -Him1l1m2 #Hl1m12 #Hm2im1
+ @ex2_intro [2: /2 width=1 by lift_lref_ge_minus/ | skip ] -Hl1m12
@lift_lref_ge_minus_eq [ >plus_minus_associative // | /2 width=1 by minus_le_minus_minus_comm/ ]
]
-| #p #d1 #e1 #e #e2 #T2 #H >(lift_inv_gref2 … H) -H /2 width=3 by lift_gref, ex2_intro/
-| #a #I #V1 #V #T1 #T #d1 #e1 #_ #_ #IHV1 #IHT1 #e #e2 #X #H #He1 #He1e2
+| #p #l1 #m1 #m #m2 #T2 #H >(lift_inv_gref2 … H) -H /2 width=3 by lift_gref, ex2_intro/
+| #a #I #V1 #V #T1 #T #l1 #m1 #_ #_ #IHV1 #IHT1 #m #m2 #X #H #Hm1 #Hm1m2
elim (lift_inv_bind2 … H) -H #V2 #T2 #HV2 #HT2 #H destruct
elim (IHV1 … HV2) -V // >plus_plus_comm_23 in HT2; #HT2
elim (IHT1 … HT2) -T /3 width=5 by lift_bind, ex2_intro/
-| #I #V1 #V #T1 #T #d1 #e1 #_ #_ #IHV1 #IHT1 #e #e2 #X #H #He1 #He1e2
+| #I #V1 #V #T1 #T #l1 #m1 #_ #_ #IHV1 #IHT1 #m #m2 #X #H #Hm1 #Hm1m2
elim (lift_inv_flat2 … H) -H #V2 #T2 #HV2 #HT2 #H destruct
elim (IHV1 … HV2) -V //
elim (IHT1 … HT2) -T /3 width=5 by lift_flat, ex2_intro/
]
qed.
-theorem lift_mono: ∀d,e,T,U1. ⬆[d,e] T ≡ U1 → ∀U2. ⬆[d,e] T ≡ U2 → U1 = U2.
-#d #e #T #U1 #H elim H -d -e -T -U1
-[ #k #d #e #X #HX
+theorem lift_mono: ∀l,m,T,U1. ⬆[l,m] T ≡ U1 → ∀U2. ⬆[l,m] T ≡ U2 → U1 = U2.
+#l #m #T #U1 #H elim H -l -m -T -U1
+[ #k #l #m #X #HX
lapply (lift_inv_sort1 … HX) -HX //
-| #i #d #e #Hid #X #HX
+| #i #l #m #Hil #X #HX
lapply (lift_inv_lref1_lt … HX ?) -HX //
-| #i #d #e #Hdi #X #HX
+| #i #l #m #Hli #X #HX
lapply (lift_inv_lref1_ge … HX ?) -HX //
-| #p #d #e #X #HX
+| #p #l #m #X #HX
lapply (lift_inv_gref1 … HX) -HX //
-| #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
+| #a #I #V1 #V2 #T1 #T2 #l #m #_ #_ #IHV12 #IHT12 #X #HX
elim (lift_inv_bind1 … HX) -HX #V #T #HV1 #HT1 #HX destruct /3 width=1 by eq_f2/
-| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
+| #I #V1 #V2 #T1 #T2 #l #m #_ #_ #IHV12 #IHT12 #X #HX
elim (lift_inv_flat1 … HX) -HX #V #T #HV1 #HT1 #HX destruct /3 width=1 by eq_f2/
]
qed-.
(* Basic_1: was: lift_free (left to right) *)
-theorem lift_trans_be: ∀d1,e1,T1,T. ⬆[d1, e1] T1 ≡ T →
- ∀d2,e2,T2. ⬆[d2, e2] T ≡ T2 →
- d1 ≤ d2 → d2 ≤ d1 + e1 → ⬆[d1, e1 + e2] T1 ≡ T2.
-#d1 #e1 #T1 #T #H elim H -d1 -e1 -T1 -T
-[ #k #d1 #e1 #d2 #e2 #T2 #HT2 #_ #_
+theorem lift_trans_be: ∀l1,m1,T1,T. ⬆[l1, m1] T1 ≡ T →
+ ∀l2,m2,T2. ⬆[l2, m2] T ≡ T2 →
+ l1 ≤ l2 → l2 ≤ l1 + m1 → ⬆[l1, m1 + m2] T1 ≡ T2.
+#l1 #m1 #T1 #T #H elim H -l1 -m1 -T1 -T
+[ #k #l1 #m1 #l2 #m2 #T2 #HT2 #_ #_
>(lift_inv_sort1 … HT2) -HT2 //
-| #i #d1 #e1 #Hid1 #d2 #e2 #T2 #HT2 #Hd12 #_
- lapply (lt_to_le_to_lt … Hid1 Hd12) -Hd12 #Hid2
- lapply (lift_inv_lref1_lt … HT2 Hid2) /2 width=1 by lift_lref_lt/
-| #i #d1 #e1 #Hid1 #d2 #e2 #T2 #HT2 #_ #Hd21
+| #i #l1 #m1 #Hil1 #l2 #m2 #T2 #HT2 #Hl12 #_
+ lapply (lt_to_le_to_lt … Hil1 Hl12) -Hl12 #Hil2
+ lapply (lift_inv_lref1_lt … HT2 Hil2) /2 width=1 by lift_lref_lt/
+| #i #l1 #m1 #Hil1 #l2 #m2 #T2 #HT2 #_ #Hl21
lapply (lift_inv_lref1_ge … HT2 ?) -HT2
- [ @(transitive_le … Hd21 ?) -Hd21 /2 width=1 by monotonic_le_plus_l/
- | -Hd21 /2 width=1 by lift_lref_ge/
+ [ @(transitive_le … Hl21 ?) -Hl21 /2 width=1 by monotonic_le_plus_l/
+ | -Hl21 /2 width=1 by lift_lref_ge/
]
-| #p #d1 #e1 #d2 #e2 #T2 #HT2 #_ #_
+| #p #l1 #m1 #l2 #m2 #T2 #HT2 #_ #_
>(lift_inv_gref1 … HT2) -HT2 //
-| #a #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hd12 #Hd21
+| #a #I #V1 #V2 #T1 #T2 #l1 #m1 #_ #_ #IHV12 #IHT12 #l2 #m2 #X #HX #Hl12 #Hl21
elim (lift_inv_bind1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct
lapply (IHV12 … HV20 ? ?) // -IHV12 -HV20 #HV10
lapply (IHT12 … HT20 ? ?) /2 width=1 by lift_bind, le_S_S/ (**) (* full auto a bit slow *)
-| #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hd12 #Hd21
+| #I #V1 #V2 #T1 #T2 #l1 #m1 #_ #_ #IHV12 #IHT12 #l2 #m2 #X #HX #Hl12 #Hl21
elim (lift_inv_flat1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct
lapply (IHV12 … HV20 ? ?) // -IHV12 -HV20 #HV10
lapply (IHT12 … HT20 ? ?) /2 width=1 by lift_flat/ (**) (* full auto a bit slow *)
qed.
(* Basic_1: was: lift_d (right to left) *)
-theorem lift_trans_le: ∀d1,e1,T1,T. ⬆[d1, e1] T1 ≡ T →
- ∀d2,e2,T2. ⬆[d2, e2] T ≡ T2 → d2 ≤ d1 →
- ∃∃T0. ⬆[d2, e2] T1 ≡ T0 & ⬆[d1 + e2, e1] T0 ≡ T2.
-#d1 #e1 #T1 #T #H elim H -d1 -e1 -T1 -T
-[ #k #d1 #e1 #d2 #e2 #X #HX #_
+theorem lift_trans_le: ∀l1,m1,T1,T. ⬆[l1, m1] T1 ≡ T →
+ ∀l2,m2,T2. ⬆[l2, m2] T ≡ T2 → l2 ≤ l1 →
+ ∃∃T0. ⬆[l2, m2] T1 ≡ T0 & ⬆[l1 + m2, m1] T0 ≡ T2.
+#l1 #m1 #T1 #T #H elim H -l1 -m1 -T1 -T
+[ #k #l1 #m1 #l2 #m2 #X #HX #_
>(lift_inv_sort1 … HX) -HX /2 width=3 by lift_sort, ex2_intro/
-| #i #d1 #e1 #Hid1 #d2 #e2 #X #HX #_
- lapply (lt_to_le_to_lt … (d1+e2) Hid1 ?) // #Hie2
- elim (lift_inv_lref1 … HX) -HX * #Hid2 #HX destruct /4 width=3 by lift_lref_ge_minus, lift_lref_lt, lt_minus_to_plus, monotonic_le_plus_l, ex2_intro/
-| #i #d1 #e1 #Hid1 #d2 #e2 #X #HX #Hd21
- lapply (transitive_le … Hd21 Hid1) -Hd21 #Hid2
+| #i #l1 #m1 #Hil1 #l2 #m2 #X #HX #_
+ lapply (lt_to_le_to_lt … (l1+m2) Hil1 ?) // #Him2
+ elim (lift_inv_lref1 … HX) -HX * #Hil2 #HX destruct /4 width=3 by lift_lref_ge_minus, lift_lref_lt, lt_minus_to_plus, monotonic_le_plus_l, ex2_intro/
+| #i #l1 #m1 #Hil1 #l2 #m2 #X #HX #Hl21
+ lapply (transitive_le … Hl21 Hil1) -Hl21 #Hil2
lapply (lift_inv_lref1_ge … HX ?) -HX /2 width=3 by transitive_le/ #HX destruct
>plus_plus_comm_23 /4 width=3 by lift_lref_ge_minus, lift_lref_ge, monotonic_le_plus_l, ex2_intro/
-| #p #d1 #e1 #d2 #e2 #X #HX #_
+| #p #l1 #m1 #l2 #m2 #X #HX #_
>(lift_inv_gref1 … HX) -HX /2 width=3 by lift_gref, ex2_intro/
-| #a #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hd21
+| #a #I #V1 #V2 #T1 #T2 #l1 #m1 #_ #_ #IHV12 #IHT12 #l2 #m2 #X #HX #Hl21
elim (lift_inv_bind1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct
elim (IHV12 … HV20) -IHV12 -HV20 //
elim (IHT12 … HT20) -IHT12 -HT20 /3 width=5 by lift_bind, le_S_S, ex2_intro/
-| #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hd21
+| #I #V1 #V2 #T1 #T2 #l1 #m1 #_ #_ #IHV12 #IHT12 #l2 #m2 #X #HX #Hl21
elim (lift_inv_flat1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct
elim (IHV12 … HV20) -IHV12 -HV20 //
elim (IHT12 … HT20) -IHT12 -HT20 /3 width=5 by lift_flat, ex2_intro/
qed.
(* Basic_1: was: lift_d (left to right) *)
-theorem lift_trans_ge: ∀d1,e1,T1,T. ⬆[d1, e1] T1 ≡ T →
- ∀d2,e2,T2. ⬆[d2, e2] T ≡ T2 → d1 + e1 ≤ d2 →
- ∃∃T0. ⬆[d2 - e1, e2] T1 ≡ T0 & ⬆[d1, e1] T0 ≡ T2.
-#d1 #e1 #T1 #T #H elim H -d1 -e1 -T1 -T
-[ #k #d1 #e1 #d2 #e2 #X #HX #_
+theorem lift_trans_ge: ∀l1,m1,T1,T. ⬆[l1, m1] T1 ≡ T →
+ ∀l2,m2,T2. ⬆[l2, m2] T ≡ T2 → l1 + m1 ≤ l2 →
+ ∃∃T0. ⬆[l2 - m1, m2] T1 ≡ T0 & ⬆[l1, m1] T0 ≡ T2.
+#l1 #m1 #T1 #T #H elim H -l1 -m1 -T1 -T
+[ #k #l1 #m1 #l2 #m2 #X #HX #_
>(lift_inv_sort1 … HX) -HX /2 width=3 by lift_sort, ex2_intro/
-| #i #d1 #e1 #Hid1 #d2 #e2 #X #HX #Hded
- lapply (lt_to_le_to_lt … (d1+e1) Hid1 ?) // #Hid1e
- lapply (lt_to_le_to_lt … (d2-e1) Hid1 ?) /2 width=1 by le_plus_to_minus_r/ #Hid2e
- lapply (lt_to_le_to_lt … Hid1e Hded) -Hid1e -Hded #Hid2
+| #i #l1 #m1 #Hil1 #l2 #m2 #X #HX #Hlml
+ lapply (lt_to_le_to_lt … (l1+m1) Hil1 ?) // #Hil1m
+ lapply (lt_to_le_to_lt … (l2-m1) Hil1 ?) /2 width=1 by le_plus_to_minus_r/ #Hil2m
+ lapply (lt_to_le_to_lt … Hil1m Hlml) -Hil1m -Hlml #Hil2
lapply (lift_inv_lref1_lt … HX ?) -HX // #HX destruct /3 width=3 by lift_lref_lt, ex2_intro/
-| #i #d1 #e1 #Hid1 #d2 #e2 #X #HX #_
- elim (lift_inv_lref1 … HX) -HX * #Hied #HX destruct /4 width=3 by lift_lref_lt, lift_lref_ge, monotonic_le_minus_l, lt_plus_to_minus_r, transitive_le, ex2_intro/
-| #p #d1 #e1 #d2 #e2 #X #HX #_
+| #i #l1 #m1 #Hil1 #l2 #m2 #X #HX #_
+ elim (lift_inv_lref1 … HX) -HX * #Himl #HX destruct /4 width=3 by lift_lref_lt, lift_lref_ge, monotonic_le_minus_l, lt_plus_to_minus_r, transitive_le, ex2_intro/
+| #p #l1 #m1 #l2 #m2 #X #HX #_
>(lift_inv_gref1 … HX) -HX /2 width=3 by lift_gref, ex2_intro/
-| #a #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hded
+| #a #I #V1 #V2 #T1 #T2 #l1 #m1 #_ #_ #IHV12 #IHT12 #l2 #m2 #X #HX #Hlml
elim (lift_inv_bind1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct
elim (IHV12 … HV20) -IHV12 -HV20 //
elim (IHT12 … HT20) -IHT12 -HT20 /2 width=1 by le_S_S/ #T
<plus_minus /3 width=5 by lift_bind, le_plus_to_minus_r, le_plus_b, ex2_intro/
-| #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hded
+| #I #V1 #V2 #T1 #T2 #l1 #m1 #_ #_ #IHV12 #IHT12 #l2 #m2 #X #HX #Hlml
elim (lift_inv_flat1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct
elim (IHV12 … HV20) -IHV12 -HV20 //
elim (IHT12 … HT20) -IHT12 -HT20 /3 width=5 by lift_flat, ex2_intro/
(* Advanced properties ******************************************************)
-lemma lift_conf_O1: ∀T,T1,d1,e1. ⬆[d1, e1] T ≡ T1 → ∀T2,e2. ⬆[0, e2] T ≡ T2 →
- ∃∃T0. ⬆[0, e2] T1 ≡ T0 & ⬆[d1 + e2, e1] T2 ≡ T0.
-#T #T1 #d1 #e1 #HT1 #T2 #e2 #HT2
-elim (lift_total T1 0 e2) #T0 #HT10
+lemma lift_conf_O1: ∀T,T1,l1,m1. ⬆[l1, m1] T ≡ T1 → ∀T2,m2. ⬆[0, m2] T ≡ T2 →
+ ∃∃T0. ⬆[0, m2] T1 ≡ T0 & ⬆[l1 + m2, m1] T2 ≡ T0.
+#T #T1 #l1 #m1 #HT1 #T2 #m2 #HT2
+elim (lift_total T1 0 m2) #T0 #HT10
elim (lift_trans_le … HT1 … HT10) -HT1 // #X #HTX #HT20
lapply (lift_mono … HTX … HT2) -T #H destruct /2 width=3 by ex2_intro/
qed.
-lemma lift_conf_be: ∀T,T1,d,e1. ⬆[d, e1] T ≡ T1 → ∀T2,e2. ⬆[d, e2] T ≡ T2 →
- e1 ≤ e2 → ⬆[d + e1, e2 - e1] T1 ≡ T2.
-#T #T1 #d #e1 #HT1 #T2 #e2 #HT2 #He12
-elim (lift_split … HT2 (d+e1) e1) -HT2 // #X #H
+lemma lift_conf_be: ∀T,T1,l,m1. ⬆[l, m1] T ≡ T1 → ∀T2,m2. ⬆[l, m2] T ≡ T2 →
+ m1 ≤ m2 → ⬆[l + m1, m2 - m1] T1 ≡ T2.
+#T #T1 #l #m1 #HT1 #T2 #m2 #HT2 #Hm12
+elim (lift_split … HT2 (l+m1) m1) -HT2 // #X #H
>(lift_mono … H … HT1) -T //
qed.
(* Main properties ***********************************************************)
-theorem liftv_mono: ∀Ts,U1s,d,e. ⬆[d,e] Ts ≡ U1s →
- ∀U2s:list term. ⬆[d,e] Ts ≡ U2s → U1s = U2s.
-#Ts #U1s #d #e #H elim H -Ts -U1s
+theorem liftv_mono: ∀Ts,U1s,l,m. ⬆[l,m] Ts ≡ U1s →
+ ∀U2s:list term. ⬆[l,m] Ts ≡ U2s → U1s = U2s.
+#Ts #U1s #l #m #H elim H -Ts -U1s
[ #U2s #H >(liftv_inv_nil1 … H) -H //
| #Ts #U1s #T #U1 #HTU1 #_ #IHTU1s #X #H destruct
elim (liftv_inv_cons1 … H) -H #U2 #U2s #HTU2 #HTU2s #H destruct
lemma nlift_lref_be_SO: ∀X,i. ⬆[i, 1] X ≡ #i → ⊥.
/2 width=7 by lift_inv_lref2_be/ qed-.
-lemma nlift_bind_sn: ∀W,d,e. (∀V. ⬆[d, e] V ≡ W → ⊥) →
- ∀a,I,U. (∀X. ⬆[d, e] X ≡ ⓑ{a,I}W.U → ⊥).
-#W #d #e #HW #a #I #U #X #H elim (lift_inv_bind2 … H) -H /2 width=2 by/
+lemma nlift_bind_sn: ∀W,l,m. (∀V. ⬆[l, m] V ≡ W → ⊥) →
+ ∀a,I,U. (∀X. ⬆[l, m] X ≡ ⓑ{a,I}W.U → ⊥).
+#W #l #m #HW #a #I #U #X #H elim (lift_inv_bind2 … H) -H /2 width=2 by/
qed-.
-lemma nlift_bind_dx: ∀U,d,e. (∀T. ⬆[d+1, e] T ≡ U → ⊥) →
- ∀a,I,W. (∀X. ⬆[d, e] X ≡ ⓑ{a,I}W.U → ⊥).
-#U #d #e #HU #a #I #W #X #H elim (lift_inv_bind2 … H) -H /2 width=2 by/
+lemma nlift_bind_dx: ∀U,l,m. (∀T. ⬆[l+1, m] T ≡ U → ⊥) →
+ ∀a,I,W. (∀X. ⬆[l, m] X ≡ ⓑ{a,I}W.U → ⊥).
+#U #l #m #HU #a #I #W #X #H elim (lift_inv_bind2 … H) -H /2 width=2 by/
qed-.
-lemma nlift_flat_sn: ∀W,d,e. (∀V. ⬆[d, e] V ≡ W → ⊥) →
- ∀I,U. (∀X. ⬆[d, e] X ≡ ⓕ{I}W.U → ⊥).
-#W #d #e #HW #I #U #X #H elim (lift_inv_flat2 … H) -H /2 width=2 by/
+lemma nlift_flat_sn: ∀W,l,m. (∀V. ⬆[l, m] V ≡ W → ⊥) →
+ ∀I,U. (∀X. ⬆[l, m] X ≡ ⓕ{I}W.U → ⊥).
+#W #l #m #HW #I #U #X #H elim (lift_inv_flat2 … H) -H /2 width=2 by/
qed-.
-lemma nlift_flat_dx: ∀U,d,e. (∀T. ⬆[d, e] T ≡ U → ⊥) →
- ∀I,W. (∀X. ⬆[d, e] X ≡ ⓕ{I}W.U → ⊥).
-#U #d #e #HU #I #W #X #H elim (lift_inv_flat2 … H) -H /2 width=2 by/
+lemma nlift_flat_dx: ∀U,l,m. (∀T. ⬆[l, m] T ≡ U → ⊥) →
+ ∀I,W. (∀X. ⬆[l, m] X ≡ ⓕ{I}W.U → ⊥).
+#U #l #m #HU #I #W #X #H elim (lift_inv_flat2 … H) -H /2 width=2 by/
qed-.
(* Inversion lemmas on negated basic relocation *****************************)
]
qed-.
-lemma nlift_inv_bind: ∀a,I,W,U,d,e. (∀X. ⬆[d, e] X ≡ ⓑ{a,I}W.U → ⊥) →
- (∀V. ⬆[d, e] V ≡ W → ⊥) ∨ (∀T. ⬆[d+1, e] T ≡ U → ⊥).
-#a #I #W #U #d #e #H elim (is_lift_dec W d e)
+lemma nlift_inv_bind: ∀a,I,W,U,l,m. (∀X. ⬆[l, m] X ≡ ⓑ{a,I}W.U → ⊥) →
+ (∀V. ⬆[l, m] V ≡ W → ⊥) ∨ (∀T. ⬆[l+1, m] T ≡ U → ⊥).
+#a #I #W #U #l #m #H elim (is_lift_dec W l m)
[ * /4 width=2 by lift_bind, or_intror/
| /4 width=2 by ex_intro, or_introl/
]
qed-.
-lemma nlift_inv_flat: ∀I,W,U,d,e. (∀X. ⬆[d, e] X ≡ ⓕ{I}W.U → ⊥) →
- (∀V. ⬆[d, e] V ≡ W → ⊥) ∨ (∀T. ⬆[d, e] T ≡ U → ⊥).
-#I #W #U #d #e #H elim (is_lift_dec W d e)
+lemma nlift_inv_flat: ∀I,W,U,l,m. (∀X. ⬆[l, m] X ≡ ⓕ{I}W.U → ⊥) →
+ (∀V. ⬆[l, m] V ≡ W → ⊥) ∨ (∀T. ⬆[l, m] T ≡ U → ⊥).
+#I #W #U #l #m #H elim (is_lift_dec W l m)
[ * /4 width=2 by lift_flat, or_intror/
| /4 width=2 by ex_intro, or_introl/
]
(* BASIC TERM VECTOR RELOCATION *********************************************)
-inductive liftv (d,e:nat) : relation (list term) ≝
-| liftv_nil : liftv d e (◊) (◊)
+inductive liftv (l,m:nat) : relation (list term) ≝
+| liftv_nil : liftv l m (◊) (◊)
| liftv_cons: ∀T1s,T2s,T1,T2.
- ⬆[d, e] T1 ≡ T2 → liftv d e T1s T2s →
- liftv d e (T1 @ T1s) (T2 @ T2s)
+ ⬆[l, m] T1 ≡ T2 → liftv l m T1s T2s →
+ liftv l m (T1 @ T1s) (T2 @ T2s)
.
-interpretation "relocation (vector)" 'RLift d e T1s T2s = (liftv d e T1s T2s).
+interpretation "relocation (vector)" 'RLift l m T1s T2s = (liftv l m T1s T2s).
(* Basic inversion lemmas ***************************************************)
-fact liftv_inv_nil1_aux: ∀T1s,T2s,d,e. ⬆[d, e] T1s ≡ T2s → T1s = ◊ → T2s = ◊.
-#T1s #T2s #d #e * -T1s -T2s //
+fact liftv_inv_nil1_aux: ∀T1s,T2s,l,m. ⬆[l, m] T1s ≡ T2s → T1s = ◊ → T2s = ◊.
+#T1s #T2s #l #m * -T1s -T2s //
#T1s #T2s #T1 #T2 #_ #_ #H destruct
qed-.
-lemma liftv_inv_nil1: ∀T2s,d,e. ⬆[d, e] ◊ ≡ T2s → T2s = ◊.
+lemma liftv_inv_nil1: ∀T2s,l,m. ⬆[l, m] ◊ ≡ T2s → T2s = ◊.
/2 width=5 by liftv_inv_nil1_aux/ qed-.
-fact liftv_inv_cons1_aux: ∀T1s,T2s,d,e. ⬆[d, e] T1s ≡ T2s →
+fact liftv_inv_cons1_aux: ∀T1s,T2s,l,m. ⬆[l, m] T1s ≡ T2s →
∀U1,U1s. T1s = U1 @ U1s →
- ∃∃U2,U2s. ⬆[d, e] U1 ≡ U2 & ⬆[d, e] U1s ≡ U2s &
+ ∃∃U2,U2s. ⬆[l, m] U1 ≡ U2 & ⬆[l, m] U1s ≡ U2s &
T2s = U2 @ U2s.
-#T1s #T2s #d #e * -T1s -T2s
+#T1s #T2s #l #m * -T1s -T2s
[ #U1 #U1s #H destruct
| #T1s #T2s #T1 #T2 #HT12 #HT12s #U1 #U1s #H destruct /2 width=5 by ex3_2_intro/
]
qed-.
-lemma liftv_inv_cons1: ∀U1,U1s,T2s,d,e. ⬆[d, e] U1 @ U1s ≡ T2s →
- ∃∃U2,U2s. ⬆[d, e] U1 ≡ U2 & ⬆[d, e] U1s ≡ U2s &
+lemma liftv_inv_cons1: ∀U1,U1s,T2s,l,m. ⬆[l, m] U1 @ U1s ≡ T2s →
+ ∃∃U2,U2s. ⬆[l, m] U1 ≡ U2 & ⬆[l, m] U1s ≡ U2s &
T2s = U2 @ U2s.
/2 width=3 by liftv_inv_cons1_aux/ qed-.
(* Basic properties *********************************************************)
-lemma liftv_total: ∀d,e. ∀T1s:list term. ∃T2s. ⬆[d, e] T1s ≡ T2s.
-#d #e #T1s elim T1s -T1s
+lemma liftv_total: ∀l,m. ∀T1s:list term. ∃T2s. ⬆[l, m] T1s ≡ T2s.
+#l #m #T1s elim T1s -T1s
[ /2 width=2 by liftv_nil, ex_intro/
| #T1 #T1s * #T2s #HT12s
- elim (lift_total T1 d e) /3 width=2 by liftv_cons, ex_intro/
+ elim (lift_total T1 l m) /3 width=2 by liftv_cons, ex_intro/
]
qed-.
(* *)
(**************************************************************************)
-include "basic_2/substitution/drop_leq.ma".
+include "basic_2/substitution/drop_lreq.ma".
include "basic_2/substitution/lpx_sn.ma".
(* SN POINTWISE EXTENSION OF A CONTEXT-SENSITIVE REALTION FOR TERMS *********)
]
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
+lemma lpx_sn_deliftable_dropable: ∀R. d_deliftable_sn R → dropable_sn (lpx_sn R).
+#R #HR #L1 #K1 #s #l #m #H elim H -L1 -K1 -l -m
+[ #l #m #Hm #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
+| #I #L1 #K1 #V1 #m #_ #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
+| #I #L1 #K1 #V1 #W1 #l #m #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
+ d_liftable R → dedropable_sn (lpx_sn R).
+#R #H1R #H2R #L1 #K1 #s #l #m #H elim H -L1 -K1 -l -m
+[ #l #m #Hm #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
+ @lreq_O2 normalize //
+| #I #L1 #K1 #V1 #m #_ #IHLK1 #K2 #HK12 elim (IHLK1 … HK12) -K1
+ /3 width=5 by drop_drop, lreq_pair, lpx_sn_pair, ex3_intro/
+| #I #L1 #K1 #V1 #W1 #l #m #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
+ elim (lift_total W2 l m) #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/
+ /3 width=6 by drop_skip, lreq_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
+fact lpx_sn_dropable_aux: ∀R,L2,K2,s,l,m. ⬇[s, l, m] L2 ≡ K2 → ∀L1. lpx_sn R L1 L2 →
+ l = 0 → ∃∃K1. ⬇[s, 0, m] L1 ≡ K1 & lpx_sn R K1 K2.
+#R #L2 #K2 #s #l #m #H elim H -L2 -K2 -l -m
+[ #l #m #Hm #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
+| #I #L2 #K2 #V2 #m #_ #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 #_
+| #I #L2 #K2 #V2 #W2 #l #m #_ #_ #_ #L1 #_
<plus_n_Sm #H destruct
]
qed-.
(* LOCAL ENVIRONMENT REFINEMENT FOR EXTENDED SUBSTITUTION *******************)
inductive lsuby: relation4 ynat ynat lenv lenv ≝
-| lsuby_atom: ∀L,d,e. lsuby d e L (⋆)
+| lsuby_atom: ∀L,l,m. lsuby l m L (⋆)
| lsuby_zero: ∀I1,I2,L1,L2,V1,V2.
lsuby 0 0 L1 L2 → lsuby 0 0 (L1.ⓑ{I1}V1) (L2.ⓑ{I2}V2)
-| lsuby_pair: ∀I1,I2,L1,L2,V,e. lsuby 0 e L1 L2 →
- lsuby 0 (⫯e) (L1.ⓑ{I1}V) (L2.ⓑ{I2}V)
-| lsuby_succ: ∀I1,I2,L1,L2,V1,V2,d,e.
- lsuby d e L1 L2 → lsuby (⫯d) e (L1.ⓑ{I1}V1) (L2.ⓑ{I2}V2)
+| lsuby_pair: ∀I1,I2,L1,L2,V,m. lsuby 0 m L1 L2 →
+ lsuby 0 (⫯m) (L1.ⓑ{I1}V) (L2.ⓑ{I2}V)
+| lsuby_succ: ∀I1,I2,L1,L2,V1,V2,l,m.
+ lsuby l m L1 L2 → lsuby (⫯l) m (L1.ⓑ{I1}V1) (L2.ⓑ{I2}V2)
.
interpretation
"local environment refinement (extended substitution)"
- 'LRSubEq L1 d e L2 = (lsuby d e L1 L2).
+ 'LRSubEq L1 l m L2 = (lsuby l m L1 L2).
(* Basic properties *********************************************************)
-lemma lsuby_pair_lt: ∀I1,I2,L1,L2,V,e. L1 ⊆[0, ⫰e] L2 → 0 < e →
- L1.ⓑ{I1}V ⊆[0, e] L2.ⓑ{I2}V.
-#I1 #I2 #L1 #L2 #V #e #HL12 #He <(ylt_inv_O1 … He) /2 width=1 by lsuby_pair/
+lemma lsuby_pair_lt: ∀I1,I2,L1,L2,V,m. L1 ⊆[0, ⫰m] L2 → 0 < m →
+ L1.ⓑ{I1}V ⊆[0, m] L2.ⓑ{I2}V.
+#I1 #I2 #L1 #L2 #V #m #HL12 #Hm <(ylt_inv_O1 … Hm) /2 width=1 by lsuby_pair/
qed.
-lemma lsuby_succ_lt: ∀I1,I2,L1,L2,V1,V2,d,e. L1 ⊆[⫰d, e] L2 → 0 < d →
- L1.ⓑ{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 lsuby_succ/
+lemma lsuby_succ_lt: ∀I1,I2,L1,L2,V1,V2,l,m. L1 ⊆[⫰l, m] L2 → 0 < l →
+ L1.ⓑ{I1}V1 ⊆[l, m] L2. ⓑ{I2}V2.
+#I1 #I2 #L1 #L2 #V1 #V2 #l #m #HL12 #Hl <(ylt_inv_O1 … Hl) /2 width=1 by lsuby_succ/
qed.
lemma lsuby_pair_O_Y: ∀L1,L2. L1 ⊆[0, ∞] L2 →
#L1 #L2 #HL12 #I1 #I2 #V lapply (lsuby_pair I1 I2 … V … HL12) -HL12 //
qed.
-lemma lsuby_refl: ∀L,d,e. L ⊆[d, e] L.
+lemma lsuby_refl: ∀L,l,m. L ⊆[l, m] L.
#L elim L -L //
-#L #I #V #IHL #d elim (ynat_cases … d) [| * #x ]
-#Hd destruct /2 width=1 by lsuby_succ/
-#e elim (ynat_cases … e) [| * #x ]
-#He destruct /2 width=1 by lsuby_zero, lsuby_pair/
+#L #I #V #IHL #l elim (ynat_cases … l) [| * #x ]
+#Hl destruct /2 width=1 by lsuby_succ/
+#m elim (ynat_cases … m) [| * #x ]
+#Hm destruct /2 width=1 by lsuby_zero, lsuby_pair/
qed.
-lemma lsuby_O2: ∀L2,L1,d. |L2| ≤ |L1| → L1 ⊆[d, yinj 0] L2.
+lemma lsuby_O2: ∀L2,L1,l. |L2| ≤ |L1| → L1 ⊆[l, yinj 0] L2.
#L2 elim L2 -L2 // #L2 #I2 #V2 #IHL2 * normalize
-[ #d #H elim (le_plus_xSy_O_false … H)
-| #L1 #I1 #V1 #d #H lapply (le_plus_to_le_r … H) -H #HL12
- elim (ynat_cases d) /3 width=1 by lsuby_zero/
+[ #l #H elim (le_plus_xSy_O_false … H)
+| #L1 #I1 #V1 #l #H lapply (le_plus_to_le_r … H) -H #HL12
+ elim (ynat_cases l) /3 width=1 by lsuby_zero/
* /3 width=1 by lsuby_succ/
]
qed.
-lemma lsuby_sym: ∀d,e,L1,L2. L1 ⊆[d, e] L2 → |L1| = |L2| → L2 ⊆[d, e] L1.
-#d #e #L1 #L2 #H elim H -d -e -L1 -L2
-[ #L1 #d #e #H >(length_inv_zero_dx … H) -L1 //
+lemma lsuby_sym: ∀l,m,L1,L2. L1 ⊆[l, m] L2 → |L1| = |L2| → L2 ⊆[l, m] L1.
+#l #m #L1 #L2 #H elim H -l -m -L1 -L2
+[ #L1 #l #m #H >(length_inv_zero_dx … H) -L1 //
| /2 width=1 by lsuby_O2/
-| #I1 #I2 #L1 #L2 #V #e #_ #IHL12 #H lapply (injective_plus_l … H)
+| #I1 #I2 #L1 #L2 #V #m #_ #IHL12 #H lapply (injective_plus_l … H)
/3 width=1 by lsuby_pair/
-| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #IHL12 #H lapply (injective_plus_l … H)
+| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #_ #IHL12 #H lapply (injective_plus_l … H)
/3 width=1 by lsuby_succ/
]
qed-.
(* Basic inversion lemmas ***************************************************)
-fact lsuby_inv_atom1_aux: ∀L1,L2,d,e. L1 ⊆[d, e] L2 → L1 = ⋆ → L2 = ⋆.
-#L1 #L2 #d #e * -L1 -L2 -d -e //
+fact lsuby_inv_atom1_aux: ∀L1,L2,l,m. L1 ⊆[l, m] L2 → L1 = ⋆ → L2 = ⋆.
+#L1 #L2 #l #m * -L1 -L2 -l -m //
[ #I1 #I2 #L1 #L2 #V1 #V2 #_ #H destruct
-| #I1 #I2 #L1 #L2 #V #e #_ #H destruct
-| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #H destruct
+| #I1 #I2 #L1 #L2 #V #m #_ #H destruct
+| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #_ #H destruct
]
qed-.
-lemma lsuby_inv_atom1: ∀L2,d,e. ⋆ ⊆[d, e] L2 → L2 = ⋆.
+lemma lsuby_inv_atom1: ∀L2,l,m. ⋆ ⊆[l, m] L2 → L2 = ⋆.
/2 width=5 by lsuby_inv_atom1_aux/ qed-.
-fact lsuby_inv_zero1_aux: ∀L1,L2,d,e. L1 ⊆[d, e] L2 →
- ∀J1,K1,W1. L1 = K1.ⓑ{J1}W1 → d = 0 → e = 0 →
+fact lsuby_inv_zero1_aux: ∀L1,L2,l,m. L1 ⊆[l, m] L2 →
+ ∀J1,K1,W1. L1 = K1.ⓑ{J1}W1 → l = 0 → m = 0 →
L2 = ⋆ ∨
∃∃J2,K2,W2. K1 ⊆[0, 0] K2 & L2 = K2.ⓑ{J2}W2.
-#L1 #L2 #d #e * -L1 -L2 -d -e /2 width=1 by or_introl/
+#L1 #L2 #l #m * -L1 -L2 -l -m /2 width=1 by or_introl/
[ #I1 #I2 #L1 #L2 #V1 #V2 #HL12 #J1 #K1 #W1 #H #_ #_ destruct
/3 width=5 by ex2_3_intro, or_intror/
-| #I1 #I2 #L1 #L2 #V #e #_ #J1 #K1 #W1 #_ #_ #H
+| #I1 #I2 #L1 #L2 #V #m #_ #J1 #K1 #W1 #_ #_ #H
elim (ysucc_inv_O_dx … H)
-| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #J1 #K1 #W1 #_ #H
+| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #_ #J1 #K1 #W1 #_ #H
elim (ysucc_inv_O_dx … H)
]
qed-.
∃∃I2,K2,V2. K1 ⊆[0, 0] K2 & L2 = K2.ⓑ{I2}V2.
/2 width=9 by lsuby_inv_zero1_aux/ qed-.
-fact lsuby_inv_pair1_aux: ∀L1,L2,d,e. L1 ⊆[d, e] L2 →
- ∀J1,K1,W. L1 = K1.ⓑ{J1}W → d = 0 → 0 < e →
+fact lsuby_inv_pair1_aux: ∀L1,L2,l,m. L1 ⊆[l, m] L2 →
+ ∀J1,K1,W. L1 = K1.ⓑ{J1}W → l = 0 → 0 < m →
L2 = ⋆ ∨
- ∃∃J2,K2. K1 ⊆[0, ⫰e] K2 & L2 = K2.ⓑ{J2}W.
-#L1 #L2 #d #e * -L1 -L2 -d -e /2 width=1 by or_introl/
+ ∃∃J2,K2. K1 ⊆[0, ⫰m] K2 & L2 = K2.ⓑ{J2}W.
+#L1 #L2 #l #m * -L1 -L2 -l -m /2 width=1 by or_introl/
[ #I1 #I2 #L1 #L2 #V1 #V2 #_ #J1 #K1 #W #_ #_ #H
elim (ylt_yle_false … H) //
-| #I1 #I2 #L1 #L2 #V #e #HL12 #J1 #K1 #W #H #_ #_ destruct
+| #I1 #I2 #L1 #L2 #V #m #HL12 #J1 #K1 #W #H #_ #_ destruct
/3 width=4 by ex2_2_intro, or_intror/
-| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #J1 #K1 #W #_ #H
+| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #_ #J1 #K1 #W #_ #H
elim (ysucc_inv_O_dx … H)
]
qed-.
-lemma lsuby_inv_pair1: ∀I1,K1,L2,V,e. K1.ⓑ{I1}V ⊆[0, e] L2 → 0 < e →
+lemma lsuby_inv_pair1: ∀I1,K1,L2,V,m. K1.ⓑ{I1}V ⊆[0, m] L2 → 0 < m →
L2 = ⋆ ∨
- ∃∃I2,K2. K1 ⊆[0, ⫰e] K2 & L2 = K2.ⓑ{I2}V.
+ ∃∃I2,K2. K1 ⊆[0, ⫰m] K2 & L2 = K2.ⓑ{I2}V.
/2 width=6 by lsuby_inv_pair1_aux/ qed-.
-fact lsuby_inv_succ1_aux: ∀L1,L2,d,e. L1 ⊆[d, e] L2 →
- ∀J1,K1,W1. L1 = K1.ⓑ{J1}W1 → 0 < d →
+fact lsuby_inv_succ1_aux: ∀L1,L2,l,m. L1 ⊆[l, m] L2 →
+ ∀J1,K1,W1. L1 = K1.ⓑ{J1}W1 → 0 < l →
L2 = ⋆ ∨
- ∃∃J2,K2,W2. K1 ⊆[⫰d, e] K2 & L2 = K2.ⓑ{J2}W2.
-#L1 #L2 #d #e * -L1 -L2 -d -e /2 width=1 by or_introl/
+ ∃∃J2,K2,W2. K1 ⊆[⫰l, m] K2 & L2 = K2.ⓑ{J2}W2.
+#L1 #L2 #l #m * -L1 -L2 -l -m /2 width=1 by or_introl/
[ #I1 #I2 #L1 #L2 #V1 #V2 #_ #J1 #K1 #W1 #_ #H
elim (ylt_yle_false … H) //
-| #I1 #I2 #L1 #L2 #V #e #_ #J1 #K1 #W1 #_ #H
+| #I1 #I2 #L1 #L2 #V #m #_ #J1 #K1 #W1 #_ #H
elim (ylt_yle_false … H) //
-| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #HL12 #J1 #K1 #W1 #H #_ destruct
+| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #HL12 #J1 #K1 #W1 #H #_ destruct
/3 width=5 by ex2_3_intro, or_intror/
]
qed-.
-lemma lsuby_inv_succ1: ∀I1,K1,L2,V1,d,e. K1.ⓑ{I1}V1 ⊆[d, e] L2 → 0 < d →
+lemma lsuby_inv_succ1: ∀I1,K1,L2,V1,l,m. K1.ⓑ{I1}V1 ⊆[l, m] L2 → 0 < l →
L2 = ⋆ ∨
- ∃∃I2,K2,V2. K1 ⊆[⫰d, e] K2 & L2 = K2.ⓑ{I2}V2.
+ ∃∃I2,K2,V2. K1 ⊆[⫰l, m] K2 & L2 = K2.ⓑ{I2}V2.
/2 width=5 by lsuby_inv_succ1_aux/ qed-.
-fact lsuby_inv_zero2_aux: ∀L1,L2,d,e. L1 ⊆[d, e] L2 →
- ∀J2,K2,W2. L2 = K2.ⓑ{J2}W2 → d = 0 → e = 0 →
+fact lsuby_inv_zero2_aux: ∀L1,L2,l,m. L1 ⊆[l, m] L2 →
+ ∀J2,K2,W2. L2 = K2.ⓑ{J2}W2 → l = 0 → m = 0 →
∃∃J1,K1,W1. K1 ⊆[0, 0] K2 & L1 = K1.ⓑ{J1}W1.
-#L1 #L2 #d #e * -L1 -L2 -d -e
-[ #L1 #d #e #J2 #K2 #W1 #H destruct
+#L1 #L2 #l #m * -L1 -L2 -l -m
+[ #L1 #l #m #J2 #K2 #W1 #H destruct
| #I1 #I2 #L1 #L2 #V1 #V2 #HL12 #J2 #K2 #W2 #H #_ #_ destruct
/2 width=5 by ex2_3_intro/
-| #I1 #I2 #L1 #L2 #V #e #_ #J2 #K2 #W2 #_ #_ #H
+| #I1 #I2 #L1 #L2 #V #m #_ #J2 #K2 #W2 #_ #_ #H
elim (ysucc_inv_O_dx … H)
-| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #J2 #K2 #W2 #_ #H
+| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #_ #J2 #K2 #W2 #_ #H
elim (ysucc_inv_O_dx … H)
]
qed-.
∃∃I1,K1,V1. K1 ⊆[0, 0] K2 & L1 = K1.ⓑ{I1}V1.
/2 width=9 by lsuby_inv_zero2_aux/ qed-.
-fact lsuby_inv_pair2_aux: ∀L1,L2,d,e. L1 ⊆[d, e] L2 →
- ∀J2,K2,W. L2 = K2.ⓑ{J2}W → d = 0 → 0 < e →
- ∃∃J1,K1. K1 ⊆[0, ⫰e] K2 & L1 = K1.ⓑ{J1}W.
-#L1 #L2 #d #e * -L1 -L2 -d -e
-[ #L1 #d #e #J2 #K2 #W #H destruct
+fact lsuby_inv_pair2_aux: ∀L1,L2,l,m. L1 ⊆[l, m] L2 →
+ ∀J2,K2,W. L2 = K2.ⓑ{J2}W → l = 0 → 0 < m →
+ ∃∃J1,K1. K1 ⊆[0, ⫰m] K2 & L1 = K1.ⓑ{J1}W.
+#L1 #L2 #l #m * -L1 -L2 -l -m
+[ #L1 #l #m #J2 #K2 #W #H destruct
| #I1 #I2 #L1 #L2 #V1 #V2 #_ #J2 #K2 #W #_ #_ #H
elim (ylt_yle_false … H) //
-| #I1 #I2 #L1 #L2 #V #e #HL12 #J2 #K2 #W #H #_ #_ destruct
+| #I1 #I2 #L1 #L2 #V #m #HL12 #J2 #K2 #W #H #_ #_ destruct
/2 width=4 by ex2_2_intro/
-| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #J2 #K2 #W #_ #H
+| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #_ #J2 #K2 #W #_ #H
elim (ysucc_inv_O_dx … H)
]
qed-.
-lemma lsuby_inv_pair2: ∀I2,K2,L1,V,e. L1 ⊆[0, e] K2.ⓑ{I2}V → 0 < e →
- ∃∃I1,K1. K1 ⊆[0, ⫰e] K2 & L1 = K1.ⓑ{I1}V.
+lemma lsuby_inv_pair2: ∀I2,K2,L1,V,m. L1 ⊆[0, m] K2.ⓑ{I2}V → 0 < m →
+ ∃∃I1,K1. K1 ⊆[0, ⫰m] K2 & L1 = K1.ⓑ{I1}V.
/2 width=6 by lsuby_inv_pair2_aux/ qed-.
-fact lsuby_inv_succ2_aux: ∀L1,L2,d,e. L1 ⊆[d, e] L2 →
- ∀J2,K2,W2. L2 = K2.ⓑ{J2}W2 → 0 < d →
- ∃∃J1,K1,W1. K1 ⊆[⫰d, e] K2 & L1 = K1.ⓑ{J1}W1.
-#L1 #L2 #d #e * -L1 -L2 -d -e
-[ #L1 #d #e #J2 #K2 #W2 #H destruct
+fact lsuby_inv_succ2_aux: ∀L1,L2,l,m. L1 ⊆[l, m] L2 →
+ ∀J2,K2,W2. L2 = K2.ⓑ{J2}W2 → 0 < l →
+ ∃∃J1,K1,W1. K1 ⊆[⫰l, m] K2 & L1 = K1.ⓑ{J1}W1.
+#L1 #L2 #l #m * -L1 -L2 -l -m
+[ #L1 #l #m #J2 #K2 #W2 #H destruct
| #I1 #I2 #L1 #L2 #V1 #V2 #_ #J2 #K2 #W2 #_ #H
elim (ylt_yle_false … H) //
-| #I1 #I2 #L1 #L2 #V #e #_ #J2 #K1 #W2 #_ #H
+| #I1 #I2 #L1 #L2 #V #m #_ #J2 #K1 #W2 #_ #H
elim (ylt_yle_false … H) //
-| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #HL12 #J2 #K2 #W2 #H #_ destruct
+| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #HL12 #J2 #K2 #W2 #H #_ destruct
/2 width=5 by ex2_3_intro/
]
qed-.
-lemma lsuby_inv_succ2: ∀I2,K2,L1,V2,d,e. L1 ⊆[d, e] K2.ⓑ{I2}V2 → 0 < d →
- ∃∃I1,K1,V1. K1 ⊆[⫰d, e] K2 & L1 = K1.ⓑ{I1}V1.
+lemma lsuby_inv_succ2: ∀I2,K2,L1,V2,l,m. L1 ⊆[l, m] K2.ⓑ{I2}V2 → 0 < l →
+ ∃∃I1,K1,V1. K1 ⊆[⫰l, m] K2 & L1 = K1.ⓑ{I1}V1.
/2 width=5 by lsuby_inv_succ2_aux/ qed-.
(* Basic forward lemmas *****************************************************)
-lemma lsuby_fwd_length: ∀L1,L2,d,e. L1 ⊆[d, e] L2 → |L2| ≤ |L1|.
-#L1 #L2 #d #e #H elim H -L1 -L2 -d -e normalize /2 width=1 by le_S_S/
+lemma lsuby_fwd_length: ∀L1,L2,l,m. L1 ⊆[l, m] L2 → |L2| ≤ |L1|.
+#L1 #L2 #l #m #H elim H -L1 -L2 -l -m normalize /2 width=1 by le_S_S/
qed-.
(* Properties on basic slicing **********************************************)
-lemma lsuby_drop_trans_be: ∀L1,L2,d,e. L1 ⊆[d, e] L2 →
+lemma lsuby_drop_trans_be: ∀L1,L2,l,m. L1 ⊆[l, m] 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
+ l ≤ i → i < l + m →
+ ∃∃I1,K1. K1 ⊆[0, ⫰(l+m-i)] K2 & ⬇[s, 0, i] L1 ≡ K1.ⓑ{I1}W.
+#L1 #L2 #l #m #H elim H -L1 -L2 -l -m
+[ #L1 #l #m #J2 #K2 #W #s #i #H
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
+| #I1 #I2 #L1 #L2 #V #m #HL12 #IHL12 #J2 #K2 #W #s #i #H #_ >yplus_O1
elim (drop_inv_O1_pair1 … H) -H * #Hi #HLK1 [ -IHL12 | -HL12 ]
[ #_ destruct -I2 >ypred_succ
/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
+ #Him elim (IHL12 … HLK1) -IHL12 -HLK1 // -Him
>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 (drop_inv_drop1_lt … HLK2 ?) -HLK2 /2 width=1 by ylt_O/
+| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #_ #IHL12 #J2 #K2 #W #s #i #HLK2 #Hli
+ elim (yle_inv_succ1 … Hli) -Hli
+ #Hli #Hi <Hi >yplus_succ1 #H lapply (ylt_inv_succ … H) -H
+ #Hilm 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, drop_drop_lt, ex2_2_intro/
]
(* Main properties **********************************************************)
-theorem lsuby_trans: ∀d,e. Transitive … (lsuby d e).
-#d #e #L1 #L2 #H elim H -L1 -L2 -d -e
-[ #L1 #d #e #X #H lapply (lsuby_inv_atom1 … H) -H
+theorem lsuby_trans: ∀l,m. Transitive … (lsuby l m).
+#l #m #L1 #L2 #H elim H -L1 -L2 -l -m
+[ #L1 #l #m #X #H lapply (lsuby_inv_atom1 … H) -H
#H destruct //
| #I1 #I2 #L1 #L #V1 #V #_ #IHL1 #X #H elim (lsuby_inv_zero1 … H) -H //
* #I2 #L2 #V2 #HL2 #H destruct /3 width=1 by lsuby_zero/
-| #I1 #I2 #L1 #L2 #V #e #_ #IHL1 #X #H elim (lsuby_inv_pair1 … H) -H //
+| #I1 #I2 #L1 #L2 #V #m #_ #IHL1 #X #H elim (lsuby_inv_pair1 … H) -H //
* #I2 #L2 #HL2 #H destruct /3 width=1 by lsuby_pair/
-| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #IHL1 #X #H elim (lsuby_inv_succ1 … H) -H //
+| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #_ #IHL1 #X #H elim (lsuby_inv_succ1 … H) -H //
* #I2 #L2 #V2 #HL2 #H destruct /3 width=1 by lsuby_succ/
]
qed-.
(* activate genv *)
inductive lstas (h): nat → relation4 genv lenv term term ≝
-| lstas_sort: ∀G,L,l,k. lstas h l G L (⋆k) (⋆((next h)^l k))
-| lstas_ldef: ∀G,L,K,V,W,U,i,l. ⬇[i] L ≡ K.ⓓV → lstas h l G K V W →
- ⬆[0, i+1] W ≡ U → lstas h l G L (#i) U
+| lstas_sort: ∀G,L,d,k. lstas h d G L (⋆k) (⋆((next h)^d k))
+| lstas_ldef: ∀G,L,K,V,W,U,i,d. ⬇[i] L ≡ K.ⓓV → lstas h d G K V W →
+ ⬆[0, i+1] W ≡ U → lstas h d G L (#i) U
| lstas_zero: ∀G,L,K,W,V,i. ⬇[i] L ≡ K.ⓛW → lstas h 0 G K W V →
lstas h 0 G L (#i) (#i)
-| lstas_succ: ∀G,L,K,W,V,U,i,l. ⬇[i] L ≡ K.ⓛW → lstas h l G K W V →
- ⬆[0, i+1] V ≡ U → lstas h (l+1) G L (#i) U
-| lstas_bind: ∀a,I,G,L,V,T,U,l. lstas h l G (L.ⓑ{I}V) T U →
- lstas h l G L (ⓑ{a,I}V.T) (ⓑ{a,I}V.U)
-| lstas_appl: ∀G,L,V,T,U,l. lstas h l G L T U → lstas h l G L (ⓐV.T) (ⓐV.U)
-| lstas_cast: ∀G,L,W,T,U,l. lstas h l G L T U → lstas h l G L (ⓝW.T) U
+| lstas_succ: ∀G,L,K,W,V,U,i,d. ⬇[i] L ≡ K.ⓛW → lstas h d G K W V →
+ ⬆[0, i+1] V ≡ U → lstas h (d+1) G L (#i) U
+| lstas_bind: ∀a,I,G,L,V,T,U,d. lstas h d G (L.ⓑ{I}V) T U →
+ lstas h d G L (ⓑ{a,I}V.T) (ⓑ{a,I}V.U)
+| lstas_appl: ∀G,L,V,T,U,d. lstas h d G L T U → lstas h d G L (ⓐV.T) (ⓐV.U)
+| lstas_cast: ∀G,L,W,T,U,d. lstas h d G L T U → lstas h d G L (ⓝW.T) U
.
interpretation "nat-iterated static type assignment (term)"
- 'StaticTypeStar h G L l T U = (lstas h l G L T U).
+ 'StaticTypeStar h G L d T U = (lstas h d G L T U).
(* Basic inversion lemmas ***************************************************)
-fact lstas_inv_sort1_aux: ∀h,G,L,T,U,l. ⦃G, L⦄ ⊢ T •*[h, l] U → ∀k0. T = ⋆k0 →
- U = ⋆((next h)^l k0).
-#h #G #L #T #U #l * -G -L -T -U -l
-[ #G #L #l #k #k0 #H destruct //
-| #G #L #K #V #W #U #i #l #_ #_ #_ #k0 #H destruct
+fact lstas_inv_sort1_aux: ∀h,G,L,T,U,d. ⦃G, L⦄ ⊢ T •*[h, d] U → ∀k0. T = ⋆k0 →
+ U = ⋆((next h)^d k0).
+#h #G #L #T #U #d * -G -L -T -U -d
+[ #G #L #d #k #k0 #H destruct //
+| #G #L #K #V #W #U #i #d #_ #_ #_ #k0 #H destruct
| #G #L #K #W #V #i #_ #_ #k0 #H destruct
-| #G #L #K #W #V #U #i #l #_ #_ #_ #k0 #H destruct
-| #a #I #G #L #V #T #U #l #_ #k0 #H destruct
-| #G #L #V #T #U #l #_ #k0 #H destruct
-| #G #L #W #T #U #l #_ #k0 #H destruct
+| #G #L #K #W #V #U #i #d #_ #_ #_ #k0 #H destruct
+| #a #I #G #L #V #T #U #d #_ #k0 #H destruct
+| #G #L #V #T #U #d #_ #k0 #H destruct
+| #G #L #W #T #U #d #_ #k0 #H destruct
qed-.
(* Basic_1: was just: sty0_gen_sort *)
-lemma lstas_inv_sort1: ∀h,G,L,X,k,l. ⦃G, L⦄ ⊢ ⋆k •*[h, l] X → X = ⋆((next h)^l k).
+lemma lstas_inv_sort1: ∀h,G,L,X,k,d. ⦃G, L⦄ ⊢ ⋆k •*[h, d] X → X = ⋆((next h)^d k).
/2 width=5 by lstas_inv_sort1_aux/
qed-.
-fact lstas_inv_lref1_aux: ∀h,G,L,T,U,l. ⦃G, L⦄ ⊢ T •*[h, l] U → ∀j. T = #j → ∨∨
- (∃∃K,V,W. ⬇[j] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V •*[h, l] W &
+fact lstas_inv_lref1_aux: ∀h,G,L,T,U,d. ⦃G, L⦄ ⊢ T •*[h, d] U → ∀j. T = #j → ∨∨
+ (∃∃K,V,W. ⬇[j] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V •*[h, d] W &
⬆[0, j+1] W ≡ U
) |
(∃∃K,W,V. ⬇[j] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W •*[h, 0] V &
- U = #j & l = 0
+ U = #j & d = 0
) |
- (∃∃K,W,V,l0. ⬇[j] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W •*[h, l0] V &
- ⬆[0, j+1] V ≡ U & l = l0+1
+ (∃∃K,W,V,d0. ⬇[j] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W •*[h, d0] V &
+ ⬆[0, j+1] V ≡ U & d = d0+1
).
-#h #G #L #T #U #l * -G -L -T -U -l
-[ #G #L #l #k #j #H destruct
-| #G #L #K #V #W #U #i #l #HLK #HVW #HWU #j #H destruct /3 width=6 by or3_intro0, ex3_3_intro/
+#h #G #L #T #U #d * -G -L -T -U -d
+[ #G #L #d #k #j #H destruct
+| #G #L #K #V #W #U #i #d #HLK #HVW #HWU #j #H destruct /3 width=6 by or3_intro0, ex3_3_intro/
| #G #L #K #W #V #i #HLK #HWV #j #H destruct /3 width=5 by or3_intro1, ex4_3_intro/
-| #G #L #K #W #V #U #i #l #HLK #HWV #HWU #j #H destruct /3 width=8 by or3_intro2, ex4_4_intro/
-| #a #I #G #L #V #T #U #l #_ #j #H destruct
-| #G #L #V #T #U #l #_ #j #H destruct
-| #G #L #W #T #U #l #_ #j #H destruct
+| #G #L #K #W #V #U #i #d #HLK #HWV #HWU #j #H destruct /3 width=8 by or3_intro2, ex4_4_intro/
+| #a #I #G #L #V #T #U #d #_ #j #H destruct
+| #G #L #V #T #U #d #_ #j #H destruct
+| #G #L #W #T #U #d #_ #j #H destruct
]
qed-.
-lemma lstas_inv_lref1: ∀h,G,L,X,i,l. ⦃G, L⦄ ⊢ #i •*[h, l] X → ∨∨
- (∃∃K,V,W. ⬇[i] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V •*[h, l] W &
+lemma lstas_inv_lref1: ∀h,G,L,X,i,d. ⦃G, L⦄ ⊢ #i •*[h, d] X → ∨∨
+ (∃∃K,V,W. ⬇[i] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V •*[h, d] W &
⬆[0, i+1] W ≡ X
) |
(∃∃K,W,V. ⬇[i] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W •*[h, 0] V &
- X = #i & l = 0
+ X = #i & d = 0
) |
- (∃∃K,W,V,l0. ⬇[i] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W •*[h, l0] V &
- ⬆[0, i+1] V ≡ X & l = l0+1
+ (∃∃K,W,V,d0. ⬇[i] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W •*[h, d0] V &
+ ⬆[0, i+1] V ≡ X & d = d0+1
).
/2 width=3 by lstas_inv_lref1_aux/
qed-.
X = #i
).
#h #G #L #X #i #H elim (lstas_inv_lref1 … H) -H * /3 width=6 by ex3_3_intro, or_introl, or_intror/
-#K #W #V #l #_ #_ #_ <plus_n_Sm #H destruct
+#K #W #V #d #_ #_ #_ <plus_n_Sm #H destruct
qed-.
(* Basic_1: was just: sty0_gen_lref *)
-lemma lstas_inv_lref1_S: ∀h,G,L,X,i,l. ⦃G, L⦄ ⊢ #i •*[h, l+1] X →
- (∃∃K,V,W. ⬇[i] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V •*[h, l+1] W &
+lemma lstas_inv_lref1_S: ∀h,G,L,X,i,d. ⦃G, L⦄ ⊢ #i •*[h, d+1] X →
+ (∃∃K,V,W. ⬇[i] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V •*[h, d+1] W &
⬆[0, i+1] W ≡ X
) ∨
- (∃∃K,W,V. ⬇[i] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W •*[h, l] V &
+ (∃∃K,W,V. ⬇[i] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W •*[h, d] V &
⬆[0, i+1] V ≡ X
).
-#h #G #L #X #i #l #H elim (lstas_inv_lref1 … H) -H * /3 width=6 by ex3_3_intro, or_introl, or_intror/
+#h #G #L #X #i #d #H elim (lstas_inv_lref1 … H) -H * /3 width=6 by ex3_3_intro, or_introl, or_intror/
#K #W #V #_ #_ #_ <plus_n_Sm #H destruct
qed-.
-fact lstas_inv_gref1_aux: ∀h,G,L,T,U,l. ⦃G, L⦄ ⊢ T •*[h, l] U → ∀p0. T = §p0 → ⊥.
-#h #G #L #T #U #l * -G -L -T -U -l
-[ #G #L #l #k #p0 #H destruct
-| #G #L #K #V #W #U #i #l #_ #_ #_ #p0 #H destruct
+fact lstas_inv_gref1_aux: ∀h,G,L,T,U,d. ⦃G, L⦄ ⊢ T •*[h, d] U → ∀p0. T = §p0 → ⊥.
+#h #G #L #T #U #d * -G -L -T -U -d
+[ #G #L #d #k #p0 #H destruct
+| #G #L #K #V #W #U #i #d #_ #_ #_ #p0 #H destruct
| #G #L #K #W #V #i #_ #_ #p0 #H destruct
-| #G #L #K #W #V #U #i #l #_ #_ #_ #p0 #H destruct
-| #a #I #G #L #V #T #U #l #_ #p0 #H destruct
-| #G #L #V #T #U #l #_ #p0 #H destruct
-| #G #L #W #T #U #l #_ #p0 #H destruct
+| #G #L #K #W #V #U #i #d #_ #_ #_ #p0 #H destruct
+| #a #I #G #L #V #T #U #d #_ #p0 #H destruct
+| #G #L #V #T #U #d #_ #p0 #H destruct
+| #G #L #W #T #U #d #_ #p0 #H destruct
qed-.
-lemma lstas_inv_gref1: ∀h,G,L,X,p,l. ⦃G, L⦄ ⊢ §p •*[h, l] X → ⊥.
+lemma lstas_inv_gref1: ∀h,G,L,X,p,d. ⦃G, L⦄ ⊢ §p •*[h, d] X → ⊥.
/2 width=9 by lstas_inv_gref1_aux/
qed-.
-fact lstas_inv_bind1_aux: ∀h,G,L,T,U,l. ⦃G, L⦄ ⊢ T •*[h, l] U → ∀b,J,X,Y. T = ⓑ{b,J}Y.X →
- ∃∃Z. ⦃G, L.ⓑ{J}Y⦄ ⊢ X •*[h, l] Z & U = ⓑ{b,J}Y.Z.
-#h #G #L #T #U #l * -G -L -T -U -l
-[ #G #L #l #k #b #J #X #Y #H destruct
-| #G #L #K #V #W #U #i #l #_ #_ #_ #b #J #X #Y #H destruct
+fact lstas_inv_bind1_aux: ∀h,G,L,T,U,d. ⦃G, L⦄ ⊢ T •*[h, d] U → ∀b,J,X,Y. T = ⓑ{b,J}Y.X →
+ ∃∃Z. ⦃G, L.ⓑ{J}Y⦄ ⊢ X •*[h, d] Z & U = ⓑ{b,J}Y.Z.
+#h #G #L #T #U #d * -G -L -T -U -d
+[ #G #L #d #k #b #J #X #Y #H destruct
+| #G #L #K #V #W #U #i #d #_ #_ #_ #b #J #X #Y #H destruct
| #G #L #K #W #V #i #_ #_ #b #J #X #Y #H destruct
-| #G #L #K #W #V #U #i #l #_ #_ #_ #b #J #X #Y #H destruct
-| #a #I #G #L #V #T #U #l #HTU #b #J #X #Y #H destruct /2 width=3 by ex2_intro/
-| #G #L #V #T #U #l #_ #b #J #X #Y #H destruct
-| #G #L #W #T #U #l #_ #b #J #X #Y #H destruct
+| #G #L #K #W #V #U #i #d #_ #_ #_ #b #J #X #Y #H destruct
+| #a #I #G #L #V #T #U #d #HTU #b #J #X #Y #H destruct /2 width=3 by ex2_intro/
+| #G #L #V #T #U #d #_ #b #J #X #Y #H destruct
+| #G #L #W #T #U #d #_ #b #J #X #Y #H destruct
]
qed-.
(* Basic_1: was just: sty0_gen_bind *)
-lemma lstas_inv_bind1: ∀h,a,I,G,L,V,T,X,l. ⦃G, L⦄ ⊢ ⓑ{a,I}V.T •*[h, l] X →
- ∃∃U. ⦃G, L.ⓑ{I}V⦄ ⊢ T •*[h, l] U & X = ⓑ{a,I}V.U.
+lemma lstas_inv_bind1: ∀h,a,I,G,L,V,T,X,d. ⦃G, L⦄ ⊢ ⓑ{a,I}V.T •*[h, d] X →
+ ∃∃U. ⦃G, L.ⓑ{I}V⦄ ⊢ T •*[h, d] U & X = ⓑ{a,I}V.U.
/2 width=3 by lstas_inv_bind1_aux/
qed-.
-fact lstas_inv_appl1_aux: ∀h,G,L,T,U,l. ⦃G, L⦄ ⊢ T •*[h, l] U → ∀X,Y. T = ⓐY.X →
- ∃∃Z. ⦃G, L⦄ ⊢ X •*[h, l] Z & U = ⓐY.Z.
-#h #G #L #T #U #l * -G -L -T -U -l
-[ #G #L #l #k #X #Y #H destruct
-| #G #L #K #V #W #U #i #l #_ #_ #_ #X #Y #H destruct
+fact lstas_inv_appl1_aux: ∀h,G,L,T,U,d. ⦃G, L⦄ ⊢ T •*[h, d] U → ∀X,Y. T = ⓐY.X →
+ ∃∃Z. ⦃G, L⦄ ⊢ X •*[h, d] Z & U = ⓐY.Z.
+#h #G #L #T #U #d * -G -L -T -U -d
+[ #G #L #d #k #X #Y #H destruct
+| #G #L #K #V #W #U #i #d #_ #_ #_ #X #Y #H destruct
| #G #L #K #W #V #i #_ #_ #X #Y #H destruct
-| #G #L #K #W #V #U #i #l #_ #_ #_ #X #Y #H destruct
-| #a #I #G #L #V #T #U #l #_ #X #Y #H destruct
-| #G #L #V #T #U #l #HTU #X #Y #H destruct /2 width=3 by ex2_intro/
-| #G #L #W #T #U #l #_ #X #Y #H destruct
+| #G #L #K #W #V #U #i #d #_ #_ #_ #X #Y #H destruct
+| #a #I #G #L #V #T #U #d #_ #X #Y #H destruct
+| #G #L #V #T #U #d #HTU #X #Y #H destruct /2 width=3 by ex2_intro/
+| #G #L #W #T #U #d #_ #X #Y #H destruct
]
qed-.
(* Basic_1: was just: sty0_gen_appl *)
-lemma lstas_inv_appl1: ∀h,G,L,V,T,X,l. ⦃G, L⦄ ⊢ ⓐV.T •*[h, l] X →
- ∃∃U. ⦃G, L⦄ ⊢ T •*[h, l] U & X = ⓐV.U.
+lemma lstas_inv_appl1: ∀h,G,L,V,T,X,d. ⦃G, L⦄ ⊢ ⓐV.T •*[h, d] X →
+ ∃∃U. ⦃G, L⦄ ⊢ T •*[h, d] U & X = ⓐV.U.
/2 width=3 by lstas_inv_appl1_aux/
qed-.
-fact lstas_inv_cast1_aux: ∀h,G,L,T,U,l. ⦃G, L⦄ ⊢ T •*[h, l] U → ∀X,Y. T = ⓝY.X →
- ⦃G, L⦄ ⊢ X •*[h, l] U.
-#h #G #L #T #U #l * -G -L -T -U -l
-[ #G #L #l #k #X #Y #H destruct
-| #G #L #K #V #W #U #i #l #_ #_ #_ #X #Y #H destruct
+fact lstas_inv_cast1_aux: ∀h,G,L,T,U,d. ⦃G, L⦄ ⊢ T •*[h, d] U → ∀X,Y. T = ⓝY.X →
+ ⦃G, L⦄ ⊢ X •*[h, d] U.
+#h #G #L #T #U #d * -G -L -T -U -d
+[ #G #L #d #k #X #Y #H destruct
+| #G #L #K #V #W #U #i #d #_ #_ #_ #X #Y #H destruct
| #G #L #K #W #V #i #_ #_ #X #Y #H destruct
-| #G #L #K #W #V #U #i #l #_ #_ #_ #X #Y #H destruct
-| #a #I #G #L #V #T #U #l #_ #X #Y #H destruct
-| #G #L #V #T #U #l #_ #X #Y #H destruct
-| #G #L #W #T #U #l #HTU #X #Y #H destruct //
+| #G #L #K #W #V #U #i #d #_ #_ #_ #X #Y #H destruct
+| #a #I #G #L #V #T #U #d #_ #X #Y #H destruct
+| #G #L #V #T #U #d #_ #X #Y #H destruct
+| #G #L #W #T #U #d #HTU #X #Y #H destruct //
]
qed-.
(* Basic_1: was just: sty0_gen_cast *)
-lemma lstas_inv_cast1: ∀h,G,L,W,T,U,l. ⦃G, L⦄ ⊢ ⓝW.T •*[h, l] U → ⦃G, L⦄ ⊢ T •*[h, l] U.
+lemma lstas_inv_cast1: ∀h,G,L,W,T,U,d. ⦃G, L⦄ ⊢ ⓝW.T •*[h, d] U → ⦃G, L⦄ ⊢ T •*[h, d] U.
/2 width=4 by lstas_inv_cast1_aux/
qed-.
(* Properties on atomic arity assignment for terms **************************)
-lemma aaa_lstas: ∀h,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀l.
- ∃∃U. ⦃G, L⦄ ⊢ T •*[h, l] U & ⦃G, L⦄ ⊢ U ⁝ A.
+lemma aaa_lstas: ∀h,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀d.
+ ∃∃U. ⦃G, L⦄ ⊢ T •*[h, d] U & ⦃G, L⦄ ⊢ U ⁝ A.
#h #G #L #T #A #H elim H -G -L -T -A
[ /2 width=3 by ex2_intro/
-| * #G #L #K #V #B #i #HLK #HV #IHV #l
- [ elim (IHV l) -IHV #W
+| * #G #L #K #V #B #i #HLK #HV #IHV #d
+ [ elim (IHV d) -IHV #W
elim (lift_total W 0 (i+1))
lapply (drop_fwd_drop2 … HLK)
/3 width=10 by lstas_ldef, aaa_lift, ex2_intro/
- | @(nat_ind_plus … l) -l
+ | @(nat_ind_plus … d) -d
[ elim (IHV 0) -IHV /3 width=7 by lstas_zero, aaa_lref, ex2_intro/
- | #l #_ elim (IHV l) -IHV #W
+ | #d #_ elim (IHV d) -IHV #W
elim (lift_total W 0 (i+1))
lapply (drop_fwd_drop2 … HLK)
/3 width=10 by lstas_succ, aaa_lift, ex2_intro/
]
]
-| #a #G #L #V #T #B #A #HV #_ #_ #IHT #l elim (IHT l) -IHT
+| #a #G #L #V #T #B #A #HV #_ #_ #IHT #d elim (IHT d) -IHT
/3 width=7 by lstas_bind, aaa_abbr, ex2_intro/
-| #a #G #L #V #T #B #A #HV #_ #_ #IHT #l elim (IHT l) -IHT
+| #a #G #L #V #T #B #A #HV #_ #_ #IHT #d elim (IHT d) -IHT
/3 width=6 by lstas_bind, aaa_abst, ex2_intro/
-| #G #L #V #T #B #A #HV #_ #_ #IHT #l elim (IHT l) -IHT
+| #G #L #V #T #B #A #HV #_ #_ #IHT #d elim (IHT d) -IHT
/3 width=6 by lstas_appl, aaa_appl, ex2_intro/
-| #G #L #W #T #A #HW #_ #_ #IHT #l elim (IHT l) -IHT
+| #G #L #W #T #A #HW #_ #_ #IHT #d elim (IHT d) -IHT
/3 width=3 by lstas_cast, aaa_cast, ex2_intro/
]
qed-.
-lemma lstas_aaa_conf: ∀h,G,L,l. Conf3 … (aaa G L) (lstas h l G L).
-#h #G #L #l #A #T #HT #U #HTU
-elim (aaa_lstas h … HT l) -HT #X #HTX
+lemma lstas_aaa_conf: ∀h,G,L,d. Conf3 … (aaa G L) (lstas h d G L).
+#h #G #L #d #A #T #HT #U #HTU
+elim (aaa_lstas h … HT d) -HT #X #HTX
lapply (lstas_mono … HTX … HTU) -T //
qed-.
(* Properties on degree assignment for terms ********************************)
-lemma da_lstas: ∀h,g,G,L,T,l1. ⦃G, L⦄ ⊢ T ▪[h, g] l1 → ∀l2.
- ∃∃U. ⦃G, L⦄ ⊢ T •*[h, l2] U & ⦃G, L⦄ ⊢ U ▪[h, g] l1-l2.
-#h #g #G #L #T #l1 #H elim H -G -L -T -l1
+lemma da_lstas: ∀h,g,G,L,T,d1. ⦃G, L⦄ ⊢ T ▪[h, g] d1 → ∀d2.
+ ∃∃U. ⦃G, L⦄ ⊢ T •*[h, d2] U & ⦃G, L⦄ ⊢ U ▪[h, g] d1-d2.
+#h #g #G #L #T #d1 #H elim H -G -L -T -d1
[ /4 width=3 by da_sort, deg_iter, ex2_intro/
-| #G #L #K #V #i #l1 #HLK #_ #IHV #l2
- elim (IHV l2) -IHV #W
+| #G #L #K #V #i #d1 #HLK #_ #IHV #d2
+ elim (IHV d2) -IHV #W
elim (lift_total W 0 (i+1))
lapply (drop_fwd_drop2 … HLK)
/3 width=10 by lstas_ldef, da_lift, ex2_intro/
-| #G #L #K #W #i #l1 #HLK #HW #IHW #l2 @(nat_ind_plus … l2) -l2
+| #G #L #K #W #i #d1 #HLK #HW #IHW #d2 @(nat_ind_plus … d2) -d2
[ elim (IHW 0) -IHW /3 width=6 by lstas_zero, da_ldec, ex2_intro/
- | #l #_ elim (IHW l) -IHW #V
+ | #d #_ elim (IHW d) -IHW #V
elim (lift_total V 0 (i+1))
lapply (drop_fwd_drop2 … HLK)
/3 width=10 by lstas_succ, da_lift, ex2_intro/
]
-| #a #I #G #L #V #T #l1 #_ #IHT #l2 elim (IHT … l2) -IHT
+| #a #I #G #L #V #T #d1 #_ #IHT #d2 elim (IHT … d2) -IHT
/3 width=6 by lstas_bind, da_bind, ex2_intro/
-| * #G #L #V #T #l1 #_ #IHT #l2 elim (IHT … l2) -IHT
+| * #G #L #V #T #d1 #_ #IHT #d2 elim (IHT … d2) -IHT
/3 width=5 by lstas_appl, lstas_cast, da_flat, ex2_intro/
]
qed-.
-lemma lstas_da_conf: ∀h,g,G,L,T,U,l2. ⦃G, L⦄ ⊢ T •*[h, l2] U →
- ∀l1. ⦃G, L⦄ ⊢ T ▪[h, g] l1 → ⦃G, L⦄ ⊢ U ▪[h, g] l1-l2.
-#h #g #G #L #T #U #l2 #HTU #l1 #HT
-elim (da_lstas … HT l2) -HT #X #HTX
+lemma lstas_da_conf: ∀h,g,G,L,T,U,d2. ⦃G, L⦄ ⊢ T •*[h, d2] U →
+ ∀d1. ⦃G, L⦄ ⊢ T ▪[h, g] d1 → ⦃G, L⦄ ⊢ U ▪[h, g] d1-d2.
+#h #g #G #L #T #U #d2 #HTU #d1 #HT
+elim (da_lstas … HT d2) -HT #X #HTX
lapply (lstas_mono … HTX … HTU) -T //
qed-.
(* inversion lemmas on degree assignment for terms **************************)
-lemma lstas_inv_da: ∀h,g,G,L,T,U,l2. ⦃G, L⦄ ⊢ T •*[h, l2] U →
- ∃∃l1. ⦃G, L⦄ ⊢ T ▪[h, g] l1 & ⦃G, L⦄ ⊢ U ▪[h, g] l1-l2.
-#h #g #G #L #T #U #l2 #H elim H -G -L -T -U -l2
-[ #G #L #l2 #k elim (deg_total h g k) /4 width=3 by da_sort, deg_iter, ex2_intro/
-| #G #L #K #V #W #U #i #l2 #HLK #_ #HWU *
+lemma lstas_inv_da: ∀h,g,G,L,T,U,d2. ⦃G, L⦄ ⊢ T •*[h, d2] U →
+ ∃∃d1. ⦃G, L⦄ ⊢ T ▪[h, g] d1 & ⦃G, L⦄ ⊢ U ▪[h, g] d1-d2.
+#h #g #G #L #T #U #d2 #H elim H -G -L -T -U -d2
+[ #G #L #d2 #k elim (deg_total h g k) /4 width=3 by da_sort, deg_iter, ex2_intro/
+| #G #L #K #V #W #U #i #d2 #HLK #_ #HWU *
lapply (drop_fwd_drop2 … HLK) /3 width=10 by da_lift, da_ldef, ex2_intro/
| #G #L #K #W #V #i #HLK #_ * /3 width=6 by da_ldec, ex2_intro/
-| #G #L #K #W #V #U #i #l2 #HLK #_ #HVU *
+| #G #L #K #W #V #U #i #d2 #HLK #_ #HVU *
lapply (drop_fwd_drop2 … HLK) /3 width=10 by da_lift, da_ldec, ex2_intro/
-| #a #I #G #L #V #T #U #l2 #_ * /3 width=3 by da_bind, ex2_intro/
-| #G #L #V #T #U #l2 #_ * /3 width=3 by da_flat, ex2_intro/
-| #G #L #W #T #U #l2 #_ * /3 width=3 by da_flat, ex2_intro/
+| #a #I #G #L #V #T #U #d2 #_ * /3 width=3 by da_bind, ex2_intro/
+| #G #L #V #T #U #d2 #_ * /3 width=3 by da_flat, ex2_intro/
+| #G #L #W #T #U #d2 #_ * /3 width=3 by da_flat, ex2_intro/
]
qed-.
-lemma lstas_inv_da_ge: ∀h,G,L,T,U,l2,l. ⦃G, L⦄ ⊢ T •*[h, l2] U →
- ∃∃g,l1. ⦃G, L⦄ ⊢ T ▪[h, g] l1 & ⦃G, L⦄ ⊢ U ▪[h, g] l1-l2 & l ≤ l1.
-#h #G #L #T #U #l2 #l #H elim H -G -L -T -U -l2
+lemma lstas_inv_da_ge: ∀h,G,L,T,U,d2,d. ⦃G, L⦄ ⊢ T •*[h, d2] U →
+ ∃∃g,d1. ⦃G, L⦄ ⊢ T ▪[h, g] d1 & ⦃G, L⦄ ⊢ U ▪[h, g] d1-d2 & d ≤ d1.
+#h #G #L #T #U #d2 #d #H elim H -G -L -T -U -d2
[ /4 width=5 by da_sort, deg_iter, ex3_2_intro/
-| #G #L #K #V #W #U #i #l2 #HLK #_ #HWU *
+| #G #L #K #V #W #U #i #d2 #HLK #_ #HWU *
lapply (drop_fwd_drop2 … HLK) /3 width=10 by da_lift, da_ldef, ex3_2_intro/
| #G #L #K #W #V #i #HLK #_ *
- #g #l1 #HW #HV #Hl1 /4 width=6 by da_ldec, lt_to_le, le_S_S, ex3_2_intro/
-| #G #L #K #W #V #U #i #l2 #HLK #_ #HVU *
+ #g #d1 #HW #HV #Hd1 /4 width=6 by da_ldec, lt_to_le, le_S_S, ex3_2_intro/
+| #G #L #K #W #V #U #i #d2 #HLK #_ #HVU *
lapply (drop_fwd_drop2 … HLK)
/4 width=10 by da_lift, da_ldec, lt_to_le, le_S_S, ex3_2_intro/
-| #a #I #G #L #V #T #U #l2 #_ * /3 width=5 by da_bind, ex3_2_intro/
-| #G #L #V #T #U #l2 #_ * /3 width=5 by da_flat, ex3_2_intro/
-| #G #L #W #T #U #l2 #_ * /3 width=5 by da_flat, ex3_2_intro/
+| #a #I #G #L #V #T #U #d2 #_ * /3 width=5 by da_bind, ex3_2_intro/
+| #G #L #V #T #U #d2 #_ * /3 width=5 by da_flat, ex3_2_intro/
+| #G #L #W #T #U #d2 #_ * /3 width=5 by da_flat, ex3_2_intro/
]
qed-.
(* Advanced inversion lemmas ************************************************)
-lemma lstas_inv_refl_pos: ∀h,G,L,T,l. ⦃G, L⦄ ⊢ T •*[h, l+1] T → ⊥.
-#h #G #L #T #l2 #H elim (lstas_inv_da_ge … (l2+1) H) -H
-#g #l1 #HT1 #HT12 #Hl21 lapply (da_mono … HT1 … HT12) -h -G -L -T
+lemma lstas_inv_refl_pos: ∀h,G,L,T,d. ⦃G, L⦄ ⊢ T •*[h, d+1] T → ⊥.
+#h #G #L #T #d2 #H elim (lstas_inv_da_ge … (d2+1) H) -H
+#g #d1 #HT1 #HT12 #Hd21 lapply (da_mono … HT1 … HT12) -h -G -L -T
#H elim (discr_x_minus_xy … H) -H
[ #H destruct /2 width=3 by le_plus_xSy_O_false/
-| -l1 <plus_n_Sm #H destruct
+| -d1 <plus_n_Sm #H destruct
]
qed-.
(* Properties on relocation *************************************************)
(* Basic_1: was just: sty0_lift *)
-lemma lstas_lift: ∀h,G,l. l_liftable (lstas h G l).
-#h #G #l #L1 #T1 #U1 #H elim H -G -L1 -T1 -U1 -l
-[ #G #L1 #l #k #L2 #s #d #e #HL21 #X1 #H1 #X2 #H2
+lemma lstas_lift: ∀h,G,d. d_liftable (lstas h G d).
+#h #G #d #L1 #T1 #U1 #H elim H -G -L1 -T1 -U1 -d
+[ #G #L1 #d #k #L2 #s #l #m #HL21 #X1 #H1 #X2 #H2
>(lift_inv_sort1 … H1) -X1
>(lift_inv_sort1 … H2) -X2 //
-| #G #L1 #K1 #V1 #W1 #W #i #l #HLK1 #_ #HW1 #IHVW1 #L2 #s #d #e #HL21 #X #H #U2 #HWU2
- elim (lift_inv_lref1 … H) * #Hid #H destruct
+| #G #L1 #K1 #V1 #W1 #W #i #d #HLK1 #_ #HW1 #IHVW1 #L2 #s #l #m #HL21 #X #H #U2 #HWU2
+ elim (lift_inv_lref1 … H) * #Hil #H destruct
[ elim (lift_trans_ge … HW1 … HWU2) -W // #W2 #HW12 #HWU2
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
+ elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hil #K2 #V2 #HK21 #HV12 #H destruct
/3 width=9 by lstas_ldef/
| lapply (lift_trans_be … HW1 … HWU2 ? ?) -W /2 width=1 by le_S/ #HW1U2
lapply (drop_trans_ge … HL21 … HLK1 ?) -L1 /3 width=9 by lstas_ldef, drop_inv_gen/
]
-| #G #L1 #K1 #V1 #W1 #i #HLK1 #_ #IHVW1 #L2 #s #d #e #HL21 #X #H #U2 #HWU2
+| #G #L1 #K1 #V1 #W1 #i #HLK1 #_ #IHVW1 #L2 #s #l #m #HL21 #X #H #U2 #HWU2
>(lift_mono … HWU2 … H) -U2
- elim (lift_inv_lref1 … H) * #Hid #H destruct
- [ elim (lift_total W1 (d-i-1) e) #W2 #HW12
+ elim (lift_inv_lref1 … H) * #Hil #H destruct
+ [ elim (lift_total W1 (l-i-1) m) #W2 #HW12
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
+ elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hil #K2 #V2 #HK21 #HV12 #H destruct
/3 width=10 by lstas_zero/
| lapply (drop_trans_ge … HL21 … HLK1 ?) -L1
/3 width=10 by lstas_zero, drop_inv_gen/
]
-| #G #L1 #K1 #W1 #V1 #W #i #l #HLK1 #_ #HW1 #IHWV1 #L2 #s #d #e #HL21 #X #H #U2 #HWU2
- elim (lift_inv_lref1 … H) * #Hid #H destruct
+| #G #L1 #K1 #W1 #V1 #W #i #d #HLK1 #_ #HW1 #IHWV1 #L2 #s #l #m #HL21 #X #H #U2 #HWU2
+ elim (lift_inv_lref1 … H) * #Hil #H destruct
[ elim (lift_trans_ge … HW1 … HWU2) -W // <minus_plus #W #HW1 #HWU2
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
+ elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hil #K2 #W2 #HK21 #HW12 #H destruct
/3 width=9 by lstas_succ/
| lapply (lift_trans_be … HW1 … HWU2 ? ?) -W /2 width=1 by le_S/ #HW1U2
lapply (drop_trans_ge … HL21 … HLK1 ?) -L1 /3 width=9 by lstas_succ, drop_inv_gen/
]
-| #a #I #G #L1 #V1 #T1 #U1 #l #_ #IHTU1 #L2 #s #d #e #HL21 #X1 #H1 #X2 #H2
+| #a #I #G #L1 #V1 #T1 #U1 #d #_ #IHTU1 #L2 #s #l #m #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 lstas_bind, drop_skip/
-| #G #L1 #V1 #T1 #U1 #l #_ #IHTU1 #L2 #s #d #e #HL21 #X1 #H1 #X2 #H2
+| #G #L1 #V1 #T1 #U1 #d #_ #IHTU1 #L2 #s #l #m #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
lapply (lift_mono … H1 … HV12) -H1 #H destruct /4 width=6 by lstas_appl/
-| #G #L1 #W1 #T1 #U1 #l #_ #IHTU1 #L2 #s #d #e #HL21 #X #H #U2 #HU12
+| #G #L1 #W1 #T1 #U1 #d #_ #IHTU1 #L2 #s #l #m #HL21 #X #H #U2 #HU12
elim (lift_inv_flat1 … H) -H #W2 #T2 #_ #HT12 #H destruct /3 width=6 by lstas_cast/
]
qed.
(* Inversion lemmas on relocation *******************************************)
(* Note: apparently this was missing in basic_1 *)
-lemma lstas_inv_lift1: ∀h,G,l. l_deliftable_sn (lstas h G l).
-#h #G #l #L2 #T2 #U2 #H elim H -G -L2 -T2 -U2 -l
-[ #G #L2 #l #k #L1 #s #d #e #_ #X #H
+lemma lstas_inv_lift1: ∀h,G,d. d_deliftable_sn (lstas h G d).
+#h #G #d #L2 #T2 #U2 #H elim H -G -L2 -T2 -U2 -d
+[ #G #L2 #d #k #L1 #s #l #m #_ #X #H
>(lift_inv_sort2 … H) -X /2 width=3 by lstas_sort, lift_sort, ex2_intro/
-| #G #L2 #K2 #V2 #W2 #W #i #l #HLK2 #HVW2 #HW2 #IHVW2 #L1 #s #d #e #HL21 #X #H
- elim (lift_inv_lref2 … H) * #Hid #H destruct [ -HVW2 | -IHVW2 ]
+| #G #L2 #K2 #V2 #W2 #W #i #d #HLK2 #HVW2 #HW2 #IHVW2 #L1 #s #l #m #HL21 #X #H
+ elim (lift_inv_lref2 … H) * #Hil #H destruct [ -HVW2 | -IHVW2 ]
[ 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 lstas_ldef, ex2_intro/
| 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/
+ elim (le_inv_plus_l … Hil) -Hil #Hlim #mi
+ elim (lift_split … HW2 l (i-m+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 lstas_ldef, le_S, ex2_intro/
]
-| #G #L2 #K2 #W2 #V2 #i #HLK2 #HWV2 #IHWV2 #L1 #s #d #e #HL21 #X #H
- elim (lift_inv_lref2 … H) * #Hid #H destruct [ -HWV2 | -IHWV2 ]
+| #G #L2 #K2 #W2 #V2 #i #HLK2 #HWV2 #IHWV2 #L1 #s #l #m #HL21 #X #H
+ elim (lift_inv_lref2 … H) * #Hil #H destruct [ -HWV2 | -IHWV2 ]
[ elim (drop_conf_lt … HL21 … HLK2) -L2 // #K1 #W1 #HLK1 #HK21 #HW12
elim (IHWV2 … HK21 … HW12) -K2
/3 width=5 by lstas_zero, lift_lref_lt, ex2_intro/
| lapply (drop_conf_ge … HL21 … HLK2 ?) -L2
/3 width=5 by lstas_zero, lift_lref_ge_minus, ex2_intro/
]
-| #G #L2 #K2 #W2 #V2 #W #i #l #HLK2 #HWV2 #HW2 #IHWV2 #L1 #s #d #e #HL21 #X #H
- elim (lift_inv_lref2 … H) * #Hid #H destruct [ -HWV2 | -IHWV2 ]
+| #G #L2 #K2 #W2 #V2 #W #i #d #HLK2 #HWV2 #HW2 #IHWV2 #L1 #s #l #m #HL21 #X #H
+ elim (lift_inv_lref2 … H) * #Hil #H destruct [ -HWV2 | -IHWV2 ]
[ elim (drop_conf_lt … HL21 … HLK2) -L2 // #K1 #W1 #HLK1 #HK21 #HW12
elim (IHWV2 … HK21 … HW12) -K2 #V1 #HV12 #HWV1
elim (lift_trans_le … HV12 … HW2) -W2 // >minus_plus <plus_minus_m_m /3 width=8 by lstas_succ, ex2_intro/
| 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/
+ elim (le_inv_plus_l … Hil) -Hil #Hlim #mi
+ elim (lift_split … HW2 l (i-m+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 lstas_succ, le_S, ex2_intro/
]
-| #a #I #G #L2 #V2 #T2 #U2 #l #_ #IHTU2 #L1 #s #d #e #HL21 #X #H
+| #a #I #G #L2 #V2 #T2 #U2 #d #_ #IHTU2 #L1 #s #l #m #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 lstas_bind, drop_skip, lift_bind, ex2_intro/
-| #G #L2 #V2 #T2 #U2 #l #_ #IHTU2 #L1 #s #d #e #HL21 #X #H
+| #G #L2 #V2 #T2 #U2 #d #_ #IHTU2 #L1 #s #l #m #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 lstas_appl, lift_flat, ex2_intro/
-| #G #L2 #W2 #T2 #U2 #l #_ #IHTU2 #L1 #s #d #e #HL21 #X #H
+| #G #L2 #W2 #T2 #U2 #d #_ #IHTU2 #L1 #s #l #m #HL21 #X #H
elim (lift_inv_flat2 … H) -H #W1 #T1 #_ #HT12 #H destruct
elim (IHTU2 … HL21 … HT12) -L2 -HT12 /3 width=3 by lstas_cast, ex2_intro/
]
(* Advanced inversion lemmas ************************************************)
-lemma lstas_split_aux: ∀h,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 •*[h, l] T2 → ∀l1,l2. l = l1 + l2 →
- ∃∃T. ⦃G, L⦄ ⊢ T1 •*[h, l1] T & ⦃G, L⦄ ⊢ T •*[h, l2] T2.
-#h #G #L #T1 #T2 #l #H elim H -G -L -T1 -T2 -l
-[ #G #L #l #k #l1 #l2 #H destruct
+lemma lstas_split_aux: ∀h,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 •*[h, d] T2 → ∀d1,d2. d = d1 + d2 →
+ ∃∃T. ⦃G, L⦄ ⊢ T1 •*[h, d1] T & ⦃G, L⦄ ⊢ T •*[h, d2] T2.
+#h #G #L #T1 #T2 #d #H elim H -G -L -T1 -T2 -d
+[ #G #L #d #k #d1 #d2 #H destruct
>commutative_plus >iter_plus /2 width=3 by lstas_sort, ex2_intro/
-| #G #L #K #V1 #V2 #U2 #i #l #HLK #_ #VU2 #IHV12 #l1 #l2 #H destruct
- elim (IHV12 l1 l2) -IHV12 // #V
+| #G #L #K #V1 #V2 #U2 #i #d #HLK #_ #VU2 #IHV12 #d1 #d2 #H destruct
+ elim (IHV12 d1 d2) -IHV12 // #V
elim (lift_total V 0 (i+1))
lapply (drop_fwd_drop2 … HLK)
/3 width=12 by lstas_lift, lstas_ldef, ex2_intro/
-| #G #L #K #W1 #W2 #i #HLK #HW12 #_ #l1 #l2 #H
+| #G #L #K #W1 #W2 #i #HLK #HW12 #_ #d1 #d2 #H
elim (zero_eq_plus … H) -H #H1 #H2 destruct
/3 width=5 by lstas_zero, ex2_intro/
-| #G #L #K #W1 #W2 #U2 #i #l #HLK #HW12 #HWU2 #IHW12 #l1 @(nat_ind_plus … l1) -l1
- [ #l2 normalize #H destruct
- elim (IHW12 0 l) -IHW12 //
+| #G #L #K #W1 #W2 #U2 #i #d #HLK #HW12 #HWU2 #IHW12 #d1 @(nat_ind_plus … d1) -d1
+ [ #d2 normalize #H destruct
+ elim (IHW12 0 d) -IHW12 //
lapply (drop_fwd_drop2 … HLK)
/3 width=8 by lstas_succ, lstas_zero, ex2_intro/
- | #l1 #_ #l2 <plus_plus_comm_23 #H lapply (injective_plus_l … H) -H #H
- elim (IHW12 … H) -l #W
+ | #d1 #_ #d2 <plus_plus_comm_23 #H lapply (injective_plus_l … H) -H #H
+ elim (IHW12 … H) -d #W
elim (lift_total W 0 (i+1))
lapply (drop_fwd_drop2 … HLK)
/3 width=12 by lstas_lift, lstas_succ, ex2_intro/
]
-| #a #I #G #L #V #T #U #l #_ #IHTU #l1 #l2 #H
- elim (IHTU … H) -l /3 width=3 by lstas_bind, ex2_intro/
-| #G #L #V #T #U #l #_ #IHTU #l1 #l2 #H
- elim (IHTU … H) -l /3 width=3 by lstas_appl, ex2_intro/
-| #G #L #W #T #U #l #_ #IHTU #l1 #l2 #H
- elim (IHTU … H) -l /3 width=3 by lstas_cast, ex2_intro/
+| #a #I #G #L #V #T #U #d #_ #IHTU #d1 #d2 #H
+ elim (IHTU … H) -d /3 width=3 by lstas_bind, ex2_intro/
+| #G #L #V #T #U #d #_ #IHTU #d1 #d2 #H
+ elim (IHTU … H) -d /3 width=3 by lstas_appl, ex2_intro/
+| #G #L #W #T #U #d #_ #IHTU #d1 #d2 #H
+ elim (IHTU … H) -d /3 width=3 by lstas_cast, ex2_intro/
]
qed-.
-lemma lstas_split: ∀h,G,L,T1,T2,l1,l2. ⦃G, L⦄ ⊢ T1 •*[h, l1 + l2] T2 →
- ∃∃T. ⦃G, L⦄ ⊢ T1 •*[h, l1] T & ⦃G, L⦄ ⊢ T •*[h, l2] T2.
+lemma lstas_split: ∀h,G,L,T1,T2,d1,d2. ⦃G, L⦄ ⊢ T1 •*[h, d1 + d2] T2 →
+ ∃∃T. ⦃G, L⦄ ⊢ T1 •*[h, d1] T & ⦃G, L⦄ ⊢ T •*[h, d2] T2.
/2 width=3 by lstas_split_aux/ qed-.
(* Advanced properties ******************************************************)
-lemma lstas_lstas: ∀h,G,L,T,T1,l1. ⦃G, L⦄ ⊢ T •*[h, l1] T1 →
- ∀l2. ∃T2. ⦃G, L⦄ ⊢ T •*[h, l2] T2.
-#h #G #L #T #T1 #l1 #H elim H -G -L -T -T1 -l1
+lemma lstas_lstas: ∀h,G,L,T,T1,d1. ⦃G, L⦄ ⊢ T •*[h, d1] T1 →
+ ∀d2. ∃T2. ⦃G, L⦄ ⊢ T •*[h, d2] T2.
+#h #G #L #T #T1 #d1 #H elim H -G -L -T -T1 -d1
[ /2 width=2 by lstas_sort, ex_intro/
-| #G #L #K #V #V1 #U1 #i #l1 #HLK #_ #HVU1 #IHV1 #l2
- elim (IHV1 l2) -IHV1 #V2
+| #G #L #K #V #V1 #U1 #i #d1 #HLK #_ #HVU1 #IHV1 #d2
+ elim (IHV1 d2) -IHV1 #V2
elim (lift_total V2 0 (i+1))
/3 width=7 by ex_intro, lstas_ldef/
-| #G #L #K #W #W1 #i #HLK #HW1 #IHW1 #l2
- @(nat_ind_plus … l2) -l2 /3 width=5 by lstas_zero, ex_intro/
- #l2 #_ elim (IHW1 l2) -IHW1 #W2
+| #G #L #K #W #W1 #i #HLK #HW1 #IHW1 #d2
+ @(nat_ind_plus … d2) -d2 /3 width=5 by lstas_zero, ex_intro/
+ #d2 #_ elim (IHW1 d2) -IHW1 #W2
elim (lift_total W2 0 (i+1))
/3 width=7 by lstas_succ, ex_intro/
-| #G #L #K #W #W1 #U1 #i #l #HLK #_ #_ #IHW1 #l2
- @(nat_ind_plus … l2) -l2
+| #G #L #K #W #W1 #U1 #i #d #HLK #_ #_ #IHW1 #d2
+ @(nat_ind_plus … d2) -d2
[ elim (IHW1 0) -IHW1 /3 width=5 by lstas_zero, ex_intro/
- | #l2 #_ elim (IHW1 l2) -IHW1
+ | #d2 #_ elim (IHW1 d2) -IHW1
#W2 elim (lift_total W2 0 (i+1)) /3 width=7 by ex_intro, lstas_succ/
]
-| #a #I #G #L #V #T #U #l #_ #IHTU #l2
- elim (IHTU l2) -IHTU /3 width=2 by lstas_bind, ex_intro/
-| #G #L #V #T #U #l #_ #IHTU #l2
- elim (IHTU l2) -IHTU /3 width=2 by lstas_appl, ex_intro/
-| #G #L #W #T #U #l #_ #IHTU #l2
- elim (IHTU l2) -IHTU /3 width=2 by lstas_cast, ex_intro/
+| #a #I #G #L #V #T #U #d #_ #IHTU #d2
+ elim (IHTU d2) -IHTU /3 width=2 by lstas_bind, ex_intro/
+| #G #L #V #T #U #d #_ #IHTU #d2
+ elim (IHTU d2) -IHTU /3 width=2 by lstas_appl, ex_intro/
+| #G #L #W #T #U #d #_ #IHTU #d2
+ elim (IHTU d2) -IHTU /3 width=2 by lstas_cast, ex_intro/
]
qed-.
(* Properties on lazy sn pointwise extensions *******************************)
-lemma lstas_llpx_sn_conf: ∀R. (∀L. reflexive … (R L)) → l_liftable R →
- ∀h,G,l. s_r_confluent1 … (lstas h l G) (llpx_sn R 0).
-#R #H1R #H2R #h #G #l #Ls #T1 #T2 #H elim H -G -Ls -T1 -T2 -l
+lemma lstas_llpx_sn_conf: ∀R. (∀L. reflexive … (R L)) → d_liftable R →
+ ∀h,G,d. s_r_confluent1 … (lstas h d G) (llpx_sn R 0).
+#R #H1R #H2R #h #G #d #Ls #T1 #T2 #H elim H -G -Ls -T1 -T2 -d
[ /3 width=4 by llpx_sn_fwd_length, llpx_sn_sort/
-| #G #Ls #Ks #V1s #V2s #W2s #i #l #HLKs #_ #HVW2s #IHV12s #Ld #H elim (llpx_sn_inv_lref_ge_sn … H … HLKs) // -H
- #Kd #V1d #HLKd #HV1s #HV1sd
+| #G #Ls #Ks #V1s #V2s #W2s #i #d #HLKs #_ #HVW2s #IHV12s #Ld #H elim (llpx_sn_inv_lref_ge_sn … H … HLKs) // -H
+ #Kd #V1l #HLKd #HV1s #HV1sd
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 #V2s #W2s #i #l #HLKs #_ #HVW2s #IHV12s #Ld #H elim (llpx_sn_inv_lref_ge_sn … H … HLKs) // -H
- #Kd #V1d #HLKd #HV1s #HV1sd
+| #G #Ls #Ks #V1s #V2s #W2s #i #d #HLKs #_ #HVW2s #IHV12s #Ld #H elim (llpx_sn_inv_lref_ge_sn … H … HLKs) // -H
+ #Kd #V1l #HLKd #HV1s #HV1sd
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 #V #T1 #T2 #l #_ #IHT12 #Ld #H elim (llpx_sn_inv_bind_O … H) -H
+| #a #I #G #Ls #V #T1 #T2 #d #_ #IHT12 #Ld #H elim (llpx_sn_inv_bind_O … H) -H
/4 width=5 by llpx_sn_bind_repl_SO, llpx_sn_bind/
-| #G #Ls #V #T1 #T2 #l #_ #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
+| #G #Ls #V #T1 #T2 #d #_ #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
/3 width=1 by llpx_sn_flat/
-| #G #Ls #V #T1 #T2 #l #_ #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
+| #G #Ls #V #T1 #T2 #d #_ #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
/3 width=1 by llpx_sn_flat/
]
qed-.
(* Main properties **********************************************************)
-theorem lstas_trans: ∀h,G,L,T1,T,l1. ⦃G, L⦄ ⊢ T1 •*[h, l1] T →
- ∀T2,l2. ⦃G, L⦄ ⊢ T •*[h, l2] T2 → ⦃G, L⦄ ⊢ T1 •*[h, l1+l2] T2.
-#h #G #L #T1 #T #l1 #H elim H -G -L -T1 -T -l1
-[ #G #L #l1 #k #X #l2 #H >(lstas_inv_sort1 … H) -X
+theorem lstas_trans: ∀h,G,L,T1,T,d1. ⦃G, L⦄ ⊢ T1 •*[h, d1] T →
+ ∀T2,d2. ⦃G, L⦄ ⊢ T •*[h, d2] T2 → ⦃G, L⦄ ⊢ T1 •*[h, d1+d2] T2.
+#h #G #L #T1 #T #d1 #H elim H -G -L -T1 -T -d1
+[ #G #L #d1 #k #X #d2 #H >(lstas_inv_sort1 … H) -X
<iter_plus /2 width=1 by lstas_sort/
-| #G #L #K #V1 #V #U #i #l1 #HLK #_ #HVU #IHV1 #U2 #l2 #HU2
+| #G #L #K #V1 #V #U #i #d1 #HLK #_ #HVU #IHV1 #U2 #d2 #HU2
lapply (drop_fwd_drop2 … HLK) #H0
elim (lstas_inv_lift1 … HU2 … H0 … HVU)
/3 width=6 by lstas_ldef/
| //
-| #G #L #K #W1 #W #U #i #l1 #HLK #_ #HWU #IHW1 #U2 #l2 #HU2
+| #G #L #K #W1 #W #U #i #d1 #HLK #_ #HWU #IHW1 #U2 #d2 #HU2
lapply (drop_fwd_drop2 … HLK) #H0
elim (lstas_inv_lift1 … HU2 … H0 … HWU)
/3 width=6 by lstas_succ/
-| #a #I #G #L #V #T1 #T #l1 #_ #IHT1 #X #l2 #H
+| #a #I #G #L #V #T1 #T #d1 #_ #IHT1 #X #d2 #H
elim (lstas_inv_bind1 … H) -H #T2 #HT2 #H destruct
/3 width=1 by lstas_bind/
-| #G #L #V #T1 #T #l1 #_ #IHT1 #X #l2 #H
+| #G #L #V #T1 #T #d1 #_ #IHT1 #X #d2 #H
elim (lstas_inv_appl1 … H) -H #T2 #HT2 #H destruct
/3 width=1 by lstas_appl/
| /3 width=1 by lstas_cast/
qed-.
(* Note: apparently this was missing in basic_1 *)
-theorem lstas_mono: ∀h,G,L,l. singlevalued … (lstas h l G L).
-#h #G #L #l #T #T1 #H elim H -G -L -T -T1 -l
-[ #G #L #l #k #X #H >(lstas_inv_sort1 … H) -X //
-| #G #L #K #V #V1 #U1 #i #l #HLK #_ #HVU1 #IHV1 #X #H
+theorem lstas_mono: ∀h,G,L,d. singlevalued … (lstas h d G L).
+#h #G #L #d #T #T1 #H elim H -G -L -T -T1 -d
+[ #G #L #d #k #X #H >(lstas_inv_sort1 … H) -X //
+| #G #L #K #V #V1 #U1 #i #d #HLK #_ #HVU1 #IHV1 #X #H
elim (lstas_inv_lref1 … H) -H *
- #K0 #V0 #W0 [3: #l0 ] #HLK0
+ #K0 #V0 #W0 [3: #d0 ] #HLK0
lapply (drop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct
#HVW0 #HX lapply (IHV1 … HVW0) -IHV1 -HVW0 #H destruct
/2 width=5 by lift_mono/
elim (lstas_inv_lref1_O … H) -H *
#K0 #V0 #W0 #HLK0
lapply (drop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct //
-| #G #L #K #W #W1 #U1 #i #l #HLK #_ #HWU1 #IHWV #X #H
+| #G #L #K #W #W1 #U1 #i #d #HLK #_ #HWU1 #IHWV #X #H
elim (lstas_inv_lref1_S … H) -H * #K0 #W0 #V0 #HLK0
lapply (drop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct
#HW0 #HX lapply (IHWV … HW0) -IHWV -HW0 #H destruct
/2 width=5 by lift_mono/
-| #a #I #G #L #V #T #U1 #l #_ #IHTU1 #X #H
+| #a #I #G #L #V #T #U1 #d #_ #IHTU1 #X #H
elim (lstas_inv_bind1 … H) -H #U2 #HTU2 #H destruct /3 width=1 by eq_f/
-| #G #L #V #T #U1 #l #_ #IHTU1 #X #H
+| #G #L #V #T #U1 #d #_ #IHTU1 #X #H
elim (lstas_inv_appl1 … H) -H #U2 #HTU2 #H destruct /3 width=1 by eq_f/
-| #G #L #W #T #U1 #l #_ #IHTU1 #U2 #H
+| #G #L #W #T #U1 #d #_ #IHTU1 #U2 #H
lapply (lstas_inv_cast1 … H) -H /2 width=1 by/
]
qed-.
(* Advanced inversion lemmas ************************************************)
(* Basic_1: was just: sty0_correct *)
-lemma lstas_correct: ∀h,G,L,T1,T,l1. ⦃G, L⦄ ⊢ T1 •*[h, l1] T →
- ∀l2. ∃T2. ⦃G, L⦄ ⊢ T •*[h, l2] T2.
-#h #G #L #T1 #T #l1 #H elim H -G -L -T1 -T -l1
+lemma lstas_correct: ∀h,G,L,T1,T,d1. ⦃G, L⦄ ⊢ T1 •*[h, d1] T →
+ ∀d2. ∃T2. ⦃G, L⦄ ⊢ T •*[h, d2] T2.
+#h #G #L #T1 #T #d1 #H elim H -G -L -T1 -T -d1
[ /2 width=2 by lstas_sort, ex_intro/
-| #G #L #K #V1 #V #U #i #l #HLK #_ #HVU #IHV1 #l2
- elim (IHV1 l2) -IHV1 #V2
+| #G #L #K #V1 #V #U #i #d #HLK #_ #HVU #IHV1 #d2
+ elim (IHV1 d2) -IHV1 #V2
elim (lift_total V2 0 (i+1))
lapply (drop_fwd_drop2 … HLK) -HLK
/3 width=11 by ex_intro, lstas_lift/
-| #G #L #K #W1 #W #i #HLK #HW1 #IHW1 #l2
- @(nat_ind_plus … l2) -l2 /3 width=5 by lstas_zero, ex_intro/
- #l2 #_ elim (IHW1 l2) -IHW1 #W2 #HW2
+| #G #L #K #W1 #W #i #HLK #HW1 #IHW1 #d2
+ @(nat_ind_plus … d2) -d2 /3 width=5 by lstas_zero, ex_intro/
+ #d2 #_ elim (IHW1 d2) -IHW1 #W2 #HW2
lapply (lstas_trans … HW1 … HW2) -W
elim (lift_total W2 0 (i+1))
/3 width=7 by lstas_succ, ex_intro/
-| #G #L #K #W1 #W #U #i #l #HLK #_ #HWU #IHW1 #l2
- elim (IHW1 l2) -IHW1 #W2
+| #G #L #K #W1 #W #U #i #d #HLK #_ #HWU #IHW1 #d2
+ elim (IHW1 d2) -IHW1 #W2
elim (lift_total W2 0 (i+1))
lapply (drop_fwd_drop2 … HLK) -HLK
/3 width=11 by ex_intro, lstas_lift/
-| #a #I #G #L #V #T #U #l #_ #IHTU #l2
- elim (IHTU l2) -IHTU /3 width=2 by lstas_bind, ex_intro/
-| #G #L #V #T #U #l #_ #IHTU #l2
- elim (IHTU l2) -IHTU /3 width=2 by lstas_appl, ex_intro/
-| #G #L #W #T #U #l #_ #IHTU #l2
- elim (IHTU l2) -IHTU /2 width=2 by ex_intro/
+| #a #I #G #L #V #T #U #d #_ #IHTU #d2
+ elim (IHTU d2) -IHTU /3 width=2 by lstas_bind, ex_intro/
+| #G #L #V #T #U #d #_ #IHTU #d2
+ elim (IHTU d2) -IHTU /3 width=2 by lstas_appl, ex_intro/
+| #G #L #W #T #U #d #_ #IHTU #d2
+ elim (IHTU d2) -IHTU /2 width=2 by ex_intro/
]
qed-.
(* more main properties *****************************************************)
-theorem lstas_conf_le: ∀h,G,L,T,U1,l1. ⦃G, L⦄ ⊢ T •*[h, l1] U1 →
- ∀U2,l2. l1 ≤ l2 → ⦃G, L⦄ ⊢ T •*[h, l2] U2 →
- ⦃G, L⦄ ⊢ U1 •*[h, l2-l1] U2.
-#h #G #L #T #U1 #l1 #HTU1 #U2 #l2 #Hl12
->(plus_minus_m_m … Hl12) in ⊢ (%→?); -Hl12 >commutative_plus #H
+theorem lstas_conf_le: ∀h,G,L,T,U1,d1. ⦃G, L⦄ ⊢ T •*[h, d1] U1 →
+ ∀U2,d2. d1 ≤ d2 → ⦃G, L⦄ ⊢ T •*[h, d2] U2 →
+ ⦃G, L⦄ ⊢ U1 •*[h, d2-d1] U2.
+#h #G #L #T #U1 #d1 #HTU1 #U2 #d2 #Hd12
+>(plus_minus_m_m … Hd12) in ⊢ (%→?); -Hd12 >commutative_plus #H
elim (lstas_split … H) -H #U #HTU
>(lstas_mono … HTU … HTU1) -T //
qed-.
-theorem lstas_conf: ∀h,G,L,T0,T1,l1. ⦃G, L⦄ ⊢ T0 •*[h, l1] T1 →
- ∀T2,l2. ⦃G, L⦄ ⊢ T0 •*[h, l2] T2 →
- ∃∃T. ⦃G, L⦄ ⊢ T1 •*[h, l2] T & ⦃G, L⦄ ⊢ T2 •*[h, l1] T.
-#h #G #L #T0 #T1 #l1 #HT01 #T2 #l2 #HT02
-elim (lstas_lstas … HT01 (l1+l2)) #T #HT0
+theorem lstas_conf: ∀h,G,L,T0,T1,d1. ⦃G, L⦄ ⊢ T0 •*[h, d1] T1 →
+ ∀T2,d2. ⦃G, L⦄ ⊢ T0 •*[h, d2] T2 →
+ ∃∃T. ⦃G, L⦄ ⊢ T1 •*[h, d2] T & ⦃G, L⦄ ⊢ T2 •*[h, d1] T.
+#h #G #L #T0 #T1 #d1 #HT01 #T2 #d2 #HT02
+elim (lstas_lstas … HT01 (d1+d2)) #T #HT0
lapply (lstas_conf_le … HT01 … HT0) // -HT01 <minus_plus_m_m_commutative
lapply (lstas_conf_le … HT02 … HT0) // -T0 <minus_plus_m_m
/2 width=3 by ex2_intro/
<section>Summary of the Specification</section>
<body>Here is a numerical acount of the specification's contents
and its timeline.
- Nodes are counted according to the "intrinsinc complexity measure"
- [F. Guidi: "Procedural Representation of CIC Proof Terms"
- Journal of Automated Reasoning 44(1-2), Springer (February 2010),
- pp. 53-78].
</body>
<table name="basic_2_sum"/>
<subsection>Stage "B"</subsection>
- <news date="In progress.">
+ <news class="alpha" date="Ongoing.">
Context-sensitive subject equivalence
for native type assignment.
</news>
<subsection>Stage "A": "Weakening the Applicability Condition"</subsection>
- <news date="2014 September 9.">
+ <news class="beta" date="2014 September 9.">
Iterated static type assignment defined (more elegantly)
as a primitive notion.
</news>
- <news date="2014 June 18.">
+ <news class="beta" date="2014 June 18.">
Preservation of stratified native validity
for context-sensitive computation on terms.
</news>
- <news date="2014 June 9.">
+ <news class="alpha" date="2014 June 9.">
Strong qrst-normalization
for simply typed terms.
</news>
- <news date="2014 April 16.">
+ <news class="alpha" date="2014 April 16.">
Lazy equivalence on local environments
addded as q-step to rst-computation on closures
(anniversary milestone).
</news>
- <news date="2014 January 20.">
+ <news class="alpha" date="2014 January 20.">
Parametrized slicing of local environments
comprises both versions of this operation
(one from basic_1, the other used in basic_2 till now).
</news>
- <news date="2013 August 7.">
+ <news class="alpha" date="2013 August 7.">
Passive support for global environments.
</news>
- <news date="2013 July 27.">
+ <news class="alpha" date="2013 July 27.">
Reaxiomatized β-reductum as in rt-reduction.
</news>
- <news date="2013 July 20.">
+ <news class="alpha" date="2013 July 20.">
Context-sensitive strong rt-normalization
for simply typed terms.
</news>
- <news date="2013 April 16.">
+ <news class="alpha" date="2013 April 16.">
Reaxiomatized substitution and reduction
commute with respect to subclosure
(anniversary milestone).
</news>
- <news date="2013 March 16.">
+ <news class="alpha" date="2013 March 16.">
Mutual recursive preservation of stratified native validity
for rst-computation on closures.
</news>
- <news date="2012 October 16.">
+ <news class="alpha" date="2012 October 16.">
Confluence for context-free parallel reduction on closures.
</news>
- <news date="2012 July 26.">
+ <news class="alpha" date="2012 July 26.">
Term binders polarized to control ζ-reduction (not released).
</news>
- <news date="2012 April 16.">
+ <news class="alpha" date="2012 April 16.">
Context-sensitive subject equivalence
for atomic arity assignment
(anniversary milestone).
</news>
- <news date="2012 March 15.">
+ <news class="alpha" date="2012 March 15.">
Context-sensitive strong normalization
for simply typed terms.
</news>
- <news date="2012 January 27.">
+ <news class="alpha" date="2012 January 27.">
Support for abstract candidates of reducibility.
</news>
- <news date="2011 September 21.">
+ <news class="alpha" date="2011 September 21.">
Confluence for context-sensitive parallel reduction on terms.
</news>
- <news date="2011 September 6.">
+ <news class="alpha" date="2011 September 6.">
Confluence for context-free parallel reduction on terms.
</news>
- <news date="2011 April 17.">
+ <news class="alpha" date="2011 April 17.">
Specification starts.
</news>
</body>
<table name="basic_2_src"/>
- <section>Physical Structure of the Specification</section>
- <body>The source files are grouped in directories,
- one for each component.
- </body>
<footer/>
</page>
]
[ { "context-sensitive rt-computation" * } {
[ "lpxs ( ⦃?,?⦄ ⊢ ➡*[?,?] ? )" "lpxs_drop" + "lpxs_lleq" + "lpxs_aaa" + "lpxs_cpxs" + "lpxs_lpxs" * ]
- [ "cpxs ( ⦃?,?⦄ ⊢ ? ➡*[?,?] ? )" "cpxs_tsts" + "cpxs_tsts_vector" + "cpxs_leq" + "cpxs_lift" + "cpxs_lleq" + "cpxs_aaa" + "cpxs_cpxs" * ]
+ [ "cpxs ( ⦃?,?⦄ ⊢ ? ➡*[?,?] ? )" "cpxs_tsts" + "cpxs_tsts_vector" + "cpxs_lreq" + "cpxs_lift" + "cpxs_lleq" + "cpxs_aaa" + "cpxs_cpxs" * ]
}
]
[ { "context-sensitive computation" * } {
]
[ { "context-sensitive rt-reduction" * } {
[ "lpx ( ⦃?,?⦄ ⊢ ➡[?,?] ? )" "lpx_drop" + "lpx_frees" + "lpx_lleq" + "lpx_aaa" * ]
- [ "cpx ( ⦃?,?⦄ ⊢ ? ➡[?,?] ? )" "cpx_leq" + "cpx_lift" + "cpx_llpx_sn" + "cpx_lleq" + "cpx_cix" * ]
+ [ "cpx ( ⦃?,?⦄ ⊢ ? ➡[?,?] ? )" "cpx_lreq" + "cpx_lift" + "cpx_llpx_sn" + "cpx_lleq" + "cpx_cix" * ]
}
]
[ { "irreducible forms for context-sensitive rt-reduction" * } {
[ { "multiple substitution" * } {
[ { "lazy equivalence" * } {
[ "fleq ( ⦃?,?,?⦄ ≡[?] ⦃?,?,?⦄ )" "fleq_fleq" * ]
- [ "lleq ( ? ≡[?,?] ? )" "lleq_alt" + "lleq_alt_rec" + "lleq_leq" + "lleq_drop" + "lleq_fqus" + "lleq_llor" + "lleq_lleq" * ]
+ [ "lleq ( ? ≡[?,?] ? )" "lleq_alt" + "lleq_alt_rec" + "lleq_lreq" + "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_drop" + "llpx_sn_lpx_sn" + "llpx_sn_frees" + "llpx_sn_llor" * ]
+ [ "llpx_sn" "llpx_sn_alt" + "llpx_sn_alt_rec" + "llpx_sn_tc" + "llpx_sn_lreq" + "llpx_sn_drop" + "llpx_sn_lpx_sn" + "llpx_sn_frees" + "llpx_sn_llor" * ]
}
]
[ { "pointwise union for local environments" * } {
}
]
[ { "context-sensitive exclusion from free variables" * } {
- [ "frees ( ? ⊢ ? ϵ 𝐅*[?]⦃?⦄ )" "frees_append" + "frees_leq" + "frees_lift" * ]
+ [ "frees ( ? ⊢ ? ϵ 𝐅*[?]⦃?⦄ )" "frees_append" + "frees_lreq" + "frees_lift" * ]
}
]
[ { "contxt-sensitive multiple rt-substitution" * } {
}
]
[ { "basic local env. slicing" * } {
- [ "drop ( ⬇[?,?,?] ? ≡ ? )" "drop_append" + "drop_leq" + "drop_drop" * ]
+ [ "drop ( ⬇[?,?,?] ? ≡ ? )" "drop_append" + "drop_lreq" + "drop_drop" * ]
}
]
[ { "basic term relocation" * } {
class "red"
[ { "grammar" * } {
[ { "equivalence for local environments" * } {
- [ "leq ( ? ⩬[?,?] ? )" "leq_leq" * ]
+ [ "lreq ( ? ⩬[?,?] ? )" "lreq_lreq" * ]
}
]
[ { "same top term structure" * } {
pp. 53-78].
</body>
<table name="ground_2_sum"/>
- <news date="2013 November 27.">
+ <news class="alpha" date="2013 November 27.">
Natural numbers with infinity.
</news>
- <news date="2011 August 10.">
+ <news class="alpha" date="2011 August 10.">
Specification starts.
</news>
</body>
<table name="ground_2_src"/>
- <section>Physical Structure of the Specification</section>
- <body>The source files are grouped in directories,
- one for each plane.
- </body>
<footer/>
</page>