+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/predeval_4.ma".
-include "basic_2/computation/cprs.ma".
-include "basic_2/computation/csx.ma".
-
-(* EVALUATION FOR CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS *************)
-
-definition cpre: relation4 genv lenv term term ≝
- λG,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡* T2 ∧ ⦃G, L⦄ ⊢ ➡ 𝐍⦃T2⦄.
-
-interpretation "evaluation for context-sensitive parallel reduction (term)"
- 'PRedEval G L T1 T2 = (cpre G L T1 T2).
-
-(* Basic properties *********************************************************)
-
-(* Basic_1: was just: nf2_sn3 *)
-lemma csx_cpre: ∀h,o,G,L,T1. ⦃G, L⦄ ⊢ ⬊*[h, o] T1 → ∃T2. ⦃G, L⦄ ⊢ T1 ➡* 𝐍⦃T2⦄.
-#h #o #G #L #T1 #H @(csx_ind … H) -T1
-#T1 #_ #IHT1 elim (cnr_dec G L T1) /3 width=3 by ex_intro, conj/
-* #T #H1T1 #H2T1 elim (IHT1 … H2T1) -IHT1 -H2T1 /2 width=2 by cpr_cpx/
-#T2 * /4 width=3 by cprs_strap2, ex_intro, conj/
-qed.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/computation/cprs_cprs.ma".
-include "basic_2/computation/cpre.ma".
-
-(* EVALUATION FOR CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS *************)
-
-(* Main properties *********************************************************)
-
-(* Basic_1: was: nf2_pr3_confluence *)
-theorem cpre_mono: ∀G,L,T,T1. ⦃G, L⦄ ⊢ T ➡* 𝐍⦃T1⦄ → ∀T2. ⦃G, L⦄ ⊢ T ➡* 𝐍⦃T2⦄ → T1 = T2.
-#G #L #T #T1 * #H1T1 #H2T1 #T2 * #H1T2 #H2T2
-elim (cprs_conf … H1T1 … H1T2) -T #T #HT1
->(cprs_inv_cnr1 … HT1 H2T1) -T1 #HT2
->(cprs_inv_cnr1 … HT2 H2T2) -T2 //
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/predstar_4.ma".
-include "basic_2/reduction/cnr.ma".
-
-(* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON TERMS **************************)
-
-(* Basic_1: includes: pr1_pr0 *)
-definition cprs: relation4 genv lenv term term ≝
- λG. LTC … (cpr G).
-
-interpretation "context-sensitive parallel computation (term)"
- 'PRedStar G L T1 T2 = (cprs G L T1 T2).
-
-(* Basic eliminators ********************************************************)
-
-lemma cprs_ind: ∀G,L,T1. ∀R:predicate term. R T1 →
- (∀T,T2. ⦃G, L⦄ ⊢ T1 ➡* T → ⦃G, L⦄ ⊢ T ➡ T2 → R T → R T2) →
- ∀T2. ⦃G, L⦄ ⊢ T1 ➡* T2 → R T2.
-#G #L #T1 #R #HT1 #IHT1 #T2 #HT12
-@(TC_star_ind … HT1 IHT1 … HT12) //
-qed-.
-
-lemma cprs_ind_dx: ∀G,L,T2. ∀R:predicate term. R T2 →
- (∀T1,T. ⦃G, L⦄ ⊢ T1 ➡ T → ⦃G, L⦄ ⊢ T ➡* T2 → R T → R T1) →
- ∀T1. ⦃G, L⦄ ⊢ T1 ➡* T2 → R T1.
-#G #L #T2 #R #HT2 #IHT2 #T1 #HT12
-@(TC_star_ind_dx … HT2 IHT2 … HT12) //
-qed-.
-
-(* Basic properties *********************************************************)
-
-(* Basic_1: was: pr3_pr2 *)
-lemma cpr_cprs: ∀G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ⦃G, L⦄ ⊢ T1 ➡* T2.
-/2 width=1 by inj/ qed.
-
-(* Basic_1: was: pr3_refl *)
-lemma cprs_refl: ∀G,L,T. ⦃G, L⦄ ⊢ T ➡* T.
-/2 width=1 by cpr_cprs/ qed.
-
-lemma cprs_strap1: ∀G,L,T1,T,T2.
- ⦃G, L⦄ ⊢ T1 ➡* T → ⦃G, L⦄ ⊢ T ➡ T2 → ⦃G, L⦄ ⊢ T1 ➡* T2.
-normalize /2 width=3 by step/ qed-.
-
-(* Basic_1: was: pr3_step *)
-lemma cprs_strap2: ∀G,L,T1,T,T2.
- ⦃G, L⦄ ⊢ T1 ➡ T → ⦃G, L⦄ ⊢ T ➡* T2 → ⦃G, L⦄ ⊢ T1 ➡* T2.
-normalize /2 width=3 by TC_strap/ qed-.
-
-lemma lsubr_cprs_trans: ∀G. lsub_trans … (cprs G) lsubr.
-/3 width=5 by lsubr_cpr_trans, LTC_lsub_trans/
-qed-.
-
-(* Basic_1: was: pr3_pr1 *)
-lemma tprs_cprs: ∀G,L,T1,T2. ⦃G, ⋆⦄ ⊢ T1 ➡* T2 → ⦃G, L⦄ ⊢ T1 ➡* T2.
-/2 width=3 by lsubr_cprs_trans/ qed.
-
-lemma cprs_bind_dx: ∀G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡ V2 → ∀I,T1,T2. ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ➡* T2 →
- ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1. T1 ➡* ⓑ{a,I}V2. T2.
-#G #L #V1 #V2 #HV12 #I #T1 #T2 #HT12 #a @(cprs_ind_dx … HT12) -T1
-/3 width=3 by cprs_strap2, cpr_cprs, cpr_pair_sn, cpr_bind/ qed.
-
-(* Basic_1: was only: pr3_thin_dx *)
-lemma cprs_flat_dx: ∀I,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡ V2 → ∀T1,T2. ⦃G, L⦄ ⊢ T1 ➡* T2 →
- ⦃G, L⦄ ⊢ ⓕ{I} V1. T1 ➡* ⓕ{I} V2. T2.
-#I #G #L #V1 #V2 #HV12 #T1 #T2 #HT12 @(cprs_ind … HT12) -T2
-/3 width=5 by cprs_strap1, cpr_flat, cpr_cprs, cpr_pair_sn/
-qed.
-
-lemma cprs_flat_sn: ∀I,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ∀V1,V2. ⦃G, L⦄ ⊢ V1 ➡* V2 →
- ⦃G, L⦄ ⊢ ⓕ{I} V1. T1 ➡* ⓕ{I} V2. T2.
-#I #G #L #T1 #T2 #HT12 #V1 #V2 #H @(cprs_ind … H) -V2
-/3 width=3 by cprs_strap1, cpr_cprs, cpr_pair_sn, cpr_flat/
-qed.
-
-lemma cprs_zeta: ∀G,L,V,T1,T,T2. ⬆[0, 1] T2 ≡ T →
- ⦃G, L.ⓓV⦄ ⊢ T1 ➡* T → ⦃G, L⦄ ⊢ +ⓓV.T1 ➡* T2.
-#G #L #V #T1 #T #T2 #HT2 #H @(cprs_ind_dx … H) -T1
-/3 width=3 by cprs_strap2, cpr_cprs, cpr_bind, cpr_zeta/
-qed.
-
-lemma cprs_eps: ∀G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡* T2 → ∀V. ⦃G, L⦄ ⊢ ⓝV.T1 ➡* T2.
-#G #L #T1 #T2 #H @(cprs_ind … H) -T2
-/3 width=3 by cprs_strap1, cpr_cprs, cpr_eps/
-qed.
-
-lemma cprs_beta_dx: ∀a,G,L,V1,V2,W1,W2,T1,T2.
- ⦃G, L⦄ ⊢ V1 ➡ V2 → ⦃G, L⦄ ⊢ W1 ➡ W2 → ⦃G, L.ⓛW1⦄ ⊢ T1 ➡* T2 →
- ⦃G, L⦄ ⊢ ⓐV1.ⓛ{a}W1.T1 ➡* ⓓ{a}ⓝW2.V2.T2.
-#a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #HV12 #HW12 * -T2
-/4 width=7 by cprs_strap1, cpr_cprs, cprs_bind_dx, cprs_flat_dx, cpr_beta/
-qed.
-
-lemma cprs_theta_dx: ∀a,G,L,V1,V,V2,W1,W2,T1,T2.
- ⦃G, L⦄ ⊢ V1 ➡ V → ⬆[0, 1] V ≡ V2 → ⦃G, L.ⓓW1⦄ ⊢ T1 ➡* T2 →
- ⦃G, L⦄ ⊢ W1 ➡ W2 → ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}W1.T1 ➡* ⓓ{a}W2.ⓐV2.T2.
-#a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV1 #HV2 * -T2
-/4 width=9 by cprs_strap1, cpr_cprs, cprs_bind_dx, cprs_flat_dx, cpr_theta/
-qed.
-
-(* Basic inversion lemmas ***************************************************)
-
-(* Basic_1: was: pr3_gen_sort *)
-lemma cprs_inv_sort1: ∀G,L,U2,s. ⦃G, L⦄ ⊢ ⋆s ➡* U2 → U2 = ⋆s.
-#G #L #U2 #s #H @(cprs_ind … H) -U2 //
-#U2 #U #_ #HU2 #IHU2 destruct
->(cpr_inv_sort1 … HU2) -HU2 //
-qed-.
-
-(* Basic_1: was: pr3_gen_cast *)
-lemma cprs_inv_cast1: ∀G,L,W1,T1,U2. ⦃G, L⦄ ⊢ ⓝW1.T1 ➡* U2 → ⦃G, L⦄ ⊢ T1 ➡* U2 ∨
- ∃∃W2,T2. ⦃G, L⦄ ⊢ W1 ➡* W2 & ⦃G, L⦄ ⊢ T1 ➡* T2 & U2 = ⓝW2.T2.
-#G #L #W1 #T1 #U2 #H @(cprs_ind … H) -U2 /3 width=5 by ex3_2_intro, or_intror/
-#U2 #U #_ #HU2 * /3 width=3 by cprs_strap1, or_introl/ *
-#W #T #HW1 #HT1 #H destruct
-elim (cpr_inv_cast1 … HU2) -HU2 /3 width=3 by cprs_strap1, or_introl/ *
-#W2 #T2 #HW2 #HT2 #H destruct /4 width=5 by cprs_strap1, ex3_2_intro, or_intror/
-qed-.
-
-(* Basic_1: was: nf2_pr3_unfold *)
-lemma cprs_inv_cnr1: ∀G,L,T,U. ⦃G, L⦄ ⊢ T ➡* U → ⦃G, L⦄ ⊢ ➡ 𝐍⦃T⦄ → T = U.
-#G #L #T #U #H @(cprs_ind_dx … H) -T //
-#T0 #T #H1T0 #_ #IHT #H2T0
-lapply (H2T0 … H1T0) -H1T0 #H destruct /2 width=1 by/
-qed-.
-
-(* Basic_1: removed theorems 13:
- pr1_head_1 pr1_head_2 pr1_comp
- clear_pr3_trans pr3_cflat pr3_gen_bind
- pr3_head_1 pr3_head_2 pr3_head_21 pr3_head_12
- pr3_iso_appl_bind pr3_iso_appls_appl_bind pr3_iso_appls_bind
-*)
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/reduction/lpr_lpr.ma".
-include "basic_2/computation/cprs_lift.ma".
-
-(* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON TERMS **************************)
-
-(* Main properties **********************************************************)
-
-(* Basic_1: was: pr3_t *)
-(* Basic_1: includes: pr1_t *)
-theorem cprs_trans: ∀G,L. Transitive … (cprs G L).
-normalize /2 width=3 by trans_TC/ qed-.
-
-(* Basic_1: was: pr3_confluence *)
-(* Basic_1: includes: pr1_confluence *)
-theorem cprs_conf: ∀G,L. confluent2 … (cprs G L) (cprs G L).
-normalize /3 width=3 by cpr_conf, TC_confluent2/ qed-.
-
-theorem cprs_bind: ∀a,I,G,L,V1,V2,T1,T2. ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ➡* T2 → ⦃G, L⦄ ⊢ V1 ➡* V2 →
- ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡* ⓑ{a,I}V2.T2.
-#a #I #G #L #V1 #V2 #T1 #T2 #HT12 #H @(cprs_ind … H) -V2
-/3 width=5 by cprs_trans, cprs_bind_dx/
-qed.
-
-(* Basic_1: was: pr3_flat *)
-theorem cprs_flat: ∀I,G,L,V1,V2,T1,T2. ⦃G, L⦄ ⊢ T1 ➡* T2 → ⦃G, L⦄ ⊢ V1 ➡* V2 →
- ⦃G, L⦄ ⊢ ⓕ{I}V1.T1 ➡* ⓕ{I}V2.T2.
-#I #G #L #V1 #V2 #T1 #T2 #HT12 #H @(cprs_ind … H) -V2
-/3 width=3 by cprs_flat_dx, cprs_strap1, cpr_pair_sn/
-qed.
-
-theorem cprs_beta_rc: ∀a,G,L,V1,V2,W1,W2,T1,T2.
- ⦃G, L⦄ ⊢ V1 ➡ V2 → ⦃G, L.ⓛW1⦄ ⊢ T1 ➡* T2 → ⦃G, L⦄ ⊢ W1 ➡* W2 →
- ⦃G, L⦄ ⊢ ⓐV1.ⓛ{a}W1.T1 ➡* ⓓ{a}ⓝW2.V2.T2.
-#a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #HV12 #HT12 #H @(cprs_ind … H) -W2 /2 width=1 by cprs_beta_dx/
-#W #W2 #_ #HW2 #IHW1 (**) (* fulla uto too slow 14s *)
-@(cprs_trans … IHW1) -IHW1 /3 width=1 by cprs_flat_dx, cprs_bind/
-qed.
-
-theorem cprs_beta: ∀a,G,L,V1,V2,W1,W2,T1,T2.
- ⦃G, L.ⓛW1⦄ ⊢ T1 ➡* T2 → ⦃G, L⦄ ⊢ W1 ➡* W2 → ⦃G, L⦄ ⊢ V1 ➡* V2 →
- ⦃G, L⦄ ⊢ ⓐV1.ⓛ{a}W1.T1 ➡* ⓓ{a}ⓝW2.V2.T2.
-#a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #HT12 #HW12 #H @(cprs_ind … H) -V2 /2 width=1 by cprs_beta_rc/
-#V #V2 #_ #HV2 #IHV1
-@(cprs_trans … IHV1) -IHV1 /3 width=1 by cprs_flat_sn, cprs_bind/
-qed.
-
-theorem cprs_theta_rc: ∀a,G,L,V1,V,V2,W1,W2,T1,T2.
- ⦃G, L⦄ ⊢ V1 ➡ V → ⬆[0, 1] V ≡ V2 → ⦃G, L.ⓓW1⦄ ⊢ T1 ➡* T2 →
- ⦃G, L⦄ ⊢ W1 ➡* W2 → ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}W1.T1 ➡* ⓓ{a}W2.ⓐV2.T2.
-#a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV1 #HV2 #HT12 #H @(cprs_ind … H) -W2
-/3 width=5 by cprs_trans, cprs_theta_dx, cprs_bind_dx/
-qed.
-
-theorem cprs_theta: ∀a,G,L,V1,V,V2,W1,W2,T1,T2.
- ⬆[0, 1] V ≡ V2 → ⦃G, L⦄ ⊢ W1 ➡* W2 → ⦃G, L.ⓓW1⦄ ⊢ T1 ➡* T2 →
- ⦃G, L⦄ ⊢ V1 ➡* V → ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}W1.T1 ➡* ⓓ{a}W2.ⓐV2.T2.
-#a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV2 #HW12 #HT12 #H @(cprs_ind_dx … H) -V1
-/3 width=3 by cprs_trans, cprs_theta_rc, cprs_flat_dx/
-qed.
-
-(* Advanced inversion lemmas ************************************************)
-
-(* Basic_1: was pr3_gen_appl *)
-lemma cprs_inv_appl1: ∀G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓐV1.T1 ➡* U2 →
- ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡* V2 & ⦃G, L⦄ ⊢ T1 ➡* T2 &
- U2 = ⓐV2. T2
- | ∃∃a,W,T. ⦃G, L⦄ ⊢ T1 ➡* ⓛ{a}W.T &
- ⦃G, L⦄ ⊢ ⓓ{a}ⓝW.V1.T ➡* U2
- | ∃∃a,V0,V2,V,T. ⦃G, L⦄ ⊢ V1 ➡* V0 & ⬆[0,1] V0 ≡ V2 &
- ⦃G, L⦄ ⊢ T1 ➡* ⓓ{a}V.T &
- ⦃G, L⦄ ⊢ ⓓ{a}V.ⓐV2.T ➡* U2.
-#G #L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /3 width=5 by or3_intro0, ex3_2_intro/
-#U #U2 #_ #HU2 * *
-[ #V0 #T0 #HV10 #HT10 #H destruct
- elim (cpr_inv_appl1 … HU2) -HU2 *
- [ #V2 #T2 #HV02 #HT02 #H destruct /4 width=5 by cprs_strap1, or3_intro0, ex3_2_intro/
- | #a #V2 #W #W2 #T #T2 #HV02 #HW2 #HT2 #H1 #H2 destruct
- lapply (cprs_strap1 … HV10 … HV02) -V0 #HV12
- lapply (lsubr_cpr_trans … HT2 (L.ⓓⓝW.V1) ?) -HT2
- /5 width=5 by cprs_bind, cprs_flat_dx, cpr_cprs, lsubr_beta, ex2_3_intro, or3_intro1/
- | #a #V #V2 #W0 #W2 #T #T2 #HV0 #HV2 #HW02 #HT2 #H1 #H2 destruct
- /5 width=10 by cprs_flat_sn, cprs_bind_dx, cprs_strap1, ex4_5_intro, or3_intro2/
- ]
-| /4 width=9 by cprs_strap1, or3_intro1, ex2_3_intro/
-| /4 width=11 by cprs_strap1, or3_intro2, ex4_5_intro/
-]
-qed-.
-
-(* Properties concerning sn parallel reduction on local environments ********)
-
-(* Basic_1: was just: pr3_pr2_pr2_t *)
-(* Basic_1: includes: pr3_pr0_pr2_t *)
-lemma lpr_cpr_trans: ∀G. c_r_transitive … (cpr G) (λ_. lpr G).
-#G #L2 #T1 #T2 #HT12 elim HT12 -G -L2 -T1 -T2
-[ /2 width=3 by/
-| #G #L2 #K2 #V0 #V2 #W2 #i #HLK2 #_ #HVW2 #IHV02 #L1 #HL12
- elim (lpr_drop_trans_O1 … HL12 … HLK2) -L2 #X #HLK1 #H
- elim (lpr_inv_pair2 … H) -H #K1 #V1 #HK12 #HV10 #H destruct
- /4 width=6 by cprs_strap2, cprs_delta/
-|3,7: /4 width=1 by lpr_pair, cprs_bind, cprs_beta/
-|4,6: /3 width=1 by cprs_flat, cprs_eps/
-|5,8: /4 width=3 by lpr_pair, cprs_zeta, cprs_theta, cprs_strap1/
-]
-qed-.
-
-lemma cpr_bind2: ∀G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡ V2 → ∀I,T1,T2. ⦃G, L.ⓑ{I}V2⦄ ⊢ T1 ➡ T2 →
- ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡* ⓑ{a,I}V2.T2.
-/4 width=5 by lpr_cpr_trans, cprs_bind_dx, lpr_pair/ qed.
-
-(* Advanced properties ******************************************************)
-
-(* Basic_1: was only: pr3_pr2_pr3_t pr3_wcpr0_t *)
-lemma lpr_cprs_trans: ∀G. c_rs_transitive … (cpr G) (λ_. lpr G).
-#G @c_r_trans_LTC1 /2 width=3 by lpr_cpr_trans/ (**) (* full auto fails *)
-qed-.
-
-(* Basic_1: was: pr3_strip *)
-(* Basic_1: includes: pr1_strip *)
-lemma cprs_strip: ∀G,L. confluent2 … (cprs G L) (cpr G L).
-normalize /4 width=3 by cpr_conf, TC_strip1/ qed-.
-
-lemma cprs_lpr_conf_dx: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡* T1 → ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 →
- ∃∃T. ⦃G, L1⦄ ⊢ T1 ➡* T & ⦃G, L1⦄ ⊢ T0 ➡* T.
-#G #L0 #T0 #T1 #H @(cprs_ind … H) -T1 /2 width=3 by ex2_intro/
-#T #T1 #_ #HT1 #IHT0 #L1 #HL01 elim (IHT0 … HL01)
-#T2 #HT2 #HT02 elim (lpr_cpr_conf_dx … HT1 … HL01) -L0
-#T3 #HT3 #HT13 elim (cprs_strip … HT2 … HT3) -T
-/4 width=5 by cprs_strap2, cprs_strap1, ex2_intro/
-qed-.
-
-lemma cprs_lpr_conf_sn: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡* T1 →
- ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 →
- ∃∃T. ⦃G, L0⦄ ⊢ T1 ➡* T & ⦃G, L1⦄ ⊢ T0 ➡* T.
-#G #L0 #T0 #T1 #HT01 #L1 #HL01 elim (cprs_lpr_conf_dx … HT01 … HL01) -HT01
-/3 width=3 by lpr_cprs_trans, ex2_intro/
-qed-.
-
-lemma cprs_bind2_dx: ∀G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡ V2 →
- ∀I,T1,T2. ⦃G, L.ⓑ{I}V2⦄ ⊢ T1 ➡* T2 →
- ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡* ⓑ{a,I}V2.T2.
-/4 width=5 by lpr_cprs_trans, cprs_bind_dx, lpr_pair/ 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/cpr_lift.ma".
-include "basic_2/computation/cprs.ma".
-
-(* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON TERMS **************************)
-
-(* Advanced properties ******************************************************)
-
-(* Note: apparently this was missing in basic_1 *)
-lemma cprs_delta: ∀G,L,K,V,V2,i.
- ⬇[i] L ≡ K.ⓓV → ⦃G, K⦄ ⊢ V ➡* V2 →
- ∀W2. ⬆[0, i + 1] V2 ≡ W2 → ⦃G, L⦄ ⊢ #i ➡* W2.
-#G #L #K #V #V2 #i #HLK #H elim H -V2 [ /3 width=6 by cpr_cprs, cpr_delta/ ]
-#V1 #V2 #_ #HV12 #IHV1 #W2 #HVW2
-lapply (drop_fwd_drop2 … HLK) -HLK #HLK
-elim (lift_total V1 0 (i+1)) /4 width=12 by cpr_lift, cprs_strap1/
-qed.
-
-(* Advanced inversion lemmas ************************************************)
-
-(* Basic_1: was: pr3_gen_lref *)
-lemma cprs_inv_lref1: ∀G,L,T2,i. ⦃G, L⦄ ⊢ #i ➡* T2 →
- T2 = #i ∨
- ∃∃K,V1,T1. ⬇[i] L ≡ K.ⓓV1 & ⦃G, K⦄ ⊢ V1 ➡* T1 &
- ⬆[0, i + 1] T1 ≡ T2.
-#G #L #T2 #i #H @(cprs_ind … H) -T2 /2 width=1 by or_introl/
-#T #T2 #_ #HT2 *
-[ #H destruct
- elim (cpr_inv_lref1 … HT2) -HT2 /2 width=1 by or_introl/
- * /4 width=6 by cpr_cprs, ex3_3_intro, or_intror/
-| * #K #V1 #T1 #HLK #HVT1 #HT1
- lapply (drop_fwd_drop2 … HLK) #H0LK
- elim (cpr_inv_lift1 … HT2 … H0LK … HT1) -H0LK -T
- /4 width=6 by cprs_strap1, ex3_3_intro, or_intror/
-]
-qed-.
-
-(* Relocation properties ****************************************************)
-
-(* Basic_1: was: pr3_lift *)
-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. d_deliftable_sn (cprs G).
-/3 width=6 by d_deliftable_sn_LTC, cpr_inv_lift1/
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/predeval_6.ma".
-include "basic_2/computation/cpxs.ma".
-include "basic_2/computation/csx.ma".
-
-(* EVALUATION FOR CONTEXT-SENSITIVE EXTENDED PARALLEL REDUCTION ON TERMS ****)
-
-definition cpxe: ∀h. sd h → relation4 genv lenv term term ≝
- λh,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 ∧ ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T2⦄.
-
-interpretation "evaluation for context-sensitive extended parallel reduction (term)"
- 'PRedEval h o G L T1 T2 = (cpxe h o G L T1 T2).
-
-(* Basic properties *********************************************************)
-
-lemma csx_cpxe: ∀h,o,G,L,T1. ⦃G, L⦄ ⊢ ⬊*[h, o] T1 → ∃T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] 𝐍⦃T2⦄.
-#h #o #G #L #T1 #H @(csx_ind … H) -T1
-#T1 #_ #IHT1 elim (cnx_dec h o G L T1) /3 width=3 by ex_intro, conj/
-* #T #H1T1 #H2T1 elim (IHT1 … H1T1 H2T1) -IHT1 -H2T1
-#T2 * /4 width=3 by cpxs_strap2, ex_intro, conj/
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/predstar_6.ma".
-include "basic_2/reduction/cnx.ma".
-include "basic_2/computation/cprs.ma".
-
-(* CONTEXT-SENSITIVE EXTENDED PARALLEL COMPUTATION ON TERMS *****************)
-
-definition cpxs: ∀h. sd h → relation4 genv lenv term term ≝
- λh,o,G. LTC … (cpx h o G).
-
-interpretation "extended context-sensitive parallel computation (term)"
- 'PRedStar h o G L T1 T2 = (cpxs h o G L T1 T2).
-
-(* Basic eliminators ********************************************************)
-
-lemma cpxs_ind: ∀h,o,G,L,T1. ∀R:predicate term. R T1 →
- (∀T,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T → ⦃G, L⦄ ⊢ T ➡[h, o] T2 → R T → R T2) →
- ∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → R T2.
-#h #o #L #G #T1 #R #HT1 #IHT1 #T2 #HT12
-@(TC_star_ind … HT1 IHT1 … HT12) //
-qed-.
-
-lemma cpxs_ind_dx: ∀h,o,G,L,T2. ∀R:predicate term. R T2 →
- (∀T1,T. ⦃G, L⦄ ⊢ T1 ➡[h, o] T → ⦃G, L⦄ ⊢ T ➡*[h, o] T2 → R T → R T1) →
- ∀T1. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → R T1.
-#h #o #G #L #T2 #R #HT2 #IHT2 #T1 #HT12
-@(TC_star_ind_dx … HT2 IHT2 … HT12) //
-qed-.
-
-(* Basic properties *********************************************************)
-
-lemma cpxs_refl: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ T ➡*[h, o] T.
-/2 width=1 by inj/ qed.
-
-lemma cpx_cpxs: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2.
-/2 width=1 by inj/ qed.
-
-lemma cpxs_strap1: ∀h,o,G,L,T1,T. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T →
- ∀T2. ⦃G, L⦄ ⊢ T ➡[h, o] T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2.
-normalize /2 width=3 by step/ qed.
-
-lemma cpxs_strap2: ∀h,o,G,L,T1,T. ⦃G, L⦄ ⊢ T1 ➡[h, o] T →
- ∀T2. ⦃G, L⦄ ⊢ T ➡*[h, o] T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2.
-normalize /2 width=3 by TC_strap/ qed.
-
-lemma lsubr_cpxs_trans: ∀h,o,G. lsub_trans … (cpxs h o G) lsubr.
-/3 width=5 by lsubr_cpx_trans, LTC_lsub_trans/
-qed-.
-
-lemma cprs_cpxs: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡* T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2.
-#h #o #G #L #T1 #T2 #H @(cprs_ind … H) -T2 /3 width=3 by cpxs_strap1, cpr_cpx/
-qed.
-
-lemma cpxs_sort: ∀h,o,G,L,s,d1. deg h o s d1 →
- ∀d2. d2 ≤ d1 → ⦃G, L⦄ ⊢ ⋆s ➡*[h, o] ⋆((next h)^d2 s).
-#h #o #G #L #s #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) s)))
-[ /3 width=3 by lt_to_le/
-| @(cpx_st … (d1-d2-1)) <plus_minus_k_k
- /2 width=1 by deg_iter, monotonic_le_minus_r/
-]
-qed.
-
-lemma cpxs_bind_dx: ∀h,o,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 →
- ∀I,T1,T2. ⦃G, L. ⓑ{I}V1⦄ ⊢ T1 ➡*[h, o] T2 →
- ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡*[h, o] ⓑ{a,I}V2.T2.
-#h #o #G #L #V1 #V2 #HV12 #I #T1 #T2 #HT12 #a @(cpxs_ind_dx … HT12) -T1
-/3 width=3 by cpxs_strap2, cpx_cpxs, cpx_pair_sn, cpx_bind/
-qed.
-
-lemma cpxs_flat_dx: ∀h,o,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 →
- ∀T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 →
- ∀I. ⦃G, L⦄ ⊢ ⓕ{I}V1.T1 ➡*[h, o] ⓕ{I}V2.T2.
-#h #o #G #L #V1 #V2 #HV12 #T1 #T2 #HT12 @(cpxs_ind … HT12) -T2
-/3 width=5 by cpxs_strap1, cpx_cpxs, cpx_pair_sn, cpx_flat/
-qed.
-
-lemma cpxs_flat_sn: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 →
- ∀V1,V2. ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 →
- ∀I. ⦃G, L⦄ ⊢ ⓕ{I}V1.T1 ➡*[h, o] ⓕ{I}V2.T2.
-#h #o #G #L #T1 #T2 #HT12 #V1 #V2 #H @(cpxs_ind … H) -V2
-/3 width=5 by cpxs_strap1, cpx_cpxs, cpx_pair_sn, cpx_flat/
-qed.
-
-lemma cpxs_pair_sn: ∀h,o,I,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 →
- ∀T. ⦃G, L⦄ ⊢ ②{I}V1.T ➡*[h, o] ②{I}V2.T.
-#h #o #I #G #L #V1 #V2 #H @(cpxs_ind … H) -V2
-/3 width=3 by cpxs_strap1, cpx_pair_sn/
-qed.
-
-lemma cpxs_zeta: ∀h,o,G,L,V,T1,T,T2. ⬆[0, 1] T2 ≡ T →
- ⦃G, L.ⓓV⦄ ⊢ T1 ➡*[h, o] T → ⦃G, L⦄ ⊢ +ⓓV.T1 ➡*[h, o] T2.
-#h #o #G #L #V #T1 #T #T2 #HT2 #H @(cpxs_ind_dx … H) -T1
-/3 width=3 by cpxs_strap2, cpx_cpxs, cpx_bind, cpx_zeta/
-qed.
-
-lemma cpxs_eps: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 →
- ∀V. ⦃G, L⦄ ⊢ ⓝV.T1 ➡*[h, o] T2.
-#h #o #G #L #T1 #T2 #H @(cpxs_ind … H) -T2
-/3 width=3 by cpxs_strap1, cpx_cpxs, cpx_eps/
-qed.
-
-lemma cpxs_ct: ∀h,o,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 →
- ∀T. ⦃G, L⦄ ⊢ ⓝV1.T ➡*[h, o] V2.
-#h #o #G #L #V1 #V2 #H @(cpxs_ind … H) -V2
-/3 width=3 by cpxs_strap1, cpx_cpxs, cpx_ct/
-qed.
-
-lemma cpxs_beta_dx: ∀h,o,a,G,L,V1,V2,W1,W2,T1,T2.
- ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 → ⦃G, L.ⓛW1⦄ ⊢ T1 ➡*[h, o] T2 → ⦃G, L⦄ ⊢ W1 ➡[h, o] W2 →
- ⦃G, L⦄ ⊢ ⓐV1.ⓛ{a}W1.T1 ➡*[h, o] ⓓ{a}ⓝW2.V2.T2.
-#h #o #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #HV12 * -T2
-/4 width=7 by cpx_cpxs, cpxs_strap1, cpxs_bind_dx, cpxs_flat_dx, cpx_beta/
-qed.
-
-lemma cpxs_theta_dx: ∀h,o,a,G,L,V1,V,V2,W1,W2,T1,T2.
- ⦃G, L⦄ ⊢ V1 ➡[h, o] V → ⬆[0, 1] V ≡ V2 → ⦃G, L.ⓓW1⦄ ⊢ T1 ➡*[h, o] T2 →
- ⦃G, L⦄ ⊢ W1 ➡[h, o] W2 → ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}W1.T1 ➡*[h, o] ⓓ{a}W2.ⓐV2.T2.
-#h #o #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV1 #HV2 * -T2
-/4 width=9 by cpx_cpxs, cpxs_strap1, cpxs_bind_dx, cpxs_flat_dx, cpx_theta/
-qed.
-
-(* Basic inversion lemmas ***************************************************)
-
-lemma cpxs_inv_sort1: ∀h,o,G,L,U2,s. ⦃G, L⦄ ⊢ ⋆s ➡*[h, o] U2 →
- ∃∃n,d. deg h o s (n+d) & U2 = ⋆((next h)^n s).
-#h #o #G #L #U2 #s #H @(cpxs_ind … H) -U2
-[ elim (deg_total h o s) #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/
- | * #d0 #Hkd0 #H destruct -d
- @(ex2_2_intro … (n+1) d0) /2 width=1 by deg_inv_prec/ >iter_SO //
- ]
-]
-qed-.
-
-lemma cpxs_inv_cast1: ∀h,o,G,L,W1,T1,U2. ⦃G, L⦄ ⊢ ⓝW1.T1 ➡*[h, o] U2 →
- ∨∨ ∃∃W2,T2. ⦃G, L⦄ ⊢ W1 ➡*[h, o] W2 & ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 & U2 = ⓝW2.T2
- | ⦃G, L⦄ ⊢ T1 ➡*[h, o] U2
- | ⦃G, L⦄ ⊢ W1 ➡*[h, o] U2.
-#h #o #G #L #W1 #T1 #U2 #H @(cpxs_ind … H) -U2 /3 width=5 by or3_intro0, ex3_2_intro/
-#U2 #U #_ #HU2 * /3 width=3 by cpxs_strap1, or3_intro1, or3_intro2/ *
-#W #T #HW1 #HT1 #H destruct
-elim (cpx_inv_cast1 … HU2) -HU2 /3 width=3 by cpxs_strap1, or3_intro1, or3_intro2/ *
-#W2 #T2 #HW2 #HT2 #H destruct
-lapply (cpxs_strap1 … HW1 … HW2) -W
-lapply (cpxs_strap1 … HT1 … HT2) -T /3 width=5 by or3_intro0, ex3_2_intro/
-qed-.
-
-lemma cpxs_inv_cnx1: ∀h,o,G,L,T,U. ⦃G, L⦄ ⊢ T ➡*[h, o] U → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ → T = U.
-#h #o #G #L #T #U #H @(cpxs_ind_dx … H) -T //
-#T0 #T #H1T0 #_ #IHT #H2T0
-lapply (H2T0 … H1T0) -H1T0 #H destruct /2 width=1 by/
-qed-.
-
-lemma cpxs_neq_inv_step_sn: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → (T1 = T2 → ⊥) →
- ∃∃T. ⦃G, L⦄ ⊢ T1 ➡[h, o] T & T1 = T → ⊥ & ⦃G, L⦄ ⊢ T ➡*[h, o] T2.
-#h #o #G #L #T1 #T2 #H @(cpxs_ind_dx … H) -T1
-[ #H elim H -H //
-| #T1 #T #H1 #H2 #IH2 #H12 elim (eq_term_dec T1 T) #H destruct
- [ -H1 -H2 /3 width=1 by/
- | -IH2 /3 width=4 by ex3_intro/ (**) (* auto fails without clear *)
- ]
-]
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/reduction/lpx_aaa.ma".
-include "basic_2/computation/cpxs.ma".
-
-(* CONTEXT-SENSITIVE EXTENDED PARALLEL COMPUTATION ON TERMS *****************)
-
-(* Properties about atomic arity assignment on terms ************************)
-
-lemma cpxs_aaa_conf: ∀h,o,G,L,T1,A. ⦃G, L⦄ ⊢ T1 ⁝ A →
- ∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → ⦃G, L⦄ ⊢ T2 ⁝ A.
-#h #o #G #L #T1 #A #HT1 #T2 #HT12
-@(TC_Conf3 … HT1 ? HT12) -A -T1 -T2 /2 width=5 by cpx_aaa_conf/
-qed-.
-
-lemma cprs_aaa_conf: ∀G,L,T1,A. ⦃G, L⦄ ⊢ T1 ⁝ A → ∀T2. ⦃G, L⦄ ⊢ T1 ➡* T2 → ⦃G, L⦄ ⊢ T2 ⁝ A.
-/3 width=5 by cpxs_aaa_conf, cprs_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/lpx_drop.ma".
-include "basic_2/computation/cpxs_lift.ma".
-
-(* CONTEXT-SENSITIVE EXTENDED PARALLEL COMPUTATION ON TERMS *****************)
-
-(* Main properties **********************************************************)
-
-theorem cpxs_trans: ∀h,o,G,L. Transitive … (cpxs h o G L).
-normalize /2 width=3 by trans_TC/ qed-.
-
-theorem cpxs_bind: ∀h,o,a,I,G,L,V1,V2,T1,T2. ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ➡*[h, o] T2 →
- ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 →
- ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡*[h, o] ⓑ{a,I}V2.T2.
-#h #o #a #I #G #L #V1 #V2 #T1 #T2 #HT12 #H @(cpxs_ind … H) -V2
-/3 width=5 by cpxs_trans, cpxs_bind_dx/
-qed.
-
-theorem cpxs_flat: ∀h,o,I,G,L,V1,V2,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 →
- ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 →
- ⦃G, L⦄ ⊢ ⓕ{I}V1.T1 ➡*[h, o] ⓕ{I}V2.T2.
-#h #o #I #G #L #V1 #V2 #T1 #T2 #HT12 #H @(cpxs_ind … H) -V2
-/3 width=5 by cpxs_trans, cpxs_flat_dx/
-qed.
-
-theorem cpxs_beta_rc: ∀h,o,a,G,L,V1,V2,W1,W2,T1,T2.
- ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 → ⦃G, L.ⓛW1⦄ ⊢ T1 ➡*[h, o] T2 → ⦃G, L⦄ ⊢ W1 ➡*[h, o] W2 →
- ⦃G, L⦄ ⊢ ⓐV1.ⓛ{a}W1.T1 ➡*[h, o] ⓓ{a}ⓝW2.V2.T2.
-#h #o #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #HV12 #HT12 #H @(cpxs_ind … H) -W2
-/4 width=5 by cpxs_trans, cpxs_beta_dx, cpxs_bind_dx, cpx_pair_sn/
-qed.
-
-theorem cpxs_beta: ∀h,o,a,G,L,V1,V2,W1,W2,T1,T2.
- ⦃G, L.ⓛW1⦄ ⊢ T1 ➡*[h, o] T2 → ⦃G, L⦄ ⊢ W1 ➡*[h, o] W2 → ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 →
- ⦃G, L⦄ ⊢ ⓐV1.ⓛ{a}W1.T1 ➡*[h, o] ⓓ{a}ⓝW2.V2.T2.
-#h #o #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #HT12 #HW12 #H @(cpxs_ind … H) -V2
-/4 width=5 by cpxs_trans, cpxs_beta_rc, cpxs_bind_dx, cpx_flat/
-qed.
-
-theorem cpxs_theta_rc: ∀h,o,a,G,L,V1,V,V2,W1,W2,T1,T2.
- ⦃G, L⦄ ⊢ V1 ➡[h, o] V → ⬆[0, 1] V ≡ V2 →
- ⦃G, L.ⓓW1⦄ ⊢ T1 ➡*[h, o] T2 → ⦃G, L⦄ ⊢ W1 ➡*[h, o] W2 →
- ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}W1.T1 ➡*[h, o] ⓓ{a}W2.ⓐV2.T2.
-#h #o #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV1 #HV2 #HT12 #H @(cpxs_ind … H) -W2
-/3 width=5 by cpxs_trans, cpxs_theta_dx, cpxs_bind_dx/
-qed.
-
-theorem cpxs_theta: ∀h,o,a,G,L,V1,V,V2,W1,W2,T1,T2.
- ⬆[0, 1] V ≡ V2 → ⦃G, L⦄ ⊢ W1 ➡*[h, o] W2 →
- ⦃G, L.ⓓW1⦄ ⊢ T1 ➡*[h, o] T2 → ⦃G, L⦄ ⊢ V1 ➡*[h, o] V →
- ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}W1.T1 ➡*[h, o] ⓓ{a}W2.ⓐV2.T2.
-#h #o #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV2 #HW12 #HT12 #H @(TC_ind_dx … V1 H) -V1
-/3 width=5 by cpxs_trans, cpxs_theta_rc, cpxs_flat_dx/
-qed.
-
-(* Advanced inversion lemmas ************************************************)
-
-lemma cpxs_inv_appl1: ∀h,o,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓐV1.T1 ➡*[h, o] U2 →
- ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 & ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 &
- U2 = ⓐV2. T2
- | ∃∃a,W,T. ⦃G, L⦄ ⊢ T1 ➡*[h, o] ⓛ{a}W.T & ⦃G, L⦄ ⊢ ⓓ{a}ⓝW.V1.T ➡*[h, o] U2
- | ∃∃a,V0,V2,V,T. ⦃G, L⦄ ⊢ V1 ➡*[h, o] V0 & ⬆[0,1] V0 ≡ V2 &
- ⦃G, L⦄ ⊢ T1 ➡*[h, o] ⓓ{a}V.T & ⦃G, L⦄ ⊢ ⓓ{a}V.ⓐV2.T ➡*[h, o] U2.
-#h #o #G #L #V1 #T1 #U2 #H @(cpxs_ind … H) -U2 [ /3 width=5 by or3_intro0, ex3_2_intro/ ]
-#U #U2 #_ #HU2 * *
-[ #V0 #T0 #HV10 #HT10 #H destruct
- elim (cpx_inv_appl1 … HU2) -HU2 *
- [ #V2 #T2 #HV02 #HT02 #H destruct /4 width=5 by cpxs_strap1, or3_intro0, ex3_2_intro/
- | #a #V2 #W #W2 #T #T2 #HV02 #HW2 #HT2 #H1 #H2 destruct
- lapply (cpxs_strap1 … HV10 … HV02) -V0 #HV12
- lapply (lsubr_cpx_trans … HT2 (L.ⓓⓝW.V1) ?) -HT2
- /5 width=5 by cpxs_bind, cpxs_flat_dx, cpx_cpxs, lsubr_beta, ex2_3_intro, or3_intro1/
- | #a #V #V2 #W0 #W2 #T #T2 #HV0 #HV2 #HW02 #HT2 #H1 #H2 destruct
- /5 width=10 by cpxs_flat_sn, cpxs_bind_dx, cpxs_strap1, ex4_5_intro, or3_intro2/
- ]
-| /4 width=9 by cpxs_strap1, or3_intro1, ex2_3_intro/
-| /4 width=11 by cpxs_strap1, or3_intro2, ex4_5_intro/
-]
-qed-.
-
-(* Properties on sn extended parallel reduction for local environments ******)
-
-lemma lpx_cpx_trans: ∀h,o,G. c_r_transitive … (cpx h o G) (λ_.lpx h o G).
-#h #o #G #L2 #T1 #T2 #HT12 elim HT12 -G -L2 -T1 -T2
-[ /2 width=3 by/
-| /3 width=2 by cpx_cpxs, cpx_st/
-| #I #G #L2 #K2 #V0 #V2 #W2 #i #HLK2 #_ #HVW2 #IHV02 #L1 #HL12
- elim (lpx_drop_trans_O1 … HL12 … HLK2) -L2 #X #HLK1 #H
- elim (lpx_inv_pair2 … H) -H #K1 #V1 #HK12 #HV10 #H destruct
- /4 width=7 by cpxs_delta, cpxs_strap2/
-|4,9: /4 width=1 by cpxs_beta, cpxs_bind, lpx_pair/
-|5,7,8: /3 width=1 by cpxs_flat, cpxs_ct, cpxs_eps/
-| /4 width=3 by cpxs_zeta, lpx_pair/
-| /4 width=3 by cpxs_theta, cpxs_strap1, lpx_pair/
-]
-qed-.
-
-lemma cpx_bind2: ∀h,o,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 →
- ∀I,T1,T2. ⦃G, L.ⓑ{I}V2⦄ ⊢ T1 ➡[h, o] T2 →
- ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡*[h, o] ⓑ{a,I}V2.T2.
-/4 width=5 by lpx_cpx_trans, cpxs_bind_dx, lpx_pair/ qed.
-
-(* Advanced properties ******************************************************)
-
-lemma lpx_cpxs_trans: ∀h,o,G. c_rs_transitive … (cpx h o G) (λ_.lpx h o G).
-#h #o #G @c_r_trans_LTC1 /2 width=3 by lpx_cpx_trans/ (**) (* full auto fails *)
-qed-.
-
-lemma cpxs_bind2_dx: ∀h,o,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 →
- ∀I,T1,T2. ⦃G, L.ⓑ{I}V2⦄ ⊢ T1 ➡*[h, o] T2 →
- ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡*[h, o] ⓑ{a,I}V2.T2.
-/4 width=5 by lpx_cpxs_trans, cpxs_bind_dx, lpx_pair/ qed.
-
-(* Properties on supclosure *************************************************)
-
-lemma fqu_cpxs_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
- ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, o] U2 → (T2 = U2 → ⊥) →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐ ⦃G2, L2, U2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
-[ #I #G #L #V1 #V2 #HV12 #_ elim (lift_total V2 0 1)
- #U2 #HVU2 @(ex3_intro … U2)
- [1,3: /3 width=7 by fqu_drop, cpxs_delta, drop_pair, drop_drop/
- | #H destruct
- lapply (lift_inv_lref2_be … HVU2 ? ?) -HVU2 //
- ]
-| #I #G #L #V1 #T #V2 #HV12 #H @(ex3_intro … (②{I}V2.T))
- [1,3: /2 width=4 by fqu_pair_sn, cpxs_pair_sn/
- | #H0 destruct /2 width=1 by/
- ]
-| #a #I #G #L #V #T1 #T2 #HT12 #H @(ex3_intro … (ⓑ{a,I}V.T2))
- [1,3: /2 width=4 by fqu_bind_dx, cpxs_bind/
- | #H0 destruct /2 width=1 by/
- ]
-| #I #G #L #V #T1 #T2 #HT12 #H @(ex3_intro … (ⓕ{I}V.T2))
- [1,3: /2 width=4 by fqu_flat_dx, cpxs_flat/
- | #H0 destruct /2 width=1 by/
- ]
-| #G #L #K #T1 #U1 #k #HLK #HTU1 #T2 #HT12 #H elim (lift_total T2 0 (k+1))
- #U2 #HTU2 @(ex3_intro … U2)
- [1,3: /2 width=10 by cpxs_lift, fqu_drop/
- | #H0 destruct /3 width=5 by lift_inj/
-]
-qed-.
-
-lemma fquq_cpxs_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, o] U2 → (T2 = U2 → ⊥) →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H12 #U2 #HTU2 #H elim (fquq_inv_gen … H12) -H12
-[ #H12 elim (fqu_cpxs_trans_neq … H12 … HTU2 H) -T2
- /3 width=4 by fqu_fquq, ex3_intro/
-| * #HG #HL #HT destruct /3 width=4 by ex3_intro/
-]
-qed-.
-
-lemma fqup_cpxs_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
- ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, o] U2 → (T2 = U2 → ⊥) →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐+ ⦃G2, L2, U2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind_dx … H) -G1 -L1 -T1
-[ #G1 #L1 #T1 #H12 #U2 #HTU2 #H elim (fqu_cpxs_trans_neq … H12 … HTU2 H) -T2
- /3 width=4 by fqu_fqup, ex3_intro/
-| #G #G1 #L #L1 #T #T1 #H1 #_ #IH12 #U2 #HTU2 #H elim (IH12 … HTU2 H) -T2
- #U1 #HTU1 #H #H12 elim (fqu_cpxs_trans_neq … H1 … HTU1 H) -T1
- /3 width=8 by fqup_strap2, ex3_intro/
-]
-qed-.
-
-lemma fqus_cpxs_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
- ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, o] U2 → (T2 = U2 → ⊥) →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐* ⦃G2, L2, U2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H12 #U2 #HTU2 #H elim (fqus_inv_gen … H12) -H12
-[ #H12 elim (fqup_cpxs_trans_neq … H12 … HTU2 H) -T2
- /3 width=4 by fqup_fqus, ex3_intro/
-| * #HG #HL #HT destruct /3 width=4 by ex3_intro/
-]
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/multiple/fqus_fqus.ma".
-include "basic_2/reduction/cpx_lift.ma".
-include "basic_2/computation/cpxs.ma".
-
-(* CONTEXT-SENSITIVE EXTENDED PARALLEL COMPUTATION ON TERMS *****************)
-
-(* Advanced properties ******************************************************)
-
-lemma cpxs_delta: ∀h,o,I,G,L,K,V,V2,i.
- ⬇[i] L ≡ K.ⓑ{I}V → ⦃G, K⦄ ⊢ V ➡*[h, o] V2 →
- ∀W2. ⬆[0, i+1] V2 ≡ W2 → ⦃G, L⦄ ⊢ #i ➡*[h, o] W2.
-#h #o #I #G #L #K #V #V2 #i #HLK #H elim H -V2
-[ /3 width=9 by cpx_cpxs, cpx_delta/
-| #V1 lapply (drop_fwd_drop2 … HLK) -HLK
- elim (lift_total V1 0 (i+1)) /4 width=12 by cpx_lift, cpxs_strap1/
-]
-qed.
-
-lemma lstas_cpxs: ∀h,o,G,L,T1,T2,d2. ⦃G, L⦄ ⊢ T1 •*[h, d2] T2 →
- ∀d1. ⦃G, L⦄ ⊢ T1 ▪[h, o] d1 → d2 ≤ d1 → ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2.
-#h #o #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 #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 #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 … 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/
-| /4 width=3 by cpxs_eps, da_inv_flat/
-]
-qed.
-
-(* Advanced inversion lemmas ************************************************)
-
-lemma cpxs_inv_lref1: ∀h,o,G,L,T2,i. ⦃G, L⦄ ⊢ #i ➡*[h, o] T2 →
- T2 = #i ∨
- ∃∃I,K,V1,T1. ⬇[i] L ≡ K.ⓑ{I}V1 & ⦃G, K⦄ ⊢ V1 ➡*[h, o] T1 &
- ⬆[0, i+1] T1 ≡ T2.
-#h #o #G #L #T2 #i #H @(cpxs_ind … H) -T2 /2 width=1 by or_introl/
-#T #T2 #_ #HT2 *
-[ #H destruct
- elim (cpx_inv_lref1 … HT2) -HT2 /2 width=1 by or_introl/
- * /4 width=7 by cpx_cpxs, ex3_4_intro, or_intror/
-| * #I #K #V1 #T1 #HLK #HVT1 #HT1
- lapply (drop_fwd_drop2 … HLK) #H0LK
- elim (cpx_inv_lift1 … HT2 … H0LK … HT1) -H0LK -T
- /4 width=7 by cpxs_strap1, ex3_4_intro, or_intror/
-]
-qed-.
-
-(* Relocation properties ****************************************************)
-
-lemma cpxs_lift: ∀h,o,G. d_liftable (cpxs h o G).
-/3 width=10 by cpx_lift, cpxs_strap1, d_liftable_LTC/ qed.
-
-lemma cpxs_inv_lift1: ∀h,o,G. d_deliftable_sn (cpxs h o G).
-/3 width=6 by d_deliftable_sn_LTC, cpx_inv_lift1/
-qed-.
-
-(* Properties on supclosure *************************************************)
-
-lemma fqu_cpxs_trans: ∀h,o,G1,G2,L1,L2,T2,U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, o] U2 →
- ∀T1. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] U1 & ⦃G1, L1, U1⦄ ⊐ ⦃G2, L2, U2⦄.
-#h #o #G1 #G2 #L1 #L2 #T2 #U2 #H @(cpxs_ind_dx … H) -T2 /2 width=3 by ex2_intro/
-#T #T2 #HT2 #_ #IHTU2 #T1 #HT1 elim (fqu_cpx_trans … HT1 … HT2) -T
-#T #HT1 #HT2 elim (IHTU2 … HT2) -T2 /3 width=3 by cpxs_strap2, ex2_intro/
-qed-.
-
-lemma fquq_cpxs_trans: ∀h,o,G1,G2,L1,L2,T2,U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, o] U2 →
- ∀T1. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] U1 & ⦃G1, L1, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
-#h #o #G1 #G2 #L1 #L2 #T2 #U2 #HTU2 #T1 #H elim (fquq_inv_gen … H) -H
-[ #HT12 elim (fqu_cpxs_trans … HTU2 … HT12) /3 width=3 by fqu_fquq, ex2_intro/
-| * #H1 #H2 #H3 destruct /2 width=3 by ex2_intro/
-]
-qed-.
-
-lemma fquq_lstas_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ∀U2,d1. ⦃G2, L2⦄ ⊢ T2 •*[h, d1] U2 →
- ∀d2. ⦃G2, L2⦄ ⊢ T2 ▪[h, o] d2 → d1 ≤ d2 →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] U1 & ⦃G1, L1, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
-/3 width=5 by fquq_cpxs_trans, lstas_cpxs/ qed-.
-
-lemma fqup_cpxs_trans: ∀h,o,G1,G2,L1,L2,T2,U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, o] U2 →
- ∀T1. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] U1 & ⦃G1, L1, U1⦄ ⊐+ ⦃G2, L2, U2⦄.
-#h #o #G1 #G2 #L1 #L2 #T2 #U2 #H @(cpxs_ind_dx … H) -T2 /2 width=3 by ex2_intro/
-#T #T2 #HT2 #_ #IHTU2 #T1 #HT1 elim (fqup_cpx_trans … HT1 … HT2) -T
-#U1 #HTU1 #H2 elim (IHTU2 … H2) -T2 /3 width=3 by cpxs_strap2, ex2_intro/
-qed-.
-
-lemma fqus_cpxs_trans: ∀h,o,G1,G2,L1,L2,T2,U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, o] U2 →
- ∀T1. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] U1 & ⦃G1, L1, U1⦄ ⊐* ⦃G2, L2, U2⦄.
-#h #o #G1 #G2 #L1 #L2 #T2 #U2 #HTU2 #T1 #H elim (fqus_inv_gen … H) -H
-[ #HT12 elim (fqup_cpxs_trans … HTU2 … HT12) /3 width=3 by fqup_fqus, ex2_intro/
-| * #H1 #H2 #H3 destruct /2 width=3 by ex2_intro/
-]
-qed-.
-
-lemma fqus_lstas_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
- ∀U2,d1. ⦃G2, L2⦄ ⊢ T2 •*[h, d1] U2 →
- ∀d2. ⦃G2, L2⦄ ⊢ T2 ▪[h, o] d2 → d1 ≤ d2 →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] 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_lleq.ma".
-include "basic_2/computation/cpxs.ma".
-
-(* CONTEXT-SENSITIVE EXTENDED PARALLEL COMPUTATION ON TERMS *****************)
-
-(* Properties on lazy equivalence for local environments ********************)
-
-lemma lleq_cpxs_trans: ∀h,o,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ➡*[h, o] T2 →
- ∀L1. L1 ≡[T1, 0] L2 → ⦃G, L1⦄ ⊢ T1 ➡*[h, o] T2.
-#h #o #G #L2 #T1 #T2 #H @(cpxs_ind_dx … H) -T1
-/4 width=6 by cpx_lleq_conf_dx, lleq_cpx_trans, cpxs_strap2/
-qed-.
-
-lemma cpxs_lleq_conf: ∀h,o,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ➡*[h, o] T2 →
- ∀L1. L2 ≡[T1, 0] L1 → ⦃G, L1⦄ ⊢ T1 ➡*[h, o] T2.
-/3 width=3 by lleq_cpxs_trans, lleq_sym/ qed-.
-
-lemma cpxs_lleq_conf_dx: ∀h,o,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ➡*[h, o] T2 →
- ∀L1. L1 ≡[T1, 0] L2 → L1 ≡[T2, 0] L2.
-#h #o #G #L2 #T1 #T2 #H @(cpxs_ind … H) -T2 /3 width=6 by cpx_lleq_conf_dx/
-qed-.
-
-lemma cpxs_lleq_conf_sn: ∀h,o,G,L1,T1,T2. ⦃G, L1⦄ ⊢ T1 ➡*[h, o] T2 →
- ∀L2. L1 ≡[T1, 0] L2 → L1 ≡[T2, 0] L2.
-/4 width=6 by cpxs_lleq_conf_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/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,o,G. lsub_trans … (cpxs h o G) (lreq 0 (∞)).
-/3 width=5 by lreq_cpx_trans, LTC_lsub_trans/
-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/tsts.ma".
-include "basic_2/computation/lpxs_cpxs.ma".
-
-(* CONTEXT-SENSITIVE EXTENDED PARALLEL COMPUTATION ON TERMS *****************)
-
-(* Forward lemmas involving same top term structure *************************)
-
-lemma cpxs_fwd_cnx: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ → ∀U. ⦃G, L⦄ ⊢ T ➡*[h, o] U → T ≂ U.
-#h #o #G #L #T #HT #U #H
->(cpxs_inv_cnx1 … H HT) -G -L -T //
-qed-.
-
-lemma cpxs_fwd_sort: ∀h,o,G,L,U,s. ⦃G, L⦄ ⊢ ⋆s ➡*[h, o] U →
- ⋆s ≂ U ∨ ⦃G, L⦄ ⊢ ⋆(next h s) ➡*[h, o] U.
-#h #o #G #L #U #s #H
-elim (cpxs_inv_sort1 … H) -H #n #d generalize in match s; -s @(nat_ind_plus … n) -n
-[ #s #_ #H -d destruct /2 width=1 by or_introl/
-| #n #IHn #s >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 Hnd; -Hnd @(nat_ind_plus … n) -n
- /4 width=3 by cpxs_strap2, cpx_st, or_intror/
- | >iter_SO >iter_n_Sm //
- ]
-]
-qed-.
-
-(* Basic_1: was just: pr3_iso_beta *)
-lemma cpxs_fwd_beta: ∀h,o,a,G,L,V,W,T,U. ⦃G, L⦄ ⊢ ⓐV.ⓛ{a}W.T ➡*[h, o] U →
- ⓐV.ⓛ{a}W.T ≂ U ∨ ⦃G, L⦄ ⊢ ⓓ{a}ⓝW.V.T ➡*[h, o] U.
-#h #o #a #G #L #V #W #T #U #H
-elim (cpxs_inv_appl1 … H) -H *
-[ #V0 #T0 #_ #_ #H destruct /2 width=1 by tsts_pair, or_introl/
-| #b #W0 #T0 #HT0 #HU
- elim (cpxs_inv_abst1 … HT0) -HT0 #W1 #T1 #HW1 #HT1 #H destruct
- lapply (lsubr_cpxs_trans … HT1 (L.ⓓⓝW.V) ?) -HT1
- /5 width=3 by cpxs_trans, cpxs_bind, cpxs_pair_sn, lsubr_beta, or_intror/
-| #b #V1 #V2 #V0 #T1 #_ #_ #HT1 #_
- elim (cpxs_inv_abst1 … HT1) -HT1 #W2 #T2 #_ #_ #H destruct
-]
-qed-.
-
-(* Note: probably this is an inversion lemma *)
-lemma cpxs_fwd_delta: ∀h,o,I,G,L,K,V1,i. ⬇[i] L ≡ K.ⓑ{I}V1 →
- ∀V2. ⬆[0, i + 1] V1 ≡ V2 →
- ∀U. ⦃G, L⦄ ⊢ #i ➡*[h, o] U →
- #i ≂ U ∨ ⦃G, L⦄ ⊢ V2 ➡*[h, o] U.
-#h #o #I #G #L #K #V1 #i #HLK #V2 #HV12 #U #H
-elim (cpxs_inv_lref1 … H) -H /2 width=1 by or_introl/
-* #I0 #K0 #V0 #U0 #HLK0 #HVU0 #HU0
-lapply (drop_mono … HLK0 … HLK) -HLK0 #H destruct
-/4 width=10 by cpxs_lift, drop_fwd_drop2, or_intror/
-qed-.
-
-lemma cpxs_fwd_theta: ∀h,o,a,G,L,V1,V,T,U. ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}V.T ➡*[h, o] U →
- ∀V2. ⬆[0, 1] V1 ≡ V2 → ⓐV1.ⓓ{a}V.T ≂ U ∨
- ⦃G, L⦄ ⊢ ⓓ{a}V.ⓐV2.T ➡*[h, o] U.
-#h #o #a #G #L #V1 #V #T #U #H #V2 #HV12
-elim (cpxs_inv_appl1 … H) -H *
-[ -HV12 #V0 #T0 #_ #_ #H destruct /2 width=1 by tsts_pair, or_introl/
-| #b #W #T0 #HT0 #HU
- elim (cpxs_inv_abbr1 … HT0) -HT0 *
- [ #V3 #T3 #_ #_ #H destruct
- | #X #HT2 #H #H0 destruct
- elim (lift_inv_bind1 … H) -H #W2 #T2 #HW2 #HT02 #H destruct
- @or_intror @(cpxs_trans … HU) -U (**) (* explicit constructor *)
- @(cpxs_trans … (+ⓓV.ⓐV2.ⓛ{b}W2.T2)) [ /3 width=1 by cpxs_flat_dx, cpxs_bind_dx/ ] -T
- @(cpxs_strap2 … (ⓐV1.ⓛ{b}W.T0)) [2: /2 width=1 by cpxs_beta_dx/ ]
- /4 width=7 by cpx_zeta, lift_bind, lift_flat/
- ]
-| #b #V3 #V4 #V0 #T0 #HV13 #HV34 #HT0 #HU
- @or_intror @(cpxs_trans … HU) -U (**) (* explicit constructor *)
- elim (cpxs_inv_abbr1 … HT0) -HT0 *
- [ #V5 #T5 #HV5 #HT5 #H destruct
- lapply (cpxs_lift … HV13 (L.ⓓV) … HV12 … HV34) -V1 -V3
- /3 width=2 by cpxs_flat, cpxs_bind, drop_drop/
- | #X #HT1 #H #H0 destruct
- elim (lift_inv_bind1 … H) -H #V5 #T5 #HV05 #HT05 #H destruct
- lapply (cpxs_lift … HV13 (L.ⓓV0) … HV12 … HV34) -V3 /2 width=2 by drop_drop/ #HV24
- @(cpxs_trans … (+ⓓV.ⓐV2.ⓓ{b}V5.T5)) [ /3 width=1 by cpxs_flat_dx, cpxs_bind_dx/ ] -T
- @(cpxs_strap2 … (ⓐV1.ⓓ{b}V0.T0)) [ /4 width=7 by cpx_zeta, lift_bind, lift_flat/ ] -V -V5 -T5
- @(cpxs_strap2 … (ⓓ{b}V0.ⓐV2.T0)) /3 width=3 by cpxs_pair_sn, cpxs_bind_dx, cpr_cpx, cpr_theta/
- ]
-]
-qed-.
-
-lemma cpxs_fwd_cast: ∀h,o,G,L,W,T,U. ⦃G, L⦄ ⊢ ⓝW.T ➡*[h, o] U →
- ∨∨ ⓝW. T ≂ U | ⦃G, L⦄ ⊢ T ➡*[h, o] U | ⦃G, L⦄ ⊢ W ➡*[h, o] U.
-#h #o #G #L #W #T #U #H
-elim (cpxs_inv_cast1 … H) -H /2 width=1 by or3_intro1, or3_intro2/ *
-#W0 #T0 #_ #_ #H destruct /2 width=1 by tsts_pair, or3_intro0/
-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/tsts_vector.ma".
-include "basic_2/substitution/lift_vector.ma".
-include "basic_2/computation/cpxs_tsts.ma".
-
-(* CONTEXT-SENSITIVE EXTENDED PARALLEL COMPUTATION ON TERMS *****************)
-
-(* Vector form of forward lemmas involving same top term structure **********)
-
-(* Basic_1: was just: nf2_iso_appls_lref *)
-lemma cpxs_fwd_cnx_vector: ∀h,o,G,L,T. 𝐒⦃T⦄ → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ →
- ∀Vs,U. ⦃G, L⦄ ⊢ ⒶVs.T ➡*[h, o] U → ⒶVs.T ≂ U.
-#h #o #G #L #T #H1T #H2T #Vs elim Vs -Vs [ @(cpxs_fwd_cnx … H2T) ] (**) (* /2 width=3 by cpxs_fwd_cnx/ does not work *)
-#V #Vs #IHVs #U #H
-elim (cpxs_inv_appl1 … H) -H *
-[ -IHVs #V0 #T0 #_ #_ #H destruct /2 width=1 by tsts_pair/
-| #a #W0 #T0 #HT0 #HU
- lapply (IHVs … HT0) -IHVs -HT0 #HT0
- elim (tsts_inv_bind_applv_simple … HT0) //
-| #a #V1 #V2 #V0 #T0 #HV1 #HV12 #HT0 #HU
- lapply (IHVs … HT0) -IHVs -HT0 #HT0
- elim (tsts_inv_bind_applv_simple … HT0) //
-]
-qed-.
-
-lemma cpxs_fwd_sort_vector: ∀h,o,G,L,s,Vs,U. ⦃G, L⦄ ⊢ ⒶVs.⋆s ➡*[h, o] U →
- ⒶVs.⋆s ≂ U ∨ ⦃G, L⦄ ⊢ ⒶVs.⋆(next h s) ➡*[h, o] U.
-#h #o #G #L #s #Vs elim Vs -Vs /2 width=1 by cpxs_fwd_sort/
-#V #Vs #IHVs #U #H
-elim (cpxs_inv_appl1 … H) -H *
-[ -IHVs #V1 #T1 #_ #_ #H destruct /2 width=1 by tsts_pair, or_introl/
-| #a #W1 #T1 #HT1 #HU
- elim (IHVs … HT1) -IHVs -HT1 #HT1
- [ elim (tsts_inv_bind_applv_simple … HT1) //
- | @or_intror (**) (* explicit constructor *)
- @(cpxs_trans … HU) -U
- @(cpxs_strap1 … (ⓐV.ⓛ{a}W1.T1)) /3 width=1 by cpxs_flat_dx, cpr_cpx, cpr_beta/
- ]
-| #a #V1 #V2 #V3 #T1 #HV01 #HV12 #HT1 #HU
- elim (IHVs … HT1) -IHVs -HT1 #HT1
- [ elim (tsts_inv_bind_applv_simple … HT1) //
- | @or_intror (**) (* explicit constructor *)
- @(cpxs_trans … HU) -U
- @(cpxs_strap1 … (ⓐV1.ⓓ{a}V3.T1)) /3 width=3 by cpxs_flat, cpr_cpx, cpr_theta/
- ]
-]
-qed-.
-
-
-(* Basic_1: was just: pr3_iso_appls_beta *)
-lemma cpxs_fwd_beta_vector: ∀h,o,a,G,L,Vs,V,W,T,U. ⦃G, L⦄ ⊢ ⒶVs.ⓐV.ⓛ{a}W.T ➡*[h, o] U →
- ⒶVs. ⓐV. ⓛ{a}W. T ≂ U ∨ ⦃G, L⦄ ⊢ ⒶVs.ⓓ{a}ⓝW.V.T ➡*[h, o] U.
-#h #o #a #G #L #Vs elim Vs -Vs /2 width=1 by cpxs_fwd_beta/
-#V0 #Vs #IHVs #V #W #T #U #H
-elim (cpxs_inv_appl1 … H) -H *
-[ -IHVs #V1 #T1 #_ #_ #H destruct /2 width=1 by tsts_pair, or_introl/
-| #b #W1 #T1 #HT1 #HU
- elim (IHVs … HT1) -IHVs -HT1 #HT1
- [ elim (tsts_inv_bind_applv_simple … HT1) //
- | @or_intror (**) (* explicit constructor *)
- @(cpxs_trans … HU) -U
- @(cpxs_strap1 … (ⓐV0.ⓛ{b}W1.T1)) /3 width=1 by cpxs_flat_dx, cpr_cpx, cpr_beta/
- ]
-| #b #V1 #V2 #V3 #T1 #HV01 #HV12 #HT1 #HU
- elim (IHVs … HT1) -IHVs -HT1 #HT1
- [ elim (tsts_inv_bind_applv_simple … HT1) //
- | @or_intror (**) (* explicit constructor *)
- @(cpxs_trans … HU) -U
- @(cpxs_strap1 … (ⓐV1.ⓓ{b}V3.T1)) /3 width=3 by cpxs_flat, cpr_cpx, cpr_theta/
- ]
-]
-qed-.
-
-lemma cpxs_fwd_delta_vector: ∀h,o,I,G,L,K,V1,i. ⬇[i] L ≡ K.ⓑ{I}V1 →
- ∀V2. ⬆[0, i + 1] V1 ≡ V2 →
- ∀Vs,U. ⦃G, L⦄ ⊢ ⒶVs.#i ➡*[h, o] U →
- ⒶVs.#i ≂ U ∨ ⦃G, L⦄ ⊢ ⒶVs.V2 ➡*[h, o] U.
-#h #o #I #G #L #K #V1 #i #HLK #V2 #HV12 #Vs elim Vs -Vs /2 width=5 by cpxs_fwd_delta/
-#V #Vs #IHVs #U #H -K -V1
-elim (cpxs_inv_appl1 … H) -H *
-[ -IHVs #V0 #T0 #_ #_ #H destruct /2 width=1 by tsts_pair, or_introl/
-| #b #W0 #T0 #HT0 #HU
- elim (IHVs … HT0) -IHVs -HT0 #HT0
- [ elim (tsts_inv_bind_applv_simple … HT0) //
- | @or_intror -i (**) (* explicit constructor *)
- @(cpxs_trans … HU) -U
- @(cpxs_strap1 … (ⓐV.ⓛ{b}W0.T0)) /3 width=1 by cpxs_flat_dx, cpr_cpx, cpr_beta/
- ]
-| #b #V0 #V1 #V3 #T0 #HV0 #HV01 #HT0 #HU
- elim (IHVs … HT0) -IHVs -HT0 #HT0
- [ elim (tsts_inv_bind_applv_simple … HT0) //
- | @or_intror -i (**) (* explicit constructor *)
- @(cpxs_trans … HU) -U
- @(cpxs_strap1 … (ⓐV0.ⓓ{b}V3.T0)) /3 width=3 by cpxs_flat, cpr_cpx, cpr_theta/
- ]
-]
-qed-.
-
-(* Basic_1: was just: pr3_iso_appls_abbr *)
-lemma cpxs_fwd_theta_vector: ∀h,o,G,L,V1c,V2c. ⬆[0, 1] V1c ≡ V2c →
- ∀a,V,T,U. ⦃G, L⦄ ⊢ ⒶV1c.ⓓ{a}V.T ➡*[h, o] U →
- ⒶV1c. ⓓ{a}V. T ≂ U ∨ ⦃G, L⦄ ⊢ ⓓ{a}V.ⒶV2c.T ➡*[h, o] U.
-#h #o #G #L #V1c #V2c * -V1c -V2c /3 width=1 by or_intror/
-#V1c #V2c #V1a #V2a #HV12a #HV12c #a
-generalize in match HV12a; -HV12a
-generalize in match V2a; -V2a
-generalize in match V1a; -V1a
-elim HV12c -V1c -V2c /2 width=1 by cpxs_fwd_theta/
-#V1c #V2c #V1b #V2b #HV12b #_ #IHV12c #V1a #V2a #HV12a #V #T #U #H
-elim (cpxs_inv_appl1 … H) -H *
-[ -IHV12c -HV12a -HV12b #V0 #T0 #_ #_ #H destruct /2 width=1 by tsts_pair, or_introl/
-| #b #W0 #T0 #HT0 #HU
- elim (IHV12c … HV12b … HT0) -IHV12c -HT0 #HT0
- [ -HV12a -HV12b -HU
- elim (tsts_inv_pair1 … HT0) #V1 #T1 #H destruct
- | @or_intror -V1c (**) (* explicit constructor *)
- @(cpxs_trans … HU) -U
- elim (cpxs_inv_abbr1 … HT0) -HT0 *
- [ -HV12a -HV12b #V1 #T1 #_ #_ #H destruct
- | -V1b #X #HT1 #H #H0 destruct
- elim (lift_inv_bind1 … H) -H #W1 #T1 #HW01 #HT01 #H destruct
- @(cpxs_trans … (+ⓓV.ⓐV2a.ⓛ{b}W1.T1)) [ /3 width=1 by cpxs_flat_dx, cpxs_bind_dx/ ] -T -V2b -V2c
- @(cpxs_strap2 … (ⓐV1a.ⓛ{b}W0.T0))
- /4 width=7 by cpxs_beta_dx, cpx_zeta, lift_bind, lift_flat/
- ]
- ]
-| #b #V0a #Va #V0 #T0 #HV10a #HV0a #HT0 #HU
- elim (IHV12c … HV12b … HT0) -HV12b -IHV12c -HT0 #HT0
- [ -HV12a -HV10a -HV0a -HU
- elim (tsts_inv_pair1 … HT0) #V1 #T1 #H destruct
- | @or_intror -V1c -V1b (**) (* explicit constructor *)
- @(cpxs_trans … HU) -U
- elim (cpxs_inv_abbr1 … HT0) -HT0 *
- [ #V1 #T1 #HV1 #HT1 #H destruct
- lapply (cpxs_lift … HV10a (L.ⓓV) (Ⓕ) … HV12a … HV0a) -V1a -V0a [ /2 width=1 by drop_drop/ ] #HV2a
- @(cpxs_trans … (ⓓ{a}V.ⓐV2a.T1)) /3 width=1 by cpxs_bind, cpxs_pair_sn, cpxs_flat_dx, cpxs_bind_dx/
- | #X #HT1 #H #H0 destruct
- elim (lift_inv_bind1 … H) -H #V1 #T1 #HW01 #HT01 #H destruct
- lapply (cpxs_lift … HV10a (L.ⓓV0) (Ⓕ) … HV12a … HV0a) -V0a [ /2 width=1 by drop_drop/ ] #HV2a
- @(cpxs_trans … (+ⓓV.ⓐV2a.ⓓ{b}V1.T1)) [ /3 width=1 by cpxs_flat_dx, cpxs_bind_dx/ ] -T -V2b -V2c
- @(cpxs_strap2 … (ⓐV1a.ⓓ{b}V0.T0)) [ /4 width=7 by cpx_zeta, lift_bind, lift_flat/ ] -V -V1 -T1
- @(cpxs_strap2 … (ⓓ{b}V0.ⓐV2a.T0)) /3 width=3 by cpxs_pair_sn, cpxs_bind_dx, cpr_cpx, cpr_theta/
- ]
- ]
-]
-qed-.
-
-(* Basic_1: was just: pr3_iso_appls_cast *)
-lemma cpxs_fwd_cast_vector: ∀h,o,G,L,Vs,W,T,U. ⦃G, L⦄ ⊢ ⒶVs.ⓝW.T ➡*[h, o] U →
- ∨∨ ⒶVs. ⓝW. T ≂ U
- | ⦃G, L⦄ ⊢ ⒶVs.T ➡*[h, o] U
- | ⦃G, L⦄ ⊢ ⒶVs.W ➡*[h, o] U.
-#h #o #G #L #Vs elim Vs -Vs /2 width=1 by cpxs_fwd_cast/
-#V #Vs #IHVs #W #T #U #H
-elim (cpxs_inv_appl1 … H) -H *
-[ -IHVs #V0 #T0 #_ #_ #H destruct /2 width=1 by tsts_pair, or3_intro0/
-| #b #W0 #T0 #HT0 #HU elim (IHVs … HT0) -IHVs -HT0 #HT0
- [ elim (tsts_inv_bind_applv_simple … HT0) //
- | @or3_intro1 -W (**) (* explicit constructor *)
- @(cpxs_trans … HU) -U
- @(cpxs_strap1 … (ⓐV.ⓛ{b}W0.T0)) /2 width=1 by cpxs_flat_dx, cpx_beta/
- | @or3_intro2 -T (**) (* explicit constructor *)
- @(cpxs_trans … HU) -U
- @(cpxs_strap1 … (ⓐV.ⓛ{b}W0.T0)) /2 width=1 by cpxs_flat_dx, cpx_beta/
- ]
-| #b #V0 #V1 #V2 #T0 #HV0 #HV01 #HT0 #HU
- elim (IHVs … HT0) -IHVs -HT0 #HT0
- [ elim (tsts_inv_bind_applv_simple … HT0) //
- | @or3_intro1 -W (**) (* explicit constructor *)
- @(cpxs_trans … HU) -U
- @(cpxs_strap1 … (ⓐV0.ⓓ{b}V2.T0)) /2 width=3 by cpxs_flat, cpx_theta/
- | @or3_intro2 -T (**) (* explicit constructor *)
- @(cpxs_trans … HU) -U
- @(cpxs_strap1 … (ⓐV0.ⓓ{b}V2.T0)) /2 width=3 by cpxs_flat, cpx_theta/
- ]
-]
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/sn_5.ma".
-include "basic_2/reduction/cnx.ma".
-
-(* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERMS ********************)
-
-definition csx: ∀h. sd h → relation3 genv lenv term ≝
- λh,o,G,L. SN … (cpx h o G L) (eq …).
-
-interpretation
- "context-sensitive extended strong normalization (term)"
- 'SN h o G L T = (csx h o G L T).
-
-(* Basic eliminators ********************************************************)
-
-lemma csx_ind: ∀h,o,G,L. ∀R:predicate term.
- (∀T1. ⦃G, L⦄ ⊢ ⬊*[h, o] T1 →
- (∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → (T1 = T2 → ⊥) → R T2) →
- R T1
- ) →
- ∀T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → R T.
-#h #o #G #L #R #H0 #T1 #H elim H -T1
-/5 width=1 by SN_intro/
-qed-.
-
-(* Basic properties *********************************************************)
-
-(* Basic_1: was just: sn3_pr2_intro *)
-lemma csx_intro: ∀h,o,G,L,T1.
- (∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → (T1 = T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊*[h, o] T2) →
- ⦃G, L⦄ ⊢ ⬊*[h, o] T1.
-/4 width=1 by SN_intro/ qed.
-
-lemma csx_cpx_trans: ∀h,o,G,L,T1. ⦃G, L⦄ ⊢ ⬊*[h, o] T1 →
- ∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → ⦃G, L⦄ ⊢ ⬊*[h, o] T2.
-#h #o #G #L #T1 #H @(csx_ind … H) -T1 #T1 #HT1 #IHT1 #T2 #HLT12
-elim (eq_term_dec T1 T2) #HT12 destruct /3 width=4 by/
-qed-.
-
-(* Basic_1: was just: sn3_nf2 *)
-lemma cnx_csx: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ → ⦃G, L⦄ ⊢ ⬊*[h, o] T.
-/2 width=1 by NF_to_SN/ qed.
-
-lemma csx_sort: ∀h,o,G,L,s. ⦃G, L⦄ ⊢ ⬊*[h, o] ⋆s.
-#h #o #G #L #s elim (deg_total h o s)
-#d generalize in match s; -s @(nat_ind_plus … d) -d /3 width=6 by cnx_csx, cnx_sort/
-#d #IHd #s #Hkd lapply (deg_next_SO … Hkd) -Hkd
-#Hkd @csx_intro #X #H #HX elim (cpx_inv_sort1 … H) -H
-[ #H destruct elim HX //
-| -HX * #d0 #_ #H destruct -d0 /2 width=1 by/
-]
-qed.
-
-(* Basic_1: was just: sn3_cast *)
-lemma csx_cast: ∀h,o,G,L,W. ⦃G, L⦄ ⊢ ⬊*[h, o] W →
- ∀T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓝW.T.
-#h #o #G #L #W #HW @(csx_ind … HW) -W #W #HW #IHW #T #HT @(csx_ind … HT) -T #T #HT #IHT
-@csx_intro #X #H1 #H2
-elim (cpx_inv_cast1 … H1) -H1
-[ * #W0 #T0 #HLW0 #HLT0 #H destruct
- elim (eq_false_inv_tpair_sn … H2) -H2
- [ /3 width=3 by csx_cpx_trans/
- | -HLW0 * #H destruct /3 width=1 by/
- ]
-|2,3: /3 width=3 by csx_cpx_trans/
-]
-qed.
-
-(* Basic forward lemmas *****************************************************)
-
-fact csx_fwd_pair_sn_aux: ∀h,o,G,L,U. ⦃G, L⦄ ⊢ ⬊*[h, o] U →
- ∀I,V,T. U = ②{I}V.T → ⦃G, L⦄ ⊢ ⬊*[h, o] V.
-#h #o #G #L #U #H elim H -H #U0 #_ #IH #I #V #T #H destruct
-@csx_intro #V2 #HLV2 #HV2
-@(IH (②{I}V2.T)) -IH /2 width=3 by cpx_pair_sn/ -HLV2
-#H destruct /2 width=1 by/
-qed-.
-
-(* Basic_1: was just: sn3_gen_head *)
-lemma csx_fwd_pair_sn: ∀h,o,I,G,L,V,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ②{I}V.T → ⦃G, L⦄ ⊢ ⬊*[h, o] V.
-/2 width=5 by csx_fwd_pair_sn_aux/ qed-.
-
-fact csx_fwd_bind_dx_aux: ∀h,o,G,L,U. ⦃G, L⦄ ⊢ ⬊*[h, o] U →
- ∀a,I,V,T. U = ⓑ{a,I}V.T → ⦃G, L.ⓑ{I}V⦄ ⊢ ⬊*[h, o] T.
-#h #o #G #L #U #H elim H -H #U0 #_ #IH #a #I #V #T #H destruct
-@csx_intro #T2 #HLT2 #HT2
-@(IH (ⓑ{a,I}V.T2)) -IH /2 width=3 by cpx_bind/ -HLT2
-#H destruct /2 width=1 by/
-qed-.
-
-(* Basic_1: was just: sn3_gen_bind *)
-lemma csx_fwd_bind_dx: ∀h,o,a,I,G,L,V,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ⓑ{a,I}V.T → ⦃G, L.ⓑ{I}V⦄ ⊢ ⬊*[h, o] T.
-/2 width=4 by csx_fwd_bind_dx_aux/ qed-.
-
-fact csx_fwd_flat_dx_aux: ∀h,o,G,L,U. ⦃G, L⦄ ⊢ ⬊*[h, o] U →
- ∀I,V,T. U = ⓕ{I}V.T → ⦃G, L⦄ ⊢ ⬊*[h, o] T.
-#h #o #G #L #U #H elim H -H #U0 #_ #IH #I #V #T #H destruct
-@csx_intro #T2 #HLT2 #HT2
-@(IH (ⓕ{I}V.T2)) -IH /2 width=3 by cpx_flat/ -HLT2
-#H destruct /2 width=1 by/
-qed-.
-
-(* Basic_1: was just: sn3_gen_flat *)
-lemma csx_fwd_flat_dx: ∀h,o,I,G,L,V,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ⓕ{I}V.T → ⦃G, L⦄ ⊢ ⬊*[h, o] T.
-/2 width=5 by csx_fwd_flat_dx_aux/ qed-.
-
-lemma csx_fwd_bind: ∀h,o,a,I,G,L,V,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ⓑ{a,I}V.T →
- ⦃G, L⦄ ⊢ ⬊*[h, o] V ∧ ⦃G, L.ⓑ{I}V⦄ ⊢ ⬊*[h, o] T.
-/3 width=3 by csx_fwd_pair_sn, csx_fwd_bind_dx, conj/ qed-.
-
-lemma csx_fwd_flat: ∀h,o,I,G,L,V,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ⓕ{I}V.T →
- ⦃G, L⦄ ⊢ ⬊*[h, o] V ∧ ⦃G, L⦄ ⊢ ⬊*[h, o] T.
-/3 width=3 by csx_fwd_pair_sn, csx_fwd_flat_dx, conj/ qed-.
-
-(* Basic_1: removed theorems 14:
- sn3_cdelta
- sn3_gen_cflat sn3_cflat sn3_cpr3_trans sn3_shift sn3_change
- sn3_appl_cast sn3_appl_beta sn3_appl_lref sn3_appl_abbr
- sn3_appl_appls sn3_bind sn3_appl_bind sn3_appls_bind
-*)
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/computation/gcp_aaa.ma".
-include "basic_2/computation/cpxs_aaa.ma".
-include "basic_2/computation/csx_tsts_vector.ma".
-
-(* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERMS ********************)
-
-(* Main properties on atomic arity assignment *******************************)
-
-theorem aaa_csx: ∀h,o,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ⦃G, L⦄ ⊢ ⬊*[h, o] T.
-#h #o #G #L #T #A #H
-@(gcr_aaa … (csx_gcp h o) (csx_gcr h o) … H)
-qed.
-
-(* Advanced eliminators *****************************************************)
-
-fact aaa_ind_csx_aux: ∀h,o,G,L,A. ∀R:predicate term.
- (∀T1. ⦃G, L⦄ ⊢ T1 ⁝ A →
- (∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → (T1 = T2 → ⊥) → R T2) → R T1
- ) →
- ∀T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → ⦃G, L⦄ ⊢ T ⁝ A → R T.
-#h #o #G #L #A #R #IH #T #H @(csx_ind … H) -T /4 width=5 by cpx_aaa_conf/
-qed-.
-
-lemma aaa_ind_csx: ∀h,o,G,L,A. ∀R:predicate term.
- (∀T1. ⦃G, L⦄ ⊢ T1 ⁝ A →
- (∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → (T1 = T2 → ⊥) → R T2) → R T1
- ) →
- ∀T. ⦃G, L⦄ ⊢ T ⁝ A → R T.
-/5 width=9 by aaa_ind_csx_aux, aaa_csx/ qed-.
-
-fact aaa_ind_csx_alt_aux: ∀h,o,G,L,A. ∀R:predicate term.
- (∀T1. ⦃G, L⦄ ⊢ T1 ⁝ A →
- (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → (T1 = T2 → ⊥) → R T2) → R T1
- ) →
- ∀T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → ⦃G, L⦄ ⊢ T ⁝ A → R T.
-#h #o #G #L #A #R #IH #T #H @(csx_ind_alt … H) -T /4 width=5 by cpxs_aaa_conf/
-qed-.
-
-lemma aaa_ind_csx_alt: ∀h,o,G,L,A. ∀R:predicate term.
- (∀T1. ⦃G, L⦄ ⊢ T1 ⁝ A →
- (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → (T1 = T2 → ⊥) → R T2) → R T1
- ) →
- ∀T. ⦃G, L⦄ ⊢ T ⁝ A → R T.
-/5 width=9 by aaa_ind_csx_alt_aux, aaa_csx/ qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/snalt_5.ma".
-include "basic_2/computation/cpxs.ma".
-include "basic_2/computation/csx.ma".
-
-(* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERMS ********************)
-
-(* alternative definition of csx *)
-definition csxa: ∀h. sd h → relation3 genv lenv term ≝
- λh,o,G,L. SN … (cpxs h o G L) (eq …).
-
-interpretation
- "context-sensitive extended strong normalization (term) alternative"
- 'SNAlt h o G L T = (csxa h o G L T).
-
-(* Basic eliminators ********************************************************)
-
-lemma csxa_ind: ∀h,o,G,L. ∀R:predicate term.
- (∀T1. ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T1 →
- (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → (T1 = T2 → ⊥) → R T2) → R T1
- ) →
- ∀T. ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T → R T.
-#h #o #G #L #R #H0 #T1 #H elim H -T1 /5 width=1 by SN_intro/
-qed-.
-
-(* Basic properties *********************************************************)
-
-lemma csx_intro_cpxs: ∀h,o,G,L,T1.
- (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → (T1 = T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊*[h, o] T2) →
- ⦃G, L⦄ ⊢ ⬊*[h, o] T1.
-/4 width=1 by cpx_cpxs, csx_intro/ qed.
-
-(* Basic_1: was just: sn3_intro *)
-lemma csxa_intro: ∀h,o,G,L,T1.
- (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → (T1 = T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T2) →
- ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T1.
-/4 width=1 by SN_intro/ qed.
-
-fact csxa_intro_aux: ∀h,o,G,L,T1. (
- ∀T,T2. ⦃G, L⦄ ⊢ T ➡*[h, o] T2 → T1 = T → (T1 = T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T2
- ) → ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T1.
-/4 width=3 by csxa_intro/ qed-.
-
-(* Basic_1: was just: sn3_pr3_trans (old version) *)
-lemma csxa_cpxs_trans: ∀h,o,G,L,T1. ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T1 →
- ∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T2.
-#h #o #G #L #T1 #H elim H -T1 #T1 #HT1 #IHT1 #T2 #HLT12
-@csxa_intro #T #HLT2 #HT2
-elim (eq_term_dec T1 T2) #HT12
-[ -IHT1 -HLT12 destruct /3 width=1 by/
-| -HT1 -HT2 /3 width=4 by/
-qed.
-
-(* Basic_1: was just: sn3_pr2_intro (old version) *)
-lemma csxa_intro_cpx: ∀h,o,G,L,T1. (
- ∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → (T1 = T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T2
- ) → ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T1.
-#h #o #G #L #T1 #H
-@csxa_intro_aux #T #T2 #H @(cpxs_ind_dx … H) -T
-[ -H #H destruct #H
- elim H //
-| #T0 #T #HLT1 #HLT2 #IHT #HT10 #HT12 destruct
- elim (eq_term_dec T0 T) #HT0
- [ -HLT1 -HLT2 -H /3 width=1 by/
- | -IHT -HT12 /4 width=3 by csxa_cpxs_trans/
- ]
-]
-qed.
-
-(* Main properties **********************************************************)
-
-theorem csx_csxa: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T.
-#h #o #G #L #T #H @(csx_ind … H) -T /4 width=1 by csxa_intro_cpx/
-qed.
-
-theorem csxa_csx: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T → ⦃G, L⦄ ⊢ ⬊*[h, o] T.
-#h #o #G #L #T #H @(csxa_ind … H) -T /4 width=1 by cpx_cpxs, csx_intro/
-qed.
-
-(* Basic_1: was just: sn3_pr3_trans *)
-lemma csx_cpxs_trans: ∀h,o,G,L,T1. ⦃G, L⦄ ⊢ ⬊*[h, o] T1 →
- ∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → ⦃G, L⦄ ⊢ ⬊*[h, o] T2.
-#h #o #G #L #T1 #HT1 #T2 #H @(cpxs_ind … H) -T2 /2 width=3 by csx_cpx_trans/
-qed-.
-
-(* Main eliminators *********************************************************)
-
-lemma csx_ind_alt: ∀h,o,G,L. ∀R:predicate term.
- (∀T1. ⦃G, L⦄ ⊢ ⬊*[h, o] T1 →
- (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → (T1 = T2 → ⊥) → R T2) → R T1
- ) →
- ∀T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → R T.
-#h #o #G #L #R #H0 #T1 #H @(csxa_ind … (csx_csxa … H)) -T1 /4 width=1 by csxa_csx/
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/computation/fpbs.ma".
-include "basic_2/computation/csx_lleq.ma".
-include "basic_2/computation/csx_lpx.ma".
-
-(* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERMS ********************)
-
-(* Advanced properties ******************************************************)
-
-lemma csx_fpb_conf: ∀h,o,G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬊*[h, o] T1 →
- ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ → ⦃G2, L2⦄ ⊢ ⬊*[h, o] T2.
-#h #o #G1 #L1 #T1 #HT1 #G2 #L2 #T2 *
-/2 width=5 by csx_cpx_trans, csx_fquq_conf, csx_lpx_conf, csx_lleq_conf/
-qed-.
-
-lemma csx_fpbs_conf: ∀h,o,G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬊*[h, o] T1 →
- ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ → ⦃G2, L2⦄ ⊢ ⬊*[h, o] T2.
-#h #o #G1 #L1 #T1 #HT1 #G2 #L2 #T2 #H @(fpbs_ind … H) -G2 -L2 -T2
-/2 width=5 by csx_fpb_conf/
-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/cnx_lift.ma".
-include "basic_2/computation/gcp.ma".
-include "basic_2/computation/csx.ma".
-
-(* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERMS ********************)
-
-(* Relocation properties ****************************************************)
-
-(* Basic_1: was just: sn3_lift *)
-lemma csx_lift: ∀h,o,G,L2,L1,T1,c,l,k. ⦃G, L1⦄ ⊢ ⬊*[h, o] T1 →
- ∀T2. ⬇[c, l, k] L2 ≡ L1 → ⬆[l, k] T1 ≡ T2 → ⦃G, L2⦄ ⊢ ⬊*[h, o] T2.
-#h #o #G #L2 #L1 #T1 #c #l #k #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
->(lift_mono … HT0 … HT12) in HT2; -T1 /2 width=1 by/
-qed.
-
-(* Basic_1: was just: sn3_gen_lift *)
-lemma csx_inv_lift: ∀h,o,G,L2,L1,T1,c,l,k. ⦃G, L1⦄ ⊢ ⬊*[h, o] T1 →
- ∀T2. ⬇[c, l, k] L1 ≡ L2 → ⬆[l, k] T2 ≡ T1 → ⦃G, L2⦄ ⊢ ⬊*[h, o] T2.
-#h #o #G #L2 #L1 #T1 #c #l #k #H elim H -T1 #T1 #_ #IHT1 #T2 #HL12 #HT21
-@csx_intro #T #HLT2 #HT2
-elim (lift_total T l k) #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/
-qed.
-
-(* Advanced inversion lemmas ************************************************)
-
-(* Basic_1: was: sn3_gen_def *)
-lemma csx_inv_lref_bind: ∀h,o,I,G,L,K,V,i. ⬇[i] L ≡ K.ⓑ{I}V →
- ⦃G, L⦄ ⊢ ⬊*[h, o] #i → ⦃G, K⦄ ⊢ ⬊*[h, o] V.
-#h #o #I #G #L #K #V #i #HLK #Hi
-elim (lift_total V 0 (i+1))
-/4 width=9 by csx_inv_lift, csx_cpx_trans, cpx_delta, drop_fwd_drop2/
-qed-.
-
-(* Advanced properties ******************************************************)
-
-(* Basic_1: was just: sn3_abbr *)
-lemma csx_lref_bind: ∀h,o,I,G,L,K,V,i. ⬇[i] L ≡ K.ⓑ{I}V → ⦃G, K⦄ ⊢ ⬊*[h, o] V → ⦃G, L⦄ ⊢ ⬊*[h, o] #i.
-#h #o #I #G #L #K #V #i #HLK #HV
-@csx_intro #X #H #Hi
-elim (cpx_inv_lref1 … H) -H
-[ #H destruct elim Hi //
-| -Hi * #I0 #K0 #V0 #V1 #HLK0 #HV01 #HV1
- lapply (drop_mono … HLK0 … HLK) -HLK #H destruct
- /3 width=8 by csx_lift, csx_cpx_trans, drop_fwd_drop2/
-]
-qed.
-
-lemma csx_appl_simple: ∀h,o,G,L,V. ⦃G, L⦄ ⊢ ⬊*[h, o] V → ∀T1.
- (∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → (T1 = T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓐV.T2) →
- 𝐒⦃T1⦄ → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓐV.T1.
-#h #o #G #L #V #H @(csx_ind … H) -V #V #_ #IHV #T1 #IHT1 #HT1
-@csx_intro #X #H1 #H2
-elim (cpx_inv_appl1_simple … H1) // -H1
-#V0 #T0 #HLV0 #HLT10 #H destruct
-elim (eq_false_inv_tpair_dx … H2) -H2
-[ -IHV -HT1 /4 width=3 by csx_cpx_trans, cpx_pair_sn/
-| -HLT10 * #H #HV0 destruct
- @IHV /4 width=3 by csx_cpx_trans, cpx_pair_sn/ (**) (* full auto 17s *)
-]
-qed.
-
-lemma csx_fqu_conf: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
- ⦃G1, L1⦄ ⊢ ⬊*[h, o] T1 → ⦃G2, L2⦄ ⊢ ⬊*[h, o] T2.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
-/2 width=8 by csx_inv_lref_bind, csx_inv_lift, csx_fwd_flat_dx, csx_fwd_bind_dx, csx_fwd_pair_sn/
-qed-.
-
-lemma csx_fquq_conf: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ⦃G1, L1⦄ ⊢ ⬊*[h, o] T1 → ⦃G2, L2⦄ ⊢ ⬊*[h, o] T2.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H12 #H elim (fquq_inv_gen … H12) -H12
-[ /2 width=5 by csx_fqu_conf/
-| * #HG #HL #HT destruct //
-]
-qed-.
-
-lemma csx_fqup_conf: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
- ⦃G1, L1⦄ ⊢ ⬊*[h, o] T1 → ⦃G2, L2⦄ ⊢ ⬊*[h, o] T2.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
-/3 width=5 by csx_fqu_conf/
-qed-.
-
-lemma csx_fqus_conf: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
- ⦃G1, L1⦄ ⊢ ⬊*[h, o] T1 → ⦃G2, L2⦄ ⊢ ⬊*[h, o] T2.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H12 #H elim (fqus_inv_gen … H12) -H12
-[ /2 width=5 by csx_fqup_conf/
-| * #HG #HL #HT destruct //
-]
-qed-.
-
-(* Main properties **********************************************************)
-
-theorem csx_gcp: ∀h,o. gcp (cpx h o) (eq …) (csx h o).
-#h #o @mk_gcp
-[ normalize /3 width=13 by cnx_lift/
-| #G #L elim (deg_total h o 0) /3 width=8 by cnx_sort_iter, ex_intro/
-| /2 width=8 by csx_lift/
-| /2 width=3 by csx_fwd_flat_dx/
-]
-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_lleq.ma".
-include "basic_2/computation/csx.ma".
-
-(* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERMS ********************)
-
-(* Properties on lazy equivalence for local environments ********************)
-
-lemma csx_lleq_conf: ∀h,o,G,L1,T. ⦃G, L1⦄ ⊢ ⬊*[h, o] T →
- ∀L2. L1 ≡[T, 0] L2 → ⦃G, L2⦄ ⊢ ⬊*[h, o] T.
-#h #o #G #L1 #T #H @(csx_ind … H) -T
-/4 width=6 by csx_intro, cpx_lleq_conf_dx, lleq_cpx_trans/
-qed-.
-
-lemma csx_lleq_trans: ∀h,o,G,L1,L2,T.
- L1 ≡[T, 0] L2 → ⦃G, L2⦄ ⊢ ⬊*[h, o] T → ⦃G, L1⦄ ⊢ ⬊*[h, o] T.
-/3 width=3 by csx_lleq_conf, 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/grammar/tsts_tsts.ma".
-include "basic_2/computation/cpxs_cpxs.ma".
-include "basic_2/computation/csx_alt.ma".
-include "basic_2/computation/csx_lift.ma".
-
-(* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERMS ********************)
-
-(* Advanced properties ******************************************************)
-
-lemma csx_lpx_conf: ∀h,o,G,L1,L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 →
- ∀T. ⦃G, L1⦄ ⊢ ⬊*[h, o] T → ⦃G, L2⦄ ⊢ ⬊*[h, o] T.
-#h #o #G #L1 #L2 #HL12 #T #H @(csx_ind_alt … H) -T
-/4 width=3 by csx_intro, lpx_cpx_trans/
-qed-.
-
-lemma csx_abst: ∀h,o,a,G,L,W. ⦃G, L⦄ ⊢ ⬊*[h, o] W →
- ∀T. ⦃G, L.ⓛW⦄ ⊢ ⬊*[h, o] T → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓛ{a}W.T.
-#h #o #a #G #L #W #HW @(csx_ind … HW) -W #W #_ #IHW #T #HT @(csx_ind … HT) -T #T #HT #IHT
-@csx_intro #X #H1 #H2
-elim (cpx_inv_abst1 … H1) -H1
-#W0 #T0 #HLW0 #HLT0 #H destruct
-elim (eq_false_inv_tpair_sn … H2) -H2
-[ -IHT #H lapply (csx_cpx_trans … HLT0) // -HT
- #HT0 lapply (csx_lpx_conf … (L.ⓛW0) … HT0) -HT0 /3 width=1 by lpx_pair/
-| -IHW -HLW0 -HT * #H destruct /3 width=1 by/
-]
-qed.
-
-lemma csx_abbr: ∀h,o,a,G,L,V. ⦃G, L⦄ ⊢ ⬊*[h, o] V →
- ∀T. ⦃G, L.ⓓV⦄ ⊢ ⬊*[h, o] T → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓓ{a}V. T.
-#h #o #a #G #L #V #HV elim HV -V #V #_ #IHV #T #HT @(csx_ind_alt … HT) -T #T #HT #IHT
-@csx_intro #X #H1 #H2
-elim (cpx_inv_abbr1 … H1) -H1 *
-[ #V1 #T1 #HLV1 #HLT1 #H destruct
- elim (eq_false_inv_tpair_sn … H2) -H2
- [ /4 width=5 by csx_cpx_trans, csx_lpx_conf, lpx_pair/
- | -IHV -HLV1 * #H destruct /3 width=1 by cpx_cpxs/
- ]
-| -IHV -IHT -H2
- /3 width=8 by csx_cpx_trans, csx_inv_lift, drop_drop/
-]
-qed.
-
-fact csx_appl_beta_aux: ∀h,o,a,G,L,U1. ⦃G, L⦄ ⊢ ⬊*[h, o] U1 →
- ∀V,W,T1. U1 = ⓓ{a}ⓝW.V.T1 → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓐV.ⓛ{a}W.T1.
-#h #o #a #G #L #X #H @(csx_ind … H) -X
-#X #HT1 #IHT1 #V #W #T1 #H1 destruct
-@csx_intro #X #H1 #H2
-elim (cpx_inv_appl1 … H1) -H1 *
-[ -HT1 #V0 #Y #HLV0 #H #H0 destruct
- elim (cpx_inv_abst1 … H) -H #W0 #T0 #HLW0 #HLT0 #H destruct
- @IHT1 -IHT1 [4: // | skip |3: #H destruct /2 width=1 by/ ] -H2
- lapply (lsubr_cpx_trans … HLT0 (L.ⓓⓝW.V) ?) -HLT0 /3 width=1 by cpx_bind, cpx_flat, lsubr_beta/
-| -IHT1 -H2 #b #V0 #W0 #W2 #T0 #T2 #HLV0 #HLW02 #HLT02 #H1 #H3 destruct
- lapply (lsubr_cpx_trans … HLT02 (L.ⓓⓝW0.V) ?) -HLT02
- /4 width=5 by csx_cpx_trans, cpx_bind, cpx_flat, lsubr_beta/
-| -HT1 -IHT1 -H2 #b #V0 #V1 #W0 #W1 #T0 #T3 #_ #_ #_ #_ #H destruct
-]
-qed-.
-
-(* Basic_1: was just: sn3_beta *)
-lemma csx_appl_beta: ∀h,o,a,G,L,V,W,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ⓓ{a}ⓝW.V.T → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓐV.ⓛ{a}W.T.
-/2 width=3 by csx_appl_beta_aux/ qed.
-
-fact csx_appl_theta_aux: ∀h,o,a,G,L,U. ⦃G, L⦄ ⊢ ⬊*[h, o] U → ∀V1,V2. ⬆[0, 1] V1 ≡ V2 →
- ∀V,T. U = ⓓ{a}V.ⓐV2.T → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓐV1.ⓓ{a}V.T.
-#h #o #a #G #L #X #H @(csx_ind_alt … H) -X #X #HVT #IHVT #V1 #V2 #HV12 #V #T #H destruct
-lapply (csx_fwd_pair_sn … HVT) #HV
-lapply (csx_fwd_bind_dx … HVT) -HVT #HVT
-@csx_intro #X #HL #H
-elim (cpx_inv_appl1 … HL) -HL *
-[ -HV #V0 #Y #HLV10 #HL #H0 destruct
- elim (cpx_inv_abbr1 … HL) -HL *
- [ #V3 #T3 #HV3 #HLT3 #H0 destruct
- elim (lift_total V0 0 1) #V4 #HV04
- elim (eq_term_dec (ⓓ{a}V.ⓐV2.T) (ⓓ{a}V3.ⓐV4.T3))
- [ -IHVT #H0 destruct
- elim (eq_false_inv_tpair_sn … H) -H
- [ -HLV10 -HV3 -HLT3 -HVT
- >(lift_inj … HV12 … HV04) -V4
- #H elim H //
- | * #_ #H elim H //
- ]
- | -H -HVT #H
- lapply (cpx_lift … HLV10 (L.ⓓV) (Ⓕ) … HV12 … HV04) -HLV10 -HV12 /2 width=1 by drop_drop/ #HV24
- @(IHVT … H … HV04) -IHVT /4 width=1 by cpx_cpxs, cpx_bind, cpx_flat/
- ]
- | -H -IHVT #T0 #HLT0 #HT0 #H0 destruct
- lapply (csx_cpx_trans … HVT (ⓐV2.T0) ?) /2 width=1 by cpx_flat/ -T #HVT0
- lapply (csx_inv_lift … L … (Ⓕ) … 1 HVT0 ? ? ?) -HVT0
- /3 width=5 by csx_cpx_trans, cpx_pair_sn, drop_drop, lift_flat/
- ]
-| -HV -HV12 -HVT -IHVT -H #b #V0 #W0 #W1 #T0 #T1 #_ #_ #_ #H destruct
-| -IHVT -H #b #V0 #V3 #W0 #W1 #T0 #T1 #HLV10 #HV03 #HLW01 #HLT01 #H1 #H2 destruct
- lapply (cpx_lift … HLV10 (L. ⓓW0) … HV12 … HV03) -HLV10 -HV12 -HV03 /2 width=2 by drop_drop/ #HLV23
- @csx_abbr /2 width=3 by csx_cpx_trans/ -HV
- @(csx_lpx_conf … (L.ⓓW0)) /2 width=1 by lpx_pair/ -W1
- /4 width=5 by csx_cpxs_trans, cpx_cpxs, cpx_flat/
-]
-qed-.
-
-lemma csx_appl_theta: ∀h,o,a,V1,V2. ⬆[0, 1] V1 ≡ V2 →
- ∀G,L,V,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ⓓ{a}V.ⓐV2.T → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓐV1.ⓓ{a}V.T.
-/2 width=5 by csx_appl_theta_aux/ qed.
-
-(* Basic_1: was just: sn3_appl_appl *)
-lemma csx_appl_simple_tsts: ∀h,o,G,L,V. ⦃G, L⦄ ⊢ ⬊*[h, o] V → ∀T1. ⦃G, L⦄ ⊢ ⬊*[h, o] T1 →
- (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → (T1 ≂ T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓐV.T2) →
- 𝐒⦃T1⦄ → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓐV.T1.
-#h #o #G #L #V #H @(csx_ind … H) -V #V #_ #IHV #T1 #H @(csx_ind … H) -T1 #T1 #H1T1 #IHT1 #H2T1 #H3T1
-@csx_intro #X #HL #H
-elim (cpx_inv_appl1_simple … HL) -HL //
-#V0 #T0 #HLV0 #HLT10 #H0 destruct
-elim (eq_false_inv_tpair_sn … H) -H
-[ -IHT1 #HV0
- @(csx_cpx_trans … (ⓐV0.T1)) /2 width=1 by cpx_flat/ -HLT10
- @IHV -IHV /4 width=3 by csx_cpx_trans, cpx_pair_sn/
-| -IHV -H1T1 -HLV0 * #H #H1T10 destruct
- elim (tsts_dec T1 T0) #H2T10
- [ @IHT1 -IHT1 /4 width=3 by cpxs_strap2, cpxs_strap1, tsts_canc_sn, simple_tsts_repl_dx/
- | -IHT1 -H3T1 -H1T10 /3 width=1 by cpx_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/computation/csx_lpx.ma".
-include "basic_2/computation/lpxs.ma".
-
-(* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERMS ********************)
-
-(* Properties on sn extended parallel computation for local environments ****)
-
-lemma csx_lpxs_conf: ∀h,o,G,L1,L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 →
- ∀T. ⦃G, L1⦄ ⊢ ⬊*[h, o] T → ⦃G, L2⦄ ⊢ ⬊*[h, o] T.
-#h #o #G #L1 #L2 #H @(lpxs_ind … H) -L2 /3 by lpxs_strap1, csx_lpx_conf/
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/computation/gcp_cr.ma".
-include "basic_2/computation/cpxs_tsts_vector.ma".
-include "basic_2/computation/csx_lpx.ma".
-include "basic_2/computation/csx_vector.ma".
-
-(* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERM VECTORS *************)
-
-(* Advanced properties ******************************************************)
-
-(* Basic_1: was just: sn3_appls_lref *)
-lemma csx_applv_cnx: ∀h,o,G,L,T. 𝐒⦃T⦄ → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ →
- ∀Vs. ⦃G, L⦄ ⊢ ⬊*[h, o] Vs → ⦃G, L⦄ ⊢ ⬊*[h, o] ⒶVs.T.
-#h #o #G #L #T #H1T #H2T #Vs elim Vs -Vs [ #_ @(cnx_csx … H2T) ] (**) (* /2 width=1/ does not work *)
-#V #Vs #IHV #H
-elim (csxv_inv_cons … H) -H #HV #HVs
-@csx_appl_simple_tsts /2 width=1 by applv_simple/ -IHV -HV -HVs
-#X #H #H0
-lapply (cpxs_fwd_cnx_vector … H) -H // -H1T -H2T #H
-elim (H0) -H0 //
-qed.
-
-lemma csx_applv_sort: ∀h,o,G,L,s,Vs. ⦃G, L⦄ ⊢ ⬊*[h, o] Vs → ⦃G, L⦄ ⊢ ⬊*[h, o] ⒶVs.⋆s.
-#h #o #G #L #s elim (deg_total h o s)
-#d generalize in match s; -s @(nat_ind_plus … d) -d [ /3 width=6 by csx_applv_cnx, cnx_sort, simple_atom/ ]
-#d #IHd #s #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
-#X #H #H0
-elim (cpxs_fwd_sort_vector … H) -H #H
-[ elim H0 -H0 //
-| -H0 @(csx_cpxs_trans … (Ⓐ(V@Vs).⋆(next h s))) /2 width=1 by cpxs_flat_dx/
-]
-qed.
-
-(* Basic_1: was just: sn3_appls_beta *)
-lemma csx_applv_beta: ∀h,o,a,G,L,Vs,V,W,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ⒶVs.ⓓ{a}ⓝW.V.T →
- ⦃G, L⦄ ⊢ ⬊*[h, o] ⒶVs. ⓐV.ⓛ{a}W.T.
-#h #o #a #G #L #Vs elim Vs -Vs /2 width=1 by csx_appl_beta/
-#V0 #Vs #IHV #V #W #T #H1T
-lapply (csx_fwd_pair_sn … H1T) #HV0
-lapply (csx_fwd_flat_dx … H1T) #H2T
-@csx_appl_simple_tsts /2 width=1 by applv_simple, simple_flat/ -IHV -HV0 -H2T
-#X #H #H0
-elim (cpxs_fwd_beta_vector … H) -H #H
-[ -H1T elim H0 -H0 //
-| -H0 /3 width=5 by csx_cpxs_trans, cpxs_flat_dx/
-]
-qed.
-
-lemma csx_applv_delta: ∀h,o,I,G,L,K,V1,i. ⬇[i] L ≡ K.ⓑ{I}V1 →
- ∀V2. ⬆[0, i + 1] V1 ≡ V2 →
- ∀Vs. ⦃G, L⦄ ⊢ ⬊*[h, o] (ⒶVs.V2) → ⦃G, L⦄ ⊢ ⬊*[h, o] (ⒶVs.#i).
-#h #o #I #G #L #K #V1 #i #HLK #V2 #HV12 #Vs elim Vs -Vs
-[ /4 width=12 by csx_inv_lift, csx_lref_bind, drop_fwd_drop2/
-| #V #Vs #IHV #H1T
- lapply (csx_fwd_pair_sn … H1T) #HV
- lapply (csx_fwd_flat_dx … H1T) #H2T
- @csx_appl_simple_tsts /2 width=1 by applv_simple, simple_atom/ -IHV -HV -H2T
- #X #H #H0
- elim (cpxs_fwd_delta_vector … HLK … HV12 … H) -HLK -HV12 -H #H
- [ -H1T elim H0 -H0 //
- | -H0 /3 width=5 by csx_cpxs_trans, cpxs_flat_dx/
- ]
-]
-qed.
-
-(* Basic_1: was just: sn3_appls_abbr *)
-lemma csx_applv_theta: ∀h,o,a,G,L,V1c,V2c. ⬆[0, 1] V1c ≡ V2c →
- ∀V,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ⓓ{a}V.ⒶV2c.T →
- ⦃G, L⦄ ⊢ ⬊*[h, o] ⒶV1c.ⓓ{a}V.T.
-#h #o #a #G #L #V1c #V2c * -V1c -V2c /2 width=1 by/
-#V1c #V2c #V1 #V2 #HV12 #H
-generalize in match HV12; -HV12 generalize in match V2; -V2 generalize in match V1; -V1
-elim H -V1c -V2c /2 width=3 by csx_appl_theta/
-#V1c #V2c #V1 #V2 #HV12 #HV12c #IHV12c #W1 #W2 #HW12 #V #T #H
-lapply (csx_appl_theta … HW12 … H) -H -HW12 #H
-lapply (csx_fwd_pair_sn … H) #HW1
-lapply (csx_fwd_flat_dx … H) #H1
-@csx_appl_simple_tsts /2 width=3 by simple_flat/ -IHV12c -HW1 -H1 #X #H1 #H2
-elim (cpxs_fwd_theta_vector … (V2@V2c) … H1) -H1 /2 width=1 by liftv_cons/ -HV12c -HV12
-[ -H #H elim H2 -H2 //
-| -H2 /3 width=5 by csx_cpxs_trans, cpxs_flat_dx/
-]
-qed.
-
-(* Basic_1: was just: sn3_appls_cast *)
-lemma csx_applv_cast: ∀h,o,G,L,Vs,W,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ⒶVs.W → ⦃G, L⦄ ⊢ ⬊*[h, o] ⒶVs.T →
- ⦃G, L⦄ ⊢ ⬊*[h, o] ⒶVs.ⓝW.T.
-#h #o #G #L #Vs elim Vs -Vs /2 width=1 by csx_cast/
-#V #Vs #IHV #W #T #H1W #H1T
-lapply (csx_fwd_pair_sn … H1W) #HV
-lapply (csx_fwd_flat_dx … H1W) #H2W
-lapply (csx_fwd_flat_dx … H1T) #H2T
-@csx_appl_simple_tsts /2 width=1 by applv_simple, simple_flat/ -IHV -HV -H2W -H2T
-#X #H #H0
-elim (cpxs_fwd_cast_vector … H) -H #H
-[ -H1W -H1T elim H0 -H0 //
-| -H1W -H0 /3 width=5 by csx_cpxs_trans, cpxs_flat_dx/
-| -H1T -H0 /3 width=5 by csx_cpxs_trans, cpxs_flat_dx/
-]
-qed.
-
-theorem csx_gcr: ∀h,o. gcr (cpx h o) (eq …) (csx h o) (csx h o).
-#h #o @mk_gcr //
-[ /3 width=1 by csx_applv_cnx/
-|2,3,6: /2 width=1 by csx_applv_beta, csx_applv_sort, csx_applv_cast/
-| /2 width=7 by csx_applv_delta/
-| #G #L #V1c #V2c #HV12c #a #V #T #H #HV
- @(csx_applv_theta … HV12c) -HV12c
- @csx_abbr //
-]
-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/term_vector.ma".
-include "basic_2/computation/csx.ma".
-
-(* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERM VECTORS *************)
-
-definition csxv: ∀h. sd h → relation3 genv lenv (list term) ≝
- λh,o,G,L. all … (csx h o G L).
-
-interpretation
- "context-sensitive strong normalization (term vector)"
- 'SN h o G L Ts = (csxv h o G L Ts).
-
-(* Basic inversion lemmas ***************************************************)
-
-lemma csxv_inv_cons: ∀h,o,G,L,T,Ts. ⦃G, L⦄ ⊢ ⬊*[h, o] T @ Ts →
- ⦃G, L⦄ ⊢ ⬊*[h, o] T ∧ ⦃G, L⦄ ⊢ ⬊*[h, o] Ts.
-normalize // qed-.
-
-(* Basic forward lemmas *****************************************************)
-
-lemma csx_fwd_applv: ∀h,o,G,L,T,Vs. ⦃G, L⦄ ⊢ ⬊*[h, o] Ⓐ Vs.T →
- ⦃G, L⦄ ⊢ ⬊*[h, o] Vs ∧ ⦃G, L⦄ ⊢ ⬊*[h, o] T.
-#h #o #G #L #T #Vs elim Vs -Vs /2 width=1 by conj/
-#V #Vs #IHVs #HVs
-lapply (csx_fwd_pair_sn … HVs) #HV
-lapply (csx_fwd_flat_dx … HVs) -HVs #HVs
-elim (IHVs HVs) -IHVs -HVs /3 width=1 by conj/
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/lazybtpredstarproper_8.ma".
-include "basic_2/reduction/fpb.ma".
-include "basic_2/computation/fpbs.ma".
-
-(* "QRST" PROPER PARALLEL COMPUTATION FOR CLOSURES **************************)
-
-definition fpbg: ∀h. sd h → tri_relation genv lenv term ≝
- λh,o,G1,L1,T1,G2,L2,T2.
- ∃∃G,L,T. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G, L, T⦄ & ⦃G, L, T⦄ ≥[h, o] ⦃G2, L2, T2⦄.
-
-interpretation "'qrst' proper parallel computation (closure)"
- 'LazyBTPRedStarProper h o G1 L1 T1 G2 L2 T2 = (fpbg h o G1 L1 T1 G2 L2 T2).
-
-(* Basic properties *********************************************************)
-
-lemma fpb_fpbg: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄.
-/2 width=5 by ex2_3_intro/ qed.
-
-lemma fpbg_fpbq_trans: ∀h,o,G1,G,G2,L1,L,L2,T1,T,T2.
- ⦃G1, L1, T1⦄ >≡[h, o] ⦃G, L, T⦄ → ⦃G, L, T⦄ ≽[h, o] ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄.
-#h #o #G1 #G #G2 #L1 #L #L2 #T1 #T #T2 *
-/3 width=9 by fpbs_strap1, 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 "basic_2/multiple/fleq_fleq.ma".
-include "basic_2/reduction/fpbq_alt.ma".
-include "basic_2/computation/fpbg.ma".
-
-(* "QRST" PROPER PARALLEL COMPUTATION FOR CLOSURES **************************)
-
-(* Properties on lazy equivalence for closures ******************************)
-
-lemma fpbg_fleq_trans: ∀h,o,G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ >≡[h, o] ⦃G, L, T⦄ →
- ∀G2,L2,T2. ⦃G, L, T⦄ ≡[0] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄.
-/3 width=5 by fpbg_fpbq_trans, fleq_fpbq/ qed-.
-
-lemma fleq_fpbg_trans: ∀h,o,G,G2,L,L2,T,T2. ⦃G, L, T⦄ >≡[h, o] ⦃G2, L2, T2⦄ →
- ∀G1,L1,T1. ⦃G1, L1, T1⦄ ≡[0] ⦃G, L, T⦄ → ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄.
-#h #o #G #G2 #L #L2 #T #T2 * #G0 #L0 #T0 #H0 #H02 #G1 #L1 #T1 #H1
-elim (fleq_fpb_trans … H1 … H0) -G -L -T
-/4 width=9 by fpbs_strap2, fleq_fpbq, ex2_3_intro/
-qed-.
-
-(* alternative definition of fpbs *******************************************)
-
-lemma fleq_fpbs: ∀h,o,G1,G2,L1,L2,T1,T2.
- ⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 * /2 width=1 by lleq_fpbs/
-qed.
-
-lemma fpbg_fwd_fpbs: ∀h,o,G1,G2,L1,L2,T1,T2.
- ⦃G1, L1, T1⦄ >≡[h,o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 *
-/3 width=5 by fpbs_strap2, fpb_fpbq/
-qed-.
-
-lemma fpbs_fpbg: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ ∨
- ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fpbs_ind … H) -G2 -L2 -T2
-[ /2 width=1 by or_introl/
-| #G #G2 #L #L2 #T #T2 #_ #H2 * #H1 @(fpbq_ind_alt … H2) -H2 #H2
- [ /3 width=5 by fleq_trans, or_introl/
- | elim (fleq_fpb_trans … H1 … H2) -G -L -T
- /4 width=5 by ex2_3_intro, or_intror, fleq_fpbs/
- | /3 width=5 by fpbg_fleq_trans, or_intror/
- | /4 width=5 by fpbg_fpbq_trans, fpb_fpbq, or_intror/
- ]
-]
-qed-.
-
-(* Advanced properties of "qrst" parallel computation on closures ***********)
-
-lemma fpbs_fpb_trans: ∀h,o,F1,F2,K1,K2,T1,T2. ⦃F1, K1, T1⦄ ≥[h, o] ⦃F2, K2, T2⦄ →
- ∀G2,L2,U2. ⦃F2, K2, T2⦄ ≻[h, o] ⦃G2, L2, U2⦄ →
- ∃∃G1,L1,U1. ⦃F1, K1, T1⦄ ≻[h, o] ⦃G1, L1, U1⦄ & ⦃G1, L1, U1⦄ ≥[h, o] ⦃G2, L2, U2⦄.
-#h #o #F1 #F2 #K1 #K2 #T1 #T2 #H elim (fpbs_fpbg … H) -H
-[ #H12 #G2 #L2 #U2 #H2 elim (fleq_fpb_trans … H12 … H2) -F2 -K2 -T2
- /3 width=5 by fleq_fpbs, ex2_3_intro/
-| * #H1 #H2 #H3 #H4 #H5 #H6 #H7 #H8 #H9
- @(ex2_3_intro … H4) -H4 /3 width=5 by fpbs_strap1, fpb_fpbq/
-]
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/computation/fpbg_fpbs.ma".
-
-(* "QRST" PROPER PARALLEL COMPUTATION FOR CLOSURES **************************)
-
-(* Main properties **********************************************************)
-
-theorem fpbg_trans: ∀h,o. tri_transitive … (fpbg h o).
-/3 width=5 by fpbg_fpbs_trans, fpbg_fwd_fpbs/ qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/computation/lpxs_lleq.ma".
-include "basic_2/computation/fpbs_lift.ma".
-include "basic_2/computation/fpbg_fleq.ma".
-
-(* "QRST" PROPER PARALLEL COMPUTATION FOR CLOSURES **************************)
-
-(* Properties on "qrst" parallel reduction on closures **********************)
-
-lemma fpb_fpbg_trans: ∀h,o,G1,G,G2,L1,L,L2,T1,T,T2.
- ⦃G1, L1, T1⦄ ≻[h, o] ⦃G, L, T⦄ → ⦃G, L, T⦄ >≡[h, o] ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄.
-/3 width=5 by fpbg_fwd_fpbs, ex2_3_intro/ qed-.
-
-lemma fpbq_fpbg_trans: ∀h,o,G1,G,G2,L1,L,L2,T1,T,T2.
- ⦃G1, L1, T1⦄ ≽[h, o] ⦃G, L, T⦄ → ⦃G, L, T⦄ >≡[h, o] ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄.
-#h #o #G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 #H2 @(fpbq_ind_alt … H1) -H1
-/2 width=5 by fleq_fpbg_trans, fpb_fpbg_trans/
-qed-.
-
-(* Properties on "qrst" parallel compuutation on closures *******************)
-
-lemma fpbs_fpbg_trans: ∀h,o,G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L, T⦄ →
- ∀G2,L2,T2. ⦃G, L, T⦄ >≡[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄.
-#h #o #G1 #G #L1 #L #T1 #T #H @(fpbs_ind … H) -G -L -T /3 width=5 by fpbq_fpbg_trans/
-qed-.
-
-(* Note: this is used in the closure proof *)
-lemma fpbg_fpbs_trans: ∀h,o,G,G2,L,L2,T,T2. ⦃G, L, T⦄ ≥[h, o] ⦃G2, L2, T2⦄ →
- ∀G1,L1,T1. ⦃G1, L1, T1⦄ >≡[h, o] ⦃G, L, T⦄ → ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄.
-#h #o #G #G2 #L #L2 #T #T2 #H @(fpbs_ind_dx … H) -G -L -T /3 width=5 by fpbg_fpbq_trans/
-qed-.
-
-(* Note: this is used in the closure proof *)
-lemma fqup_fpbg: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim (fqup_inv_step_sn … H) -H
-/3 width=5 by fqus_fpbs, fpb_fqu, ex2_3_intro/
-qed.
-
-lemma cpxs_fpbg: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 →
- (T1 = T2 → ⊥) → ⦃G, L, T1⦄ >≡[h, o] ⦃G, L, T2⦄.
-#h #o #G #L #T1 #T2 #H #H0 elim (cpxs_neq_inv_step_sn … H … H0) -H -H0
-/4 width=5 by cpxs_fpbs, fpb_cpx, ex2_3_intro/
-qed.
-
-lemma lstas_fpbg: ∀h,o,G,L,T1,T2,d2. ⦃G, L⦄ ⊢ T1 •*[h, d2] T2 → (T1 = T2 → ⊥) →
- ∀d1. d2 ≤ d1 → ⦃G, L⦄ ⊢ T1 ▪[h, o] d1 → ⦃G, L, T1⦄ >≡[h, o] ⦃G, L, T2⦄.
-/3 width=5 by lstas_cpxs, cpxs_fpbg/ qed.
-
-lemma lpxs_fpbg: ∀h,o,G,L1,L2,T. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 →
- (L1 ≡[T, 0] L2 → ⊥) → ⦃G, L1, T⦄ >≡[h, o] ⦃G, L2, T⦄.
-#h #o #G #L1 #L2 #T #H #H0 elim (lpxs_nlleq_inv_step_sn … H … H0) -H -H0
-/4 width=5 by fpb_lpx, lpxs_lleq_fpbs, 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 "basic_2/reduction/fpb_lift.ma".
-include "basic_2/computation/fpbg.ma".
-
-(* "QRST" PARALLEL COMPUTATION FOR CLOSURES *********************************)
-
-(* Advanced properties ******************************************************)
-
-lemma sta_fpbg: ∀h,o,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 ▪[h, o] d+1 →
- ⦃G, L⦄ ⊢ T1 •*[h, 1] T2 → ⦃G, L, T1⦄ >≡[h, o] ⦃G, L, T2⦄.
-/4 width=2 by fpb_fpbg, sta_fpb/ qed.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/btpredstar_8.ma".
-include "basic_2/multiple/fqus.ma".
-include "basic_2/reduction/fpbq.ma".
-include "basic_2/computation/cpxs.ma".
-include "basic_2/computation/lpxs.ma".
-
-(* "QRST" PARALLEL COMPUTATION FOR CLOSURES *********************************)
-
-definition fpbs: ∀h. sd h → tri_relation genv lenv term ≝
- λh,o. tri_TC … (fpbq h o).
-
-interpretation "'qrst' parallel computation (closure)"
- 'BTPRedStar h o G1 L1 T1 G2 L2 T2 = (fpbs h o G1 L1 T1 G2 L2 T2).
-
-(* Basic eliminators ********************************************************)
-
-lemma fpbs_ind: ∀h,o,G1,L1,T1. ∀R:relation3 genv lenv term. R G1 L1 T1 →
- (∀G,G2,L,L2,T,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L, T⦄ → ⦃G, L, T⦄ ≽[h, o] ⦃G2, L2, T2⦄ → R G L T → R G2 L2 T2) →
- ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ → R G2 L2 T2.
-/3 width=8 by tri_TC_star_ind/ qed-.
-
-lemma fpbs_ind_dx: ∀h,o,G2,L2,T2. ∀R:relation3 genv lenv term. R G2 L2 T2 →
- (∀G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ ≽[h, o] ⦃G, L, T⦄ → ⦃G, L, T⦄ ≥[h, o] ⦃G2, L2, T2⦄ → R G L T → R G1 L1 T1) →
- ∀G1,L1,T1. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ → R G1 L1 T1.
-/3 width=8 by tri_TC_star_ind_dx/ qed-.
-
-(* Basic properties *********************************************************)
-
-lemma fpbs_refl: ∀h,o. tri_reflexive … (fpbs h o).
-/2 width=1 by tri_inj/ qed.
-
-lemma fpbq_fpbs: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
-/2 width=1 by tri_inj/ qed.
-
-lemma fpbs_strap1: ∀h,o,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L, T⦄ →
- ⦃G, L, T⦄ ≽[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
-/2 width=5 by tri_step/ qed-.
-
-lemma fpbs_strap2: ∀h,o,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ≽[h, o] ⦃G, L, T⦄ →
- ⦃G, L, T⦄ ≥[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
-/2 width=5 by tri_TC_strap/ qed-.
-
-lemma fqup_fpbs: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
-/4 width=5 by fqu_fquq, fpbq_fquq, tri_step/
-qed.
-
-lemma fqus_fpbs: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqus_ind … H) -G2 -L2 -T2
-/3 width=5 by fpbq_fquq, tri_step/
-qed.
-
-lemma cpxs_fpbs: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → ⦃G, L, T1⦄ ≥[h, o] ⦃G, L, T2⦄.
-#h #o #G #L #T1 #T2 #H @(cpxs_ind … H) -T2
-/3 width=5 by fpbq_cpx, fpbs_strap1/
-qed.
-
-lemma lpxs_fpbs: ∀h,o,G,L1,L2,T. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → ⦃G, L1, T⦄ ≥[h, o] ⦃G, L2, T⦄.
-#h #o #G #L1 #L2 #T #H @(lpxs_ind … H) -L2
-/3 width=5 by fpbq_lpx, fpbs_strap1/
-qed.
-
-lemma lleq_fpbs: ∀h,o,G,L1,L2,T. L1 ≡[T, 0] L2 → ⦃G, L1, T⦄ ≥[h, o] ⦃G, L2, T⦄.
-/3 width=1 by fpbq_fpbs, fpbq_lleq/ qed.
-
-lemma cprs_fpbs: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡* T2 → ⦃G, L, T1⦄ ≥[h, o] ⦃G, L, T2⦄.
-/3 width=1 by cprs_cpxs, cpxs_fpbs/ qed.
-
-lemma lprs_fpbs: ∀h,o,G,L1,L2,T. ⦃G, L1⦄ ⊢ ➡* L2 → ⦃G, L1, T⦄ ≥[h, o] ⦃G, L2, T⦄.
-/3 width=1 by lprs_lpxs, lpxs_fpbs/ qed.
-
-lemma fpbs_fqus_trans: ∀h,o,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L, T⦄ →
- ⦃G, L, T⦄ ⊐* ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
-#h #o #G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 #H @(fqus_ind … H) -G2 -L2 -T2
-/3 width=5 by fpbs_strap1, fpbq_fquq/
-qed-.
-
-lemma fpbs_fqup_trans: ∀h,o,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L, T⦄ →
- ⦃G, L, T⦄ ⊐+ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
-/3 width=5 by fpbs_fqus_trans, fqup_fqus/ qed-.
-
-lemma fpbs_cpxs_trans: ∀h,o,G1,G,L1,L,T1,T,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L, T⦄ →
- ⦃G, L⦄ ⊢ T ➡*[h, o] T2 → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L, T2⦄.
-#h #o #G1 #G #L1 #L #T1 #T #T2 #H1 #H @(cpxs_ind … H) -T2
-/3 width=5 by fpbs_strap1, fpbq_cpx/
-qed-.
-
-lemma fpbs_lpxs_trans: ∀h,o,G1,G,L1,L,L2,T1,T. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L, T⦄ →
- ⦃G, L⦄ ⊢ ➡*[h, o] L2 → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L2, T⦄.
-#h #o #G1 #G #L1 #L #L2 #T1 #T #H1 #H @(lpxs_ind … H) -L2
-/3 width=5 by fpbs_strap1, fpbq_lpx/
-qed-.
-
-lemma fpbs_lleq_trans: ∀h,o,G1,G,L1,L,L2,T1,T. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L, T⦄ →
- L ≡[T, 0] L2 → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L2, T⦄.
-/3 width=5 by fpbs_strap1, fpbq_lleq/ qed-.
-
-lemma fqus_fpbs_trans: ∀h,o,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G, L, T⦄ ≥[h, o] ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ⊐* ⦃G, L, T⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
-#h #o #G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 #H @(fqus_ind_dx … H) -G1 -L1 -T1
-/3 width=5 by fpbs_strap2, fpbq_fquq/
-qed-.
-
-lemma cpxs_fpbs_trans: ∀h,o,G1,G2,L1,L2,T1,T,T2. ⦃G1, L1, T⦄ ≥[h, o] ⦃G2, L2, T2⦄ →
- ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] T → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T #T2 #H1 #H @(cpxs_ind_dx … H) -T1
-/3 width=5 by fpbs_strap2, fpbq_cpx/
-qed-.
-
-lemma lpxs_fpbs_trans: ∀h,o,G1,G2,L1,L,L2,T1,T2. ⦃G1, L, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ →
- ⦃G1, L1⦄ ⊢ ➡*[h, o] L → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
-#h #o #G1 #G2 #L1 #L #L2 #T1 #T2 #H1 #H @(lpxs_ind_dx … H) -L1
-/3 width=5 by fpbs_strap2, fpbq_lpx/
-qed-.
-
-lemma lleq_fpbs_trans: ∀h,o,G1,G2,L1,L,L2,T1,T2. ⦃G1, L, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ →
- L1 ≡[T1, 0] L → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
-/3 width=5 by fpbs_strap2, fpbq_lleq/ qed-.
-
-lemma cpxs_fqus_fpbs: ∀h,o,G1,G2,L1,L2,T1,T,T2. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] T →
- ⦃G1, L1, T⦄ ⊐* ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
-/3 width=5 by fpbs_fqus_trans, cpxs_fpbs/ qed.
-
-lemma cpxs_fqup_fpbs: ∀h,o,G1,G2,L1,L2,T1,T,T2. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] T →
- ⦃G1, L1, T⦄ ⊐+ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
-/3 width=5 by fpbs_fqup_trans, cpxs_fpbs/ qed.
-
-lemma fqus_lpxs_fpbs: ∀h,o,G1,G2,L1,L,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L, T2⦄ →
- ⦃G2, L⦄ ⊢ ➡*[h, o] L2 → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
-/3 width=3 by fpbs_lpxs_trans, fqus_fpbs/ qed.
-
-lemma cpxs_fqus_lpxs_fpbs: ∀h,o,G1,G2,L1,L,L2,T1,T,T2. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] T →
- ⦃G1, L1, T⦄ ⊐* ⦃G2, L, T2⦄ → ⦃G2, L⦄ ⊢ ➡*[h, o] L2 → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
-/3 width=5 by cpxs_fqus_fpbs, fpbs_lpxs_trans/ qed.
-
-lemma lpxs_lleq_fpbs: ∀h,o,G,L1,L,L2,T. ⦃G, L1⦄ ⊢ ➡*[h, o] L →
- L ≡[T, 0] L2 → ⦃G, L1, T⦄ ≥[h, o] ⦃G, L2, T⦄.
-/3 width=3 by lpxs_fpbs_trans, lleq_fpbs/ qed.
-
-(* Note: this is used in the closure proof *)
-lemma cpr_lpr_fpbs: ∀h,o,G,L1,L2,T1,T2. ⦃G, L1⦄ ⊢ T1 ➡ T2 → ⦃G, L1⦄ ⊢ ➡ L2 →
- ⦃G, L1, T1⦄ ≥[h, o] ⦃G, L2, T2⦄.
-/4 width=5 by fpbs_strap1, fpbq_fpbs, lpr_fpbq, cpr_fpbq/
-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/fpbq_aaa.ma".
-include "basic_2/computation/fpbs.ma".
-
-(* "QRST" PARALLEL COMPUTATION FOR CLOSURES *********************************)
-
-(* Properties on atomic arity assignment for terms **************************)
-
-lemma fpbs_aaa_conf: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ →
- ∀A1. ⦃G1, L1⦄ ⊢ T1 ⁝ A1 → ∃A2. ⦃G2, L2⦄ ⊢ T2 ⁝ A2.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fpbs_ind … H) -G2 -L2 -T2 /2 width=2 by ex_intro/
-#G #G2 #L #L2 #T #T2 #_ #H2 #IH1 #A #HA elim (IH1 … HA) -IH1 -A
-/2 width=8 by fpbq_aaa_conf/
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/btpredstaralt_8.ma".
-include "basic_2/multiple/lleq_fqus.ma".
-include "basic_2/computation/cpxs_lleq.ma".
-include "basic_2/computation/lpxs_lleq.ma".
-include "basic_2/computation/fpbs.ma".
-
-(* "QREST" PARALLEL COMPUTATION FOR CLOSURES ********************************)
-
-(* Note: alternative definition of fpbs *)
-definition fpbsa: ∀h. sd h → tri_relation genv lenv term ≝
- λh,o,G1,L1,T1,G2,L2,T2.
- ∃∃L0,L,T. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] T &
- ⦃G1, L1, T⦄ ⊐* ⦃G2, L0, T2⦄ &
- ⦃G2, L0⦄ ⊢ ➡*[h, o] L & L ≡[T2, 0] L2.
-
-interpretation "'big tree' parallel computation (closure) alternative"
- 'BTPRedStarAlt h o G1 L1 T1 G2 L2 T2 = (fpbsa h o G1 L1 T1 G2 L2 T2).
-
-(* Basic properties *********************************************************)
-
-lemma fpb_fpbsa_trans: ∀h,o,G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ ≽[h, o] ⦃G, L, T⦄ →
- ∀G2,L2,T2. ⦃G, L, T⦄ ≥≥[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥≥[h, o] ⦃G2, L2, T2⦄.
-#h #o #G1 #G #L1 #L #T1 #T * -G -L -T [ #G #L #T #HG1 | #T #HT1 | #L #HL1 | #L #HL1 ]
-#G2 #L2 #T2 * #L00 #L0 #T0 #HT0 #HG2 #HL00 #HL02
-[ elim (fquq_cpxs_trans … HT0 … HG1) -T
- /3 width=7 by fqus_strap2, ex4_3_intro/
-| /3 width=7 by cpxs_strap2, ex4_3_intro/
-| lapply (lpx_cpxs_trans … HT0 … HL1) -HT0 #HT10
- elim (lpx_fqus_trans … HG2 … HL1) -L
- /3 width=7 by lpxs_strap2, cpxs_trans, ex4_3_intro/
-| lapply (lleq_cpxs_trans … HT0 … HL1) -HT0 #HT0
- lapply (cpxs_lleq_conf_sn … HT0 … HL1) -HL1 #HL1
- elim (lleq_fqus_trans … HG2 … HL1) -L #K00 #HG12 #HKL00
- elim (lleq_lpxs_trans … HL00 … HKL00) -L00
- /3 width=9 by lleq_trans, ex4_3_intro/
-]
-qed-.
-
-(* Main properties **********************************************************)
-
-theorem fpbs_fpbsa: ∀h,o,G1,G2,L1,L2,T1,T2.
- ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥≥[h, o] ⦃G2, L2, T2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fpbs_ind_dx … H) -G1 -L1 -T1
-/2 width=7 by fpb_fpbsa_trans, ex4_3_intro/
-qed.
-
-(* Main inversion lemmas ****************************************************)
-
-theorem fpbsa_inv_fpbs: ∀h,o,G1,G2,L1,L2,T1,T2.
- ⦃G1, L1, T1⦄ ≥≥[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 *
-/3 width=5 by cpxs_fqus_lpxs_fpbs, fpbs_strap1, fpbq_lleq/
-qed-.
-
-(* Advanced properties ******************************************************)
-
-lemma fpbs_intro_alt: ∀h,o,G1,G2,L1,L0,L,L2,T1,T,T2.
- ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] T → ⦃G1, L1, T⦄ ⊐* ⦃G2, L0, T2⦄ →
- ⦃G2, L0⦄ ⊢ ➡*[h, o] L → L ≡[T2, 0] L2 → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ .
-/3 width=7 by fpbsa_inv_fpbs, ex4_3_intro/ qed.
-
-(* Advanced inversion lemmas *************************************************)
-
-lemma fpbs_inv_alt: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ →
- ∃∃L0,L,T. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] T &
- ⦃G1, L1, T⦄ ⊐* ⦃G2, L0, T2⦄ &
- ⦃G2, L0⦄ ⊢ ➡*[h, o] L & L ≡[T2, 0] L2.
-/2 width=1 by fpbs_fpbsa/ 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/fpbq_alt.ma".
-include "basic_2/computation/fpbs_alt.ma".
-
-(* "QRST" PARALLEL COMPUTATION FOR CLOSURES *********************************)
-
-(* Properties on extended context-sensitive parallel computation for terms **)
-
-lemma fpbs_cpx_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ →
- ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡[h, o] U2 → (T2 = U2 → ⊥) →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ≥[h, o] ⦃G2, L2, U2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #U2 #HTU2 #HnTU2 elim (fpbs_inv_alt … H) -H
-#L00 #L0 #T0 #HT10 #H10 #HL00 #HL02 lapply (lleq_cpx_trans … HTU2 … HL02) -HTU2
-#HTU2 lapply (cpx_lleq_conf_sn … HTU2 … HL02) -HL02
-#HL02 lapply (lpxs_cpx_trans … HTU2 … HL00) -HTU2
-#HTU2 elim (fqus_cpxs_trans_neq … H10 … HTU2 HnTU2) -H10 -HTU2 -HnTU2
-#U0 #HTU0 #HnTU0 #HU02 elim (eq_term_dec T1 T0) #HnT10 destruct
-[ -HT10 elim (cpxs_neq_inv_step_sn … HTU0 HnTU0) -HTU0 -HnTU0
-| -HnTU0 elim (cpxs_neq_inv_step_sn … HT10 HnT10) -HT10 -HnT10
-]
-/4 width=10 by fpbs_intro_alt, cpxs_trans, ex3_intro/
-qed-.
-
-(* Properties on "rst" proper parallel reduction on closures ****************)
-
-lemma fpb_fpbs: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
-/3 width=1 by fpbq_fpbs, fpb_fpbq/ qed.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/computation/fpbs.ma".
-
-(* "QRST" PARALLEL COMPUTATION FOR CLOSURES *********************************)
-
-(* Main properties **********************************************************)
-
-theorem fpbs_trans: ∀h,o. tri_transitive … (fpbs h o).
-/2 width=5 by tri_TC_transitive/ qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/computation/cpxs_lift.ma".
-include "basic_2/computation/fpbs.ma".
-
-(* "QRST" PARALLEL COMPUTATION FOR CLOSURES *********************************)
-
-(* Advanced properties ******************************************************)
-
-lemma lstas_fpbs: ∀h,o,G,L,T1,T2,d2. ⦃G, L⦄ ⊢ T1 •*[h, d2] T2 →
- ∀d1. d2 ≤ d1 → ⦃G, L⦄ ⊢ T1 ▪[h, o] d1 → ⦃G, L, T1⦄ ≥[h, o] ⦃G, L, T2⦄.
-/3 width=5 by cpxs_fpbs, lstas_cpxs/ qed.
-
-lemma sta_fpbs: ∀h,o,G,L,T,U,d.
- ⦃G, L⦄ ⊢ T ▪[h, o] d+1 → ⦃G, L⦄ ⊢ T •*[h, 1] U →
- ⦃G, L, T⦄ ≥[h, o] ⦃G, L, U⦄.
-/2 width=5 by lstas_fpbs/ qed.
-
-(* Note: this is used in the closure proof *)
-lemma cpr_lpr_sta_fpbs: ∀h,o,G,L1,L2,T1,T2,U2,d2.
- ⦃G, L1⦄ ⊢ T1 ➡ T2 → ⦃G, L1⦄ ⊢ ➡ L2 →
- ⦃G, L2⦄ ⊢ T2 ▪[h, o] d2+1 → ⦃G, L2⦄ ⊢ T2 •*[h, 1] U2 →
- ⦃G, L1, T1⦄ ≥[h, o] ⦃G, L2, U2⦄.
-/4 width=5 by fpbs_strap1, cpr_lpr_fpbs, sta_cpx, fpbq_cpx/ qed.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/btsn_5.ma".
-include "basic_2/reduction/fpb.ma".
-include "basic_2/computation/csx.ma".
-
-(* "QRST" STRONGLY NORMALIZING CLOSURES *************************************)
-
-inductive fsb (h) (o): relation3 genv lenv term ≝
-| fsb_intro: ∀G1,L1,T1. (
- ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ → fsb h o G2 L2 T2
- ) → fsb h o G1 L1 T1
-.
-
-interpretation
- "'qrst' strong normalization (closure)"
- 'BTSN h o G L T = (fsb h o G L T).
-
-(* Basic eliminators ********************************************************)
-
-lemma fsb_ind_alt: ∀h,o. ∀R: relation3 …. (
- ∀G1,L1,T1. ⦥[h,o] ⦃G1, L1, T1⦄ → (
- ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ → R G2 L2 T2
- ) → R G1 L1 T1
- ) →
- ∀G,L,T. ⦥[h, o] ⦃G, L, T⦄ → R G L T.
-#h #o #R #IH #G #L #T #H elim H -G -L -T
-/4 width=1 by fsb_intro/
-qed-.
-
-(* Basic inversion lemmas ***************************************************)
-
-lemma fsb_inv_csx: ∀h,o,G,L,T. ⦥[h, o] ⦃G, L, T⦄ → ⦃G, L⦄ ⊢ ⬊*[h, o] T.
-#h #o #G #L #T #H elim H -G -L -T /5 width=1 by csx_intro, fpb_cpx/
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/computation/fpbs_aaa.ma".
-include "basic_2/computation/csx_aaa.ma".
-include "basic_2/computation/fsb_csx.ma".
-
-(* "QRST" STRONGLY NORMALIZING CLOSURES *************************************)
-
-(* Main properties **********************************************************)
-
-(* Note: this is the "big tree" theorem ("RST" version) *)
-theorem aaa_fsb: ∀h,o,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ⦥[h, o] ⦃G, L, T⦄.
-/3 width=2 by aaa_csx, csx_fsb/ qed.
-
-(* Note: this is the "big tree" theorem ("QRST" version) *)
-theorem aaa_fsba: ∀h,o,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ⦥⦥[h, o] ⦃G, L, T⦄.
-/3 width=2 by fsb_fsba, aaa_fsb/ qed.
-
-(* Advanced eliminators on atomica arity assignment for terms ***************)
-
-fact aaa_ind_fpb_aux: ∀h,o. ∀R:relation3 genv lenv term.
- (∀G1,L1,T1,A. ⦃G1, L1⦄ ⊢ T1 ⁝ A →
- (∀G2,L2,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
- R G1 L1 T1
- ) →
- ∀G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → ∀A. ⦃G, L⦄ ⊢ T ⁝ A → R G L T.
-#h #o #R #IH #G #L #T #H @(csx_ind_fpb … H) -G -L -T
-#G1 #L1 #T1 #H1 #IH1 #A1 #HTA1 @IH -IH //
-#G2 #L2 #T2 #H12 elim (fpbs_aaa_conf h o … G2 … L2 … T2 … HTA1) -A1
-/2 width=2 by fpb_fpbs/
-qed-.
-
-lemma aaa_ind_fpb: ∀h,o. ∀R:relation3 genv lenv term.
- (∀G1,L1,T1,A. ⦃G1, L1⦄ ⊢ T1 ⁝ A →
- (∀G2,L2,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
- R G1 L1 T1
- ) →
- ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → R G L T.
-/4 width=4 by aaa_ind_fpb_aux, aaa_csx/ qed-.
-
-fact aaa_ind_fpbg_aux: ∀h,o. ∀R:relation3 genv lenv term.
- (∀G1,L1,T1,A. ⦃G1, L1⦄ ⊢ T1 ⁝ A →
- (∀G2,L2,T2. ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
- R G1 L1 T1
- ) →
- ∀G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → ∀A. ⦃G, L⦄ ⊢ T ⁝ A → R G L T.
-#h #o #R #IH #G #L #T #H @(csx_ind_fpbg … H) -G -L -T
-#G1 #L1 #T1 #H1 #IH1 #A1 #HTA1 @IH -IH //
-#G2 #L2 #T2 #H12 elim (fpbs_aaa_conf h o … G2 … L2 … T2 … HTA1) -A1
-/2 width=2 by fpbg_fwd_fpbs/
-qed-.
-
-lemma aaa_ind_fpbg: ∀h,o. ∀R:relation3 genv lenv term.
- (∀G1,L1,T1,A. ⦃G1, L1⦄ ⊢ T1 ⁝ A →
- (∀G2,L2,T2. ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
- R G1 L1 T1
- ) →
- ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → R G L T.
-/4 width=4 by aaa_ind_fpbg_aux, aaa_csx/ qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/btsnalt_5.ma".
-include "basic_2/computation/fpbg_fpbs.ma".
-include "basic_2/computation/fsb.ma".
-
-(* "QRST" STRONGLY NORMALIZING CLOSURES *************************************)
-
-(* Note: alternative definition of fsb *)
-inductive fsba (h) (o): relation3 genv lenv term ≝
-| fsba_intro: ∀G1,L1,T1. (
- ∀G2,L2,T2. ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄ → fsba h o G2 L2 T2
- ) → fsba h o G1 L1 T1.
-
-interpretation
- "'big tree' strong normalization (closure) alternative"
- 'BTSNAlt h o G L T = (fsba h o G L T).
-
-(* Basic eliminators ********************************************************)
-
-lemma fsba_ind_alt: ∀h,o. ∀R: relation3 …. (
- ∀G1,L1,T1. ⦥⦥[h,o] ⦃G1, L1, T1⦄ → (
- ∀G2,L2,T2. ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄ → R G2 L2 T2
- ) → R G1 L1 T1
- ) →
- ∀G,L,T. ⦥⦥[h, o] ⦃G, L, T⦄ → R G L T.
-#h #o #R #IH #G #L #T #H elim H -G -L -T
-/4 width=1 by fsba_intro/
-qed-.
-
-(* Basic properties *********************************************************)
-
-lemma fsba_fpbs_trans: ∀h,o,G1,L1,T1. ⦥⦥[h, o] ⦃G1, L1, T1⦄ →
- ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ → ⦥⦥[h, o] ⦃G2, L2, T2⦄.
-#h #o #G1 #L1 #T1 #H @(fsba_ind_alt … H) -G1 -L1 -T1
-/4 width=5 by fsba_intro, fpbs_fpbg_trans/
-qed-.
-
-(* Main properties **********************************************************)
-
-theorem fsb_fsba: ∀h,o,G,L,T. ⦥[h, o] ⦃G, L, T⦄ → ⦥⦥[h, o] ⦃G, L, T⦄.
-#h #o #G #L #T #H @(fsb_ind_alt … H) -G -L -T
-#G1 #L1 #T1 #_ #IH @fsba_intro
-#G2 #L2 #T2 * /3 width=5 by fsba_fpbs_trans/
-qed.
-
-(* Main inversion lemmas ****************************************************)
-
-theorem fsba_inv_fsb: ∀h,o,G,L,T. ⦥⦥[h, o] ⦃G, L, T⦄ → ⦥[h, o] ⦃G, L, T⦄.
-#h #o #G #L #T #H @(fsba_ind_alt … H) -G -L -T
-/4 width=1 by fsb_intro, fpb_fpbg/
-qed-.
-
-(* Advanced properties ******************************************************)
-
-lemma fsb_fpbs_trans: ∀h,o,G1,L1,T1. ⦥[h, o] ⦃G1, L1, T1⦄ →
- ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ → ⦥[h, o] ⦃G2, L2, T2⦄.
-/4 width=5 by fsba_inv_fsb, fsb_fsba, fsba_fpbs_trans/ qed-.
-
-(* Advanced eliminators *****************************************************)
-
-lemma fsb_ind_fpbg: ∀h,o. ∀R:relation3 genv lenv term.
- (∀G1,L1,T1. ⦥[h, o] ⦃G1, L1, T1⦄ →
- (∀G2,L2,T2. ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
- R G1 L1 T1
- ) →
- ∀G1,L1,T1. ⦥[h, o] ⦃G1, L1, T1⦄ → R G1 L1 T1.
-#h #o #R #IH #G1 #L1 #T1 #H @(fsba_ind_alt h o … G1 L1 T1)
-/3 width=1 by fsba_inv_fsb, fsb_fsba/
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/computation/fpbs_fpb.ma".
-include "basic_2/computation/fpbs_fpbs.ma".
-include "basic_2/computation/csx_fpbs.ma".
-include "basic_2/computation/lsx_csx.ma".
-include "basic_2/computation/fsb_alt.ma".
-
-(* "QRST" STRONGLY NORMALIZING CLOSURES *************************************)
-
-(* Advanced propreties on context-sensitive extended normalizing terms ******)
-
-lemma csx_fsb_fpbs: ∀h,o,G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬊*[h, o] T1 →
- ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ → ⦥[h, o] ⦃G2, L2, T2⦄.
-#h #o #G1 #L1 #T1 #H @(csx_ind … H) -T1
-#T1 #HT1 #IHc #G2 #L2 #T2 @(fqup_wf_ind … G2 L2 T2) -G2 -L2 -T2
-#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 #_ #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 @(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 ]
- [ #H destruct /5 width=5 by fsb_fpbs_trans, lpxs_fpbs, fpbs_fqup_trans, lpx_lpxs/
- | #HnT04 #HT04 #H04 #HL43 elim (cpxs_neq_inv_step_sn … HT04 HnT04) -HT04 -HnT04
- #T2 #HT02 #HnT02 #HT24 elim (fpbs_cpx_trans_neq … H10 … HT02 HnT02) -T0
- lapply (fpbs_intro_alt … G3 … L4 … L3 L3 … T3 … HT24 ? ? ?) -HT24
- /3 width=8 by fpbs_trans, lpx_lpxs, fqup_fqus/ (**) (* full auto too slow *)
- ]
- ]
-]
-qed.
-
-lemma csx_fsb: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → ⦥[h, o] ⦃G, L, T⦄.
-/2 width=5 by csx_fsb_fpbs/ qed.
-
-(* Advanced eliminators *****************************************************)
-
-lemma csx_ind_fpb: ∀h,o. ∀R:relation3 genv lenv term.
- (∀G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬊*[h, o] T1 →
- (∀G2,L2,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
- R G1 L1 T1
- ) →
- ∀G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → R G L T.
-/4 width=4 by fsb_inv_csx, csx_fsb, fsb_ind_alt/ qed-.
-
-lemma csx_ind_fpbg: ∀h,o. ∀R:relation3 genv lenv term.
- (∀G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬊*[h, o] T1 →
- (∀G2,L2,T2. ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
- R G1 L1 T1
- ) →
- ∀G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → R G L T.
-/4 width=4 by fsb_inv_csx, csx_fsb, fsb_ind_fpbg/ 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/genv.ma".
-include "basic_2/multiple/drops.ma".
-
-(* GENERIC COMPUTATION PROPERTIES *******************************************)
-
-definition nf ≝ λRR:relation4 genv lenv term term. λRS:relation term.
- λG,L,T. NF … (RR G L) RS T.
-
-definition candidate: Type[0] ≝ relation3 genv lenv term.
-
-definition CP0 ≝ λRR:relation4 genv lenv term term. λRS:relation term.
- ∀G. d_liftable1 (nf RR RS G) (Ⓕ).
-
-definition CP1 ≝ λRR:relation4 genv lenv term term. λRS:relation term.
- ∀G,L. ∃s. NF … (RR G L) RS (⋆s).
-
-definition CP2 ≝ λRP:candidate. ∀G. d_liftable1 (RP G) (Ⓕ).
-
-definition CP3 ≝ λRP:candidate.
- ∀G,L,T,s. RP G L (ⓐ⋆s.T) → RP G L T.
-
-(* requirements for generic computation properties *)
-record gcp (RR:relation4 genv lenv term term) (RS:relation term) (RP:candidate) : Prop ≝
-{ cp0: CP0 RR RS;
- cp1: CP1 RR RS;
- cp2: CP2 RP;
- cp3: CP3 RP
-}.
-
-(* Basic properties *********************************************************)
-
-(* Basic_1: was: nf2_lift1 *)
-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. 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. d_liftables1_all (RP G) (Ⓕ).
-#RR #RS #RP #H #G @d1_liftables_liftables_all /2 width=7 by gcp2_lifts/
-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/lifts_lifts.ma".
-include "basic_2/multiple/drops_drops.ma".
-include "basic_2/static/aaa_lifts.ma".
-include "basic_2/static/aaa_aaa.ma".
-include "basic_2/computation/lsubc_drops.ma".
-
-(* GENERIC COMPUTATION PROPERTIES *******************************************)
-
-(* Main properties **********************************************************)
-
-(* Basic_1: was: sc3_arity_csubc *)
-theorem acr_aaa_csubc_lifts: ∀RR,RS,RP.
- gcp RR RS RP → gcr RR RS RP RP →
- ∀G,L1,T,A. ⦃G, L1⦄ ⊢ T ⁝ A → ∀L0,cs. ⬇*[Ⓕ, cs] L0 ≡ L1 →
- ∀T0. ⬆*[cs] T ≡ T0 → ∀L2. G ⊢ L2 ⫃[RP] L0 →
- ⦃G, L2, T0⦄ ϵ[RP] 〚A〛.
-#RR #RS #RP #H1RP #H2RP #G #L1 #T #A #H elim H -G -L1 -T -A
-[ #G #L #s #L0 #cs #HL0 #X #H #L2 #HL20
- >(lifts_inv_sort1 … H) -H
- lapply (acr_gcr … H1RP H2RP (⓪)) #HAtom
- lapply (c4 … HAtom G L2 (◊)) /2 width=1 by/
-| #I #G #L1 #K1 #V1 #B #i #HLK1 #HKV1B #IHB #L0 #cs #HL01 #X #H #L2 #HL20
- 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 #cs1 #i0 #HL0 #H #Hi0 #Hcs1
- >(at_mono … Hi1 … Hi0) -i1
- elim (drops_inv_skip2 … Hcs1 … H) -cs1 #K0 #V0 #cs0 #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 … Hcs0 … HV10 … HV0) -HV10 #V2 #HV12 #HV2
- lapply (c5 … HB ? G ? ? (◊) … HV0 HLK2) /3 width=7 by drops_cons, lifts_cons/ (* Note: uses IHB HL20 V2 HV0 *)
- | -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
- lapply (aaa_mono … H2KV0A … HKV0B) #H destruct -H2KV0A -HKV0B
- elim (lift_total V0 0 (i0 +1)) #V3 #HV03
- elim (lift_total V2 0 (i0 +1)) #V #HV2
- lapply (c5 … HB ? G ? ? (◊) … (ⓝV3.V) … HLK2) /2 width=1 by lift_flat/
- lapply (c7 … HB G L2 (◊)) /3 width=7 by gcr_lift/
- ]
-| #a #G #L #V #T #B #A #_ #_ #IHB #IHA #L0 #cs #HL0 #X #H #L2 #HL20
- elim (lifts_inv_bind1 … H) -H #V0 #T0 #HV0 #HT0 #H destruct
- lapply (acr_gcr … H1RP H2RP A) #HA
- lapply (acr_gcr … H1RP H2RP B) #HB
- lapply (c1 … HB) -HB #HB
- lapply (c6 … HA G L2 (◊) (◊)) /4 width=5 by lsubc_pair, drops_skip, liftv_nil/
-| #a #G #L #W #T #B #A #HLWB #_ #IHB #IHA #L0 #cs #HL0 #X #H #L2 #HL02
- elim (lifts_inv_bind1 … H) -H #W0 #T0 #HW0 #HT0 #H destruct
- @(acr_abst … H1RP H2RP) /2 width=5 by/
- #L3 #V3 #W3 #T3 #cs3 #HL32 #HW03 #HT03 #H1B #H2B
- elim (drops_lsubc_trans … H1RP … HL32 … HL02) -L2 #L2 #HL32 #HL20
- lapply (aaa_lifts … L2 W3 … (cs @@ cs3) … HLWB) -HLWB /2 width=4 by drops_trans, lifts_trans/ #HLW2B
- @(IHA (L2. ⓛW3) … (cs + 1 @@ cs3 + 1)) -IHA
- /3 width=5 by lsubc_beta, drops_trans, drops_skip, lifts_trans/
-| #G #L #V #T #B #A #_ #_ #IHB #IHA #L0 #cs #HL0 #X #H #L2 #HL20
- elim (lifts_inv_flat1 … H) -H #V0 #T0 #HV0 #HT0 #H destruct
- /3 width=10 by drops_nil, lifts_nil/
-| #G #L #V #T #A #_ #_ #IH1A #IH2A #L0 #cs #HL0 #X #H #L2 #HL20
- elim (lifts_inv_flat1 … H) -H #V0 #T0 #HV0 #HT0 #H destruct
- lapply (acr_gcr … H1RP H2RP A) #HA
- lapply (c7 … HA G L2 (◊)) /3 width=5 by/
-]
-qed.
-
-(* Basic_1: was: sc3_arity *)
-lemma acr_aaa: ∀RR,RS,RP. gcp RR RS RP → gcr RR RS RP RP →
- ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ⦃G, L, T⦄ ϵ[RP] 〚A〛.
-/2 width=8 by drops_nil, lifts_nil, acr_aaa_csubc_lifts/ qed.
-
-lemma gcr_aaa: ∀RR,RS,RP. gcp RR RS RP → gcr RR RS RP RP →
- ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → RP G L T.
-#RR #RS #RP #H1RP #H2RP #G #L #T #A #HT
-lapply (acr_gcr … H1RP H2RP A) #HA
-@(c1 … HA) /2 width=4 by acr_aaa/
-qed.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/ineint_5.ma".
-include "basic_2/grammar/aarity.ma".
-include "basic_2/multiple/mr2_mr2.ma".
-include "basic_2/multiple/lifts_lift_vector.ma".
-include "basic_2/multiple/drops_drop.ma".
-include "basic_2/computation/gcp.ma".
-
-(* GENERIC COMPUTATION PROPERTIES *******************************************)
-
-(* Note: this is Girard's CR1 *)
-definition S1 ≝ λRP,C:candidate.
- ∀G,L,T. C G L T → RP G L T.
-
-(* Note: this is Tait's iii, or Girard's CR4 *)
-definition S2 ≝ λRR:relation4 genv lenv term term. λRS:relation term. λRP,C:candidate.
- ∀G,L,Vs. all … (RP G L) Vs →
- ∀T. 𝐒⦃T⦄ → NF … (RR G L) RS T → C G L (ⒶVs.T).
-
-(* Note: this generalizes Tait's ii *)
-definition S3 ≝ λC:candidate.
- ∀a,G,L,Vs,V,T,W.
- C G L (ⒶVs.ⓓ{a}ⓝW.V.T) → C G L (ⒶVs.ⓐV.ⓛ{a}W.T).
-
-definition S4 ≝ λRP,C:candidate.
- ∀G,L,Vs. all … (RP G L) Vs → ∀s. C G L (ⒶVs.⋆s).
-
-definition S5 ≝ λC:candidate. ∀I,G,L,K,Vs,V1,V2,i.
- C G L (ⒶVs.V2) → ⬆[0, i+1] V1 ≡ V2 →
- ⬇[i] L ≡ K.ⓑ{I}V1 → C G L (ⒶVs.#i).
-
-definition S6 ≝ λRP,C:candidate.
- ∀G,L,V1c,V2c. ⬆[0, 1] V1c ≡ V2c →
- ∀a,V,T. C G (L.ⓓV) (ⒶV2c.T) → RP G L V → C G L (ⒶV1c.ⓓ{a}V.T).
-
-definition S7 ≝ λC:candidate.
- ∀G,L,Vs,T,W. C G L (ⒶVs.T) → C G L (ⒶVs.W) → C G L (ⒶVs.ⓝW.T).
-
-(* requirements for the generic reducibility candidate *)
-record gcr (RR:relation4 genv lenv term term) (RS:relation term) (RP,C:candidate) : Prop ≝
-{ c1: S1 RP C;
- c2: S2 RR RS RP C;
- c3: S3 C;
- c4: S4 RP C;
- c5: S5 C;
- c6: S6 RP C;
- c7: S7 C
-}.
-
-(* the functional construction for candidates *)
-definition cfun: candidate → candidate → candidate ≝
- λC1,C2,G,K,T. ∀L,W,U,cs.
- ⬇*[Ⓕ, cs] L ≡ K → ⬆*[cs] T ≡ U → C1 G L W → C2 G L (ⓐW.U).
-
-(* the reducibility candidate associated to an atomic arity *)
-rec definition acr (RP:candidate) (A:aarity) on A: candidate ≝
-match A with
-[ AAtom ⇒ RP
-| APair B A ⇒ cfun (acr RP B) (acr RP A)
-].
-
-interpretation
- "candidate of reducibility of an atomic arity (abstract)"
- 'InEInt RP G L T A = (acr RP A G L T).
-
-(* Basic properties *********************************************************)
-
-(* Basic 1: was: sc3_lift *)
-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. 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:
- sc3_sn3 sc3_abst sc3_appl sc3_abbr sc3_bind sc3_cast
-*)
-lemma acr_gcr: ∀RR,RS,RP. gcp RR RS RP → gcr RR RS RP RP →
- ∀A. gcr RR RS RP (acr RP A).
-#RR #RS #RP #H1RP #H2RP #A elim A -A //
-#B #A #IHB #IHA @mk_gcr
-[ #G #L #T #H
- elim (cp1 … H1RP G L) #s #HK
- lapply (H L (⋆s) T (◊) ? ? ?) -H //
- [ lapply (c2 … IHB G L (◊) … HK) //
- | /3 width=6 by c1, cp3/
- ]
-| #G #L #Vs #HVs #T #H1T #H2T #L0 #V0 #X #cs #HL0 #H #HB
- elim (lifts_inv_applv1 … H) -H #V0c #T0 #HV0c #HT0 #H destruct
- lapply (c1 … IHB … HB) #HV0
- @(c2 … IHA … (V0 @ V0c))
- /3 width=14 by gcp2_lifts_all, gcp2_lifts, gcp0_lifts, lifts_simple_dx, conj/
-| #a #G #L #Vs #U #T #W #HA #L0 #V0 #X #cs #HL0 #H #HB
- elim (lifts_inv_applv1 … H) -H #V0c #Y #HV0c #HY #H destruct
- elim (lifts_inv_flat1 … HY) -HY #U0 #X #HU0 #HX #H destruct
- elim (lifts_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT0 #H destruct
- @(c3 … IHA … (V0 @ V0c)) /5 width=6 by lifts_applv, lifts_flat, lifts_bind/
-| #G #L #Vs #HVs #s #L0 #V0 #X #cs #HL0 #H #HB
- elim (lifts_inv_applv1 … H) -H #V0c #Y #HV0c #HY #H destruct
- >(lifts_inv_sort1 … HY) -Y
- lapply (c1 … IHB … HB) #HV0
- @(c4 … IHA … (V0 @ V0c)) /3 width=7 by gcp2_lifts_all, conj/
-| #I #G #L #K #Vs #V1 #V2 #i #HA #HV12 #HLK #L0 #V0 #X #cs #HL0 #H #HB
- elim (lifts_inv_applv1 … H) -H #V0c #Y #HV0c #HY #H destruct
- elim (lifts_inv_lref1 … HY) -HY #i0 #Hi0 #H destruct
- elim (drops_drop_trans … HL0 … HLK) #X #cs0 #i1 #HL02 #H #Hi1 #Hcs0
- >(at_mono … Hi1 … Hi0) in HL02; -i1 #HL02
- elim (drops_inv_skip2 … Hcs0 … H) -H -cs0 #L2 #W1 #cs0 #Hcs0 #HLK #HVW1 #H destruct
- elim (lift_total W1 0 (i0 + 1)) #W2 #HW12
- elim (lifts_lift_trans … Hcs0 … HVW1 … HW12) // -Hcs0 -Hi0 #V3 #HV13 #HVW2
- >(lift_mono … HV13 … HV12) in HVW2; -V3 #HVW2
- @(c5 … IHA … (V0 @ V0c) … HW12 HL02) /3 width=5 by lifts_applv/
-| #G #L #V1c #V2c #HV12c #a #V #T #HA #HV #L0 #V10 #X #cs #HL0 #H #HB
- elim (lifts_inv_applv1 … H) -H #V10c #Y #HV10c #HY #H destruct
- elim (lifts_inv_bind1 … HY) -HY #V0 #T0 #HV0 #HT0 #H destruct
- elim (lift_total V10 0 1) #V20 #HV120
- elim (liftv_total 0 1 V10c) #V20c #HV120c
- @(c6 … IHA … (V10 @ V10c) (V20 @ V20c)) /3 width=7 by gcp2_lifts, liftv_cons/
- @(HA … (cs + 1)) /2 width=2 by drops_skip/
- [ @lifts_applv //
- elim (liftsv_liftv_trans_le … HV10c … HV120c) -V10c #V10c #HV10c #HV120c
- >(liftv_mono … HV12c … HV10c) -V1c //
- | @(gcr_lift … H1RP … HB … HV120) /2 width=2 by drop_drop/
- ]
-| #G #L #Vs #T #W #HA #HW #L0 #V0 #X #cs #HL0 #H #HB
- elim (lifts_inv_applv1 … H) -H #V0c #Y #HV0c #HY #H destruct
- elim (lifts_inv_flat1 … HY) -HY #W0 #T0 #HW0 #HT0 #H destruct
- @(c7 … IHA … (V0 @ V0c)) /3 width=5 by lifts_applv/
-]
-qed.
-
-lemma acr_abst: ∀RR,RS,RP. gcp RR RS RP → gcr RR RS RP RP →
- ∀a,G,L,W,T,A,B. ⦃G, L, W⦄ ϵ[RP] 〚B〛 → (
- ∀L0,V0,W0,T0,cs. ⬇*[Ⓕ, cs] L0 ≡ L → ⬆*[cs] W ≡ W0 → ⬆*[cs + 1] T ≡ T0 →
- ⦃G, L0, V0⦄ ϵ[RP] 〚B〛 → ⦃G, L0, W0⦄ ϵ[RP] 〚B〛 → ⦃G, L0.ⓓⓝW0.V0, T0⦄ ϵ[RP] 〚A〛
- ) →
- ⦃G, L, ⓛ{a}W.T⦄ ϵ[RP] 〚②B.A〛.
-#RR #RS #RP #H1RP #H2RP #a #G #L #W #T #A #B #HW #HA #L0 #V0 #X #cs #HL0 #H #HB
-lapply (acr_gcr … H1RP H2RP A) #HCA
-lapply (acr_gcr … H1RP H2RP B) #HCB
-elim (lifts_inv_bind1 … H) -H #W0 #T0 #HW0 #HT0 #H destruct
-lapply (gcr_lifts … H1RP … HL0 … HW0 HW) -HW #HW0
-lapply (c3 … HCA … a G L0 (◊)) #H @H -H
-lapply (c6 … HCA G L0 (◊) (◊) ?) // #H @H -H
-[ @(HA … HL0) //
-| lapply (c1 … HCB) -HCB #HCB
- lapply (c7 … H2RP G L0 (◊)) /3 width=1 by/
-]
-qed.
-
-(* Basic_1: removed theorems 2: sc3_arity_gen sc3_repl *)
-(* Basic_1: removed local theorems 1: sc3_sn3_abst *)
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/cosn_5.ma".
-include "basic_2/computation/lsx.ma".
-
-(* SN EXTENDED STRONGLY CONORMALIZING LOCAL ENVIRONMENTS ********************)
-
-inductive lcosx (h) (o) (G): relation2 ynat lenv ≝
-| lcosx_sort: ∀l. lcosx h o G l (⋆)
-| lcosx_skip: ∀I,L,T. lcosx h o G 0 L → lcosx h o G 0 (L.ⓑ{I}T)
-| lcosx_pair: ∀I,L,T,l. G ⊢ ⬊*[h, o, T, l] L →
- lcosx h o G l L → lcosx h o G (⫯l) (L.ⓑ{I}T)
-.
-
-interpretation
- "sn extended strong conormalization (local environment)"
- 'CoSN h o l G L = (lcosx h o G l L).
-
-(* Basic properties *********************************************************)
-
-lemma lcosx_O: ∀h,o,G,L. G ⊢ ~⬊*[h, o, 0] L.
-#h #o #G #L elim L /2 width=1 by lcosx_skip/
-qed.
-
-lemma lcosx_drop_trans_lt: ∀h,o,G,L,l. G ⊢ ~⬊*[h, o, l] L →
- ∀I,K,V,i. ⬇[i] L ≡ K.ⓑ{I}V → i < l →
- G ⊢ ~⬊*[h, o, ⫰(l-i)] K ∧ G ⊢ ⬊*[h, o, V, ⫰(l-i)] K.
-#h #o #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 #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 … Hil ?) -Hil /2 width=1 by ylt_inj/ >ypred_succ #Hil
- elim (IHL … HLK ?) -IHL -HLK <yminus_inj >yminus_SO2 //
- <(ypred_succ l) in ⊢ (%→%→?); >yminus_pred /2 width=1 by ylt_inj, conj/
- ]
-]
-qed-.
-
-(* Basic inversion lemmas ***************************************************)
-
-fact lcosx_inv_succ_aux: ∀h,o,G,L,x. G ⊢ ~⬊*[h, o, x] L → ∀l. x = ⫯l →
- L = ⋆ ∨
- ∃∃I,K,V. L = K.ⓑ{I}V & G ⊢ ~⬊*[h, o, l] K &
- G ⊢ ⬊*[h, o, V, l] K.
-#h #o #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 #l #HT #HL #x #H <(ysucc_inv_inj … H) -x
- /3 width=6 by ex3_3_intro, or_intror/
-]
-qed-.
-
-lemma lcosx_inv_succ: ∀h,o,G,L,l. G ⊢ ~⬊*[h, o, ⫯l] L → L = ⋆ ∨
- ∃∃I,K,V. L = K.ⓑ{I}V & G ⊢ ~⬊*[h, o, l] K &
- G ⊢ ⬊*[h, o, V, l] K.
-/2 width=3 by lcosx_inv_succ_aux/ qed-.
-
-lemma lcosx_inv_pair: ∀h,o,I,G,L,T,l. G ⊢ ~⬊*[h, o, ⫯l] L.ⓑ{I}T →
- G ⊢ ~⬊*[h, o, l] L ∧ G ⊢ ⬊*[h, o, T, l] L.
-#h #o #I #G #L #T #l #H elim (lcosx_inv_succ … H) -H
-[ #H destruct
-| * #Z #Y #X #H destruct /2 width=1 by conj/
-]
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/computation/lsx_drop.ma".
-include "basic_2/computation/lsx_lpx.ma".
-include "basic_2/computation/lsx_lpxs.ma".
-include "basic_2/computation/lcosx.ma".
-
-(* SN EXTENDED STRONGLY CONORMALIZING LOCAL ENVIRONMENTS ********************)
-
-(* Properties on extended context-sensitive parallel reduction for term *****)
-
-lemma lsx_cpx_trans_lcosx: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 →
- ∀l. G ⊢ ~⬊*[h, o, l] L →
- G ⊢ ⬊*[h, o, T1, l] L → G ⊢ ⬊*[h, o, T2, l] L.
-#h #o #G #L #T1 #T2 #H elim H -G -L -T1 -T2 //
-[ #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_succ1/
- 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 -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 #l #HL #H
- elim (lsx_inv_bind … H) -H #HV1 #HT1
- @lsx_bind /2 width=2 by/ (**) (* explicit constructor *)
- @(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 #l #HL #H
- elim (lsx_inv_bind … H) -H #HV #HU1
- <(ypred_succ l) <yminus_SO2
- /4 width=7 by lcosx_pair, lsx_inv_lift_ge, drop_drop/
-| #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 #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 #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_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 [ <yplus_SO2 /3 width=7 by lsx_lift_ge, drop_drop/ ]
- @(lsx_lreq_conf … (L.ⓓW1)) /3 width=1 by lcosx_pair, lreq_succ/
-]
-qed-.
-
-lemma lsx_cpx_trans_O: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 →
- G ⊢ ⬊*[h, o, T1, 0] L → G ⊢ ⬊*[h, o, T2, 0] L.
-/2 width=3 by lsx_cpx_trans_lcosx/ qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/predsnstar_3.ma".
-include "basic_2/substitution/lpx_sn_tc.ma".
-include "basic_2/reduction/lpr.ma".
-
-(* SN PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS ****************************)
-
-definition lprs: relation3 genv lenv lenv ≝
- λG. TC … (lpr G).
-
-interpretation "parallel computation (local environment, sn variant)"
- 'PRedSnStar G L1 L2 = (lprs G L1 L2).
-
-(* Basic eliminators ********************************************************)
-
-lemma lprs_ind: ∀G,L1. ∀R:predicate lenv. R L1 →
- (∀L,L2. ⦃G, L1⦄ ⊢ ➡* L → ⦃G, L⦄ ⊢ ➡ L2 → R L → R L2) →
- ∀L2. ⦃G, L1⦄ ⊢ ➡* L2 → R L2.
-#G #L1 #R #HL1 #IHL1 #L2 #HL12
-@(TC_star_ind … HL1 IHL1 … HL12) //
-qed-.
-
-lemma lprs_ind_dx: ∀G,L2. ∀R:predicate lenv. R L2 →
- (∀L1,L. ⦃G, L1⦄ ⊢ ➡ L → ⦃G, L⦄ ⊢ ➡* L2 → R L → R L1) →
- ∀L1. ⦃G, L1⦄ ⊢ ➡* L2 → R L1.
-#G #L2 #R #HL2 #IHL2 #L1 #HL12
-@(TC_star_ind_dx … HL2 IHL2 … HL12) //
-qed-.
-
-(* Basic properties *********************************************************)
-
-lemma lpr_lprs: ∀G,L1,L2. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L1⦄ ⊢ ➡* L2.
-/2 width=1 by inj/ qed.
-
-lemma lprs_refl: ∀G,L. ⦃G, L⦄ ⊢ ➡* L.
-/2 width=1 by lpr_lprs/ qed.
-
-lemma lprs_strap1: ∀G,L1,L,L2. ⦃G, L1⦄ ⊢ ➡* L → ⦃G, L⦄ ⊢ ➡ L2 → ⦃G, L1⦄ ⊢ ➡* L2.
-/2 width=3 by step/ qed-.
-
-lemma lprs_strap2: ∀G,L1,L,L2. ⦃G, L1⦄ ⊢ ➡ L → ⦃G, L⦄ ⊢ ➡* L2 → ⦃G, L1⦄ ⊢ ➡* L2.
-/2 width=3 by TC_strap/ qed-.
-
-lemma lprs_pair_refl: ∀G,L1,L2. ⦃G, L1⦄ ⊢ ➡* L2 → ∀I,V. ⦃G, L1.ⓑ{I}V⦄ ⊢ ➡* L2.ⓑ{I}V.
-/2 width=1 by TC_lpx_sn_pair_refl/ qed.
-
-(* Basic inversion lemmas ***************************************************)
-
-lemma lprs_inv_atom1: ∀G,L2. ⦃G, ⋆⦄ ⊢ ➡* L2 → L2 = ⋆.
-/2 width=2 by TC_lpx_sn_inv_atom1/ qed-.
-
-lemma lprs_inv_atom2: ∀G,L1. ⦃G, L1⦄ ⊢ ➡* ⋆ → L1 = ⋆.
-/2 width=2 by TC_lpx_sn_inv_atom2/ qed-.
-
-(* Basic forward lemmas *****************************************************)
-
-lemma lprs_fwd_length: ∀G,L1,L2. ⦃G, L1⦄ ⊢ ➡* L2 → |L1| = |L2|.
-/2 width=2 by TC_lpx_sn_fwd_length/ qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/computation/cprs_cprs.ma".
-include "basic_2/computation/lprs.ma".
-
-(* SN PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS ****************************)
-
-(* Advanced properties ******************************************************)
-
-lemma lprs_pair: ∀I,G,L1,L2. ⦃G, L1⦄ ⊢ ➡* L2 →
- ∀V1,V2. ⦃G, L1⦄ ⊢ V1 ➡* V2 → ⦃G, L1.ⓑ{I}V1⦄ ⊢ ➡* L2.ⓑ{I}V2.
-/2 width=1 by TC_lpx_sn_pair/ qed.
-
-(* Advanced inversion lemmas ************************************************)
-
-lemma lprs_inv_pair1: ∀I,G,K1,L2,V1. ⦃G, K1.ⓑ{I}V1⦄ ⊢ ➡* L2 →
- ∃∃K2,V2. ⦃G, K1⦄ ⊢ ➡* K2 & ⦃G, K1⦄ ⊢ V1 ➡* V2 &
- L2 = K2.ⓑ{I}V2.
-/3 width=3 by TC_lpx_sn_inv_pair1, lpr_cprs_trans/ qed-.
-
-lemma lprs_inv_pair2: ∀I,G,L1,K2,V2. ⦃G, L1⦄ ⊢ ➡* K2.ⓑ{I}V2 →
- ∃∃K1,V1. ⦃G, K1⦄ ⊢ ➡* K2 & ⦃G, K1⦄ ⊢ V1 ➡* V2 &
- L1 = K1.ⓑ{I}V1.
-/3 width=3 by TC_lpx_sn_inv_pair2, lpr_cprs_trans/ qed-.
-
-(* Advanced eliminators *****************************************************)
-
-lemma lprs_ind_alt: ∀G. ∀R:relation lenv.
- R (⋆) (⋆) → (
- ∀I,K1,K2,V1,V2.
- ⦃G, K1⦄ ⊢ ➡* K2 → ⦃G, K1⦄ ⊢ V1 ➡* V2 →
- R K1 K2 → R (K1.ⓑ{I}V1) (K2.ⓑ{I}V2)
- ) →
- ∀L1,L2. ⦃G, L1⦄ ⊢ ➡* L2 → R L1 L2.
-/3 width=4 by TC_lpx_sn_ind, lpr_cprs_trans/ qed-.
-
-(* Properties on context-sensitive parallel computation for terms ***********)
-
-lemma lprs_cpr_trans: ∀G. c_r_transitive … (cpr G) (λ_. lprs G).
-/3 width=5 by c_r_trans_LTC2, lpr_cprs_trans/ qed-.
-
-(* Basic_1: was just: pr3_pr3_pr3_t *)
-(* Note: alternative proof /3 width=5 by s_r_trans_LTC1, lprs_cpr_trans/ *)
-lemma lprs_cprs_trans: ∀G. c_rs_transitive … (cpr G) (λ_. lprs G).
-#G @c_r_to_c_rs_trans @c_r_trans_LTC2
-@c_rs_trans_TC1 /2 width=3 by lpr_cprs_trans/ (**) (* full auto too slow *)
-qed-.
-
-lemma lprs_cprs_conf_dx: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡* T1 →
- ∀L1. ⦃G, L0⦄ ⊢ ➡* L1 →
- ∃∃T. ⦃G, L1⦄ ⊢ T1 ➡* T & ⦃G, L1⦄ ⊢ T0 ➡* T.
-#G #L0 #T0 #T1 #HT01 #L1 #H @(lprs_ind … H) -L1 /2 width=3 by ex2_intro/
-#L #L1 #_ #HL1 * #T #HT1 #HT0 -L0
-elim (cprs_lpr_conf_dx … HT1 … HL1) -HT1 #T2 #HT2
-elim (cprs_lpr_conf_dx … HT0 … HL1) -L #T3 #HT3
-elim (cprs_conf … HT2 … HT3) -T
-/3 width=5 by cprs_trans, ex2_intro/
-qed-.
-
-lemma lprs_cpr_conf_dx: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡ T1 →
- ∀L1. ⦃G, L0⦄ ⊢ ➡* L1 →
- ∃∃T. ⦃G, L1⦄ ⊢ T1 ➡* T & ⦃G, L1⦄ ⊢ T0 ➡* T.
-/3 width=3 by lprs_cprs_conf_dx, cpr_cprs/ qed-.
-
-(* Note: this can be proved on its own using lprs_ind_dx *)
-lemma lprs_cprs_conf_sn: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡* T1 →
- ∀L1. ⦃G, L0⦄ ⊢ ➡* L1 →
- ∃∃T. ⦃G, L0⦄ ⊢ T1 ➡* T & ⦃G, L1⦄ ⊢ T0 ➡* T.
-#G #L0 #T0 #T1 #HT01 #L1 #HL01
-elim (lprs_cprs_conf_dx … HT01 … HL01) -HT01
-/3 width=3 by lprs_cprs_trans, ex2_intro/
-qed-.
-
-lemma lprs_cpr_conf_sn: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡ T1 →
- ∀L1. ⦃G, L0⦄ ⊢ ➡* L1 →
- ∃∃T. ⦃G, L0⦄ ⊢ T1 ➡* T & ⦃G, L1⦄ ⊢ T0 ➡* T.
-/3 width=3 by lprs_cprs_conf_sn, cpr_cprs/ qed-.
-
-lemma cprs_bind2: ∀G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡* V2 →
- ∀I,T1,T2. ⦃G, L.ⓑ{I}V2⦄ ⊢ T1 ➡* T2 →
- ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡* ⓑ{a,I}V2.T2.
-/4 width=5 by lprs_cprs_trans, lprs_pair, cprs_bind/ qed.
-
-(* Inversion lemmas on context-sensitive parallel computation for terms *****)
-
-(* Basic_1: was: pr3_gen_abst *)
-lemma cprs_inv_abst1: ∀a,G,L,W1,T1,U2. ⦃G, L⦄ ⊢ ⓛ{a}W1.T1 ➡* U2 →
- ∃∃W2,T2. ⦃G, L⦄ ⊢ W1 ➡* W2 & ⦃G, L.ⓛW1⦄ ⊢ T1 ➡* T2 &
- U2 = ⓛ{a}W2.T2.
-#a #G #L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /2 width=5 by ex3_2_intro/
-#U0 #U2 #_ #HU02 * #V0 #T0 #HV10 #HT10 #H destruct
-elim (cpr_inv_abst1 … HU02) -HU02 #V2 #T2 #HV02 #HT02 #H destruct
-lapply (lprs_cpr_trans … HT02 (L.ⓛV1) ?)
-/3 width=5 by lprs_pair, cprs_trans, cprs_strap1, ex3_2_intro/
-qed-.
-
-lemma cprs_inv_abst: ∀a,G,L,W1,W2,T1,T2. ⦃G, L⦄ ⊢ ⓛ{a}W1.T1 ➡* ⓛ{a}W2.T2 →
- ⦃G, L⦄ ⊢ W1 ➡* W2 ∧ ⦃G, L.ⓛW1⦄ ⊢ T1 ➡* T2.
-#a #G #L #W1 #W2 #T1 #T2 #H elim (cprs_inv_abst1 … H) -H
-#W #T #HW1 #HT1 #H destruct /2 width=1 by conj/
-qed-.
-
-(* Basic_1: was pr3_gen_abbr *)
-lemma cprs_inv_abbr1: ∀a,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓓ{a}V1.T1 ➡* U2 → (
- ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡* V2 & ⦃G, L.ⓓV1⦄ ⊢ T1 ➡* T2 &
- U2 = ⓓ{a}V2.T2
- ) ∨
- ∃∃T2. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡* T2 & ⬆[0, 1] U2 ≡ T2 & a = true.
-#a #G #L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /3 width=5 by ex3_2_intro, or_introl/
-#U0 #U2 #_ #HU02 * *
-[ #V0 #T0 #HV10 #HT10 #H destruct
- elim (cpr_inv_abbr1 … HU02) -HU02 *
- [ #V2 #T2 #HV02 #HT02 #H destruct
- lapply (lprs_cpr_trans … HT02 (L.ⓓV1) ?)
- /4 width=5 by lprs_pair, cprs_trans, cprs_strap1, ex3_2_intro, or_introl/
- | #T2 #HT02 #HUT2
- lapply (lprs_cpr_trans … HT02 (L.ⓓV1) ?) -HT02
- /4 width=3 by lprs_pair, cprs_trans, ex3_intro, or_intror/
- ]
-| #U1 #HTU1 #HU01 elim (lift_total U2 0 1)
- #U #HU2 lapply (cpr_lift … HU02 (L.ⓓV1) … HU01 … HU2) -U0
- /4 width=3 by cprs_strap1, drop_drop, ex3_intro, or_intror/
-]
-qed-.
-
-(* More advanced properties *************************************************)
-
-lemma lprs_pair2: ∀I,G,L1,L2. ⦃G, L1⦄ ⊢ ➡* L2 →
- ∀V1,V2. ⦃G, L2⦄ ⊢ V1 ➡* V2 → ⦃G, L1.ⓑ{I}V1⦄ ⊢ ➡* L2.ⓑ{I}V2.
-/3 width=3 by lprs_pair, lprs_cprs_trans/ qed.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/reduction/lpr_drop.ma".
-include "basic_2/computation/lprs.ma".
-
-(* SN PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS ****************************)
-
-(* Properties on local environment slicing ***********************************)
-
-lemma lprs_drop_conf: ∀G. dropable_sn (lprs G).
-/3 width=3 by dropable_sn_TC, lpr_drop_conf/ qed-.
-
-lemma drop_lprs_trans: ∀G. dedropable_sn (lprs G).
-/3 width=3 by dedropable_sn_TC, drop_lpr_trans/ qed-.
-
-lemma lprs_drop_trans_O1: ∀G. dropable_dx (lprs G).
-/3 width=3 by dropable_dx_TC, lpr_drop_trans_O1/ qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/reduction/lpr_lpr.ma".
-include "basic_2/computation/lprs.ma".
-
-(* SN PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS ****************************)
-
-(* Advanced properties ******************************************************)
-
-lemma lprs_strip: ∀G. confluent2 … (lprs G) (lpr G).
-/3 width=3 by TC_strip1, lpr_conf/ qed-.
-
-(* Main properties **********************************************************)
-
-theorem lprs_conf: ∀G. confluent2 … (lprs G) (lprs G).
-/3 width=3 by TC_confluent2, lpr_conf/ qed-.
-
-theorem lprs_trans: ∀G. Transitive … (lprs G).
-/2 width=3 by trans_TC/ qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/predsnstar_5.ma".
-include "basic_2/reduction/lpx.ma".
-include "basic_2/computation/lprs.ma".
-
-(* SN EXTENDED PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS *******************)
-
-definition lpxs: ∀h. sd h → relation3 genv lenv lenv ≝
- λh,o,G. TC … (lpx h o G).
-
-interpretation "extended parallel computation (local environment, sn variant)"
- 'PRedSnStar h o G L1 L2 = (lpxs h o G L1 L2).
-
-(* Basic eliminators ********************************************************)
-
-lemma lpxs_ind: ∀h,o,G,L1. ∀R:predicate lenv. R L1 →
- (∀L,L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L → ⦃G, L⦄ ⊢ ➡[h, o] L2 → R L → R L2) →
- ∀L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → R L2.
-#h #o #G #L1 #R #HL1 #IHL1 #L2 #HL12
-@(TC_star_ind … HL1 IHL1 … HL12) //
-qed-.
-
-lemma lpxs_ind_dx: ∀h,o,G,L2. ∀R:predicate lenv. R L2 →
- (∀L1,L. ⦃G, L1⦄ ⊢ ➡[h, o] L → ⦃G, L⦄ ⊢ ➡*[h, o] L2 → R L → R L1) →
- ∀L1. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → R L1.
-#h #o #G #L2 #R #HL2 #IHL2 #L1 #HL12
-@(TC_star_ind_dx … HL2 IHL2 … HL12) //
-qed-.
-
-(* Basic properties *********************************************************)
-
-lemma lprs_lpxs: ∀h,o,G,L1,L2. ⦃G, L1⦄ ⊢ ➡* L2 → ⦃G, L1⦄ ⊢ ➡*[h, o] L2.
-/3 width=3 by lpr_lpx, monotonic_TC/ qed.
-
-lemma lpx_lpxs: ∀h,o,G,L1,L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 → ⦃G, L1⦄ ⊢ ➡*[h, o] L2.
-/2 width=1 by inj/ qed.
-
-lemma lpxs_refl: ∀h,o,G,L. ⦃G, L⦄ ⊢ ➡*[h, o] L.
-/2 width=1 by lprs_lpxs/ qed.
-
-lemma lpxs_strap1: ∀h,o,G,L1,L,L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L → ⦃G, L⦄ ⊢ ➡[h, o] L2 → ⦃G, L1⦄ ⊢ ➡*[h, o] L2.
-/2 width=3 by step/ qed.
-
-lemma lpxs_strap2: ∀h,o,G,L1,L,L2. ⦃G, L1⦄ ⊢ ➡[h, o] L → ⦃G, L⦄ ⊢ ➡*[h, o] L2 → ⦃G, L1⦄ ⊢ ➡*[h, o] L2.
-/2 width=3 by TC_strap/ qed.
-
-lemma lpxs_pair_refl: ∀h,o,G,L1,L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → ∀I,V. ⦃G, L1.ⓑ{I}V⦄ ⊢ ➡*[h, o] L2.ⓑ{I}V.
-/2 width=1 by TC_lpx_sn_pair_refl/ qed.
-
-(* Basic inversion lemmas ***************************************************)
-
-lemma lpxs_inv_atom1: ∀h,o,G,L2. ⦃G, ⋆⦄ ⊢ ➡*[h, o] L2 → L2 = ⋆.
-/2 width=2 by TC_lpx_sn_inv_atom1/ qed-.
-
-lemma lpxs_inv_atom2: ∀h,o,G,L1. ⦃G, L1⦄ ⊢ ➡*[h, o] ⋆ → L1 = ⋆.
-/2 width=2 by TC_lpx_sn_inv_atom2/ qed-.
-
-(* Basic forward lemmas *****************************************************)
-
-lemma lpxs_fwd_length: ∀h,o,G,L1,L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → |L1| = |L2|.
-/2 width=2 by TC_lpx_sn_fwd_length/ qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/reduction/lpx_aaa.ma".
-include "basic_2/computation/lpxs.ma".
-
-(* SN EXTENDED PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS *******************)
-
-(* Properties about atomic arity assignment on terms ************************)
-
-lemma lpxs_aaa_conf: ∀h,o,G,L1,T,A. ⦃G, L1⦄ ⊢ T ⁝ A →
- ∀L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → ⦃G, L2⦄ ⊢ T ⁝ A.
-#h #o #G #L1 #T #A #HT #L2 #HL12
-@(TC_Conf3 … (λL,A. ⦃G, L⦄ ⊢ T ⁝ A) … HT ? HL12) /2 width=5 by lpx_aaa_conf/
-qed-.
-
-lemma lprs_aaa_conf: ∀G,L1,T,A. ⦃G, L1⦄ ⊢ T ⁝ A →
- ∀L2. ⦃G, L1⦄ ⊢ ➡* L2 → ⦃G, L2⦄ ⊢ T ⁝ A.
-/3 width=5 by lprs_lpxs, lpxs_aaa_conf/ qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/computation/cpxs_cpxs.ma".
-include "basic_2/computation/lpxs.ma".
-
-(* SN EXTENDED PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS *******************)
-
-(* Advanced properties ******************************************************)
-
-lemma lpxs_pair: ∀h,o,I,G,L1,L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 →
- ∀V1,V2. ⦃G, L1⦄ ⊢ V1 ➡*[h, o] V2 →
- ⦃G, L1.ⓑ{I}V1⦄ ⊢ ➡*[h, o] L2.ⓑ{I}V2.
-/2 width=1 by TC_lpx_sn_pair/ qed.
-
-(* Advanced inversion lemmas ************************************************)
-
-lemma lpxs_inv_pair1: ∀h,o,I,G,K1,L2,V1. ⦃G, K1.ⓑ{I}V1⦄ ⊢ ➡*[h, o] L2 →
- ∃∃K2,V2. ⦃G, K1⦄ ⊢ ➡*[h, o] K2 & ⦃G, K1⦄ ⊢ V1 ➡*[h, o] V2 & L2 = K2.ⓑ{I}V2.
-/3 width=3 by TC_lpx_sn_inv_pair1, lpx_cpxs_trans/ qed-.
-
-lemma lpxs_inv_pair2: ∀h,o,I,G,L1,K2,V2. ⦃G, L1⦄ ⊢ ➡*[h, o] K2.ⓑ{I}V2 →
- ∃∃K1,V1. ⦃G, K1⦄ ⊢ ➡*[h, o] K2 & ⦃G, K1⦄ ⊢ V1 ➡*[h, o] V2 & L1 = K1.ⓑ{I}V1.
-/3 width=3 by TC_lpx_sn_inv_pair2, lpx_cpxs_trans/ qed-.
-
-(* Advanced eliminators *****************************************************)
-
-lemma lpxs_ind_alt: ∀h,o,G. ∀R:relation lenv.
- R (⋆) (⋆) → (
- ∀I,K1,K2,V1,V2.
- ⦃G, K1⦄ ⊢ ➡*[h, o] K2 → ⦃G, K1⦄ ⊢ V1 ➡*[h, o] V2 →
- R K1 K2 → R (K1.ⓑ{I}V1) (K2.ⓑ{I}V2)
- ) →
- ∀L1,L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → R L1 L2.
-/3 width=4 by TC_lpx_sn_ind, lpx_cpxs_trans/ qed-.
-
-(* Properties on context-sensitive extended parallel computation for terms **)
-
-lemma lpxs_cpx_trans: ∀h,o,G. c_r_transitive … (cpx h o G) (λ_.lpxs h o G).
-/3 width=5 by c_r_trans_LTC2, lpx_cpxs_trans/ qed-.
-
-(* Note: alternative proof: /3 width=5 by s_r_trans_TC1, lpxs_cpx_trans/ *)
-lemma lpxs_cpxs_trans: ∀h,o,G. c_rs_transitive … (cpx h o G) (λ_.lpxs h o G).
-#h #o #G @c_r_to_c_rs_trans @c_r_trans_LTC2
-@c_rs_trans_TC1 /2 width=3 by lpx_cpxs_trans/ (**) (* full auto too slow *)
-qed-.
-
-lemma cpxs_bind2: ∀h,o,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 →
- ∀I,T1,T2. ⦃G, L.ⓑ{I}V2⦄ ⊢ T1 ➡*[h, o] T2 →
- ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡*[h, o] ⓑ{a,I}V2.T2.
-/4 width=5 by lpxs_cpxs_trans, lpxs_pair, cpxs_bind/ qed.
-
-(* Inversion lemmas on context-sensitive ext parallel computation for terms *)
-
-lemma cpxs_inv_abst1: ∀h,o,a,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓛ{a}V1.T1 ➡*[h, o] U2 →
- ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 & ⦃G, L.ⓛV1⦄ ⊢ T1 ➡*[h, o] T2 &
- U2 = ⓛ{a}V2.T2.
-#h #o #a #G #L #V1 #T1 #U2 #H @(cpxs_ind … H) -U2 /2 width=5 by ex3_2_intro/
-#U0 #U2 #_ #HU02 * #V0 #T0 #HV10 #HT10 #H destruct
-elim (cpx_inv_abst1 … HU02) -HU02 #V2 #T2 #HV02 #HT02 #H destruct
-lapply (lpxs_cpx_trans … HT02 (L.ⓛV1) ?)
-/3 width=5 by lpxs_pair, cpxs_trans, cpxs_strap1, ex3_2_intro/
-qed-.
-
-lemma cpxs_inv_abbr1: ∀h,o,a,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓓ{a}V1.T1 ➡*[h, o] U2 → (
- ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 & ⦃G, L.ⓓV1⦄ ⊢ T1 ➡*[h, o] T2 &
- U2 = ⓓ{a}V2.T2
- ) ∨
- ∃∃T2. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡*[h, o] T2 & ⬆[0, 1] U2 ≡ T2 & a = true.
-#h #o #a #G #L #V1 #T1 #U2 #H @(cpxs_ind … H) -U2 /3 width=5 by ex3_2_intro, or_introl/
-#U0 #U2 #_ #HU02 * *
-[ #V0 #T0 #HV10 #HT10 #H destruct
- elim (cpx_inv_abbr1 … HU02) -HU02 *
- [ #V2 #T2 #HV02 #HT02 #H destruct
- lapply (lpxs_cpx_trans … HT02 (L.ⓓV1) ?)
- /4 width=5 by lpxs_pair, cpxs_trans, cpxs_strap1, ex3_2_intro, or_introl/
- | #T2 #HT02 #HUT2
- lapply (lpxs_cpx_trans … HT02 (L.ⓓV1) ?) -HT02
- /4 width=3 by lpxs_pair, cpxs_trans, ex3_intro, or_intror/
- ]
-| #U1 #HTU1 #HU01
- elim (lift_total U2 0 1) #U #HU2
- /6 width=12 by cpxs_strap1, cpx_lift, drop_drop, ex3_intro, or_intror/
-]
-qed-.
-
-(* More advanced properties *************************************************)
-
-lemma lpxs_pair2: ∀h,o,I,G,L1,L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 →
- ∀V1,V2. ⦃G, L2⦄ ⊢ V1 ➡*[h, o] V2 → ⦃G, L1.ⓑ{I}V1⦄ ⊢ ➡*[h, o] L2.ⓑ{I}V2.
-/3 width=3 by lpxs_pair, lpxs_cpxs_trans/ qed.
-
-(* Properties on supclosure *************************************************)
-
-lemma lpx_fqup_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
- ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 →
- ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡*[h, o] T & ⦃G1, K1, T⦄ ⊐+ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
-[ #G2 #L2 #T2 #H12 #K1 #HKL1 elim (lpx_fqu_trans … H12 … HKL1) -L1
- /3 width=5 by cpx_cpxs, fqu_fqup, ex3_2_intro/
-| #G #G2 #L #L2 #T #T2 #_ #H2 #IH1 #K1 #HLK1 elim (IH1 … HLK1) -L1
- #L0 #T0 #HT10 #HT0 #HL0 elim (lpx_fqu_trans … H2 … HL0) -L
- #L #T3 #HT3 #HT32 #HL2 elim (fqup_cpx_trans … HT0 … HT3) -T
- /3 width=7 by cpxs_strap1, fqup_strap1, ex3_2_intro/
-]
-qed-.
-
-lemma lpx_fqus_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
- ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 →
- ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡*[h, o] T & ⦃G1, K1, T⦄ ⊐* ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqus_ind … H) -G2 -L2 -T2 [ /2 width=5 by ex3_2_intro/ ]
-#G #G2 #L #L2 #T #T2 #_ #H2 #IH1 #K1 #HLK1 elim (IH1 … HLK1) -L1
-#L0 #T0 #HT10 #HT0 #HL0 elim (lpx_fquq_trans … H2 … HL0) -L
-#L #T3 #HT3 #HT32 #HL2 elim (fqus_cpx_trans … HT0 … HT3) -T
-/3 width=7 by cpxs_strap1, fqus_strap1, ex3_2_intro/
-qed-.
-
-lemma lpxs_fquq_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ∀K1. ⦃G1, K1⦄ ⊢ ➡*[h, o] L1 →
- ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡*[h, o] T & ⦃G1, K1, T⦄ ⊐⸮ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡*[h, o] L2.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #HT12 #K1 #H @(lpxs_ind_dx … H) -K1
-[ /2 width=5 by ex3_2_intro/
-| #K1 #K #HK1 #_ * #L #T #HT1 #HT2 #HL2 -HT12
- lapply (lpx_cpxs_trans … HT1 … HK1) -HT1
- elim (lpx_fquq_trans … HT2 … HK1) -K
- /3 width=7 by lpxs_strap2, cpxs_strap1, ex3_2_intro/
-]
-qed-.
-
-lemma lpxs_fqup_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
- ∀K1. ⦃G1, K1⦄ ⊢ ➡*[h, o] L1 →
- ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡*[h, o] T & ⦃G1, K1, T⦄ ⊐+ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡*[h, o] L2.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #HT12 #K1 #H @(lpxs_ind_dx … H) -K1
-[ /2 width=5 by ex3_2_intro/
-| #K1 #K #HK1 #_ * #L #T #HT1 #HT2 #HL2 -HT12
- lapply (lpx_cpxs_trans … HT1 … HK1) -HT1
- elim (lpx_fqup_trans … HT2 … HK1) -K
- /3 width=7 by lpxs_strap2, cpxs_trans, ex3_2_intro/
-]
-qed-.
-
-lemma lpxs_fqus_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
- ∀K1. ⦃G1, K1⦄ ⊢ ➡*[h, o] L1 →
- ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡*[h, o] T & ⦃G1, K1, T⦄ ⊐* ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡*[h, o] L2.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqus_ind … H) -G2 -L2 -T2 /2 width=5 by ex3_2_intro/
-#G #G2 #L #L2 #T #T2 #_ #H2 #IH1 #K1 #HLK1 elim (IH1 … HLK1) -L1
-#L0 #T0 #HT10 #HT0 #HL0 elim (lpxs_fquq_trans … H2 … HL0) -L
-#L #T3 #HT3 #HT32 #HL2 elim (fqus_cpxs_trans … HT3 … HT0) -T
-/3 width=7 by cpxs_trans, fqus_strap1, ex3_2_intro/
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/reduction/lpx_drop.ma".
-include "basic_2/computation/lpxs.ma".
-
-(* SN EXTENDED PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS *******************)
-
-(* Properties on local environment slicing ***********************************)
-
-lemma lpxs_drop_conf: ∀h,o,G. dropable_sn (lpxs h o G).
-/3 width=3 by dropable_sn_TC, lpx_drop_conf/ qed-.
-
-lemma drop_lpxs_trans: ∀h,o,G. dedropable_sn (lpxs h o G).
-/3 width=3 by dedropable_sn_TC, drop_lpx_trans/ qed-.
-
-lemma lpxs_drop_trans_O1: ∀h,o,G. dropable_dx (lpxs h o G).
-/3 width=3 by dropable_dx_TC, lpx_drop_trans_O1/ qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/multiple/lleq_lleq.ma".
-include "basic_2/reduction/lpx_lleq.ma".
-include "basic_2/computation/cpxs_lreq.ma".
-include "basic_2/computation/lpxs_drop.ma".
-include "basic_2/computation/lpxs_cpxs.ma".
-
-(* SN EXTENDED PARALLEL COMPUTATION FOR LOCAL ENVIRONMENTS ******************)
-
-(* Properties on lazy equivalence for local environments ********************)
-
-lemma lleq_lpxs_trans: ∀h,o,G,L2,K2. ⦃G, L2⦄ ⊢ ➡*[h, o] K2 →
- ∀L1,T,l. L1 ≡[T, l] L2 →
- ∃∃K1. ⦃G, L1⦄ ⊢ ➡*[h, o] K1 & K1 ≡[T, l] K2.
-#h #o #G #L2 #K2 #H @(lpxs_ind … H) -K2 /2 width=3 by ex2_intro/
-#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,o,G,L1,L2,T,l. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → (L1 ≡[T, l] L2 → ⊥) →
- ∃∃L,L0. ⦃G, L1⦄ ⊢ ➡[h, o] L & L1 ≡[T, l] L → ⊥ & ⦃G, L⦄ ⊢ ➡*[h, o] L0 & L0 ≡[T, l] L2.
-#h #o #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 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
- #L #H1 #H lapply (nlleq_lleq_div … H … H2) -H2
- #H2 elim (lleq_lpxs_trans … H3 … H) -L0
- /3 width=8 by lleq_trans, ex4_2_intro/
- | -H12 -IH2 /3 width=6 by ex4_2_intro/
- ]
-]
-qed-.
-
-lemma lpxs_lleq_fqu_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
- ∀K1. ⦃G1, K1⦄ ⊢ ➡*[h, o] L1 → K1 ≡[T1, 0] L1 →
- ∃∃K2. ⦃G1, K1, T1⦄ ⊐ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡*[h, o] L2 & K2 ≡[T2, 0] L2.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
-[ #I #G1 #L1 #V1 #X #H1 #H2 elim (lpxs_inv_pair2 … H1) -H1
- #K0 #V0 #H1KL1 #_ #H destruct
- elim (lleq_inv_lref_ge_dx … H2 ? I L1 V1) -H2 //
- #K1 #H #H2KL1 lapply (drop_inv_O2 … H) -H #H destruct
- /2 width=4 by fqu_lref_O, ex3_intro/
-| * [ #a ] #I #G1 #L1 #V1 #T1 #K1 #HLK1 #H
- [ elim (lleq_inv_bind … H)
- | elim (lleq_inv_flat … H)
- ] -H /2 width=4 by fqu_pair_sn, ex3_intro/
-| #a #I #G1 #L1 #V1 #T1 #K1 #HLK1 #H elim (lleq_inv_bind_O … H) -H
- /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 #k #HL1 #HTU1 #K1 #H1KL1 #H2KL1
- elim (drop_O1_le (Ⓕ) (k+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
- /3 width=4 by fqu_drop, ex3_intro/
- | lapply (drop_fwd_length_le2 … HL1) -L -T1 -o
- lapply (lleq_fwd_length … H2KL1) //
- ]
-]
-qed-.
-
-lemma lpxs_lleq_fquq_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ∀K1. ⦃G1, K1⦄ ⊢ ➡*[h, o] L1 → K1 ≡[T1, 0] L1 →
- ∃∃K2. ⦃G1, K1, T1⦄ ⊐⸮ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡*[h, o] L2 & K2 ≡[T2, 0] L2.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #K1 #H1KL1 #H2KL1
-elim (fquq_inv_gen … H) -H
-[ #H elim (lpxs_lleq_fqu_trans … H … H1KL1 H2KL1) -L1
- /3 width=4 by fqu_fquq, ex3_intro/
-| * #HG #HL #HT destruct /2 width=4 by ex3_intro/
-]
-qed-.
-
-lemma lpxs_lleq_fqup_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
- ∀K1. ⦃G1, K1⦄ ⊢ ➡*[h, o] L1 → K1 ≡[T1, 0] L1 →
- ∃∃K2. ⦃G1, K1, T1⦄ ⊐+ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡*[h, o] L2 & K2 ≡[T2, 0] L2.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
-[ #G2 #L2 #T2 #H #K1 #H1KL1 #H2KL1 elim (lpxs_lleq_fqu_trans … H … H1KL1 H2KL1) -L1
- /3 width=4 by fqu_fqup, ex3_intro/
-| #G #G2 #L #L2 #T #T2 #_ #HT2 #IHT1 #K1 #H1KL1 #H2KL1 elim (IHT1 … H2KL1) // -L1
- #K #HT1 #H1KL #H2KL elim (lpxs_lleq_fqu_trans … HT2 … H1KL H2KL) -L
- /3 width=5 by fqup_strap1, ex3_intro/
-]
-qed-.
-
-lemma lpxs_lleq_fqus_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
- ∀K1. ⦃G1, K1⦄ ⊢ ➡*[h, o] L1 → K1 ≡[T1, 0] L1 →
- ∃∃K2. ⦃G1, K1, T1⦄ ⊐* ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡*[h, o] L2 & K2 ≡[T2, 0] L2.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #K1 #H1KL1 #H2KL1
-elim (fqus_inv_gen … H) -H
-[ #H elim (lpxs_lleq_fqup_trans … H … H1KL1 H2KL1) -L1
- /3 width=4 by fqup_fqus, ex3_intro/
-| * #HG #HL #HT destruct /2 width=4 by ex3_intro/
-]
-qed-.
-
-fact lreq_lpxs_trans_lleq_aux: ∀h,o,G,L1,L0,l,k. L1 ⩬[l, k] L0 → k = ∞ →
- ∀L2. ⦃G, L0⦄ ⊢ ➡*[h, o] L2 →
- ∃∃L. L ⩬[l, k] L2 & ⦃G, L1⦄ ⊢ ➡*[h, o] L &
- (∀T. L0 ≡[T, l] L2 ↔ L1 ≡[T, l] L).
-#h #o #G #L1 #L0 #l #k #H elim H -L1 -L0 -l -k
-[ #l #k #_ #L2 #H >(lpxs_inv_atom1 … H) -H
- /3 width=5 by ex3_intro, conj/
-| #I1 #I0 #L1 #L0 #V1 #V0 #_ #_ #Hm destruct
-| #I #L1 #L0 #V1 #k #HL10 #IHL10 #Hm #Y #H
- elim (lpxs_inv_pair1 … H) -H #L2 #V2 #HL02 #HV02 #H destruct
- 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, lreq_cpxs_trans, lreq_pair/
- #T elim (IH T) #HL0dx #HL0sn
- @conj #H @(lleq_lreq_repl … H) -H /3 width=1 by lreq_sym, lreq_pair_O_Y/
-| #I1 #I0 #L1 #L0 #V1 #V0 #l #k #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, lreq_succ/
- #T elim (IH T) #HL0dx #HL0sn
- @conj #H @(lleq_lreq_repl … H) -H /3 width=1 by lreq_sym, lreq_succ/
-]
-qed-.
-
-lemma lreq_lpxs_trans_lleq: ∀h,o,G,L1,L0,l. L1 ⩬[l, ∞] L0 →
- ∀L2. ⦃G, L0⦄ ⊢ ➡*[h, o] L2 →
- ∃∃L. L ⩬[l, ∞] L2 & ⦃G, L1⦄ ⊢ ➡*[h, o] L &
- (∀T. L0 ≡[T, l] L2 ↔ L1 ≡[T, l] L).
-/2 width=1 by lreq_lpxs_trans_lleq_aux/ qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/computation/lpxs.ma".
-
-(* SN EXTENDED PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS *******************)
-
-(* Main properties **********************************************************)
-
-theorem lpxs_trans: ∀h,o,G. Transitive … (lpxs h o G).
-/2 width=3 by trans_TC/ qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/lrsubeqc_4.ma".
-include "basic_2/static/lsubr.ma".
-include "basic_2/static/aaa.ma".
-include "basic_2/computation/gcp_cr.ma".
-
-(* LOCAL ENVIRONMENT REFINEMENT FOR GENERIC REDUCIBILITY ********************)
-
-inductive lsubc (RP) (G): relation lenv ≝
-| lsubc_atom: lsubc RP G (⋆) (⋆)
-| lsubc_pair: ∀I,L1,L2,V. lsubc RP G L1 L2 → lsubc RP G (L1.ⓑ{I}V) (L2.ⓑ{I}V)
-| lsubc_beta: ∀L1,L2,V,W,A. ⦃G, L1, V⦄ ϵ[RP] 〚A〛 → ⦃G, L1, W⦄ ϵ[RP] 〚A〛 → ⦃G, L2⦄ ⊢ W ⁝ A →
- lsubc RP G L1 L2 → lsubc RP G (L1. ⓓⓝW.V) (L2.ⓛW)
-.
-
-interpretation
- "local environment refinement (generic reducibility)"
- 'LRSubEqC RP G L1 L2 = (lsubc RP G L1 L2).
-
-(* Basic inversion lemmas ***************************************************)
-
-fact lsubc_inv_atom1_aux: ∀RP,G,L1,L2. G ⊢ L1 ⫃[RP] L2 → L1 = ⋆ → L2 = ⋆.
-#RP #G #L1 #L2 * -L1 -L2
-[ //
-| #I #L1 #L2 #V #_ #H destruct
-| #L1 #L2 #V #W #A #_ #_ #_ #_ #H destruct
-]
-qed-.
-
-(* Basic_1: was just: csubc_gen_sort_r *)
-lemma lsubc_inv_atom1: ∀RP,G,L2. G ⊢ ⋆ ⫃[RP] L2 → L2 = ⋆.
-/2 width=5 by lsubc_inv_atom1_aux/ qed-.
-
-fact lsubc_inv_pair1_aux: ∀RP,G,L1,L2. G ⊢ L1 ⫃[RP] L2 → ∀I,K1,X. L1 = K1.ⓑ{I}X →
- (∃∃K2. G ⊢ K1 ⫃[RP] K2 & L2 = K2.ⓑ{I}X) ∨
- ∃∃K2,V,W,A. ⦃G, K1, V⦄ ϵ[RP] 〚A〛 & ⦃G, K1, W⦄ ϵ[RP] 〚A〛 & ⦃G, K2⦄ ⊢ W ⁝ A &
- G ⊢ K1 ⫃[RP] K2 &
- L2 = K2. ⓛW & X = ⓝW.V & I = Abbr.
-#RP #G #L1 #L2 * -L1 -L2
-[ #I #K1 #V #H destruct
-| #J #L1 #L2 #V #HL12 #I #K1 #W #H destruct /3 width=3 by ex2_intro, or_introl/
-| #L1 #L2 #V1 #W2 #A #HV1 #H1W2 #H2W2 #HL12 #I #K1 #V #H destruct /3 width=10 by ex7_4_intro, or_intror/
-]
-qed-.
-
-(* Basic_1: was: csubc_gen_head_r *)
-lemma lsubc_inv_pair1: ∀RP,I,G,K1,L2,X. G ⊢ K1.ⓑ{I}X ⫃[RP] L2 →
- (∃∃K2. G ⊢ K1 ⫃[RP] K2 & L2 = K2.ⓑ{I}X) ∨
- ∃∃K2,V,W,A. ⦃G, K1, V⦄ ϵ[RP] 〚A〛 & ⦃G, K1, W⦄ ϵ[RP] 〚A〛 & ⦃G, K2⦄ ⊢ W ⁝ A &
- G ⊢ K1 ⫃[RP] K2 &
- L2 = K2.ⓛW & X = ⓝW.V & I = Abbr.
-/2 width=3 by lsubc_inv_pair1_aux/ qed-.
-
-fact lsubc_inv_atom2_aux: ∀RP,G,L1,L2. G ⊢ L1 ⫃[RP] L2 → L2 = ⋆ → L1 = ⋆.
-#RP #G #L1 #L2 * -L1 -L2
-[ //
-| #I #L1 #L2 #V #_ #H destruct
-| #L1 #L2 #V #W #A #_ #_ #_ #_ #H destruct
-]
-qed-.
-
-(* Basic_1: was just: csubc_gen_sort_l *)
-lemma lsubc_inv_atom2: ∀RP,G,L1. G ⊢ L1 ⫃[RP] ⋆ → L1 = ⋆.
-/2 width=5 by lsubc_inv_atom2_aux/ qed-.
-
-fact lsubc_inv_pair2_aux: ∀RP,G,L1,L2. G ⊢ L1 ⫃[RP] L2 → ∀I,K2,W. L2 = K2.ⓑ{I} W →
- (∃∃K1. G ⊢ K1 ⫃[RP] K2 & L1 = K1. ⓑ{I} W) ∨
- ∃∃K1,V,A. ⦃G, K1, V⦄ ϵ[RP] 〚A〛 & ⦃G, K1, W⦄ ϵ[RP] 〚A〛 & ⦃G, K2⦄ ⊢ W ⁝ A &
- G ⊢ K1 ⫃[RP] K2 &
- L1 = K1.ⓓⓝW.V & I = Abst.
-#RP #G #L1 #L2 * -L1 -L2
-[ #I #K2 #W #H destruct
-| #J #L1 #L2 #V #HL12 #I #K2 #W #H destruct /3 width=3 by ex2_intro, or_introl/
-| #L1 #L2 #V1 #W2 #A #HV1 #H1W2 #H2W2 #HL12 #I #K2 #W #H destruct /3 width=8 by ex6_3_intro, or_intror/
-]
-qed-.
-
-(* Basic_1: was just: csubc_gen_head_l *)
-lemma lsubc_inv_pair2: ∀RP,I,G,L1,K2,W. G ⊢ L1 ⫃[RP] K2.ⓑ{I} W →
- (∃∃K1. G ⊢ K1 ⫃[RP] K2 & L1 = K1.ⓑ{I} W) ∨
- ∃∃K1,V,A. ⦃G, K1, V⦄ ϵ[RP] 〚A〛 & ⦃G, K1, W⦄ ϵ[RP] 〚A〛 & ⦃G, K2⦄ ⊢ W ⁝ A &
- G ⊢ K1 ⫃[RP] K2 &
- L1 = K1.ⓓⓝW.V & I = Abst.
-/2 width=3 by lsubc_inv_pair2_aux/ qed-.
-
-(* Basic forward lemmas *****************************************************)
-
-lemma lsubc_fwd_lsubr: ∀RP,G,L1,L2. G ⊢ L1 ⫃[RP] L2 → L1 ⫃ L2.
-#RP #G #L1 #L2 #H elim H -L1 -L2 /2 width=1 by lsubr_pair, lsubr_beta/
-qed-.
-
-(* Basic properties *********************************************************)
-
-(* Basic_1: was just: csubc_refl *)
-lemma lsubc_refl: ∀RP,G,L. G ⊢ L ⫃[RP] L.
-#RP #G #L elim L -L /2 width=1 by lsubc_pair/
-qed.
-
-(* Basic_1: removed theorems 3:
- csubc_clear_conf csubc_getl_conf csubc_csuba
-*)
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/static/aaa_lift.ma".
-include "basic_2/computation/lsubc.ma".
-
-(* LOCAL ENVIRONMENT REFINEMENT FOR GENERIC REDUCIBILITY ********************)
-
-(* Properties concerning basic local environment slicing ********************)
-
-(* Basic_1: was: csubc_drop_conf_O *)
-(* Note: the constant 0 can not be generalized *)
-lemma lsubc_drop_O1_trans: ∀RP,G,L1,L2. G ⊢ L1 ⫃[RP] L2 → ∀K2,c,k. ⬇[c, 0, k] L2 ≡ K2 →
- ∃∃K1. ⬇[c, 0, k] L1 ≡ K1 & G ⊢ K1 ⫃[RP] K2.
-#RP #G #L1 #L2 #H elim H -L1 -L2
-[ #X #c #k #H elim (drop_inv_atom1 … H) -H /4 width=3 by drop_atom, ex2_intro/
-| #I #L1 #L2 #V #_ #IHL12 #X #c #k #H
- elim (drop_inv_O1_pair1 … H) -H * #Hm #H destruct
- [ elim (IHL12 L2 c 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 #c #k #H
- elim (drop_inv_O1_pair1 … H) -H * #Hm #H destruct
- [ elim (IHL12 L2 c 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/
- ]
-]
-qed-.
-
-(* Basic_1: was: csubc_drop_conf_rev *)
-lemma drop_lsubc_trans: ∀RR,RS,RP. gcp RR RS RP →
- ∀G,L1,K1,l,k. ⬇[Ⓕ, l, k] L1 ≡ K1 → ∀K2. G ⊢ K1 ⫃[RP] K2 →
- ∃∃L2. G ⊢ L1 ⫃[RP] L2 & ⬇[Ⓕ, l, k] L2 ≡ K2.
-#RR #RS #RP #Hgcp #G #L1 #K1 #l #k #H elim H -L1 -K1 -l -k
-[ #l #k #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 #k #_ #IHLK1 #K2 #HK12
- elim (IHLK1 … HK12) -K1 /3 width=5 by lsubc_pair, drop_drop, ex2_intro/
-| #I #L1 #K1 #V1 #V2 #l #k #HLK1 #HV21 #IHLK1 #X #H
- elim (lsubc_inv_pair1 … H) -H *
- [ #K2 #HK12 #H destruct
- elim (IHLK1 … HK12) -K1 /3 width=5 by lsubc_pair, drop_skip, ex2_intro/
- | #K2 #V #W2 #A #HV2 #H1W2 #H2W2 #HK12 #H1 #H2 #H3 destruct
- elim (lift_inv_flat1 … HV21) -HV21 #W3 #V3 #HW23 #HV3 #H destruct
- elim (IHLK1 … HK12) #K #HL1K #HK2
- lapply (gcr_lift … Hgcp … HV2 … HLK1 … HV3) -HV2
- lapply (gcr_lift … Hgcp … H1W2 … HLK1 … HW23) -H1W2
- /4 width=11 by lsubc_beta, aaa_lift, drop_skip, ex2_intro/
- ]
-]
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/computation/lsubc_drop.ma".
-
-(* LOCAL ENVIRONMENT REFINEMENT FOR GENERIC REDUCIBILITY ********************)
-
-(* Properties concerning generic local environment slicing ******************)
-
-(* Basic_1: was: csubc_drop1_conf_rev *)
-lemma drops_lsubc_trans: ∀RR,RS,RP. gcp RR RS RP →
- ∀G,L1,K1,cs. ⬇*[Ⓕ, cs] L1 ≡ K1 → ∀K2. G ⊢ K1 ⫃[RP] K2 →
- ∃∃L2. G ⊢ L1 ⫃[RP] L2 & ⬇*[Ⓕ, cs] L2 ≡ K2.
-#RR #RS #RP #Hgcp #G #L1 #K1 #cs #H elim H -L1 -K1 -cs
-[ /2 width=3 by drops_nil, ex2_intro/
-| #L1 #L #K1 #cs #l #k #_ #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/
-]
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/static/lsuba.ma".
-include "basic_2/computation/gcp_aaa.ma".
-
-(* LOCAL ENVIRONMENT REFINEMENT FOR GENERIC REDUCIBILITY ********************)
-
-(* properties concerning lenv refinement for atomic arity assignment ********)
-
-lemma lsuba_lsubc: ∀RR,RS,RP. gcp RR RS RP → gcr RR RS RP RP →
- ∀G,L1,L2. G ⊢ L1 ⫃⁝ L2 → G ⊢ L1 ⫃[RP] L2.
-#RR #RS #RP #H1RP #H2RP #G #L1 #L2 #H elim H -L1 -L2 /2 width=1 by lsubc_pair/
-#L1 #L2 #W #V #A #H elim (aaa_inv_cast … H) -H /3 width=4 by acr_aaa, lsubc_beta/
-qed.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/sn_6.ma".
-include "basic_2/multiple/lleq.ma".
-include "basic_2/reduction/lpx.ma".
-
-(* SN EXTENDED STRONGLY NORMALIZING LOCAL ENVIRONMENTS **********************)
-
-definition lsx: ∀h. sd h → relation4 ynat term genv lenv ≝
- λh,o,l,T,G. SN … (lpx h o G) (lleq l T).
-
-interpretation
- "extended strong normalization (local environment)"
- 'SN h o l T G L = (lsx h o T l G L).
-
-(* Basic eliminators ********************************************************)
-
-lemma lsx_ind: ∀h,o,G,T,l. ∀R:predicate lenv.
- (∀L1. G ⊢ ⬊*[h, o, T, l] L1 →
- (∀L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 → (L1 ≡[T, l] L2 → ⊥) → R L2) →
- R L1
- ) →
- ∀L. G ⊢ ⬊*[h, o, T, l] L → R L.
-#h #o #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,o,G,L1,T,l.
- (∀L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 → (L1 ≡[T, l] L2 → ⊥) → G ⊢ ⬊*[h, o, T, l] L2) →
- G ⊢ ⬊*[h, o, T, l] L1.
-/5 width=1 by lleq_sym, SN_intro/ qed.
-
-lemma lsx_atom: ∀h,o,G,T,l. G ⊢ ⬊*[h, o, T, l] ⋆.
-#h #o #G #T #l @lsx_intro
-#X #H #HT lapply (lpx_inv_atom1 … H) -H
-#H destruct elim HT -HT //
-qed.
-
-lemma lsx_sort: ∀h,o,G,L,l,s. G ⊢ ⬊*[h, o, ⋆s, l] L.
-#h #o #G #L1 #l #s @lsx_intro
-#L2 #HL12 #H elim H -H
-/3 width=4 by lpx_fwd_length, lleq_sort/
-qed.
-
-lemma lsx_gref: ∀h,o,G,L,l,p. G ⊢ ⬊*[h, o, §p, l] L.
-#h #o #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,o,G,L,T,U,lt,l,k. lt ≤ yinj l + yinj k →
- ⬆[l, k] T ≡ U → G ⊢ ⬊*[h, o, U, lt] L → G ⊢ ⬊*[h, o, U, l] L.
-#h #o #G #L #T #U #lt #l #k #Hltlm #HTU #H @(lsx_ind … H) -L
-/5 width=7 by lsx_intro, lleq_ge_up/
-qed-.
-
-lemma lsx_ge: ∀h,o,G,L,T,l1,l2. l1 ≤ l2 →
- G ⊢ ⬊*[h, o, T, l1] L → G ⊢ ⬊*[h, o, T, l2] L.
-#h #o #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,o,a,I,G,L,V,T,l. G ⊢ ⬊*[h, o, ⓑ{a,I}V.T, l] L →
- G ⊢ ⬊*[h, o, V, l] L.
-#h #o #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,o,I,G,L,V,T,l. G ⊢ ⬊*[h, o, ⓕ{I}V.T, l] L →
- G ⊢ ⬊*[h, o, V, l] L.
-#h #o #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,o,I,G,L,V,T,l. G ⊢ ⬊*[h, o, ⓕ{I}V.T, l] L →
- G ⊢ ⬊*[h, o, T, l] L.
-#h #o #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,o,I,G,L,V,T,l. G ⊢ ⬊*[h, o, ②{I}V.T, l] L →
- G ⊢ ⬊*[h, o, V, l] L.
-#h #o * /2 width=4 by lsx_fwd_bind_sn, lsx_fwd_flat_sn/
-qed-.
-
-(* Basic inversion lemmas ***************************************************)
-
-lemma lsx_inv_flat: ∀h,o,I,G,L,V,T,l. G ⊢ ⬊*[h, o, ⓕ{I}V.T, l] L →
- G ⊢ ⬊*[h, o, V, l] L ∧ G ⊢ ⬊*[h, o, T, l] L.
-/3 width=3 by lsx_fwd_flat_sn, lsx_fwd_flat_dx, conj/ qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/snalt_6.ma".
-include "basic_2/computation/lpxs_lleq.ma".
-include "basic_2/computation/lsx.ma".
-
-(* SN EXTENDED STRONGLY NORMALIZING LOCAL ENVIRONMENTS **********************)
-
-(* alternative definition of lsx *)
-definition lsxa: ∀h. sd h → relation4 ynat term genv lenv ≝
- λh,o,l,T,G. SN … (lpxs h o G) (lleq l T).
-
-interpretation
- "extended strong normalization (local environment) alternative"
- 'SNAlt h o l T G L = (lsxa h o T l G L).
-
-(* Basic eliminators ********************************************************)
-
-lemma lsxa_ind: ∀h,o,G,T,l. ∀R:predicate lenv.
- (∀L1. G ⊢ ⬊⬊*[h, o, T, l] L1 →
- (∀L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → (L1 ≡[T, l] L2 → ⊥) → R L2) →
- R L1
- ) →
- ∀L. G ⊢ ⬊⬊*[h, o, T, l] L → R L.
-#h #o #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,o,G,L1,T,l.
- (∀L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → (L1 ≡[T, l] L2 → ⊥) → G ⊢ ⬊⬊*[h, o, T, l] L2) →
- G ⊢ ⬊⬊*[h, o, T, l] L1.
-/5 width=1 by lleq_sym, SN_intro/ qed.
-
-fact lsxa_intro_aux: ∀h,o,G,L1,T,l.
- (∀L,L2. ⦃G, L⦄ ⊢ ➡*[h, o] L2 → L1 ≡[T, l] L → (L1 ≡[T, l] L2 → ⊥) → G ⊢ ⬊⬊*[h, o, T, l] L2) →
- G ⊢ ⬊⬊*[h, o, T, l] L1.
-/4 width=3 by lsxa_intro/ qed-.
-
-lemma lsxa_lleq_trans: ∀h,o,T,G,L1,l. G ⊢ ⬊⬊*[h, o, T, l] L1 →
- ∀L2. L1 ≡[T, l] L2 → G ⊢ ⬊⬊*[h, o, T, l] L2.
-#h #o #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,o,T,G,L1,l. G ⊢ ⬊⬊*[h, o, T, l] L1 →
- ∀L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → G ⊢ ⬊⬊*[h, o, T, l] L2.
-#h #o #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,o,G,L1,T,l.
- (∀L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 → (L1 ≡[T, l] L2 → ⊥) → G ⊢ ⬊⬊*[h, o, T, l] L2) →
- G ⊢ ⬊⬊*[h, o, T, l] L1.
-#h #o #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 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/
-]
-qed-.
-
-(* Main properties **********************************************************)
-
-theorem lsx_lsxa: ∀h,o,G,L,T,l. G ⊢ ⬊*[h, o, T, l] L → G ⊢ ⬊⬊*[h, o, T, l] L.
-#h #o #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,o,G,L,T,l. G ⊢ ⬊⬊*[h, o, T, l] L → G ⊢ ⬊*[h, o, T, l] L.
-#h #o #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,o,G,L1,T,l.
- (∀L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → (L1 ≡[T, l] L2 → ⊥) → G ⊢ ⬊*[h, o, T, l] L2) →
- G ⊢ ⬊*[h, o, T, l] L1.
-/6 width=1 by lsxa_inv_lsx, lsx_lsxa, lsxa_intro/ qed.
-
-lemma lsx_lpxs_trans: ∀h,o,G,L1,T,l. G ⊢ ⬊*[h, o, T, l] L1 →
- ∀L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → G ⊢ ⬊*[h, o, T, l] L2.
-/4 width=3 by lsxa_inv_lsx, lsx_lsxa, lsxa_lpxs_trans/ qed-.
-
-(* Advanced eliminators *****************************************************)
-
-lemma lsx_ind_alt: ∀h,o,G,T,l. ∀R:predicate lenv.
- (∀L1. G ⊢ ⬊*[h, o, T, l] L1 →
- (∀L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → (L1 ≡[T, l] L2 → ⊥) → R L2) →
- R L1
- ) →
- ∀L. G ⊢ ⬊*[h, o, T, l] L → R L.
-#h #o #G #T #l #R #IH #L #H @(lsxa_ind h o G T l … L)
-/4 width=1 by lsxa_inv_lsx, lsx_lsxa/
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/computation/csx_lpxs.ma".
-include "basic_2/computation/lcosx_cpx.ma".
-
-(* SN EXTENDED STRONGLY NORMALIZING LOCAL ENVIRONMENTS **********************)
-
-(* Advanced properties ******************************************************)
-
-lemma lsx_lref_be_lpxs: ∀h,o,I,G,K1,V,i,l. l ≤ yinj i → ⦃G, K1⦄ ⊢ ⬊*[h, o] V →
- ∀K2. G ⊢ ⬊*[h, o, V, 0] K2 → ⦃G, K1⦄ ⊢ ➡*[h, o] K2 →
- ∀L2. ⬇[i] L2 ≡ K2.ⓑ{I}V → G ⊢ ⬊*[h, o, #i, l] L2.
-#h #o #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
-#Y #H #HLK2 elim (lpx_inv_pair1 … H) -H
-#K2 #V2 #HK02 #HV02 #H destruct
-elim (eq_term_dec V0 V2) #HnV02 destruct [ -IHV0 -HV02 -HK0 | -IHK0 -HnL02 -HLK0 ]
-[ /4 width=8 by lpxs_strap1, lleq_lref/
-| @(IHV0 … HnV02 … HLK2) -IHV0 -HnV02 -HLK2
- /3 width=4 by lsx_cpx_trans_O, lsx_lpx_trans, lpxs_cpx_trans, lpxs_strap1/ (**) (* full auto too slow *)
-]
-qed.
-
-lemma lsx_lref_be: ∀h,o,I,G,K,V,i,l. l ≤ yinj i → ⦃G, K⦄ ⊢ ⬊*[h, o] V →
- G ⊢ ⬊*[h, o, V, 0] K →
- ∀L. ⬇[i] L ≡ K.ⓑ{I}V → G ⊢ ⬊*[h, o, #i, l] L.
-/2 width=8 by lsx_lref_be_lpxs/ qed.
-
-(* Main properties **********************************************************)
-
-theorem csx_lsx: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → ∀l. G ⊢ ⬊*[h, o, T, l] L.
-#h #o #G #L #T @(fqup_wf_ind_eq … G L T) -G -L -T
-#Z #Y #X #IH #G #L * * //
-[ #i #HG #HL #HT #H #l destruct
- elim (lt_or_ge i (|L|)) /2 width=1 by lsx_lref_free/
- 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 #l destruct
- elim (csx_fwd_bind … H) -H /3 width=1 by lsx_bind/
-| #I #V #T #HG #HL #HT #H #l destruct
- elim (csx_fwd_flat … H) -H /3 width=1 by lsx_flat/
-]
-qed.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/multiple/lleq_drop.ma".
-include "basic_2/reduction/lpx_drop.ma".
-include "basic_2/computation/lsx.ma".
-
-(* SN EXTENDED STRONGLY NORMALIZING LOCAL ENVIRONMENTS **********************)
-
-(* Advanced properties ******************************************************)
-
-lemma lsx_lref_free: ∀h,o,G,L,l,i. |L| ≤ i → G ⊢ ⬊*[h, o, #i, l] L.
-#h #o #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,o,G,L,l,i. yinj i < l → G ⊢ ⬊*[h, o, #i, l] L.
-#h #o #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,o,I,G,L,l,i. l ≤ yinj i → G ⊢ ⬊*[h, o, #i, l] L →
- ∀K,V. ⬇[i] L ≡ K.ⓑ{I}V → G ⊢ ⬊*[h, o, V, 0] K.
-#h #o #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 (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/
-qed-.
-
-(* Properties on relocation *************************************************)
-
-lemma lsx_lift_le: ∀h,o,G,K,T,U,lt,l,k. lt ≤ yinj l →
- ⬆[l, k] T ≡ U → G ⊢ ⬊*[h, o, T, lt] K →
- ∀L. ⬇[Ⓕ, l, k] L ≡ K → G ⊢ ⬊*[h, o, U, lt] L.
-#h #o #G #K #T #U #lt #l #k #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,o,G,K,T,U,lt,l,k. yinj l ≤ lt →
- ⬆[l, k] T ≡ U → G ⊢ ⬊*[h, o, T, lt] K →
- ∀L. ⬇[Ⓕ, l, k] L ≡ K → G ⊢ ⬊*[h, o, U, lt + k] L.
-#h #o #G #K #T #U #lt #l #k #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/
-qed-.
-
-(* Inversion lemmas on relocation *******************************************)
-
-lemma lsx_inv_lift_le: ∀h,o,G,L,T,U,lt,l,k. lt ≤ yinj l →
- ⬆[l, k] T ≡ U → G ⊢ ⬊*[h, o, U, lt] L →
- ∀K. ⬇[Ⓕ, l, k] L ≡ K → G ⊢ ⬊*[h, o, T, lt] K.
-#h #o #G #L #T #U #lt #l #k #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,o,G,L,T,U,lt,l,k. yinj l ≤ lt → lt ≤ l + k →
- ⬆[l, k] T ≡ U → G ⊢ ⬊*[h, o, U, lt] L →
- ∀K. ⬇[Ⓕ, l, k] L ≡ K → G ⊢ ⬊*[h, o, T, l] K.
-#h #o #G #L #T #U #lt #l #k #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,o,G,L,T,U,lt,l,k. yinj l + yinj k ≤ lt →
- ⬆[l, k] T ≡ U → G ⊢ ⬊*[h, o, U, lt] L →
- ∀K. ⬇[Ⓕ, l, k] L ≡ K → G ⊢ ⬊*[h, o, T, lt-k] K.
-#h #o #G #L #T #U #lt #l #k #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/
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/multiple/lleq_lleq.ma".
-include "basic_2/reduction/lpx_lleq.ma".
-include "basic_2/computation/lsx.ma".
-
-(* SN EXTENDED STRONGLY NORMALIZING LOCAL ENVIRONMENTS **********************)
-
-(* Advanced properties ******************************************************)
-
-lemma lsx_lleq_trans: ∀h,o,T,G,L1,l. G ⊢ ⬊*[h, o, T, l] L1 →
- ∀L2. L1 ≡[T, l] L2 → G ⊢ ⬊*[h, o, T, l] L2.
-#h #o #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,o,T,G,L1,l. G ⊢ ⬊*[h, o, T, l] L1 →
- ∀L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 → G ⊢ ⬊*[h, o, T, l] L2.
-#h #o #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_lreq_conf: ∀h,o,G,L1,T,l. G ⊢ ⬊*[h, o, T, l] L1 →
- ∀L2. L1 ⩬[l, ∞] L2 → G ⊢ ⬊*[h, o, T, l] L2.
-#h #o #G #L1 #T #l #H @(lsx_ind … H) -L1
-#L1 #_ #IHL1 #L2 #HL12 @lsx_intro
-#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,o,a,I,G,L,V,T,l. G ⊢ ⬊*[h, o, ⓑ{a,I}V.T, l] L →
- G ⊢ ⬊*[h, o, T, ⫯l] L.ⓑ{I}V.
-#h #o #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_lreq_conf … (L2.ⓑ{I}V1)) /2 width=1 by lreq_succ/
-@IHL1 // #H @HT -IHL1 -HL12 -HT
-@(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,o,a,I,G,L,V,T,l. G ⊢ ⬊*[h, o, ⓑ{a, I}V.T, l] L →
- G ⊢ ⬊*[h, o, V, l] L ∧ G ⊢ ⬊*[h, o, T, ⫯l] L.ⓑ{I}V.
-/3 width=4 by lsx_fwd_bind_sn, lsx_fwd_bind_dx, conj/ qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/computation/lpxs_lpxs.ma".
-include "basic_2/computation/lsx_alt.ma".
-
-(* SN EXTENDED STRONGLY NORMALIZING LOCAL ENVIRONMENTS **********************)
-
-(* Advanced properties ******************************************************)
-
-fact lsx_bind_lpxs_aux: ∀h,o,a,I,G,L1,V,l. G ⊢ ⬊*[h, o, V, l] L1 →
- ∀Y,T. G ⊢ ⬊*[h, o, T, ⫯l] Y →
- ∀L2. Y = L2.ⓑ{I}V → ⦃G, L1⦄ ⊢ ➡*[h, o] L2 →
- G ⊢ ⬊*[h, o, ⓑ{a,I}V.T, l] L2.
-#h #o #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 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/
- ]
-| /3 width=4 by lpxs_pair/
-]
-qed-.
-
-lemma lsx_bind: ∀h,o,a,I,G,L,V,l. G ⊢ ⬊*[h, o, V, l] L →
- ∀T. G ⊢ ⬊*[h, o, T, ⫯l] L.ⓑ{I}V →
- G ⊢ ⬊*[h, o, ⓑ{a,I}V.T, l] L.
-/2 width=3 by lsx_bind_lpxs_aux/ qed.
-
-lemma lsx_flat_lpxs: ∀h,o,I,G,L1,V,l. G ⊢ ⬊*[h, o, V, l] L1 →
- ∀L2,T. G ⊢ ⬊*[h, o, T, l] L2 → ⦃G, L1⦄ ⊢ ➡*[h, o] L2 →
- G ⊢ ⬊*[h, o, ⓕ{I}V.T, l] L2.
-#h #o #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 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/
- ]
-| /3 width=1 by/
-]
-qed-.
-
-lemma lsx_flat: ∀h,o,I,G,L,V,l. G ⊢ ⬊*[h, o, V, l] L →
- ∀T. G ⊢ ⬊*[h, o, T, l] L → G ⊢ ⬊*[h, o, ⓕ{I}V.T, l] L.
-/2 width=3 by lsx_flat_lpxs/ qed.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/dpredstar_7.ma".
-include "basic_2/static/da.ma".
-include "basic_2/computation/cprs.ma".
-
-(* STRATIFIED DECOMPOSED PARALLEL COMPUTATION ON TERMS **********************)
-
-definition scpds: ∀h. sd h → nat → relation4 genv lenv term term ≝
- λh,o,d2,G,L,T1,T2.
- ∃∃T,d1. d2 ≤ d1 & ⦃G, L⦄ ⊢ T1 ▪[h, o] d1 & ⦃G, L⦄ ⊢ T1 •*[h, d2] T & ⦃G, L⦄ ⊢ T ➡* T2.
-
-interpretation "stratified decomposed parallel computation (term)"
- 'DPRedStar h o d G L T1 T2 = (scpds h o d G L T1 T2).
-
-(* Basic properties *********************************************************)
-
-lemma sta_cprs_scpds: ∀h,o,G,L,T1,T,T2,d. ⦃G, L⦄ ⊢ T1 ▪[h, o] d+1 → ⦃G, L⦄ ⊢ T1 •*[h, 1] T →
- ⦃G, L⦄ ⊢ T ➡* T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, o, 1] T2.
-/2 width=6 by ex4_2_intro/ qed.
-
-lemma lstas_scpds: ∀h,o,G,L,T1,T2,d1. ⦃G, L⦄ ⊢ T1 ▪[h, o] d1 →
- ∀d2. d2 ≤ d1 → ⦃G, L⦄ ⊢ T1 •*[h, d2] T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d2] T2.
-/2 width=6 by ex4_2_intro/ qed.
-
-lemma scpds_strap1: ∀h,o,G,L,T1,T,T2,d.
- ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d] T → ⦃G, L⦄ ⊢ T ➡ T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d] T2.
-#h #o #G #L #T1 #T #T2 #d * /3 width=8 by cprs_strap1, ex4_2_intro/
-qed.
-
-(* Basic forward lemmas *****************************************************)
-
-lemma scpds_fwd_cprs: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 •*➡*[h, o, 0] T2 →
- ⦃G, L⦄ ⊢ T1 ➡* T2.
-#h #o #G #L #T1 #T2 * /3 width=3 by cprs_strap2, lstas_cpr/
-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/unfold/lstas_aaa.ma".
-include "basic_2/computation/cpxs_aaa.ma".
-include "basic_2/computation/scpds.ma".
-
-(* STRATIFIED DECOMPOSED PARALLEL COMPUTATION ON TERMS **********************)
-
-(* Properties on atomic arity assignment for terms **************************)
-
-lemma scpds_aaa_conf: ∀h,o,G,L,d. Conf3 … (aaa G L) (scpds h o d G L).
-#h #o #G #L #d #A #T #HT #U * /3 width=6 by lstas_aaa_conf, cprs_aaa_conf/
-qed.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/static/da_lift.ma".
-include "basic_2/unfold/lstas_lift.ma".
-include "basic_2/computation/cprs_lift.ma".
-include "basic_2/computation/scpds.ma".
-
-(* STRATIFIED DECOMPOSED PARALLEL COMPUTATION ON TERMS **********************)
-
-(* Relocation properties ****************************************************)
-
-lemma scpds_lift: ∀h,o,G,d. d_liftable (scpds h o d G).
-#h #o #G #d2 #K #T1 #T2 * #T #d1 #Hd21 #Hd1 #HT1 #HT2 #L #c #l #k
-elim (lift_total T l k)
-/3 width=15 by cprs_lift, da_lift, lstas_lift, ex4_2_intro/
-qed.
-
-lemma scpds_inv_lift1: ∀h,o,G,d. d_deliftable_sn (scpds h o d G).
-#h #o #G #d2 #L #U1 #U2 * #U #d1 #Hd21 #Hd1 #HU1 #HU2 #K #c #l #k #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/
-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/unfold/lstas_da.ma".
-include "basic_2/computation/lprs_cprs.ma".
-include "basic_2/computation/cpxs_cpxs.ma".
-include "basic_2/computation/scpds.ma".
-
-(* STRATIFIED DECOMPOSED PARALLEL COMPUTATION ON TERMS **********************)
-
-(* Advanced properties ******************************************************)
-
-lemma scpds_strap2: ∀h,o,G,L,T1,T,T2,d1,d. ⦃G, L⦄ ⊢ T1 ▪[h, o] d1+1 →
- ⦃G, L⦄ ⊢ T1 •*[h, 1] T → ⦃G, L⦄ ⊢ T •*➡*[h, o, d] T2 →
- ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d+1] T2.
-#h #o #G #L #T1 #T #T2 #d1 #d #Hd1 #HT1 *
-#T0 #d0 #Hd0 #HTd0 #HT0 #HT02
-lapply (lstas_da_conf … HT1 … Hd1) <minus_plus_k_k #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,o,G,L,T1,T,T2,d.
- ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d] T → ⦃G, L⦄ ⊢ T ➡* T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d] T2.
-#h #o #G #L #T1 #T #T2 #d * /3 width=8 by cprs_trans, ex4_2_intro/
-qed-.
-
-lemma lstas_scpds_trans: ∀h,o,G,L,T1,T,T2,d1,d2,d.
- d2 ≤ d1 → ⦃G, L⦄ ⊢ T1 ▪[h, o] d1 →
- ⦃G, L⦄ ⊢ T1 •*[h, d2] T → ⦃G, L⦄ ⊢ T •*➡*[h, o, d] T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d2+d] T2.
-#h #o #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,o,a,G,L,V1,T1,U2,d. ⦃G, L⦄ ⊢ ⓛ{a}V1.T1 •*➡*[h, o, d] U2 →
- ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡* V2 & ⦃G, L.ⓛV1⦄ ⊢ T1 •*➡*[h, o, d] T2 &
- U2 = ⓛ{a}V2.T2.
-#h #o #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,o,a1,a2,G,L,V1,W2,T1,T2,d. ⦃G, L⦄ ⊢ ⓓ{a1}V1.T1 •*➡*[h, o, d] ⓛ{a2}W2.T2 →
- ∃∃T. ⦃G, L.ⓓV1⦄ ⊢ T1 •*➡*[h, o, d] T & ⬆[0, 1] ⓛ{a2}W2.T2 ≡ T & a1 = true.
-#h #o #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
-| /3 width=6 by ex4_2_intro, ex3_intro/
-]
-qed-.
-
-lemma scpds_inv_lstas_eq: ∀h,o,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d] T2 →
- ∀T. ⦃G, L⦄ ⊢ T1 •*[h, d] T → ⦃G, L⦄ ⊢ T ➡* T2.
-#h #o #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,o,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d] T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2.
-#h #o #G #L #T1 #T2 #d * /3 width=5 by cpxs_trans, lstas_cpxs, cprs_cpxs/
-qed-.
-
-(* Main properties **********************************************************)
-
-theorem scpds_conf_eq: ∀h,o,G,L,T0,T1,d. ⦃G, L⦄ ⊢ T0 •*➡*[h, o, d] T1 →
- ∀T2. ⦃G, L⦄ ⊢ T0 •*➡*[h, o, d] T2 →
- ∃∃T. ⦃G, L⦄ ⊢ T1 ➡* T & ⦃G, L⦄ ⊢ T2 ➡* T.
-#h #o #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-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/prednotreducible_3.ma".
+include "basic_2/reduction/crr.ma".
+
+(* IRREDUCIBLE TERMS FOR CONTEXT-SENSITIVE REDUCTION ************************)
+
+definition cir: relation3 genv lenv term ≝ λG,L,T. ⦃G, L⦄ ⊢ ➡ 𝐑⦃T⦄ → ⊥.
+
+interpretation "irreducibility for context-sensitive reduction (term)"
+ 'PRedNotReducible G L T = (cir G L T).
+
+(* Basic inversion lemmas ***************************************************)
+
+lemma cir_inv_delta: ∀G,L,K,V,i. ⬇[i] L ≡ K.ⓓV → ⦃G, L⦄ ⊢ ➡ 𝐈⦃#i⦄ → ⊥.
+/3 width=3 by crr_delta/ qed-.
+
+lemma cir_inv_ri2: ∀I,G,L,V,T. ri2 I → ⦃G, L⦄ ⊢ ➡ 𝐈⦃②{I}V.T⦄ → ⊥.
+/3 width=1 by crr_ri2/ qed-.
+
+lemma cir_inv_ib2: ∀a,I,G,L,V,T. ib2 a I → ⦃G, L⦄ ⊢ ➡ 𝐈⦃ⓑ{a,I}V.T⦄ →
+ ⦃G, L⦄ ⊢ ➡ 𝐈⦃V⦄ ∧ ⦃G, L.ⓑ{I}V⦄ ⊢ ➡ 𝐈⦃T⦄.
+/4 width=1 by crr_ib2_sn, crr_ib2_dx, conj/ qed-.
+
+lemma cir_inv_bind: ∀a,I,G,L,V,T. ⦃G, L⦄ ⊢ ➡ 𝐈⦃ⓑ{a,I}V.T⦄ →
+ ∧∧ ⦃G, L⦄ ⊢ ➡ 𝐈⦃V⦄ & ⦃G, L.ⓑ{I}V⦄ ⊢ ➡ 𝐈⦃T⦄ & ib2 a I.
+#a * [ elim a -a ]
+#G #L #V #T #H [ elim H -H /3 width=1 by crr_ri2, or_introl/ ]
+elim (cir_inv_ib2 … H) -H /3 width=1 by and3_intro, or_introl/
+qed-.
+
+lemma cir_inv_appl: ∀G,L,V,T. ⦃G, L⦄ ⊢ ➡ 𝐈⦃ⓐV.T⦄ →
+ ∧∧ ⦃G, L⦄ ⊢ ➡ 𝐈⦃V⦄ & ⦃G, L⦄ ⊢ ➡ 𝐈⦃T⦄ & 𝐒⦃T⦄.
+#G #L #V #T #HVT @and3_intro /3 width=1 by crr_appl_sn, crr_appl_dx/
+generalize in match HVT; -HVT elim T -T //
+* // #a * #U #T #_ #_ #H elim H -H /2 width=1 by crr_beta, crr_theta/
+qed-.
+
+lemma cir_inv_flat: ∀I,G,L,V,T. ⦃G, L⦄ ⊢ ➡ 𝐈⦃ⓕ{I}V.T⦄ →
+ ∧∧ ⦃G, L⦄ ⊢ ➡ 𝐈⦃V⦄ & ⦃G, L⦄ ⊢ ➡ 𝐈⦃T⦄ & 𝐒⦃T⦄ & I = Appl.
+* #G #L #V #T #H
+[ elim (cir_inv_appl … H) -H /2 width=1 by and4_intro/
+| elim (cir_inv_ri2 … H) -H //
+]
+qed-.
+
+(* Basic properties *********************************************************)
+
+lemma cir_sort: ∀G,L,s. ⦃G, L⦄ ⊢ ➡ 𝐈⦃⋆s⦄.
+/2 width=4 by crr_inv_sort/ qed.
+
+lemma cir_gref: ∀G,L,p. ⦃G, L⦄ ⊢ ➡ 𝐈⦃§p⦄.
+/2 width=4 by crr_inv_gref/ qed.
+
+lemma tir_atom: ∀G,I. ⦃G, ⋆⦄ ⊢ ➡ 𝐈⦃⓪{I}⦄.
+/2 width=3 by trr_inv_atom/ qed.
+
+lemma cir_ib2: ∀a,I,G,L,V,T.
+ ib2 a I → ⦃G, L⦄ ⊢ ➡ 𝐈⦃V⦄ → ⦃G, L.ⓑ{I}V⦄ ⊢ ➡ 𝐈⦃T⦄ → ⦃G, L⦄ ⊢ ➡ 𝐈⦃ⓑ{a,I}V.T⦄.
+#a #I #G #L #V #T #HI #HV #HT #H
+elim (crr_inv_ib2 … HI H) -HI -H /2 width=1 by/
+qed.
+
+lemma cir_appl: ∀G,L,V,T. ⦃G, L⦄ ⊢ ➡ 𝐈⦃V⦄ → ⦃G, L⦄ ⊢ ➡ 𝐈⦃T⦄ → 𝐒⦃T⦄ → ⦃G, L⦄ ⊢ ➡ 𝐈⦃ⓐV.T⦄.
+#G #L #V #T #HV #HT #H1 #H2
+elim (crr_inv_appl … H2) -H2 /2 width=1 by/
+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/crr_lift.ma".
+include "basic_2/reduction/cir.ma".
+
+(* IRREDUCIBLE TERMS FOR CONTEXT-SENSITIVE REDUCTION ************************)
+
+(* Properties on relocation *************************************************)
+
+lemma cir_lift: ∀G,K,T. ⦃G, K⦄ ⊢ ➡ 𝐈⦃T⦄ → ∀L,c,l,k. ⬇[c, l, k] L ≡ K →
+ ∀U. ⬆[l, k] T ≡ U → ⦃G, L⦄ ⊢ ➡ 𝐈⦃U⦄.
+/3 width=8 by crr_inv_lift/ qed.
+
+lemma cir_inv_lift: ∀G,L,U. ⦃G, L⦄ ⊢ ➡ 𝐈⦃U⦄ → ∀K,c,l,k. ⬇[c, l, k] L ≡ K →
+ ∀T. ⬆[l, k] T ≡ U → ⦃G, K⦄ ⊢ ➡ 𝐈⦃T⦄.
+/3 width=8 by crr_lift/ 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 *)
+(* *)
+(**************************************************************************)
+
+(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
+
+notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ ➡ 𝐈 break ⦃ term 46 T ⦄ )"
+ non associative with precedence 45
+ for @{ 'PRedNotReducible $G $L $T }.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/prednotreducible_5.ma".
+include "basic_2/reduction/cir.ma".
+include "basic_2/reduction/crx.ma".
+
+(* IRREDUCIBLE TERMS FOR CONTEXT-SENSITIVE EXTENDED REDUCTION ***************)
+
+definition cix: ∀h. sd h → relation3 genv lenv term ≝
+ λh,o,G,L,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃T⦄ → ⊥.
+
+interpretation "irreducibility for context-sensitive extended reduction (term)"
+ 'PRedNotReducible h o G L T = (cix h o G L T).
+
+(* Basic inversion lemmas ***************************************************)
+
+lemma cix_inv_sort: ∀h,o,G,L,s,d. deg h o s (d+1) → ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃⋆s⦄ → ⊥.
+/3 width=2 by crx_sort/ qed-.
+
+lemma cix_inv_delta: ∀h,o,I,G,L,K,V,i. ⬇[i] L ≡ K.ⓑ{I}V → ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃#i⦄ → ⊥.
+/3 width=4 by crx_delta/ qed-.
+
+lemma cix_inv_ri2: ∀h,o,I,G,L,V,T. ri2 I → ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃②{I}V.T⦄ → ⊥.
+/3 width=1 by crx_ri2/ qed-.
+
+lemma cix_inv_ib2: ∀h,o,a,I,G,L,V,T. ib2 a I → ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃ⓑ{a,I}V.T⦄ →
+ ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃V⦄ ∧ ⦃G, L.ⓑ{I}V⦄ ⊢ ➡[h, o] 𝐈⦃T⦄.
+/4 width=1 by crx_ib2_sn, crx_ib2_dx, conj/ qed-.
+
+lemma cix_inv_bind: ∀h,o,a,I,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃ⓑ{a,I}V.T⦄ →
+ ∧∧ ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃V⦄ & ⦃G, L.ⓑ{I}V⦄ ⊢ ➡[h, o] 𝐈⦃T⦄ & ib2 a I.
+#h #o #a * [ elim a -a ]
+#G #L #V #T #H [ elim H -H /3 width=1 by crx_ri2, or_introl/ ]
+elim (cix_inv_ib2 … H) -H /3 width=1 by and3_intro, or_introl/
+qed-.
+
+lemma cix_inv_appl: ∀h,o,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃ⓐV.T⦄ →
+ ∧∧ ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃V⦄ & ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃T⦄ & 𝐒⦃T⦄.
+#h #o #G #L #V #T #HVT @and3_intro /3 width=1 by crx_appl_sn, crx_appl_dx/
+generalize in match HVT; -HVT elim T -T //
+* // #a * #U #T #_ #_ #H elim H -H /2 width=1 by crx_beta, crx_theta/
+qed-.
+
+lemma cix_inv_flat: ∀h,o,I,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃ⓕ{I}V.T⦄ →
+ ∧∧ ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃V⦄ & ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃T⦄ & 𝐒⦃T⦄ & I = Appl.
+#h #o * #G #L #V #T #H
+[ elim (cix_inv_appl … H) -H /2 width=1 by and4_intro/
+| elim (cix_inv_ri2 … H) -H //
+]
+qed-.
+
+(* Basic forward lemmas *****************************************************)
+
+lemma cix_inv_cir: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃T⦄ → ⦃G, L⦄ ⊢ ➡ 𝐈⦃T⦄.
+/3 width=1 by crr_crx/ qed-.
+
+(* Basic properties *********************************************************)
+
+lemma cix_sort: ∀h,o,G,L,s. deg h o s 0 → ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃⋆s⦄.
+#h #o #G #L #s #Hk #H elim (crx_inv_sort … H) -L #d #Hkd
+lapply (deg_mono … Hk Hkd) -h -s <plus_n_Sm #H destruct
+qed.
+
+lemma tix_lref: ∀h,o,G,i. ⦃G, ⋆⦄ ⊢ ➡[h, o] 𝐈⦃#i⦄.
+#h #o #G #i #H elim (trx_inv_atom … H) -H #s #d #_ #H destruct
+qed.
+
+lemma cix_gref: ∀h,o,G,L,p. ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃§p⦄.
+#h #o #G #L #p #H elim (crx_inv_gref … H)
+qed.
+
+lemma cix_ib2: ∀h,o,a,I,G,L,V,T. ib2 a I → ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃V⦄ → ⦃G, L.ⓑ{I}V⦄ ⊢ ➡[h, o] 𝐈⦃T⦄ →
+ ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃ⓑ{a,I}V.T⦄.
+#h #o #a #I #G #L #V #T #HI #HV #HT #H
+elim (crx_inv_ib2 … HI H) -HI -H /2 width=1 by/
+qed.
+
+lemma cix_appl: ∀h,o,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃V⦄ → ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃T⦄ → 𝐒⦃T⦄ → ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃ⓐV.T⦄.
+#h #o #G #L #V #T #HV #HT #H1 #H2
+elim (crx_inv_appl … H2) -H2 /2 width=1 by/
+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/crx_lift.ma".
+include "basic_2/reduction/cix.ma".
+
+(* IRREDUCIBLE TERMS FOR CONTEXT-SENSITIVE EXTENDED REDUCTION ***************)
+
+(* Advanced properties ******************************************************)
+
+lemma cix_lref: ∀h,o,G,L,i. ⬇[i] L ≡ ⋆ → ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃#i⦄.
+#h #o #G #L #i #HL #H elim (crx_inv_lref … H) -h #I #K #V #HLK
+lapply (drop_mono … HLK … HL) -L -i #H destruct
+qed.
+
+(* Properties on relocation *************************************************)
+
+lemma cix_lift: ∀h,o,G,K,T. ⦃G, K⦄ ⊢ ➡[h, o] 𝐈⦃T⦄ → ∀L,c,l,k. ⬇[c, l, k] L ≡ K →
+ ∀U. ⬆[l, k] T ≡ U → ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃U⦄.
+/3 width=8 by crx_inv_lift/ qed.
+
+lemma cix_inv_lift: ∀h,o,G,L,U. ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃U⦄ → ∀K,c,l,k. ⬇[c, l, k] L ≡ K →
+ ∀T. ⬆[l, k] T ≡ U → ⦃G, K⦄ ⊢ ➡[h, o] 𝐈⦃T⦄.
+/3 width=8 by crx_lift/ 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 *)
+(* *)
+(**************************************************************************)
+
+(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
+
+notation "hvbox( ⦃ term 46 G, break term 46 L ⦄ ⊢ ➡ break [ term 46 o , break term 46 h ] 𝐈 break ⦃ term 46 T ⦄ )"
+ non associative with precedence 45
+ for @{ 'PRedNotReducible $h $o $G $L $T }.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/prednormal_3.ma".
+include "basic_2/reduction/cpr.ma".
+
+(* NORMAL TERMS FOR CONTEXT-SENSITIVE REDUCTION *****************************)
+
+definition cnr: relation3 genv lenv term ≝ λG,L. NF … (cpr G L) (eq …).
+
+interpretation
+ "normality for context-sensitive reduction (term)"
+ 'PRedNormal G L T = (cnr G L T).
+
+(* Basic inversion lemmas ***************************************************)
+
+lemma cnr_inv_delta: ∀G,L,K,V,i. ⬇[i] L ≡ K.ⓓV → ⦃G, L⦄ ⊢ ➡ 𝐍⦃#i⦄ → ⊥.
+#G #L #K #V #i #HLK #H
+elim (lift_total V 0 (i+1)) #W #HVW
+lapply (H W ?) -H [ /3 width=6 by cpr_delta/ ] -HLK #H destruct
+elim (lift_inv_lref2_be … HVW) -HVW /2 width=1 by ylt_inj/
+qed-.
+
+lemma cnr_inv_abst: ∀a,G,L,V,T. ⦃G, L⦄ ⊢ ➡ 𝐍⦃ⓛ{a}V.T⦄ → ⦃G, L⦄ ⊢ ➡ 𝐍⦃V⦄ ∧ ⦃G, L.ⓛV⦄ ⊢ ➡ 𝐍⦃T⦄.
+#a #G #L #V1 #T1 #HVT1 @conj
+[ #V2 #HV2 lapply (HVT1 (ⓛ{a}V2.T1) ?) -HVT1 /2 width=2 by cpr_pair_sn/ -HV2 #H destruct //
+| #T2 #HT2 lapply (HVT1 (ⓛ{a}V1.T2) ?) -HVT1 /2 width=2 by cpr_bind/ -HT2 #H destruct //
+]
+qed-.
+
+lemma cnr_inv_abbr: ∀G,L,V,T. ⦃G, L⦄ ⊢ ➡ 𝐍⦃-ⓓV.T⦄ → ⦃G, L⦄ ⊢ ➡ 𝐍⦃V⦄ ∧ ⦃G, L.ⓓV⦄ ⊢ ➡ 𝐍⦃T⦄.
+#G #L #V1 #T1 #HVT1 @conj
+[ #V2 #HV2 lapply (HVT1 (-ⓓV2.T1) ?) -HVT1 /2 width=2 by cpr_pair_sn/ -HV2 #H destruct //
+| #T2 #HT2 lapply (HVT1 (-ⓓV1.T2) ?) -HVT1 /2 width=2 by cpr_bind/ -HT2 #H destruct //
+]
+qed-.
+
+lemma cnr_inv_zeta: ∀G,L,V,T. ⦃G, L⦄ ⊢ ➡ 𝐍⦃+ⓓV.T⦄ → ⊥.
+#G #L #V #T #H elim (is_lift_dec T 0 1)
+[ * #U #HTU
+ lapply (H U ?) -H /2 width=3 by cpr_zeta/ #H destruct
+ elim (lift_inv_pair_xy_y … HTU)
+| #HT
+ elim (cpr_delift G (⋆) V T (⋆. ⓓV) 0) //
+ #T2 #T1 #HT2 #HT12 lapply (H (+ⓓV.T2) ?) -H /4 width=1 by tpr_cpr, cpr_bind/ -HT2
+ #H destruct /3 width=2 by ex_intro/
+]
+qed-.
+
+lemma cnr_inv_appl: ∀G,L,V,T. ⦃G, L⦄ ⊢ ➡ 𝐍⦃ⓐV.T⦄ → ∧∧ ⦃G, L⦄ ⊢ ➡ 𝐍⦃V⦄ & ⦃G, L⦄ ⊢ ➡ 𝐍⦃T⦄ & 𝐒⦃T⦄.
+#G #L #V1 #T1 #HVT1 @and3_intro
+[ #V2 #HV2 lapply (HVT1 (ⓐV2.T1) ?) -HVT1 /2 width=1 by cpr_pair_sn/ -HV2 #H destruct //
+| #T2 #HT2 lapply (HVT1 (ⓐV1.T2) ?) -HVT1 /2 width=1 by cpr_flat/ -HT2 #H destruct //
+| generalize in match HVT1; -HVT1 elim T1 -T1 * // #a * #W1 #U1 #_ #_ #H
+ [ elim (lift_total V1 0 1) #V2 #HV12
+ lapply (H (ⓓ{a}W1.ⓐV2.U1) ?) -H /3 width=3 by tpr_cpr, cpr_theta/ -HV12 #H destruct
+ | lapply (H (ⓓ{a}ⓝW1.V1.U1) ?) -H /3 width=1 by tpr_cpr, cpr_beta/ #H destruct
+]
+qed-.
+
+lemma cnr_inv_eps: ∀G,L,V,T. ⦃G, L⦄ ⊢ ➡ 𝐍⦃ⓝV.T⦄ → ⊥.
+#G #L #V #T #H lapply (H T ?) -H
+/2 width=4 by cpr_eps, discr_tpair_xy_y/
+qed-.
+
+(* Basic properties *********************************************************)
+
+(* Basic_1: was: nf2_sort *)
+lemma cnr_sort: ∀G,L,s. ⦃G, L⦄ ⊢ ➡ 𝐍⦃⋆s⦄.
+#G #L #s #X #H
+>(cpr_inv_sort1 … H) //
+qed.
+
+lemma cnr_lref_free: ∀G,L,i. |L| ≤ i → ⦃G, L⦄ ⊢ ➡ 𝐍⦃#i⦄.
+#G #L #i #Hi #X #H elim (cpr_inv_lref1 … H) -H // *
+#K #V1 #V2 #HLK lapply (drop_fwd_length_lt2 … HLK) -HLK
+#H elim (lt_refl_false i) /2 width=3 by lt_to_le_to_lt/
+qed.
+
+(* Basic_1: was only: nf2_csort_lref *)
+lemma cnr_lref_atom: ∀G,L,i. ⬇[i] L ≡ ⋆ → ⦃G, L⦄ ⊢ ➡ 𝐍⦃#i⦄.
+#G #L #i #HL @cnr_lref_free >(drop_fwd_length … HL) -HL //
+qed.
+
+(* Basic_1: was: nf2_abst *)
+lemma cnr_abst: ∀a,G,L,W,T. ⦃G, L⦄ ⊢ ➡ 𝐍⦃W⦄ → ⦃G, L.ⓛW⦄ ⊢ ➡ 𝐍⦃T⦄ → ⦃G, L⦄ ⊢ ➡ 𝐍⦃ⓛ{a}W.T⦄.
+#a #G #L #W #T #HW #HT #X #H
+elim (cpr_inv_abst1 … H) -H #W0 #T0 #HW0 #HT0 #H destruct
+>(HW … HW0) -W0 >(HT … HT0) -T0 //
+qed.
+
+(* Basic_1: was only: nf2_appl_lref *)
+lemma cnr_appl_simple: ∀G,L,V,T. ⦃G, L⦄ ⊢ ➡ 𝐍⦃V⦄ → ⦃G, L⦄ ⊢ ➡ 𝐍⦃T⦄ → 𝐒⦃T⦄ → ⦃G, L⦄ ⊢ ➡ 𝐍⦃ⓐV.T⦄.
+#G #L #V #T #HV #HT #HS #X #H
+elim (cpr_inv_appl1_simple … H) -H // #V0 #T0 #HV0 #HT0 #H destruct
+>(HV … HV0) -V0 >(HT … HT0) -T0 //
+qed.
+
+(* Basic_1: was: nf2_dec *)
+axiom cnr_dec: ∀G,L,T1. ⦃G, L⦄ ⊢ ➡ 𝐍⦃T1⦄ ∨
+ ∃∃T2. ⦃G, L⦄ ⊢ T1 ➡ T2 & (T1 = T2 → ⊥).
+
+(* Basic_1: removed theorems 1: nf2_abst_shift *)
--- /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/cpr_cir.ma".
+include "basic_2/reduction/cnr_crr.ma".
+
+(* NORMAL TERMS FOR CONTEXT-SENSITIVE REDUCTION *****************************)
+
+(* Main properties on irreducibility ****************************************)
+
+theorem cir_cnr: ∀G,L,T. ⦃G, L⦄ ⊢ ➡ 𝐈⦃T⦄ → ⦃G, L⦄ ⊢ ➡ 𝐍⦃T⦄.
+/2 width=4 by cpr_fwd_cir/ qed.
+
+(* Main inversion lemmas on irreducibility **********************************)
+
+theorem cnr_inv_cir: ∀G,L,T. ⦃G, L⦄ ⊢ ➡ 𝐍⦃T⦄ → ⦃G, L⦄ ⊢ ➡ 𝐈⦃T⦄.
+/2 width=5 by cnr_inv_crr/ 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/crr.ma".
+include "basic_2/reduction/cnr.ma".
+
+(* NORMAL TERMS FOR CONTEXT-SENSITIVE REDUCTION *****************************)
+
+(* Advanced inversion lemmas on reducibility ********************************)
+
+(* Note: this property is unusual *)
+lemma cnr_inv_crr: ∀G,L,T. ⦃G, L⦄ ⊢ ➡ 𝐑⦃T⦄ → ⦃G, L⦄ ⊢ ➡ 𝐍⦃T⦄ → ⊥.
+#G #L #T #H elim H -L -T
+[ #L #K #V #i #HLK #H
+ elim (cnr_inv_delta … HLK H)
+| #L #V #T #_ #IHV #H
+ elim (cnr_inv_appl … H) -H /2 width=1 by/
+| #L #V #T #_ #IHT #H
+ elim (cnr_inv_appl … H) -H /2 width=1 by/
+| #I #L #V #T * #H1 #H2 destruct
+ [ elim (cnr_inv_zeta … H2)
+ | elim (cnr_inv_eps … H2)
+ ]
+|5,6: #a * [ elim a ] #L #V #T * #H1 #_ #IH #H2 destruct
+ [1,3: elim (cnr_inv_abbr … H2) -H2 /2 width=1 by/
+ |*: elim (cnr_inv_abst … H2) -H2 /2 width=1 by/
+ ]
+| #a #L #V #W #T #H
+ elim (cnr_inv_appl … H) -H #_ #_ #H
+ elim (simple_inv_bind … H)
+| #a #L #V #W #T #H
+ elim (cnr_inv_appl … H) -H #_ #_ #H
+ elim (simple_inv_bind … H)
+]
+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/cpr_lift.ma".
+include "basic_2/reduction/cnr.ma".
+
+(* NORMAL TERMS FOR CONTEXT-SENSITIVE REDUCTION *****************************)
+
+(* Advanced properties ******************************************************)
+
+(* Basic_1: was: nf2_lref_abst *)
+lemma cnr_lref_abst: ∀G,L,K,V,i. ⬇[i] L ≡ K. ⓛV → ⦃G, L⦄ ⊢ ➡ 𝐍⦃#i⦄.
+#G #L #K #V #i #HLK #X #H
+elim (cpr_inv_lref1 … H) -H // *
+#K0 #V1 #V2 #HLK0 #_ #_
+lapply (drop_mono … HLK … HLK0) -L #H destruct
+qed.
+
+(* Relocation properties ****************************************************)
+
+(* Basic_1: was: nf2_lift *)
+lemma cnr_lift: ∀G,L0,L,T,T0,c,l,k. ⦃G, L⦄ ⊢ ➡ 𝐍⦃T⦄ →
+ ⬇[c, l, k] L0 ≡ L → ⬆[l, k] T ≡ T0 → ⦃G, L0⦄ ⊢ ➡ 𝐍⦃T0⦄.
+#G #L0 #L #T #T0 #c #l #k #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,c,l,k. ⦃G, L0⦄ ⊢ ➡ 𝐍⦃T0⦄ →
+ ⬇[c, l, k] L0 ≡ L → ⬆[l, k] T ≡ T0 → ⦃G, L⦄ ⊢ ➡ 𝐍⦃T⦄.
+#G #L0 #L #T #T0 #c #l #k #HLT0 #HL0 #HT0 #X #H
+elim (lift_total X l k) #X0 #HX0
+lapply (cpr_lift … H … HL0 … HT0 … HX0) -L #HTX0
+>(HLT0 … HTX0) in HX0; -L0 -X0 #H
+>(lift_inj … H … HT0) -T0 -X -c -l -k //
+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 *)
+(* *)
+(**************************************************************************)
+
+(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
+
+notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ ➡ 𝐍 break ⦃ term 46 T ⦄ )"
+ non associative with precedence 45
+ for @{ 'PRedNormal $G $L $T }.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/prednormal_5.ma".
+include "basic_2/reduction/cnr.ma".
+include "basic_2/reduction/cpx.ma".
+
+(* NORMAL TERMS FOR CONTEXT-SENSITIVE EXTENDED REDUCTION ********************)
+
+definition cnx: ∀h. sd h → relation3 genv lenv term ≝
+ λh,o,G,L. NF … (cpx h o G L) (eq …).
+
+interpretation
+ "normality for context-sensitive extended reduction (term)"
+ 'PRedNormal h o L T = (cnx h o L T).
+
+(* Basic inversion lemmas ***************************************************)
+
+lemma cnx_inv_sort: ∀h,o,G,L,s. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃⋆s⦄ → deg h o s 0.
+#h #o #G #L #s #H elim (deg_total h o s)
+#d @(nat_ind_plus … d) -d // #d #_ #Hkd
+lapply (H (⋆(next h s)) ?) -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 s) >H -H #H elim (lt_refl_false … H)
+qed-.
+
+lemma cnx_inv_delta: ∀h,o,I,G,L,K,V,i. ⬇[i] L ≡ K.ⓑ{I}V → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃#i⦄ → ⊥.
+#h #o #I #G #L #K #V #i #HLK #H
+elim (lift_total V 0 (i+1)) #W #HVW
+lapply (H W ?) -H [ /3 width=7 by cpx_delta/ ] -HLK #H destruct
+elim (lift_inv_lref2_be … HVW) -HVW /2 width=1 by ylt_inj/
+qed-.
+
+lemma cnx_inv_abst: ∀h,o,a,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃ⓛ{a}V.T⦄ →
+ ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃V⦄ ∧ ⦃G, L.ⓛV⦄ ⊢ ➡[h, o] 𝐍⦃T⦄.
+#h #o #a #G #L #V1 #T1 #HVT1 @conj
+[ #V2 #HV2 lapply (HVT1 (ⓛ{a}V2.T1) ?) -HVT1 /2 width=2 by cpx_pair_sn/ -HV2 #H destruct //
+| #T2 #HT2 lapply (HVT1 (ⓛ{a}V1.T2) ?) -HVT1 /2 width=2 by cpx_bind/ -HT2 #H destruct //
+]
+qed-.
+
+lemma cnx_inv_abbr: ∀h,o,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃-ⓓV.T⦄ →
+ ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃V⦄ ∧ ⦃G, L.ⓓV⦄ ⊢ ➡[h, o] 𝐍⦃T⦄.
+#h #o #G #L #V1 #T1 #HVT1 @conj
+[ #V2 #HV2 lapply (HVT1 (-ⓓV2.T1) ?) -HVT1 /2 width=2 by cpx_pair_sn/ -HV2 #H destruct //
+| #T2 #HT2 lapply (HVT1 (-ⓓV1.T2) ?) -HVT1 /2 width=2 by cpx_bind/ -HT2 #H destruct //
+]
+qed-.
+
+lemma cnx_inv_zeta: ∀h,o,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃+ⓓV.T⦄ → ⊥.
+#h #o #G #L #V #T #H elim (is_lift_dec T 0 1)
+[ * #U #HTU
+ lapply (H U ?) -H /2 width=3 by cpx_zeta/ #H destruct
+ elim (lift_inv_pair_xy_y … HTU)
+| #HT
+ elim (cpr_delift G(⋆) V T (⋆.ⓓV) 0) // #T2 #T1 #HT2 #HT12
+ lapply (H (+ⓓV.T2) ?) -H /5 width=1 by cpr_cpx, tpr_cpr, cpr_bind/ -HT2
+ #H destruct /3 width=2 by ex_intro/
+]
+qed-.
+
+lemma cnx_inv_appl: ∀h,o,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃ⓐV.T⦄ →
+ ∧∧ ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃V⦄ & ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ & 𝐒⦃T⦄.
+#h #o #G #L #V1 #T1 #HVT1 @and3_intro
+[ #V2 #HV2 lapply (HVT1 (ⓐV2.T1) ?) -HVT1 /2 width=1 by cpx_pair_sn/ -HV2 #H destruct //
+| #T2 #HT2 lapply (HVT1 (ⓐV1.T2) ?) -HVT1 /2 width=1 by cpx_flat/ -HT2 #H destruct //
+| generalize in match HVT1; -HVT1 elim T1 -T1 * // #a * #W1 #U1 #_ #_ #H
+ [ elim (lift_total V1 0 1) #V2 #HV12
+ lapply (H (ⓓ{a}W1.ⓐV2.U1) ?) -H /3 width=3 by cpr_cpx, cpr_theta/ -HV12 #H destruct
+ | lapply (H (ⓓ{a}ⓝW1.V1.U1) ?) -H /3 width=1 by cpr_cpx, cpr_beta/ #H destruct
+ ]
+]
+qed-.
+
+lemma cnx_inv_eps: ∀h,o,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃ⓝV.T⦄ → ⊥.
+#h #o #G #L #V #T #H lapply (H T ?) -H
+/2 width=4 by cpx_eps, discr_tpair_xy_y/
+qed-.
+
+(* Basic forward lemmas *****************************************************)
+
+lemma cnx_fwd_cnr: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ → ⦃G, L⦄ ⊢ ➡ 𝐍⦃T⦄.
+#h #o #G #L #T #H #U #HTU
+@H /2 width=1 by cpr_cpx/ (**) (* auto fails because a δ-expansion gets in the way *)
+qed-.
+
+(* Basic properties *********************************************************)
+
+lemma cnx_sort: ∀h,o,G,L,s. deg h o s 0 → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃⋆s⦄.
+#h #o #G #L #s #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,o,G,L,s,d. deg h o s d → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃⋆((next h)^d s)⦄.
+#h #o #G #L #s #d #Hkd
+lapply (deg_iter … d Hkd) -Hkd <minus_n_n /2 width=6 by cnx_sort/
+qed.
+
+lemma cnx_lref_free: ∀h,o,G,L,i. |L| ≤ i → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃#i⦄.
+#h #o #G #L #i #Hi #X #H elim (cpx_inv_lref1 … H) -H // *
+#I #K #V1 #V2 #HLK lapply (drop_fwd_length_lt2 … HLK) -HLK
+#H elim (lt_refl_false i) /2 width=3 by lt_to_le_to_lt/
+qed.
+
+lemma cnx_lref_atom: ∀h,o,G,L,i. ⬇[i] L ≡ ⋆ → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃#i⦄.
+#h #o #G #L #i #HL @cnx_lref_free >(drop_fwd_length … HL) -HL //
+qed.
+
+lemma cnx_abst: ∀h,o,a,G,L,W,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃W⦄ → ⦃G, L.ⓛW⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ →
+ ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃ⓛ{a}W.T⦄.
+#h #o #a #G #L #W #T #HW #HT #X #H
+elim (cpx_inv_abst1 … H) -H #W0 #T0 #HW0 #HT0 #H destruct
+>(HW … HW0) -W0 >(HT … HT0) -T0 //
+qed.
+
+lemma cnx_appl_simple: ∀h,o,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃V⦄ → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ → 𝐒⦃T⦄ →
+ ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃ⓐV.T⦄.
+#h #o #G #L #V #T #HV #HT #HS #X #H
+elim (cpx_inv_appl1_simple … H) -H // #V0 #T0 #HV0 #HT0 #H destruct
+>(HV … HV0) -V0 >(HT … HT0) -T0 //
+qed.
+
+axiom cnx_dec: ∀h,o,G,L,T1. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T1⦄ ∨
+ ∃∃T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 & (T1 = T2 → ⊥).
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/reduction/cpx_cix.ma".
+include "basic_2/reduction/cnx_crx.ma".
+
+(* NORMAL TERMS FOR CONTEXT-SENSITIVE EXTENDED REDUCTION ********************)
+
+(* Main properties on irreducibility ****************************************)
+
+theorem cix_cnx: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃T⦄ → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄.
+/2 width=6 by cpx_fwd_cix/ qed.
+
+(* Main inversion lemmas on irreducibility **********************************)
+
+theorem cnx_inv_cix: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ → ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃T⦄.
+/2 width=7 by cnx_inv_crx/ 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/crx.ma".
+include "basic_2/reduction/cnx.ma".
+
+(* NORMAL TERMS FOR CONTEXT-SENSITIVE EXTENDED REDUCTION ********************)
+
+(* Advanced inversion lemmas on reducibility ********************************)
+
+(* Note: this property is unusual *)
+lemma cnx_inv_crx: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃T⦄ → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ → ⊥.
+#h #o #G #L #T #H elim H -L -T
+[ #L #s #d #Hkd #H
+ lapply (cnx_inv_sort … H) -H #H
+ lapply (deg_mono … H Hkd) -h -L -s <plus_n_Sm #H destruct
+| #I #L #K #V #i #HLK #H
+ elim (cnx_inv_delta … HLK H)
+| #L #V #T #_ #IHV #H
+ elim (cnx_inv_appl … H) -H /2 width=1 by/
+| #L #V #T #_ #IHT #H
+ elim (cnx_inv_appl … H) -H /2 width=1 by/
+| #I #L #V #T * #H1 #H2 destruct
+ [ elim (cnx_inv_zeta … H2)
+ | elim (cnx_inv_eps … H2)
+ ]
+|6,7: #a * [ elim a ] #L #V #T * #H1 #_ #IH #H2 destruct
+ [1,3: elim (cnx_inv_abbr … H2) -H2 /2 width=1 by/
+ |*: elim (cnx_inv_abst … H2) -H2 /2 width=1 by/
+ ]
+| #a #L #V #W #T #H
+ elim (cnx_inv_appl … H) -H #_ #_ #H
+ elim (simple_inv_bind … H)
+| #a #L #V #W #T #H
+ elim (cnx_inv_appl … H) -H #_ #_ #H
+ elim (simple_inv_bind … H)
+]
+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_lift.ma".
+include "basic_2/reduction/cnx.ma".
+
+(* NORMAL TERMS FOR CONTEXT-SENSITIVE EXTENDED REDUCTION ********************)
+
+(* Relocation properties ****************************************************)
+
+lemma cnx_lift: ∀h,o,G,L0,L,T,T0,c,l,k. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ → ⬇[c, l, k] L0 ≡ L →
+ ⬆[l, k] T ≡ T0 → ⦃G, L0⦄ ⊢ ➡[h, o] 𝐍⦃T0⦄.
+#h #o #G #L0 #L #T #T0 #c #l #k #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,o,G,L0,L,T,T0,c,l,k. ⦃G, L0⦄ ⊢ ➡[h, o] 𝐍⦃T0⦄ → ⬇[c, l, k] L0 ≡ L →
+ ⬆[l, k] T ≡ T0 → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄.
+#h #o #G #L0 #L #T #T0 #c #l #k #HLT0 #HL0 #HT0 #X #H
+elim (lift_total X l k) #X0 #HX0
+lapply (cpx_lift … H … HL0 … HT0 … HX0) -L #HTX0
+>(HLT0 … HTX0) in HX0; -L0 -X0 #H
+>(lift_inj … H … HT0) -T0 -X -l -k //
+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 *)
+(* *)
+(**************************************************************************)
+
+(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
+
+notation "hvbox( ⦃ term 46 G, break term 46 L ⦄ ⊢ ➡ break [ term 46 o , break term 46 h ] 𝐍 break ⦃ term 46 T ⦄ )"
+ non associative with precedence 45
+ for @{ 'PRedNormal $h $o $G $L $T }.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/predreducible_3.ma".
+include "basic_2/grammar/genv.ma".
+include "basic_2/substitution/drop.ma".
+
+(* REDUCIBLE TERMS FOR CONTEXT-SENSITIVE REDUCTION **************************)
+
+(* reducible binary items *)
+definition ri2: predicate item2 ≝
+ λI. I = Bind2 true Abbr ∨ I = Flat2 Cast.
+
+(* irreducible binary binders *)
+definition ib2: relation2 bool bind2 ≝
+ λa,I. I = Abst ∨ Bind2 a I = Bind2 false Abbr.
+
+(* activate genv *)
+(* reducible terms *)
+inductive crr (G:genv): relation2 lenv term ≝
+| crr_delta : ∀L,K,V,i. ⬇[i] L ≡ K.ⓓV → crr G L (#i)
+| crr_appl_sn: ∀L,V,T. crr G L V → crr G L (ⓐV.T)
+| crr_appl_dx: ∀L,V,T. crr G L T → crr G L (ⓐV.T)
+| crr_ri2 : ∀I,L,V,T. ri2 I → crr G L (②{I}V.T)
+| crr_ib2_sn : ∀a,I,L,V,T. ib2 a I → crr G L V → crr G L (ⓑ{a,I}V.T)
+| crr_ib2_dx : ∀a,I,L,V,T. ib2 a I → crr G (L.ⓑ{I}V) T → crr G L (ⓑ{a,I}V.T)
+| crr_beta : ∀a,L,V,W,T. crr G L (ⓐV.ⓛ{a}W.T)
+| crr_theta : ∀a,L,V,W,T. crr G L (ⓐV.ⓓ{a}W.T)
+.
+
+interpretation
+ "reducibility for context-sensitive reduction (term)"
+ 'PRedReducible G L T = (crr G L T).
+
+(* Basic inversion lemmas ***************************************************)
+
+fact crr_inv_sort_aux: ∀G,L,T,s. ⦃G, L⦄ ⊢ ➡ 𝐑⦃T⦄ → T = ⋆s → ⊥.
+#G #L #T #s0 * -L -T
+[ #L #K #V #i #HLK #H destruct
+| #L #V #T #_ #H destruct
+| #L #V #T #_ #H destruct
+| #I #L #V #T #_ #H destruct
+| #a #I #L #V #T #_ #_ #H destruct
+| #a #I #L #V #T #_ #_ #H destruct
+| #a #L #V #W #T #H destruct
+| #a #L #V #W #T #H destruct
+]
+qed-.
+
+lemma crr_inv_sort: ∀G,L,s. ⦃G, L⦄ ⊢ ➡ 𝐑⦃⋆s⦄ → ⊥.
+/2 width=6 by crr_inv_sort_aux/ qed-.
+
+fact crr_inv_lref_aux: ∀G,L,T,i. ⦃G, L⦄ ⊢ ➡ 𝐑⦃T⦄ → T = #i →
+ ∃∃K,V. ⬇[i] L ≡ K.ⓓV.
+#G #L #T #j * -L -T
+[ #L #K #V #i #HLK #H destruct /2 width=3 by ex1_2_intro/
+| #L #V #T #_ #H destruct
+| #L #V #T #_ #H destruct
+| #I #L #V #T #_ #H destruct
+| #a #I #L #V #T #_ #_ #H destruct
+| #a #I #L #V #T #_ #_ #H destruct
+| #a #L #V #W #T #H destruct
+| #a #L #V #W #T #H destruct
+]
+qed-.
+
+lemma crr_inv_lref: ∀G,L,i. ⦃G, L⦄ ⊢ ➡ 𝐑⦃#i⦄ → ∃∃K,V. ⬇[i] L ≡ K.ⓓV.
+/2 width=4 by crr_inv_lref_aux/ qed-.
+
+fact crr_inv_gref_aux: ∀G,L,T,p. ⦃G, L⦄ ⊢ ➡ 𝐑⦃T⦄ → T = §p → ⊥.
+#G #L #T #q * -L -T
+[ #L #K #V #i #HLK #H destruct
+| #L #V #T #_ #H destruct
+| #L #V #T #_ #H destruct
+| #I #L #V #T #_ #H destruct
+| #a #I #L #V #T #_ #_ #H destruct
+| #a #I #L #V #T #_ #_ #H destruct
+| #a #L #V #W #T #H destruct
+| #a #L #V #W #T #H destruct
+]
+qed-.
+
+lemma crr_inv_gref: ∀G,L,p. ⦃G, L⦄ ⊢ ➡ 𝐑⦃§p⦄ → ⊥.
+/2 width=6 by crr_inv_gref_aux/ qed-.
+
+lemma trr_inv_atom: ∀G,I. ⦃G, ⋆⦄ ⊢ ➡ 𝐑⦃⓪{I}⦄ → ⊥.
+#G * #i #H
+[ elim (crr_inv_sort … H)
+| elim (crr_inv_lref … H) -H #L #V #H
+ elim (drop_inv_atom1 … H) -H #H destruct
+| elim (crr_inv_gref … H)
+]
+qed-.
+
+fact crr_inv_ib2_aux: ∀a,I,G,L,W,U,T. ib2 a I → ⦃G, L⦄ ⊢ ➡ 𝐑⦃T⦄ → T = ⓑ{a,I}W.U →
+ ⦃G, L⦄ ⊢ ➡ 𝐑⦃W⦄ ∨ ⦃G, L.ⓑ{I}W⦄ ⊢ ➡ 𝐑⦃U⦄.
+#G #b #J #L #W0 #U #T #HI * -L -T
+[ #L #K #V #i #_ #H destruct
+| #L #V #T #_ #H destruct
+| #L #V #T #_ #H destruct
+| #I #L #V #T #H1 #H2 destruct
+ elim H1 -H1 #H destruct
+ elim HI -HI #H destruct
+| #a #I #L #V #T #_ #HV #H destruct /2 width=1 by or_introl/
+| #a #I #L #V #T #_ #HT #H destruct /2 width=1 by or_intror/
+| #a #L #V #W #T #H destruct
+| #a #L #V #W #T #H destruct
+]
+qed-.
+
+lemma crr_inv_ib2: ∀a,I,G,L,W,T. ib2 a I → ⦃G, L⦄ ⊢ ➡ 𝐑⦃ⓑ{a,I}W.T⦄ →
+ ⦃G, L⦄ ⊢ ➡ 𝐑⦃W⦄ ∨ ⦃G, L.ⓑ{I}W⦄ ⊢ ➡ 𝐑⦃T⦄.
+/2 width=5 by crr_inv_ib2_aux/ qed-.
+
+fact crr_inv_appl_aux: ∀G,L,W,U,T. ⦃G, L⦄ ⊢ ➡ 𝐑⦃T⦄ → T = ⓐW.U →
+ ∨∨ ⦃G, L⦄ ⊢ ➡ 𝐑⦃W⦄ | ⦃G, L⦄ ⊢ ➡ 𝐑⦃U⦄ | (𝐒⦃U⦄ → ⊥).
+#G #L #W0 #U #T * -L -T
+[ #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/
+| #I #L #V #T #H1 #H2 destruct
+ elim H1 -H1 #H destruct
+| #a #I #L #V #T #_ #_ #H destruct
+| #a #I #L #V #T #_ #_ #H destruct
+| #a #L #V #W #T #H destruct
+ @or3_intro2 #H elim (simple_inv_bind … H)
+| #a #L #V #W #T #H destruct
+ @or3_intro2 #H elim (simple_inv_bind … H)
+]
+qed-.
+
+lemma crr_inv_appl: ∀G,L,V,T. ⦃G, L⦄ ⊢ ➡ 𝐑⦃ⓐV.T⦄ →
+ ∨∨ ⦃G, L⦄ ⊢ ➡ 𝐑⦃V⦄ | ⦃G, L⦄ ⊢ ➡ 𝐑⦃T⦄ | (𝐒⦃T⦄ → ⊥).
+/2 width=3 by crr_inv_appl_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_drop.ma".
+include "basic_2/reduction/crr.ma".
+
+(* REDUCIBLE TERMS FOR CONTEXT-SENSITIVE REDUCTION **************************)
+
+(* Properties on relocation *************************************************)
+
+lemma crr_lift: ∀G,K,T. ⦃G, K⦄ ⊢ ➡ 𝐑⦃T⦄ → ∀L,c,l,k. ⬇[c, l, k] L ≡ K →
+ ∀U. ⬆[l, k] T ≡ U → ⦃G, L⦄ ⊢ ➡ 𝐑⦃U⦄.
+#G #K #T #H elim H -K -T
+[ #K #K0 #V #i #HK0 #L #c #l #k #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 #c #l #k #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 #c #l #k #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 #c #l #k #_ #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 #c #l #k #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 #c #l #k #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 #c #l #k #_ #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 #c #l #k #_ #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,c,l,k. ⬇[c, l, k] L ≡ K →
+ ∀T. ⬆[l, k] T ≡ U → ⦃G, K⦄ ⊢ ➡ 𝐑⦃T⦄.
+#G #L #U #H elim H -L -U
+[ #L #L0 #W #i #HK0 #K #c #l #k #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 #c #l #k #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 #c #l #k #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 #c #l #k #_ #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 #c #l #k #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 #c #l #k #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 #c #l #k #_ #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 #c #l #k #_ #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/
+]
+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 *)
+(* *)
+(**************************************************************************)
+
+(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
+
+notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ ➡ 𝐑 break ⦃ term 46 T ⦄ )"
+ non associative with precedence 45
+ for @{ 'PRedReducible $G $L $T }.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/predreducible_5.ma".
+include "basic_2/static/sd.ma".
+include "basic_2/reduction/crr.ma".
+
+(* REDUCIBLE TERMS FOR CONTEXT-SENSITIVE EXTENDED REDUCTION *****************)
+
+(* activate genv *)
+(* extended reducible terms *)
+inductive crx (h) (o) (G:genv): relation2 lenv term ≝
+| crx_sort : ∀L,s,d. deg h o s (d+1) → crx h o G L (⋆s)
+| crx_delta : ∀I,L,K,V,i. ⬇[i] L ≡ K.ⓑ{I}V → crx h o G L (#i)
+| crx_appl_sn: ∀L,V,T. crx h o G L V → crx h o G L (ⓐV.T)
+| crx_appl_dx: ∀L,V,T. crx h o G L T → crx h o G L (ⓐV.T)
+| crx_ri2 : ∀I,L,V,T. ri2 I → crx h o G L (②{I}V.T)
+| crx_ib2_sn : ∀a,I,L,V,T. ib2 a I → crx h o G L V → crx h o G L (ⓑ{a,I}V.T)
+| crx_ib2_dx : ∀a,I,L,V,T. ib2 a I → crx h o G (L.ⓑ{I}V) T → crx h o G L (ⓑ{a,I}V.T)
+| crx_beta : ∀a,L,V,W,T. crx h o G L (ⓐV. ⓛ{a}W.T)
+| crx_theta : ∀a,L,V,W,T. crx h o G L (ⓐV. ⓓ{a}W.T)
+.
+
+interpretation
+ "reducibility for context-sensitive extended reduction (term)"
+ 'PRedReducible h o G L T = (crx h o G L T).
+
+(* Basic properties *********************************************************)
+
+lemma crr_crx: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ➡ 𝐑⦃T⦄ → ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃T⦄.
+#h #o #G #L #T #H elim H -L -T
+/2 width=4 by crx_delta, crx_appl_sn, crx_appl_dx, crx_ri2, crx_ib2_sn, crx_ib2_dx, crx_beta, crx_theta/
+qed.
+
+(* Basic inversion lemmas ***************************************************)
+
+fact crx_inv_sort_aux: ∀h,o,G,L,T,s. ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃T⦄ → T = ⋆s →
+ ∃d. deg h o s (d+1).
+#h #o #G #L #T #s0 * -L -T
+[ #L #s #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
+| #I #L #V #T #_ #H destruct
+| #a #I #L #V #T #_ #_ #H destruct
+| #a #I #L #V #T #_ #_ #H destruct
+| #a #L #V #W #T #H destruct
+| #a #L #V #W #T #H destruct
+]
+qed-.
+
+lemma crx_inv_sort: ∀h,o,G,L,s. ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃⋆s⦄ → ∃d. deg h o s (d+1).
+/2 width=5 by crx_inv_sort_aux/ qed-.
+
+fact crx_inv_lref_aux: ∀h,o,G,L,T,i. ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃T⦄ → T = #i →
+ ∃∃I,K,V. ⬇[i] L ≡ K.ⓑ{I}V.
+#h #o #G #L #T #j * -L -T
+[ #L #s #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
+| #I #L #V #T #_ #H destruct
+| #a #I #L #V #T #_ #_ #H destruct
+| #a #I #L #V #T #_ #_ #H destruct
+| #a #L #V #W #T #H destruct
+| #a #L #V #W #T #H destruct
+]
+qed-.
+
+lemma crx_inv_lref: ∀h,o,G,L,i. ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃#i⦄ → ∃∃I,K,V. ⬇[i] L ≡ K.ⓑ{I}V.
+/2 width=6 by crx_inv_lref_aux/ qed-.
+
+fact crx_inv_gref_aux: ∀h,o,G,L,T,p. ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃T⦄ → T = §p → ⊥.
+#h #o #G #L #T #q * -L -T
+[ #L #s #d #_ #H destruct
+| #I #L #K #V #i #HLK #H destruct
+| #L #V #T #_ #H destruct
+| #L #V #T #_ #H destruct
+| #I #L #V #T #_ #H destruct
+| #a #I #L #V #T #_ #_ #H destruct
+| #a #I #L #V #T #_ #_ #H destruct
+| #a #L #V #W #T #H destruct
+| #a #L #V #W #T #H destruct
+]
+qed-.
+
+lemma crx_inv_gref: ∀h,o,G,L,p. ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃§p⦄ → ⊥.
+/2 width=8 by crx_inv_gref_aux/ qed-.
+
+lemma trx_inv_atom: ∀h,o,I,G. ⦃G, ⋆⦄ ⊢ ➡[h, o] 𝐑⦃⓪{I}⦄ →
+ ∃∃s,d. deg h o s (d+1) & I = Sort s.
+#h #o * #i #G #H
+[ elim (crx_inv_sort … H) -H /2 width=4 by ex2_2_intro/
+| elim (crx_inv_lref … H) -H #I #L #V #H
+ elim (drop_inv_atom1 … H) -H #H destruct
+| elim (crx_inv_gref … H)
+]
+qed-.
+
+fact crx_inv_ib2_aux: ∀h,o,a,I,G,L,W,U,T. ib2 a I → ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃T⦄ →
+ T = ⓑ{a,I}W.U → ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃W⦄ ∨ ⦃G, L.ⓑ{I}W⦄ ⊢ ➡[h, o] 𝐑⦃U⦄.
+#h #o #b #J #G #L #W0 #U #T #HI * -L -T
+[ #L #s #d #_ #H destruct
+| #I #L #K #V #i #_ #H destruct
+| #L #V #T #_ #H destruct
+| #L #V #T #_ #H destruct
+| #I #L #V #T #H1 #H2 destruct
+ elim H1 -H1 #H destruct
+ elim HI -HI #H destruct
+| #a #I #L #V #T #_ #HV #H destruct /2 width=1 by or_introl/
+| #a #I #L #V #T #_ #HT #H destruct /2 width=1 by or_intror/
+| #a #L #V #W #T #H destruct
+| #a #L #V #W #T #H destruct
+]
+qed-.
+
+lemma crx_inv_ib2: ∀h,o,a,I,G,L,W,T. ib2 a I → ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃ⓑ{a,I}W.T⦄ →
+ ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃W⦄ ∨ ⦃G, L.ⓑ{I}W⦄ ⊢ ➡[h, o] 𝐑⦃T⦄.
+/2 width=5 by crx_inv_ib2_aux/ qed-.
+
+fact crx_inv_appl_aux: ∀h,o,G,L,W,U,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃T⦄ → T = ⓐW.U →
+ ∨∨ ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃W⦄ | ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃U⦄ | (𝐒⦃U⦄ → ⊥).
+#h #o #G #L #W0 #U #T * -L -T
+[ #L #s #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/
+| #I #L #V #T #H1 #H2 destruct
+ elim H1 -H1 #H destruct
+| #a #I #L #V #T #_ #_ #H destruct
+| #a #I #L #V #T #_ #_ #H destruct
+| #a #L #V #W #T #H destruct
+ @or3_intro2 #H elim (simple_inv_bind … H)
+| #a #L #V #W #T #H destruct
+ @or3_intro2 #H elim (simple_inv_bind … H)
+]
+qed-.
+
+lemma crx_inv_appl: ∀h,o,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃ⓐV.T⦄ →
+ ∨∨ ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃V⦄ | ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃T⦄ | (𝐒⦃T⦄ → ⊥).
+/2 width=3 by crx_inv_appl_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_drop.ma".
+include "basic_2/reduction/crx.ma".
+
+(* REDUCIBLE TERMS FOR CONTEXT-SENSITIVE EXTENDED REDUCTION *****************)
+
+(* Properties on relocation *************************************************)
+
+lemma crx_lift: ∀h,o,G,K,T. ⦃G, K⦄ ⊢ ➡[h, o] 𝐑⦃T⦄ → ∀L,c,l,k. ⬇[c, l, k] L ≡ K →
+ ∀U. ⬆[l, k] T ≡ U → ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃U⦄.
+#h #o #G #K #T #H elim H -K -T
+[ #K #s #d #Hkd #L #c #l #k #_ #X #H
+ >(lift_inv_sort1 … H) -X /2 width=2 by crx_sort/
+| #I #K #K0 #V #i #HK0 #L #c #l #k #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 #c #l #k #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 #c #l #k #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 #c #l #k #_ #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 #c #l #k #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 #c #l #k #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 #c #l #k #_ #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 #c #l #k #_ #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,o,G,L,U. ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃U⦄ → ∀K,c,l,k. ⬇[c, l, k] L ≡ K →
+ ∀T. ⬆[l, k] T ≡ U → ⦃G, K⦄ ⊢ ➡[h, o] 𝐑⦃T⦄.
+#h #o #G #L #U #H elim H -L -U
+[ #L #s #d #Hkd #K #c #l #k #_ #X #H
+ >(lift_inv_sort2 … H) -X /2 width=2 by crx_sort/
+| #I #L #L0 #W #i #HK0 #K #c #l #k #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 #c #l #k #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 #c #l #k #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 #c #l #k #_ #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 #c #l #k #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 #c #l #k #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 #c #l #k #_ #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 #c #l #k #_ #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/
+]
+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 *)
+(* *)
+(**************************************************************************)
+
+(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
+
+notation "hvbox( ⦃ term 46 G, break term 46 L ⦄ ⊢ ➡ break [ term 46 o , break term 46 h ] 𝐑 break ⦃ term 46 T ⦄ )"
+ non associative with precedence 45
+ for @{ 'PRedReducible $h $o $G $L $T }.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/statictypestar_6.ma".
+include "basic_2/grammar/genv.ma".
+include "basic_2/substitution/drop.ma".
+include "basic_2/static/sh.ma".
+
+(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************)
+
+(* activate genv *)
+inductive lstas (h): nat → relation4 genv lenv term term ≝
+| 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,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 d T U = (lstas h d G L T U).
+
+(* Basic inversion lemmas ***************************************************)
+
+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 #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,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,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 & d = 0
+ ) |
+ (∃∃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 #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 #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,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 & d = 0
+ ) |
+ (∃∃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-.
+
+lemma lstas_inv_lref1_O: ∀h,G,L,X,i. ⦃G, L⦄ ⊢ #i •*[h, 0] X →
+ (∃∃K,V,W. ⬇[i] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V •*[h, 0] W &
+ ⬆[0, i+1] W ≡ X
+ ) ∨
+ (∃∃K,W,V. ⬇[i] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W •*[h, 0] V &
+ 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 #d #_ #_ #_ <plus_n_Sm #H destruct
+qed-.
+
+(* Basic_1: was just: sty0_gen_lref *)
+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, d] V &
+ ⬆[0, i+1] V ≡ X
+ ).
+#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,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 #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,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,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 #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,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,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 #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,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,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 #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,d. ⦃G, L⦄ ⊢ ⓝW.T •*[h, d] U → ⦃G, L⦄ ⊢ T •*[h, d] U.
+/2 width=4 by lstas_inv_cast1_aux/
+qed-.
+
+(* Basic_1: removed theorems 7:
+ sty1_abbr sty1_appl sty1_bind sty1_cast2
+ sty1_correct sty1_lift sty1_trans
+*)
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/static/aaa_lift.ma".
+include "basic_2/unfold/lstas_lstas.ma".
+
+(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************)
+
+(* Properties on atomic arity assignment for terms **************************)
+
+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 #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 … d) -d
+ [ elim (IHV 0) -IHV /3 width=7 by lstas_zero, aaa_lref, ex2_intro/
+ | #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 #d elim (IHT d) -IHT
+ /3 width=7 by lstas_bind, aaa_abbr, ex2_intro/
+| #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 #d elim (IHT d) -IHT
+ /3 width=6 by lstas_appl, aaa_appl, ex2_intro/
+| #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,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-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/static/da_da.ma".
+include "basic_2/unfold/lstas_lstas.ma".
+
+(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************)
+
+(* Properties on degree assignment for terms ********************************)
+
+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 #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 #d1 #HLK #HW #IHW #d2 @(nat_ind_plus … d2) -d2
+ [ elim (IHW 0) -IHW /3 width=6 by lstas_zero, da_ldec, ex2_intro/
+ | #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 #d1 #_ #IHT #d2 elim (IHT … d2) -IHT
+ /3 width=6 by lstas_bind, da_bind, ex2_intro/
+| * #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,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,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 #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 #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,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 #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 #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 #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,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/
+| -d1 <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 "ground_2/ynat/ynat_max.ma".
+include "basic_2/substitution/drop_drop.ma".
+include "basic_2/unfold/lstas.ma".
+
+(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************)
+
+(* Properties on relocation *************************************************)
+
+(* Basic_1: was just: sty0_lift *)
+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 #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 /2 width=1 by ylt_fwd_le_succ1/ #W2 #HW12 #HWU2
+ elim (drop_trans_le … HL21 … HLK1) -L1 /2 width=2 by ylt_fwd_le/ #X #HLK2 #H
+ elim (drop_inv_skip2 … H) -H /2 width=1 by ylt_to_minus/ -Hil #K2 #V2 #HK21 #HV12 #H destruct
+ /3 width=9 by lstas_ldef/
+ | lapply (lift_trans_be … HW1 … HWU2 ? ?) -W /2 width=1 by yle_succ_dx/ #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 #l #m #HL21 #X #H #U2 #HWU2
+ >(lift_mono … HWU2 … H) -U2
+ 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 ylt_fwd_le/ #X #HLK2 #H
+ elim (drop_inv_skip2 … H) -H /2 width=1 by ylt_to_minus/ -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 #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 /2 width=1 by ylt_fwd_le_succ1/ #W #HW1 #HWU2
+ elim (drop_trans_le … HL21 … HLK1) -L1 /2 width=2 by ylt_fwd_le/ #X #HLK2 #H
+ elim (drop_inv_skip2 … H) -H /2 width=1 by ylt_to_minus/ -Hil #K2 #W2 #HK21 #HW12 #H destruct
+ /3 width=9 by lstas_succ/
+ | lapply (lift_trans_be … HW1 … HWU2 ? ?) -W /2 width=1 by yle_succ_dx/ #HW1U2
+ lapply (drop_trans_ge … HL21 … HLK1 ?) -L1 /3 width=9 by lstas_succ, drop_inv_gen/
+ ]
+| #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 #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 #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,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 #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 // <yminus_succ2 <yplus_inj >yplus_SO2 >ymax_pre_sn /3 width=8 by lstas_ldef, ylt_fwd_le_succ1, ex2_intro/
+ | lapply (drop_conf_ge … HL21 … HLK2 ?) -L2 // #HL1K2
+ elim (yle_inv_plus_inj2 … Hil) -Hil #Hlim #mi
+ elim (lift_split … HW2 l (i-m+1)) -HW2 /2 width=1 by yle_succ_dx, le_S_S/
+ #W0 #HW20 <le_plus_minus_comm /2 width=1 by yle_inv_inj/ >minus_minus_m_m /3 width=8 by lstas_ldef, yle_inv_inj, le_S, ex2_intro/
+ ]
+| #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 #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 // <yminus_succ2 <yplus_inj >yplus_SO2 >ymax_pre_sn /3 width=8 by lstas_succ, ylt_fwd_le_succ1, ex2_intro/
+ | lapply (drop_conf_ge … HL21 … HLK2 ?) -L2 // #HL1K2
+ elim (yle_inv_plus_inj2 … Hil) -Hil #Hlim #mi
+ elim (lift_split … HW2 l (i-m+1)) -HW2 /2 width=1 by yle_succ_dx, le_S_S/
+ #W0 #HW20 <le_plus_minus_comm /2 width=1 by yle_inv_inj/ >minus_minus_m_m /3 width=8 by lstas_succ, yle_inv_inj, le_S, ex2_intro/
+ ]
+| #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 #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 #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/
+]
+qed-.
+
+(* Advanced inversion lemmas ************************************************)
+
+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 #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 #_ #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 #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/
+ | #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 #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,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,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 #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 #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 #d #HLK #_ #_ #IHW1 #d2
+ @(nat_ind_plus … d2) -d2
+ [ elim (IHW1 0) -IHW1 /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 ex_intro, lstas_succ/
+ ]
+| #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-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/multiple/llpx_sn_drop.ma".
+include "basic_2/unfold/lstas.ma".
+
+(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************)
+
+(* Properties on lazy sn pointwise extensions *******************************)
+
+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 #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 #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 #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 #d #_ #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
+ /3 width=1 by llpx_sn_flat/
+| #G #Ls #V #T1 #T2 #d #_ #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
+ /3 width=1 by llpx_sn_flat/
+]
+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/unfold/lstas_lift.ma".
+
+(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************)
+
+(* Main properties **********************************************************)
+
+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 #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 #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 #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 #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,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: #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/
+| #G #L #K #W #W1 #i #HLK #_ #_ #X #H
+ 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 #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 #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 #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 #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,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 #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 #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 #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 #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,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,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/
+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 *)
+(* *)
+(**************************************************************************)
+
+(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
+
+notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 •* break [ term 46 o , break term 46 n ] break term 46 T2 )"
+ non associative with precedence 45
+ for @{ 'StaticTypeStar $h $G $L $n $T1 $T2 }.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/unfold_4.ma".
+include "basic_2/grammar/lenv_append.ma".
+include "basic_2/grammar/genv.ma".
+include "basic_2/substitution/drop.ma".
+
+(* CONTEXT-SENSITIVE UNFOLD FOR TERMS ***************************************)
+
+(* activate genv *)
+inductive unfold: relation4 genv lenv term lenv ≝
+| unfold_sort: ∀G,L,k. unfold G L (⋆k) L
+| unfold_lref: ∀I,G,L1,L2,K1,K2,V,i. ⬇[i] L1 ≡ K1. ⓑ{I}V →
+ unfold G K1 V K2 → ⬇[Ⓣ, |L2|, i] L2 ≡ K2 →
+ unfold G L1 (#i) (L1@@L2)
+| unfold_bind: ∀a,I,G,L1,L2,V,T.
+ unfold G (L1.ⓑ{I}V) T L2 → unfold G L1 (ⓑ{a,I}V.T) L2
+| unfold_flat: ∀I,G,L1,L2,V,T.
+ unfold G L1 T L2 → unfold G L1 (ⓕ{I}V.T) L2
+.
+
+interpretation "context-sensitive unfold (term)"
+ 'Unfold G L1 T L2 = (unfold G L1 T L2).
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
+
+notation "hvbox( ⦃ term 46 G, break term 46 L1 ⦄ ⊢ ⧫ * break term 46 T ≡ break term 46 L2 )"
+ non associative with precedence 45
+ for @{ 'Unfold $G $L1 $T $L2 }.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-
-notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ ➡ 𝐍 break ⦃ term 46 T ⦄ )"
- non associative with precedence 45
- for @{ 'PRedNormal $G $L $T }.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-
-notation "hvbox( ⦃ term 46 G, break term 46 L ⦄ ⊢ ➡ break [ term 46 o , break term 46 h ] 𝐍 break ⦃ term 46 T ⦄ )"
- non associative with precedence 45
- for @{ 'PRedNormal $h $o $G $L $T }.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-
-notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ ➡ 𝐈 break ⦃ term 46 T ⦄ )"
- non associative with precedence 45
- for @{ 'PRedNotReducible $G $L $T }.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-
-notation "hvbox( ⦃ term 46 G, break term 46 L ⦄ ⊢ ➡ break [ term 46 o , break term 46 h ] 𝐈 break ⦃ term 46 T ⦄ )"
- non associative with precedence 45
- for @{ 'PRedNotReducible $h $o $G $L $T }.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-
-notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ ➡ 𝐑 break ⦃ term 46 T ⦄ )"
- non associative with precedence 45
- for @{ 'PRedReducible $G $L $T }.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-
-notation "hvbox( ⦃ term 46 G, break term 46 L ⦄ ⊢ ➡ break [ term 46 o , break term 46 h ] 𝐑 break ⦃ term 46 T ⦄ )"
- non associative with precedence 45
- for @{ 'PRedReducible $h $o $G $L $T }.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-
-notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 •* break [ term 46 o , break term 46 n ] break term 46 T2 )"
- non associative with precedence 45
- for @{ 'StaticTypeStar $h $G $L $n $T1 $T2 }.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-
-notation "hvbox( ⦃ term 46 G, break term 46 L1 ⦄ ⊢ ⧫ * break term 46 T ≡ break term 46 L2 )"
- non associative with precedence 45
- for @{ 'Unfold $G $L1 $T $L2 }.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/prednotreducible_3.ma".
-include "basic_2/reduction/crr.ma".
-
-(* IRREDUCIBLE TERMS FOR CONTEXT-SENSITIVE REDUCTION ************************)
-
-definition cir: relation3 genv lenv term ≝ λG,L,T. ⦃G, L⦄ ⊢ ➡ 𝐑⦃T⦄ → ⊥.
-
-interpretation "irreducibility for context-sensitive reduction (term)"
- 'PRedNotReducible G L T = (cir G L T).
-
-(* Basic inversion lemmas ***************************************************)
-
-lemma cir_inv_delta: ∀G,L,K,V,i. ⬇[i] L ≡ K.ⓓV → ⦃G, L⦄ ⊢ ➡ 𝐈⦃#i⦄ → ⊥.
-/3 width=3 by crr_delta/ qed-.
-
-lemma cir_inv_ri2: ∀I,G,L,V,T. ri2 I → ⦃G, L⦄ ⊢ ➡ 𝐈⦃②{I}V.T⦄ → ⊥.
-/3 width=1 by crr_ri2/ qed-.
-
-lemma cir_inv_ib2: ∀a,I,G,L,V,T. ib2 a I → ⦃G, L⦄ ⊢ ➡ 𝐈⦃ⓑ{a,I}V.T⦄ →
- ⦃G, L⦄ ⊢ ➡ 𝐈⦃V⦄ ∧ ⦃G, L.ⓑ{I}V⦄ ⊢ ➡ 𝐈⦃T⦄.
-/4 width=1 by crr_ib2_sn, crr_ib2_dx, conj/ qed-.
-
-lemma cir_inv_bind: ∀a,I,G,L,V,T. ⦃G, L⦄ ⊢ ➡ 𝐈⦃ⓑ{a,I}V.T⦄ →
- ∧∧ ⦃G, L⦄ ⊢ ➡ 𝐈⦃V⦄ & ⦃G, L.ⓑ{I}V⦄ ⊢ ➡ 𝐈⦃T⦄ & ib2 a I.
-#a * [ elim a -a ]
-#G #L #V #T #H [ elim H -H /3 width=1 by crr_ri2, or_introl/ ]
-elim (cir_inv_ib2 … H) -H /3 width=1 by and3_intro, or_introl/
-qed-.
-
-lemma cir_inv_appl: ∀G,L,V,T. ⦃G, L⦄ ⊢ ➡ 𝐈⦃ⓐV.T⦄ →
- ∧∧ ⦃G, L⦄ ⊢ ➡ 𝐈⦃V⦄ & ⦃G, L⦄ ⊢ ➡ 𝐈⦃T⦄ & 𝐒⦃T⦄.
-#G #L #V #T #HVT @and3_intro /3 width=1 by crr_appl_sn, crr_appl_dx/
-generalize in match HVT; -HVT elim T -T //
-* // #a * #U #T #_ #_ #H elim H -H /2 width=1 by crr_beta, crr_theta/
-qed-.
-
-lemma cir_inv_flat: ∀I,G,L,V,T. ⦃G, L⦄ ⊢ ➡ 𝐈⦃ⓕ{I}V.T⦄ →
- ∧∧ ⦃G, L⦄ ⊢ ➡ 𝐈⦃V⦄ & ⦃G, L⦄ ⊢ ➡ 𝐈⦃T⦄ & 𝐒⦃T⦄ & I = Appl.
-* #G #L #V #T #H
-[ elim (cir_inv_appl … H) -H /2 width=1 by and4_intro/
-| elim (cir_inv_ri2 … H) -H //
-]
-qed-.
-
-(* Basic properties *********************************************************)
-
-lemma cir_sort: ∀G,L,s. ⦃G, L⦄ ⊢ ➡ 𝐈⦃⋆s⦄.
-/2 width=4 by crr_inv_sort/ qed.
-
-lemma cir_gref: ∀G,L,p. ⦃G, L⦄ ⊢ ➡ 𝐈⦃§p⦄.
-/2 width=4 by crr_inv_gref/ qed.
-
-lemma tir_atom: ∀G,I. ⦃G, ⋆⦄ ⊢ ➡ 𝐈⦃⓪{I}⦄.
-/2 width=3 by trr_inv_atom/ qed.
-
-lemma cir_ib2: ∀a,I,G,L,V,T.
- ib2 a I → ⦃G, L⦄ ⊢ ➡ 𝐈⦃V⦄ → ⦃G, L.ⓑ{I}V⦄ ⊢ ➡ 𝐈⦃T⦄ → ⦃G, L⦄ ⊢ ➡ 𝐈⦃ⓑ{a,I}V.T⦄.
-#a #I #G #L #V #T #HI #HV #HT #H
-elim (crr_inv_ib2 … HI H) -HI -H /2 width=1 by/
-qed.
-
-lemma cir_appl: ∀G,L,V,T. ⦃G, L⦄ ⊢ ➡ 𝐈⦃V⦄ → ⦃G, L⦄ ⊢ ➡ 𝐈⦃T⦄ → 𝐒⦃T⦄ → ⦃G, L⦄ ⊢ ➡ 𝐈⦃ⓐV.T⦄.
-#G #L #V #T #HV #HT #H1 #H2
-elim (crr_inv_appl … H2) -H2 /2 width=1 by/
-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/crr_lift.ma".
-include "basic_2/reduction/cir.ma".
-
-(* IRREDUCIBLE TERMS FOR CONTEXT-SENSITIVE REDUCTION ************************)
-
-(* Properties on relocation *************************************************)
-
-lemma cir_lift: ∀G,K,T. ⦃G, K⦄ ⊢ ➡ 𝐈⦃T⦄ → ∀L,c,l,k. ⬇[c, l, k] L ≡ K →
- ∀U. ⬆[l, k] T ≡ U → ⦃G, L⦄ ⊢ ➡ 𝐈⦃U⦄.
-/3 width=8 by crr_inv_lift/ qed.
-
-lemma cir_inv_lift: ∀G,L,U. ⦃G, L⦄ ⊢ ➡ 𝐈⦃U⦄ → ∀K,c,l,k. ⬇[c, l, k] L ≡ K →
- ∀T. ⬆[l, k] T ≡ U → ⦃G, K⦄ ⊢ ➡ 𝐈⦃T⦄.
-/3 width=8 by crr_lift/ qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/prednotreducible_5.ma".
-include "basic_2/reduction/cir.ma".
-include "basic_2/reduction/crx.ma".
-
-(* IRREDUCIBLE TERMS FOR CONTEXT-SENSITIVE EXTENDED REDUCTION ***************)
-
-definition cix: ∀h. sd h → relation3 genv lenv term ≝
- λh,o,G,L,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃T⦄ → ⊥.
-
-interpretation "irreducibility for context-sensitive extended reduction (term)"
- 'PRedNotReducible h o G L T = (cix h o G L T).
-
-(* Basic inversion lemmas ***************************************************)
-
-lemma cix_inv_sort: ∀h,o,G,L,s,d. deg h o s (d+1) → ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃⋆s⦄ → ⊥.
-/3 width=2 by crx_sort/ qed-.
-
-lemma cix_inv_delta: ∀h,o,I,G,L,K,V,i. ⬇[i] L ≡ K.ⓑ{I}V → ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃#i⦄ → ⊥.
-/3 width=4 by crx_delta/ qed-.
-
-lemma cix_inv_ri2: ∀h,o,I,G,L,V,T. ri2 I → ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃②{I}V.T⦄ → ⊥.
-/3 width=1 by crx_ri2/ qed-.
-
-lemma cix_inv_ib2: ∀h,o,a,I,G,L,V,T. ib2 a I → ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃ⓑ{a,I}V.T⦄ →
- ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃V⦄ ∧ ⦃G, L.ⓑ{I}V⦄ ⊢ ➡[h, o] 𝐈⦃T⦄.
-/4 width=1 by crx_ib2_sn, crx_ib2_dx, conj/ qed-.
-
-lemma cix_inv_bind: ∀h,o,a,I,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃ⓑ{a,I}V.T⦄ →
- ∧∧ ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃V⦄ & ⦃G, L.ⓑ{I}V⦄ ⊢ ➡[h, o] 𝐈⦃T⦄ & ib2 a I.
-#h #o #a * [ elim a -a ]
-#G #L #V #T #H [ elim H -H /3 width=1 by crx_ri2, or_introl/ ]
-elim (cix_inv_ib2 … H) -H /3 width=1 by and3_intro, or_introl/
-qed-.
-
-lemma cix_inv_appl: ∀h,o,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃ⓐV.T⦄ →
- ∧∧ ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃V⦄ & ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃T⦄ & 𝐒⦃T⦄.
-#h #o #G #L #V #T #HVT @and3_intro /3 width=1 by crx_appl_sn, crx_appl_dx/
-generalize in match HVT; -HVT elim T -T //
-* // #a * #U #T #_ #_ #H elim H -H /2 width=1 by crx_beta, crx_theta/
-qed-.
-
-lemma cix_inv_flat: ∀h,o,I,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃ⓕ{I}V.T⦄ →
- ∧∧ ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃V⦄ & ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃T⦄ & 𝐒⦃T⦄ & I = Appl.
-#h #o * #G #L #V #T #H
-[ elim (cix_inv_appl … H) -H /2 width=1 by and4_intro/
-| elim (cix_inv_ri2 … H) -H //
-]
-qed-.
-
-(* Basic forward lemmas *****************************************************)
-
-lemma cix_inv_cir: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃T⦄ → ⦃G, L⦄ ⊢ ➡ 𝐈⦃T⦄.
-/3 width=1 by crr_crx/ qed-.
-
-(* Basic properties *********************************************************)
-
-lemma cix_sort: ∀h,o,G,L,s. deg h o s 0 → ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃⋆s⦄.
-#h #o #G #L #s #Hk #H elim (crx_inv_sort … H) -L #d #Hkd
-lapply (deg_mono … Hk Hkd) -h -s <plus_n_Sm #H destruct
-qed.
-
-lemma tix_lref: ∀h,o,G,i. ⦃G, ⋆⦄ ⊢ ➡[h, o] 𝐈⦃#i⦄.
-#h #o #G #i #H elim (trx_inv_atom … H) -H #s #d #_ #H destruct
-qed.
-
-lemma cix_gref: ∀h,o,G,L,p. ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃§p⦄.
-#h #o #G #L #p #H elim (crx_inv_gref … H)
-qed.
-
-lemma cix_ib2: ∀h,o,a,I,G,L,V,T. ib2 a I → ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃V⦄ → ⦃G, L.ⓑ{I}V⦄ ⊢ ➡[h, o] 𝐈⦃T⦄ →
- ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃ⓑ{a,I}V.T⦄.
-#h #o #a #I #G #L #V #T #HI #HV #HT #H
-elim (crx_inv_ib2 … HI H) -HI -H /2 width=1 by/
-qed.
-
-lemma cix_appl: ∀h,o,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃V⦄ → ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃T⦄ → 𝐒⦃T⦄ → ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃ⓐV.T⦄.
-#h #o #G #L #V #T #HV #HT #H1 #H2
-elim (crx_inv_appl … H2) -H2 /2 width=1 by/
-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/crx_lift.ma".
-include "basic_2/reduction/cix.ma".
-
-(* IRREDUCIBLE TERMS FOR CONTEXT-SENSITIVE EXTENDED REDUCTION ***************)
-
-(* Advanced properties ******************************************************)
-
-lemma cix_lref: ∀h,o,G,L,i. ⬇[i] L ≡ ⋆ → ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃#i⦄.
-#h #o #G #L #i #HL #H elim (crx_inv_lref … H) -h #I #K #V #HLK
-lapply (drop_mono … HLK … HL) -L -i #H destruct
-qed.
-
-(* Properties on relocation *************************************************)
-
-lemma cix_lift: ∀h,o,G,K,T. ⦃G, K⦄ ⊢ ➡[h, o] 𝐈⦃T⦄ → ∀L,c,l,k. ⬇[c, l, k] L ≡ K →
- ∀U. ⬆[l, k] T ≡ U → ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃U⦄.
-/3 width=8 by crx_inv_lift/ qed.
-
-lemma cix_inv_lift: ∀h,o,G,L,U. ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃U⦄ → ∀K,c,l,k. ⬇[c, l, k] L ≡ K →
- ∀T. ⬆[l, k] T ≡ U → ⦃G, K⦄ ⊢ ➡[h, o] 𝐈⦃T⦄.
-/3 width=8 by crx_lift/ qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/prednormal_3.ma".
-include "basic_2/reduction/cpr.ma".
-
-(* NORMAL TERMS FOR CONTEXT-SENSITIVE REDUCTION *****************************)
-
-definition cnr: relation3 genv lenv term ≝ λG,L. NF … (cpr G L) (eq …).
-
-interpretation
- "normality for context-sensitive reduction (term)"
- 'PRedNormal G L T = (cnr G L T).
-
-(* Basic inversion lemmas ***************************************************)
-
-lemma cnr_inv_delta: ∀G,L,K,V,i. ⬇[i] L ≡ K.ⓓV → ⦃G, L⦄ ⊢ ➡ 𝐍⦃#i⦄ → ⊥.
-#G #L #K #V #i #HLK #H
-elim (lift_total V 0 (i+1)) #W #HVW
-lapply (H W ?) -H [ /3 width=6 by cpr_delta/ ] -HLK #H destruct
-elim (lift_inv_lref2_be … HVW) -HVW /2 width=1 by ylt_inj/
-qed-.
-
-lemma cnr_inv_abst: ∀a,G,L,V,T. ⦃G, L⦄ ⊢ ➡ 𝐍⦃ⓛ{a}V.T⦄ → ⦃G, L⦄ ⊢ ➡ 𝐍⦃V⦄ ∧ ⦃G, L.ⓛV⦄ ⊢ ➡ 𝐍⦃T⦄.
-#a #G #L #V1 #T1 #HVT1 @conj
-[ #V2 #HV2 lapply (HVT1 (ⓛ{a}V2.T1) ?) -HVT1 /2 width=2 by cpr_pair_sn/ -HV2 #H destruct //
-| #T2 #HT2 lapply (HVT1 (ⓛ{a}V1.T2) ?) -HVT1 /2 width=2 by cpr_bind/ -HT2 #H destruct //
-]
-qed-.
-
-lemma cnr_inv_abbr: ∀G,L,V,T. ⦃G, L⦄ ⊢ ➡ 𝐍⦃-ⓓV.T⦄ → ⦃G, L⦄ ⊢ ➡ 𝐍⦃V⦄ ∧ ⦃G, L.ⓓV⦄ ⊢ ➡ 𝐍⦃T⦄.
-#G #L #V1 #T1 #HVT1 @conj
-[ #V2 #HV2 lapply (HVT1 (-ⓓV2.T1) ?) -HVT1 /2 width=2 by cpr_pair_sn/ -HV2 #H destruct //
-| #T2 #HT2 lapply (HVT1 (-ⓓV1.T2) ?) -HVT1 /2 width=2 by cpr_bind/ -HT2 #H destruct //
-]
-qed-.
-
-lemma cnr_inv_zeta: ∀G,L,V,T. ⦃G, L⦄ ⊢ ➡ 𝐍⦃+ⓓV.T⦄ → ⊥.
-#G #L #V #T #H elim (is_lift_dec T 0 1)
-[ * #U #HTU
- lapply (H U ?) -H /2 width=3 by cpr_zeta/ #H destruct
- elim (lift_inv_pair_xy_y … HTU)
-| #HT
- elim (cpr_delift G (⋆) V T (⋆. ⓓV) 0) //
- #T2 #T1 #HT2 #HT12 lapply (H (+ⓓV.T2) ?) -H /4 width=1 by tpr_cpr, cpr_bind/ -HT2
- #H destruct /3 width=2 by ex_intro/
-]
-qed-.
-
-lemma cnr_inv_appl: ∀G,L,V,T. ⦃G, L⦄ ⊢ ➡ 𝐍⦃ⓐV.T⦄ → ∧∧ ⦃G, L⦄ ⊢ ➡ 𝐍⦃V⦄ & ⦃G, L⦄ ⊢ ➡ 𝐍⦃T⦄ & 𝐒⦃T⦄.
-#G #L #V1 #T1 #HVT1 @and3_intro
-[ #V2 #HV2 lapply (HVT1 (ⓐV2.T1) ?) -HVT1 /2 width=1 by cpr_pair_sn/ -HV2 #H destruct //
-| #T2 #HT2 lapply (HVT1 (ⓐV1.T2) ?) -HVT1 /2 width=1 by cpr_flat/ -HT2 #H destruct //
-| generalize in match HVT1; -HVT1 elim T1 -T1 * // #a * #W1 #U1 #_ #_ #H
- [ elim (lift_total V1 0 1) #V2 #HV12
- lapply (H (ⓓ{a}W1.ⓐV2.U1) ?) -H /3 width=3 by tpr_cpr, cpr_theta/ -HV12 #H destruct
- | lapply (H (ⓓ{a}ⓝW1.V1.U1) ?) -H /3 width=1 by tpr_cpr, cpr_beta/ #H destruct
-]
-qed-.
-
-lemma cnr_inv_eps: ∀G,L,V,T. ⦃G, L⦄ ⊢ ➡ 𝐍⦃ⓝV.T⦄ → ⊥.
-#G #L #V #T #H lapply (H T ?) -H
-/2 width=4 by cpr_eps, discr_tpair_xy_y/
-qed-.
-
-(* Basic properties *********************************************************)
-
-(* Basic_1: was: nf2_sort *)
-lemma cnr_sort: ∀G,L,s. ⦃G, L⦄ ⊢ ➡ 𝐍⦃⋆s⦄.
-#G #L #s #X #H
->(cpr_inv_sort1 … H) //
-qed.
-
-lemma cnr_lref_free: ∀G,L,i. |L| ≤ i → ⦃G, L⦄ ⊢ ➡ 𝐍⦃#i⦄.
-#G #L #i #Hi #X #H elim (cpr_inv_lref1 … H) -H // *
-#K #V1 #V2 #HLK lapply (drop_fwd_length_lt2 … HLK) -HLK
-#H elim (lt_refl_false i) /2 width=3 by lt_to_le_to_lt/
-qed.
-
-(* Basic_1: was only: nf2_csort_lref *)
-lemma cnr_lref_atom: ∀G,L,i. ⬇[i] L ≡ ⋆ → ⦃G, L⦄ ⊢ ➡ 𝐍⦃#i⦄.
-#G #L #i #HL @cnr_lref_free >(drop_fwd_length … HL) -HL //
-qed.
-
-(* Basic_1: was: nf2_abst *)
-lemma cnr_abst: ∀a,G,L,W,T. ⦃G, L⦄ ⊢ ➡ 𝐍⦃W⦄ → ⦃G, L.ⓛW⦄ ⊢ ➡ 𝐍⦃T⦄ → ⦃G, L⦄ ⊢ ➡ 𝐍⦃ⓛ{a}W.T⦄.
-#a #G #L #W #T #HW #HT #X #H
-elim (cpr_inv_abst1 … H) -H #W0 #T0 #HW0 #HT0 #H destruct
->(HW … HW0) -W0 >(HT … HT0) -T0 //
-qed.
-
-(* Basic_1: was only: nf2_appl_lref *)
-lemma cnr_appl_simple: ∀G,L,V,T. ⦃G, L⦄ ⊢ ➡ 𝐍⦃V⦄ → ⦃G, L⦄ ⊢ ➡ 𝐍⦃T⦄ → 𝐒⦃T⦄ → ⦃G, L⦄ ⊢ ➡ 𝐍⦃ⓐV.T⦄.
-#G #L #V #T #HV #HT #HS #X #H
-elim (cpr_inv_appl1_simple … H) -H // #V0 #T0 #HV0 #HT0 #H destruct
->(HV … HV0) -V0 >(HT … HT0) -T0 //
-qed.
-
-(* Basic_1: was: nf2_dec *)
-axiom cnr_dec: ∀G,L,T1. ⦃G, L⦄ ⊢ ➡ 𝐍⦃T1⦄ ∨
- ∃∃T2. ⦃G, L⦄ ⊢ T1 ➡ T2 & (T1 = T2 → ⊥).
-
-(* Basic_1: removed theorems 1: nf2_abst_shift *)
+++ /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/cpr_cir.ma".
-include "basic_2/reduction/cnr_crr.ma".
-
-(* NORMAL TERMS FOR CONTEXT-SENSITIVE REDUCTION *****************************)
-
-(* Main properties on irreducibility ****************************************)
-
-theorem cir_cnr: ∀G,L,T. ⦃G, L⦄ ⊢ ➡ 𝐈⦃T⦄ → ⦃G, L⦄ ⊢ ➡ 𝐍⦃T⦄.
-/2 width=4 by cpr_fwd_cir/ qed.
-
-(* Main inversion lemmas on irreducibility **********************************)
-
-theorem cnr_inv_cir: ∀G,L,T. ⦃G, L⦄ ⊢ ➡ 𝐍⦃T⦄ → ⦃G, L⦄ ⊢ ➡ 𝐈⦃T⦄.
-/2 width=5 by cnr_inv_crr/ 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/crr.ma".
-include "basic_2/reduction/cnr.ma".
-
-(* NORMAL TERMS FOR CONTEXT-SENSITIVE REDUCTION *****************************)
-
-(* Advanced inversion lemmas on reducibility ********************************)
-
-(* Note: this property is unusual *)
-lemma cnr_inv_crr: ∀G,L,T. ⦃G, L⦄ ⊢ ➡ 𝐑⦃T⦄ → ⦃G, L⦄ ⊢ ➡ 𝐍⦃T⦄ → ⊥.
-#G #L #T #H elim H -L -T
-[ #L #K #V #i #HLK #H
- elim (cnr_inv_delta … HLK H)
-| #L #V #T #_ #IHV #H
- elim (cnr_inv_appl … H) -H /2 width=1 by/
-| #L #V #T #_ #IHT #H
- elim (cnr_inv_appl … H) -H /2 width=1 by/
-| #I #L #V #T * #H1 #H2 destruct
- [ elim (cnr_inv_zeta … H2)
- | elim (cnr_inv_eps … H2)
- ]
-|5,6: #a * [ elim a ] #L #V #T * #H1 #_ #IH #H2 destruct
- [1,3: elim (cnr_inv_abbr … H2) -H2 /2 width=1 by/
- |*: elim (cnr_inv_abst … H2) -H2 /2 width=1 by/
- ]
-| #a #L #V #W #T #H
- elim (cnr_inv_appl … H) -H #_ #_ #H
- elim (simple_inv_bind … H)
-| #a #L #V #W #T #H
- elim (cnr_inv_appl … H) -H #_ #_ #H
- elim (simple_inv_bind … H)
-]
-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/cpr_lift.ma".
-include "basic_2/reduction/cnr.ma".
-
-(* NORMAL TERMS FOR CONTEXT-SENSITIVE REDUCTION *****************************)
-
-(* Advanced properties ******************************************************)
-
-(* Basic_1: was: nf2_lref_abst *)
-lemma cnr_lref_abst: ∀G,L,K,V,i. ⬇[i] L ≡ K. ⓛV → ⦃G, L⦄ ⊢ ➡ 𝐍⦃#i⦄.
-#G #L #K #V #i #HLK #X #H
-elim (cpr_inv_lref1 … H) -H // *
-#K0 #V1 #V2 #HLK0 #_ #_
-lapply (drop_mono … HLK … HLK0) -L #H destruct
-qed.
-
-(* Relocation properties ****************************************************)
-
-(* Basic_1: was: nf2_lift *)
-lemma cnr_lift: ∀G,L0,L,T,T0,c,l,k. ⦃G, L⦄ ⊢ ➡ 𝐍⦃T⦄ →
- ⬇[c, l, k] L0 ≡ L → ⬆[l, k] T ≡ T0 → ⦃G, L0⦄ ⊢ ➡ 𝐍⦃T0⦄.
-#G #L0 #L #T #T0 #c #l #k #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,c,l,k. ⦃G, L0⦄ ⊢ ➡ 𝐍⦃T0⦄ →
- ⬇[c, l, k] L0 ≡ L → ⬆[l, k] T ≡ T0 → ⦃G, L⦄ ⊢ ➡ 𝐍⦃T⦄.
-#G #L0 #L #T #T0 #c #l #k #HLT0 #HL0 #HT0 #X #H
-elim (lift_total X l k) #X0 #HX0
-lapply (cpr_lift … H … HL0 … HT0 … HX0) -L #HTX0
->(HLT0 … HTX0) in HX0; -L0 -X0 #H
->(lift_inj … H … HT0) -T0 -X -c -l -k //
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/prednormal_5.ma".
-include "basic_2/reduction/cnr.ma".
-include "basic_2/reduction/cpx.ma".
-
-(* NORMAL TERMS FOR CONTEXT-SENSITIVE EXTENDED REDUCTION ********************)
-
-definition cnx: ∀h. sd h → relation3 genv lenv term ≝
- λh,o,G,L. NF … (cpx h o G L) (eq …).
-
-interpretation
- "normality for context-sensitive extended reduction (term)"
- 'PRedNormal h o L T = (cnx h o L T).
-
-(* Basic inversion lemmas ***************************************************)
-
-lemma cnx_inv_sort: ∀h,o,G,L,s. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃⋆s⦄ → deg h o s 0.
-#h #o #G #L #s #H elim (deg_total h o s)
-#d @(nat_ind_plus … d) -d // #d #_ #Hkd
-lapply (H (⋆(next h s)) ?) -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 s) >H -H #H elim (lt_refl_false … H)
-qed-.
-
-lemma cnx_inv_delta: ∀h,o,I,G,L,K,V,i. ⬇[i] L ≡ K.ⓑ{I}V → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃#i⦄ → ⊥.
-#h #o #I #G #L #K #V #i #HLK #H
-elim (lift_total V 0 (i+1)) #W #HVW
-lapply (H W ?) -H [ /3 width=7 by cpx_delta/ ] -HLK #H destruct
-elim (lift_inv_lref2_be … HVW) -HVW /2 width=1 by ylt_inj/
-qed-.
-
-lemma cnx_inv_abst: ∀h,o,a,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃ⓛ{a}V.T⦄ →
- ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃V⦄ ∧ ⦃G, L.ⓛV⦄ ⊢ ➡[h, o] 𝐍⦃T⦄.
-#h #o #a #G #L #V1 #T1 #HVT1 @conj
-[ #V2 #HV2 lapply (HVT1 (ⓛ{a}V2.T1) ?) -HVT1 /2 width=2 by cpx_pair_sn/ -HV2 #H destruct //
-| #T2 #HT2 lapply (HVT1 (ⓛ{a}V1.T2) ?) -HVT1 /2 width=2 by cpx_bind/ -HT2 #H destruct //
-]
-qed-.
-
-lemma cnx_inv_abbr: ∀h,o,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃-ⓓV.T⦄ →
- ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃V⦄ ∧ ⦃G, L.ⓓV⦄ ⊢ ➡[h, o] 𝐍⦃T⦄.
-#h #o #G #L #V1 #T1 #HVT1 @conj
-[ #V2 #HV2 lapply (HVT1 (-ⓓV2.T1) ?) -HVT1 /2 width=2 by cpx_pair_sn/ -HV2 #H destruct //
-| #T2 #HT2 lapply (HVT1 (-ⓓV1.T2) ?) -HVT1 /2 width=2 by cpx_bind/ -HT2 #H destruct //
-]
-qed-.
-
-lemma cnx_inv_zeta: ∀h,o,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃+ⓓV.T⦄ → ⊥.
-#h #o #G #L #V #T #H elim (is_lift_dec T 0 1)
-[ * #U #HTU
- lapply (H U ?) -H /2 width=3 by cpx_zeta/ #H destruct
- elim (lift_inv_pair_xy_y … HTU)
-| #HT
- elim (cpr_delift G(⋆) V T (⋆.ⓓV) 0) // #T2 #T1 #HT2 #HT12
- lapply (H (+ⓓV.T2) ?) -H /5 width=1 by cpr_cpx, tpr_cpr, cpr_bind/ -HT2
- #H destruct /3 width=2 by ex_intro/
-]
-qed-.
-
-lemma cnx_inv_appl: ∀h,o,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃ⓐV.T⦄ →
- ∧∧ ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃V⦄ & ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ & 𝐒⦃T⦄.
-#h #o #G #L #V1 #T1 #HVT1 @and3_intro
-[ #V2 #HV2 lapply (HVT1 (ⓐV2.T1) ?) -HVT1 /2 width=1 by cpx_pair_sn/ -HV2 #H destruct //
-| #T2 #HT2 lapply (HVT1 (ⓐV1.T2) ?) -HVT1 /2 width=1 by cpx_flat/ -HT2 #H destruct //
-| generalize in match HVT1; -HVT1 elim T1 -T1 * // #a * #W1 #U1 #_ #_ #H
- [ elim (lift_total V1 0 1) #V2 #HV12
- lapply (H (ⓓ{a}W1.ⓐV2.U1) ?) -H /3 width=3 by cpr_cpx, cpr_theta/ -HV12 #H destruct
- | lapply (H (ⓓ{a}ⓝW1.V1.U1) ?) -H /3 width=1 by cpr_cpx, cpr_beta/ #H destruct
- ]
-]
-qed-.
-
-lemma cnx_inv_eps: ∀h,o,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃ⓝV.T⦄ → ⊥.
-#h #o #G #L #V #T #H lapply (H T ?) -H
-/2 width=4 by cpx_eps, discr_tpair_xy_y/
-qed-.
-
-(* Basic forward lemmas *****************************************************)
-
-lemma cnx_fwd_cnr: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ → ⦃G, L⦄ ⊢ ➡ 𝐍⦃T⦄.
-#h #o #G #L #T #H #U #HTU
-@H /2 width=1 by cpr_cpx/ (**) (* auto fails because a δ-expansion gets in the way *)
-qed-.
-
-(* Basic properties *********************************************************)
-
-lemma cnx_sort: ∀h,o,G,L,s. deg h o s 0 → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃⋆s⦄.
-#h #o #G #L #s #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,o,G,L,s,d. deg h o s d → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃⋆((next h)^d s)⦄.
-#h #o #G #L #s #d #Hkd
-lapply (deg_iter … d Hkd) -Hkd <minus_n_n /2 width=6 by cnx_sort/
-qed.
-
-lemma cnx_lref_free: ∀h,o,G,L,i. |L| ≤ i → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃#i⦄.
-#h #o #G #L #i #Hi #X #H elim (cpx_inv_lref1 … H) -H // *
-#I #K #V1 #V2 #HLK lapply (drop_fwd_length_lt2 … HLK) -HLK
-#H elim (lt_refl_false i) /2 width=3 by lt_to_le_to_lt/
-qed.
-
-lemma cnx_lref_atom: ∀h,o,G,L,i. ⬇[i] L ≡ ⋆ → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃#i⦄.
-#h #o #G #L #i #HL @cnx_lref_free >(drop_fwd_length … HL) -HL //
-qed.
-
-lemma cnx_abst: ∀h,o,a,G,L,W,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃W⦄ → ⦃G, L.ⓛW⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ →
- ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃ⓛ{a}W.T⦄.
-#h #o #a #G #L #W #T #HW #HT #X #H
-elim (cpx_inv_abst1 … H) -H #W0 #T0 #HW0 #HT0 #H destruct
->(HW … HW0) -W0 >(HT … HT0) -T0 //
-qed.
-
-lemma cnx_appl_simple: ∀h,o,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃V⦄ → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ → 𝐒⦃T⦄ →
- ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃ⓐV.T⦄.
-#h #o #G #L #V #T #HV #HT #HS #X #H
-elim (cpx_inv_appl1_simple … H) -H // #V0 #T0 #HV0 #HT0 #H destruct
->(HV … HV0) -V0 >(HT … HT0) -T0 //
-qed.
-
-axiom cnx_dec: ∀h,o,G,L,T1. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T1⦄ ∨
- ∃∃T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 & (T1 = T2 → ⊥).
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/reduction/cpx_cix.ma".
-include "basic_2/reduction/cnx_crx.ma".
-
-(* NORMAL TERMS FOR CONTEXT-SENSITIVE EXTENDED REDUCTION ********************)
-
-(* Main properties on irreducibility ****************************************)
-
-theorem cix_cnx: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃T⦄ → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄.
-/2 width=6 by cpx_fwd_cix/ qed.
-
-(* Main inversion lemmas on irreducibility **********************************)
-
-theorem cnx_inv_cix: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ → ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃T⦄.
-/2 width=7 by cnx_inv_crx/ 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/crx.ma".
-include "basic_2/reduction/cnx.ma".
-
-(* NORMAL TERMS FOR CONTEXT-SENSITIVE EXTENDED REDUCTION ********************)
-
-(* Advanced inversion lemmas on reducibility ********************************)
-
-(* Note: this property is unusual *)
-lemma cnx_inv_crx: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃T⦄ → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ → ⊥.
-#h #o #G #L #T #H elim H -L -T
-[ #L #s #d #Hkd #H
- lapply (cnx_inv_sort … H) -H #H
- lapply (deg_mono … H Hkd) -h -L -s <plus_n_Sm #H destruct
-| #I #L #K #V #i #HLK #H
- elim (cnx_inv_delta … HLK H)
-| #L #V #T #_ #IHV #H
- elim (cnx_inv_appl … H) -H /2 width=1 by/
-| #L #V #T #_ #IHT #H
- elim (cnx_inv_appl … H) -H /2 width=1 by/
-| #I #L #V #T * #H1 #H2 destruct
- [ elim (cnx_inv_zeta … H2)
- | elim (cnx_inv_eps … H2)
- ]
-|6,7: #a * [ elim a ] #L #V #T * #H1 #_ #IH #H2 destruct
- [1,3: elim (cnx_inv_abbr … H2) -H2 /2 width=1 by/
- |*: elim (cnx_inv_abst … H2) -H2 /2 width=1 by/
- ]
-| #a #L #V #W #T #H
- elim (cnx_inv_appl … H) -H #_ #_ #H
- elim (simple_inv_bind … H)
-| #a #L #V #W #T #H
- elim (cnx_inv_appl … H) -H #_ #_ #H
- elim (simple_inv_bind … H)
-]
-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_lift.ma".
-include "basic_2/reduction/cnx.ma".
-
-(* NORMAL TERMS FOR CONTEXT-SENSITIVE EXTENDED REDUCTION ********************)
-
-(* Relocation properties ****************************************************)
-
-lemma cnx_lift: ∀h,o,G,L0,L,T,T0,c,l,k. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ → ⬇[c, l, k] L0 ≡ L →
- ⬆[l, k] T ≡ T0 → ⦃G, L0⦄ ⊢ ➡[h, o] 𝐍⦃T0⦄.
-#h #o #G #L0 #L #T #T0 #c #l #k #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,o,G,L0,L,T,T0,c,l,k. ⦃G, L0⦄ ⊢ ➡[h, o] 𝐍⦃T0⦄ → ⬇[c, l, k] L0 ≡ L →
- ⬆[l, k] T ≡ T0 → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄.
-#h #o #G #L0 #L #T #T0 #c #l #k #HLT0 #HL0 #HT0 #X #H
-elim (lift_total X l k) #X0 #HX0
-lapply (cpx_lift … H … HL0 … HT0 … HX0) -L #HTX0
->(HLT0 … HTX0) in HX0; -L0 -X0 #H
->(lift_inj … H … HT0) -T0 -X -l -k //
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/pred_4.ma".
-include "basic_2/static/lsubr.ma".
-include "basic_2/unfold/lstas.ma".
-
-(* CONTEXT-SENSITIVE PARALLEL REDUCTION FOR TERMS ***************************)
-
-(* activate genv *)
-(* Basic_1: includes: pr0_delta1 pr2_delta1 pr2_thin_dx *)
-(* Note: cpr_flat: does not hold in basic_1 *)
-inductive cpr: relation4 genv lenv term term ≝
-| cpr_atom : ∀I,G,L. cpr G L (⓪{I}) (⓪{I})
-| cpr_delta: ∀G,L,K,V,V2,W2,i.
- ⬇[i] L ≡ K. ⓓV → cpr G K V V2 →
- ⬆[0, i + 1] V2 ≡ W2 → cpr G L (#i) W2
-| cpr_bind : ∀a,I,G,L,V1,V2,T1,T2.
- cpr G L V1 V2 → cpr G (L.ⓑ{I}V1) T1 T2 →
- cpr G L (ⓑ{a,I}V1.T1) (ⓑ{a,I}V2.T2)
-| cpr_flat : ∀I,G,L,V1,V2,T1,T2.
- cpr G L V1 V2 → cpr G L T1 T2 →
- cpr G L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2)
-| cpr_zeta : ∀G,L,V,T1,T,T2. cpr G (L.ⓓV) T1 T →
- ⬆[0, 1] T2 ≡ T → cpr G L (+ⓓV.T1) T2
-| cpr_eps : ∀G,L,V,T1,T2. cpr G L T1 T2 → cpr G L (ⓝV.T1) T2
-| cpr_beta : ∀a,G,L,V1,V2,W1,W2,T1,T2.
- cpr G L V1 V2 → cpr G L W1 W2 → cpr G (L.ⓛW1) T1 T2 →
- cpr G L (ⓐV1.ⓛ{a}W1.T1) (ⓓ{a}ⓝW2.V2.T2)
-| cpr_theta: ∀a,G,L,V1,V,V2,W1,W2,T1,T2.
- cpr G L V1 V → ⬆[0, 1] V ≡ V2 → cpr G L W1 W2 → cpr G (L.ⓓW1) T1 T2 →
- cpr G L (ⓐV1.ⓓ{a}W1.T1) (ⓓ{a}W2.ⓐV2.T2)
-.
-
-interpretation "context-sensitive parallel reduction (term)"
- 'PRed G L T1 T2 = (cpr G L T1 T2).
-
-(* Basic properties *********************************************************)
-
-lemma lsubr_cpr_trans: ∀G. lsub_trans … (cpr G) lsubr.
-#G #L1 #T1 #T2 #H elim H -G -L1 -T1 -T2
-[ //
-| #G #L1 #K1 #V1 #V2 #W2 #i #HLK1 #_ #HVW2 #IHV12 #L2 #HL12
- elim (lsubr_fwd_drop2_abbr … HL12 … HLK1) -L1 *
- /3 width=6 by cpr_delta/
-|3,7: /4 width=1 by lsubr_pair, cpr_bind, cpr_beta/
-|4,6: /3 width=1 by cpr_flat, cpr_eps/
-|5,8: /4 width=3 by lsubr_pair, cpr_zeta, cpr_theta/
-]
-qed-.
-
-(* Basic_1: was by definition: pr2_free *)
-lemma tpr_cpr: ∀G,T1,T2. ⦃G, ⋆⦄ ⊢ T1 ➡ T2 → ∀L. ⦃G, L⦄ ⊢ T1 ➡ T2.
-#G #T1 #T2 #HT12 #L
-lapply (lsubr_cpr_trans … HT12 L ?) //
-qed.
-
-(* Basic_1: includes by definition: pr0_refl *)
-lemma cpr_refl: ∀G,T,L. ⦃G, L⦄ ⊢ T ➡ T.
-#G #T elim T -T // * /2 width=1 by cpr_bind, cpr_flat/
-qed.
-
-(* Basic_1: was: pr2_head_1 *)
-lemma cpr_pair_sn: ∀I,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡ V2 →
- ∀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,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 #l #HLK elim (lt_or_eq_or_gt i l)
- #Hil [1,3: /4 width=4 by lift_lref_ge_minus, lift_lref_lt, ylt_inj, yle_inj, 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 #l #HLK
- elim (IHW1 … HLK) -IHW1 #W2 #W #HW12 #HW2
- [ 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,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 #K #V1 #V2 #W2 #i #d #HLK #_ #HVW2 #IHV12 #H destruct
- /3 width=6 by cpr_delta/
-| #G #L #K #V1 #V2 #W2 #i #d #_ #_ #_ #_ <plus_n_Sm #H destruct
-]
-qed-.
-
-lemma lstas_cpr: ∀h,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 •*[h, 0] T2 → ⦃G, L⦄ ⊢ T1 ➡ T2.
-/2 width=4 by lstas_cpr_aux/ qed.
-
-(* Basic inversion lemmas ***************************************************)
-
-fact cpr_inv_atom1_aux: ∀G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ∀I. T1 = ⓪{I} →
- T2 = ⓪{I} ∨
- ∃∃K,V,V2,i. ⬇[i] L ≡ K. ⓓV & ⦃G, K⦄ ⊢ V ➡ V2 &
- ⬆[O, i + 1] V2 ≡ T2 & I = LRef i.
-#G #L #T1 #T2 * -G -L -T1 -T2
-[ #I #G #L #J #H destruct /2 width=1 by or_introl/
-| #L #G #K #V #V2 #T2 #i #HLK #HV2 #HVT2 #J #H destruct /3 width=8 by ex4_4_intro, or_intror/
-| #a #I #G #L #V1 #V2 #T1 #T2 #_ #_ #J #H destruct
-| #I #G #L #V1 #V2 #T1 #T2 #_ #_ #J #H destruct
-| #G #L #V #T1 #T #T2 #_ #_ #J #H destruct
-| #G #L #V #T1 #T2 #_ #J #H destruct
-| #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #J #H destruct
-| #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #J #H destruct
-]
-qed-.
-
-lemma cpr_inv_atom1: ∀I,G,L,T2. ⦃G, L⦄ ⊢ ⓪{I} ➡ T2 →
- T2 = ⓪{I} ∨
- ∃∃K,V,V2,i. ⬇[i] L ≡ K. ⓓV & ⦃G, K⦄ ⊢ V ➡ V2 &
- ⬆[O, i + 1] V2 ≡ T2 & I = LRef i.
-/2 width=3 by cpr_inv_atom1_aux/ qed-.
-
-(* Basic_1: includes: pr0_gen_sort pr2_gen_sort *)
-lemma cpr_inv_sort1: ∀G,L,T2,s. ⦃G, L⦄ ⊢ ⋆s ➡ T2 → T2 = ⋆s.
-#G #L #T2 #s #H
-elim (cpr_inv_atom1 … H) -H //
-* #K #V #V2 #i #_ #_ #_ #H destruct
-qed-.
-
-(* Basic_1: includes: pr0_gen_lref pr2_gen_lref *)
-lemma cpr_inv_lref1: ∀G,L,T2,i. ⦃G, L⦄ ⊢ #i ➡ T2 →
- T2 = #i ∨
- ∃∃K,V,V2. ⬇[i] L ≡ K. ⓓV & ⦃G, K⦄ ⊢ V ➡ V2 &
- ⬆[O, i + 1] V2 ≡ T2.
-#G #L #T2 #i #H
-elim (cpr_inv_atom1 … H) -H /2 width=1 by or_introl/
-* #K #V #V2 #j #HLK #HV2 #HVT2 #H destruct /3 width=6 by ex3_3_intro, or_intror/
-qed-.
-
-lemma cpr_inv_gref1: ∀G,L,T2,p. ⦃G, L⦄ ⊢ §p ➡ T2 → T2 = §p.
-#G #L #T2 #p #H
-elim (cpr_inv_atom1 … H) -H //
-* #K #V #V2 #i #_ #_ #_ #H destruct
-qed-.
-
-fact cpr_inv_bind1_aux: ∀G,L,U1,U2. ⦃G, L⦄ ⊢ U1 ➡ U2 →
- ∀a,I,V1,T1. U1 = ⓑ{a,I}V1. T1 → (
- ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ➡ T2 &
- U2 = ⓑ{a,I}V2.T2
- ) ∨
- ∃∃T. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡ T & ⬆[0, 1] U2 ≡ T &
- a = true & I = Abbr.
-#G #L #U1 #U2 * -L -U1 -U2
-[ #I #G #L #b #J #W1 #U1 #H destruct
-| #L #G #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
-| #G #L #V #T1 #T #T2 #HT1 #HT2 #b #J #W #U1 #H destruct /3 width=3 by ex4_intro, or_intror/
-| #G #L #V #T1 #T2 #_ #b #J #W #U1 #H destruct
-| #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #b #J #W #U1 #H destruct
-| #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #b #J #W #U1 #H destruct
-]
-qed-.
-
-lemma cpr_inv_bind1: ∀a,I,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡ U2 → (
- ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ➡ T2 &
- U2 = ⓑ{a,I}V2.T2
- ) ∨
- ∃∃T. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡ T & ⬆[0, 1] U2 ≡ T &
- a = true & I = Abbr.
-/2 width=3 by cpr_inv_bind1_aux/ qed-.
-
-(* Basic_1: includes: pr0_gen_abbr pr2_gen_abbr *)
-lemma cpr_inv_abbr1: ∀a,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓓ{a}V1.T1 ➡ U2 → (
- ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L. ⓓV1⦄ ⊢ T1 ➡ T2 &
- U2 = ⓓ{a}V2.T2
- ) ∨
- ∃∃T. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡ T & ⬆[0, 1] U2 ≡ T & a = true.
-#a #G #L #V1 #T1 #U2 #H
-elim (cpr_inv_bind1 … H) -H *
-/3 width=5 by ex3_2_intro, ex3_intro, or_introl, or_intror/
-qed-.
-
-(* Basic_1: includes: pr0_gen_abst pr2_gen_abst *)
-lemma cpr_inv_abst1: ∀a,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓛ{a}V1.T1 ➡ U2 →
- ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L.ⓛV1⦄ ⊢ T1 ➡ T2 &
- U2 = ⓛ{a}V2.T2.
-#a #G #L #V1 #T1 #U2 #H
-elim (cpr_inv_bind1 … H) -H *
-[ /3 width=5 by ex3_2_intro/
-| #T #_ #_ #_ #H destruct
-]
-qed-.
-
-fact cpr_inv_flat1_aux: ∀G,L,U,U2. ⦃G, L⦄ ⊢ U ➡ U2 →
- ∀I,V1,U1. U = ⓕ{I}V1.U1 →
- ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L⦄ ⊢ U1 ➡ T2 &
- U2 = ⓕ{I} V2. T2
- | (⦃G, L⦄ ⊢ U1 ➡ U2 ∧ I = Cast)
- | ∃∃a,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L⦄ ⊢ W1 ➡ W2 &
- ⦃G, L.ⓛW1⦄ ⊢ T1 ➡ T2 & U1 = ⓛ{a}W1.T1 &
- U2 = ⓓ{a}ⓝW2.V2.T2 & I = Appl
- | ∃∃a,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡ V & ⬆[0,1] V ≡ V2 &
- ⦃G, L⦄ ⊢ W1 ➡ W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ➡ T2 &
- U1 = ⓓ{a}W1.T1 &
- U2 = ⓓ{a}W2.ⓐV2.T2 & I = Appl.
-#G #L #U #U2 * -L -U -U2
-[ #I #G #L #J #W1 #U1 #H destruct
-| #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 or4_intro0, ex3_2_intro/
-| #G #L #V #T1 #T #T2 #_ #_ #J #W #U1 #H destruct
-| #G #L #V #T1 #T2 #HT12 #J #W #U1 #H destruct /3 width=1 by or4_intro1, conj/
-| #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #HV12 #HW12 #HT12 #J #W #U1 #H destruct /3 width=11 by or4_intro2, ex6_6_intro/
-| #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV1 #HV2 #HW12 #HT12 #J #W #U1 #H destruct /3 width=13 by or4_intro3, ex7_7_intro/
-]
-qed-.
-
-lemma cpr_inv_flat1: ∀I,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓕ{I}V1.U1 ➡ U2 →
- ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L⦄ ⊢ U1 ➡ T2 &
- U2 = ⓕ{I}V2.T2
- | (⦃G, L⦄ ⊢ U1 ➡ U2 ∧ I = Cast)
- | ∃∃a,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L⦄ ⊢ W1 ➡ W2 &
- ⦃G, L.ⓛW1⦄ ⊢ T1 ➡ T2 & U1 = ⓛ{a}W1.T1 &
- U2 = ⓓ{a}ⓝW2.V2.T2 & I = Appl
- | ∃∃a,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡ V & ⬆[0,1] V ≡ V2 &
- ⦃G, L⦄ ⊢ W1 ➡ W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ➡ T2 &
- U1 = ⓓ{a}W1.T1 &
- U2 = ⓓ{a}W2.ⓐV2.T2 & I = Appl.
-/2 width=3 by cpr_inv_flat1_aux/ qed-.
-
-(* Basic_1: includes: pr0_gen_appl pr2_gen_appl *)
-lemma cpr_inv_appl1: ∀G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓐV1.U1 ➡ U2 →
- ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L⦄ ⊢ U1 ➡ T2 &
- U2 = ⓐV2.T2
- | ∃∃a,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L⦄ ⊢ W1 ➡ W2 &
- ⦃G, L.ⓛW1⦄ ⊢ T1 ➡ T2 &
- U1 = ⓛ{a}W1.T1 & U2 = ⓓ{a}ⓝW2.V2.T2
- | ∃∃a,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡ V & ⬆[0,1] V ≡ V2 &
- ⦃G, L⦄ ⊢ W1 ➡ W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ➡ T2 &
- U1 = ⓓ{a}W1.T1 & U2 = ⓓ{a}W2.ⓐV2.T2.
-#G #L #V1 #U1 #U2 #H elim (cpr_inv_flat1 … H) -H *
-[ /3 width=5 by or3_intro0, ex3_2_intro/
-| #_ #H destruct
-| /3 width=11 by or3_intro1, ex5_6_intro/
-| /3 width=13 by or3_intro2, ex6_7_intro/
-]
-qed-.
-
-(* Note: the main property of simple terms *)
-lemma cpr_inv_appl1_simple: ∀G,L,V1,T1,U. ⦃G, L⦄ ⊢ ⓐV1. T1 ➡ U → 𝐒⦃T1⦄ →
- ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L⦄ ⊢ T1 ➡ T2 &
- U = ⓐV2. T2.
-#G #L #V1 #T1 #U #H #HT1
-elim (cpr_inv_appl1 … H) -H *
-[ /2 width=5 by ex3_2_intro/
-| #a #V2 #W1 #W2 #U1 #U2 #_ #_ #_ #H #_ destruct
- elim (simple_inv_bind … HT1)
-| #a #V #V2 #W1 #W2 #U1 #U2 #_ #_ #_ #_ #H #_ destruct
- elim (simple_inv_bind … HT1)
-]
-qed-.
-
-(* Basic_1: includes: pr0_gen_cast pr2_gen_cast *)
-lemma cpr_inv_cast1: ∀G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓝ V1. U1 ➡ U2 → (
- ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L⦄ ⊢ U1 ➡ T2 &
- U2 = ⓝ V2. T2
- ) ∨ ⦃G, L⦄ ⊢ U1 ➡ U2.
-#G #L #V1 #U1 #U2 #H elim (cpr_inv_flat1 … H) -H *
-[ /3 width=5 by ex3_2_intro, or_introl/
-| /2 width=1 by or_intror/
-| #a #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #_ #H destruct
-| #a #V #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #_ #_ #H destruct
-]
-qed-.
-
-(* Basic forward lemmas *****************************************************)
-
-lemma cpr_fwd_bind1_minus: ∀I,G,L,V1,T1,T. ⦃G, L⦄ ⊢ -ⓑ{I}V1.T1 ➡ T → ∀b.
- ∃∃V2,T2. ⦃G, L⦄ ⊢ ⓑ{b,I}V1.T1 ➡ ⓑ{b,I}V2.T2 &
- T = -ⓑ{I}V2.T2.
-#I #G #L #V1 #T1 #T #H #b
-elim (cpr_inv_bind1 … H) -H *
-[ #V2 #T2 #HV12 #HT12 #H destruct /3 width=4 by cpr_bind, ex2_2_intro/
-| #T2 #_ #_ #H destruct
-]
-qed-.
-
-(* Basic_1: removed theorems 11:
- pr0_subst0_back pr0_subst0_fwd pr0_subst0
- pr2_head_2 pr2_cflat clear_pr2_trans
- pr2_gen_csort pr2_gen_cflat pr2_gen_cbind
- pr2_gen_ctail pr2_ctail
-*)
-(* Basic_1: removed local theorems 4:
- pr0_delta_eps pr0_cong_delta
- pr2_free_free pr2_free_delta
-*)
+++ /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/cir.ma".
-include "basic_2/reduction/cpr.ma".
-
-(* CONTEXT-SENSITIVE PARALLEL REDUCTION FOR TERMS ***************************)
-
-(* Advanced forward lemmas on irreducibility ********************************)
-
-lemma cpr_fwd_cir: ∀G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ⦃G, L⦄ ⊢ ➡ 𝐈⦃T1⦄ → T2 = T1.
-#G #L #T1 #T2 #H elim H -G -L -T1 -T2
-[ //
-| #G #L #K #V1 #V2 #W2 #i #HLK #_ #HVW2 #IHV12 #H
- elim (cir_inv_delta … HLK) //
-| #a * #G #L #V1 #V2 #T1 #T2 #_ #_ #IHV1 #IHT1 #H
- [ elim (cir_inv_bind … H) -H #HV1 #HT1 * #H destruct
- lapply (IHV1 … HV1) -IHV1 -HV1 #H destruct
- lapply (IHT1 … HT1) -IHT1 #H destruct //
- | elim (cir_inv_ib2 … H) -H /3 width=2 by or_introl, eq_f2/
- ]
-| * #G #L #V1 #V2 #T1 #T2 #_ #_ #IHV1 #IHT1 #H
- [ elim (cir_inv_appl … H) -H #HV1 #HT1 #_
- >IHV1 -IHV1 // -HV1 >IHT1 -IHT1 //
- | elim (cir_inv_ri2 … H) /2 width=1 by/
- ]
-| #G #L #V1 #T1 #T #T2 #_ #_ #_ #H
- elim (cir_inv_ri2 … H) /2 width=1 by or_introl/
-| #G #L #V1 #T1 #T2 #_ #_ #H
- elim (cir_inv_ri2 … H) /2 width=1 by/
-| #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #_ #_ #H
- elim (cir_inv_appl … H) -H #_ #_ #H
- elim (simple_inv_bind … H)
-| #a #G #L #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #_ #_ #_ #H
- elim (cir_inv_appl … H) -H #_ #_ #H
- elim (simple_inv_bind … H)
-]
-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_max.ma".
-include "basic_2/substitution/drop_drop.ma".
-include "basic_2/reduction/cpr.ma".
-
-(* CONTEXT-SENSITIVE PARALLEL REDUCTION FOR TERMS ***************************)
-
-(* Relocation properties ****************************************************)
-
-(* Basic_1: includes: pr0_lift pr2_lift *)
-lemma cpr_lift: ∀G. d_liftable (cpr G).
-#G #K #T1 #T2 #H elim H -G -K -T1 -T2
-[ #I #G #K #L #c #l #k #_ #U1 #H1 #U2 #H2
- >(lift_mono … H1 … H2) -H1 -H2 //
-| #G #K #KV #V #V2 #W2 #i #HKV #HV2 #HVW2 #IHV2 #L #c #l #k #HLK #U1 #H #U2 #HWU2
- elim (lift_inv_lref1 … H) * #Hil #H destruct
- [ elim (lift_trans_ge … HVW2 … HWU2) -W2 /2 width=1 by ylt_fwd_le_succ1/ #W2 #HVW2 #HWU2
- elim (drop_trans_le … HLK … HKV) -K /2 width=2 by ylt_fwd_le/ #X #HLK #H
- elim (drop_inv_skip2 … H) -H /2 width=1 by ylt_to_minus/ -Hil
- #K #Y #HKV #HVY #H destruct /3 width=9 by cpr_delta/
- | lapply (lift_trans_be … HVW2 … HWU2 ? ?) -W2 /2 width=1 by yle_succ_dx/ >plus_plus_comm_23 #HVU2
- 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 #c #l #k #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 #c #l #k #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 #c #l #k #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 #c #l #k #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 #c #l #k #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 #c #l #k #HLK #X1 #HX1 #X2 #HX2
- elim (lift_inv_flat1 … HX1) -HX1 #V0 #X #HV10 #HX #HX1 destruct
- elim (lift_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT10 #HX destruct
- elim (lift_inv_bind1 … HX2) -HX2 #W3 #X #HW23 #HX #HX2 destruct
- elim (lift_inv_flat1 … HX) -HX #V3 #T3 #HV3 #HT23 #HX destruct
- elim (lift_trans_ge … HV2 … HV3) -V2 /4 width=6 by cpr_theta, drop_skip/
-]
-qed.
-
-(* Basic_1: includes: pr0_gen_lift pr2_gen_lift *)
-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 #c #l #k #_ #T1 #H
- [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3 by 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 #c #l #k #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 // <yminus_succ2 <yplus_inj >yplus_SO2 >ymax_pre_sn /3 width=8 by cpr_delta, ylt_fwd_le_succ1, ex2_intro/
- | elim (yle_inv_plus_inj2 … Hil) #Hlim #Hmi
- lapply (yle_inv_inj … Hmi) -Hmi #Hmi
- lapply (drop_conf_ge … HLK … HLV ?) -L // #HKLV
- elim (lift_split … HVW2 l (i - k + 1)) -HVW2 [2,3,4: /2 width=1 by yle_succ_dx, 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=8 by cpr_delta, ex2_intro/
- ]
-| #a #I #G #L #V1 #V2 #U1 #U2 #_ #_ #IHV12 #IHU12 #K #c #l #k #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 #c #l #k #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 #c #l #k #HLK #X #H
- elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct
- elim (IHU1 (K.ⓓW1) c … 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 #c #l #k #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 #c #l #k #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) c … 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 #c #l #k #HLK #X #HX
- elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct
- elim (lift_inv_bind2 … HY) -HY #W0 #T0 #HW01 #HT01 #HY destruct
- elim (IHV1 … HLK … HV01) -V1 #V3 #HV3 #HV03
- elim (IHT12 (K.ⓓW0) c … HT01) -T1 /2 width=1 by drop_skip/ #T3 #HT32 #HT03
- elim (IHW12 … HLK … HW01) -W1
- elim (lift_trans_le … HV3 … HV2) -V
- /4 width=9 by cpr_theta, lift_flat, lift_bind, ex2_intro/
-]
-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_drop.ma".
-include "basic_2/reduction/cpr.ma".
-
-(* CONTEXT-SENSITIVE PARALLEL REDUCTION FOR TERMS ***************************)
-
-(* Properties on lazy sn pointwise extensions *******************************)
-
-lemma cpr_llpx_sn_conf: ∀R. (∀L. reflexive … (R L)) → d_liftable R → d_deliftable_sn R →
- ∀G. c_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 #V1c #V2c #W2c #i #HLKs #_ #HVW2c #IHV12c #Ld #H elim (llpx_sn_inv_lref_ge_sn … H … HLKs) // -H
- #Kd #V1l #HLKd #HV1c #HV1sd
- lapply (drop_fwd_drop2 … HLKs) -HLKs #HLKs
- lapply (drop_fwd_drop2 … HLKd) -HLKd #HLKd
- @(llpx_sn_lift_le … HLKs HLKd … HVW2c) -HLKs -HLKd -HVW2c /2 width=1 by/ (**) (* full auto too slow *)
-| #a #I #G #Ls #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #Ld #H elim (llpx_sn_inv_bind_O … H) -H
- /4 width=5 by llpx_sn_bind_repl_SO, llpx_sn_bind/
-| #I #G #Ls #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
- /3 width=1 by llpx_sn_flat/
-| #G #Ls #V #T1 #T2 #T #_ #HT2 #IHT12 #Ld #H elim (llpx_sn_inv_bind_O … H) -H
- /3 width=10 by llpx_sn_inv_lift_le, drop_drop/
-| #G #Ls #V #T1 #T2 #_ #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H /2 width=1 by/
-| #a #G #Ls #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
- #HV1 #H elim (llpx_sn_inv_bind_O … H) -H
- /4 width=5 by llpx_sn_bind_repl_SO, llpx_sn_flat, llpx_sn_bind/
-| #a #G #Ls #V1 #V2 #V #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
- #HV1 #H elim (llpx_sn_inv_bind_O … H) -H //
- #HW1 #HT1 @llpx_sn_bind_O /2 width=1 by/ @llpx_sn_flat (**) (* full auto too slow *)
- [ /3 width=10 by llpx_sn_lift_le, drop_drop/
- | /3 width=4 by llpx_sn_bind_repl_O/
-]
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/pred_6.ma".
-include "basic_2/static/sd.ma".
-include "basic_2/reduction/cpr.ma".
-
-(* CONTEXT-SENSITIVE EXTENDED PARALLEL REDUCTION FOR TERMS ******************)
-
-(* avtivate genv *)
-inductive cpx (h) (o): relation4 genv lenv term term ≝
-| cpx_atom : ∀I,G,L. cpx h o G L (⓪{I}) (⓪{I})
-| cpx_st : ∀G,L,s,d. deg h o s (d+1) → cpx h o G L (⋆s) (⋆(next h s))
-| cpx_delta: ∀I,G,L,K,V,V2,W2,i.
- ⬇[i] L ≡ K.ⓑ{I}V → cpx h o G K V V2 →
- ⬆[0, i+1] V2 ≡ W2 → cpx h o G L (#i) W2
-| cpx_bind : ∀a,I,G,L,V1,V2,T1,T2.
- cpx h o G L V1 V2 → cpx h o G (L.ⓑ{I}V1) T1 T2 →
- cpx h o G L (ⓑ{a,I}V1.T1) (ⓑ{a,I}V2.T2)
-| cpx_flat : ∀I,G,L,V1,V2,T1,T2.
- cpx h o G L V1 V2 → cpx h o G L T1 T2 →
- cpx h o G L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2)
-| cpx_zeta : ∀G,L,V,T1,T,T2. cpx h o G (L.ⓓV) T1 T →
- ⬆[0, 1] T2 ≡ T → cpx h o G L (+ⓓV.T1) T2
-| cpx_eps : ∀G,L,V,T1,T2. cpx h o G L T1 T2 → cpx h o G L (ⓝV.T1) T2
-| cpx_ct : ∀G,L,V1,V2,T. cpx h o G L V1 V2 → cpx h o G L (ⓝV1.T) V2
-| cpx_beta : ∀a,G,L,V1,V2,W1,W2,T1,T2.
- cpx h o G L V1 V2 → cpx h o G L W1 W2 → cpx h o G (L.ⓛW1) T1 T2 →
- cpx h o G L (ⓐV1.ⓛ{a}W1.T1) (ⓓ{a}ⓝW2.V2.T2)
-| cpx_theta: ∀a,G,L,V1,V,V2,W1,W2,T1,T2.
- cpx h o G L V1 V → ⬆[0, 1] V ≡ V2 → cpx h o G L W1 W2 →
- cpx h o G (L.ⓓW1) T1 T2 →
- cpx h o G L (ⓐV1.ⓓ{a}W1.T1) (ⓓ{a}W2.ⓐV2.T2)
-.
-
-interpretation
- "context-sensitive extended parallel reduction (term)"
- 'PRed h o G L T1 T2 = (cpx h o G L T1 T2).
-
-(* Basic properties *********************************************************)
-
-lemma lsubr_cpx_trans: ∀h,o,G. lsub_trans … (cpx h o G) lsubr.
-#h #o #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 (lsubr_fwd_drop2_pair … HL12 … HLK1) -HL12 -HLK1 *
- /4 width=7 by cpx_delta, cpx_ct/
-|4,9: /4 width=1 by cpx_bind, cpx_beta, lsubr_pair/
-|5,7,8: /3 width=1 by cpx_flat, cpx_eps, cpx_ct/
-|6,10: /4 width=3 by cpx_zeta, cpx_theta, lsubr_pair/
-]
-qed-.
-
-(* Note: this is "∀h,g,L. reflexive … (cpx h g L)" *)
-lemma cpx_refl: ∀h,o,G,T,L. ⦃G, L⦄ ⊢ T ➡[h, o] T.
-#h #o #G #T elim T -T // * /2 width=1 by cpx_bind, cpx_flat/
-qed.
-
-lemma cpr_cpx: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ⦃G, L⦄ ⊢ T1 ➡[h, o] T2.
-#h #o #G #L #T1 #T2 #H elim H -L -T1 -T2
-/2 width=7 by cpx_delta, cpx_bind, cpx_flat, cpx_zeta, cpx_eps, cpx_beta, cpx_theta/
-qed.
-
-lemma cpx_pair_sn: ∀h,o,I,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 →
- ∀T. ⦃G, L⦄ ⊢ ②{I}V1.T ➡[h, o] ②{I}V2.T.
-#h #o * /2 width=1 by cpx_bind, cpx_flat/
-qed.
-
-lemma cpx_delift: ∀h,o,I,G,K,V,T1,L,l. ⬇[l] L ≡ (K.ⓑ{I}V) →
- ∃∃T2,T. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 & ⬆[l, 1] T ≡ T2.
-#h #o #I #G #K #V #T1 elim T1 -T1
-[ * #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: /4 width=4 by cpx_atom, lift_lref_ge_minus, lift_lref_lt, ylt_inj, yle_inj, 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 #l #HLK
- elim (IHW1 … HLK) -IHW1 #W2 #W #HW12 #HW2
- [ 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/
- ]
-]
-qed-.
-
-(* Basic inversion lemmas ***************************************************)
-
-fact cpx_inv_atom1_aux: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → ∀J. T1 = ⓪{J} →
- ∨∨ T2 = ⓪{J}
- | ∃∃s,d. deg h o s (d+1) & T2 = ⋆(next h s) & J = Sort s
- | ∃∃I,K,V,V2,i. ⬇[i] L ≡ K.ⓑ{I}V & ⦃G, K⦄ ⊢ V ➡[h, o] V2 &
- ⬆[O, i+1] V2 ≡ T2 & J = LRef i.
-#G #h #o #L #T1 #T2 * -L -T1 -T2
-[ #I #G #L #J #H destruct /2 width=1 by or3_intro0/
-| #G #L #s #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
-| #G #L #V #T1 #T #T2 #_ #_ #J #H destruct
-| #G #L #V #T1 #T2 #_ #J #H destruct
-| #G #L #V1 #V2 #T #_ #J #H destruct
-| #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #J #H destruct
-| #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #J #H destruct
-]
-qed-.
-
-lemma cpx_inv_atom1: ∀h,o,J,G,L,T2. ⦃G, L⦄ ⊢ ⓪{J} ➡[h, o] T2 →
- ∨∨ T2 = ⓪{J}
- | ∃∃s,d. deg h o s (d+1) & T2 = ⋆(next h s) & J = Sort s
- | ∃∃I,K,V,V2,i. ⬇[i] L ≡ K.ⓑ{I}V & ⦃G, K⦄ ⊢ V ➡[h, o] V2 &
- ⬆[O, i+1] V2 ≡ T2 & J = LRef i.
-/2 width=3 by cpx_inv_atom1_aux/ qed-.
-
-lemma cpx_inv_sort1: ∀h,o,G,L,T2,s. ⦃G, L⦄ ⊢ ⋆s ➡[h, o] T2 → T2 = ⋆s ∨
- ∃∃d. deg h o s (d+1) & T2 = ⋆(next h s).
-#h #o #G #L #T2 #s #H
-elim (cpx_inv_atom1 … H) -H /2 width=1 by or_introl/ *
-[ #s0 #d0 #Hkd0 #H1 #H2 destruct /3 width=4 by ex2_intro, or_intror/
-| #I #K #V #V2 #i #_ #_ #_ #H destruct
-]
-qed-.
-
-lemma cpx_inv_lref1: ∀h,o,G,L,T2,i. ⦃G, L⦄ ⊢ #i ➡[h, o] T2 →
- T2 = #i ∨
- ∃∃I,K,V,V2. ⬇[i] L ≡ K. ⓑ{I}V & ⦃G, K⦄ ⊢ V ➡[h, o] V2 &
- ⬆[O, i+1] V2 ≡ T2.
-#h #o #G #L #T2 #i #H
-elim (cpx_inv_atom1 … H) -H /2 width=1 by or_introl/ *
-[ #s #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_lref1_ge: ∀h,o,G,L,T2,i. ⦃G, L⦄ ⊢ #i ➡[h, o] T2 → |L| ≤ i → T2 = #i.
-#h #o #G #L #T2 #i #H elim (cpx_inv_lref1 … H) -H // *
-#I #K #V1 #V2 #HLK #_ #_ #HL -h -G -V2 lapply (drop_fwd_length_lt2 … HLK) -K -I -V1
-#H elim (lt_refl_false i) /2 width=3 by lt_to_le_to_lt/
-qed-.
-
-lemma cpx_inv_gref1: ∀h,o,G,L,T2,p. ⦃G, L⦄ ⊢ §p ➡[h, o] T2 → T2 = §p.
-#h #o #G #L #T2 #p #H
-elim (cpx_inv_atom1 … H) -H // *
-[ #s #d #_ #_ #H destruct
-| #I #K #V #V2 #i #_ #_ #_ #H destruct
-]
-qed-.
-
-fact cpx_inv_bind1_aux: ∀h,o,G,L,U1,U2. ⦃G, L⦄ ⊢ U1 ➡[h, o] U2 →
- ∀a,J,V1,T1. U1 = ⓑ{a,J}V1.T1 → (
- ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L.ⓑ{J}V1⦄ ⊢ T1 ➡[h, o] T2 &
- U2 = ⓑ{a,J}V2.T2
- ) ∨
- ∃∃T. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡[h, o] T & ⬆[0, 1] U2 ≡ T &
- a = true & J = Abbr.
-#h #o #G #L #U1 #U2 * -L -U1 -U2
-[ #I #G #L #b #J #W #U1 #H destruct
-| #G #L #s #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
-| #G #L #V #T1 #T #T2 #HT1 #HT2 #b #J #W #U1 #H destruct /3 width=3 by ex4_intro, or_intror/
-| #G #L #V #T1 #T2 #_ #b #J #W #U1 #H destruct
-| #G #L #V1 #V2 #T #_ #b #J #W #U1 #H destruct
-| #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #b #J #W #U1 #H destruct
-| #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #b #J #W #U1 #H destruct
-]
-qed-.
-
-lemma cpx_inv_bind1: ∀h,o,a,I,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡[h, o] U2 → (
- ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ➡[h, o] T2 &
- U2 = ⓑ{a,I} V2. T2
- ) ∨
- ∃∃T. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡[h, o] T & ⬆[0, 1] U2 ≡ T &
- a = true & I = Abbr.
-/2 width=3 by cpx_inv_bind1_aux/ qed-.
-
-lemma cpx_inv_abbr1: ∀h,o,a,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓓ{a}V1.T1 ➡[h, o] U2 → (
- ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L.ⓓV1⦄ ⊢ T1 ➡[h, o] T2 &
- U2 = ⓓ{a} V2. T2
- ) ∨
- ∃∃T. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡[h, o] T & ⬆[0, 1] U2 ≡ T & a = true.
-#h #o #a #G #L #V1 #T1 #U2 #H
-elim (cpx_inv_bind1 … H) -H * /3 width=5 by ex3_2_intro, ex3_intro, or_introl, or_intror/
-qed-.
-
-lemma cpx_inv_abst1: ∀h,o,a,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓛ{a}V1.T1 ➡[h, o] U2 →
- ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L.ⓛV1⦄ ⊢ T1 ➡[h, o] T2 &
- U2 = ⓛ{a} V2. T2.
-#h #o #a #G #L #V1 #T1 #U2 #H
-elim (cpx_inv_bind1 … H) -H *
-[ /3 width=5 by ex3_2_intro/
-| #T #_ #_ #_ #H destruct
-]
-qed-.
-
-fact cpx_inv_flat1_aux: ∀h,o,G,L,U,U2. ⦃G, L⦄ ⊢ U ➡[h, o] U2 →
- ∀J,V1,U1. U = ⓕ{J}V1.U1 →
- ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L⦄ ⊢ U1 ➡[h, o] T2 &
- U2 = ⓕ{J}V2.T2
- | (⦃G, L⦄ ⊢ U1 ➡[h, o] U2 ∧ J = Cast)
- | (⦃G, L⦄ ⊢ V1 ➡[h, o] U2 ∧ J = Cast)
- | ∃∃a,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L⦄ ⊢ W1 ➡[h, o] W2 &
- ⦃G, L.ⓛW1⦄ ⊢ T1 ➡[h, o] T2 &
- U1 = ⓛ{a}W1.T1 &
- U2 = ⓓ{a}ⓝW2.V2.T2 & J = Appl
- | ∃∃a,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V & ⬆[0,1] V ≡ V2 &
- ⦃G, L⦄ ⊢ W1 ➡[h, o] W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ➡[h, o] T2 &
- U1 = ⓓ{a}W1.T1 &
- U2 = ⓓ{a}W2.ⓐV2.T2 & J = Appl.
-#h #o #G #L #U #U2 * -L -U -U2
-[ #I #G #L #J #W #U1 #H destruct
-| #G #L #s #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/
-| #G #L #V #T1 #T #T2 #_ #_ #J #W #U1 #H destruct
-| #G #L #V #T1 #T2 #HT12 #J #W #U1 #H destruct /3 width=1 by or5_intro1, conj/
-| #G #L #V1 #V2 #T #HV12 #J #W #U1 #H destruct /3 width=1 by or5_intro2, conj/
-| #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #HV12 #HW12 #HT12 #J #W #U1 #H destruct /3 width=11 by or5_intro3, ex6_6_intro/
-| #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV1 #HV2 #HW12 #HT12 #J #W #U1 #H destruct /3 width=13 by or5_intro4, ex7_7_intro/
-]
-qed-.
-
-lemma cpx_inv_flat1: ∀h,o,I,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓕ{I}V1.U1 ➡[h, o] U2 →
- ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L⦄ ⊢ U1 ➡[h, o] T2 &
- U2 = ⓕ{I} V2. T2
- | (⦃G, L⦄ ⊢ U1 ➡[h, o] U2 ∧ I = Cast)
- | (⦃G, L⦄ ⊢ V1 ➡[h, o] U2 ∧ I = Cast)
- | ∃∃a,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L⦄ ⊢ W1 ➡[h, o] W2 &
- ⦃G, L.ⓛW1⦄ ⊢ T1 ➡[h, o] T2 &
- U1 = ⓛ{a}W1.T1 &
- U2 = ⓓ{a}ⓝW2.V2.T2 & I = Appl
- | ∃∃a,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V & ⬆[0,1] V ≡ V2 &
- ⦃G, L⦄ ⊢ W1 ➡[h, o] W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ➡[h, o] T2 &
- U1 = ⓓ{a}W1.T1 &
- U2 = ⓓ{a}W2.ⓐV2.T2 & I = Appl.
-/2 width=3 by cpx_inv_flat1_aux/ qed-.
-
-lemma cpx_inv_appl1: ∀h,o,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓐ V1.U1 ➡[h, o] U2 →
- ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L⦄ ⊢ U1 ➡[h, o] T2 &
- U2 = ⓐ V2. T2
- | ∃∃a,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L⦄ ⊢ W1 ➡[h, o] W2 &
- ⦃G, L.ⓛW1⦄ ⊢ T1 ➡[h, o] T2 &
- U1 = ⓛ{a}W1.T1 & U2 = ⓓ{a}ⓝW2.V2.T2
- | ∃∃a,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V & ⬆[0,1] V ≡ V2 &
- ⦃G, L⦄ ⊢ W1 ➡[h, o] W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ➡[h, o] T2 &
- U1 = ⓓ{a}W1.T1 & U2 = ⓓ{a}W2. ⓐV2. T2.
-#h #o #G #L #V1 #U1 #U2 #H elim (cpx_inv_flat1 … H) -H *
-[ /3 width=5 by or3_intro0, ex3_2_intro/
-|2,3: #_ #H destruct
-| /3 width=11 by or3_intro1, ex5_6_intro/
-| /3 width=13 by or3_intro2, ex6_7_intro/
-]
-qed-.
-
-(* Note: the main property of simple terms *)
-lemma cpx_inv_appl1_simple: ∀h,o,G,L,V1,T1,U. ⦃G, L⦄ ⊢ ⓐV1.T1 ➡[h, o] U → 𝐒⦃T1⦄ →
- ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 &
- U = ⓐV2.T2.
-#h #o #G #L #V1 #T1 #U #H #HT1
-elim (cpx_inv_appl1 … H) -H *
-[ /2 width=5 by ex3_2_intro/
-| #a #V2 #W1 #W2 #U1 #U2 #_ #_ #_ #H #_ destruct
- elim (simple_inv_bind … HT1)
-| #a #V #V2 #W1 #W2 #U1 #U2 #_ #_ #_ #_ #H #_ destruct
- elim (simple_inv_bind … HT1)
-]
-qed-.
-
-lemma cpx_inv_cast1: ∀h,o,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓝV1.U1 ➡[h, o] U2 →
- ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L⦄ ⊢ U1 ➡[h, o] T2 &
- U2 = ⓝ V2. T2
- | ⦃G, L⦄ ⊢ U1 ➡[h, o] U2
- | ⦃G, L⦄ ⊢ V1 ➡[h, o] U2.
-#h #o #G #L #V1 #U1 #U2 #H elim (cpx_inv_flat1 … H) -H *
-[ /3 width=5 by or3_intro0, ex3_2_intro/
-|2,3: /2 width=1 by or3_intro1, or3_intro2/
-| #a #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #_ #H destruct
-| #a #V #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #_ #_ #H destruct
-]
-qed-.
-
-(* Basic forward lemmas *****************************************************)
-
-lemma cpx_fwd_bind1_minus: ∀h,o,I,G,L,V1,T1,T. ⦃G, L⦄ ⊢ -ⓑ{I}V1.T1 ➡[h, o] T → ∀b.
- ∃∃V2,T2. ⦃G, L⦄ ⊢ ⓑ{b,I}V1.T1 ➡[h, o] ⓑ{b,I}V2.T2 &
- T = -ⓑ{I}V2.T2.
-#h #o #I #G #L #V1 #T1 #T #H #b
-elim (cpx_inv_bind1 … H) -H *
-[ #V2 #T2 #HV12 #HT12 #H destruct /3 width=4 by cpx_bind, ex2_2_intro/
-| #T2 #_ #_ #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/reduction/cix.ma".
-include "basic_2/reduction/cpx.ma".
-
-(* CONTEXT-SENSITIVE EXTENDED PARALLEL REDUCTION FOR TERMS ******************)
-
-(* Advanced forward lemmas on irreducibility ********************************)
-
-lemma cpx_fwd_cix: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃T1⦄ → T2 = T1.
-#h #o #G #L #T1 #T2 #H elim H -G -L -T1 -T2
-[ //
-| #G #L #s #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
- [ elim (cix_inv_bind … H) -H #HV1 #HT1 * #H destruct
- lapply (IHV1 … HV1) -IHV1 -HV1 #H destruct
- lapply (IHT1 … HT1) -IHT1 #H destruct //
- | elim (cix_inv_ib2 … H) -H /3 width=2 by or_introl, eq_f2/
- ]
-| * #G #L #V1 #V2 #T1 #T2 #_ #_ #IHV1 #IHT1 #H
- [ elim (cix_inv_appl … H) -H #HV1 #HT1 #_
- >IHV1 -IHV1 // -HV1 >IHT1 -IHT1 //
- | elim (cix_inv_ri2 … H) /2 width=1 by/
- ]
-| #G #L #V1 #T1 #T #T2 #_ #_ #_ #H
- elim (cix_inv_ri2 … H) /2 width=1 by or_introl/
-| #G #L #V1 #T1 #T2 #_ #_ #H
- elim (cix_inv_ri2 … H) /2 width=1 by/
-| #G #L #V1 #V2 #T #_ #_ #H
- elim (cix_inv_ri2 … H) /2 width=1 by/
-| #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #_ #_ #H
- elim (cix_inv_appl … H) -H #_ #_ #H
- elim (simple_inv_bind … H)
-| #a #G #L #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #_ #_ #_ #H
- elim (cix_inv_appl … H) -H #_ #_ #H
- elim (simple_inv_bind … H)
-]
-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_max.ma".
-include "basic_2/substitution/drop_drop.ma".
-include "basic_2/multiple/fqus_alt.ma".
-include "basic_2/static/da.ma".
-include "basic_2/reduction/cpx.ma".
-
-(* CONTEXT-SENSITIVE EXTENDED PARALLEL REDUCTION FOR TERMS ******************)
-
-(* Advanced properties ******************************************************)
-
-fact sta_cpx_aux: ∀h,o,G,L,T1,T2,d2,d1. ⦃G, L⦄ ⊢ T1 •*[h, d2] T2 → d2 = 1 →
- ⦃G, L⦄ ⊢ T1 ▪[h, o] d1+1 → ⦃G, L⦄ ⊢ T1 ➡[h, o] T2.
-#h #o #G #L #T1 #T2 #d2 #d1 #H elim H -G -L -T1 -T2 -d2
-[ #G #L #d2 #s #H0 destruct normalize
- /3 width=4 by cpx_st, da_inv_sort/
-| #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 #d2 #HLK #HV12 #HVW2 #_ #H0 #H
- lapply (discr_plus_xy_y … H0) -H0 #H0 destruct
- 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/
-| /4 width=3 by cpx_flat, da_inv_flat/
-| /4 width=3 by cpx_eps, da_inv_flat/
-]
-qed-.
-
-lemma sta_cpx: ∀h,o,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 •*[h, 1] T2 →
- ⦃G, L⦄ ⊢ T1 ▪[h, o] d+1 → ⦃G, L⦄ ⊢ T1 ➡[h, o] T2.
-/2 width=3 by sta_cpx_aux/ qed.
-
-(* Relocation properties ****************************************************)
-
-lemma cpx_lift: ∀h,o,G. d_liftable (cpx h o G).
-#h #o #G #K #T1 #T2 #H elim H -G -K -T1 -T2
-[ #I #G #K #L #c #l #k #_ #U1 #H1 #U2 #H2
- >(lift_mono … H1 … H2) -H1 -H2 //
-| #G #K #s #d #Hkd #L #c #l #k #_ #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 #c #l #k #HLK #U1 #H #U2 #HWU2
- elim (lift_inv_lref1 … H) * #Hil #H destruct
- [ elim (lift_trans_ge … HVW2 … HWU2) -W2 /2 width=1 by ylt_fwd_le_succ1/ #W2 #HVW2 #HWU2
- elim (drop_trans_le … HLK … HKV) -K /2 width=2 by ylt_fwd_le/ #X #HLK #H
- elim (drop_inv_skip2 … H) -H /2 width=1 by ylt_to_minus/ -Hil
- #K #Y #HKV #HVY #H destruct /3 width=10 by cpx_delta/
- | lapply (lift_trans_be … HVW2 … HWU2 ? ?) -W2 /2 width=1 by yle_succ_dx/ >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 #c #l #k #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 #c #l #k #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 #c #l #k #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 #c #l #k #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 #c #l #k #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 #c #l #k #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 #c #l #k #HLK #X1 #HX1 #X2 #HX2
- elim (lift_inv_flat1 … HX1) -HX1 #V0 #X #HV10 #HX #HX1 destruct
- elim (lift_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT10 #HX destruct
- elim (lift_inv_bind1 … HX2) -HX2 #W3 #X #HW23 #HX #HX2 destruct
- elim (lift_inv_flat1 … HX) -HX #V3 #T3 #HV3 #HT23 #HX destruct
- elim (lift_trans_ge … HV2 … HV3) -V2 /4 width=6 by cpx_theta, drop_skip/
-]
-qed.
-
-lemma cpx_inv_lift1: ∀h,o,G. d_deliftable_sn (cpx h o G).
-#h #o #G #L #U1 #U2 #H elim H -G -L -U1 -U2
-[ * #i #G #L #K #c #l #k #_ #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 * #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 #s #d #Hkd #K #c #l #k #_ #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 #c #l #k #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 // <yminus_succ2 <yplus_inj >yplus_SO2 >ymax_pre_sn /3 width=9 by cpx_delta, ylt_fwd_le_succ1, ex2_intro/
- | elim (yle_inv_plus_inj2 … Hil) #Hlim #Hmi
- lapply (yle_inv_inj … Hmi) -Hmi #Hmi
- lapply (drop_conf_ge … HLK … HLV ?) -L // #HKLV
- elim (lift_split … HVW2 l (i - k + 1)) -HVW2 /3 width=1 by yle_succ, yle_pred_sn, 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 #c #l #k #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 #c #l #k #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 #c #l #k #HLK #X #H
- elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct
- elim (IHU1 (K.ⓓW1) c … HTU1) /2 width=1 by drop_skip/ -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 #c #l #k #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 #c #l #k #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 #c #l #k #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) c … 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 #c #l #k #HLK #X #HX
- elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct
- elim (lift_inv_bind2 … HY) -HY #W0 #T0 #HW01 #HT01 #HY destruct
- elim (IHV1 … HLK … HV01) -V1 #V3 #HV3 #HV03
- elim (IHT12 (K.ⓓW0) c … HT01) -T1 /2 width=1 by drop_skip/ #T3 #HT32 #HT03
- elim (IHW12 … HLK … HW01) -W1 #W3 #HW32 #HW03
- elim (lift_trans_le … HV3 … HV2) -V
- /4 width=9 by cpx_theta, lift_bind, lift_flat, ex2_intro/
-]
-qed-.
-
-(* Properties on supclosure *************************************************)
-
-lemma fqu_cpx_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
- ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡[h, o] U2 →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] U1 & ⦃G1, L1, U1⦄ ⊐ ⦃G2, L2, U2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
-/3 width=3 by fqu_pair_sn, fqu_bind_dx, fqu_flat_dx, cpx_pair_sn, cpx_bind, cpx_flat, ex2_intro/
-[ #I #G #L #V2 #U2 #HVU2
- elim (lift_total U2 0 1)
- /4 width=7 by fqu_drop, cpx_delta, drop_pair, drop_drop, ex2_intro/
-| #G #L #K #T1 #U1 #k #HLK1 #HTU1 #T2 #HTU2
- elim (lift_total T2 0 (k+1))
- /3 width=11 by cpx_lift, fqu_drop, ex2_intro/
-]
-qed-.
-
-lemma fqu_sta_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
- ∀U2. ⦃G2, L2⦄ ⊢ T2 •*[h, 1] U2 →
- ∀d. ⦃G2, L2⦄ ⊢ T2 ▪[h, o] d+1 →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] U1 & ⦃G1, L1, U1⦄ ⊐ ⦃G2, L2, U2⦄.
-/3 width=5 by fqu_cpx_trans, sta_cpx/ qed-.
-
-lemma fquq_cpx_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡[h, o] U2 →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] U1 & ⦃G1, L1, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #U2 #HTU2 elim (fquq_inv_gen … H) -H
-[ #HT12 elim (fqu_cpx_trans … HT12 … HTU2) /3 width=3 by fqu_fquq, ex2_intro/
-| * #H1 #H2 #H3 destruct /2 width=3 by ex2_intro/
-]
-qed-.
-
-lemma fquq_sta_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ∀U2. ⦃G2, L2⦄ ⊢ T2 •*[h, 1] U2 →
- ∀d. ⦃G2, L2⦄ ⊢ T2 ▪[h, o] d+1 →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] U1 & ⦃G1, L1, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
-/3 width=5 by fquq_cpx_trans, sta_cpx/ qed-.
-
-lemma fqup_cpx_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
- ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡[h, o] U2 →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] U1 & ⦃G1, L1, U1⦄ ⊐+ ⦃G2, L2, U2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
-[ #G2 #L2 #T2 #H12 #U2 #HTU2 elim (fqu_cpx_trans … H12 … HTU2) -T2
- /3 width=3 by fqu_fqup, ex2_intro/
-| #G #G2 #L #L2 #T #T2 #_ #HT2 #IHT1 #U2 #HTU2
- elim (fqu_cpx_trans … HT2 … HTU2) -T2 #T2 #HT2 #HTU2
- elim (IHT1 … HT2) -T /3 width=7 by fqup_strap1, ex2_intro/
-]
-qed-.
-
-lemma fqus_cpx_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
- ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡[h, o] U2 →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] U1 & ⦃G1, L1, U1⦄ ⊐* ⦃G2, L2, U2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #U2 #HTU2 elim (fqus_inv_gen … H) -H
-[ #HT12 elim (fqup_cpx_trans … HT12 … HTU2) /3 width=3 by fqup_fqus, ex2_intro/
-| * #H1 #H2 #H3 destruct /2 width=3 by ex2_intro/
-]
-qed-.
-
-lemma fqu_cpx_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
- ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡[h, o] U2 → (T2 = U2 → ⊥) →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐ ⦃G2, L2, U2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
-[ #I #G #L #V1 #V2 #HV12 #_ elim (lift_total V2 0 1)
- #U2 #HVU2 @(ex3_intro … U2)
- [1,3: /3 width=7 by fqu_drop, cpx_delta, drop_pair, drop_drop/
- | #H destruct
- lapply (lift_inv_lref2_be … HVU2 ? ?) -HVU2 //
- ]
-| #I #G #L #V1 #T #V2 #HV12 #H @(ex3_intro … (②{I}V2.T))
- [1,3: /2 width=4 by fqu_pair_sn, cpx_pair_sn/
- | #H0 destruct /2 width=1 by/
- ]
-| #a #I #G #L #V #T1 #T2 #HT12 #H @(ex3_intro … (ⓑ{a,I}V.T2))
- [1,3: /2 width=4 by fqu_bind_dx, cpx_bind/
- | #H0 destruct /2 width=1 by/
- ]
-| #I #G #L #V #T1 #T2 #HT12 #H @(ex3_intro … (ⓕ{I}V.T2))
- [1,3: /2 width=4 by fqu_flat_dx, cpx_flat/
- | #H0 destruct /2 width=1 by/
- ]
-| #G #L #K #T1 #U1 #k #HLK #HTU1 #T2 #HT12 #H elim (lift_total T2 0 (k+1))
- #U2 #HTU2 @(ex3_intro … U2)
- [1,3: /2 width=10 by cpx_lift, fqu_drop/
- | #H0 destruct /3 width=5 by lift_inj/
-]
-qed-.
-
-lemma fquq_cpx_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡[h, o] U2 → (T2 = U2 → ⊥) →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H12 #U2 #HTU2 #H elim (fquq_inv_gen … H12) -H12
-[ #H12 elim (fqu_cpx_trans_neq … H12 … HTU2 H) -T2
- /3 width=4 by fqu_fquq, ex3_intro/
-| * #HG #HL #HT destruct /3 width=4 by ex3_intro/
-]
-qed-.
-
-lemma fqup_cpx_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
- ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡[h, o] U2 → (T2 = U2 → ⊥) →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐+ ⦃G2, L2, U2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind_dx … H) -G1 -L1 -T1
-[ #G1 #L1 #T1 #H12 #U2 #HTU2 #H elim (fqu_cpx_trans_neq … H12 … HTU2 H) -T2
- /3 width=4 by fqu_fqup, ex3_intro/
-| #G #G1 #L #L1 #T #T1 #H1 #_ #IH12 #U2 #HTU2 #H elim (IH12 … HTU2 H) -T2
- #U1 #HTU1 #H #H12 elim (fqu_cpx_trans_neq … H1 … HTU1 H) -T1
- /3 width=8 by fqup_strap2, ex3_intro/
-]
-qed-.
-
-lemma fqus_cpx_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
- ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡[h, o] U2 → (T2 = U2 → ⊥) →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐* ⦃G2, L2, U2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H12 #U2 #HTU2 #H elim (fqus_inv_gen … H12) -H12
-[ #H12 elim (fqup_cpx_trans_neq … H12 … HTU2 H) -T2
- /3 width=4 by fqup_fqus, ex3_intro/
-| * #HG #HL #HT destruct /3 width=4 by ex3_intro/
-]
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/multiple/lleq_drop.ma".
-include "basic_2/reduction/cpx_llpx_sn.ma".
-
-(* CONTEXT-SENSITIVE EXTENDED PARALLEL REDUCTION FOR TERMS ******************)
-
-(* Properties on lazy equivalence for local environments ********************)
-
-lemma lleq_cpx_trans: ∀h,o,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ➡[h, o] T2 →
- ∀L1. L1 ≡[T1, 0] L2 → ⦃G, L1⦄ ⊢ T1 ➡[h, o] T2.
-#h #o #G #L2 #T1 #T2 #H elim H -G -L2 -T1 -T2 /2 width=2 by cpx_st/
-[ #I #G #L2 #K2 #V1 #V2 #W2 #i #HLK2 #_ #HVW2 #IHV12 #L1 #H elim (lleq_fwd_lref_dx … H … HLK2) -L2
- [ #H elim (ylt_yle_false … H) //
- | * /3 width=7 by cpx_delta/
- ]
-| #a #I #G #L2 #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L1 #H elim (lleq_inv_bind_O … H) -H
- /3 width=1 by cpx_bind/
-| #I #G #L2 #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L1 #H elim (lleq_inv_flat … H) -H
- /3 width=1 by cpx_flat/
-| #G #L2 #V2 #T1 #T #T2 #_ #HT2 #IHT1 #L1 #H elim (lleq_inv_bind_O … H) -H
- /3 width=3 by cpx_zeta/
-| #G #L2 #W1 #T1 #T2 #_ #IHT12 #L1 #H elim (lleq_inv_flat … H) -H
- /3 width=1 by cpx_eps/
-| #G #L2 #W1 #W2 #T1 #_ #IHW12 #L1 #H elim (lleq_inv_flat … H) -H
- /3 width=1 by cpx_ct/
-| #a #G #L1 #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #L1 #H elim (lleq_inv_flat … H) -H
- #HV1 #H elim (lleq_inv_bind_O … H) -H /3 width=1 by cpx_beta/
-| #a #G #L1 #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV1 #IHW12 #IHT12 #L1 #H elim (lleq_inv_flat … H) -H
- #HV1 #H elim (lleq_inv_bind_O … H) -H /3 width=3 by cpx_theta/
-]
-qed-.
-
-lemma cpx_lleq_conf: ∀h,o,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ➡[h, o] T2 →
- ∀L1. L2 ≡[T1, 0] L1 → ⦃G, L1⦄ ⊢ T1 ➡[h, o] T2.
-/3 width=3 by lleq_cpx_trans, lleq_sym/ qed-.
-
-lemma cpx_lleq_conf_sn: ∀h,o,G. c_r_confluent1 … (cpx h o G) (lleq 0).
-/3 width=6 by cpx_llpx_sn_conf, lift_mono, ex2_intro/ qed-.
-
-lemma cpx_lleq_conf_dx: ∀h,o,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ➡[h, o] T2 →
- ∀L1. L1 ≡[T1, 0] L2 → L1 ≡[T2, 0] L2.
-/4 width=6 by cpx_lleq_conf_sn, 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_drop.ma".
-include "basic_2/reduction/cpx.ma".
-
-(* CONTEXT-SENSITIVE EXTENDED PARALLEL REDUCTION FOR TERMS ******************)
-
-(* Properties on lazy sn pointwise extensions *******************************)
-
-(* Note: lemma 1000 *)
-lemma cpx_llpx_sn_conf: ∀R. (∀L. reflexive … (R L)) → d_liftable R → d_deliftable_sn R →
- ∀h,o,G. c_r_confluent1 … (cpx h o G) (llpx_sn R 0).
-#R #H1R #H2R #H3R #h #o #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 #V1c #V2c #W2c #i #HLKs #_ #HVW2c #IHV12c #Ld #H elim (llpx_sn_inv_lref_ge_sn … H … HLKs) // -H
- #Kd #V1l #HLKd #HV1c #HV1sd
- lapply (drop_fwd_drop2 … HLKs) -HLKs #HLKs
- lapply (drop_fwd_drop2 … HLKd) -HLKd #HLKd
- @(llpx_sn_lift_le … HLKs HLKd … HVW2c) -HLKs -HLKd -HVW2c /2 width=1 by/ (**) (* full auto too slow *)
-| #a #I #G #Ls #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #Ld #H elim (llpx_sn_inv_bind_O … H) -H
- /4 width=5 by llpx_sn_bind_repl_SO, llpx_sn_bind/
-| #I #G #Ls #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
- /3 width=1 by llpx_sn_flat/
-| #G #Ls #V #T1 #T2 #T #_ #HT2 #IHT12 #Ld #H elim (llpx_sn_inv_bind_O … H) -H
- /3 width=10 by llpx_sn_inv_lift_le, drop_drop/
-| #G #Ls #V #T1 #T2 #_ #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H /2 width=1 by/
-| #G #Ls #V1 #V2 #T #_ #IHV12 #Ld #H elim (llpx_sn_inv_flat … H) -H /2 width=1 by/
-| #a #G #Ls #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
- #HV1 #H elim (llpx_sn_inv_bind_O … H) -H
- /4 width=5 by llpx_sn_bind_repl_SO, llpx_sn_flat, llpx_sn_bind/
-| #a #G #Ls #V1 #V2 #V #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
- #HV1 #H elim (llpx_sn_inv_bind_O … H) -H //
- #HW1 #HT1 @llpx_sn_bind_O /2 width=1 by/ @llpx_sn_flat (**) (* full auto too slow *)
- [ /3 width=10 by llpx_sn_lift_le, drop_drop/
- | /3 width=4 by llpx_sn_bind_repl_O/
-]
-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_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,o,G. lsub_trans … (cpx h o G) (lreq 0 (∞)).
-#h #o #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-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/predreducible_3.ma".
-include "basic_2/grammar/genv.ma".
-include "basic_2/substitution/drop.ma".
-
-(* REDUCIBLE TERMS FOR CONTEXT-SENSITIVE REDUCTION **************************)
-
-(* reducible binary items *)
-definition ri2: predicate item2 ≝
- λI. I = Bind2 true Abbr ∨ I = Flat2 Cast.
-
-(* irreducible binary binders *)
-definition ib2: relation2 bool bind2 ≝
- λa,I. I = Abst ∨ Bind2 a I = Bind2 false Abbr.
-
-(* activate genv *)
-(* reducible terms *)
-inductive crr (G:genv): relation2 lenv term ≝
-| crr_delta : ∀L,K,V,i. ⬇[i] L ≡ K.ⓓV → crr G L (#i)
-| crr_appl_sn: ∀L,V,T. crr G L V → crr G L (ⓐV.T)
-| crr_appl_dx: ∀L,V,T. crr G L T → crr G L (ⓐV.T)
-| crr_ri2 : ∀I,L,V,T. ri2 I → crr G L (②{I}V.T)
-| crr_ib2_sn : ∀a,I,L,V,T. ib2 a I → crr G L V → crr G L (ⓑ{a,I}V.T)
-| crr_ib2_dx : ∀a,I,L,V,T. ib2 a I → crr G (L.ⓑ{I}V) T → crr G L (ⓑ{a,I}V.T)
-| crr_beta : ∀a,L,V,W,T. crr G L (ⓐV.ⓛ{a}W.T)
-| crr_theta : ∀a,L,V,W,T. crr G L (ⓐV.ⓓ{a}W.T)
-.
-
-interpretation
- "reducibility for context-sensitive reduction (term)"
- 'PRedReducible G L T = (crr G L T).
-
-(* Basic inversion lemmas ***************************************************)
-
-fact crr_inv_sort_aux: ∀G,L,T,s. ⦃G, L⦄ ⊢ ➡ 𝐑⦃T⦄ → T = ⋆s → ⊥.
-#G #L #T #s0 * -L -T
-[ #L #K #V #i #HLK #H destruct
-| #L #V #T #_ #H destruct
-| #L #V #T #_ #H destruct
-| #I #L #V #T #_ #H destruct
-| #a #I #L #V #T #_ #_ #H destruct
-| #a #I #L #V #T #_ #_ #H destruct
-| #a #L #V #W #T #H destruct
-| #a #L #V #W #T #H destruct
-]
-qed-.
-
-lemma crr_inv_sort: ∀G,L,s. ⦃G, L⦄ ⊢ ➡ 𝐑⦃⋆s⦄ → ⊥.
-/2 width=6 by crr_inv_sort_aux/ qed-.
-
-fact crr_inv_lref_aux: ∀G,L,T,i. ⦃G, L⦄ ⊢ ➡ 𝐑⦃T⦄ → T = #i →
- ∃∃K,V. ⬇[i] L ≡ K.ⓓV.
-#G #L #T #j * -L -T
-[ #L #K #V #i #HLK #H destruct /2 width=3 by ex1_2_intro/
-| #L #V #T #_ #H destruct
-| #L #V #T #_ #H destruct
-| #I #L #V #T #_ #H destruct
-| #a #I #L #V #T #_ #_ #H destruct
-| #a #I #L #V #T #_ #_ #H destruct
-| #a #L #V #W #T #H destruct
-| #a #L #V #W #T #H destruct
-]
-qed-.
-
-lemma crr_inv_lref: ∀G,L,i. ⦃G, L⦄ ⊢ ➡ 𝐑⦃#i⦄ → ∃∃K,V. ⬇[i] L ≡ K.ⓓV.
-/2 width=4 by crr_inv_lref_aux/ qed-.
-
-fact crr_inv_gref_aux: ∀G,L,T,p. ⦃G, L⦄ ⊢ ➡ 𝐑⦃T⦄ → T = §p → ⊥.
-#G #L #T #q * -L -T
-[ #L #K #V #i #HLK #H destruct
-| #L #V #T #_ #H destruct
-| #L #V #T #_ #H destruct
-| #I #L #V #T #_ #H destruct
-| #a #I #L #V #T #_ #_ #H destruct
-| #a #I #L #V #T #_ #_ #H destruct
-| #a #L #V #W #T #H destruct
-| #a #L #V #W #T #H destruct
-]
-qed-.
-
-lemma crr_inv_gref: ∀G,L,p. ⦃G, L⦄ ⊢ ➡ 𝐑⦃§p⦄ → ⊥.
-/2 width=6 by crr_inv_gref_aux/ qed-.
-
-lemma trr_inv_atom: ∀G,I. ⦃G, ⋆⦄ ⊢ ➡ 𝐑⦃⓪{I}⦄ → ⊥.
-#G * #i #H
-[ elim (crr_inv_sort … H)
-| elim (crr_inv_lref … H) -H #L #V #H
- elim (drop_inv_atom1 … H) -H #H destruct
-| elim (crr_inv_gref … H)
-]
-qed-.
-
-fact crr_inv_ib2_aux: ∀a,I,G,L,W,U,T. ib2 a I → ⦃G, L⦄ ⊢ ➡ 𝐑⦃T⦄ → T = ⓑ{a,I}W.U →
- ⦃G, L⦄ ⊢ ➡ 𝐑⦃W⦄ ∨ ⦃G, L.ⓑ{I}W⦄ ⊢ ➡ 𝐑⦃U⦄.
-#G #b #J #L #W0 #U #T #HI * -L -T
-[ #L #K #V #i #_ #H destruct
-| #L #V #T #_ #H destruct
-| #L #V #T #_ #H destruct
-| #I #L #V #T #H1 #H2 destruct
- elim H1 -H1 #H destruct
- elim HI -HI #H destruct
-| #a #I #L #V #T #_ #HV #H destruct /2 width=1 by or_introl/
-| #a #I #L #V #T #_ #HT #H destruct /2 width=1 by or_intror/
-| #a #L #V #W #T #H destruct
-| #a #L #V #W #T #H destruct
-]
-qed-.
-
-lemma crr_inv_ib2: ∀a,I,G,L,W,T. ib2 a I → ⦃G, L⦄ ⊢ ➡ 𝐑⦃ⓑ{a,I}W.T⦄ →
- ⦃G, L⦄ ⊢ ➡ 𝐑⦃W⦄ ∨ ⦃G, L.ⓑ{I}W⦄ ⊢ ➡ 𝐑⦃T⦄.
-/2 width=5 by crr_inv_ib2_aux/ qed-.
-
-fact crr_inv_appl_aux: ∀G,L,W,U,T. ⦃G, L⦄ ⊢ ➡ 𝐑⦃T⦄ → T = ⓐW.U →
- ∨∨ ⦃G, L⦄ ⊢ ➡ 𝐑⦃W⦄ | ⦃G, L⦄ ⊢ ➡ 𝐑⦃U⦄ | (𝐒⦃U⦄ → ⊥).
-#G #L #W0 #U #T * -L -T
-[ #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/
-| #I #L #V #T #H1 #H2 destruct
- elim H1 -H1 #H destruct
-| #a #I #L #V #T #_ #_ #H destruct
-| #a #I #L #V #T #_ #_ #H destruct
-| #a #L #V #W #T #H destruct
- @or3_intro2 #H elim (simple_inv_bind … H)
-| #a #L #V #W #T #H destruct
- @or3_intro2 #H elim (simple_inv_bind … H)
-]
-qed-.
-
-lemma crr_inv_appl: ∀G,L,V,T. ⦃G, L⦄ ⊢ ➡ 𝐑⦃ⓐV.T⦄ →
- ∨∨ ⦃G, L⦄ ⊢ ➡ 𝐑⦃V⦄ | ⦃G, L⦄ ⊢ ➡ 𝐑⦃T⦄ | (𝐒⦃T⦄ → ⊥).
-/2 width=3 by crr_inv_appl_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_drop.ma".
-include "basic_2/reduction/crr.ma".
-
-(* REDUCIBLE TERMS FOR CONTEXT-SENSITIVE REDUCTION **************************)
-
-(* Properties on relocation *************************************************)
-
-lemma crr_lift: ∀G,K,T. ⦃G, K⦄ ⊢ ➡ 𝐑⦃T⦄ → ∀L,c,l,k. ⬇[c, l, k] L ≡ K →
- ∀U. ⬆[l, k] T ≡ U → ⦃G, L⦄ ⊢ ➡ 𝐑⦃U⦄.
-#G #K #T #H elim H -K -T
-[ #K #K0 #V #i #HK0 #L #c #l #k #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 #c #l #k #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 #c #l #k #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 #c #l #k #_ #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 #c #l #k #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 #c #l #k #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 #c #l #k #_ #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 #c #l #k #_ #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,c,l,k. ⬇[c, l, k] L ≡ K →
- ∀T. ⬆[l, k] T ≡ U → ⦃G, K⦄ ⊢ ➡ 𝐑⦃T⦄.
-#G #L #U #H elim H -L -U
-[ #L #L0 #W #i #HK0 #K #c #l #k #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 #c #l #k #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 #c #l #k #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 #c #l #k #_ #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 #c #l #k #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 #c #l #k #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 #c #l #k #_ #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 #c #l #k #_ #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/
-]
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/predreducible_5.ma".
-include "basic_2/static/sd.ma".
-include "basic_2/reduction/crr.ma".
-
-(* REDUCIBLE TERMS FOR CONTEXT-SENSITIVE EXTENDED REDUCTION *****************)
-
-(* activate genv *)
-(* extended reducible terms *)
-inductive crx (h) (o) (G:genv): relation2 lenv term ≝
-| crx_sort : ∀L,s,d. deg h o s (d+1) → crx h o G L (⋆s)
-| crx_delta : ∀I,L,K,V,i. ⬇[i] L ≡ K.ⓑ{I}V → crx h o G L (#i)
-| crx_appl_sn: ∀L,V,T. crx h o G L V → crx h o G L (ⓐV.T)
-| crx_appl_dx: ∀L,V,T. crx h o G L T → crx h o G L (ⓐV.T)
-| crx_ri2 : ∀I,L,V,T. ri2 I → crx h o G L (②{I}V.T)
-| crx_ib2_sn : ∀a,I,L,V,T. ib2 a I → crx h o G L V → crx h o G L (ⓑ{a,I}V.T)
-| crx_ib2_dx : ∀a,I,L,V,T. ib2 a I → crx h o G (L.ⓑ{I}V) T → crx h o G L (ⓑ{a,I}V.T)
-| crx_beta : ∀a,L,V,W,T. crx h o G L (ⓐV. ⓛ{a}W.T)
-| crx_theta : ∀a,L,V,W,T. crx h o G L (ⓐV. ⓓ{a}W.T)
-.
-
-interpretation
- "reducibility for context-sensitive extended reduction (term)"
- 'PRedReducible h o G L T = (crx h o G L T).
-
-(* Basic properties *********************************************************)
-
-lemma crr_crx: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ➡ 𝐑⦃T⦄ → ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃T⦄.
-#h #o #G #L #T #H elim H -L -T
-/2 width=4 by crx_delta, crx_appl_sn, crx_appl_dx, crx_ri2, crx_ib2_sn, crx_ib2_dx, crx_beta, crx_theta/
-qed.
-
-(* Basic inversion lemmas ***************************************************)
-
-fact crx_inv_sort_aux: ∀h,o,G,L,T,s. ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃T⦄ → T = ⋆s →
- ∃d. deg h o s (d+1).
-#h #o #G #L #T #s0 * -L -T
-[ #L #s #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
-| #I #L #V #T #_ #H destruct
-| #a #I #L #V #T #_ #_ #H destruct
-| #a #I #L #V #T #_ #_ #H destruct
-| #a #L #V #W #T #H destruct
-| #a #L #V #W #T #H destruct
-]
-qed-.
-
-lemma crx_inv_sort: ∀h,o,G,L,s. ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃⋆s⦄ → ∃d. deg h o s (d+1).
-/2 width=5 by crx_inv_sort_aux/ qed-.
-
-fact crx_inv_lref_aux: ∀h,o,G,L,T,i. ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃T⦄ → T = #i →
- ∃∃I,K,V. ⬇[i] L ≡ K.ⓑ{I}V.
-#h #o #G #L #T #j * -L -T
-[ #L #s #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
-| #I #L #V #T #_ #H destruct
-| #a #I #L #V #T #_ #_ #H destruct
-| #a #I #L #V #T #_ #_ #H destruct
-| #a #L #V #W #T #H destruct
-| #a #L #V #W #T #H destruct
-]
-qed-.
-
-lemma crx_inv_lref: ∀h,o,G,L,i. ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃#i⦄ → ∃∃I,K,V. ⬇[i] L ≡ K.ⓑ{I}V.
-/2 width=6 by crx_inv_lref_aux/ qed-.
-
-fact crx_inv_gref_aux: ∀h,o,G,L,T,p. ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃T⦄ → T = §p → ⊥.
-#h #o #G #L #T #q * -L -T
-[ #L #s #d #_ #H destruct
-| #I #L #K #V #i #HLK #H destruct
-| #L #V #T #_ #H destruct
-| #L #V #T #_ #H destruct
-| #I #L #V #T #_ #H destruct
-| #a #I #L #V #T #_ #_ #H destruct
-| #a #I #L #V #T #_ #_ #H destruct
-| #a #L #V #W #T #H destruct
-| #a #L #V #W #T #H destruct
-]
-qed-.
-
-lemma crx_inv_gref: ∀h,o,G,L,p. ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃§p⦄ → ⊥.
-/2 width=8 by crx_inv_gref_aux/ qed-.
-
-lemma trx_inv_atom: ∀h,o,I,G. ⦃G, ⋆⦄ ⊢ ➡[h, o] 𝐑⦃⓪{I}⦄ →
- ∃∃s,d. deg h o s (d+1) & I = Sort s.
-#h #o * #i #G #H
-[ elim (crx_inv_sort … H) -H /2 width=4 by ex2_2_intro/
-| elim (crx_inv_lref … H) -H #I #L #V #H
- elim (drop_inv_atom1 … H) -H #H destruct
-| elim (crx_inv_gref … H)
-]
-qed-.
-
-fact crx_inv_ib2_aux: ∀h,o,a,I,G,L,W,U,T. ib2 a I → ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃T⦄ →
- T = ⓑ{a,I}W.U → ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃W⦄ ∨ ⦃G, L.ⓑ{I}W⦄ ⊢ ➡[h, o] 𝐑⦃U⦄.
-#h #o #b #J #G #L #W0 #U #T #HI * -L -T
-[ #L #s #d #_ #H destruct
-| #I #L #K #V #i #_ #H destruct
-| #L #V #T #_ #H destruct
-| #L #V #T #_ #H destruct
-| #I #L #V #T #H1 #H2 destruct
- elim H1 -H1 #H destruct
- elim HI -HI #H destruct
-| #a #I #L #V #T #_ #HV #H destruct /2 width=1 by or_introl/
-| #a #I #L #V #T #_ #HT #H destruct /2 width=1 by or_intror/
-| #a #L #V #W #T #H destruct
-| #a #L #V #W #T #H destruct
-]
-qed-.
-
-lemma crx_inv_ib2: ∀h,o,a,I,G,L,W,T. ib2 a I → ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃ⓑ{a,I}W.T⦄ →
- ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃W⦄ ∨ ⦃G, L.ⓑ{I}W⦄ ⊢ ➡[h, o] 𝐑⦃T⦄.
-/2 width=5 by crx_inv_ib2_aux/ qed-.
-
-fact crx_inv_appl_aux: ∀h,o,G,L,W,U,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃T⦄ → T = ⓐW.U →
- ∨∨ ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃W⦄ | ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃U⦄ | (𝐒⦃U⦄ → ⊥).
-#h #o #G #L #W0 #U #T * -L -T
-[ #L #s #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/
-| #I #L #V #T #H1 #H2 destruct
- elim H1 -H1 #H destruct
-| #a #I #L #V #T #_ #_ #H destruct
-| #a #I #L #V #T #_ #_ #H destruct
-| #a #L #V #W #T #H destruct
- @or3_intro2 #H elim (simple_inv_bind … H)
-| #a #L #V #W #T #H destruct
- @or3_intro2 #H elim (simple_inv_bind … H)
-]
-qed-.
-
-lemma crx_inv_appl: ∀h,o,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃ⓐV.T⦄ →
- ∨∨ ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃V⦄ | ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃T⦄ | (𝐒⦃T⦄ → ⊥).
-/2 width=3 by crx_inv_appl_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_drop.ma".
-include "basic_2/reduction/crx.ma".
-
-(* REDUCIBLE TERMS FOR CONTEXT-SENSITIVE EXTENDED REDUCTION *****************)
-
-(* Properties on relocation *************************************************)
-
-lemma crx_lift: ∀h,o,G,K,T. ⦃G, K⦄ ⊢ ➡[h, o] 𝐑⦃T⦄ → ∀L,c,l,k. ⬇[c, l, k] L ≡ K →
- ∀U. ⬆[l, k] T ≡ U → ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃U⦄.
-#h #o #G #K #T #H elim H -K -T
-[ #K #s #d #Hkd #L #c #l #k #_ #X #H
- >(lift_inv_sort1 … H) -X /2 width=2 by crx_sort/
-| #I #K #K0 #V #i #HK0 #L #c #l #k #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 #c #l #k #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 #c #l #k #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 #c #l #k #_ #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 #c #l #k #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 #c #l #k #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 #c #l #k #_ #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 #c #l #k #_ #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,o,G,L,U. ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃U⦄ → ∀K,c,l,k. ⬇[c, l, k] L ≡ K →
- ∀T. ⬆[l, k] T ≡ U → ⦃G, K⦄ ⊢ ➡[h, o] 𝐑⦃T⦄.
-#h #o #G #L #U #H elim H -L -U
-[ #L #s #d #Hkd #K #c #l #k #_ #X #H
- >(lift_inv_sort2 … H) -X /2 width=2 by crx_sort/
-| #I #L #L0 #W #i #HK0 #K #c #l #k #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 #c #l #k #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 #c #l #k #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 #c #l #k #_ #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 #c #l #k #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 #c #l #k #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 #c #l #k #_ #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 #c #l #k #_ #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/
-]
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/btpredproper_8.ma".
-include "basic_2/substitution/fqu.ma".
-include "basic_2/multiple/lleq.ma".
-include "basic_2/reduction/lpx.ma".
-
-(* "RST" PROPER PARALLEL COMPUTATION FOR CLOSURES ***************************)
-
-inductive fpb (h) (o) (G1) (L1) (T1): relation3 genv lenv term ≝
-| fpb_fqu: ∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → fpb h o G1 L1 T1 G2 L2 T2
-| fpb_cpx: ∀T2. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] T2 → (T1 = T2 → ⊥) → fpb h o G1 L1 T1 G1 L1 T2
-| fpb_lpx: ∀L2. ⦃G1, L1⦄ ⊢ ➡[h, o] L2 → (L1 ≡[T1, 0] L2 → ⊥) → fpb h o G1 L1 T1 G1 L2 T1
-.
-
-interpretation
- "'rst' proper parallel reduction (closure)"
- 'BTPRedProper h o G1 L1 T1 G2 L2 T2 = (fpb h o G1 L1 T1 G2 L2 T2).
-
-(* Basic properties *********************************************************)
-
-lemma cpr_fpb: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → (T1 = T2 → ⊥) →
- ⦃G, L, T1⦄ ≻[h, o] ⦃G, L, T2⦄.
-/3 width=1 by fpb_cpx, cpr_cpx/ qed.
-
-lemma lpr_fpb: ∀h,o,G,L1,L2,T. ⦃G, L1⦄ ⊢ ➡ L2 → (L1 ≡[T, 0] L2 → ⊥) →
- ⦃G, L1, T⦄ ≻[h, o] ⦃G, L2, T⦄.
-/3 width=1 by fpb_lpx, lpr_lpx/ 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/fleq.ma".
-include "basic_2/reduction/fpb_lleq.ma".
-
-(* "RST" PROPER PARALLEL COMPUTATION FOR CLOSURES ***************************)
-
-(* Properties on lazy equivalence for closures ******************************)
-
-lemma fleq_fpb_trans: ∀h,o,F1,F2,K1,K2,T1,T2. ⦃F1, K1, T1⦄ ≡[0] ⦃F2, K2, T2⦄ →
- ∀G2,L2,U2. ⦃F2, K2, T2⦄ ≻[h, o] ⦃G2, L2, U2⦄ →
- ∃∃G1,L1,U1. ⦃F1, K1, T1⦄ ≻[h, o] ⦃G1, L1, U1⦄ & ⦃G1, L1, U1⦄ ≡[0] ⦃G2, L2, U2⦄.
-#h #o #F1 #F2 #K1 #K2 #T1 #T2 * -F2 -K2 -T2
-#K2 #HK12 #G2 #L2 #U2 #H12 elim (lleq_fpb_trans … HK12 … H12) -K2
-/3 width=5 by fleq_intro, ex2_3_intro/
-qed-.
-
-(* Inversion lemmas on lazy equivalence for closures ************************)
-
-lemma fpb_inv_fleq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → ⊥.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2
-[ #G2 #L2 #T2 #H12 #H elim (fleq_inv_gen … H) -H
- /3 width=8 by lleq_fwd_length, fqu_inv_eq/
-| #T2 #_ #HnT #H elim (fleq_inv_gen … H) -H /2 width=1 by/
-| #L2 #_ #HnL #H elim (fleq_inv_gen … H) -H /2 width=1 by/
-]
-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/unfold/lstas_da.ma".
-include "basic_2/reduction/cpx_lift.ma".
-include "basic_2/reduction/fpb.ma".
-
-(* "RST" PROPER PARALLEL COMPUTATION FOR CLOSURES ***************************)
-
-(* Advanced properties ******************************************************)
-
-lemma sta_fpb: ∀h,o,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 ▪[h, o] d+1 →
- ⦃G, L⦄ ⊢ T1 •*[h, 1] T2 → ⦃G, L, T1⦄ ≻[h, o] ⦃G, L, T2⦄.
-#h #o #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.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/multiple/lleq_fqus.ma".
-include "basic_2/multiple/lleq_lleq.ma".
-include "basic_2/reduction/lpx_lleq.ma".
-include "basic_2/reduction/fpb.ma".
-
-(* "RST" PROPER PARALLEL COMPUTATION FOR CLOSURES ***************************)
-
-(* Properties on lazy equivalence for local environments ********************)
-
-lemma lleq_fpb_trans: ∀h,o,F,K1,K2,T. K1 ≡[T, 0] K2 →
- ∀G,L2,U. ⦃F, K2, T⦄ ≻[h, o] ⦃G, L2, U⦄ →
- ∃∃L1. ⦃F, K1, T⦄ ≻[h, o] ⦃G, L1, U⦄ & L1 ≡[U, 0] L2.
-#h #o #F #K1 #K2 #T #HT #G #L2 #U * -G -L2 -U
-[ #G #L2 #U #H2 elim (lleq_fqu_trans … H2 … HT) -K2
- /3 width=3 by fpb_fqu, ex2_intro/
-| /4 width=10 by fpb_cpx, cpx_lleq_conf_sn, lleq_cpx_trans, ex2_intro/
-| #L2 #HKL2 #HnKL2 elim (lleq_lpx_trans … HKL2 … HT) -HKL2
- /6 width=3 by fpb_lpx, lleq_canc_sn, ex2_intro/ (* 2 lleq_canc_sn *)
-]
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/btpred_8.ma".
-include "basic_2/substitution/fquq.ma".
-include "basic_2/multiple/lleq.ma".
-include "basic_2/reduction/lpx.ma".
-
-(* "QRST" PARALLEL REDUCTION FOR CLOSURES ***********************************)
-
-inductive fpbq (h) (o) (G1) (L1) (T1): relation3 genv lenv term ≝
-| fpbq_fquq: ∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → fpbq h o G1 L1 T1 G2 L2 T2
-| fpbq_cpx : ∀T2. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] T2 → fpbq h o G1 L1 T1 G1 L1 T2
-| fpbq_lpx : ∀L2. ⦃G1, L1⦄ ⊢ ➡[h, o] L2 → fpbq h o G1 L1 T1 G1 L2 T1
-| fpbq_lleq: ∀L2. L1 ≡[T1, 0] L2 → fpbq h o G1 L1 T1 G1 L2 T1
-.
-
-interpretation
- "'qrst' parallel reduction (closure)"
- 'BTPRed h o G1 L1 T1 G2 L2 T2 = (fpbq h o G1 L1 T1 G2 L2 T2).
-
-(* Basic properties *********************************************************)
-
-lemma fpbq_refl: ∀h,o. tri_reflexive … (fpbq h o).
-/2 width=1 by fpbq_cpx/ qed.
-
-lemma cpr_fpbq: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ⦃G, L, T1⦄ ≽[h, o] ⦃G, L, T2⦄.
-/3 width=1 by fpbq_cpx, cpr_cpx/ qed.
-
-lemma lpr_fpbq: ∀h,o,G,L1,L2,T. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L1, T⦄ ≽[h, o] ⦃G, L2, T⦄.
-/3 width=1 by fpbq_lpx, lpr_lpx/ qed.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/static/aaa_fqus.ma".
-include "basic_2/static/aaa_lleq.ma".
-include "basic_2/reduction/lpx_aaa.ma".
-include "basic_2/reduction/fpbq.ma".
-
-(* "QRST" PARALLEL REDUCTION FOR CLOSURES ***********************************)
-
-(* Properties on atomic arity assignment for terms **************************)
-
-lemma fpbq_aaa_conf: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ →
- ∀A1. ⦃G1, L1⦄ ⊢ T1 ⁝ A1 → ∃A2. ⦃G2, L2⦄ ⊢ T2 ⁝ A2.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2
-/3 width=6 by aaa_lleq_conf, lpx_aaa_conf, cpx_aaa_conf, aaa_fquq_conf, ex_intro/
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/btpredalt_8.ma".
-include "basic_2/reduction/fpb_fleq.ma".
-include "basic_2/reduction/fpbq.ma".
-
-(* "QRST" PARALLEL REDUCTION FOR CLOSURES ***********************************)
-
-(* alternative definition of fpbq *)
-definition fpbqa: ∀h. sd h → tri_relation genv lenv term ≝
- λh,o,G1,L1,T1,G2,L2,T2.
- ⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ ∨ ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄.
-
-interpretation
- "'qrst' parallel reduction (closure) alternative"
- 'BTPRedAlt h o G1 L1 T1 G2 L2 T2 = (fpbqa h o G1 L1 T1 G2 L2 T2).
-
-(* Basic properties *********************************************************)
-
-lemma fleq_fpbq: ∀h,o,G1,G2,L1,L2,T1,T2.
- ⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 * /2 width=1 by fpbq_lleq/
-qed.
-
-lemma fpb_fpbq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2
-/3 width=1 by fpbq_fquq, fpbq_cpx, fpbq_lpx, fqu_fquq/
-qed.
-
-(* Main properties **********************************************************)
-
-theorem fpbq_fpbqa: ∀h,o,G1,G2,L1,L2,T1,T2.
- ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ≽≽[h, o] ⦃G2, L2, T2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2
-[ #G2 #L2 #T2 #H elim (fquq_inv_gen … H) -H
- [ /3 width=1 by fpb_fqu, or_intror/
- | * #H1 #H2 #H3 destruct /2 width=1 by or_introl/
- ]
-| #T2 #HT12 elim (eq_term_dec T1 T2)
- #HnT12 destruct /4 width=1 by fpb_cpx, or_intror, or_introl/
-| #L2 #HL12 elim (lleq_dec … T1 L1 L2 0)
- /4 width=1 by fpb_lpx, fleq_intro, or_intror, or_introl/
-| /3 width=1 by fleq_intro, or_introl/
-]
-qed.
-
-theorem fpbqa_inv_fpbq: ∀h,o,G1,G2,L1,L2,T1,T2.
- ⦃G1, L1, T1⦄ ≽≽[h, o] ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 * /2 width=1 by fleq_fpbq, fpb_fpbq/
-qed-.
-
-(* Advanced eliminators *****************************************************)
-
-lemma fpbq_ind_alt: ∀h,o,G1,G2,L1,L2,T1,T2. ∀R:Prop.
- (⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → R) →
- (⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ → R) →
- ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ → R.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #R #HI1 #HI2 #H elim (fpbq_fpbqa … H) /2 width=1 by/
-qed-.
-
-(* aternative definition of fpb *********************************************)
-
-lemma fpb_fpbq_alt: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ ∧ (⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → ⊥).
-/3 width=10 by fpb_fpbq, fpb_inv_fleq, conj/ qed.
-
-lemma fpbq_inv_fpb_alt: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ →
- (⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → ⊥) → ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #H0 @(fpbq_ind_alt … H) -H //
-#H elim H0 -H0 //
-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_lift.ma".
-include "basic_2/reduction/fpbq.ma".
-
-(* "QRST" PARALLEL REDUCTION FOR CLOSURES ***********************************)
-
-(* Advanced properties ******************************************************)
-
-lemma sta_fpbq: ∀h,o,G,L,T1,T2,d.
- ⦃G, L⦄ ⊢ T1 ▪[h, o] d+1 → ⦃G, L⦄ ⊢ T1 •*[h, 1] T2 →
- ⦃G, L, T1⦄ ≽[h, o] ⦃G, L, T2⦄.
-/3 width=4 by fpbq_cpx, sta_cpx/ qed.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/predsn_3.ma".
-include "basic_2/substitution/lpx_sn.ma".
-include "basic_2/reduction/cpr.ma".
-
-(* SN PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS *****************************)
-
-definition lpr: relation3 genv lenv lenv ≝ λG. lpx_sn (cpr G).
-
-interpretation "parallel reduction (local environment, sn variant)"
- 'PRedSn G L1 L2 = (lpr G L1 L2).
-
-(* Basic inversion lemmas ***************************************************)
-
-(* Basic_1: includes: wcpr0_gen_sort *)
-lemma lpr_inv_atom1: ∀G,L2. ⦃G, ⋆⦄ ⊢ ➡ L2 → L2 = ⋆.
-/2 width=4 by lpx_sn_inv_atom1_aux/ qed-.
-
-(* Basic_1: includes: wcpr0_gen_head *)
-lemma lpr_inv_pair1: ∀I,G,K1,V1,L2. ⦃G, K1.ⓑ{I}V1⦄ ⊢ ➡ L2 →
- ∃∃K2,V2. ⦃G, K1⦄ ⊢ ➡ K2 & ⦃G, K1⦄ ⊢ V1 ➡ V2 & L2 = K2.ⓑ{I}V2.
-/2 width=3 by lpx_sn_inv_pair1_aux/ qed-.
-
-lemma lpr_inv_atom2: ∀G,L1. ⦃G, L1⦄ ⊢ ➡ ⋆ → L1 = ⋆.
-/2 width=4 by lpx_sn_inv_atom2_aux/ qed-.
-
-lemma lpr_inv_pair2: ∀I,G,L1,K2,V2. ⦃G, L1⦄ ⊢ ➡ K2.ⓑ{I}V2 →
- ∃∃K1,V1. ⦃G, K1⦄ ⊢ ➡ K2 & ⦃G, K1⦄ ⊢ V1 ➡ V2 & L1 = K1. ⓑ{I} V1.
-/2 width=3 by lpx_sn_inv_pair2_aux/ qed-.
-
-(* Basic properties *********************************************************)
-
-(* Note: lemma 250 *)
-lemma lpr_refl: ∀G,L. ⦃G, L⦄ ⊢ ➡ L.
-/2 width=1 by lpx_sn_refl/ qed.
-
-lemma lpr_pair: ∀I,G,K1,K2,V1,V2. ⦃G, K1⦄ ⊢ ➡ K2 → ⦃G, K1⦄ ⊢ V1 ➡ V2 →
- ⦃G, K1.ⓑ{I}V1⦄ ⊢ ➡ K2.ⓑ{I}V2.
-/2 width=1 by lpx_sn_pair/ qed.
-
-(* Basic forward lemmas *****************************************************)
-
-lemma lpr_fwd_length: ∀G,L1,L2. ⦃G, L1⦄ ⊢ ➡ L2 → |L1| = |L2|.
-/2 width=2 by lpx_sn_fwd_length/ qed-.
-
-(* Basic_1: removed theorems 3: wcpr0_getl wcpr0_getl_back
- pr0_subst1_back
-*)
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/substitution/lpx_sn_drop.ma".
-include "basic_2/substitution/fquq_alt.ma".
-include "basic_2/reduction/cpr_lift.ma".
-include "basic_2/reduction/lpr.ma".
-
-(* SN PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS *****************************)
-
-(* Properties on local environment slicing ***********************************)
-
-(* Basic_1: includes: wcpr0_drop *)
-lemma lpr_drop_conf: ∀G. dropable_sn (lpr G).
-/3 width=6 by lpx_sn_deliftable_dropable, cpr_inv_lift1/ qed-.
-
-(* Basic_1: includes: wcpr0_drop_back *)
-lemma drop_lpr_trans: ∀G. dedropable_sn (lpr G).
-/3 width=10 by lpx_sn_liftable_dedropable, cpr_lift/ qed-.
-
-lemma lpr_drop_trans_O1: ∀G. dropable_dx (lpr G).
-/2 width=3 by lpx_sn_dropable/ qed-.
-
-(* Properties on context-sensitive parallel reduction for terms *************)
-
-lemma fqu_cpr_trans_dx: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
- ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡ U2 →
- ∃∃L,U1. ⦃G1, L1⦄ ⊢ ➡ L & ⦃G1, L⦄ ⊢ T1 ➡ U1 & ⦃G1, L, U1⦄ ⊐ ⦃G2, L2, U2⦄.
-#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
-/3 width=5 by fqu_lref_O, fqu_pair_sn, fqu_bind_dx, fqu_flat_dx, lpr_pair, cpr_pair_sn, cpr_atom, cpr_bind, cpr_flat, ex3_2_intro/
-#G #L #K #U #T #k #HLK #HUT #U2 #HU2
-elim (lift_total U2 0 (k+1)) #T2 #HUT2
-lapply (cpr_lift … HU2 … HLK … HUT … HUT2) -HU2 -HUT /3 width=9 by fqu_drop, ex3_2_intro/
-qed-.
-
-lemma fquq_cpr_trans_dx: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡ U2 →
- ∃∃L,U1. ⦃G1, L1⦄ ⊢ ➡ L & ⦃G1, L⦄ ⊢ T1 ➡ U1 & ⦃G1, L, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
-#G1 #G2 #L1 #L2 #T1 #T2 #H #U2 #HTU2 elim (fquq_inv_gen … H) -H
-[ #HT12 elim (fqu_cpr_trans_dx … HT12 … HTU2) /3 width=5 by fqu_fquq, ex3_2_intro/
-| * #H1 #H2 #H3 destruct /2 width=5 by ex3_2_intro/
-]
-qed-.
-
-lemma fqu_cpr_trans_sn: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
- ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡ U2 →
- ∃∃L,U1. ⦃G1, L1⦄ ⊢ ➡ L & ⦃G1, L1⦄ ⊢ T1 ➡ U1 & ⦃G1, L, U1⦄ ⊐ ⦃G2, L2, U2⦄.
-#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
-/3 width=5 by fqu_lref_O, fqu_pair_sn, fqu_bind_dx, fqu_flat_dx, lpr_pair, cpr_pair_sn, cpr_atom, cpr_bind, cpr_flat, ex3_2_intro/
-#G #L #K #U #T #k #HLK #HUT #U2 #HU2
-elim (lift_total U2 0 (k+1)) #T2 #HUT2
-lapply (cpr_lift … HU2 … HLK … HUT … HUT2) -HU2 -HUT /3 width=9 by fqu_drop, ex3_2_intro/
-qed-.
-
-lemma fquq_cpr_trans_sn: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡ U2 →
- ∃∃L,U1. ⦃G1, L1⦄ ⊢ ➡ L & ⦃G1, L1⦄ ⊢ T1 ➡ U1 & ⦃G1, L, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
-#G1 #G2 #L1 #L2 #T1 #T2 #H #U2 #HTU2 elim (fquq_inv_gen … H) -H
-[ #HT12 elim (fqu_cpr_trans_sn … HT12 … HTU2) /3 width=5 by fqu_fquq, ex3_2_intro/
-| * #H1 #H2 #H3 destruct /2 width=5 by ex3_2_intro/
-]
-qed-.
-
-lemma fqu_lpr_trans: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
- ∀K2. ⦃G2, L2⦄ ⊢ ➡ K2 →
- ∃∃K1,T. ⦃G1, L1⦄ ⊢ ➡ K1 & ⦃G1, L1⦄ ⊢ T1 ➡ T & ⦃G1, K1, T⦄ ⊐ ⦃G2, K2, T2⦄.
-#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
-/3 width=5 by fqu_lref_O, fqu_pair_sn, fqu_flat_dx, lpr_pair, ex3_2_intro/
-[ #a #I #G2 #L2 #V2 #T2 #X #H elim (lpr_inv_pair1 … H) -H
- #K2 #W2 #HLK2 #HVW2 #H destruct
- /3 width=5 by fqu_fquq, cpr_pair_sn, fqu_bind_dx, ex3_2_intro/
-| #G #L1 #K1 #T1 #U1 #k #HLK1 #HTU1 #K2 #HK12
- elim (drop_lpr_trans … HLK1 … HK12) -HK12
- /3 width=7 by fqu_drop, ex3_2_intro/
-]
-qed-.
-
-lemma fquq_lpr_trans: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ∀K2. ⦃G2, L2⦄ ⊢ ➡ K2 →
- ∃∃K1,T. ⦃G1, L1⦄ ⊢ ➡ K1 & ⦃G1, L1⦄ ⊢ T1 ➡ T & ⦃G1, K1, T⦄ ⊐⸮ ⦃G2, K2, T2⦄.
-#G1 #G2 #L1 #L2 #T1 #T2 #H #K2 #HLK2 elim (fquq_inv_gen … H) -H
-[ #HT12 elim (fqu_lpr_trans … HT12 … HLK2) /3 width=5 by fqu_fquq, ex3_2_intro/
-| * #H1 #H2 #H3 destruct /2 width=5 by ex3_2_intro/
-]
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/substitution/lpx_sn_lpx_sn.ma".
-include "basic_2/multiple/fqup.ma".
-include "basic_2/reduction/lpr_drop.ma".
-
-(* SN PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS *****************************)
-
-(* Main properties on context-sensitive parallel reduction for terms ********)
-
-fact cpr_conf_lpr_atom_atom:
- ∀I,G,L1,L2. ∃∃T. ⦃G, L1⦄ ⊢ ⓪{I} ➡ T & ⦃G, L2⦄ ⊢ ⓪{I} ➡ T.
-/2 width=3 by cpr_atom, ex2_intro/ qed-.
-
-fact cpr_conf_lpr_atom_delta:
- ∀G,L0,i. (
- ∀L,T. ⦃G, L0, #i⦄ ⊐+ ⦃G, L, T⦄ →
- ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
- ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
- ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
- ) →
- ∀K0,V0. ⬇[i] L0 ≡ K0.ⓓV0 →
- ∀V2. ⦃G, K0⦄ ⊢ V0 ➡ V2 → ∀T2. ⬆[O, i + 1] V2 ≡ T2 →
- ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
- ∃∃T. ⦃G, L1⦄ ⊢ #i ➡ T & ⦃G, L2⦄ ⊢ T2 ➡ T.
-#G #L0 #i #IH #K0 #V0 #HLK0 #V2 #HV02 #T2 #HVT2 #L1 #HL01 #L2 #HL02
-elim (lpr_drop_conf … HLK0 … HL01) -HL01 #X1 #H1 #HLK1
-elim (lpr_inv_pair1 … H1) -H1 #K1 #V1 #HK01 #HV01 #H destruct
-elim (lpr_drop_conf … HLK0 … HL02) -HL02 #X2 #H2 #HLK2
-elim (lpr_inv_pair1 … H2) -H2 #K2 #W2 #HK02 #_ #H destruct
-lapply (drop_fwd_drop2 … HLK2) -W2 #HLK2
-lapply (fqup_lref … G … HLK0) -HLK0 #HLK0
-elim (IH … HLK0 … HV01 … HV02 … HK01 … HK02) -L0 -K0 -V0 #V #HV1 #HV2
-elim (lift_total V 0 (i+1))
-/3 width=12 by cpr_lift, cpr_delta, ex2_intro/
-qed-.
-
-(* Basic_1: includes: pr0_delta_delta pr2_delta_delta *)
-fact cpr_conf_lpr_delta_delta:
- ∀G,L0,i. (
- ∀L,T. ⦃G, L0, #i⦄ ⊐+ ⦃G, L, T⦄ →
- ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
- ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
- ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
- ) →
- ∀K0,V0. ⬇[i] L0 ≡ K0.ⓓV0 →
- ∀V1. ⦃G, K0⦄ ⊢ V0 ➡ V1 → ∀T1. ⬆[O, i + 1] V1 ≡ T1 →
- ∀KX,VX. ⬇[i] L0 ≡ KX.ⓓVX →
- ∀V2. ⦃G, KX⦄ ⊢ VX ➡ V2 → ∀T2. ⬆[O, i + 1] V2 ≡ T2 →
- ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
- ∃∃T. ⦃G, L1⦄ ⊢ T1 ➡ T & ⦃G, L2⦄ ⊢ T2 ➡ T.
-#G #L0 #i #IH #K0 #V0 #HLK0 #V1 #HV01 #T1 #HVT1
-#KX #VX #H #V2 #HV02 #T2 #HVT2 #L1 #HL01 #L2 #HL02
-lapply (drop_mono … H … HLK0) -H #H destruct
-elim (lpr_drop_conf … HLK0 … HL01) -HL01 #X1 #H1 #HLK1
-elim (lpr_inv_pair1 … H1) -H1 #K1 #W1 #HK01 #_ #H destruct
-lapply (drop_fwd_drop2 … HLK1) -W1 #HLK1
-elim (lpr_drop_conf … HLK0 … HL02) -HL02 #X2 #H2 #HLK2
-elim (lpr_inv_pair1 … H2) -H2 #K2 #W2 #HK02 #_ #H destruct
-lapply (drop_fwd_drop2 … HLK2) -W2 #HLK2
-lapply (fqup_lref … G … HLK0) -HLK0 #HLK0
-elim (IH … HLK0 … HV01 … HV02 … HK01 … HK02) -L0 -K0 -V0 #V #HV1 #HV2
-elim (lift_total V 0 (i+1)) /3 width=12 by cpr_lift, ex2_intro/
-qed-.
-
-fact cpr_conf_lpr_bind_bind:
- ∀a,I,G,L0,V0,T0. (
- ∀L,T. ⦃G, L0, ⓑ{a,I}V0.T0⦄ ⊐+ ⦃G, L, T⦄ →
- ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
- ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
- ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
- ) →
- ∀V1. ⦃G, L0⦄ ⊢ V0 ➡ V1 → ∀T1. ⦃G, L0.ⓑ{I}V0⦄ ⊢ T0 ➡ T1 →
- ∀V2. ⦃G, L0⦄ ⊢ V0 ➡ V2 → ∀T2. ⦃G, L0.ⓑ{I}V0⦄ ⊢ T0 ➡ T2 →
- ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
- ∃∃T. ⦃G, L1⦄ ⊢ ⓑ{a,I}V1.T1 ➡ T & ⦃G, L2⦄ ⊢ ⓑ{a,I}V2.T2 ➡ T.
-#a #I #G #L0 #V0 #T0 #IH #V1 #HV01 #T1 #HT01
-#V2 #HV02 #T2 #HT02 #L1 #HL01 #L2 #HL02
-elim (IH … HV01 … HV02 … HL01 … HL02) //
-elim (IH … HT01 … HT02 (L1.ⓑ{I}V1) … (L2.ⓑ{I}V2)) -IH
-/3 width=5 by lpr_pair, cpr_bind, ex2_intro/
-qed-.
-
-fact cpr_conf_lpr_bind_zeta:
- ∀G,L0,V0,T0. (
- ∀L,T. ⦃G, L0, +ⓓV0.T0⦄ ⊐+ ⦃G, L, T⦄ →
- ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
- ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
- ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
- ) →
- ∀V1. ⦃G, L0⦄ ⊢ V0 ➡ V1 → ∀T1. ⦃G, L0.ⓓV0⦄ ⊢ T0 ➡ T1 →
- ∀T2. ⦃G, L0.ⓓV0⦄ ⊢ T0 ➡ T2 → ∀X2. ⬆[O, 1] X2 ≡ T2 →
- ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
- ∃∃T. ⦃G, L1⦄ ⊢ +ⓓV1.T1 ➡ T & ⦃G, L2⦄ ⊢ X2 ➡ T.
-#G #L0 #V0 #T0 #IH #V1 #HV01 #T1 #HT01
-#T2 #HT02 #X2 #HXT2 #L1 #HL01 #L2 #HL02
-elim (IH … HT01 … HT02 (L1.ⓓV1) … (L2.ⓓV1)) -IH -HT01 -HT02 /2 width=1 by lpr_pair/ -L0 -V0 -T0 #T #HT1 #HT2
-elim (cpr_inv_lift1 … HT2 L2 … HXT2) -T2 /3 width=3 by cpr_zeta, drop_drop, ex2_intro/
-qed-.
-
-fact cpr_conf_lpr_zeta_zeta:
- ∀G,L0,V0,T0. (
- ∀L,T. ⦃G, L0, +ⓓV0.T0⦄ ⊐+ ⦃G, L, T⦄ →
- ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
- ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
- ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
- ) →
- ∀T1. ⦃G, L0.ⓓV0⦄ ⊢ T0 ➡ T1 → ∀X1. ⬆[O, 1] X1 ≡ T1 →
- ∀T2. ⦃G, L0.ⓓV0⦄ ⊢ T0 ➡ T2 → ∀X2. ⬆[O, 1] X2 ≡ T2 →
- ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
- ∃∃T. ⦃G, L1⦄ ⊢ X1 ➡ T & ⦃G, L2⦄ ⊢ X2 ➡ T.
-#G #L0 #V0 #T0 #IH #T1 #HT01 #X1 #HXT1
-#T2 #HT02 #X2 #HXT2 #L1 #HL01 #L2 #HL02
-elim (IH … HT01 … HT02 (L1.ⓓV0) … (L2.ⓓV0)) -IH -HT01 -HT02 /2 width=1 by lpr_pair/ -L0 -T0 #T #HT1 #HT2
-elim (cpr_inv_lift1 … HT1 L1 … HXT1) -T1 /2 width=2 by drop_drop/ #T1 #HT1 #HXT1
-elim (cpr_inv_lift1 … HT2 L2 … HXT2) -T2 /2 width=2 by drop_drop/ #T2 #HT2 #HXT2
-lapply (lift_inj … HT2 … HT1) -T #H destruct /2 width=3 by ex2_intro/
-qed-.
-
-fact cpr_conf_lpr_flat_flat:
- ∀I,G,L0,V0,T0. (
- ∀L,T. ⦃G, L0, ⓕ{I}V0.T0⦄ ⊐+ ⦃G, L, T⦄ →
- ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
- ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
- ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
- ) →
- ∀V1. ⦃G, L0⦄ ⊢ V0 ➡ V1 → ∀T1. ⦃G, L0⦄ ⊢ T0 ➡ T1 →
- ∀V2. ⦃G, L0⦄ ⊢ V0 ➡ V2 → ∀T2. ⦃G, L0⦄ ⊢ T0 ➡ T2 →
- ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
- ∃∃T. ⦃G, L1⦄ ⊢ ⓕ{I}V1.T1 ➡ T & ⦃G, L2⦄ ⊢ ⓕ{I}V2.T2 ➡ T.
-#I #G #L0 #V0 #T0 #IH #V1 #HV01 #T1 #HT01
-#V2 #HV02 #T2 #HT02 #L1 #HL01 #L2 #HL02
-elim (IH … HV01 … HV02 … HL01 … HL02) //
-elim (IH … HT01 … HT02 … HL01 … HL02) /3 width=5 by cpr_flat, ex2_intro/
-qed-.
-
-fact cpr_conf_lpr_flat_eps:
- ∀G,L0,V0,T0. (
- ∀L,T. ⦃G, L0, ⓝV0.T0⦄ ⊐+ ⦃G, L, T⦄ →
- ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
- ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
- ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
- ) →
- ∀V1,T1. ⦃G, L0⦄ ⊢ T0 ➡ T1 → ∀T2. ⦃G, L0⦄ ⊢ T0 ➡ T2 →
- ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
- ∃∃T. ⦃G, L1⦄ ⊢ ⓝV1.T1 ➡ T & ⦃G, L2⦄ ⊢ T2 ➡ T.
-#G #L0 #V0 #T0 #IH #V1 #T1 #HT01
-#T2 #HT02 #L1 #HL01 #L2 #HL02
-elim (IH … HT01 … HT02 … HL01 … HL02) // -L0 -V0 -T0 /3 width=3 by cpr_eps, ex2_intro/
-qed-.
-
-fact cpr_conf_lpr_eps_eps:
- ∀G,L0,V0,T0. (
- ∀L,T. ⦃G, L0, ⓝV0.T0⦄ ⊐+ ⦃G, L, T⦄ →
- ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
- ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
- ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
- ) →
- ∀T1. ⦃G, L0⦄ ⊢ T0 ➡ T1 → ∀T2. ⦃G, L0⦄ ⊢ T0 ➡ T2 →
- ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
- ∃∃T. ⦃G, L1⦄ ⊢ T1 ➡ T & ⦃G, L2⦄ ⊢ T2 ➡ T.
-#G #L0 #V0 #T0 #IH #T1 #HT01
-#T2 #HT02 #L1 #HL01 #L2 #HL02
-elim (IH … HT01 … HT02 … HL01 … HL02) // -L0 -V0 -T0 /2 width=3 by ex2_intro/
-qed-.
-
-fact cpr_conf_lpr_flat_beta:
- ∀a,G,L0,V0,W0,T0. (
- ∀L,T. ⦃G, L0, ⓐV0.ⓛ{a}W0.T0⦄ ⊐+ ⦃G, L, T⦄ →
- ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
- ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
- ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
- ) →
- ∀V1. ⦃G, L0⦄ ⊢ V0 ➡ V1 → ∀T1. ⦃G, L0⦄ ⊢ ⓛ{a}W0.T0 ➡ T1 →
- ∀V2. ⦃G, L0⦄ ⊢ V0 ➡ V2 → ∀W2. ⦃G, L0⦄ ⊢ W0 ➡ W2 → ∀T2. ⦃G, L0.ⓛW0⦄ ⊢ T0 ➡ T2 →
- ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
- ∃∃T. ⦃G, L1⦄ ⊢ ⓐV1.T1 ➡ T & ⦃G, L2⦄ ⊢ ⓓ{a}ⓝW2.V2.T2 ➡ T.
-#a #G #L0 #V0 #W0 #T0 #IH #V1 #HV01 #X #H
-#V2 #HV02 #W2 #HW02 #T2 #HT02 #L1 #HL01 #L2 #HL02
-elim (cpr_inv_abst1 … H) -H #W1 #T1 #HW01 #HT01 #H destruct
-elim (IH … HV01 … HV02 … HL01 … HL02) -HV01 -HV02 /2 width=1 by/ #V #HV1 #HV2
-elim (IH … HW01 … HW02 … HL01 … HL02) /2 width=1 by/ #W #HW1 #HW2
-elim (IH … HT01 … HT02 (L1.ⓛW1) … (L2.ⓛW2)) /2 width=1 by lpr_pair/ -L0 -V0 -W0 -T0 #T #HT1 #HT2
-lapply (lsubr_cpr_trans … HT2 (L2.ⓓⓝW2.V2) ?) -HT2 /2 width=1 by lsubr_beta/ (**) (* full auto not tried *)
-/4 width=5 by cpr_bind, cpr_flat, cpr_beta, ex2_intro/
-qed-.
-
-(* Basic-1: includes:
- pr0_cong_upsilon_refl pr0_cong_upsilon_zeta
- pr0_cong_upsilon_cong pr0_cong_upsilon_delta
-*)
-fact cpr_conf_lpr_flat_theta:
- ∀a,G,L0,V0,W0,T0. (
- ∀L,T. ⦃G, L0, ⓐV0.ⓓ{a}W0.T0⦄ ⊐+ ⦃G, L, T⦄ →
- ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
- ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
- ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
- ) →
- ∀V1. ⦃G, L0⦄ ⊢ V0 ➡ V1 → ∀T1. ⦃G, L0⦄ ⊢ ⓓ{a}W0.T0 ➡ T1 →
- ∀V2. ⦃G, L0⦄ ⊢ V0 ➡ V2 → ∀U2. ⬆[O, 1] V2 ≡ U2 →
- ∀W2. ⦃G, L0⦄ ⊢ W0 ➡ W2 → ∀T2. ⦃G, L0.ⓓW0⦄ ⊢ T0 ➡ T2 →
- ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
- ∃∃T. ⦃G, L1⦄ ⊢ ⓐV1.T1 ➡ T & ⦃G, L2⦄ ⊢ ⓓ{a}W2.ⓐU2.T2 ➡ T.
-#a #G #L0 #V0 #W0 #T0 #IH #V1 #HV01 #X #H
-#V2 #HV02 #U2 #HVU2 #W2 #HW02 #T2 #HT02 #L1 #HL01 #L2 #HL02
-elim (IH … HV01 … HV02 … HL01 … HL02) -HV01 -HV02 /2 width=1 by/ #V #HV1 #HV2
-elim (lift_total V 0 1) #U #HVU
-lapply (cpr_lift … HV2 (L2.ⓓW2) … HVU2 … HVU) -HVU2 /2 width=2 by drop_drop/ #HU2
-elim (cpr_inv_abbr1 … H) -H *
-[ #W1 #T1 #HW01 #HT01 #H destruct
- elim (IH … HW01 … HW02 … HL01 … HL02) /2 width=1 by/
- elim (IH … HT01 … HT02 (L1.ⓓW1) … (L2.ⓓW2)) /2 width=1 by lpr_pair/ -L0 -V0 -W0 -T0
- /4 width=7 by cpr_bind, cpr_flat, cpr_theta, ex2_intro/
-| #T1 #HT01 #HXT1 #H destruct
- elim (IH … HT01 … HT02 (L1.ⓓW2) … (L2.ⓓW2)) /2 width=1 by lpr_pair/ -L0 -V0 -W0 -T0 #T #HT1 #HT2
- elim (cpr_inv_lift1 … HT1 L1 … HXT1) -HXT1
- /4 width=9 by cpr_flat, cpr_zeta, drop_drop, lift_flat, ex2_intro/
-]
-qed-.
-
-fact cpr_conf_lpr_beta_beta:
- ∀a,G,L0,V0,W0,T0. (
- ∀L,T. ⦃G, L0, ⓐV0.ⓛ{a}W0.T0⦄ ⊐+ ⦃G, L, T⦄ →
- ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
- ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
- ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
- ) →
- ∀V1. ⦃G, L0⦄ ⊢ V0 ➡ V1 → ∀W1. ⦃G, L0⦄ ⊢ W0 ➡ W1 → ∀T1. ⦃G, L0.ⓛW0⦄ ⊢ T0 ➡ T1 →
- ∀V2. ⦃G, L0⦄ ⊢ V0 ➡ V2 → ∀W2. ⦃G, L0⦄ ⊢ W0 ➡ W2 → ∀T2. ⦃G, L0.ⓛW0⦄ ⊢ T0 ➡ T2 →
- ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
- ∃∃T. ⦃G, L1⦄ ⊢ ⓓ{a}ⓝW1.V1.T1 ➡ T & ⦃G, L2⦄ ⊢ ⓓ{a}ⓝW2.V2.T2 ➡ T.
-#a #G #L0 #V0 #W0 #T0 #IH #V1 #HV01 #W1 #HW01 #T1 #HT01
-#V2 #HV02 #W2 #HW02 #T2 #HT02 #L1 #HL01 #L2 #HL02
-elim (IH … HV01 … HV02 … HL01 … HL02) -HV01 -HV02 /2 width=1 by/ #V #HV1 #HV2
-elim (IH … HW01 … HW02 … HL01 … HL02) /2 width=1 by/ #W #HW1 #HW2
-elim (IH … HT01 … HT02 (L1.ⓛW1) … (L2.ⓛW2)) /2 width=1 by lpr_pair/ -L0 -V0 -W0 -T0 #T #HT1 #HT2
-lapply (lsubr_cpr_trans … HT1 (L1.ⓓⓝW1.V1) ?) -HT1 /2 width=1 by lsubr_beta/
-lapply (lsubr_cpr_trans … HT2 (L2.ⓓⓝW2.V2) ?) -HT2 /2 width=1 by lsubr_beta/
-/4 width=5 by cpr_bind, cpr_flat, ex2_intro/ (**) (* full auto not tried *)
-qed-.
-
-(* Basic_1: was: pr0_upsilon_upsilon *)
-fact cpr_conf_lpr_theta_theta:
- ∀a,G,L0,V0,W0,T0. (
- ∀L,T. ⦃G, L0, ⓐV0.ⓓ{a}W0.T0⦄ ⊐+ ⦃G, L, T⦄ →
- ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
- ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
- ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
- ) →
- ∀V1. ⦃G, L0⦄ ⊢ V0 ➡ V1 → ∀U1. ⬆[O, 1] V1 ≡ U1 →
- ∀W1. ⦃G, L0⦄ ⊢ W0 ➡ W1 → ∀T1. ⦃G, L0.ⓓW0⦄ ⊢ T0 ➡ T1 →
- ∀V2. ⦃G, L0⦄ ⊢ V0 ➡ V2 → ∀U2. ⬆[O, 1] V2 ≡ U2 →
- ∀W2. ⦃G, L0⦄ ⊢ W0 ➡ W2 → ∀T2. ⦃G, L0.ⓓW0⦄ ⊢ T0 ➡ T2 →
- ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
- ∃∃T. ⦃G, L1⦄ ⊢ ⓓ{a}W1.ⓐU1.T1 ➡ T & ⦃G, L2⦄ ⊢ ⓓ{a}W2.ⓐU2.T2 ➡ T.
-#a #G #L0 #V0 #W0 #T0 #IH #V1 #HV01 #U1 #HVU1 #W1 #HW01 #T1 #HT01
-#V2 #HV02 #U2 #HVU2 #W2 #HW02 #T2 #HT02 #L1 #HL01 #L2 #HL02
-elim (IH … HV01 … HV02 … HL01 … HL02) -HV01 -HV02 /2 width=1 by/ #V #HV1 #HV2
-elim (IH … HW01 … HW02 … HL01 … HL02) /2 width=1 by/
-elim (IH … HT01 … HT02 (L1.ⓓW1) … (L2.ⓓW2)) /2 width=1 by lpr_pair/ -L0 -V0 -W0 -T0
-elim (lift_total V 0 1) #U #HVU
-lapply (cpr_lift … HV1 (L1.ⓓW1) … HVU1 … HVU) -HVU1 /2 width=2 by drop_drop/
-lapply (cpr_lift … HV2 (L2.ⓓW2) … HVU2 … HVU) -HVU2 /2 width=2 by drop_drop/
-/4 width=7 by cpr_bind, cpr_flat, ex2_intro/ (**) (* full auto not tried *)
-qed-.
-
-theorem cpr_conf_lpr: ∀G. lpx_sn_confluent (cpr G) (cpr G).
-#G #L0 #T0 @(fqup_wf_ind_eq … G L0 T0) -G -L0 -T0 #G #L #T #IH #G0 #L0 * [| * ]
-[ #I0 #HG #HL #HT #T1 #H1 #T2 #H2 #L1 #HL01 #L2 #HL02 destruct
- elim (cpr_inv_atom1 … H1) -H1
- elim (cpr_inv_atom1 … H2) -H2
- [ #H2 #H1 destruct
- /2 width=1 by cpr_conf_lpr_atom_atom/
- | * #K0 #V0 #V2 #i2 #HLK0 #HV02 #HVT2 #H2 #H1 destruct
- /3 width=10 by cpr_conf_lpr_atom_delta/
- | #H2 * #K0 #V0 #V1 #i1 #HLK0 #HV01 #HVT1 #H1 destruct
- /4 width=10 by ex2_commute, cpr_conf_lpr_atom_delta/
- | * #X #Y #V2 #z #H #HV02 #HVT2 #H2
- * #K0 #V0 #V1 #i #HLK0 #HV01 #HVT1 #H1 destruct
- /3 width=17 by cpr_conf_lpr_delta_delta/
- ]
-| #a #I #V0 #T0 #HG #HL #HT #X1 #H1 #X2 #H2 #L1 #HL01 #L2 #HL02 destruct
- elim (cpr_inv_bind1 … H1) -H1 *
- [ #V1 #T1 #HV01 #HT01 #H1
- | #T1 #HT01 #HXT1 #H11 #H12
- ]
- elim (cpr_inv_bind1 … H2) -H2 *
- [1,3: #V2 #T2 #HV02 #HT02 #H2
- |2,4: #T2 #HT02 #HXT2 #H21 #H22
- ] destruct
- [ /3 width=10 by cpr_conf_lpr_bind_bind/
- | /4 width=11 by ex2_commute, cpr_conf_lpr_bind_zeta/
- | /3 width=11 by cpr_conf_lpr_bind_zeta/
- | /3 width=12 by cpr_conf_lpr_zeta_zeta/
- ]
-| #I #V0 #T0 #HG #HL #HT #X1 #H1 #X2 #H2 #L1 #HL01 #L2 #HL02 destruct
- elim (cpr_inv_flat1 … H1) -H1 *
- [ #V1 #T1 #HV01 #HT01 #H1
- | #HX1 #H1
- | #a1 #V1 #Y1 #W1 #Z1 #T1 #HV01 #HYW1 #HZT1 #H11 #H12 #H13
- | #a1 #V1 #U1 #Y1 #W1 #Z1 #T1 #HV01 #HVU1 #HYW1 #HZT1 #H11 #H12 #H13
- ]
- elim (cpr_inv_flat1 … H2) -H2 *
- [1,5,9,13: #V2 #T2 #HV02 #HT02 #H2
- |2,6,10,14: #HX2 #H2
- |3,7,11,15: #a2 #V2 #Y2 #W2 #Z2 #T2 #HV02 #HYW2 #HZT2 #H21 #H22 #H23
- |4,8,12,16: #a2 #V2 #U2 #Y2 #W2 #Z2 #T2 #HV02 #HVU2 #HYW2 #HZT2 #H21 #H22 #H23
- ] destruct
- [ /3 width=10 by cpr_conf_lpr_flat_flat/
- | /4 width=8 by ex2_commute, cpr_conf_lpr_flat_eps/
- | /4 width=12 by ex2_commute, cpr_conf_lpr_flat_beta/
- | /4 width=14 by ex2_commute, cpr_conf_lpr_flat_theta/
- | /3 width=8 by cpr_conf_lpr_flat_eps/
- | /3 width=7 by cpr_conf_lpr_eps_eps/
- | /3 width=12 by cpr_conf_lpr_flat_beta/
- | /3 width=13 by cpr_conf_lpr_beta_beta/
- | /3 width=14 by cpr_conf_lpr_flat_theta/
- | /3 width=17 by cpr_conf_lpr_theta_theta/
- ]
-]
-qed-.
-
-(* Basic_1: includes: pr0_confluence pr2_confluence *)
-theorem cpr_conf: ∀G,L. confluent … (cpr G L).
-/2 width=6 by cpr_conf_lpr/ qed-.
-
-(* Properties on context-sensitive parallel reduction for terms *************)
-
-lemma lpr_cpr_conf_dx: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡ T1 → ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 →
- ∃∃T. ⦃G, L1⦄ ⊢ T0 ➡ T & ⦃G, L1⦄ ⊢ T1 ➡ T.
-#G #L0 #T0 #T1 #HT01 #L1 #HL01
-elim (cpr_conf_lpr … HT01 T0 … HL01 … HL01) /2 width=3 by ex2_intro/
-qed-.
-
-lemma lpr_cpr_conf_sn: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡ T1 → ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 →
- ∃∃T. ⦃G, L1⦄ ⊢ T0 ➡ T & ⦃G, L0⦄ ⊢ T1 ➡ T.
-#G #L0 #T0 #T1 #HT01 #L1 #HL01
-elim (cpr_conf_lpr … HT01 T0 … L0 … HL01) /2 width=3 by ex2_intro/
-qed-.
-
-(* Main properties **********************************************************)
-
-theorem lpr_conf: ∀G. confluent … (lpr G).
-/3 width=6 by lpx_sn_conf, cpr_conf_lpr/
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/predsn_5.ma".
-include "basic_2/reduction/lpr.ma".
-include "basic_2/reduction/cpx.ma".
-
-(* SN EXTENDED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS ********************)
-
-definition lpx: ∀h. sd h → relation3 genv lenv lenv ≝
- λh,o,G. lpx_sn (cpx h o G).
-
-interpretation "extended parallel reduction (local environment, sn variant)"
- 'PRedSn h o G L1 L2 = (lpx h o G L1 L2).
-
-(* Basic inversion lemmas ***************************************************)
-
-lemma lpx_inv_atom1: ∀h,o,G,L2. ⦃G, ⋆⦄ ⊢ ➡[h, o] L2 → L2 = ⋆.
-/2 width=4 by lpx_sn_inv_atom1_aux/ qed-.
-
-lemma lpx_inv_pair1: ∀h,o,I,G,K1,V1,L2. ⦃G, K1.ⓑ{I}V1⦄ ⊢ ➡[h, o] L2 →
- ∃∃K2,V2. ⦃G, K1⦄ ⊢ ➡[h, o] K2 & ⦃G, K1⦄ ⊢ V1 ➡[h, o] V2 &
- L2 = K2. ⓑ{I} V2.
-/2 width=3 by lpx_sn_inv_pair1_aux/ qed-.
-
-lemma lpx_inv_atom2: ∀h,o,G,L1. ⦃G, L1⦄ ⊢ ➡[h, o] ⋆ → L1 = ⋆.
-/2 width=4 by lpx_sn_inv_atom2_aux/ qed-.
-
-lemma lpx_inv_pair2: ∀h,o,I,G,L1,K2,V2. ⦃G, L1⦄ ⊢ ➡[h, o] K2.ⓑ{I}V2 →
- ∃∃K1,V1. ⦃G, K1⦄ ⊢ ➡[h, o] K2 & ⦃G, K1⦄ ⊢ V1 ➡[h, o] V2 &
- L1 = K1. ⓑ{I} V1.
-/2 width=3 by lpx_sn_inv_pair2_aux/ qed-.
-
-lemma lpx_inv_pair: ∀h,o,I1,I2,G,L1,L2,V1,V2. ⦃G, L1.ⓑ{I1}V1⦄ ⊢ ➡[h, o] L2.ⓑ{I2}V2 →
- ∧∧ ⦃G, L1⦄ ⊢ ➡[h, o] L2 & ⦃G, L1⦄ ⊢ V1 ➡[h, o] V2 & I1 = I2.
-/2 width=1 by lpx_sn_inv_pair/ qed-.
-
-(* Basic properties *********************************************************)
-
-lemma lpx_refl: ∀h,o,G,L. ⦃G, L⦄ ⊢ ➡[h, o] L.
-/2 width=1 by lpx_sn_refl/ qed.
-
-lemma lpx_pair: ∀h,o,I,G,K1,K2,V1,V2. ⦃G, K1⦄ ⊢ ➡[h, o] K2 → ⦃G, K1⦄ ⊢ V1 ➡[h, o] V2 →
- ⦃G, K1.ⓑ{I}V1⦄ ⊢ ➡[h, o] K2.ⓑ{I}V2.
-/2 width=1 by lpx_sn_pair/ qed.
-
-lemma lpr_lpx: ∀h,o,G,L1,L2. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L1⦄ ⊢ ➡[h, o] L2.
-#h #o #G #L1 #L2 #H elim H -L1 -L2 /3 width=1 by lpx_pair, cpr_cpx/
-qed.
-
-(* Basic forward lemmas *****************************************************)
-
-lemma lpx_fwd_length: ∀h,o,G,L1,L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 → |L1| = |L2|.
-/2 width=2 by lpx_sn_fwd_length/ qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/static/aaa_lift.ma".
-include "basic_2/static/lsuba_aaa.ma".
-include "basic_2/reduction/lpx_drop.ma".
-
-(* SN EXTENDED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS ********************)
-
-(* Properties on atomic arity assignment for terms **************************)
-
-(* Note: lemma 500 *)
-lemma cpx_lpx_aaa_conf: ∀h,o,G,L1,T1,A. ⦃G, L1⦄ ⊢ T1 ⁝ A →
- ∀T2. ⦃G, L1⦄ ⊢ T1 ➡[h, o] T2 →
- ∀L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 → ⦃G, L2⦄ ⊢ T2 ⁝ A.
-#h #o #G #L1 #T1 #A #H elim H -G -L1 -T1 -A
-[ #o #L1 #s #X #H
- elim (cpx_inv_sort1 … H) -H // * //
-| #I #G #L1 #K1 #V1 #B #i #HLK1 #_ #IHV1 #X #H #L2 #HL12
- elim (cpx_inv_lref1 … H) -H
- [ #H destruct
- elim (lpx_drop_conf … HLK1 … HL12) -L1 #X #H #HLK2
- elim (lpx_inv_pair1 … H) -H
- #K2 #V2 #HK12 #HV12 #H destruct /3 width=6 by aaa_lref/
- | * #J #Y #Z #V2 #H #HV12 #HV2
- lapply (drop_mono … H … HLK1) -H #H destruct
- elim (lpx_drop_conf … HLK1 … HL12) -L1 #Z #H #HLK2
- elim (lpx_inv_pair1 … H) -H #K2 #V0 #HK12 #_ #H destruct
- /3 width=8 by aaa_lift, drop_fwd_drop2/
- ]
-| #a #G #L1 #V1 #T1 #B #A #_ #_ #IHV1 #IHT1 #X #H #L2 #HL12
- elim (cpx_inv_abbr1 … H) -H *
- [ #V2 #T2 #HV12 #HT12 #H destruct /4 width=2 by lpx_pair, aaa_abbr/
- | #T2 #HT12 #HT2 #H destruct -IHV1
- /4 width=8 by lpx_pair, aaa_inv_lift, drop_drop/
- ]
-| #a #G #L1 #V1 #T1 #B #A #_ #_ #IHV1 #IHT1 #X #H #L2 #HL12
- elim (cpx_inv_abst1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct
- /4 width=1 by lpx_pair, aaa_abst/
-| #G #L1 #V1 #T1 #B #A #_ #_ #IHV1 #IHT1 #X #H #L2 #HL12
- elim (cpx_inv_appl1 … H) -H *
- [ #V2 #T2 #HV12 #HT12 #H destruct /3 width=3 by aaa_appl/
- | #b #V2 #W1 #W2 #U1 #U2 #HV12 #HW12 #HU12 #H1 #H2 destruct
- lapply (IHV1 … HV12 … HL12) -IHV1 -HV12 #HV2
- lapply (IHT1 (ⓛ{b}W2.U2) … HL12) -IHT1 /2 width=1 by cpx_bind/ -L1 #H
- elim (aaa_inv_abst … H) -H #B0 #A0 #HW1 #HU2 #H destruct
- /5 width=6 by lsuba_aaa_trans, lsuba_beta, aaa_abbr, aaa_cast/
- | #b #V #V2 #W1 #W2 #U1 #U2 #HV1 #HV2 #HW12 #HU12 #H1 #H2 destruct
- lapply (aaa_lift G L2 … B … (L2.ⓓW2) … HV2) -HV2 /2 width=2 by drop_drop/ #HV2
- lapply (IHT1 (ⓓ{b}W2.U2) … HL12) -IHT1 /2 width=1 by cpx_bind/ -L1 #H
- elim (aaa_inv_abbr … H) -H /3 width=3 by aaa_abbr, aaa_appl/
- ]
-| #G #L1 #V1 #T1 #A #_ #_ #IHV1 #IHT1 #X #H #L2 #HL12
- elim (cpx_inv_cast1 … H) -H
- [ * #V2 #T2 #HV12 #HT12 #H destruct /3 width=1 by aaa_cast/
- | -IHV1 /2 width=1 by/
- | -IHT1 /2 width=1 by/
- ]
-]
-qed-.
-
-lemma cpx_aaa_conf: ∀h,o,G,L,T1,A. ⦃G, L⦄ ⊢ T1 ⁝ A → ∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → ⦃G, L⦄ ⊢ T2 ⁝ A.
-/2 width=7 by cpx_lpx_aaa_conf/ qed-.
-
-lemma lpx_aaa_conf: ∀h,o,G,L1,T,A. ⦃G, L1⦄ ⊢ T ⁝ A → ∀L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 → ⦃G, L2⦄ ⊢ T ⁝ A.
-/2 width=7 by cpx_lpx_aaa_conf/ qed-.
-
-lemma cpr_aaa_conf: ∀G,L,T1,A. ⦃G, L⦄ ⊢ T1 ⁝ A → ∀T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ⦃G, L⦄ ⊢ T2 ⁝ A.
-/3 width=5 by cpx_aaa_conf, cpr_cpx/ qed-.
-
-lemma lpr_aaa_conf: ∀G,L1,T,A. ⦃G, L1⦄ ⊢ T ⁝ A → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L2⦄ ⊢ T ⁝ A.
-/3 width=5 by lpx_aaa_conf, lpr_lpx/ qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/substitution/lpx_sn_drop.ma".
-include "basic_2/reduction/cpx_lift.ma".
-include "basic_2/reduction/lpx.ma".
-
-(* SN EXTENDED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS ********************)
-
-(* Properties on local environment slicing ***********************************)
-
-lemma lpx_drop_conf: ∀h,o,G. dropable_sn (lpx h o G).
-/3 width=6 by lpx_sn_deliftable_dropable, cpx_inv_lift1/ qed-.
-
-lemma drop_lpx_trans: ∀h,o,G. dedropable_sn (lpx h o G).
-/3 width=10 by lpx_sn_liftable_dedropable, cpx_lift/ qed-.
-
-lemma lpx_drop_trans_O1: ∀h,o,G. dropable_dx (lpx h o G).
-/2 width=3 by lpx_sn_dropable/ qed-.
-
-(* Properties on supclosure *************************************************)
-
-lemma fqu_lpx_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
- ∀K2. ⦃G2, L2⦄ ⊢ ➡[h, o] K2 →
- ∃∃K1,T. ⦃G1, L1⦄ ⊢ ➡[h, o] K1 & ⦃G1, L1⦄ ⊢ T1 ➡[h, o] T & ⦃G1, K1, T⦄ ⊐ ⦃G2, K2, T2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
-/3 width=5 by fqu_lref_O, fqu_pair_sn, fqu_flat_dx, lpx_pair, ex3_2_intro/
-[ #a #I #G2 #L2 #V2 #T2 #X #H elim (lpx_inv_pair1 … H) -H
- #K2 #W2 #HLK2 #HVW2 #H destruct
- /3 width=5 by cpx_pair_sn, fqu_bind_dx, ex3_2_intro/
-| #G #L1 #K1 #T1 #U1 #k #HLK1 #HTU1 #K2 #HK12
- elim (drop_lpx_trans … HLK1 … HK12) -HK12
- /3 width=7 by fqu_drop, ex3_2_intro/
-]
-qed-.
-
-lemma fquq_lpx_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ∀K2. ⦃G2, L2⦄ ⊢ ➡[h, o] K2 →
- ∃∃K1,T. ⦃G1, L1⦄ ⊢ ➡[h, o] K1 & ⦃G1, L1⦄ ⊢ T1 ➡[h, o] T & ⦃G1, K1, T⦄ ⊐⸮ ⦃G2, K2, T2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #K2 #HLK2 elim (fquq_inv_gen … H) -H
-[ #HT12 elim (fqu_lpx_trans … HT12 … HLK2) /3 width=5 by fqu_fquq, ex3_2_intro/
-| * #H1 #H2 #H3 destruct /2 width=5 by ex3_2_intro/
-]
-qed-.
-
-lemma lpx_fqu_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
- ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 →
- ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡[h, o] T & ⦃G1, K1, T⦄ ⊐ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
-/3 width=7 by fqu_pair_sn, fqu_bind_dx, fqu_flat_dx, lpx_pair, ex3_2_intro/
-[ #I #G1 #L1 #V1 #X #H elim (lpx_inv_pair2 … H) -H
- #K1 #W1 #HKL1 #HWV1 #H destruct elim (lift_total V1 0 1)
- /4 width=7 by cpx_delta, fqu_drop, drop_drop, ex3_2_intro/
-| #G #L1 #K1 #T1 #U1 #k #HLK1 #HTU1 #L0 #HL01
- elim (lpx_drop_trans_O1 … HL01 … HLK1) -L1
- /3 width=5 by fqu_drop, ex3_2_intro/
-]
-qed-.
-
-lemma lpx_fquq_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 →
- ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡[h, o] T & ⦃G1, K1, T⦄ ⊐⸮ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #K1 #HKL1 elim (fquq_inv_gen … H) -H
-[ #HT12 elim (lpx_fqu_trans … HT12 … HKL1) /3 width=5 by fqu_fquq, ex3_2_intro/
-| * #H1 #H2 #H3 destruct /2 width=5 by ex3_2_intro/
-]
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/multiple/frees_lreq.ma".
-include "basic_2/multiple/frees_lift.ma".
-include "basic_2/reduction/lpx_drop.ma".
-
-(* SN EXTENDED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS ********************)
-
-(* Properties on context-sensitive free variables ***************************)
-
-lemma lpx_cpx_frees_trans: ∀h,o,G,L1,U1,U2. ⦃G, L1⦄ ⊢ U1 ➡[h, o] U2 →
- ∀L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 →
- ∀i. L2 ⊢ i ϵ 𝐅*[0]⦃U2⦄ → L1 ⊢ i ϵ 𝐅*[0]⦃U1⦄.
-#h #o #G #L1 #U1 @(fqup_wf_ind_eq … G L1 U1) -G -L1 -U1
-#G0 #L0 #U0 #IH #G #L1 * *
-[ -IH #s #HG #HL #HU #U2 #H1 #L2 #_ #i #H2 elim (cpx_inv_sort1 … H1) -H1
- [| * #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 (lpx_drop_trans_O1 … HL12 … HLK2) -HL12 #Y #HLK1 #H
- elim (lpx_inv_pair2 … H) -H #K1 #W1 #HK12 #HW12 #H destruct
- /4 width=11 by frees_lref_be, fqup_lref/
- | * #I #K1 #W1 #W0 #HLK1 #HW10 #HW0U2
- lapply (drop_fwd_drop2 … HLK1) #H0
- elim (lpx_drop_conf … H0 … HL12) -H0 -HL12 #K2 #HK12 #HLK2
- elim (ylt_split i (j+1)) >yplus_SO2 #Hji
- [ -IH elim (frees_inv_lift_be … H2 … HLK2 … HW0U2) /2 width=1 by ylt_fwd_succ2/
- | lapply (frees_inv_lift_ge … H2 … HLK2 … HW0U2 ?) -L2 -U2 // destruct
- /4 width=11 by frees_lref_be, fqup_lref, yle_succ1_inj/
- ]
- ]
-| -IH #p #HG #HL #HU #U2 #H1 >(cpx_inv_gref1 … H1) -H1 destruct
- #L2 #_ #i #H2 elim (frees_inv_gref … H2)
-| #a #I #W1 #U1 #HG #HL #HU #X #HX #L2 #HL12 #i #Hi destruct
- elim (cpx_inv_bind1 … HX) -HX *
- [ #W2 #U2 #HW12 #HU12 #H destruct
- elim (frees_inv_bind_O … Hi) -Hi
- /4 width=7 by frees_bind_dx_O, frees_bind_sn, lpx_pair/
- | #U2 #HU12 #HXU2 #H1 #H2 destruct
- lapply (frees_lift_ge … Hi (L2.ⓓW1) (Ⓕ) … HXU2 ?)
- /4 width=7 by frees_bind_dx_O, lpx_pair, drop_drop/
- ]
-| #I #W1 #X1 #HG #HL #HU #X2 #HX2 #L2 #HL12 #i #Hi destruct
- elim (cpx_inv_flat1 … HX2) -HX2 *
- [ #W2 #U2 #HW12 #HU12 #H destruct
- elim (frees_inv_flat … Hi) -Hi /3 width=7 by frees_flat_dx, frees_flat_sn/
- | #HU12 #H destruct /3 width=7 by frees_flat_dx/
- | #HW12 #H destruct /3 width=7 by frees_flat_sn/
- | #b #W2 #V1 #V2 #U1 #U2 #HW12 #HV12 #HU12 #H1 #H2 #H3 destruct
- 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 (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/
- ]
- | #b #W2 #W0 #V1 #V2 #U1 #U2 #HW12 #HW20 #HV12 #HU12 #H1 #H2 #H3 destruct
- elim (frees_inv_bind_O … Hi) -Hi #Hi
- [ /4 width=7 by frees_flat_dx, frees_bind_sn/
- | elim (frees_inv_flat … Hi) -Hi
- [ #HW0 lapply (frees_inv_lift_ge … HW0 L2 (Ⓕ) … HW20 ?) -W0
- /3 width=7 by frees_flat_sn, drop_drop/
- | /5 width=7 by frees_bind_dx_O, frees_flat_dx, lpx_pair/
- ]
- ]
- ]
-]
-qed-.
-
-lemma cpx_frees_trans: ∀h,o,G. frees_trans (cpx h o G).
-/2 width=8 by lpx_cpx_frees_trans/ qed-.
-
-lemma lpx_frees_trans: ∀h,o,G,L1,L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 →
- ∀U,i. L2 ⊢ i ϵ 𝐅*[0]⦃U⦄ → L1 ⊢ i ϵ 𝐅*[0]⦃U⦄.
-/2 width=8 by lpx_cpx_frees_trans/ 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/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_lreq.ma".
-include "basic_2/multiple/lleq_llor.ma".
-include "basic_2/reduction/cpx_lreq.ma".
-include "basic_2/reduction/cpx_lleq.ma".
-include "basic_2/reduction/lpx_frees.ma".
-
-(* SN EXTENDED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS ********************)
-
-(* Properties on lazy equivalence for local environments ********************)
-
-(* Note: contains a proof of llpx_cpx_conf *)
-lemma lleq_lpx_trans: ∀h,o,G,L2,K2. ⦃G, L2⦄ ⊢ ➡[h, o] K2 →
- ∀L1,T,l. L1 ≡[T, l] L2 →
- ∃∃K1. ⦃G, L1⦄ ⊢ ➡[h, o] K1 & K1 ≡[T, l] K2.
-#h #o #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 … 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 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=5 by llpx_sn_llor_fwd_sn, ex2_intro/
-]
-qed-.
-
-lemma lpx_lleq_fqu_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
- ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 → K1 ≡[T1, 0] L1 →
- ∃∃K2. ⦃G1, K1, T1⦄ ⊐ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2 & K2 ≡[T2, 0] L2.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
-[ #I #G1 #L1 #V1 #X #H1 #H2 elim (lpx_inv_pair2 … H1) -H1
- #K0 #V0 #H1KL1 #_ #H destruct
- elim (lleq_inv_lref_ge_dx … H2 ? I L1 V1) -H2 //
- #K1 #H #H2KL1 lapply (drop_inv_O2 … H) -H #H destruct
- /2 width=4 by fqu_lref_O, ex3_intro/
-| * [ #a ] #I #G1 #L1 #V1 #T1 #K1 #HLK1 #H
- [ elim (lleq_inv_bind … H)
- | elim (lleq_inv_flat … H)
- ] -H /2 width=4 by fqu_pair_sn, ex3_intro/
-| #a #I #G1 #L1 #V1 #T1 #K1 #HLK1 #H elim (lleq_inv_bind_O … H) -H
- /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 #k #HL1 #HTU1 #K1 #H1KL1 #H2KL1
- elim (drop_O1_le (Ⓕ) (k+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
- /3 width=4 by fqu_drop, ex3_intro/
- | lapply (drop_fwd_length_le2 … HL1) -L -T1 -o
- lapply (lleq_fwd_length … H2KL1) //
- ]
-]
-qed-.
-
-lemma lpx_lleq_fquq_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 → K1 ≡[T1, 0] L1 →
- ∃∃K2. ⦃G1, K1, T1⦄ ⊐⸮ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2 & K2 ≡[T2, 0] L2.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #K1 #H1KL1 #H2KL1
-elim (fquq_inv_gen … H) -H
-[ #H elim (lpx_lleq_fqu_trans … H … H1KL1 H2KL1) -L1
- /3 width=4 by fqu_fquq, ex3_intro/
-| * #HG #HL #HT destruct /2 width=4 by ex3_intro/
-]
-qed-.
-
-lemma lpx_lleq_fqup_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
- ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 → K1 ≡[T1, 0] L1 →
- ∃∃K2. ⦃G1, K1, T1⦄ ⊐+ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2 & K2 ≡[T2, 0] L2.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
-[ #G2 #L2 #T2 #H #K1 #H1KL1 #H2KL1 elim (lpx_lleq_fqu_trans … H … H1KL1 H2KL1) -L1
- /3 width=4 by fqu_fqup, ex3_intro/
-| #G #G2 #L #L2 #T #T2 #_ #HT2 #IHT1 #K1 #H1KL1 #H2KL1 elim (IHT1 … H2KL1) // -L1
- #K #HT1 #H1KL #H2KL elim (lpx_lleq_fqu_trans … HT2 … H1KL H2KL) -L
- /3 width=5 by fqup_strap1, ex3_intro/
-]
-qed-.
-
-lemma lpx_lleq_fqus_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
- ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 → K1 ≡[T1, 0] L1 →
- ∃∃K2. ⦃G1, K1, T1⦄ ⊐* ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2 & K2 ≡[T2, 0] L2.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #K1 #H1KL1 #H2KL1
-elim (fqus_inv_gen … H) -H
-[ #H elim (lpx_lleq_fqup_trans … H … H1KL1 H2KL1) -L1
- /3 width=4 by fqup_fqus, ex3_intro/
-| * #HG #HL #HT destruct /2 width=4 by ex3_intro/
-]
-qed-.
-
-fact lreq_lpx_trans_lleq_aux: ∀h,o,G,L1,L0,l,k. L1 ⩬[l, k] L0 → k = ∞ →
- ∀L2. ⦃G, L0⦄ ⊢ ➡[h, o] L2 →
- ∃∃L. L ⩬[l, k] L2 & ⦃G, L1⦄ ⊢ ➡[h, o] L &
- (∀T. L0 ≡[T, l] L2 ↔ L1 ≡[T, l] L).
-#h #o #G #L1 #L0 #l #k #H elim H -L1 -L0 -l -k
-[ #l #k #_ #L2 #H >(lpx_inv_atom1 … H) -H
- /3 width=5 by ex3_intro, conj/
-| #I1 #I0 #L1 #L0 #V1 #V0 #_ #_ #Hm destruct
-| #I #L1 #L0 #V1 #k #HL10 #IHL10 #Hm #Y #H
- elim (lpx_inv_pair1 … H) -H #L2 #V2 #HL02 #HV02 #H destruct
- 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, lreq_cpx_trans, lreq_pair/
- #T elim (IH T) #HL0dx #HL0sn
- @conj #H @(lleq_lreq_repl … H) -H /3 width=1 by lreq_sym, lreq_pair_O_Y/
-| #I1 #I0 #L1 #L0 #V1 #V0 #l #k #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, lreq_succ/
- #T elim (IH T) #HL0dx #HL0sn
- @conj #H @(lleq_lreq_repl … H) -H /3 width=1 by lreq_sym, lreq_succ/
-]
-qed-.
-
-lemma lreq_lpx_trans_lleq: ∀h,o,G,L1,L0,l. L1 ⩬[l, ∞] L0 →
- ∀L2. ⦃G, L0⦄ ⊢ ➡[h, o] L2 →
- ∃∃L. L ⩬[l, ∞] L2 & ⦃G, L1⦄ ⊢ ➡[h, o] L &
- (∀T. L0 ≡[T, l] L2 ↔ L1 ≡[T, l] L).
-/2 width=1 by lreq_lpx_trans_lleq_aux/ qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/predeval_4.ma".
+include "basic_2/computation/cprs.ma".
+include "basic_2/computation/csx.ma".
+
+(* EVALUATION FOR CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS *************)
+
+definition cpre: relation4 genv lenv term term ≝
+ λG,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡* T2 ∧ ⦃G, L⦄ ⊢ ➡ 𝐍⦃T2⦄.
+
+interpretation "evaluation for context-sensitive parallel reduction (term)"
+ 'PRedEval G L T1 T2 = (cpre G L T1 T2).
+
+(* Basic properties *********************************************************)
+
+(* Basic_1: was just: nf2_sn3 *)
+lemma csx_cpre: ∀h,o,G,L,T1. ⦃G, L⦄ ⊢ ⬊*[h, o] T1 → ∃T2. ⦃G, L⦄ ⊢ T1 ➡* 𝐍⦃T2⦄.
+#h #o #G #L #T1 #H @(csx_ind … H) -T1
+#T1 #_ #IHT1 elim (cnr_dec G L T1) /3 width=3 by ex_intro, conj/
+* #T #H1T1 #H2T1 elim (IHT1 … H2T1) -IHT1 -H2T1 /2 width=2 by cpr_cpx/
+#T2 * /4 width=3 by cprs_strap2, ex_intro, conj/
+qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/computation/cprs_cprs.ma".
+include "basic_2/computation/cpre.ma".
+
+(* EVALUATION FOR CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS *************)
+
+(* Main properties *********************************************************)
+
+(* Basic_1: was: nf2_pr3_confluence *)
+theorem cpre_mono: ∀G,L,T,T1. ⦃G, L⦄ ⊢ T ➡* 𝐍⦃T1⦄ → ∀T2. ⦃G, L⦄ ⊢ T ➡* 𝐍⦃T2⦄ → T1 = T2.
+#G #L #T #T1 * #H1T1 #H2T1 #T2 * #H1T2 #H2T2
+elim (cprs_conf … H1T1 … H1T2) -T #T #HT1
+>(cprs_inv_cnr1 … HT1 H2T1) -T1 #HT2
+>(cprs_inv_cnr1 … HT2 H2T2) -T2 //
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/predstar_4.ma".
+include "basic_2/reduction/cnr.ma".
+
+(* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON TERMS **************************)
+
+(* Basic_1: includes: pr1_pr0 *)
+definition cprs: relation4 genv lenv term term ≝
+ λG. LTC … (cpr G).
+
+interpretation "context-sensitive parallel computation (term)"
+ 'PRedStar G L T1 T2 = (cprs G L T1 T2).
+
+(* Basic eliminators ********************************************************)
+
+lemma cprs_ind: ∀G,L,T1. ∀R:predicate term. R T1 →
+ (∀T,T2. ⦃G, L⦄ ⊢ T1 ➡* T → ⦃G, L⦄ ⊢ T ➡ T2 → R T → R T2) →
+ ∀T2. ⦃G, L⦄ ⊢ T1 ➡* T2 → R T2.
+#G #L #T1 #R #HT1 #IHT1 #T2 #HT12
+@(TC_star_ind … HT1 IHT1 … HT12) //
+qed-.
+
+lemma cprs_ind_dx: ∀G,L,T2. ∀R:predicate term. R T2 →
+ (∀T1,T. ⦃G, L⦄ ⊢ T1 ➡ T → ⦃G, L⦄ ⊢ T ➡* T2 → R T → R T1) →
+ ∀T1. ⦃G, L⦄ ⊢ T1 ➡* T2 → R T1.
+#G #L #T2 #R #HT2 #IHT2 #T1 #HT12
+@(TC_star_ind_dx … HT2 IHT2 … HT12) //
+qed-.
+
+(* Basic properties *********************************************************)
+
+(* Basic_1: was: pr3_pr2 *)
+lemma cpr_cprs: ∀G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ⦃G, L⦄ ⊢ T1 ➡* T2.
+/2 width=1 by inj/ qed.
+
+(* Basic_1: was: pr3_refl *)
+lemma cprs_refl: ∀G,L,T. ⦃G, L⦄ ⊢ T ➡* T.
+/2 width=1 by cpr_cprs/ qed.
+
+lemma cprs_strap1: ∀G,L,T1,T,T2.
+ ⦃G, L⦄ ⊢ T1 ➡* T → ⦃G, L⦄ ⊢ T ➡ T2 → ⦃G, L⦄ ⊢ T1 ➡* T2.
+normalize /2 width=3 by step/ qed-.
+
+(* Basic_1: was: pr3_step *)
+lemma cprs_strap2: ∀G,L,T1,T,T2.
+ ⦃G, L⦄ ⊢ T1 ➡ T → ⦃G, L⦄ ⊢ T ➡* T2 → ⦃G, L⦄ ⊢ T1 ➡* T2.
+normalize /2 width=3 by TC_strap/ qed-.
+
+lemma lsubr_cprs_trans: ∀G. lsub_trans … (cprs G) lsubr.
+/3 width=5 by lsubr_cpr_trans, LTC_lsub_trans/
+qed-.
+
+(* Basic_1: was: pr3_pr1 *)
+lemma tprs_cprs: ∀G,L,T1,T2. ⦃G, ⋆⦄ ⊢ T1 ➡* T2 → ⦃G, L⦄ ⊢ T1 ➡* T2.
+/2 width=3 by lsubr_cprs_trans/ qed.
+
+lemma cprs_bind_dx: ∀G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡ V2 → ∀I,T1,T2. ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ➡* T2 →
+ ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1. T1 ➡* ⓑ{a,I}V2. T2.
+#G #L #V1 #V2 #HV12 #I #T1 #T2 #HT12 #a @(cprs_ind_dx … HT12) -T1
+/3 width=3 by cprs_strap2, cpr_cprs, cpr_pair_sn, cpr_bind/ qed.
+
+(* Basic_1: was only: pr3_thin_dx *)
+lemma cprs_flat_dx: ∀I,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡ V2 → ∀T1,T2. ⦃G, L⦄ ⊢ T1 ➡* T2 →
+ ⦃G, L⦄ ⊢ ⓕ{I} V1. T1 ➡* ⓕ{I} V2. T2.
+#I #G #L #V1 #V2 #HV12 #T1 #T2 #HT12 @(cprs_ind … HT12) -T2
+/3 width=5 by cprs_strap1, cpr_flat, cpr_cprs, cpr_pair_sn/
+qed.
+
+lemma cprs_flat_sn: ∀I,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ∀V1,V2. ⦃G, L⦄ ⊢ V1 ➡* V2 →
+ ⦃G, L⦄ ⊢ ⓕ{I} V1. T1 ➡* ⓕ{I} V2. T2.
+#I #G #L #T1 #T2 #HT12 #V1 #V2 #H @(cprs_ind … H) -V2
+/3 width=3 by cprs_strap1, cpr_cprs, cpr_pair_sn, cpr_flat/
+qed.
+
+lemma cprs_zeta: ∀G,L,V,T1,T,T2. ⬆[0, 1] T2 ≡ T →
+ ⦃G, L.ⓓV⦄ ⊢ T1 ➡* T → ⦃G, L⦄ ⊢ +ⓓV.T1 ➡* T2.
+#G #L #V #T1 #T #T2 #HT2 #H @(cprs_ind_dx … H) -T1
+/3 width=3 by cprs_strap2, cpr_cprs, cpr_bind, cpr_zeta/
+qed.
+
+lemma cprs_eps: ∀G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡* T2 → ∀V. ⦃G, L⦄ ⊢ ⓝV.T1 ➡* T2.
+#G #L #T1 #T2 #H @(cprs_ind … H) -T2
+/3 width=3 by cprs_strap1, cpr_cprs, cpr_eps/
+qed.
+
+lemma cprs_beta_dx: ∀a,G,L,V1,V2,W1,W2,T1,T2.
+ ⦃G, L⦄ ⊢ V1 ➡ V2 → ⦃G, L⦄ ⊢ W1 ➡ W2 → ⦃G, L.ⓛW1⦄ ⊢ T1 ➡* T2 →
+ ⦃G, L⦄ ⊢ ⓐV1.ⓛ{a}W1.T1 ➡* ⓓ{a}ⓝW2.V2.T2.
+#a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #HV12 #HW12 * -T2
+/4 width=7 by cprs_strap1, cpr_cprs, cprs_bind_dx, cprs_flat_dx, cpr_beta/
+qed.
+
+lemma cprs_theta_dx: ∀a,G,L,V1,V,V2,W1,W2,T1,T2.
+ ⦃G, L⦄ ⊢ V1 ➡ V → ⬆[0, 1] V ≡ V2 → ⦃G, L.ⓓW1⦄ ⊢ T1 ➡* T2 →
+ ⦃G, L⦄ ⊢ W1 ➡ W2 → ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}W1.T1 ➡* ⓓ{a}W2.ⓐV2.T2.
+#a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV1 #HV2 * -T2
+/4 width=9 by cprs_strap1, cpr_cprs, cprs_bind_dx, cprs_flat_dx, cpr_theta/
+qed.
+
+(* Basic inversion lemmas ***************************************************)
+
+(* Basic_1: was: pr3_gen_sort *)
+lemma cprs_inv_sort1: ∀G,L,U2,s. ⦃G, L⦄ ⊢ ⋆s ➡* U2 → U2 = ⋆s.
+#G #L #U2 #s #H @(cprs_ind … H) -U2 //
+#U2 #U #_ #HU2 #IHU2 destruct
+>(cpr_inv_sort1 … HU2) -HU2 //
+qed-.
+
+(* Basic_1: was: pr3_gen_cast *)
+lemma cprs_inv_cast1: ∀G,L,W1,T1,U2. ⦃G, L⦄ ⊢ ⓝW1.T1 ➡* U2 → ⦃G, L⦄ ⊢ T1 ➡* U2 ∨
+ ∃∃W2,T2. ⦃G, L⦄ ⊢ W1 ➡* W2 & ⦃G, L⦄ ⊢ T1 ➡* T2 & U2 = ⓝW2.T2.
+#G #L #W1 #T1 #U2 #H @(cprs_ind … H) -U2 /3 width=5 by ex3_2_intro, or_intror/
+#U2 #U #_ #HU2 * /3 width=3 by cprs_strap1, or_introl/ *
+#W #T #HW1 #HT1 #H destruct
+elim (cpr_inv_cast1 … HU2) -HU2 /3 width=3 by cprs_strap1, or_introl/ *
+#W2 #T2 #HW2 #HT2 #H destruct /4 width=5 by cprs_strap1, ex3_2_intro, or_intror/
+qed-.
+
+(* Basic_1: was: nf2_pr3_unfold *)
+lemma cprs_inv_cnr1: ∀G,L,T,U. ⦃G, L⦄ ⊢ T ➡* U → ⦃G, L⦄ ⊢ ➡ 𝐍⦃T⦄ → T = U.
+#G #L #T #U #H @(cprs_ind_dx … H) -T //
+#T0 #T #H1T0 #_ #IHT #H2T0
+lapply (H2T0 … H1T0) -H1T0 #H destruct /2 width=1 by/
+qed-.
+
+(* Basic_1: removed theorems 13:
+ pr1_head_1 pr1_head_2 pr1_comp
+ clear_pr3_trans pr3_cflat pr3_gen_bind
+ pr3_head_1 pr3_head_2 pr3_head_21 pr3_head_12
+ pr3_iso_appl_bind pr3_iso_appls_appl_bind pr3_iso_appls_bind
+*)
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/reduction/lpr_lpr.ma".
+include "basic_2/computation/cprs_lift.ma".
+
+(* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON TERMS **************************)
+
+(* Main properties **********************************************************)
+
+(* Basic_1: was: pr3_t *)
+(* Basic_1: includes: pr1_t *)
+theorem cprs_trans: ∀G,L. Transitive … (cprs G L).
+normalize /2 width=3 by trans_TC/ qed-.
+
+(* Basic_1: was: pr3_confluence *)
+(* Basic_1: includes: pr1_confluence *)
+theorem cprs_conf: ∀G,L. confluent2 … (cprs G L) (cprs G L).
+normalize /3 width=3 by cpr_conf, TC_confluent2/ qed-.
+
+theorem cprs_bind: ∀a,I,G,L,V1,V2,T1,T2. ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ➡* T2 → ⦃G, L⦄ ⊢ V1 ➡* V2 →
+ ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡* ⓑ{a,I}V2.T2.
+#a #I #G #L #V1 #V2 #T1 #T2 #HT12 #H @(cprs_ind … H) -V2
+/3 width=5 by cprs_trans, cprs_bind_dx/
+qed.
+
+(* Basic_1: was: pr3_flat *)
+theorem cprs_flat: ∀I,G,L,V1,V2,T1,T2. ⦃G, L⦄ ⊢ T1 ➡* T2 → ⦃G, L⦄ ⊢ V1 ➡* V2 →
+ ⦃G, L⦄ ⊢ ⓕ{I}V1.T1 ➡* ⓕ{I}V2.T2.
+#I #G #L #V1 #V2 #T1 #T2 #HT12 #H @(cprs_ind … H) -V2
+/3 width=3 by cprs_flat_dx, cprs_strap1, cpr_pair_sn/
+qed.
+
+theorem cprs_beta_rc: ∀a,G,L,V1,V2,W1,W2,T1,T2.
+ ⦃G, L⦄ ⊢ V1 ➡ V2 → ⦃G, L.ⓛW1⦄ ⊢ T1 ➡* T2 → ⦃G, L⦄ ⊢ W1 ➡* W2 →
+ ⦃G, L⦄ ⊢ ⓐV1.ⓛ{a}W1.T1 ➡* ⓓ{a}ⓝW2.V2.T2.
+#a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #HV12 #HT12 #H @(cprs_ind … H) -W2 /2 width=1 by cprs_beta_dx/
+#W #W2 #_ #HW2 #IHW1 (**) (* fulla uto too slow 14s *)
+@(cprs_trans … IHW1) -IHW1 /3 width=1 by cprs_flat_dx, cprs_bind/
+qed.
+
+theorem cprs_beta: ∀a,G,L,V1,V2,W1,W2,T1,T2.
+ ⦃G, L.ⓛW1⦄ ⊢ T1 ➡* T2 → ⦃G, L⦄ ⊢ W1 ➡* W2 → ⦃G, L⦄ ⊢ V1 ➡* V2 →
+ ⦃G, L⦄ ⊢ ⓐV1.ⓛ{a}W1.T1 ➡* ⓓ{a}ⓝW2.V2.T2.
+#a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #HT12 #HW12 #H @(cprs_ind … H) -V2 /2 width=1 by cprs_beta_rc/
+#V #V2 #_ #HV2 #IHV1
+@(cprs_trans … IHV1) -IHV1 /3 width=1 by cprs_flat_sn, cprs_bind/
+qed.
+
+theorem cprs_theta_rc: ∀a,G,L,V1,V,V2,W1,W2,T1,T2.
+ ⦃G, L⦄ ⊢ V1 ➡ V → ⬆[0, 1] V ≡ V2 → ⦃G, L.ⓓW1⦄ ⊢ T1 ➡* T2 →
+ ⦃G, L⦄ ⊢ W1 ➡* W2 → ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}W1.T1 ➡* ⓓ{a}W2.ⓐV2.T2.
+#a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV1 #HV2 #HT12 #H @(cprs_ind … H) -W2
+/3 width=5 by cprs_trans, cprs_theta_dx, cprs_bind_dx/
+qed.
+
+theorem cprs_theta: ∀a,G,L,V1,V,V2,W1,W2,T1,T2.
+ ⬆[0, 1] V ≡ V2 → ⦃G, L⦄ ⊢ W1 ➡* W2 → ⦃G, L.ⓓW1⦄ ⊢ T1 ➡* T2 →
+ ⦃G, L⦄ ⊢ V1 ➡* V → ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}W1.T1 ➡* ⓓ{a}W2.ⓐV2.T2.
+#a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV2 #HW12 #HT12 #H @(cprs_ind_dx … H) -V1
+/3 width=3 by cprs_trans, cprs_theta_rc, cprs_flat_dx/
+qed.
+
+(* Advanced inversion lemmas ************************************************)
+
+(* Basic_1: was pr3_gen_appl *)
+lemma cprs_inv_appl1: ∀G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓐV1.T1 ➡* U2 →
+ ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡* V2 & ⦃G, L⦄ ⊢ T1 ➡* T2 &
+ U2 = ⓐV2. T2
+ | ∃∃a,W,T. ⦃G, L⦄ ⊢ T1 ➡* ⓛ{a}W.T &
+ ⦃G, L⦄ ⊢ ⓓ{a}ⓝW.V1.T ➡* U2
+ | ∃∃a,V0,V2,V,T. ⦃G, L⦄ ⊢ V1 ➡* V0 & ⬆[0,1] V0 ≡ V2 &
+ ⦃G, L⦄ ⊢ T1 ➡* ⓓ{a}V.T &
+ ⦃G, L⦄ ⊢ ⓓ{a}V.ⓐV2.T ➡* U2.
+#G #L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /3 width=5 by or3_intro0, ex3_2_intro/
+#U #U2 #_ #HU2 * *
+[ #V0 #T0 #HV10 #HT10 #H destruct
+ elim (cpr_inv_appl1 … HU2) -HU2 *
+ [ #V2 #T2 #HV02 #HT02 #H destruct /4 width=5 by cprs_strap1, or3_intro0, ex3_2_intro/
+ | #a #V2 #W #W2 #T #T2 #HV02 #HW2 #HT2 #H1 #H2 destruct
+ lapply (cprs_strap1 … HV10 … HV02) -V0 #HV12
+ lapply (lsubr_cpr_trans … HT2 (L.ⓓⓝW.V1) ?) -HT2
+ /5 width=5 by cprs_bind, cprs_flat_dx, cpr_cprs, lsubr_beta, ex2_3_intro, or3_intro1/
+ | #a #V #V2 #W0 #W2 #T #T2 #HV0 #HV2 #HW02 #HT2 #H1 #H2 destruct
+ /5 width=10 by cprs_flat_sn, cprs_bind_dx, cprs_strap1, ex4_5_intro, or3_intro2/
+ ]
+| /4 width=9 by cprs_strap1, or3_intro1, ex2_3_intro/
+| /4 width=11 by cprs_strap1, or3_intro2, ex4_5_intro/
+]
+qed-.
+
+(* Properties concerning sn parallel reduction on local environments ********)
+
+(* Basic_1: was just: pr3_pr2_pr2_t *)
+(* Basic_1: includes: pr3_pr0_pr2_t *)
+lemma lpr_cpr_trans: ∀G. c_r_transitive … (cpr G) (λ_. lpr G).
+#G #L2 #T1 #T2 #HT12 elim HT12 -G -L2 -T1 -T2
+[ /2 width=3 by/
+| #G #L2 #K2 #V0 #V2 #W2 #i #HLK2 #_ #HVW2 #IHV02 #L1 #HL12
+ elim (lpr_drop_trans_O1 … HL12 … HLK2) -L2 #X #HLK1 #H
+ elim (lpr_inv_pair2 … H) -H #K1 #V1 #HK12 #HV10 #H destruct
+ /4 width=6 by cprs_strap2, cprs_delta/
+|3,7: /4 width=1 by lpr_pair, cprs_bind, cprs_beta/
+|4,6: /3 width=1 by cprs_flat, cprs_eps/
+|5,8: /4 width=3 by lpr_pair, cprs_zeta, cprs_theta, cprs_strap1/
+]
+qed-.
+
+lemma cpr_bind2: ∀G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡ V2 → ∀I,T1,T2. ⦃G, L.ⓑ{I}V2⦄ ⊢ T1 ➡ T2 →
+ ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡* ⓑ{a,I}V2.T2.
+/4 width=5 by lpr_cpr_trans, cprs_bind_dx, lpr_pair/ qed.
+
+(* Advanced properties ******************************************************)
+
+(* Basic_1: was only: pr3_pr2_pr3_t pr3_wcpr0_t *)
+lemma lpr_cprs_trans: ∀G. c_rs_transitive … (cpr G) (λ_. lpr G).
+#G @c_r_trans_LTC1 /2 width=3 by lpr_cpr_trans/ (**) (* full auto fails *)
+qed-.
+
+(* Basic_1: was: pr3_strip *)
+(* Basic_1: includes: pr1_strip *)
+lemma cprs_strip: ∀G,L. confluent2 … (cprs G L) (cpr G L).
+normalize /4 width=3 by cpr_conf, TC_strip1/ qed-.
+
+lemma cprs_lpr_conf_dx: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡* T1 → ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 →
+ ∃∃T. ⦃G, L1⦄ ⊢ T1 ➡* T & ⦃G, L1⦄ ⊢ T0 ➡* T.
+#G #L0 #T0 #T1 #H @(cprs_ind … H) -T1 /2 width=3 by ex2_intro/
+#T #T1 #_ #HT1 #IHT0 #L1 #HL01 elim (IHT0 … HL01)
+#T2 #HT2 #HT02 elim (lpr_cpr_conf_dx … HT1 … HL01) -L0
+#T3 #HT3 #HT13 elim (cprs_strip … HT2 … HT3) -T
+/4 width=5 by cprs_strap2, cprs_strap1, ex2_intro/
+qed-.
+
+lemma cprs_lpr_conf_sn: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡* T1 →
+ ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 →
+ ∃∃T. ⦃G, L0⦄ ⊢ T1 ➡* T & ⦃G, L1⦄ ⊢ T0 ➡* T.
+#G #L0 #T0 #T1 #HT01 #L1 #HL01 elim (cprs_lpr_conf_dx … HT01 … HL01) -HT01
+/3 width=3 by lpr_cprs_trans, ex2_intro/
+qed-.
+
+lemma cprs_bind2_dx: ∀G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡ V2 →
+ ∀I,T1,T2. ⦃G, L.ⓑ{I}V2⦄ ⊢ T1 ➡* T2 →
+ ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡* ⓑ{a,I}V2.T2.
+/4 width=5 by lpr_cprs_trans, cprs_bind_dx, lpr_pair/ 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/cpr_lift.ma".
+include "basic_2/computation/cprs.ma".
+
+(* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON TERMS **************************)
+
+(* Advanced properties ******************************************************)
+
+(* Note: apparently this was missing in basic_1 *)
+lemma cprs_delta: ∀G,L,K,V,V2,i.
+ ⬇[i] L ≡ K.ⓓV → ⦃G, K⦄ ⊢ V ➡* V2 →
+ ∀W2. ⬆[0, i + 1] V2 ≡ W2 → ⦃G, L⦄ ⊢ #i ➡* W2.
+#G #L #K #V #V2 #i #HLK #H elim H -V2 [ /3 width=6 by cpr_cprs, cpr_delta/ ]
+#V1 #V2 #_ #HV12 #IHV1 #W2 #HVW2
+lapply (drop_fwd_drop2 … HLK) -HLK #HLK
+elim (lift_total V1 0 (i+1)) /4 width=12 by cpr_lift, cprs_strap1/
+qed.
+
+(* Advanced inversion lemmas ************************************************)
+
+(* Basic_1: was: pr3_gen_lref *)
+lemma cprs_inv_lref1: ∀G,L,T2,i. ⦃G, L⦄ ⊢ #i ➡* T2 →
+ T2 = #i ∨
+ ∃∃K,V1,T1. ⬇[i] L ≡ K.ⓓV1 & ⦃G, K⦄ ⊢ V1 ➡* T1 &
+ ⬆[0, i + 1] T1 ≡ T2.
+#G #L #T2 #i #H @(cprs_ind … H) -T2 /2 width=1 by or_introl/
+#T #T2 #_ #HT2 *
+[ #H destruct
+ elim (cpr_inv_lref1 … HT2) -HT2 /2 width=1 by or_introl/
+ * /4 width=6 by cpr_cprs, ex3_3_intro, or_intror/
+| * #K #V1 #T1 #HLK #HVT1 #HT1
+ lapply (drop_fwd_drop2 … HLK) #H0LK
+ elim (cpr_inv_lift1 … HT2 … H0LK … HT1) -H0LK -T
+ /4 width=6 by cprs_strap1, ex3_3_intro, or_intror/
+]
+qed-.
+
+(* Relocation properties ****************************************************)
+
+(* Basic_1: was: pr3_lift *)
+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. d_deliftable_sn (cprs G).
+/3 width=6 by d_deliftable_sn_LTC, cpr_inv_lift1/
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/predeval_6.ma".
+include "basic_2/computation/cpxs.ma".
+include "basic_2/computation/csx.ma".
+
+(* EVALUATION FOR CONTEXT-SENSITIVE EXTENDED PARALLEL REDUCTION ON TERMS ****)
+
+definition cpxe: ∀h. sd h → relation4 genv lenv term term ≝
+ λh,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 ∧ ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T2⦄.
+
+interpretation "evaluation for context-sensitive extended parallel reduction (term)"
+ 'PRedEval h o G L T1 T2 = (cpxe h o G L T1 T2).
+
+(* Basic properties *********************************************************)
+
+lemma csx_cpxe: ∀h,o,G,L,T1. ⦃G, L⦄ ⊢ ⬊*[h, o] T1 → ∃T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] 𝐍⦃T2⦄.
+#h #o #G #L #T1 #H @(csx_ind … H) -T1
+#T1 #_ #IHT1 elim (cnx_dec h o G L T1) /3 width=3 by ex_intro, conj/
+* #T #H1T1 #H2T1 elim (IHT1 … H1T1 H2T1) -IHT1 -H2T1
+#T2 * /4 width=3 by cpxs_strap2, ex_intro, conj/
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/predstar_6.ma".
+include "basic_2/reduction/cnx.ma".
+include "basic_2/computation/cprs.ma".
+
+(* CONTEXT-SENSITIVE EXTENDED PARALLEL COMPUTATION ON TERMS *****************)
+
+definition cpxs: ∀h. sd h → relation4 genv lenv term term ≝
+ λh,o,G. LTC … (cpx h o G).
+
+interpretation "extended context-sensitive parallel computation (term)"
+ 'PRedStar h o G L T1 T2 = (cpxs h o G L T1 T2).
+
+(* Basic eliminators ********************************************************)
+
+lemma cpxs_ind: ∀h,o,G,L,T1. ∀R:predicate term. R T1 →
+ (∀T,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T → ⦃G, L⦄ ⊢ T ➡[h, o] T2 → R T → R T2) →
+ ∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → R T2.
+#h #o #L #G #T1 #R #HT1 #IHT1 #T2 #HT12
+@(TC_star_ind … HT1 IHT1 … HT12) //
+qed-.
+
+lemma cpxs_ind_dx: ∀h,o,G,L,T2. ∀R:predicate term. R T2 →
+ (∀T1,T. ⦃G, L⦄ ⊢ T1 ➡[h, o] T → ⦃G, L⦄ ⊢ T ➡*[h, o] T2 → R T → R T1) →
+ ∀T1. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → R T1.
+#h #o #G #L #T2 #R #HT2 #IHT2 #T1 #HT12
+@(TC_star_ind_dx … HT2 IHT2 … HT12) //
+qed-.
+
+(* Basic properties *********************************************************)
+
+lemma cpxs_refl: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ T ➡*[h, o] T.
+/2 width=1 by inj/ qed.
+
+lemma cpx_cpxs: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2.
+/2 width=1 by inj/ qed.
+
+lemma cpxs_strap1: ∀h,o,G,L,T1,T. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T →
+ ∀T2. ⦃G, L⦄ ⊢ T ➡[h, o] T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2.
+normalize /2 width=3 by step/ qed.
+
+lemma cpxs_strap2: ∀h,o,G,L,T1,T. ⦃G, L⦄ ⊢ T1 ➡[h, o] T →
+ ∀T2. ⦃G, L⦄ ⊢ T ➡*[h, o] T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2.
+normalize /2 width=3 by TC_strap/ qed.
+
+lemma lsubr_cpxs_trans: ∀h,o,G. lsub_trans … (cpxs h o G) lsubr.
+/3 width=5 by lsubr_cpx_trans, LTC_lsub_trans/
+qed-.
+
+lemma cprs_cpxs: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡* T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2.
+#h #o #G #L #T1 #T2 #H @(cprs_ind … H) -T2 /3 width=3 by cpxs_strap1, cpr_cpx/
+qed.
+
+lemma cpxs_sort: ∀h,o,G,L,s,d1. deg h o s d1 →
+ ∀d2. d2 ≤ d1 → ⦃G, L⦄ ⊢ ⋆s ➡*[h, o] ⋆((next h)^d2 s).
+#h #o #G #L #s #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) s)))
+[ /3 width=3 by lt_to_le/
+| @(cpx_st … (d1-d2-1)) <plus_minus_k_k
+ /2 width=1 by deg_iter, monotonic_le_minus_r/
+]
+qed.
+
+lemma cpxs_bind_dx: ∀h,o,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 →
+ ∀I,T1,T2. ⦃G, L. ⓑ{I}V1⦄ ⊢ T1 ➡*[h, o] T2 →
+ ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡*[h, o] ⓑ{a,I}V2.T2.
+#h #o #G #L #V1 #V2 #HV12 #I #T1 #T2 #HT12 #a @(cpxs_ind_dx … HT12) -T1
+/3 width=3 by cpxs_strap2, cpx_cpxs, cpx_pair_sn, cpx_bind/
+qed.
+
+lemma cpxs_flat_dx: ∀h,o,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 →
+ ∀T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 →
+ ∀I. ⦃G, L⦄ ⊢ ⓕ{I}V1.T1 ➡*[h, o] ⓕ{I}V2.T2.
+#h #o #G #L #V1 #V2 #HV12 #T1 #T2 #HT12 @(cpxs_ind … HT12) -T2
+/3 width=5 by cpxs_strap1, cpx_cpxs, cpx_pair_sn, cpx_flat/
+qed.
+
+lemma cpxs_flat_sn: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 →
+ ∀V1,V2. ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 →
+ ∀I. ⦃G, L⦄ ⊢ ⓕ{I}V1.T1 ➡*[h, o] ⓕ{I}V2.T2.
+#h #o #G #L #T1 #T2 #HT12 #V1 #V2 #H @(cpxs_ind … H) -V2
+/3 width=5 by cpxs_strap1, cpx_cpxs, cpx_pair_sn, cpx_flat/
+qed.
+
+lemma cpxs_pair_sn: ∀h,o,I,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 →
+ ∀T. ⦃G, L⦄ ⊢ ②{I}V1.T ➡*[h, o] ②{I}V2.T.
+#h #o #I #G #L #V1 #V2 #H @(cpxs_ind … H) -V2
+/3 width=3 by cpxs_strap1, cpx_pair_sn/
+qed.
+
+lemma cpxs_zeta: ∀h,o,G,L,V,T1,T,T2. ⬆[0, 1] T2 ≡ T →
+ ⦃G, L.ⓓV⦄ ⊢ T1 ➡*[h, o] T → ⦃G, L⦄ ⊢ +ⓓV.T1 ➡*[h, o] T2.
+#h #o #G #L #V #T1 #T #T2 #HT2 #H @(cpxs_ind_dx … H) -T1
+/3 width=3 by cpxs_strap2, cpx_cpxs, cpx_bind, cpx_zeta/
+qed.
+
+lemma cpxs_eps: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 →
+ ∀V. ⦃G, L⦄ ⊢ ⓝV.T1 ➡*[h, o] T2.
+#h #o #G #L #T1 #T2 #H @(cpxs_ind … H) -T2
+/3 width=3 by cpxs_strap1, cpx_cpxs, cpx_eps/
+qed.
+
+lemma cpxs_ct: ∀h,o,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 →
+ ∀T. ⦃G, L⦄ ⊢ ⓝV1.T ➡*[h, o] V2.
+#h #o #G #L #V1 #V2 #H @(cpxs_ind … H) -V2
+/3 width=3 by cpxs_strap1, cpx_cpxs, cpx_ct/
+qed.
+
+lemma cpxs_beta_dx: ∀h,o,a,G,L,V1,V2,W1,W2,T1,T2.
+ ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 → ⦃G, L.ⓛW1⦄ ⊢ T1 ➡*[h, o] T2 → ⦃G, L⦄ ⊢ W1 ➡[h, o] W2 →
+ ⦃G, L⦄ ⊢ ⓐV1.ⓛ{a}W1.T1 ➡*[h, o] ⓓ{a}ⓝW2.V2.T2.
+#h #o #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #HV12 * -T2
+/4 width=7 by cpx_cpxs, cpxs_strap1, cpxs_bind_dx, cpxs_flat_dx, cpx_beta/
+qed.
+
+lemma cpxs_theta_dx: ∀h,o,a,G,L,V1,V,V2,W1,W2,T1,T2.
+ ⦃G, L⦄ ⊢ V1 ➡[h, o] V → ⬆[0, 1] V ≡ V2 → ⦃G, L.ⓓW1⦄ ⊢ T1 ➡*[h, o] T2 →
+ ⦃G, L⦄ ⊢ W1 ➡[h, o] W2 → ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}W1.T1 ➡*[h, o] ⓓ{a}W2.ⓐV2.T2.
+#h #o #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV1 #HV2 * -T2
+/4 width=9 by cpx_cpxs, cpxs_strap1, cpxs_bind_dx, cpxs_flat_dx, cpx_theta/
+qed.
+
+(* Basic inversion lemmas ***************************************************)
+
+lemma cpxs_inv_sort1: ∀h,o,G,L,U2,s. ⦃G, L⦄ ⊢ ⋆s ➡*[h, o] U2 →
+ ∃∃n,d. deg h o s (n+d) & U2 = ⋆((next h)^n s).
+#h #o #G #L #U2 #s #H @(cpxs_ind … H) -U2
+[ elim (deg_total h o s) #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/
+ | * #d0 #Hkd0 #H destruct -d
+ @(ex2_2_intro … (n+1) d0) /2 width=1 by deg_inv_prec/ >iter_SO //
+ ]
+]
+qed-.
+
+lemma cpxs_inv_cast1: ∀h,o,G,L,W1,T1,U2. ⦃G, L⦄ ⊢ ⓝW1.T1 ➡*[h, o] U2 →
+ ∨∨ ∃∃W2,T2. ⦃G, L⦄ ⊢ W1 ➡*[h, o] W2 & ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 & U2 = ⓝW2.T2
+ | ⦃G, L⦄ ⊢ T1 ➡*[h, o] U2
+ | ⦃G, L⦄ ⊢ W1 ➡*[h, o] U2.
+#h #o #G #L #W1 #T1 #U2 #H @(cpxs_ind … H) -U2 /3 width=5 by or3_intro0, ex3_2_intro/
+#U2 #U #_ #HU2 * /3 width=3 by cpxs_strap1, or3_intro1, or3_intro2/ *
+#W #T #HW1 #HT1 #H destruct
+elim (cpx_inv_cast1 … HU2) -HU2 /3 width=3 by cpxs_strap1, or3_intro1, or3_intro2/ *
+#W2 #T2 #HW2 #HT2 #H destruct
+lapply (cpxs_strap1 … HW1 … HW2) -W
+lapply (cpxs_strap1 … HT1 … HT2) -T /3 width=5 by or3_intro0, ex3_2_intro/
+qed-.
+
+lemma cpxs_inv_cnx1: ∀h,o,G,L,T,U. ⦃G, L⦄ ⊢ T ➡*[h, o] U → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ → T = U.
+#h #o #G #L #T #U #H @(cpxs_ind_dx … H) -T //
+#T0 #T #H1T0 #_ #IHT #H2T0
+lapply (H2T0 … H1T0) -H1T0 #H destruct /2 width=1 by/
+qed-.
+
+lemma cpxs_neq_inv_step_sn: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → (T1 = T2 → ⊥) →
+ ∃∃T. ⦃G, L⦄ ⊢ T1 ➡[h, o] T & T1 = T → ⊥ & ⦃G, L⦄ ⊢ T ➡*[h, o] T2.
+#h #o #G #L #T1 #T2 #H @(cpxs_ind_dx … H) -T1
+[ #H elim H -H //
+| #T1 #T #H1 #H2 #IH2 #H12 elim (eq_term_dec T1 T) #H destruct
+ [ -H1 -H2 /3 width=1 by/
+ | -IH2 /3 width=4 by ex3_intro/ (**) (* auto fails without clear *)
+ ]
+]
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/reduction/lpx_aaa.ma".
+include "basic_2/computation/cpxs.ma".
+
+(* CONTEXT-SENSITIVE EXTENDED PARALLEL COMPUTATION ON TERMS *****************)
+
+(* Properties about atomic arity assignment on terms ************************)
+
+lemma cpxs_aaa_conf: ∀h,o,G,L,T1,A. ⦃G, L⦄ ⊢ T1 ⁝ A →
+ ∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → ⦃G, L⦄ ⊢ T2 ⁝ A.
+#h #o #G #L #T1 #A #HT1 #T2 #HT12
+@(TC_Conf3 … HT1 ? HT12) -A -T1 -T2 /2 width=5 by cpx_aaa_conf/
+qed-.
+
+lemma cprs_aaa_conf: ∀G,L,T1,A. ⦃G, L⦄ ⊢ T1 ⁝ A → ∀T2. ⦃G, L⦄ ⊢ T1 ➡* T2 → ⦃G, L⦄ ⊢ T2 ⁝ A.
+/3 width=5 by cpxs_aaa_conf, cprs_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/lpx_drop.ma".
+include "basic_2/computation/cpxs_lift.ma".
+
+(* CONTEXT-SENSITIVE EXTENDED PARALLEL COMPUTATION ON TERMS *****************)
+
+(* Main properties **********************************************************)
+
+theorem cpxs_trans: ∀h,o,G,L. Transitive … (cpxs h o G L).
+normalize /2 width=3 by trans_TC/ qed-.
+
+theorem cpxs_bind: ∀h,o,a,I,G,L,V1,V2,T1,T2. ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ➡*[h, o] T2 →
+ ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 →
+ ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡*[h, o] ⓑ{a,I}V2.T2.
+#h #o #a #I #G #L #V1 #V2 #T1 #T2 #HT12 #H @(cpxs_ind … H) -V2
+/3 width=5 by cpxs_trans, cpxs_bind_dx/
+qed.
+
+theorem cpxs_flat: ∀h,o,I,G,L,V1,V2,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 →
+ ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 →
+ ⦃G, L⦄ ⊢ ⓕ{I}V1.T1 ➡*[h, o] ⓕ{I}V2.T2.
+#h #o #I #G #L #V1 #V2 #T1 #T2 #HT12 #H @(cpxs_ind … H) -V2
+/3 width=5 by cpxs_trans, cpxs_flat_dx/
+qed.
+
+theorem cpxs_beta_rc: ∀h,o,a,G,L,V1,V2,W1,W2,T1,T2.
+ ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 → ⦃G, L.ⓛW1⦄ ⊢ T1 ➡*[h, o] T2 → ⦃G, L⦄ ⊢ W1 ➡*[h, o] W2 →
+ ⦃G, L⦄ ⊢ ⓐV1.ⓛ{a}W1.T1 ➡*[h, o] ⓓ{a}ⓝW2.V2.T2.
+#h #o #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #HV12 #HT12 #H @(cpxs_ind … H) -W2
+/4 width=5 by cpxs_trans, cpxs_beta_dx, cpxs_bind_dx, cpx_pair_sn/
+qed.
+
+theorem cpxs_beta: ∀h,o,a,G,L,V1,V2,W1,W2,T1,T2.
+ ⦃G, L.ⓛW1⦄ ⊢ T1 ➡*[h, o] T2 → ⦃G, L⦄ ⊢ W1 ➡*[h, o] W2 → ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 →
+ ⦃G, L⦄ ⊢ ⓐV1.ⓛ{a}W1.T1 ➡*[h, o] ⓓ{a}ⓝW2.V2.T2.
+#h #o #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #HT12 #HW12 #H @(cpxs_ind … H) -V2
+/4 width=5 by cpxs_trans, cpxs_beta_rc, cpxs_bind_dx, cpx_flat/
+qed.
+
+theorem cpxs_theta_rc: ∀h,o,a,G,L,V1,V,V2,W1,W2,T1,T2.
+ ⦃G, L⦄ ⊢ V1 ➡[h, o] V → ⬆[0, 1] V ≡ V2 →
+ ⦃G, L.ⓓW1⦄ ⊢ T1 ➡*[h, o] T2 → ⦃G, L⦄ ⊢ W1 ➡*[h, o] W2 →
+ ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}W1.T1 ➡*[h, o] ⓓ{a}W2.ⓐV2.T2.
+#h #o #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV1 #HV2 #HT12 #H @(cpxs_ind … H) -W2
+/3 width=5 by cpxs_trans, cpxs_theta_dx, cpxs_bind_dx/
+qed.
+
+theorem cpxs_theta: ∀h,o,a,G,L,V1,V,V2,W1,W2,T1,T2.
+ ⬆[0, 1] V ≡ V2 → ⦃G, L⦄ ⊢ W1 ➡*[h, o] W2 →
+ ⦃G, L.ⓓW1⦄ ⊢ T1 ➡*[h, o] T2 → ⦃G, L⦄ ⊢ V1 ➡*[h, o] V →
+ ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}W1.T1 ➡*[h, o] ⓓ{a}W2.ⓐV2.T2.
+#h #o #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV2 #HW12 #HT12 #H @(TC_ind_dx … V1 H) -V1
+/3 width=5 by cpxs_trans, cpxs_theta_rc, cpxs_flat_dx/
+qed.
+
+(* Advanced inversion lemmas ************************************************)
+
+lemma cpxs_inv_appl1: ∀h,o,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓐV1.T1 ➡*[h, o] U2 →
+ ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 & ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 &
+ U2 = ⓐV2. T2
+ | ∃∃a,W,T. ⦃G, L⦄ ⊢ T1 ➡*[h, o] ⓛ{a}W.T & ⦃G, L⦄ ⊢ ⓓ{a}ⓝW.V1.T ➡*[h, o] U2
+ | ∃∃a,V0,V2,V,T. ⦃G, L⦄ ⊢ V1 ➡*[h, o] V0 & ⬆[0,1] V0 ≡ V2 &
+ ⦃G, L⦄ ⊢ T1 ➡*[h, o] ⓓ{a}V.T & ⦃G, L⦄ ⊢ ⓓ{a}V.ⓐV2.T ➡*[h, o] U2.
+#h #o #G #L #V1 #T1 #U2 #H @(cpxs_ind … H) -U2 [ /3 width=5 by or3_intro0, ex3_2_intro/ ]
+#U #U2 #_ #HU2 * *
+[ #V0 #T0 #HV10 #HT10 #H destruct
+ elim (cpx_inv_appl1 … HU2) -HU2 *
+ [ #V2 #T2 #HV02 #HT02 #H destruct /4 width=5 by cpxs_strap1, or3_intro0, ex3_2_intro/
+ | #a #V2 #W #W2 #T #T2 #HV02 #HW2 #HT2 #H1 #H2 destruct
+ lapply (cpxs_strap1 … HV10 … HV02) -V0 #HV12
+ lapply (lsubr_cpx_trans … HT2 (L.ⓓⓝW.V1) ?) -HT2
+ /5 width=5 by cpxs_bind, cpxs_flat_dx, cpx_cpxs, lsubr_beta, ex2_3_intro, or3_intro1/
+ | #a #V #V2 #W0 #W2 #T #T2 #HV0 #HV2 #HW02 #HT2 #H1 #H2 destruct
+ /5 width=10 by cpxs_flat_sn, cpxs_bind_dx, cpxs_strap1, ex4_5_intro, or3_intro2/
+ ]
+| /4 width=9 by cpxs_strap1, or3_intro1, ex2_3_intro/
+| /4 width=11 by cpxs_strap1, or3_intro2, ex4_5_intro/
+]
+qed-.
+
+(* Properties on sn extended parallel reduction for local environments ******)
+
+lemma lpx_cpx_trans: ∀h,o,G. c_r_transitive … (cpx h o G) (λ_.lpx h o G).
+#h #o #G #L2 #T1 #T2 #HT12 elim HT12 -G -L2 -T1 -T2
+[ /2 width=3 by/
+| /3 width=2 by cpx_cpxs, cpx_st/
+| #I #G #L2 #K2 #V0 #V2 #W2 #i #HLK2 #_ #HVW2 #IHV02 #L1 #HL12
+ elim (lpx_drop_trans_O1 … HL12 … HLK2) -L2 #X #HLK1 #H
+ elim (lpx_inv_pair2 … H) -H #K1 #V1 #HK12 #HV10 #H destruct
+ /4 width=7 by cpxs_delta, cpxs_strap2/
+|4,9: /4 width=1 by cpxs_beta, cpxs_bind, lpx_pair/
+|5,7,8: /3 width=1 by cpxs_flat, cpxs_ct, cpxs_eps/
+| /4 width=3 by cpxs_zeta, lpx_pair/
+| /4 width=3 by cpxs_theta, cpxs_strap1, lpx_pair/
+]
+qed-.
+
+lemma cpx_bind2: ∀h,o,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 →
+ ∀I,T1,T2. ⦃G, L.ⓑ{I}V2⦄ ⊢ T1 ➡[h, o] T2 →
+ ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡*[h, o] ⓑ{a,I}V2.T2.
+/4 width=5 by lpx_cpx_trans, cpxs_bind_dx, lpx_pair/ qed.
+
+(* Advanced properties ******************************************************)
+
+lemma lpx_cpxs_trans: ∀h,o,G. c_rs_transitive … (cpx h o G) (λ_.lpx h o G).
+#h #o #G @c_r_trans_LTC1 /2 width=3 by lpx_cpx_trans/ (**) (* full auto fails *)
+qed-.
+
+lemma cpxs_bind2_dx: ∀h,o,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 →
+ ∀I,T1,T2. ⦃G, L.ⓑ{I}V2⦄ ⊢ T1 ➡*[h, o] T2 →
+ ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡*[h, o] ⓑ{a,I}V2.T2.
+/4 width=5 by lpx_cpxs_trans, cpxs_bind_dx, lpx_pair/ qed.
+
+(* Properties on supclosure *************************************************)
+
+lemma fqu_cpxs_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, o] U2 → (T2 = U2 → ⊥) →
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐ ⦃G2, L2, U2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
+[ #I #G #L #V1 #V2 #HV12 #_ elim (lift_total V2 0 1)
+ #U2 #HVU2 @(ex3_intro … U2)
+ [1,3: /3 width=7 by fqu_drop, cpxs_delta, drop_pair, drop_drop/
+ | #H destruct
+ lapply (lift_inv_lref2_be … HVU2 ? ?) -HVU2 //
+ ]
+| #I #G #L #V1 #T #V2 #HV12 #H @(ex3_intro … (②{I}V2.T))
+ [1,3: /2 width=4 by fqu_pair_sn, cpxs_pair_sn/
+ | #H0 destruct /2 width=1 by/
+ ]
+| #a #I #G #L #V #T1 #T2 #HT12 #H @(ex3_intro … (ⓑ{a,I}V.T2))
+ [1,3: /2 width=4 by fqu_bind_dx, cpxs_bind/
+ | #H0 destruct /2 width=1 by/
+ ]
+| #I #G #L #V #T1 #T2 #HT12 #H @(ex3_intro … (ⓕ{I}V.T2))
+ [1,3: /2 width=4 by fqu_flat_dx, cpxs_flat/
+ | #H0 destruct /2 width=1 by/
+ ]
+| #G #L #K #T1 #U1 #k #HLK #HTU1 #T2 #HT12 #H elim (lift_total T2 0 (k+1))
+ #U2 #HTU2 @(ex3_intro … U2)
+ [1,3: /2 width=10 by cpxs_lift, fqu_drop/
+ | #H0 destruct /3 width=5 by lift_inj/
+]
+qed-.
+
+lemma fquq_cpxs_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, o] U2 → (T2 = U2 → ⊥) →
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H12 #U2 #HTU2 #H elim (fquq_inv_gen … H12) -H12
+[ #H12 elim (fqu_cpxs_trans_neq … H12 … HTU2 H) -T2
+ /3 width=4 by fqu_fquq, ex3_intro/
+| * #HG #HL #HT destruct /3 width=4 by ex3_intro/
+]
+qed-.
+
+lemma fqup_cpxs_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, o] U2 → (T2 = U2 → ⊥) →
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐+ ⦃G2, L2, U2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind_dx … H) -G1 -L1 -T1
+[ #G1 #L1 #T1 #H12 #U2 #HTU2 #H elim (fqu_cpxs_trans_neq … H12 … HTU2 H) -T2
+ /3 width=4 by fqu_fqup, ex3_intro/
+| #G #G1 #L #L1 #T #T1 #H1 #_ #IH12 #U2 #HTU2 #H elim (IH12 … HTU2 H) -T2
+ #U1 #HTU1 #H #H12 elim (fqu_cpxs_trans_neq … H1 … HTU1 H) -T1
+ /3 width=8 by fqup_strap2, ex3_intro/
+]
+qed-.
+
+lemma fqus_cpxs_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, o] U2 → (T2 = U2 → ⊥) →
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐* ⦃G2, L2, U2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H12 #U2 #HTU2 #H elim (fqus_inv_gen … H12) -H12
+[ #H12 elim (fqup_cpxs_trans_neq … H12 … HTU2 H) -T2
+ /3 width=4 by fqup_fqus, ex3_intro/
+| * #HG #HL #HT destruct /3 width=4 by ex3_intro/
+]
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/multiple/fqus_fqus.ma".
+include "basic_2/reduction/cpx_lift.ma".
+include "basic_2/computation/cpxs.ma".
+
+(* CONTEXT-SENSITIVE EXTENDED PARALLEL COMPUTATION ON TERMS *****************)
+
+(* Advanced properties ******************************************************)
+
+lemma cpxs_delta: ∀h,o,I,G,L,K,V,V2,i.
+ ⬇[i] L ≡ K.ⓑ{I}V → ⦃G, K⦄ ⊢ V ➡*[h, o] V2 →
+ ∀W2. ⬆[0, i+1] V2 ≡ W2 → ⦃G, L⦄ ⊢ #i ➡*[h, o] W2.
+#h #o #I #G #L #K #V #V2 #i #HLK #H elim H -V2
+[ /3 width=9 by cpx_cpxs, cpx_delta/
+| #V1 lapply (drop_fwd_drop2 … HLK) -HLK
+ elim (lift_total V1 0 (i+1)) /4 width=12 by cpx_lift, cpxs_strap1/
+]
+qed.
+
+lemma lstas_cpxs: ∀h,o,G,L,T1,T2,d2. ⦃G, L⦄ ⊢ T1 •*[h, d2] T2 →
+ ∀d1. ⦃G, L⦄ ⊢ T1 ▪[h, o] d1 → d2 ≤ d1 → ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2.
+#h #o #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 #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 #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 … 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/
+| /4 width=3 by cpxs_eps, da_inv_flat/
+]
+qed.
+
+(* Advanced inversion lemmas ************************************************)
+
+lemma cpxs_inv_lref1: ∀h,o,G,L,T2,i. ⦃G, L⦄ ⊢ #i ➡*[h, o] T2 →
+ T2 = #i ∨
+ ∃∃I,K,V1,T1. ⬇[i] L ≡ K.ⓑ{I}V1 & ⦃G, K⦄ ⊢ V1 ➡*[h, o] T1 &
+ ⬆[0, i+1] T1 ≡ T2.
+#h #o #G #L #T2 #i #H @(cpxs_ind … H) -T2 /2 width=1 by or_introl/
+#T #T2 #_ #HT2 *
+[ #H destruct
+ elim (cpx_inv_lref1 … HT2) -HT2 /2 width=1 by or_introl/
+ * /4 width=7 by cpx_cpxs, ex3_4_intro, or_intror/
+| * #I #K #V1 #T1 #HLK #HVT1 #HT1
+ lapply (drop_fwd_drop2 … HLK) #H0LK
+ elim (cpx_inv_lift1 … HT2 … H0LK … HT1) -H0LK -T
+ /4 width=7 by cpxs_strap1, ex3_4_intro, or_intror/
+]
+qed-.
+
+(* Relocation properties ****************************************************)
+
+lemma cpxs_lift: ∀h,o,G. d_liftable (cpxs h o G).
+/3 width=10 by cpx_lift, cpxs_strap1, d_liftable_LTC/ qed.
+
+lemma cpxs_inv_lift1: ∀h,o,G. d_deliftable_sn (cpxs h o G).
+/3 width=6 by d_deliftable_sn_LTC, cpx_inv_lift1/
+qed-.
+
+(* Properties on supclosure *************************************************)
+
+lemma fqu_cpxs_trans: ∀h,o,G1,G2,L1,L2,T2,U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, o] U2 →
+ ∀T1. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] U1 & ⦃G1, L1, U1⦄ ⊐ ⦃G2, L2, U2⦄.
+#h #o #G1 #G2 #L1 #L2 #T2 #U2 #H @(cpxs_ind_dx … H) -T2 /2 width=3 by ex2_intro/
+#T #T2 #HT2 #_ #IHTU2 #T1 #HT1 elim (fqu_cpx_trans … HT1 … HT2) -T
+#T #HT1 #HT2 elim (IHTU2 … HT2) -T2 /3 width=3 by cpxs_strap2, ex2_intro/
+qed-.
+
+lemma fquq_cpxs_trans: ∀h,o,G1,G2,L1,L2,T2,U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, o] U2 →
+ ∀T1. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] U1 & ⦃G1, L1, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
+#h #o #G1 #G2 #L1 #L2 #T2 #U2 #HTU2 #T1 #H elim (fquq_inv_gen … H) -H
+[ #HT12 elim (fqu_cpxs_trans … HTU2 … HT12) /3 width=3 by fqu_fquq, ex2_intro/
+| * #H1 #H2 #H3 destruct /2 width=3 by ex2_intro/
+]
+qed-.
+
+lemma fquq_lstas_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
+ ∀U2,d1. ⦃G2, L2⦄ ⊢ T2 •*[h, d1] U2 →
+ ∀d2. ⦃G2, L2⦄ ⊢ T2 ▪[h, o] d2 → d1 ≤ d2 →
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] U1 & ⦃G1, L1, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
+/3 width=5 by fquq_cpxs_trans, lstas_cpxs/ qed-.
+
+lemma fqup_cpxs_trans: ∀h,o,G1,G2,L1,L2,T2,U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, o] U2 →
+ ∀T1. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] U1 & ⦃G1, L1, U1⦄ ⊐+ ⦃G2, L2, U2⦄.
+#h #o #G1 #G2 #L1 #L2 #T2 #U2 #H @(cpxs_ind_dx … H) -T2 /2 width=3 by ex2_intro/
+#T #T2 #HT2 #_ #IHTU2 #T1 #HT1 elim (fqup_cpx_trans … HT1 … HT2) -T
+#U1 #HTU1 #H2 elim (IHTU2 … H2) -T2 /3 width=3 by cpxs_strap2, ex2_intro/
+qed-.
+
+lemma fqus_cpxs_trans: ∀h,o,G1,G2,L1,L2,T2,U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, o] U2 →
+ ∀T1. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] U1 & ⦃G1, L1, U1⦄ ⊐* ⦃G2, L2, U2⦄.
+#h #o #G1 #G2 #L1 #L2 #T2 #U2 #HTU2 #T1 #H elim (fqus_inv_gen … H) -H
+[ #HT12 elim (fqup_cpxs_trans … HTU2 … HT12) /3 width=3 by fqup_fqus, ex2_intro/
+| * #H1 #H2 #H3 destruct /2 width=3 by ex2_intro/
+]
+qed-.
+
+lemma fqus_lstas_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
+ ∀U2,d1. ⦃G2, L2⦄ ⊢ T2 •*[h, d1] U2 →
+ ∀d2. ⦃G2, L2⦄ ⊢ T2 ▪[h, o] d2 → d1 ≤ d2 →
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] 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_lleq.ma".
+include "basic_2/computation/cpxs.ma".
+
+(* CONTEXT-SENSITIVE EXTENDED PARALLEL COMPUTATION ON TERMS *****************)
+
+(* Properties on lazy equivalence for local environments ********************)
+
+lemma lleq_cpxs_trans: ∀h,o,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ➡*[h, o] T2 →
+ ∀L1. L1 ≡[T1, 0] L2 → ⦃G, L1⦄ ⊢ T1 ➡*[h, o] T2.
+#h #o #G #L2 #T1 #T2 #H @(cpxs_ind_dx … H) -T1
+/4 width=6 by cpx_lleq_conf_dx, lleq_cpx_trans, cpxs_strap2/
+qed-.
+
+lemma cpxs_lleq_conf: ∀h,o,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ➡*[h, o] T2 →
+ ∀L1. L2 ≡[T1, 0] L1 → ⦃G, L1⦄ ⊢ T1 ➡*[h, o] T2.
+/3 width=3 by lleq_cpxs_trans, lleq_sym/ qed-.
+
+lemma cpxs_lleq_conf_dx: ∀h,o,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ➡*[h, o] T2 →
+ ∀L1. L1 ≡[T1, 0] L2 → L1 ≡[T2, 0] L2.
+#h #o #G #L2 #T1 #T2 #H @(cpxs_ind … H) -T2 /3 width=6 by cpx_lleq_conf_dx/
+qed-.
+
+lemma cpxs_lleq_conf_sn: ∀h,o,G,L1,T1,T2. ⦃G, L1⦄ ⊢ T1 ➡*[h, o] T2 →
+ ∀L2. L1 ≡[T1, 0] L2 → L1 ≡[T2, 0] L2.
+/4 width=6 by cpxs_lleq_conf_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/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,o,G. lsub_trans … (cpxs h o G) (lreq 0 (∞)).
+/3 width=5 by lreq_cpx_trans, LTC_lsub_trans/
+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/tsts.ma".
+include "basic_2/computation/lpxs_cpxs.ma".
+
+(* CONTEXT-SENSITIVE EXTENDED PARALLEL COMPUTATION ON TERMS *****************)
+
+(* Forward lemmas involving same top term structure *************************)
+
+lemma cpxs_fwd_cnx: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ → ∀U. ⦃G, L⦄ ⊢ T ➡*[h, o] U → T ≂ U.
+#h #o #G #L #T #HT #U #H
+>(cpxs_inv_cnx1 … H HT) -G -L -T //
+qed-.
+
+lemma cpxs_fwd_sort: ∀h,o,G,L,U,s. ⦃G, L⦄ ⊢ ⋆s ➡*[h, o] U →
+ ⋆s ≂ U ∨ ⦃G, L⦄ ⊢ ⋆(next h s) ➡*[h, o] U.
+#h #o #G #L #U #s #H
+elim (cpxs_inv_sort1 … H) -H #n #d generalize in match s; -s @(nat_ind_plus … n) -n
+[ #s #_ #H -d destruct /2 width=1 by or_introl/
+| #n #IHn #s >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 Hnd; -Hnd @(nat_ind_plus … n) -n
+ /4 width=3 by cpxs_strap2, cpx_st, or_intror/
+ | >iter_SO >iter_n_Sm //
+ ]
+]
+qed-.
+
+(* Basic_1: was just: pr3_iso_beta *)
+lemma cpxs_fwd_beta: ∀h,o,a,G,L,V,W,T,U. ⦃G, L⦄ ⊢ ⓐV.ⓛ{a}W.T ➡*[h, o] U →
+ ⓐV.ⓛ{a}W.T ≂ U ∨ ⦃G, L⦄ ⊢ ⓓ{a}ⓝW.V.T ➡*[h, o] U.
+#h #o #a #G #L #V #W #T #U #H
+elim (cpxs_inv_appl1 … H) -H *
+[ #V0 #T0 #_ #_ #H destruct /2 width=1 by tsts_pair, or_introl/
+| #b #W0 #T0 #HT0 #HU
+ elim (cpxs_inv_abst1 … HT0) -HT0 #W1 #T1 #HW1 #HT1 #H destruct
+ lapply (lsubr_cpxs_trans … HT1 (L.ⓓⓝW.V) ?) -HT1
+ /5 width=3 by cpxs_trans, cpxs_bind, cpxs_pair_sn, lsubr_beta, or_intror/
+| #b #V1 #V2 #V0 #T1 #_ #_ #HT1 #_
+ elim (cpxs_inv_abst1 … HT1) -HT1 #W2 #T2 #_ #_ #H destruct
+]
+qed-.
+
+(* Note: probably this is an inversion lemma *)
+lemma cpxs_fwd_delta: ∀h,o,I,G,L,K,V1,i. ⬇[i] L ≡ K.ⓑ{I}V1 →
+ ∀V2. ⬆[0, i + 1] V1 ≡ V2 →
+ ∀U. ⦃G, L⦄ ⊢ #i ➡*[h, o] U →
+ #i ≂ U ∨ ⦃G, L⦄ ⊢ V2 ➡*[h, o] U.
+#h #o #I #G #L #K #V1 #i #HLK #V2 #HV12 #U #H
+elim (cpxs_inv_lref1 … H) -H /2 width=1 by or_introl/
+* #I0 #K0 #V0 #U0 #HLK0 #HVU0 #HU0
+lapply (drop_mono … HLK0 … HLK) -HLK0 #H destruct
+/4 width=10 by cpxs_lift, drop_fwd_drop2, or_intror/
+qed-.
+
+lemma cpxs_fwd_theta: ∀h,o,a,G,L,V1,V,T,U. ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}V.T ➡*[h, o] U →
+ ∀V2. ⬆[0, 1] V1 ≡ V2 → ⓐV1.ⓓ{a}V.T ≂ U ∨
+ ⦃G, L⦄ ⊢ ⓓ{a}V.ⓐV2.T ➡*[h, o] U.
+#h #o #a #G #L #V1 #V #T #U #H #V2 #HV12
+elim (cpxs_inv_appl1 … H) -H *
+[ -HV12 #V0 #T0 #_ #_ #H destruct /2 width=1 by tsts_pair, or_introl/
+| #b #W #T0 #HT0 #HU
+ elim (cpxs_inv_abbr1 … HT0) -HT0 *
+ [ #V3 #T3 #_ #_ #H destruct
+ | #X #HT2 #H #H0 destruct
+ elim (lift_inv_bind1 … H) -H #W2 #T2 #HW2 #HT02 #H destruct
+ @or_intror @(cpxs_trans … HU) -U (**) (* explicit constructor *)
+ @(cpxs_trans … (+ⓓV.ⓐV2.ⓛ{b}W2.T2)) [ /3 width=1 by cpxs_flat_dx, cpxs_bind_dx/ ] -T
+ @(cpxs_strap2 … (ⓐV1.ⓛ{b}W.T0)) [2: /2 width=1 by cpxs_beta_dx/ ]
+ /4 width=7 by cpx_zeta, lift_bind, lift_flat/
+ ]
+| #b #V3 #V4 #V0 #T0 #HV13 #HV34 #HT0 #HU
+ @or_intror @(cpxs_trans … HU) -U (**) (* explicit constructor *)
+ elim (cpxs_inv_abbr1 … HT0) -HT0 *
+ [ #V5 #T5 #HV5 #HT5 #H destruct
+ lapply (cpxs_lift … HV13 (L.ⓓV) … HV12 … HV34) -V1 -V3
+ /3 width=2 by cpxs_flat, cpxs_bind, drop_drop/
+ | #X #HT1 #H #H0 destruct
+ elim (lift_inv_bind1 … H) -H #V5 #T5 #HV05 #HT05 #H destruct
+ lapply (cpxs_lift … HV13 (L.ⓓV0) … HV12 … HV34) -V3 /2 width=2 by drop_drop/ #HV24
+ @(cpxs_trans … (+ⓓV.ⓐV2.ⓓ{b}V5.T5)) [ /3 width=1 by cpxs_flat_dx, cpxs_bind_dx/ ] -T
+ @(cpxs_strap2 … (ⓐV1.ⓓ{b}V0.T0)) [ /4 width=7 by cpx_zeta, lift_bind, lift_flat/ ] -V -V5 -T5
+ @(cpxs_strap2 … (ⓓ{b}V0.ⓐV2.T0)) /3 width=3 by cpxs_pair_sn, cpxs_bind_dx, cpr_cpx, cpr_theta/
+ ]
+]
+qed-.
+
+lemma cpxs_fwd_cast: ∀h,o,G,L,W,T,U. ⦃G, L⦄ ⊢ ⓝW.T ➡*[h, o] U →
+ ∨∨ ⓝW. T ≂ U | ⦃G, L⦄ ⊢ T ➡*[h, o] U | ⦃G, L⦄ ⊢ W ➡*[h, o] U.
+#h #o #G #L #W #T #U #H
+elim (cpxs_inv_cast1 … H) -H /2 width=1 by or3_intro1, or3_intro2/ *
+#W0 #T0 #_ #_ #H destruct /2 width=1 by tsts_pair, or3_intro0/
+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/tsts_vector.ma".
+include "basic_2/substitution/lift_vector.ma".
+include "basic_2/computation/cpxs_tsts.ma".
+
+(* CONTEXT-SENSITIVE EXTENDED PARALLEL COMPUTATION ON TERMS *****************)
+
+(* Vector form of forward lemmas involving same top term structure **********)
+
+(* Basic_1: was just: nf2_iso_appls_lref *)
+lemma cpxs_fwd_cnx_vector: ∀h,o,G,L,T. 𝐒⦃T⦄ → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ →
+ ∀Vs,U. ⦃G, L⦄ ⊢ ⒶVs.T ➡*[h, o] U → ⒶVs.T ≂ U.
+#h #o #G #L #T #H1T #H2T #Vs elim Vs -Vs [ @(cpxs_fwd_cnx … H2T) ] (**) (* /2 width=3 by cpxs_fwd_cnx/ does not work *)
+#V #Vs #IHVs #U #H
+elim (cpxs_inv_appl1 … H) -H *
+[ -IHVs #V0 #T0 #_ #_ #H destruct /2 width=1 by tsts_pair/
+| #a #W0 #T0 #HT0 #HU
+ lapply (IHVs … HT0) -IHVs -HT0 #HT0
+ elim (tsts_inv_bind_applv_simple … HT0) //
+| #a #V1 #V2 #V0 #T0 #HV1 #HV12 #HT0 #HU
+ lapply (IHVs … HT0) -IHVs -HT0 #HT0
+ elim (tsts_inv_bind_applv_simple … HT0) //
+]
+qed-.
+
+lemma cpxs_fwd_sort_vector: ∀h,o,G,L,s,Vs,U. ⦃G, L⦄ ⊢ ⒶVs.⋆s ➡*[h, o] U →
+ ⒶVs.⋆s ≂ U ∨ ⦃G, L⦄ ⊢ ⒶVs.⋆(next h s) ➡*[h, o] U.
+#h #o #G #L #s #Vs elim Vs -Vs /2 width=1 by cpxs_fwd_sort/
+#V #Vs #IHVs #U #H
+elim (cpxs_inv_appl1 … H) -H *
+[ -IHVs #V1 #T1 #_ #_ #H destruct /2 width=1 by tsts_pair, or_introl/
+| #a #W1 #T1 #HT1 #HU
+ elim (IHVs … HT1) -IHVs -HT1 #HT1
+ [ elim (tsts_inv_bind_applv_simple … HT1) //
+ | @or_intror (**) (* explicit constructor *)
+ @(cpxs_trans … HU) -U
+ @(cpxs_strap1 … (ⓐV.ⓛ{a}W1.T1)) /3 width=1 by cpxs_flat_dx, cpr_cpx, cpr_beta/
+ ]
+| #a #V1 #V2 #V3 #T1 #HV01 #HV12 #HT1 #HU
+ elim (IHVs … HT1) -IHVs -HT1 #HT1
+ [ elim (tsts_inv_bind_applv_simple … HT1) //
+ | @or_intror (**) (* explicit constructor *)
+ @(cpxs_trans … HU) -U
+ @(cpxs_strap1 … (ⓐV1.ⓓ{a}V3.T1)) /3 width=3 by cpxs_flat, cpr_cpx, cpr_theta/
+ ]
+]
+qed-.
+
+
+(* Basic_1: was just: pr3_iso_appls_beta *)
+lemma cpxs_fwd_beta_vector: ∀h,o,a,G,L,Vs,V,W,T,U. ⦃G, L⦄ ⊢ ⒶVs.ⓐV.ⓛ{a}W.T ➡*[h, o] U →
+ ⒶVs. ⓐV. ⓛ{a}W. T ≂ U ∨ ⦃G, L⦄ ⊢ ⒶVs.ⓓ{a}ⓝW.V.T ➡*[h, o] U.
+#h #o #a #G #L #Vs elim Vs -Vs /2 width=1 by cpxs_fwd_beta/
+#V0 #Vs #IHVs #V #W #T #U #H
+elim (cpxs_inv_appl1 … H) -H *
+[ -IHVs #V1 #T1 #_ #_ #H destruct /2 width=1 by tsts_pair, or_introl/
+| #b #W1 #T1 #HT1 #HU
+ elim (IHVs … HT1) -IHVs -HT1 #HT1
+ [ elim (tsts_inv_bind_applv_simple … HT1) //
+ | @or_intror (**) (* explicit constructor *)
+ @(cpxs_trans … HU) -U
+ @(cpxs_strap1 … (ⓐV0.ⓛ{b}W1.T1)) /3 width=1 by cpxs_flat_dx, cpr_cpx, cpr_beta/
+ ]
+| #b #V1 #V2 #V3 #T1 #HV01 #HV12 #HT1 #HU
+ elim (IHVs … HT1) -IHVs -HT1 #HT1
+ [ elim (tsts_inv_bind_applv_simple … HT1) //
+ | @or_intror (**) (* explicit constructor *)
+ @(cpxs_trans … HU) -U
+ @(cpxs_strap1 … (ⓐV1.ⓓ{b}V3.T1)) /3 width=3 by cpxs_flat, cpr_cpx, cpr_theta/
+ ]
+]
+qed-.
+
+lemma cpxs_fwd_delta_vector: ∀h,o,I,G,L,K,V1,i. ⬇[i] L ≡ K.ⓑ{I}V1 →
+ ∀V2. ⬆[0, i + 1] V1 ≡ V2 →
+ ∀Vs,U. ⦃G, L⦄ ⊢ ⒶVs.#i ➡*[h, o] U →
+ ⒶVs.#i ≂ U ∨ ⦃G, L⦄ ⊢ ⒶVs.V2 ➡*[h, o] U.
+#h #o #I #G #L #K #V1 #i #HLK #V2 #HV12 #Vs elim Vs -Vs /2 width=5 by cpxs_fwd_delta/
+#V #Vs #IHVs #U #H -K -V1
+elim (cpxs_inv_appl1 … H) -H *
+[ -IHVs #V0 #T0 #_ #_ #H destruct /2 width=1 by tsts_pair, or_introl/
+| #b #W0 #T0 #HT0 #HU
+ elim (IHVs … HT0) -IHVs -HT0 #HT0
+ [ elim (tsts_inv_bind_applv_simple … HT0) //
+ | @or_intror -i (**) (* explicit constructor *)
+ @(cpxs_trans … HU) -U
+ @(cpxs_strap1 … (ⓐV.ⓛ{b}W0.T0)) /3 width=1 by cpxs_flat_dx, cpr_cpx, cpr_beta/
+ ]
+| #b #V0 #V1 #V3 #T0 #HV0 #HV01 #HT0 #HU
+ elim (IHVs … HT0) -IHVs -HT0 #HT0
+ [ elim (tsts_inv_bind_applv_simple … HT0) //
+ | @or_intror -i (**) (* explicit constructor *)
+ @(cpxs_trans … HU) -U
+ @(cpxs_strap1 … (ⓐV0.ⓓ{b}V3.T0)) /3 width=3 by cpxs_flat, cpr_cpx, cpr_theta/
+ ]
+]
+qed-.
+
+(* Basic_1: was just: pr3_iso_appls_abbr *)
+lemma cpxs_fwd_theta_vector: ∀h,o,G,L,V1c,V2c. ⬆[0, 1] V1c ≡ V2c →
+ ∀a,V,T,U. ⦃G, L⦄ ⊢ ⒶV1c.ⓓ{a}V.T ➡*[h, o] U →
+ ⒶV1c. ⓓ{a}V. T ≂ U ∨ ⦃G, L⦄ ⊢ ⓓ{a}V.ⒶV2c.T ➡*[h, o] U.
+#h #o #G #L #V1c #V2c * -V1c -V2c /3 width=1 by or_intror/
+#V1c #V2c #V1a #V2a #HV12a #HV12c #a
+generalize in match HV12a; -HV12a
+generalize in match V2a; -V2a
+generalize in match V1a; -V1a
+elim HV12c -V1c -V2c /2 width=1 by cpxs_fwd_theta/
+#V1c #V2c #V1b #V2b #HV12b #_ #IHV12c #V1a #V2a #HV12a #V #T #U #H
+elim (cpxs_inv_appl1 … H) -H *
+[ -IHV12c -HV12a -HV12b #V0 #T0 #_ #_ #H destruct /2 width=1 by tsts_pair, or_introl/
+| #b #W0 #T0 #HT0 #HU
+ elim (IHV12c … HV12b … HT0) -IHV12c -HT0 #HT0
+ [ -HV12a -HV12b -HU
+ elim (tsts_inv_pair1 … HT0) #V1 #T1 #H destruct
+ | @or_intror -V1c (**) (* explicit constructor *)
+ @(cpxs_trans … HU) -U
+ elim (cpxs_inv_abbr1 … HT0) -HT0 *
+ [ -HV12a -HV12b #V1 #T1 #_ #_ #H destruct
+ | -V1b #X #HT1 #H #H0 destruct
+ elim (lift_inv_bind1 … H) -H #W1 #T1 #HW01 #HT01 #H destruct
+ @(cpxs_trans … (+ⓓV.ⓐV2a.ⓛ{b}W1.T1)) [ /3 width=1 by cpxs_flat_dx, cpxs_bind_dx/ ] -T -V2b -V2c
+ @(cpxs_strap2 … (ⓐV1a.ⓛ{b}W0.T0))
+ /4 width=7 by cpxs_beta_dx, cpx_zeta, lift_bind, lift_flat/
+ ]
+ ]
+| #b #V0a #Va #V0 #T0 #HV10a #HV0a #HT0 #HU
+ elim (IHV12c … HV12b … HT0) -HV12b -IHV12c -HT0 #HT0
+ [ -HV12a -HV10a -HV0a -HU
+ elim (tsts_inv_pair1 … HT0) #V1 #T1 #H destruct
+ | @or_intror -V1c -V1b (**) (* explicit constructor *)
+ @(cpxs_trans … HU) -U
+ elim (cpxs_inv_abbr1 … HT0) -HT0 *
+ [ #V1 #T1 #HV1 #HT1 #H destruct
+ lapply (cpxs_lift … HV10a (L.ⓓV) (Ⓕ) … HV12a … HV0a) -V1a -V0a [ /2 width=1 by drop_drop/ ] #HV2a
+ @(cpxs_trans … (ⓓ{a}V.ⓐV2a.T1)) /3 width=1 by cpxs_bind, cpxs_pair_sn, cpxs_flat_dx, cpxs_bind_dx/
+ | #X #HT1 #H #H0 destruct
+ elim (lift_inv_bind1 … H) -H #V1 #T1 #HW01 #HT01 #H destruct
+ lapply (cpxs_lift … HV10a (L.ⓓV0) (Ⓕ) … HV12a … HV0a) -V0a [ /2 width=1 by drop_drop/ ] #HV2a
+ @(cpxs_trans … (+ⓓV.ⓐV2a.ⓓ{b}V1.T1)) [ /3 width=1 by cpxs_flat_dx, cpxs_bind_dx/ ] -T -V2b -V2c
+ @(cpxs_strap2 … (ⓐV1a.ⓓ{b}V0.T0)) [ /4 width=7 by cpx_zeta, lift_bind, lift_flat/ ] -V -V1 -T1
+ @(cpxs_strap2 … (ⓓ{b}V0.ⓐV2a.T0)) /3 width=3 by cpxs_pair_sn, cpxs_bind_dx, cpr_cpx, cpr_theta/
+ ]
+ ]
+]
+qed-.
+
+(* Basic_1: was just: pr3_iso_appls_cast *)
+lemma cpxs_fwd_cast_vector: ∀h,o,G,L,Vs,W,T,U. ⦃G, L⦄ ⊢ ⒶVs.ⓝW.T ➡*[h, o] U →
+ ∨∨ ⒶVs. ⓝW. T ≂ U
+ | ⦃G, L⦄ ⊢ ⒶVs.T ➡*[h, o] U
+ | ⦃G, L⦄ ⊢ ⒶVs.W ➡*[h, o] U.
+#h #o #G #L #Vs elim Vs -Vs /2 width=1 by cpxs_fwd_cast/
+#V #Vs #IHVs #W #T #U #H
+elim (cpxs_inv_appl1 … H) -H *
+[ -IHVs #V0 #T0 #_ #_ #H destruct /2 width=1 by tsts_pair, or3_intro0/
+| #b #W0 #T0 #HT0 #HU elim (IHVs … HT0) -IHVs -HT0 #HT0
+ [ elim (tsts_inv_bind_applv_simple … HT0) //
+ | @or3_intro1 -W (**) (* explicit constructor *)
+ @(cpxs_trans … HU) -U
+ @(cpxs_strap1 … (ⓐV.ⓛ{b}W0.T0)) /2 width=1 by cpxs_flat_dx, cpx_beta/
+ | @or3_intro2 -T (**) (* explicit constructor *)
+ @(cpxs_trans … HU) -U
+ @(cpxs_strap1 … (ⓐV.ⓛ{b}W0.T0)) /2 width=1 by cpxs_flat_dx, cpx_beta/
+ ]
+| #b #V0 #V1 #V2 #T0 #HV0 #HV01 #HT0 #HU
+ elim (IHVs … HT0) -IHVs -HT0 #HT0
+ [ elim (tsts_inv_bind_applv_simple … HT0) //
+ | @or3_intro1 -W (**) (* explicit constructor *)
+ @(cpxs_trans … HU) -U
+ @(cpxs_strap1 … (ⓐV0.ⓓ{b}V2.T0)) /2 width=3 by cpxs_flat, cpx_theta/
+ | @or3_intro2 -T (**) (* explicit constructor *)
+ @(cpxs_trans … HU) -U
+ @(cpxs_strap1 … (ⓐV0.ⓓ{b}V2.T0)) /2 width=3 by cpxs_flat, cpx_theta/
+ ]
+]
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/sn_5.ma".
+include "basic_2/reduction/cnx.ma".
+
+(* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERMS ********************)
+
+definition csx: ∀h. sd h → relation3 genv lenv term ≝
+ λh,o,G,L. SN … (cpx h o G L) (eq …).
+
+interpretation
+ "context-sensitive extended strong normalization (term)"
+ 'SN h o G L T = (csx h o G L T).
+
+(* Basic eliminators ********************************************************)
+
+lemma csx_ind: ∀h,o,G,L. ∀R:predicate term.
+ (∀T1. ⦃G, L⦄ ⊢ ⬊*[h, o] T1 →
+ (∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → (T1 = T2 → ⊥) → R T2) →
+ R T1
+ ) →
+ ∀T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → R T.
+#h #o #G #L #R #H0 #T1 #H elim H -T1
+/5 width=1 by SN_intro/
+qed-.
+
+(* Basic properties *********************************************************)
+
+(* Basic_1: was just: sn3_pr2_intro *)
+lemma csx_intro: ∀h,o,G,L,T1.
+ (∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → (T1 = T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊*[h, o] T2) →
+ ⦃G, L⦄ ⊢ ⬊*[h, o] T1.
+/4 width=1 by SN_intro/ qed.
+
+lemma csx_cpx_trans: ∀h,o,G,L,T1. ⦃G, L⦄ ⊢ ⬊*[h, o] T1 →
+ ∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → ⦃G, L⦄ ⊢ ⬊*[h, o] T2.
+#h #o #G #L #T1 #H @(csx_ind … H) -T1 #T1 #HT1 #IHT1 #T2 #HLT12
+elim (eq_term_dec T1 T2) #HT12 destruct /3 width=4 by/
+qed-.
+
+(* Basic_1: was just: sn3_nf2 *)
+lemma cnx_csx: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ → ⦃G, L⦄ ⊢ ⬊*[h, o] T.
+/2 width=1 by NF_to_SN/ qed.
+
+lemma csx_sort: ∀h,o,G,L,s. ⦃G, L⦄ ⊢ ⬊*[h, o] ⋆s.
+#h #o #G #L #s elim (deg_total h o s)
+#d generalize in match s; -s @(nat_ind_plus … d) -d /3 width=6 by cnx_csx, cnx_sort/
+#d #IHd #s #Hkd lapply (deg_next_SO … Hkd) -Hkd
+#Hkd @csx_intro #X #H #HX elim (cpx_inv_sort1 … H) -H
+[ #H destruct elim HX //
+| -HX * #d0 #_ #H destruct -d0 /2 width=1 by/
+]
+qed.
+
+(* Basic_1: was just: sn3_cast *)
+lemma csx_cast: ∀h,o,G,L,W. ⦃G, L⦄ ⊢ ⬊*[h, o] W →
+ ∀T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓝW.T.
+#h #o #G #L #W #HW @(csx_ind … HW) -W #W #HW #IHW #T #HT @(csx_ind … HT) -T #T #HT #IHT
+@csx_intro #X #H1 #H2
+elim (cpx_inv_cast1 … H1) -H1
+[ * #W0 #T0 #HLW0 #HLT0 #H destruct
+ elim (eq_false_inv_tpair_sn … H2) -H2
+ [ /3 width=3 by csx_cpx_trans/
+ | -HLW0 * #H destruct /3 width=1 by/
+ ]
+|2,3: /3 width=3 by csx_cpx_trans/
+]
+qed.
+
+(* Basic forward lemmas *****************************************************)
+
+fact csx_fwd_pair_sn_aux: ∀h,o,G,L,U. ⦃G, L⦄ ⊢ ⬊*[h, o] U →
+ ∀I,V,T. U = ②{I}V.T → ⦃G, L⦄ ⊢ ⬊*[h, o] V.
+#h #o #G #L #U #H elim H -H #U0 #_ #IH #I #V #T #H destruct
+@csx_intro #V2 #HLV2 #HV2
+@(IH (②{I}V2.T)) -IH /2 width=3 by cpx_pair_sn/ -HLV2
+#H destruct /2 width=1 by/
+qed-.
+
+(* Basic_1: was just: sn3_gen_head *)
+lemma csx_fwd_pair_sn: ∀h,o,I,G,L,V,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ②{I}V.T → ⦃G, L⦄ ⊢ ⬊*[h, o] V.
+/2 width=5 by csx_fwd_pair_sn_aux/ qed-.
+
+fact csx_fwd_bind_dx_aux: ∀h,o,G,L,U. ⦃G, L⦄ ⊢ ⬊*[h, o] U →
+ ∀a,I,V,T. U = ⓑ{a,I}V.T → ⦃G, L.ⓑ{I}V⦄ ⊢ ⬊*[h, o] T.
+#h #o #G #L #U #H elim H -H #U0 #_ #IH #a #I #V #T #H destruct
+@csx_intro #T2 #HLT2 #HT2
+@(IH (ⓑ{a,I}V.T2)) -IH /2 width=3 by cpx_bind/ -HLT2
+#H destruct /2 width=1 by/
+qed-.
+
+(* Basic_1: was just: sn3_gen_bind *)
+lemma csx_fwd_bind_dx: ∀h,o,a,I,G,L,V,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ⓑ{a,I}V.T → ⦃G, L.ⓑ{I}V⦄ ⊢ ⬊*[h, o] T.
+/2 width=4 by csx_fwd_bind_dx_aux/ qed-.
+
+fact csx_fwd_flat_dx_aux: ∀h,o,G,L,U. ⦃G, L⦄ ⊢ ⬊*[h, o] U →
+ ∀I,V,T. U = ⓕ{I}V.T → ⦃G, L⦄ ⊢ ⬊*[h, o] T.
+#h #o #G #L #U #H elim H -H #U0 #_ #IH #I #V #T #H destruct
+@csx_intro #T2 #HLT2 #HT2
+@(IH (ⓕ{I}V.T2)) -IH /2 width=3 by cpx_flat/ -HLT2
+#H destruct /2 width=1 by/
+qed-.
+
+(* Basic_1: was just: sn3_gen_flat *)
+lemma csx_fwd_flat_dx: ∀h,o,I,G,L,V,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ⓕ{I}V.T → ⦃G, L⦄ ⊢ ⬊*[h, o] T.
+/2 width=5 by csx_fwd_flat_dx_aux/ qed-.
+
+lemma csx_fwd_bind: ∀h,o,a,I,G,L,V,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ⓑ{a,I}V.T →
+ ⦃G, L⦄ ⊢ ⬊*[h, o] V ∧ ⦃G, L.ⓑ{I}V⦄ ⊢ ⬊*[h, o] T.
+/3 width=3 by csx_fwd_pair_sn, csx_fwd_bind_dx, conj/ qed-.
+
+lemma csx_fwd_flat: ∀h,o,I,G,L,V,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ⓕ{I}V.T →
+ ⦃G, L⦄ ⊢ ⬊*[h, o] V ∧ ⦃G, L⦄ ⊢ ⬊*[h, o] T.
+/3 width=3 by csx_fwd_pair_sn, csx_fwd_flat_dx, conj/ qed-.
+
+(* Basic_1: removed theorems 14:
+ sn3_cdelta
+ sn3_gen_cflat sn3_cflat sn3_cpr3_trans sn3_shift sn3_change
+ sn3_appl_cast sn3_appl_beta sn3_appl_lref sn3_appl_abbr
+ sn3_appl_appls sn3_bind sn3_appl_bind sn3_appls_bind
+*)
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/computation/gcp_aaa.ma".
+include "basic_2/computation/cpxs_aaa.ma".
+include "basic_2/computation/csx_tsts_vector.ma".
+
+(* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERMS ********************)
+
+(* Main properties on atomic arity assignment *******************************)
+
+theorem aaa_csx: ∀h,o,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ⦃G, L⦄ ⊢ ⬊*[h, o] T.
+#h #o #G #L #T #A #H
+@(gcr_aaa … (csx_gcp h o) (csx_gcr h o) … H)
+qed.
+
+(* Advanced eliminators *****************************************************)
+
+fact aaa_ind_csx_aux: ∀h,o,G,L,A. ∀R:predicate term.
+ (∀T1. ⦃G, L⦄ ⊢ T1 ⁝ A →
+ (∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → (T1 = T2 → ⊥) → R T2) → R T1
+ ) →
+ ∀T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → ⦃G, L⦄ ⊢ T ⁝ A → R T.
+#h #o #G #L #A #R #IH #T #H @(csx_ind … H) -T /4 width=5 by cpx_aaa_conf/
+qed-.
+
+lemma aaa_ind_csx: ∀h,o,G,L,A. ∀R:predicate term.
+ (∀T1. ⦃G, L⦄ ⊢ T1 ⁝ A →
+ (∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → (T1 = T2 → ⊥) → R T2) → R T1
+ ) →
+ ∀T. ⦃G, L⦄ ⊢ T ⁝ A → R T.
+/5 width=9 by aaa_ind_csx_aux, aaa_csx/ qed-.
+
+fact aaa_ind_csx_alt_aux: ∀h,o,G,L,A. ∀R:predicate term.
+ (∀T1. ⦃G, L⦄ ⊢ T1 ⁝ A →
+ (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → (T1 = T2 → ⊥) → R T2) → R T1
+ ) →
+ ∀T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → ⦃G, L⦄ ⊢ T ⁝ A → R T.
+#h #o #G #L #A #R #IH #T #H @(csx_ind_alt … H) -T /4 width=5 by cpxs_aaa_conf/
+qed-.
+
+lemma aaa_ind_csx_alt: ∀h,o,G,L,A. ∀R:predicate term.
+ (∀T1. ⦃G, L⦄ ⊢ T1 ⁝ A →
+ (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → (T1 = T2 → ⊥) → R T2) → R T1
+ ) →
+ ∀T. ⦃G, L⦄ ⊢ T ⁝ A → R T.
+/5 width=9 by aaa_ind_csx_alt_aux, aaa_csx/ qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/snalt_5.ma".
+include "basic_2/computation/cpxs.ma".
+include "basic_2/computation/csx.ma".
+
+(* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERMS ********************)
+
+(* alternative definition of csx *)
+definition csxa: ∀h. sd h → relation3 genv lenv term ≝
+ λh,o,G,L. SN … (cpxs h o G L) (eq …).
+
+interpretation
+ "context-sensitive extended strong normalization (term) alternative"
+ 'SNAlt h o G L T = (csxa h o G L T).
+
+(* Basic eliminators ********************************************************)
+
+lemma csxa_ind: ∀h,o,G,L. ∀R:predicate term.
+ (∀T1. ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T1 →
+ (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → (T1 = T2 → ⊥) → R T2) → R T1
+ ) →
+ ∀T. ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T → R T.
+#h #o #G #L #R #H0 #T1 #H elim H -T1 /5 width=1 by SN_intro/
+qed-.
+
+(* Basic properties *********************************************************)
+
+lemma csx_intro_cpxs: ∀h,o,G,L,T1.
+ (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → (T1 = T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊*[h, o] T2) →
+ ⦃G, L⦄ ⊢ ⬊*[h, o] T1.
+/4 width=1 by cpx_cpxs, csx_intro/ qed.
+
+(* Basic_1: was just: sn3_intro *)
+lemma csxa_intro: ∀h,o,G,L,T1.
+ (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → (T1 = T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T2) →
+ ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T1.
+/4 width=1 by SN_intro/ qed.
+
+fact csxa_intro_aux: ∀h,o,G,L,T1. (
+ ∀T,T2. ⦃G, L⦄ ⊢ T ➡*[h, o] T2 → T1 = T → (T1 = T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T2
+ ) → ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T1.
+/4 width=3 by csxa_intro/ qed-.
+
+(* Basic_1: was just: sn3_pr3_trans (old version) *)
+lemma csxa_cpxs_trans: ∀h,o,G,L,T1. ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T1 →
+ ∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T2.
+#h #o #G #L #T1 #H elim H -T1 #T1 #HT1 #IHT1 #T2 #HLT12
+@csxa_intro #T #HLT2 #HT2
+elim (eq_term_dec T1 T2) #HT12
+[ -IHT1 -HLT12 destruct /3 width=1 by/
+| -HT1 -HT2 /3 width=4 by/
+qed.
+
+(* Basic_1: was just: sn3_pr2_intro (old version) *)
+lemma csxa_intro_cpx: ∀h,o,G,L,T1. (
+ ∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → (T1 = T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T2
+ ) → ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T1.
+#h #o #G #L #T1 #H
+@csxa_intro_aux #T #T2 #H @(cpxs_ind_dx … H) -T
+[ -H #H destruct #H
+ elim H //
+| #T0 #T #HLT1 #HLT2 #IHT #HT10 #HT12 destruct
+ elim (eq_term_dec T0 T) #HT0
+ [ -HLT1 -HLT2 -H /3 width=1 by/
+ | -IHT -HT12 /4 width=3 by csxa_cpxs_trans/
+ ]
+]
+qed.
+
+(* Main properties **********************************************************)
+
+theorem csx_csxa: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T.
+#h #o #G #L #T #H @(csx_ind … H) -T /4 width=1 by csxa_intro_cpx/
+qed.
+
+theorem csxa_csx: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T → ⦃G, L⦄ ⊢ ⬊*[h, o] T.
+#h #o #G #L #T #H @(csxa_ind … H) -T /4 width=1 by cpx_cpxs, csx_intro/
+qed.
+
+(* Basic_1: was just: sn3_pr3_trans *)
+lemma csx_cpxs_trans: ∀h,o,G,L,T1. ⦃G, L⦄ ⊢ ⬊*[h, o] T1 →
+ ∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → ⦃G, L⦄ ⊢ ⬊*[h, o] T2.
+#h #o #G #L #T1 #HT1 #T2 #H @(cpxs_ind … H) -T2 /2 width=3 by csx_cpx_trans/
+qed-.
+
+(* Main eliminators *********************************************************)
+
+lemma csx_ind_alt: ∀h,o,G,L. ∀R:predicate term.
+ (∀T1. ⦃G, L⦄ ⊢ ⬊*[h, o] T1 →
+ (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → (T1 = T2 → ⊥) → R T2) → R T1
+ ) →
+ ∀T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → R T.
+#h #o #G #L #R #H0 #T1 #H @(csxa_ind … (csx_csxa … H)) -T1 /4 width=1 by csxa_csx/
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/computation/fpbs.ma".
+include "basic_2/computation/csx_lleq.ma".
+include "basic_2/computation/csx_lpx.ma".
+
+(* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERMS ********************)
+
+(* Advanced properties ******************************************************)
+
+lemma csx_fpb_conf: ∀h,o,G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬊*[h, o] T1 →
+ ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ → ⦃G2, L2⦄ ⊢ ⬊*[h, o] T2.
+#h #o #G1 #L1 #T1 #HT1 #G2 #L2 #T2 *
+/2 width=5 by csx_cpx_trans, csx_fquq_conf, csx_lpx_conf, csx_lleq_conf/
+qed-.
+
+lemma csx_fpbs_conf: ∀h,o,G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬊*[h, o] T1 →
+ ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ → ⦃G2, L2⦄ ⊢ ⬊*[h, o] T2.
+#h #o #G1 #L1 #T1 #HT1 #G2 #L2 #T2 #H @(fpbs_ind … H) -G2 -L2 -T2
+/2 width=5 by csx_fpb_conf/
+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/cnx_lift.ma".
+include "basic_2/computation/gcp.ma".
+include "basic_2/computation/csx.ma".
+
+(* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERMS ********************)
+
+(* Relocation properties ****************************************************)
+
+(* Basic_1: was just: sn3_lift *)
+lemma csx_lift: ∀h,o,G,L2,L1,T1,c,l,k. ⦃G, L1⦄ ⊢ ⬊*[h, o] T1 →
+ ∀T2. ⬇[c, l, k] L2 ≡ L1 → ⬆[l, k] T1 ≡ T2 → ⦃G, L2⦄ ⊢ ⬊*[h, o] T2.
+#h #o #G #L2 #L1 #T1 #c #l #k #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
+>(lift_mono … HT0 … HT12) in HT2; -T1 /2 width=1 by/
+qed.
+
+(* Basic_1: was just: sn3_gen_lift *)
+lemma csx_inv_lift: ∀h,o,G,L2,L1,T1,c,l,k. ⦃G, L1⦄ ⊢ ⬊*[h, o] T1 →
+ ∀T2. ⬇[c, l, k] L1 ≡ L2 → ⬆[l, k] T2 ≡ T1 → ⦃G, L2⦄ ⊢ ⬊*[h, o] T2.
+#h #o #G #L2 #L1 #T1 #c #l #k #H elim H -T1 #T1 #_ #IHT1 #T2 #HL12 #HT21
+@csx_intro #T #HLT2 #HT2
+elim (lift_total T l k) #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/
+qed.
+
+(* Advanced inversion lemmas ************************************************)
+
+(* Basic_1: was: sn3_gen_def *)
+lemma csx_inv_lref_bind: ∀h,o,I,G,L,K,V,i. ⬇[i] L ≡ K.ⓑ{I}V →
+ ⦃G, L⦄ ⊢ ⬊*[h, o] #i → ⦃G, K⦄ ⊢ ⬊*[h, o] V.
+#h #o #I #G #L #K #V #i #HLK #Hi
+elim (lift_total V 0 (i+1))
+/4 width=9 by csx_inv_lift, csx_cpx_trans, cpx_delta, drop_fwd_drop2/
+qed-.
+
+(* Advanced properties ******************************************************)
+
+(* Basic_1: was just: sn3_abbr *)
+lemma csx_lref_bind: ∀h,o,I,G,L,K,V,i. ⬇[i] L ≡ K.ⓑ{I}V → ⦃G, K⦄ ⊢ ⬊*[h, o] V → ⦃G, L⦄ ⊢ ⬊*[h, o] #i.
+#h #o #I #G #L #K #V #i #HLK #HV
+@csx_intro #X #H #Hi
+elim (cpx_inv_lref1 … H) -H
+[ #H destruct elim Hi //
+| -Hi * #I0 #K0 #V0 #V1 #HLK0 #HV01 #HV1
+ lapply (drop_mono … HLK0 … HLK) -HLK #H destruct
+ /3 width=8 by csx_lift, csx_cpx_trans, drop_fwd_drop2/
+]
+qed.
+
+lemma csx_appl_simple: ∀h,o,G,L,V. ⦃G, L⦄ ⊢ ⬊*[h, o] V → ∀T1.
+ (∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → (T1 = T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓐV.T2) →
+ 𝐒⦃T1⦄ → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓐV.T1.
+#h #o #G #L #V #H @(csx_ind … H) -V #V #_ #IHV #T1 #IHT1 #HT1
+@csx_intro #X #H1 #H2
+elim (cpx_inv_appl1_simple … H1) // -H1
+#V0 #T0 #HLV0 #HLT10 #H destruct
+elim (eq_false_inv_tpair_dx … H2) -H2
+[ -IHV -HT1 /4 width=3 by csx_cpx_trans, cpx_pair_sn/
+| -HLT10 * #H #HV0 destruct
+ @IHV /4 width=3 by csx_cpx_trans, cpx_pair_sn/ (**) (* full auto 17s *)
+]
+qed.
+
+lemma csx_fqu_conf: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+ ⦃G1, L1⦄ ⊢ ⬊*[h, o] T1 → ⦃G2, L2⦄ ⊢ ⬊*[h, o] T2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
+/2 width=8 by csx_inv_lref_bind, csx_inv_lift, csx_fwd_flat_dx, csx_fwd_bind_dx, csx_fwd_pair_sn/
+qed-.
+
+lemma csx_fquq_conf: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
+ ⦃G1, L1⦄ ⊢ ⬊*[h, o] T1 → ⦃G2, L2⦄ ⊢ ⬊*[h, o] T2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H12 #H elim (fquq_inv_gen … H12) -H12
+[ /2 width=5 by csx_fqu_conf/
+| * #HG #HL #HT destruct //
+]
+qed-.
+
+lemma csx_fqup_conf: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
+ ⦃G1, L1⦄ ⊢ ⬊*[h, o] T1 → ⦃G2, L2⦄ ⊢ ⬊*[h, o] T2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
+/3 width=5 by csx_fqu_conf/
+qed-.
+
+lemma csx_fqus_conf: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
+ ⦃G1, L1⦄ ⊢ ⬊*[h, o] T1 → ⦃G2, L2⦄ ⊢ ⬊*[h, o] T2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H12 #H elim (fqus_inv_gen … H12) -H12
+[ /2 width=5 by csx_fqup_conf/
+| * #HG #HL #HT destruct //
+]
+qed-.
+
+(* Main properties **********************************************************)
+
+theorem csx_gcp: ∀h,o. gcp (cpx h o) (eq …) (csx h o).
+#h #o @mk_gcp
+[ normalize /3 width=13 by cnx_lift/
+| #G #L elim (deg_total h o 0) /3 width=8 by cnx_sort_iter, ex_intro/
+| /2 width=8 by csx_lift/
+| /2 width=3 by csx_fwd_flat_dx/
+]
+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_lleq.ma".
+include "basic_2/computation/csx.ma".
+
+(* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERMS ********************)
+
+(* Properties on lazy equivalence for local environments ********************)
+
+lemma csx_lleq_conf: ∀h,o,G,L1,T. ⦃G, L1⦄ ⊢ ⬊*[h, o] T →
+ ∀L2. L1 ≡[T, 0] L2 → ⦃G, L2⦄ ⊢ ⬊*[h, o] T.
+#h #o #G #L1 #T #H @(csx_ind … H) -T
+/4 width=6 by csx_intro, cpx_lleq_conf_dx, lleq_cpx_trans/
+qed-.
+
+lemma csx_lleq_trans: ∀h,o,G,L1,L2,T.
+ L1 ≡[T, 0] L2 → ⦃G, L2⦄ ⊢ ⬊*[h, o] T → ⦃G, L1⦄ ⊢ ⬊*[h, o] T.
+/3 width=3 by csx_lleq_conf, 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/grammar/tsts_tsts.ma".
+include "basic_2/computation/cpxs_cpxs.ma".
+include "basic_2/computation/csx_alt.ma".
+include "basic_2/computation/csx_lift.ma".
+
+(* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERMS ********************)
+
+(* Advanced properties ******************************************************)
+
+lemma csx_lpx_conf: ∀h,o,G,L1,L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 →
+ ∀T. ⦃G, L1⦄ ⊢ ⬊*[h, o] T → ⦃G, L2⦄ ⊢ ⬊*[h, o] T.
+#h #o #G #L1 #L2 #HL12 #T #H @(csx_ind_alt … H) -T
+/4 width=3 by csx_intro, lpx_cpx_trans/
+qed-.
+
+lemma csx_abst: ∀h,o,a,G,L,W. ⦃G, L⦄ ⊢ ⬊*[h, o] W →
+ ∀T. ⦃G, L.ⓛW⦄ ⊢ ⬊*[h, o] T → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓛ{a}W.T.
+#h #o #a #G #L #W #HW @(csx_ind … HW) -W #W #_ #IHW #T #HT @(csx_ind … HT) -T #T #HT #IHT
+@csx_intro #X #H1 #H2
+elim (cpx_inv_abst1 … H1) -H1
+#W0 #T0 #HLW0 #HLT0 #H destruct
+elim (eq_false_inv_tpair_sn … H2) -H2
+[ -IHT #H lapply (csx_cpx_trans … HLT0) // -HT
+ #HT0 lapply (csx_lpx_conf … (L.ⓛW0) … HT0) -HT0 /3 width=1 by lpx_pair/
+| -IHW -HLW0 -HT * #H destruct /3 width=1 by/
+]
+qed.
+
+lemma csx_abbr: ∀h,o,a,G,L,V. ⦃G, L⦄ ⊢ ⬊*[h, o] V →
+ ∀T. ⦃G, L.ⓓV⦄ ⊢ ⬊*[h, o] T → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓓ{a}V. T.
+#h #o #a #G #L #V #HV elim HV -V #V #_ #IHV #T #HT @(csx_ind_alt … HT) -T #T #HT #IHT
+@csx_intro #X #H1 #H2
+elim (cpx_inv_abbr1 … H1) -H1 *
+[ #V1 #T1 #HLV1 #HLT1 #H destruct
+ elim (eq_false_inv_tpair_sn … H2) -H2
+ [ /4 width=5 by csx_cpx_trans, csx_lpx_conf, lpx_pair/
+ | -IHV -HLV1 * #H destruct /3 width=1 by cpx_cpxs/
+ ]
+| -IHV -IHT -H2
+ /3 width=8 by csx_cpx_trans, csx_inv_lift, drop_drop/
+]
+qed.
+
+fact csx_appl_beta_aux: ∀h,o,a,G,L,U1. ⦃G, L⦄ ⊢ ⬊*[h, o] U1 →
+ ∀V,W,T1. U1 = ⓓ{a}ⓝW.V.T1 → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓐV.ⓛ{a}W.T1.
+#h #o #a #G #L #X #H @(csx_ind … H) -X
+#X #HT1 #IHT1 #V #W #T1 #H1 destruct
+@csx_intro #X #H1 #H2
+elim (cpx_inv_appl1 … H1) -H1 *
+[ -HT1 #V0 #Y #HLV0 #H #H0 destruct
+ elim (cpx_inv_abst1 … H) -H #W0 #T0 #HLW0 #HLT0 #H destruct
+ @IHT1 -IHT1 [4: // | skip |3: #H destruct /2 width=1 by/ ] -H2
+ lapply (lsubr_cpx_trans … HLT0 (L.ⓓⓝW.V) ?) -HLT0 /3 width=1 by cpx_bind, cpx_flat, lsubr_beta/
+| -IHT1 -H2 #b #V0 #W0 #W2 #T0 #T2 #HLV0 #HLW02 #HLT02 #H1 #H3 destruct
+ lapply (lsubr_cpx_trans … HLT02 (L.ⓓⓝW0.V) ?) -HLT02
+ /4 width=5 by csx_cpx_trans, cpx_bind, cpx_flat, lsubr_beta/
+| -HT1 -IHT1 -H2 #b #V0 #V1 #W0 #W1 #T0 #T3 #_ #_ #_ #_ #H destruct
+]
+qed-.
+
+(* Basic_1: was just: sn3_beta *)
+lemma csx_appl_beta: ∀h,o,a,G,L,V,W,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ⓓ{a}ⓝW.V.T → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓐV.ⓛ{a}W.T.
+/2 width=3 by csx_appl_beta_aux/ qed.
+
+fact csx_appl_theta_aux: ∀h,o,a,G,L,U. ⦃G, L⦄ ⊢ ⬊*[h, o] U → ∀V1,V2. ⬆[0, 1] V1 ≡ V2 →
+ ∀V,T. U = ⓓ{a}V.ⓐV2.T → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓐV1.ⓓ{a}V.T.
+#h #o #a #G #L #X #H @(csx_ind_alt … H) -X #X #HVT #IHVT #V1 #V2 #HV12 #V #T #H destruct
+lapply (csx_fwd_pair_sn … HVT) #HV
+lapply (csx_fwd_bind_dx … HVT) -HVT #HVT
+@csx_intro #X #HL #H
+elim (cpx_inv_appl1 … HL) -HL *
+[ -HV #V0 #Y #HLV10 #HL #H0 destruct
+ elim (cpx_inv_abbr1 … HL) -HL *
+ [ #V3 #T3 #HV3 #HLT3 #H0 destruct
+ elim (lift_total V0 0 1) #V4 #HV04
+ elim (eq_term_dec (ⓓ{a}V.ⓐV2.T) (ⓓ{a}V3.ⓐV4.T3))
+ [ -IHVT #H0 destruct
+ elim (eq_false_inv_tpair_sn … H) -H
+ [ -HLV10 -HV3 -HLT3 -HVT
+ >(lift_inj … HV12 … HV04) -V4
+ #H elim H //
+ | * #_ #H elim H //
+ ]
+ | -H -HVT #H
+ lapply (cpx_lift … HLV10 (L.ⓓV) (Ⓕ) … HV12 … HV04) -HLV10 -HV12 /2 width=1 by drop_drop/ #HV24
+ @(IHVT … H … HV04) -IHVT /4 width=1 by cpx_cpxs, cpx_bind, cpx_flat/
+ ]
+ | -H -IHVT #T0 #HLT0 #HT0 #H0 destruct
+ lapply (csx_cpx_trans … HVT (ⓐV2.T0) ?) /2 width=1 by cpx_flat/ -T #HVT0
+ lapply (csx_inv_lift … L … (Ⓕ) … 1 HVT0 ? ? ?) -HVT0
+ /3 width=5 by csx_cpx_trans, cpx_pair_sn, drop_drop, lift_flat/
+ ]
+| -HV -HV12 -HVT -IHVT -H #b #V0 #W0 #W1 #T0 #T1 #_ #_ #_ #H destruct
+| -IHVT -H #b #V0 #V3 #W0 #W1 #T0 #T1 #HLV10 #HV03 #HLW01 #HLT01 #H1 #H2 destruct
+ lapply (cpx_lift … HLV10 (L. ⓓW0) … HV12 … HV03) -HLV10 -HV12 -HV03 /2 width=2 by drop_drop/ #HLV23
+ @csx_abbr /2 width=3 by csx_cpx_trans/ -HV
+ @(csx_lpx_conf … (L.ⓓW0)) /2 width=1 by lpx_pair/ -W1
+ /4 width=5 by csx_cpxs_trans, cpx_cpxs, cpx_flat/
+]
+qed-.
+
+lemma csx_appl_theta: ∀h,o,a,V1,V2. ⬆[0, 1] V1 ≡ V2 →
+ ∀G,L,V,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ⓓ{a}V.ⓐV2.T → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓐV1.ⓓ{a}V.T.
+/2 width=5 by csx_appl_theta_aux/ qed.
+
+(* Basic_1: was just: sn3_appl_appl *)
+lemma csx_appl_simple_tsts: ∀h,o,G,L,V. ⦃G, L⦄ ⊢ ⬊*[h, o] V → ∀T1. ⦃G, L⦄ ⊢ ⬊*[h, o] T1 →
+ (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → (T1 ≂ T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓐV.T2) →
+ 𝐒⦃T1⦄ → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓐV.T1.
+#h #o #G #L #V #H @(csx_ind … H) -V #V #_ #IHV #T1 #H @(csx_ind … H) -T1 #T1 #H1T1 #IHT1 #H2T1 #H3T1
+@csx_intro #X #HL #H
+elim (cpx_inv_appl1_simple … HL) -HL //
+#V0 #T0 #HLV0 #HLT10 #H0 destruct
+elim (eq_false_inv_tpair_sn … H) -H
+[ -IHT1 #HV0
+ @(csx_cpx_trans … (ⓐV0.T1)) /2 width=1 by cpx_flat/ -HLT10
+ @IHV -IHV /4 width=3 by csx_cpx_trans, cpx_pair_sn/
+| -IHV -H1T1 -HLV0 * #H #H1T10 destruct
+ elim (tsts_dec T1 T0) #H2T10
+ [ @IHT1 -IHT1 /4 width=3 by cpxs_strap2, cpxs_strap1, tsts_canc_sn, simple_tsts_repl_dx/
+ | -IHT1 -H3T1 -H1T10 /3 width=1 by cpx_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/computation/csx_lpx.ma".
+include "basic_2/computation/lpxs.ma".
+
+(* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERMS ********************)
+
+(* Properties on sn extended parallel computation for local environments ****)
+
+lemma csx_lpxs_conf: ∀h,o,G,L1,L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 →
+ ∀T. ⦃G, L1⦄ ⊢ ⬊*[h, o] T → ⦃G, L2⦄ ⊢ ⬊*[h, o] T.
+#h #o #G #L1 #L2 #H @(lpxs_ind … H) -L2 /3 by lpxs_strap1, csx_lpx_conf/
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/computation/gcp_cr.ma".
+include "basic_2/computation/cpxs_tsts_vector.ma".
+include "basic_2/computation/csx_lpx.ma".
+include "basic_2/computation/csx_vector.ma".
+
+(* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERM VECTORS *************)
+
+(* Advanced properties ******************************************************)
+
+(* Basic_1: was just: sn3_appls_lref *)
+lemma csx_applv_cnx: ∀h,o,G,L,T. 𝐒⦃T⦄ → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ →
+ ∀Vs. ⦃G, L⦄ ⊢ ⬊*[h, o] Vs → ⦃G, L⦄ ⊢ ⬊*[h, o] ⒶVs.T.
+#h #o #G #L #T #H1T #H2T #Vs elim Vs -Vs [ #_ @(cnx_csx … H2T) ] (**) (* /2 width=1/ does not work *)
+#V #Vs #IHV #H
+elim (csxv_inv_cons … H) -H #HV #HVs
+@csx_appl_simple_tsts /2 width=1 by applv_simple/ -IHV -HV -HVs
+#X #H #H0
+lapply (cpxs_fwd_cnx_vector … H) -H // -H1T -H2T #H
+elim (H0) -H0 //
+qed.
+
+lemma csx_applv_sort: ∀h,o,G,L,s,Vs. ⦃G, L⦄ ⊢ ⬊*[h, o] Vs → ⦃G, L⦄ ⊢ ⬊*[h, o] ⒶVs.⋆s.
+#h #o #G #L #s elim (deg_total h o s)
+#d generalize in match s; -s @(nat_ind_plus … d) -d [ /3 width=6 by csx_applv_cnx, cnx_sort, simple_atom/ ]
+#d #IHd #s #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
+#X #H #H0
+elim (cpxs_fwd_sort_vector … H) -H #H
+[ elim H0 -H0 //
+| -H0 @(csx_cpxs_trans … (Ⓐ(V@Vs).⋆(next h s))) /2 width=1 by cpxs_flat_dx/
+]
+qed.
+
+(* Basic_1: was just: sn3_appls_beta *)
+lemma csx_applv_beta: ∀h,o,a,G,L,Vs,V,W,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ⒶVs.ⓓ{a}ⓝW.V.T →
+ ⦃G, L⦄ ⊢ ⬊*[h, o] ⒶVs. ⓐV.ⓛ{a}W.T.
+#h #o #a #G #L #Vs elim Vs -Vs /2 width=1 by csx_appl_beta/
+#V0 #Vs #IHV #V #W #T #H1T
+lapply (csx_fwd_pair_sn … H1T) #HV0
+lapply (csx_fwd_flat_dx … H1T) #H2T
+@csx_appl_simple_tsts /2 width=1 by applv_simple, simple_flat/ -IHV -HV0 -H2T
+#X #H #H0
+elim (cpxs_fwd_beta_vector … H) -H #H
+[ -H1T elim H0 -H0 //
+| -H0 /3 width=5 by csx_cpxs_trans, cpxs_flat_dx/
+]
+qed.
+
+lemma csx_applv_delta: ∀h,o,I,G,L,K,V1,i. ⬇[i] L ≡ K.ⓑ{I}V1 →
+ ∀V2. ⬆[0, i + 1] V1 ≡ V2 →
+ ∀Vs. ⦃G, L⦄ ⊢ ⬊*[h, o] (ⒶVs.V2) → ⦃G, L⦄ ⊢ ⬊*[h, o] (ⒶVs.#i).
+#h #o #I #G #L #K #V1 #i #HLK #V2 #HV12 #Vs elim Vs -Vs
+[ /4 width=12 by csx_inv_lift, csx_lref_bind, drop_fwd_drop2/
+| #V #Vs #IHV #H1T
+ lapply (csx_fwd_pair_sn … H1T) #HV
+ lapply (csx_fwd_flat_dx … H1T) #H2T
+ @csx_appl_simple_tsts /2 width=1 by applv_simple, simple_atom/ -IHV -HV -H2T
+ #X #H #H0
+ elim (cpxs_fwd_delta_vector … HLK … HV12 … H) -HLK -HV12 -H #H
+ [ -H1T elim H0 -H0 //
+ | -H0 /3 width=5 by csx_cpxs_trans, cpxs_flat_dx/
+ ]
+]
+qed.
+
+(* Basic_1: was just: sn3_appls_abbr *)
+lemma csx_applv_theta: ∀h,o,a,G,L,V1c,V2c. ⬆[0, 1] V1c ≡ V2c →
+ ∀V,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ⓓ{a}V.ⒶV2c.T →
+ ⦃G, L⦄ ⊢ ⬊*[h, o] ⒶV1c.ⓓ{a}V.T.
+#h #o #a #G #L #V1c #V2c * -V1c -V2c /2 width=1 by/
+#V1c #V2c #V1 #V2 #HV12 #H
+generalize in match HV12; -HV12 generalize in match V2; -V2 generalize in match V1; -V1
+elim H -V1c -V2c /2 width=3 by csx_appl_theta/
+#V1c #V2c #V1 #V2 #HV12 #HV12c #IHV12c #W1 #W2 #HW12 #V #T #H
+lapply (csx_appl_theta … HW12 … H) -H -HW12 #H
+lapply (csx_fwd_pair_sn … H) #HW1
+lapply (csx_fwd_flat_dx … H) #H1
+@csx_appl_simple_tsts /2 width=3 by simple_flat/ -IHV12c -HW1 -H1 #X #H1 #H2
+elim (cpxs_fwd_theta_vector … (V2@V2c) … H1) -H1 /2 width=1 by liftv_cons/ -HV12c -HV12
+[ -H #H elim H2 -H2 //
+| -H2 /3 width=5 by csx_cpxs_trans, cpxs_flat_dx/
+]
+qed.
+
+(* Basic_1: was just: sn3_appls_cast *)
+lemma csx_applv_cast: ∀h,o,G,L,Vs,W,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ⒶVs.W → ⦃G, L⦄ ⊢ ⬊*[h, o] ⒶVs.T →
+ ⦃G, L⦄ ⊢ ⬊*[h, o] ⒶVs.ⓝW.T.
+#h #o #G #L #Vs elim Vs -Vs /2 width=1 by csx_cast/
+#V #Vs #IHV #W #T #H1W #H1T
+lapply (csx_fwd_pair_sn … H1W) #HV
+lapply (csx_fwd_flat_dx … H1W) #H2W
+lapply (csx_fwd_flat_dx … H1T) #H2T
+@csx_appl_simple_tsts /2 width=1 by applv_simple, simple_flat/ -IHV -HV -H2W -H2T
+#X #H #H0
+elim (cpxs_fwd_cast_vector … H) -H #H
+[ -H1W -H1T elim H0 -H0 //
+| -H1W -H0 /3 width=5 by csx_cpxs_trans, cpxs_flat_dx/
+| -H1T -H0 /3 width=5 by csx_cpxs_trans, cpxs_flat_dx/
+]
+qed.
+
+theorem csx_gcr: ∀h,o. gcr (cpx h o) (eq …) (csx h o) (csx h o).
+#h #o @mk_gcr //
+[ /3 width=1 by csx_applv_cnx/
+|2,3,6: /2 width=1 by csx_applv_beta, csx_applv_sort, csx_applv_cast/
+| /2 width=7 by csx_applv_delta/
+| #G #L #V1c #V2c #HV12c #a #V #T #H #HV
+ @(csx_applv_theta … HV12c) -HV12c
+ @csx_abbr //
+]
+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/term_vector.ma".
+include "basic_2/computation/csx.ma".
+
+(* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERM VECTORS *************)
+
+definition csxv: ∀h. sd h → relation3 genv lenv (list term) ≝
+ λh,o,G,L. all … (csx h o G L).
+
+interpretation
+ "context-sensitive strong normalization (term vector)"
+ 'SN h o G L Ts = (csxv h o G L Ts).
+
+(* Basic inversion lemmas ***************************************************)
+
+lemma csxv_inv_cons: ∀h,o,G,L,T,Ts. ⦃G, L⦄ ⊢ ⬊*[h, o] T @ Ts →
+ ⦃G, L⦄ ⊢ ⬊*[h, o] T ∧ ⦃G, L⦄ ⊢ ⬊*[h, o] Ts.
+normalize // qed-.
+
+(* Basic forward lemmas *****************************************************)
+
+lemma csx_fwd_applv: ∀h,o,G,L,T,Vs. ⦃G, L⦄ ⊢ ⬊*[h, o] Ⓐ Vs.T →
+ ⦃G, L⦄ ⊢ ⬊*[h, o] Vs ∧ ⦃G, L⦄ ⊢ ⬊*[h, o] T.
+#h #o #G #L #T #Vs elim Vs -Vs /2 width=1 by conj/
+#V #Vs #IHVs #HVs
+lapply (csx_fwd_pair_sn … HVs) #HV
+lapply (csx_fwd_flat_dx … HVs) -HVs #HVs
+elim (IHVs HVs) -IHVs -HVs /3 width=1 by conj/
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/lazybtpredstarproper_8.ma".
+include "basic_2/reduction/fpb.ma".
+include "basic_2/computation/fpbs.ma".
+
+(* "QRST" PROPER PARALLEL COMPUTATION FOR CLOSURES **************************)
+
+definition fpbg: ∀h. sd h → tri_relation genv lenv term ≝
+ λh,o,G1,L1,T1,G2,L2,T2.
+ ∃∃G,L,T. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G, L, T⦄ & ⦃G, L, T⦄ ≥[h, o] ⦃G2, L2, T2⦄.
+
+interpretation "'qrst' proper parallel computation (closure)"
+ 'LazyBTPRedStarProper h o G1 L1 T1 G2 L2 T2 = (fpbg h o G1 L1 T1 G2 L2 T2).
+
+(* Basic properties *********************************************************)
+
+lemma fpb_fpbg: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄.
+/2 width=5 by ex2_3_intro/ qed.
+
+lemma fpbg_fpbq_trans: ∀h,o,G1,G,G2,L1,L,L2,T1,T,T2.
+ ⦃G1, L1, T1⦄ >≡[h, o] ⦃G, L, T⦄ → ⦃G, L, T⦄ ≽[h, o] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G #G2 #L1 #L #L2 #T1 #T #T2 *
+/3 width=9 by fpbs_strap1, 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 "basic_2/multiple/fleq_fleq.ma".
+include "basic_2/reduction/fpbq_alt.ma".
+include "basic_2/computation/fpbg.ma".
+
+(* "QRST" PROPER PARALLEL COMPUTATION FOR CLOSURES **************************)
+
+(* Properties on lazy equivalence for closures ******************************)
+
+lemma fpbg_fleq_trans: ∀h,o,G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ >≡[h, o] ⦃G, L, T⦄ →
+ ∀G2,L2,T2. ⦃G, L, T⦄ ≡[0] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄.
+/3 width=5 by fpbg_fpbq_trans, fleq_fpbq/ qed-.
+
+lemma fleq_fpbg_trans: ∀h,o,G,G2,L,L2,T,T2. ⦃G, L, T⦄ >≡[h, o] ⦃G2, L2, T2⦄ →
+ ∀G1,L1,T1. ⦃G1, L1, T1⦄ ≡[0] ⦃G, L, T⦄ → ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄.
+#h #o #G #G2 #L #L2 #T #T2 * #G0 #L0 #T0 #H0 #H02 #G1 #L1 #T1 #H1
+elim (fleq_fpb_trans … H1 … H0) -G -L -T
+/4 width=9 by fpbs_strap2, fleq_fpbq, ex2_3_intro/
+qed-.
+
+(* alternative definition of fpbs *******************************************)
+
+lemma fleq_fpbs: ∀h,o,G1,G2,L1,L2,T1,T2.
+ ⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 * /2 width=1 by lleq_fpbs/
+qed.
+
+lemma fpbg_fwd_fpbs: ∀h,o,G1,G2,L1,L2,T1,T2.
+ ⦃G1, L1, T1⦄ >≡[h,o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 *
+/3 width=5 by fpbs_strap2, fpb_fpbq/
+qed-.
+
+lemma fpbs_fpbg: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ ∨
+ ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fpbs_ind … H) -G2 -L2 -T2
+[ /2 width=1 by or_introl/
+| #G #G2 #L #L2 #T #T2 #_ #H2 * #H1 @(fpbq_ind_alt … H2) -H2 #H2
+ [ /3 width=5 by fleq_trans, or_introl/
+ | elim (fleq_fpb_trans … H1 … H2) -G -L -T
+ /4 width=5 by ex2_3_intro, or_intror, fleq_fpbs/
+ | /3 width=5 by fpbg_fleq_trans, or_intror/
+ | /4 width=5 by fpbg_fpbq_trans, fpb_fpbq, or_intror/
+ ]
+]
+qed-.
+
+(* Advanced properties of "qrst" parallel computation on closures ***********)
+
+lemma fpbs_fpb_trans: ∀h,o,F1,F2,K1,K2,T1,T2. ⦃F1, K1, T1⦄ ≥[h, o] ⦃F2, K2, T2⦄ →
+ ∀G2,L2,U2. ⦃F2, K2, T2⦄ ≻[h, o] ⦃G2, L2, U2⦄ →
+ ∃∃G1,L1,U1. ⦃F1, K1, T1⦄ ≻[h, o] ⦃G1, L1, U1⦄ & ⦃G1, L1, U1⦄ ≥[h, o] ⦃G2, L2, U2⦄.
+#h #o #F1 #F2 #K1 #K2 #T1 #T2 #H elim (fpbs_fpbg … H) -H
+[ #H12 #G2 #L2 #U2 #H2 elim (fleq_fpb_trans … H12 … H2) -F2 -K2 -T2
+ /3 width=5 by fleq_fpbs, ex2_3_intro/
+| * #H1 #H2 #H3 #H4 #H5 #H6 #H7 #H8 #H9
+ @(ex2_3_intro … H4) -H4 /3 width=5 by fpbs_strap1, fpb_fpbq/
+]
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/computation/fpbg_fpbs.ma".
+
+(* "QRST" PROPER PARALLEL COMPUTATION FOR CLOSURES **************************)
+
+(* Main properties **********************************************************)
+
+theorem fpbg_trans: ∀h,o. tri_transitive … (fpbg h o).
+/3 width=5 by fpbg_fpbs_trans, fpbg_fwd_fpbs/ qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/computation/lpxs_lleq.ma".
+include "basic_2/computation/fpbs_lift.ma".
+include "basic_2/computation/fpbg_fleq.ma".
+
+(* "QRST" PROPER PARALLEL COMPUTATION FOR CLOSURES **************************)
+
+(* Properties on "qrst" parallel reduction on closures **********************)
+
+lemma fpb_fpbg_trans: ∀h,o,G1,G,G2,L1,L,L2,T1,T,T2.
+ ⦃G1, L1, T1⦄ ≻[h, o] ⦃G, L, T⦄ → ⦃G, L, T⦄ >≡[h, o] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄.
+/3 width=5 by fpbg_fwd_fpbs, ex2_3_intro/ qed-.
+
+lemma fpbq_fpbg_trans: ∀h,o,G1,G,G2,L1,L,L2,T1,T,T2.
+ ⦃G1, L1, T1⦄ ≽[h, o] ⦃G, L, T⦄ → ⦃G, L, T⦄ >≡[h, o] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 #H2 @(fpbq_ind_alt … H1) -H1
+/2 width=5 by fleq_fpbg_trans, fpb_fpbg_trans/
+qed-.
+
+(* Properties on "qrst" parallel compuutation on closures *******************)
+
+lemma fpbs_fpbg_trans: ∀h,o,G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L, T⦄ →
+ ∀G2,L2,T2. ⦃G, L, T⦄ >≡[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G #L1 #L #T1 #T #H @(fpbs_ind … H) -G -L -T /3 width=5 by fpbq_fpbg_trans/
+qed-.
+
+(* Note: this is used in the closure proof *)
+lemma fpbg_fpbs_trans: ∀h,o,G,G2,L,L2,T,T2. ⦃G, L, T⦄ ≥[h, o] ⦃G2, L2, T2⦄ →
+ ∀G1,L1,T1. ⦃G1, L1, T1⦄ >≡[h, o] ⦃G, L, T⦄ → ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄.
+#h #o #G #G2 #L #L2 #T #T2 #H @(fpbs_ind_dx … H) -G -L -T /3 width=5 by fpbg_fpbq_trans/
+qed-.
+
+(* Note: this is used in the closure proof *)
+lemma fqup_fpbg: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim (fqup_inv_step_sn … H) -H
+/3 width=5 by fqus_fpbs, fpb_fqu, ex2_3_intro/
+qed.
+
+lemma cpxs_fpbg: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 →
+ (T1 = T2 → ⊥) → ⦃G, L, T1⦄ >≡[h, o] ⦃G, L, T2⦄.
+#h #o #G #L #T1 #T2 #H #H0 elim (cpxs_neq_inv_step_sn … H … H0) -H -H0
+/4 width=5 by cpxs_fpbs, fpb_cpx, ex2_3_intro/
+qed.
+
+lemma lstas_fpbg: ∀h,o,G,L,T1,T2,d2. ⦃G, L⦄ ⊢ T1 •*[h, d2] T2 → (T1 = T2 → ⊥) →
+ ∀d1. d2 ≤ d1 → ⦃G, L⦄ ⊢ T1 ▪[h, o] d1 → ⦃G, L, T1⦄ >≡[h, o] ⦃G, L, T2⦄.
+/3 width=5 by lstas_cpxs, cpxs_fpbg/ qed.
+
+lemma lpxs_fpbg: ∀h,o,G,L1,L2,T. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 →
+ (L1 ≡[T, 0] L2 → ⊥) → ⦃G, L1, T⦄ >≡[h, o] ⦃G, L2, T⦄.
+#h #o #G #L1 #L2 #T #H #H0 elim (lpxs_nlleq_inv_step_sn … H … H0) -H -H0
+/4 width=5 by fpb_lpx, lpxs_lleq_fpbs, 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 "basic_2/reduction/fpb_lift.ma".
+include "basic_2/computation/fpbg.ma".
+
+(* "QRST" PARALLEL COMPUTATION FOR CLOSURES *********************************)
+
+(* Advanced properties ******************************************************)
+
+lemma sta_fpbg: ∀h,o,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 ▪[h, o] d+1 →
+ ⦃G, L⦄ ⊢ T1 •*[h, 1] T2 → ⦃G, L, T1⦄ >≡[h, o] ⦃G, L, T2⦄.
+/4 width=2 by fpb_fpbg, sta_fpb/ qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/btpredstar_8.ma".
+include "basic_2/multiple/fqus.ma".
+include "basic_2/reduction/fpbq.ma".
+include "basic_2/computation/cpxs.ma".
+include "basic_2/computation/lpxs.ma".
+
+(* "QRST" PARALLEL COMPUTATION FOR CLOSURES *********************************)
+
+definition fpbs: ∀h. sd h → tri_relation genv lenv term ≝
+ λh,o. tri_TC … (fpbq h o).
+
+interpretation "'qrst' parallel computation (closure)"
+ 'BTPRedStar h o G1 L1 T1 G2 L2 T2 = (fpbs h o G1 L1 T1 G2 L2 T2).
+
+(* Basic eliminators ********************************************************)
+
+lemma fpbs_ind: ∀h,o,G1,L1,T1. ∀R:relation3 genv lenv term. R G1 L1 T1 →
+ (∀G,G2,L,L2,T,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L, T⦄ → ⦃G, L, T⦄ ≽[h, o] ⦃G2, L2, T2⦄ → R G L T → R G2 L2 T2) →
+ ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ → R G2 L2 T2.
+/3 width=8 by tri_TC_star_ind/ qed-.
+
+lemma fpbs_ind_dx: ∀h,o,G2,L2,T2. ∀R:relation3 genv lenv term. R G2 L2 T2 →
+ (∀G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ ≽[h, o] ⦃G, L, T⦄ → ⦃G, L, T⦄ ≥[h, o] ⦃G2, L2, T2⦄ → R G L T → R G1 L1 T1) →
+ ∀G1,L1,T1. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ → R G1 L1 T1.
+/3 width=8 by tri_TC_star_ind_dx/ qed-.
+
+(* Basic properties *********************************************************)
+
+lemma fpbs_refl: ∀h,o. tri_reflexive … (fpbs h o).
+/2 width=1 by tri_inj/ qed.
+
+lemma fpbq_fpbs: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
+/2 width=1 by tri_inj/ qed.
+
+lemma fpbs_strap1: ∀h,o,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L, T⦄ →
+ ⦃G, L, T⦄ ≽[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
+/2 width=5 by tri_step/ qed-.
+
+lemma fpbs_strap2: ∀h,o,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ≽[h, o] ⦃G, L, T⦄ →
+ ⦃G, L, T⦄ ≥[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
+/2 width=5 by tri_TC_strap/ qed-.
+
+lemma fqup_fpbs: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
+/4 width=5 by fqu_fquq, fpbq_fquq, tri_step/
+qed.
+
+lemma fqus_fpbs: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqus_ind … H) -G2 -L2 -T2
+/3 width=5 by fpbq_fquq, tri_step/
+qed.
+
+lemma cpxs_fpbs: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → ⦃G, L, T1⦄ ≥[h, o] ⦃G, L, T2⦄.
+#h #o #G #L #T1 #T2 #H @(cpxs_ind … H) -T2
+/3 width=5 by fpbq_cpx, fpbs_strap1/
+qed.
+
+lemma lpxs_fpbs: ∀h,o,G,L1,L2,T. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → ⦃G, L1, T⦄ ≥[h, o] ⦃G, L2, T⦄.
+#h #o #G #L1 #L2 #T #H @(lpxs_ind … H) -L2
+/3 width=5 by fpbq_lpx, fpbs_strap1/
+qed.
+
+lemma lleq_fpbs: ∀h,o,G,L1,L2,T. L1 ≡[T, 0] L2 → ⦃G, L1, T⦄ ≥[h, o] ⦃G, L2, T⦄.
+/3 width=1 by fpbq_fpbs, fpbq_lleq/ qed.
+
+lemma cprs_fpbs: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡* T2 → ⦃G, L, T1⦄ ≥[h, o] ⦃G, L, T2⦄.
+/3 width=1 by cprs_cpxs, cpxs_fpbs/ qed.
+
+lemma lprs_fpbs: ∀h,o,G,L1,L2,T. ⦃G, L1⦄ ⊢ ➡* L2 → ⦃G, L1, T⦄ ≥[h, o] ⦃G, L2, T⦄.
+/3 width=1 by lprs_lpxs, lpxs_fpbs/ qed.
+
+lemma fpbs_fqus_trans: ∀h,o,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L, T⦄ →
+ ⦃G, L, T⦄ ⊐* ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 #H @(fqus_ind … H) -G2 -L2 -T2
+/3 width=5 by fpbs_strap1, fpbq_fquq/
+qed-.
+
+lemma fpbs_fqup_trans: ∀h,o,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L, T⦄ →
+ ⦃G, L, T⦄ ⊐+ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
+/3 width=5 by fpbs_fqus_trans, fqup_fqus/ qed-.
+
+lemma fpbs_cpxs_trans: ∀h,o,G1,G,L1,L,T1,T,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L, T⦄ →
+ ⦃G, L⦄ ⊢ T ➡*[h, o] T2 → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L, T2⦄.
+#h #o #G1 #G #L1 #L #T1 #T #T2 #H1 #H @(cpxs_ind … H) -T2
+/3 width=5 by fpbs_strap1, fpbq_cpx/
+qed-.
+
+lemma fpbs_lpxs_trans: ∀h,o,G1,G,L1,L,L2,T1,T. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L, T⦄ →
+ ⦃G, L⦄ ⊢ ➡*[h, o] L2 → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L2, T⦄.
+#h #o #G1 #G #L1 #L #L2 #T1 #T #H1 #H @(lpxs_ind … H) -L2
+/3 width=5 by fpbs_strap1, fpbq_lpx/
+qed-.
+
+lemma fpbs_lleq_trans: ∀h,o,G1,G,L1,L,L2,T1,T. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L, T⦄ →
+ L ≡[T, 0] L2 → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L2, T⦄.
+/3 width=5 by fpbs_strap1, fpbq_lleq/ qed-.
+
+lemma fqus_fpbs_trans: ∀h,o,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G, L, T⦄ ≥[h, o] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ⊐* ⦃G, L, T⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 #H @(fqus_ind_dx … H) -G1 -L1 -T1
+/3 width=5 by fpbs_strap2, fpbq_fquq/
+qed-.
+
+lemma cpxs_fpbs_trans: ∀h,o,G1,G2,L1,L2,T1,T,T2. ⦃G1, L1, T⦄ ≥[h, o] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] T → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T #T2 #H1 #H @(cpxs_ind_dx … H) -T1
+/3 width=5 by fpbs_strap2, fpbq_cpx/
+qed-.
+
+lemma lpxs_fpbs_trans: ∀h,o,G1,G2,L1,L,L2,T1,T2. ⦃G1, L, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1⦄ ⊢ ➡*[h, o] L → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G2 #L1 #L #L2 #T1 #T2 #H1 #H @(lpxs_ind_dx … H) -L1
+/3 width=5 by fpbs_strap2, fpbq_lpx/
+qed-.
+
+lemma lleq_fpbs_trans: ∀h,o,G1,G2,L1,L,L2,T1,T2. ⦃G1, L, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ →
+ L1 ≡[T1, 0] L → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
+/3 width=5 by fpbs_strap2, fpbq_lleq/ qed-.
+
+lemma cpxs_fqus_fpbs: ∀h,o,G1,G2,L1,L2,T1,T,T2. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] T →
+ ⦃G1, L1, T⦄ ⊐* ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
+/3 width=5 by fpbs_fqus_trans, cpxs_fpbs/ qed.
+
+lemma cpxs_fqup_fpbs: ∀h,o,G1,G2,L1,L2,T1,T,T2. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] T →
+ ⦃G1, L1, T⦄ ⊐+ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
+/3 width=5 by fpbs_fqup_trans, cpxs_fpbs/ qed.
+
+lemma fqus_lpxs_fpbs: ∀h,o,G1,G2,L1,L,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L, T2⦄ →
+ ⦃G2, L⦄ ⊢ ➡*[h, o] L2 → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
+/3 width=3 by fpbs_lpxs_trans, fqus_fpbs/ qed.
+
+lemma cpxs_fqus_lpxs_fpbs: ∀h,o,G1,G2,L1,L,L2,T1,T,T2. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] T →
+ ⦃G1, L1, T⦄ ⊐* ⦃G2, L, T2⦄ → ⦃G2, L⦄ ⊢ ➡*[h, o] L2 → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
+/3 width=5 by cpxs_fqus_fpbs, fpbs_lpxs_trans/ qed.
+
+lemma lpxs_lleq_fpbs: ∀h,o,G,L1,L,L2,T. ⦃G, L1⦄ ⊢ ➡*[h, o] L →
+ L ≡[T, 0] L2 → ⦃G, L1, T⦄ ≥[h, o] ⦃G, L2, T⦄.
+/3 width=3 by lpxs_fpbs_trans, lleq_fpbs/ qed.
+
+(* Note: this is used in the closure proof *)
+lemma cpr_lpr_fpbs: ∀h,o,G,L1,L2,T1,T2. ⦃G, L1⦄ ⊢ T1 ➡ T2 → ⦃G, L1⦄ ⊢ ➡ L2 →
+ ⦃G, L1, T1⦄ ≥[h, o] ⦃G, L2, T2⦄.
+/4 width=5 by fpbs_strap1, fpbq_fpbs, lpr_fpbq, cpr_fpbq/
+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/fpbq_aaa.ma".
+include "basic_2/computation/fpbs.ma".
+
+(* "QRST" PARALLEL COMPUTATION FOR CLOSURES *********************************)
+
+(* Properties on atomic arity assignment for terms **************************)
+
+lemma fpbs_aaa_conf: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ →
+ ∀A1. ⦃G1, L1⦄ ⊢ T1 ⁝ A1 → ∃A2. ⦃G2, L2⦄ ⊢ T2 ⁝ A2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fpbs_ind … H) -G2 -L2 -T2 /2 width=2 by ex_intro/
+#G #G2 #L #L2 #T #T2 #_ #H2 #IH1 #A #HA elim (IH1 … HA) -IH1 -A
+/2 width=8 by fpbq_aaa_conf/
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/btpredstaralt_8.ma".
+include "basic_2/multiple/lleq_fqus.ma".
+include "basic_2/computation/cpxs_lleq.ma".
+include "basic_2/computation/lpxs_lleq.ma".
+include "basic_2/computation/fpbs.ma".
+
+(* "QREST" PARALLEL COMPUTATION FOR CLOSURES ********************************)
+
+(* Note: alternative definition of fpbs *)
+definition fpbsa: ∀h. sd h → tri_relation genv lenv term ≝
+ λh,o,G1,L1,T1,G2,L2,T2.
+ ∃∃L0,L,T. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] T &
+ ⦃G1, L1, T⦄ ⊐* ⦃G2, L0, T2⦄ &
+ ⦃G2, L0⦄ ⊢ ➡*[h, o] L & L ≡[T2, 0] L2.
+
+interpretation "'big tree' parallel computation (closure) alternative"
+ 'BTPRedStarAlt h o G1 L1 T1 G2 L2 T2 = (fpbsa h o G1 L1 T1 G2 L2 T2).
+
+(* Basic properties *********************************************************)
+
+lemma fpb_fpbsa_trans: ∀h,o,G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ ≽[h, o] ⦃G, L, T⦄ →
+ ∀G2,L2,T2. ⦃G, L, T⦄ ≥≥[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥≥[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G #L1 #L #T1 #T * -G -L -T [ #G #L #T #HG1 | #T #HT1 | #L #HL1 | #L #HL1 ]
+#G2 #L2 #T2 * #L00 #L0 #T0 #HT0 #HG2 #HL00 #HL02
+[ elim (fquq_cpxs_trans … HT0 … HG1) -T
+ /3 width=7 by fqus_strap2, ex4_3_intro/
+| /3 width=7 by cpxs_strap2, ex4_3_intro/
+| lapply (lpx_cpxs_trans … HT0 … HL1) -HT0 #HT10
+ elim (lpx_fqus_trans … HG2 … HL1) -L
+ /3 width=7 by lpxs_strap2, cpxs_trans, ex4_3_intro/
+| lapply (lleq_cpxs_trans … HT0 … HL1) -HT0 #HT0
+ lapply (cpxs_lleq_conf_sn … HT0 … HL1) -HL1 #HL1
+ elim (lleq_fqus_trans … HG2 … HL1) -L #K00 #HG12 #HKL00
+ elim (lleq_lpxs_trans … HL00 … HKL00) -L00
+ /3 width=9 by lleq_trans, ex4_3_intro/
+]
+qed-.
+
+(* Main properties **********************************************************)
+
+theorem fpbs_fpbsa: ∀h,o,G1,G2,L1,L2,T1,T2.
+ ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥≥[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fpbs_ind_dx … H) -G1 -L1 -T1
+/2 width=7 by fpb_fpbsa_trans, ex4_3_intro/
+qed.
+
+(* Main inversion lemmas ****************************************************)
+
+theorem fpbsa_inv_fpbs: ∀h,o,G1,G2,L1,L2,T1,T2.
+ ⦃G1, L1, T1⦄ ≥≥[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 *
+/3 width=5 by cpxs_fqus_lpxs_fpbs, fpbs_strap1, fpbq_lleq/
+qed-.
+
+(* Advanced properties ******************************************************)
+
+lemma fpbs_intro_alt: ∀h,o,G1,G2,L1,L0,L,L2,T1,T,T2.
+ ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] T → ⦃G1, L1, T⦄ ⊐* ⦃G2, L0, T2⦄ →
+ ⦃G2, L0⦄ ⊢ ➡*[h, o] L → L ≡[T2, 0] L2 → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ .
+/3 width=7 by fpbsa_inv_fpbs, ex4_3_intro/ qed.
+
+(* Advanced inversion lemmas *************************************************)
+
+lemma fpbs_inv_alt: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ →
+ ∃∃L0,L,T. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] T &
+ ⦃G1, L1, T⦄ ⊐* ⦃G2, L0, T2⦄ &
+ ⦃G2, L0⦄ ⊢ ➡*[h, o] L & L ≡[T2, 0] L2.
+/2 width=1 by fpbs_fpbsa/ 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/fpbq_alt.ma".
+include "basic_2/computation/fpbs_alt.ma".
+
+(* "QRST" PARALLEL COMPUTATION FOR CLOSURES *********************************)
+
+(* Properties on extended context-sensitive parallel computation for terms **)
+
+lemma fpbs_cpx_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡[h, o] U2 → (T2 = U2 → ⊥) →
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ≥[h, o] ⦃G2, L2, U2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #U2 #HTU2 #HnTU2 elim (fpbs_inv_alt … H) -H
+#L00 #L0 #T0 #HT10 #H10 #HL00 #HL02 lapply (lleq_cpx_trans … HTU2 … HL02) -HTU2
+#HTU2 lapply (cpx_lleq_conf_sn … HTU2 … HL02) -HL02
+#HL02 lapply (lpxs_cpx_trans … HTU2 … HL00) -HTU2
+#HTU2 elim (fqus_cpxs_trans_neq … H10 … HTU2 HnTU2) -H10 -HTU2 -HnTU2
+#U0 #HTU0 #HnTU0 #HU02 elim (eq_term_dec T1 T0) #HnT10 destruct
+[ -HT10 elim (cpxs_neq_inv_step_sn … HTU0 HnTU0) -HTU0 -HnTU0
+| -HnTU0 elim (cpxs_neq_inv_step_sn … HT10 HnT10) -HT10 -HnT10
+]
+/4 width=10 by fpbs_intro_alt, cpxs_trans, ex3_intro/
+qed-.
+
+(* Properties on "rst" proper parallel reduction on closures ****************)
+
+lemma fpb_fpbs: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
+/3 width=1 by fpbq_fpbs, fpb_fpbq/ qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/computation/fpbs.ma".
+
+(* "QRST" PARALLEL COMPUTATION FOR CLOSURES *********************************)
+
+(* Main properties **********************************************************)
+
+theorem fpbs_trans: ∀h,o. tri_transitive … (fpbs h o).
+/2 width=5 by tri_TC_transitive/ qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/computation/cpxs_lift.ma".
+include "basic_2/computation/fpbs.ma".
+
+(* "QRST" PARALLEL COMPUTATION FOR CLOSURES *********************************)
+
+(* Advanced properties ******************************************************)
+
+lemma lstas_fpbs: ∀h,o,G,L,T1,T2,d2. ⦃G, L⦄ ⊢ T1 •*[h, d2] T2 →
+ ∀d1. d2 ≤ d1 → ⦃G, L⦄ ⊢ T1 ▪[h, o] d1 → ⦃G, L, T1⦄ ≥[h, o] ⦃G, L, T2⦄.
+/3 width=5 by cpxs_fpbs, lstas_cpxs/ qed.
+
+lemma sta_fpbs: ∀h,o,G,L,T,U,d.
+ ⦃G, L⦄ ⊢ T ▪[h, o] d+1 → ⦃G, L⦄ ⊢ T •*[h, 1] U →
+ ⦃G, L, T⦄ ≥[h, o] ⦃G, L, U⦄.
+/2 width=5 by lstas_fpbs/ qed.
+
+(* Note: this is used in the closure proof *)
+lemma cpr_lpr_sta_fpbs: ∀h,o,G,L1,L2,T1,T2,U2,d2.
+ ⦃G, L1⦄ ⊢ T1 ➡ T2 → ⦃G, L1⦄ ⊢ ➡ L2 →
+ ⦃G, L2⦄ ⊢ T2 ▪[h, o] d2+1 → ⦃G, L2⦄ ⊢ T2 •*[h, 1] U2 →
+ ⦃G, L1, T1⦄ ≥[h, o] ⦃G, L2, U2⦄.
+/4 width=5 by fpbs_strap1, cpr_lpr_fpbs, sta_cpx, fpbq_cpx/ qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/btsn_5.ma".
+include "basic_2/reduction/fpb.ma".
+include "basic_2/computation/csx.ma".
+
+(* "QRST" STRONGLY NORMALIZING CLOSURES *************************************)
+
+inductive fsb (h) (o): relation3 genv lenv term ≝
+| fsb_intro: ∀G1,L1,T1. (
+ ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ → fsb h o G2 L2 T2
+ ) → fsb h o G1 L1 T1
+.
+
+interpretation
+ "'qrst' strong normalization (closure)"
+ 'BTSN h o G L T = (fsb h o G L T).
+
+(* Basic eliminators ********************************************************)
+
+lemma fsb_ind_alt: ∀h,o. ∀R: relation3 …. (
+ ∀G1,L1,T1. ⦥[h,o] ⦃G1, L1, T1⦄ → (
+ ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ → R G2 L2 T2
+ ) → R G1 L1 T1
+ ) →
+ ∀G,L,T. ⦥[h, o] ⦃G, L, T⦄ → R G L T.
+#h #o #R #IH #G #L #T #H elim H -G -L -T
+/4 width=1 by fsb_intro/
+qed-.
+
+(* Basic inversion lemmas ***************************************************)
+
+lemma fsb_inv_csx: ∀h,o,G,L,T. ⦥[h, o] ⦃G, L, T⦄ → ⦃G, L⦄ ⊢ ⬊*[h, o] T.
+#h #o #G #L #T #H elim H -G -L -T /5 width=1 by csx_intro, fpb_cpx/
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/computation/fpbs_aaa.ma".
+include "basic_2/computation/csx_aaa.ma".
+include "basic_2/computation/fsb_csx.ma".
+
+(* "QRST" STRONGLY NORMALIZING CLOSURES *************************************)
+
+(* Main properties **********************************************************)
+
+(* Note: this is the "big tree" theorem ("RST" version) *)
+theorem aaa_fsb: ∀h,o,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ⦥[h, o] ⦃G, L, T⦄.
+/3 width=2 by aaa_csx, csx_fsb/ qed.
+
+(* Note: this is the "big tree" theorem ("QRST" version) *)
+theorem aaa_fsba: ∀h,o,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ⦥⦥[h, o] ⦃G, L, T⦄.
+/3 width=2 by fsb_fsba, aaa_fsb/ qed.
+
+(* Advanced eliminators on atomica arity assignment for terms ***************)
+
+fact aaa_ind_fpb_aux: ∀h,o. ∀R:relation3 genv lenv term.
+ (∀G1,L1,T1,A. ⦃G1, L1⦄ ⊢ T1 ⁝ A →
+ (∀G2,L2,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
+ R G1 L1 T1
+ ) →
+ ∀G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → ∀A. ⦃G, L⦄ ⊢ T ⁝ A → R G L T.
+#h #o #R #IH #G #L #T #H @(csx_ind_fpb … H) -G -L -T
+#G1 #L1 #T1 #H1 #IH1 #A1 #HTA1 @IH -IH //
+#G2 #L2 #T2 #H12 elim (fpbs_aaa_conf h o … G2 … L2 … T2 … HTA1) -A1
+/2 width=2 by fpb_fpbs/
+qed-.
+
+lemma aaa_ind_fpb: ∀h,o. ∀R:relation3 genv lenv term.
+ (∀G1,L1,T1,A. ⦃G1, L1⦄ ⊢ T1 ⁝ A →
+ (∀G2,L2,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
+ R G1 L1 T1
+ ) →
+ ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → R G L T.
+/4 width=4 by aaa_ind_fpb_aux, aaa_csx/ qed-.
+
+fact aaa_ind_fpbg_aux: ∀h,o. ∀R:relation3 genv lenv term.
+ (∀G1,L1,T1,A. ⦃G1, L1⦄ ⊢ T1 ⁝ A →
+ (∀G2,L2,T2. ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
+ R G1 L1 T1
+ ) →
+ ∀G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → ∀A. ⦃G, L⦄ ⊢ T ⁝ A → R G L T.
+#h #o #R #IH #G #L #T #H @(csx_ind_fpbg … H) -G -L -T
+#G1 #L1 #T1 #H1 #IH1 #A1 #HTA1 @IH -IH //
+#G2 #L2 #T2 #H12 elim (fpbs_aaa_conf h o … G2 … L2 … T2 … HTA1) -A1
+/2 width=2 by fpbg_fwd_fpbs/
+qed-.
+
+lemma aaa_ind_fpbg: ∀h,o. ∀R:relation3 genv lenv term.
+ (∀G1,L1,T1,A. ⦃G1, L1⦄ ⊢ T1 ⁝ A →
+ (∀G2,L2,T2. ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
+ R G1 L1 T1
+ ) →
+ ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → R G L T.
+/4 width=4 by aaa_ind_fpbg_aux, aaa_csx/ qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/btsnalt_5.ma".
+include "basic_2/computation/fpbg_fpbs.ma".
+include "basic_2/computation/fsb.ma".
+
+(* "QRST" STRONGLY NORMALIZING CLOSURES *************************************)
+
+(* Note: alternative definition of fsb *)
+inductive fsba (h) (o): relation3 genv lenv term ≝
+| fsba_intro: ∀G1,L1,T1. (
+ ∀G2,L2,T2. ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄ → fsba h o G2 L2 T2
+ ) → fsba h o G1 L1 T1.
+
+interpretation
+ "'big tree' strong normalization (closure) alternative"
+ 'BTSNAlt h o G L T = (fsba h o G L T).
+
+(* Basic eliminators ********************************************************)
+
+lemma fsba_ind_alt: ∀h,o. ∀R: relation3 …. (
+ ∀G1,L1,T1. ⦥⦥[h,o] ⦃G1, L1, T1⦄ → (
+ ∀G2,L2,T2. ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄ → R G2 L2 T2
+ ) → R G1 L1 T1
+ ) →
+ ∀G,L,T. ⦥⦥[h, o] ⦃G, L, T⦄ → R G L T.
+#h #o #R #IH #G #L #T #H elim H -G -L -T
+/4 width=1 by fsba_intro/
+qed-.
+
+(* Basic properties *********************************************************)
+
+lemma fsba_fpbs_trans: ∀h,o,G1,L1,T1. ⦥⦥[h, o] ⦃G1, L1, T1⦄ →
+ ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ → ⦥⦥[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #L1 #T1 #H @(fsba_ind_alt … H) -G1 -L1 -T1
+/4 width=5 by fsba_intro, fpbs_fpbg_trans/
+qed-.
+
+(* Main properties **********************************************************)
+
+theorem fsb_fsba: ∀h,o,G,L,T. ⦥[h, o] ⦃G, L, T⦄ → ⦥⦥[h, o] ⦃G, L, T⦄.
+#h #o #G #L #T #H @(fsb_ind_alt … H) -G -L -T
+#G1 #L1 #T1 #_ #IH @fsba_intro
+#G2 #L2 #T2 * /3 width=5 by fsba_fpbs_trans/
+qed.
+
+(* Main inversion lemmas ****************************************************)
+
+theorem fsba_inv_fsb: ∀h,o,G,L,T. ⦥⦥[h, o] ⦃G, L, T⦄ → ⦥[h, o] ⦃G, L, T⦄.
+#h #o #G #L #T #H @(fsba_ind_alt … H) -G -L -T
+/4 width=1 by fsb_intro, fpb_fpbg/
+qed-.
+
+(* Advanced properties ******************************************************)
+
+lemma fsb_fpbs_trans: ∀h,o,G1,L1,T1. ⦥[h, o] ⦃G1, L1, T1⦄ →
+ ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ → ⦥[h, o] ⦃G2, L2, T2⦄.
+/4 width=5 by fsba_inv_fsb, fsb_fsba, fsba_fpbs_trans/ qed-.
+
+(* Advanced eliminators *****************************************************)
+
+lemma fsb_ind_fpbg: ∀h,o. ∀R:relation3 genv lenv term.
+ (∀G1,L1,T1. ⦥[h, o] ⦃G1, L1, T1⦄ →
+ (∀G2,L2,T2. ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
+ R G1 L1 T1
+ ) →
+ ∀G1,L1,T1. ⦥[h, o] ⦃G1, L1, T1⦄ → R G1 L1 T1.
+#h #o #R #IH #G1 #L1 #T1 #H @(fsba_ind_alt h o … G1 L1 T1)
+/3 width=1 by fsba_inv_fsb, fsb_fsba/
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/computation/fpbs_fpb.ma".
+include "basic_2/computation/fpbs_fpbs.ma".
+include "basic_2/computation/csx_fpbs.ma".
+include "basic_2/computation/lsx_csx.ma".
+include "basic_2/computation/fsb_alt.ma".
+
+(* "QRST" STRONGLY NORMALIZING CLOSURES *************************************)
+
+(* Advanced propreties on context-sensitive extended normalizing terms ******)
+
+lemma csx_fsb_fpbs: ∀h,o,G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬊*[h, o] T1 →
+ ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ → ⦥[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #L1 #T1 #H @(csx_ind … H) -T1
+#T1 #HT1 #IHc #G2 #L2 #T2 @(fqup_wf_ind … G2 L2 T2) -G2 -L2 -T2
+#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 #_ #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 @(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 ]
+ [ #H destruct /5 width=5 by fsb_fpbs_trans, lpxs_fpbs, fpbs_fqup_trans, lpx_lpxs/
+ | #HnT04 #HT04 #H04 #HL43 elim (cpxs_neq_inv_step_sn … HT04 HnT04) -HT04 -HnT04
+ #T2 #HT02 #HnT02 #HT24 elim (fpbs_cpx_trans_neq … H10 … HT02 HnT02) -T0
+ lapply (fpbs_intro_alt … G3 … L4 … L3 L3 … T3 … HT24 ? ? ?) -HT24
+ /3 width=8 by fpbs_trans, lpx_lpxs, fqup_fqus/ (**) (* full auto too slow *)
+ ]
+ ]
+]
+qed.
+
+lemma csx_fsb: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → ⦥[h, o] ⦃G, L, T⦄.
+/2 width=5 by csx_fsb_fpbs/ qed.
+
+(* Advanced eliminators *****************************************************)
+
+lemma csx_ind_fpb: ∀h,o. ∀R:relation3 genv lenv term.
+ (∀G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬊*[h, o] T1 →
+ (∀G2,L2,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
+ R G1 L1 T1
+ ) →
+ ∀G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → R G L T.
+/4 width=4 by fsb_inv_csx, csx_fsb, fsb_ind_alt/ qed-.
+
+lemma csx_ind_fpbg: ∀h,o. ∀R:relation3 genv lenv term.
+ (∀G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬊*[h, o] T1 →
+ (∀G2,L2,T2. ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
+ R G1 L1 T1
+ ) →
+ ∀G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → R G L T.
+/4 width=4 by fsb_inv_csx, csx_fsb, fsb_ind_fpbg/ 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/genv.ma".
+include "basic_2/multiple/drops.ma".
+
+(* GENERIC COMPUTATION PROPERTIES *******************************************)
+
+definition nf ≝ λRR:relation4 genv lenv term term. λRS:relation term.
+ λG,L,T. NF … (RR G L) RS T.
+
+definition candidate: Type[0] ≝ relation3 genv lenv term.
+
+definition CP0 ≝ λRR:relation4 genv lenv term term. λRS:relation term.
+ ∀G. d_liftable1 (nf RR RS G) (Ⓕ).
+
+definition CP1 ≝ λRR:relation4 genv lenv term term. λRS:relation term.
+ ∀G,L. ∃s. NF … (RR G L) RS (⋆s).
+
+definition CP2 ≝ λRP:candidate. ∀G. d_liftable1 (RP G) (Ⓕ).
+
+definition CP3 ≝ λRP:candidate.
+ ∀G,L,T,s. RP G L (ⓐ⋆s.T) → RP G L T.
+
+(* requirements for generic computation properties *)
+record gcp (RR:relation4 genv lenv term term) (RS:relation term) (RP:candidate) : Prop ≝
+{ cp0: CP0 RR RS;
+ cp1: CP1 RR RS;
+ cp2: CP2 RP;
+ cp3: CP3 RP
+}.
+
+(* Basic properties *********************************************************)
+
+(* Basic_1: was: nf2_lift1 *)
+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. 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. d_liftables1_all (RP G) (Ⓕ).
+#RR #RS #RP #H #G @d1_liftables_liftables_all /2 width=7 by gcp2_lifts/
+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/lifts_lifts.ma".
+include "basic_2/multiple/drops_drops.ma".
+include "basic_2/static/aaa_lifts.ma".
+include "basic_2/static/aaa_aaa.ma".
+include "basic_2/computation/lsubc_drops.ma".
+
+(* GENERIC COMPUTATION PROPERTIES *******************************************)
+
+(* Main properties **********************************************************)
+
+(* Basic_1: was: sc3_arity_csubc *)
+theorem acr_aaa_csubc_lifts: ∀RR,RS,RP.
+ gcp RR RS RP → gcr RR RS RP RP →
+ ∀G,L1,T,A. ⦃G, L1⦄ ⊢ T ⁝ A → ∀L0,cs. ⬇*[Ⓕ, cs] L0 ≡ L1 →
+ ∀T0. ⬆*[cs] T ≡ T0 → ∀L2. G ⊢ L2 ⫃[RP] L0 →
+ ⦃G, L2, T0⦄ ϵ[RP] 〚A〛.
+#RR #RS #RP #H1RP #H2RP #G #L1 #T #A #H elim H -G -L1 -T -A
+[ #G #L #s #L0 #cs #HL0 #X #H #L2 #HL20
+ >(lifts_inv_sort1 … H) -H
+ lapply (acr_gcr … H1RP H2RP (⓪)) #HAtom
+ lapply (c4 … HAtom G L2 (◊)) /2 width=1 by/
+| #I #G #L1 #K1 #V1 #B #i #HLK1 #HKV1B #IHB #L0 #cs #HL01 #X #H #L2 #HL20
+ 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 #cs1 #i0 #HL0 #H #Hi0 #Hcs1
+ >(at_mono … Hi1 … Hi0) -i1
+ elim (drops_inv_skip2 … Hcs1 … H) -cs1 #K0 #V0 #cs0 #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 … Hcs0 … HV10 … HV0) -HV10 #V2 #HV12 #HV2
+ lapply (c5 … HB ? G ? ? (◊) … HV0 HLK2) /3 width=7 by drops_cons, lifts_cons/ (* Note: uses IHB HL20 V2 HV0 *)
+ | -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
+ lapply (aaa_mono … H2KV0A … HKV0B) #H destruct -H2KV0A -HKV0B
+ elim (lift_total V0 0 (i0 +1)) #V3 #HV03
+ elim (lift_total V2 0 (i0 +1)) #V #HV2
+ lapply (c5 … HB ? G ? ? (◊) … (ⓝV3.V) … HLK2) /2 width=1 by lift_flat/
+ lapply (c7 … HB G L2 (◊)) /3 width=7 by gcr_lift/
+ ]
+| #a #G #L #V #T #B #A #_ #_ #IHB #IHA #L0 #cs #HL0 #X #H #L2 #HL20
+ elim (lifts_inv_bind1 … H) -H #V0 #T0 #HV0 #HT0 #H destruct
+ lapply (acr_gcr … H1RP H2RP A) #HA
+ lapply (acr_gcr … H1RP H2RP B) #HB
+ lapply (c1 … HB) -HB #HB
+ lapply (c6 … HA G L2 (◊) (◊)) /4 width=5 by lsubc_pair, drops_skip, liftv_nil/
+| #a #G #L #W #T #B #A #HLWB #_ #IHB #IHA #L0 #cs #HL0 #X #H #L2 #HL02
+ elim (lifts_inv_bind1 … H) -H #W0 #T0 #HW0 #HT0 #H destruct
+ @(acr_abst … H1RP H2RP) /2 width=5 by/
+ #L3 #V3 #W3 #T3 #cs3 #HL32 #HW03 #HT03 #H1B #H2B
+ elim (drops_lsubc_trans … H1RP … HL32 … HL02) -L2 #L2 #HL32 #HL20
+ lapply (aaa_lifts … L2 W3 … (cs @@ cs3) … HLWB) -HLWB /2 width=4 by drops_trans, lifts_trans/ #HLW2B
+ @(IHA (L2. ⓛW3) … (cs + 1 @@ cs3 + 1)) -IHA
+ /3 width=5 by lsubc_beta, drops_trans, drops_skip, lifts_trans/
+| #G #L #V #T #B #A #_ #_ #IHB #IHA #L0 #cs #HL0 #X #H #L2 #HL20
+ elim (lifts_inv_flat1 … H) -H #V0 #T0 #HV0 #HT0 #H destruct
+ /3 width=10 by drops_nil, lifts_nil/
+| #G #L #V #T #A #_ #_ #IH1A #IH2A #L0 #cs #HL0 #X #H #L2 #HL20
+ elim (lifts_inv_flat1 … H) -H #V0 #T0 #HV0 #HT0 #H destruct
+ lapply (acr_gcr … H1RP H2RP A) #HA
+ lapply (c7 … HA G L2 (◊)) /3 width=5 by/
+]
+qed.
+
+(* Basic_1: was: sc3_arity *)
+lemma acr_aaa: ∀RR,RS,RP. gcp RR RS RP → gcr RR RS RP RP →
+ ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ⦃G, L, T⦄ ϵ[RP] 〚A〛.
+/2 width=8 by drops_nil, lifts_nil, acr_aaa_csubc_lifts/ qed.
+
+lemma gcr_aaa: ∀RR,RS,RP. gcp RR RS RP → gcr RR RS RP RP →
+ ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → RP G L T.
+#RR #RS #RP #H1RP #H2RP #G #L #T #A #HT
+lapply (acr_gcr … H1RP H2RP A) #HA
+@(c1 … HA) /2 width=4 by acr_aaa/
+qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/ineint_5.ma".
+include "basic_2/grammar/aarity.ma".
+include "basic_2/multiple/mr2_mr2.ma".
+include "basic_2/multiple/lifts_lift_vector.ma".
+include "basic_2/multiple/drops_drop.ma".
+include "basic_2/computation/gcp.ma".
+
+(* GENERIC COMPUTATION PROPERTIES *******************************************)
+
+(* Note: this is Girard's CR1 *)
+definition S1 ≝ λRP,C:candidate.
+ ∀G,L,T. C G L T → RP G L T.
+
+(* Note: this is Tait's iii, or Girard's CR4 *)
+definition S2 ≝ λRR:relation4 genv lenv term term. λRS:relation term. λRP,C:candidate.
+ ∀G,L,Vs. all … (RP G L) Vs →
+ ∀T. 𝐒⦃T⦄ → NF … (RR G L) RS T → C G L (ⒶVs.T).
+
+(* Note: this generalizes Tait's ii *)
+definition S3 ≝ λC:candidate.
+ ∀a,G,L,Vs,V,T,W.
+ C G L (ⒶVs.ⓓ{a}ⓝW.V.T) → C G L (ⒶVs.ⓐV.ⓛ{a}W.T).
+
+definition S4 ≝ λRP,C:candidate.
+ ∀G,L,Vs. all … (RP G L) Vs → ∀s. C G L (ⒶVs.⋆s).
+
+definition S5 ≝ λC:candidate. ∀I,G,L,K,Vs,V1,V2,i.
+ C G L (ⒶVs.V2) → ⬆[0, i+1] V1 ≡ V2 →
+ ⬇[i] L ≡ K.ⓑ{I}V1 → C G L (ⒶVs.#i).
+
+definition S6 ≝ λRP,C:candidate.
+ ∀G,L,V1c,V2c. ⬆[0, 1] V1c ≡ V2c →
+ ∀a,V,T. C G (L.ⓓV) (ⒶV2c.T) → RP G L V → C G L (ⒶV1c.ⓓ{a}V.T).
+
+definition S7 ≝ λC:candidate.
+ ∀G,L,Vs,T,W. C G L (ⒶVs.T) → C G L (ⒶVs.W) → C G L (ⒶVs.ⓝW.T).
+
+(* requirements for the generic reducibility candidate *)
+record gcr (RR:relation4 genv lenv term term) (RS:relation term) (RP,C:candidate) : Prop ≝
+{ c1: S1 RP C;
+ c2: S2 RR RS RP C;
+ c3: S3 C;
+ c4: S4 RP C;
+ c5: S5 C;
+ c6: S6 RP C;
+ c7: S7 C
+}.
+
+(* the functional construction for candidates *)
+definition cfun: candidate → candidate → candidate ≝
+ λC1,C2,G,K,T. ∀L,W,U,cs.
+ ⬇*[Ⓕ, cs] L ≡ K → ⬆*[cs] T ≡ U → C1 G L W → C2 G L (ⓐW.U).
+
+(* the reducibility candidate associated to an atomic arity *)
+rec definition acr (RP:candidate) (A:aarity) on A: candidate ≝
+match A with
+[ AAtom ⇒ RP
+| APair B A ⇒ cfun (acr RP B) (acr RP A)
+].
+
+interpretation
+ "candidate of reducibility of an atomic arity (abstract)"
+ 'InEInt RP G L T A = (acr RP A G L T).
+
+(* Basic properties *********************************************************)
+
+(* Basic 1: was: sc3_lift *)
+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. 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:
+ sc3_sn3 sc3_abst sc3_appl sc3_abbr sc3_bind sc3_cast
+*)
+lemma acr_gcr: ∀RR,RS,RP. gcp RR RS RP → gcr RR RS RP RP →
+ ∀A. gcr RR RS RP (acr RP A).
+#RR #RS #RP #H1RP #H2RP #A elim A -A //
+#B #A #IHB #IHA @mk_gcr
+[ #G #L #T #H
+ elim (cp1 … H1RP G L) #s #HK
+ lapply (H L (⋆s) T (◊) ? ? ?) -H //
+ [ lapply (c2 … IHB G L (◊) … HK) //
+ | /3 width=6 by c1, cp3/
+ ]
+| #G #L #Vs #HVs #T #H1T #H2T #L0 #V0 #X #cs #HL0 #H #HB
+ elim (lifts_inv_applv1 … H) -H #V0c #T0 #HV0c #HT0 #H destruct
+ lapply (c1 … IHB … HB) #HV0
+ @(c2 … IHA … (V0 @ V0c))
+ /3 width=14 by gcp2_lifts_all, gcp2_lifts, gcp0_lifts, lifts_simple_dx, conj/
+| #a #G #L #Vs #U #T #W #HA #L0 #V0 #X #cs #HL0 #H #HB
+ elim (lifts_inv_applv1 … H) -H #V0c #Y #HV0c #HY #H destruct
+ elim (lifts_inv_flat1 … HY) -HY #U0 #X #HU0 #HX #H destruct
+ elim (lifts_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT0 #H destruct
+ @(c3 … IHA … (V0 @ V0c)) /5 width=6 by lifts_applv, lifts_flat, lifts_bind/
+| #G #L #Vs #HVs #s #L0 #V0 #X #cs #HL0 #H #HB
+ elim (lifts_inv_applv1 … H) -H #V0c #Y #HV0c #HY #H destruct
+ >(lifts_inv_sort1 … HY) -Y
+ lapply (c1 … IHB … HB) #HV0
+ @(c4 … IHA … (V0 @ V0c)) /3 width=7 by gcp2_lifts_all, conj/
+| #I #G #L #K #Vs #V1 #V2 #i #HA #HV12 #HLK #L0 #V0 #X #cs #HL0 #H #HB
+ elim (lifts_inv_applv1 … H) -H #V0c #Y #HV0c #HY #H destruct
+ elim (lifts_inv_lref1 … HY) -HY #i0 #Hi0 #H destruct
+ elim (drops_drop_trans … HL0 … HLK) #X #cs0 #i1 #HL02 #H #Hi1 #Hcs0
+ >(at_mono … Hi1 … Hi0) in HL02; -i1 #HL02
+ elim (drops_inv_skip2 … Hcs0 … H) -H -cs0 #L2 #W1 #cs0 #Hcs0 #HLK #HVW1 #H destruct
+ elim (lift_total W1 0 (i0 + 1)) #W2 #HW12
+ elim (lifts_lift_trans … Hcs0 … HVW1 … HW12) // -Hcs0 -Hi0 #V3 #HV13 #HVW2
+ >(lift_mono … HV13 … HV12) in HVW2; -V3 #HVW2
+ @(c5 … IHA … (V0 @ V0c) … HW12 HL02) /3 width=5 by lifts_applv/
+| #G #L #V1c #V2c #HV12c #a #V #T #HA #HV #L0 #V10 #X #cs #HL0 #H #HB
+ elim (lifts_inv_applv1 … H) -H #V10c #Y #HV10c #HY #H destruct
+ elim (lifts_inv_bind1 … HY) -HY #V0 #T0 #HV0 #HT0 #H destruct
+ elim (lift_total V10 0 1) #V20 #HV120
+ elim (liftv_total 0 1 V10c) #V20c #HV120c
+ @(c6 … IHA … (V10 @ V10c) (V20 @ V20c)) /3 width=7 by gcp2_lifts, liftv_cons/
+ @(HA … (cs + 1)) /2 width=2 by drops_skip/
+ [ @lifts_applv //
+ elim (liftsv_liftv_trans_le … HV10c … HV120c) -V10c #V10c #HV10c #HV120c
+ >(liftv_mono … HV12c … HV10c) -V1c //
+ | @(gcr_lift … H1RP … HB … HV120) /2 width=2 by drop_drop/
+ ]
+| #G #L #Vs #T #W #HA #HW #L0 #V0 #X #cs #HL0 #H #HB
+ elim (lifts_inv_applv1 … H) -H #V0c #Y #HV0c #HY #H destruct
+ elim (lifts_inv_flat1 … HY) -HY #W0 #T0 #HW0 #HT0 #H destruct
+ @(c7 … IHA … (V0 @ V0c)) /3 width=5 by lifts_applv/
+]
+qed.
+
+lemma acr_abst: ∀RR,RS,RP. gcp RR RS RP → gcr RR RS RP RP →
+ ∀a,G,L,W,T,A,B. ⦃G, L, W⦄ ϵ[RP] 〚B〛 → (
+ ∀L0,V0,W0,T0,cs. ⬇*[Ⓕ, cs] L0 ≡ L → ⬆*[cs] W ≡ W0 → ⬆*[cs + 1] T ≡ T0 →
+ ⦃G, L0, V0⦄ ϵ[RP] 〚B〛 → ⦃G, L0, W0⦄ ϵ[RP] 〚B〛 → ⦃G, L0.ⓓⓝW0.V0, T0⦄ ϵ[RP] 〚A〛
+ ) →
+ ⦃G, L, ⓛ{a}W.T⦄ ϵ[RP] 〚②B.A〛.
+#RR #RS #RP #H1RP #H2RP #a #G #L #W #T #A #B #HW #HA #L0 #V0 #X #cs #HL0 #H #HB
+lapply (acr_gcr … H1RP H2RP A) #HCA
+lapply (acr_gcr … H1RP H2RP B) #HCB
+elim (lifts_inv_bind1 … H) -H #W0 #T0 #HW0 #HT0 #H destruct
+lapply (gcr_lifts … H1RP … HL0 … HW0 HW) -HW #HW0
+lapply (c3 … HCA … a G L0 (◊)) #H @H -H
+lapply (c6 … HCA G L0 (◊) (◊) ?) // #H @H -H
+[ @(HA … HL0) //
+| lapply (c1 … HCB) -HCB #HCB
+ lapply (c7 … H2RP G L0 (◊)) /3 width=1 by/
+]
+qed.
+
+(* Basic_1: removed theorems 2: sc3_arity_gen sc3_repl *)
+(* Basic_1: removed local theorems 1: sc3_sn3_abst *)
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/cosn_5.ma".
+include "basic_2/computation/lsx.ma".
+
+(* SN EXTENDED STRONGLY CONORMALIZING LOCAL ENVIRONMENTS ********************)
+
+inductive lcosx (h) (o) (G): relation2 ynat lenv ≝
+| lcosx_sort: ∀l. lcosx h o G l (⋆)
+| lcosx_skip: ∀I,L,T. lcosx h o G 0 L → lcosx h o G 0 (L.ⓑ{I}T)
+| lcosx_pair: ∀I,L,T,l. G ⊢ ⬊*[h, o, T, l] L →
+ lcosx h o G l L → lcosx h o G (⫯l) (L.ⓑ{I}T)
+.
+
+interpretation
+ "sn extended strong conormalization (local environment)"
+ 'CoSN h o l G L = (lcosx h o G l L).
+
+(* Basic properties *********************************************************)
+
+lemma lcosx_O: ∀h,o,G,L. G ⊢ ~⬊*[h, o, 0] L.
+#h #o #G #L elim L /2 width=1 by lcosx_skip/
+qed.
+
+lemma lcosx_drop_trans_lt: ∀h,o,G,L,l. G ⊢ ~⬊*[h, o, l] L →
+ ∀I,K,V,i. ⬇[i] L ≡ K.ⓑ{I}V → i < l →
+ G ⊢ ~⬊*[h, o, ⫰(l-i)] K ∧ G ⊢ ⬊*[h, o, V, ⫰(l-i)] K.
+#h #o #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 #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 … Hil ?) -Hil /2 width=1 by ylt_inj/ >ypred_succ #Hil
+ elim (IHL … HLK ?) -IHL -HLK <yminus_inj >yminus_SO2 //
+ <(ypred_succ l) in ⊢ (%→%→?); >yminus_pred /2 width=1 by ylt_inj, conj/
+ ]
+]
+qed-.
+
+(* Basic inversion lemmas ***************************************************)
+
+fact lcosx_inv_succ_aux: ∀h,o,G,L,x. G ⊢ ~⬊*[h, o, x] L → ∀l. x = ⫯l →
+ L = ⋆ ∨
+ ∃∃I,K,V. L = K.ⓑ{I}V & G ⊢ ~⬊*[h, o, l] K &
+ G ⊢ ⬊*[h, o, V, l] K.
+#h #o #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 #l #HT #HL #x #H <(ysucc_inv_inj … H) -x
+ /3 width=6 by ex3_3_intro, or_intror/
+]
+qed-.
+
+lemma lcosx_inv_succ: ∀h,o,G,L,l. G ⊢ ~⬊*[h, o, ⫯l] L → L = ⋆ ∨
+ ∃∃I,K,V. L = K.ⓑ{I}V & G ⊢ ~⬊*[h, o, l] K &
+ G ⊢ ⬊*[h, o, V, l] K.
+/2 width=3 by lcosx_inv_succ_aux/ qed-.
+
+lemma lcosx_inv_pair: ∀h,o,I,G,L,T,l. G ⊢ ~⬊*[h, o, ⫯l] L.ⓑ{I}T →
+ G ⊢ ~⬊*[h, o, l] L ∧ G ⊢ ⬊*[h, o, T, l] L.
+#h #o #I #G #L #T #l #H elim (lcosx_inv_succ … H) -H
+[ #H destruct
+| * #Z #Y #X #H destruct /2 width=1 by conj/
+]
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/computation/lsx_drop.ma".
+include "basic_2/computation/lsx_lpx.ma".
+include "basic_2/computation/lsx_lpxs.ma".
+include "basic_2/computation/lcosx.ma".
+
+(* SN EXTENDED STRONGLY CONORMALIZING LOCAL ENVIRONMENTS ********************)
+
+(* Properties on extended context-sensitive parallel reduction for term *****)
+
+lemma lsx_cpx_trans_lcosx: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 →
+ ∀l. G ⊢ ~⬊*[h, o, l] L →
+ G ⊢ ⬊*[h, o, T1, l] L → G ⊢ ⬊*[h, o, T2, l] L.
+#h #o #G #L #T1 #T2 #H elim H -G -L -T1 -T2 //
+[ #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_succ1/
+ 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 -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 #l #HL #H
+ elim (lsx_inv_bind … H) -H #HV1 #HT1
+ @lsx_bind /2 width=2 by/ (**) (* explicit constructor *)
+ @(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 #l #HL #H
+ elim (lsx_inv_bind … H) -H #HV #HU1
+ <(ypred_succ l) <yminus_SO2
+ /4 width=7 by lcosx_pair, lsx_inv_lift_ge, drop_drop/
+| #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 #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 #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_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 [ <yplus_SO2 /3 width=7 by lsx_lift_ge, drop_drop/ ]
+ @(lsx_lreq_conf … (L.ⓓW1)) /3 width=1 by lcosx_pair, lreq_succ/
+]
+qed-.
+
+lemma lsx_cpx_trans_O: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 →
+ G ⊢ ⬊*[h, o, T1, 0] L → G ⊢ ⬊*[h, o, T2, 0] L.
+/2 width=3 by lsx_cpx_trans_lcosx/ qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/predsnstar_3.ma".
+include "basic_2/substitution/lpx_sn_tc.ma".
+include "basic_2/reduction/lpr.ma".
+
+(* SN PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS ****************************)
+
+definition lprs: relation3 genv lenv lenv ≝
+ λG. TC … (lpr G).
+
+interpretation "parallel computation (local environment, sn variant)"
+ 'PRedSnStar G L1 L2 = (lprs G L1 L2).
+
+(* Basic eliminators ********************************************************)
+
+lemma lprs_ind: ∀G,L1. ∀R:predicate lenv. R L1 →
+ (∀L,L2. ⦃G, L1⦄ ⊢ ➡* L → ⦃G, L⦄ ⊢ ➡ L2 → R L → R L2) →
+ ∀L2. ⦃G, L1⦄ ⊢ ➡* L2 → R L2.
+#G #L1 #R #HL1 #IHL1 #L2 #HL12
+@(TC_star_ind … HL1 IHL1 … HL12) //
+qed-.
+
+lemma lprs_ind_dx: ∀G,L2. ∀R:predicate lenv. R L2 →
+ (∀L1,L. ⦃G, L1⦄ ⊢ ➡ L → ⦃G, L⦄ ⊢ ➡* L2 → R L → R L1) →
+ ∀L1. ⦃G, L1⦄ ⊢ ➡* L2 → R L1.
+#G #L2 #R #HL2 #IHL2 #L1 #HL12
+@(TC_star_ind_dx … HL2 IHL2 … HL12) //
+qed-.
+
+(* Basic properties *********************************************************)
+
+lemma lpr_lprs: ∀G,L1,L2. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L1⦄ ⊢ ➡* L2.
+/2 width=1 by inj/ qed.
+
+lemma lprs_refl: ∀G,L. ⦃G, L⦄ ⊢ ➡* L.
+/2 width=1 by lpr_lprs/ qed.
+
+lemma lprs_strap1: ∀G,L1,L,L2. ⦃G, L1⦄ ⊢ ➡* L → ⦃G, L⦄ ⊢ ➡ L2 → ⦃G, L1⦄ ⊢ ➡* L2.
+/2 width=3 by step/ qed-.
+
+lemma lprs_strap2: ∀G,L1,L,L2. ⦃G, L1⦄ ⊢ ➡ L → ⦃G, L⦄ ⊢ ➡* L2 → ⦃G, L1⦄ ⊢ ➡* L2.
+/2 width=3 by TC_strap/ qed-.
+
+lemma lprs_pair_refl: ∀G,L1,L2. ⦃G, L1⦄ ⊢ ➡* L2 → ∀I,V. ⦃G, L1.ⓑ{I}V⦄ ⊢ ➡* L2.ⓑ{I}V.
+/2 width=1 by TC_lpx_sn_pair_refl/ qed.
+
+(* Basic inversion lemmas ***************************************************)
+
+lemma lprs_inv_atom1: ∀G,L2. ⦃G, ⋆⦄ ⊢ ➡* L2 → L2 = ⋆.
+/2 width=2 by TC_lpx_sn_inv_atom1/ qed-.
+
+lemma lprs_inv_atom2: ∀G,L1. ⦃G, L1⦄ ⊢ ➡* ⋆ → L1 = ⋆.
+/2 width=2 by TC_lpx_sn_inv_atom2/ qed-.
+
+(* Basic forward lemmas *****************************************************)
+
+lemma lprs_fwd_length: ∀G,L1,L2. ⦃G, L1⦄ ⊢ ➡* L2 → |L1| = |L2|.
+/2 width=2 by TC_lpx_sn_fwd_length/ qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/computation/cprs_cprs.ma".
+include "basic_2/computation/lprs.ma".
+
+(* SN PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS ****************************)
+
+(* Advanced properties ******************************************************)
+
+lemma lprs_pair: ∀I,G,L1,L2. ⦃G, L1⦄ ⊢ ➡* L2 →
+ ∀V1,V2. ⦃G, L1⦄ ⊢ V1 ➡* V2 → ⦃G, L1.ⓑ{I}V1⦄ ⊢ ➡* L2.ⓑ{I}V2.
+/2 width=1 by TC_lpx_sn_pair/ qed.
+
+(* Advanced inversion lemmas ************************************************)
+
+lemma lprs_inv_pair1: ∀I,G,K1,L2,V1. ⦃G, K1.ⓑ{I}V1⦄ ⊢ ➡* L2 →
+ ∃∃K2,V2. ⦃G, K1⦄ ⊢ ➡* K2 & ⦃G, K1⦄ ⊢ V1 ➡* V2 &
+ L2 = K2.ⓑ{I}V2.
+/3 width=3 by TC_lpx_sn_inv_pair1, lpr_cprs_trans/ qed-.
+
+lemma lprs_inv_pair2: ∀I,G,L1,K2,V2. ⦃G, L1⦄ ⊢ ➡* K2.ⓑ{I}V2 →
+ ∃∃K1,V1. ⦃G, K1⦄ ⊢ ➡* K2 & ⦃G, K1⦄ ⊢ V1 ➡* V2 &
+ L1 = K1.ⓑ{I}V1.
+/3 width=3 by TC_lpx_sn_inv_pair2, lpr_cprs_trans/ qed-.
+
+(* Advanced eliminators *****************************************************)
+
+lemma lprs_ind_alt: ∀G. ∀R:relation lenv.
+ R (⋆) (⋆) → (
+ ∀I,K1,K2,V1,V2.
+ ⦃G, K1⦄ ⊢ ➡* K2 → ⦃G, K1⦄ ⊢ V1 ➡* V2 →
+ R K1 K2 → R (K1.ⓑ{I}V1) (K2.ⓑ{I}V2)
+ ) →
+ ∀L1,L2. ⦃G, L1⦄ ⊢ ➡* L2 → R L1 L2.
+/3 width=4 by TC_lpx_sn_ind, lpr_cprs_trans/ qed-.
+
+(* Properties on context-sensitive parallel computation for terms ***********)
+
+lemma lprs_cpr_trans: ∀G. c_r_transitive … (cpr G) (λ_. lprs G).
+/3 width=5 by c_r_trans_LTC2, lpr_cprs_trans/ qed-.
+
+(* Basic_1: was just: pr3_pr3_pr3_t *)
+(* Note: alternative proof /3 width=5 by s_r_trans_LTC1, lprs_cpr_trans/ *)
+lemma lprs_cprs_trans: ∀G. c_rs_transitive … (cpr G) (λ_. lprs G).
+#G @c_r_to_c_rs_trans @c_r_trans_LTC2
+@c_rs_trans_TC1 /2 width=3 by lpr_cprs_trans/ (**) (* full auto too slow *)
+qed-.
+
+lemma lprs_cprs_conf_dx: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡* T1 →
+ ∀L1. ⦃G, L0⦄ ⊢ ➡* L1 →
+ ∃∃T. ⦃G, L1⦄ ⊢ T1 ➡* T & ⦃G, L1⦄ ⊢ T0 ➡* T.
+#G #L0 #T0 #T1 #HT01 #L1 #H @(lprs_ind … H) -L1 /2 width=3 by ex2_intro/
+#L #L1 #_ #HL1 * #T #HT1 #HT0 -L0
+elim (cprs_lpr_conf_dx … HT1 … HL1) -HT1 #T2 #HT2
+elim (cprs_lpr_conf_dx … HT0 … HL1) -L #T3 #HT3
+elim (cprs_conf … HT2 … HT3) -T
+/3 width=5 by cprs_trans, ex2_intro/
+qed-.
+
+lemma lprs_cpr_conf_dx: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡ T1 →
+ ∀L1. ⦃G, L0⦄ ⊢ ➡* L1 →
+ ∃∃T. ⦃G, L1⦄ ⊢ T1 ➡* T & ⦃G, L1⦄ ⊢ T0 ➡* T.
+/3 width=3 by lprs_cprs_conf_dx, cpr_cprs/ qed-.
+
+(* Note: this can be proved on its own using lprs_ind_dx *)
+lemma lprs_cprs_conf_sn: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡* T1 →
+ ∀L1. ⦃G, L0⦄ ⊢ ➡* L1 →
+ ∃∃T. ⦃G, L0⦄ ⊢ T1 ➡* T & ⦃G, L1⦄ ⊢ T0 ➡* T.
+#G #L0 #T0 #T1 #HT01 #L1 #HL01
+elim (lprs_cprs_conf_dx … HT01 … HL01) -HT01
+/3 width=3 by lprs_cprs_trans, ex2_intro/
+qed-.
+
+lemma lprs_cpr_conf_sn: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡ T1 →
+ ∀L1. ⦃G, L0⦄ ⊢ ➡* L1 →
+ ∃∃T. ⦃G, L0⦄ ⊢ T1 ➡* T & ⦃G, L1⦄ ⊢ T0 ➡* T.
+/3 width=3 by lprs_cprs_conf_sn, cpr_cprs/ qed-.
+
+lemma cprs_bind2: ∀G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡* V2 →
+ ∀I,T1,T2. ⦃G, L.ⓑ{I}V2⦄ ⊢ T1 ➡* T2 →
+ ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡* ⓑ{a,I}V2.T2.
+/4 width=5 by lprs_cprs_trans, lprs_pair, cprs_bind/ qed.
+
+(* Inversion lemmas on context-sensitive parallel computation for terms *****)
+
+(* Basic_1: was: pr3_gen_abst *)
+lemma cprs_inv_abst1: ∀a,G,L,W1,T1,U2. ⦃G, L⦄ ⊢ ⓛ{a}W1.T1 ➡* U2 →
+ ∃∃W2,T2. ⦃G, L⦄ ⊢ W1 ➡* W2 & ⦃G, L.ⓛW1⦄ ⊢ T1 ➡* T2 &
+ U2 = ⓛ{a}W2.T2.
+#a #G #L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /2 width=5 by ex3_2_intro/
+#U0 #U2 #_ #HU02 * #V0 #T0 #HV10 #HT10 #H destruct
+elim (cpr_inv_abst1 … HU02) -HU02 #V2 #T2 #HV02 #HT02 #H destruct
+lapply (lprs_cpr_trans … HT02 (L.ⓛV1) ?)
+/3 width=5 by lprs_pair, cprs_trans, cprs_strap1, ex3_2_intro/
+qed-.
+
+lemma cprs_inv_abst: ∀a,G,L,W1,W2,T1,T2. ⦃G, L⦄ ⊢ ⓛ{a}W1.T1 ➡* ⓛ{a}W2.T2 →
+ ⦃G, L⦄ ⊢ W1 ➡* W2 ∧ ⦃G, L.ⓛW1⦄ ⊢ T1 ➡* T2.
+#a #G #L #W1 #W2 #T1 #T2 #H elim (cprs_inv_abst1 … H) -H
+#W #T #HW1 #HT1 #H destruct /2 width=1 by conj/
+qed-.
+
+(* Basic_1: was pr3_gen_abbr *)
+lemma cprs_inv_abbr1: ∀a,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓓ{a}V1.T1 ➡* U2 → (
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡* V2 & ⦃G, L.ⓓV1⦄ ⊢ T1 ➡* T2 &
+ U2 = ⓓ{a}V2.T2
+ ) ∨
+ ∃∃T2. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡* T2 & ⬆[0, 1] U2 ≡ T2 & a = true.
+#a #G #L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /3 width=5 by ex3_2_intro, or_introl/
+#U0 #U2 #_ #HU02 * *
+[ #V0 #T0 #HV10 #HT10 #H destruct
+ elim (cpr_inv_abbr1 … HU02) -HU02 *
+ [ #V2 #T2 #HV02 #HT02 #H destruct
+ lapply (lprs_cpr_trans … HT02 (L.ⓓV1) ?)
+ /4 width=5 by lprs_pair, cprs_trans, cprs_strap1, ex3_2_intro, or_introl/
+ | #T2 #HT02 #HUT2
+ lapply (lprs_cpr_trans … HT02 (L.ⓓV1) ?) -HT02
+ /4 width=3 by lprs_pair, cprs_trans, ex3_intro, or_intror/
+ ]
+| #U1 #HTU1 #HU01 elim (lift_total U2 0 1)
+ #U #HU2 lapply (cpr_lift … HU02 (L.ⓓV1) … HU01 … HU2) -U0
+ /4 width=3 by cprs_strap1, drop_drop, ex3_intro, or_intror/
+]
+qed-.
+
+(* More advanced properties *************************************************)
+
+lemma lprs_pair2: ∀I,G,L1,L2. ⦃G, L1⦄ ⊢ ➡* L2 →
+ ∀V1,V2. ⦃G, L2⦄ ⊢ V1 ➡* V2 → ⦃G, L1.ⓑ{I}V1⦄ ⊢ ➡* L2.ⓑ{I}V2.
+/3 width=3 by lprs_pair, lprs_cprs_trans/ qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/reduction/lpr_drop.ma".
+include "basic_2/computation/lprs.ma".
+
+(* SN PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS ****************************)
+
+(* Properties on local environment slicing ***********************************)
+
+lemma lprs_drop_conf: ∀G. dropable_sn (lprs G).
+/3 width=3 by dropable_sn_TC, lpr_drop_conf/ qed-.
+
+lemma drop_lprs_trans: ∀G. dedropable_sn (lprs G).
+/3 width=3 by dedropable_sn_TC, drop_lpr_trans/ qed-.
+
+lemma lprs_drop_trans_O1: ∀G. dropable_dx (lprs G).
+/3 width=3 by dropable_dx_TC, lpr_drop_trans_O1/ qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/reduction/lpr_lpr.ma".
+include "basic_2/computation/lprs.ma".
+
+(* SN PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS ****************************)
+
+(* Advanced properties ******************************************************)
+
+lemma lprs_strip: ∀G. confluent2 … (lprs G) (lpr G).
+/3 width=3 by TC_strip1, lpr_conf/ qed-.
+
+(* Main properties **********************************************************)
+
+theorem lprs_conf: ∀G. confluent2 … (lprs G) (lprs G).
+/3 width=3 by TC_confluent2, lpr_conf/ qed-.
+
+theorem lprs_trans: ∀G. Transitive … (lprs G).
+/2 width=3 by trans_TC/ qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/predsnstar_5.ma".
+include "basic_2/reduction/lpx.ma".
+include "basic_2/computation/lprs.ma".
+
+(* SN EXTENDED PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS *******************)
+
+definition lpxs: ∀h. sd h → relation3 genv lenv lenv ≝
+ λh,o,G. TC … (lpx h o G).
+
+interpretation "extended parallel computation (local environment, sn variant)"
+ 'PRedSnStar h o G L1 L2 = (lpxs h o G L1 L2).
+
+(* Basic eliminators ********************************************************)
+
+lemma lpxs_ind: ∀h,o,G,L1. ∀R:predicate lenv. R L1 →
+ (∀L,L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L → ⦃G, L⦄ ⊢ ➡[h, o] L2 → R L → R L2) →
+ ∀L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → R L2.
+#h #o #G #L1 #R #HL1 #IHL1 #L2 #HL12
+@(TC_star_ind … HL1 IHL1 … HL12) //
+qed-.
+
+lemma lpxs_ind_dx: ∀h,o,G,L2. ∀R:predicate lenv. R L2 →
+ (∀L1,L. ⦃G, L1⦄ ⊢ ➡[h, o] L → ⦃G, L⦄ ⊢ ➡*[h, o] L2 → R L → R L1) →
+ ∀L1. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → R L1.
+#h #o #G #L2 #R #HL2 #IHL2 #L1 #HL12
+@(TC_star_ind_dx … HL2 IHL2 … HL12) //
+qed-.
+
+(* Basic properties *********************************************************)
+
+lemma lprs_lpxs: ∀h,o,G,L1,L2. ⦃G, L1⦄ ⊢ ➡* L2 → ⦃G, L1⦄ ⊢ ➡*[h, o] L2.
+/3 width=3 by lpr_lpx, monotonic_TC/ qed.
+
+lemma lpx_lpxs: ∀h,o,G,L1,L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 → ⦃G, L1⦄ ⊢ ➡*[h, o] L2.
+/2 width=1 by inj/ qed.
+
+lemma lpxs_refl: ∀h,o,G,L. ⦃G, L⦄ ⊢ ➡*[h, o] L.
+/2 width=1 by lprs_lpxs/ qed.
+
+lemma lpxs_strap1: ∀h,o,G,L1,L,L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L → ⦃G, L⦄ ⊢ ➡[h, o] L2 → ⦃G, L1⦄ ⊢ ➡*[h, o] L2.
+/2 width=3 by step/ qed.
+
+lemma lpxs_strap2: ∀h,o,G,L1,L,L2. ⦃G, L1⦄ ⊢ ➡[h, o] L → ⦃G, L⦄ ⊢ ➡*[h, o] L2 → ⦃G, L1⦄ ⊢ ➡*[h, o] L2.
+/2 width=3 by TC_strap/ qed.
+
+lemma lpxs_pair_refl: ∀h,o,G,L1,L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → ∀I,V. ⦃G, L1.ⓑ{I}V⦄ ⊢ ➡*[h, o] L2.ⓑ{I}V.
+/2 width=1 by TC_lpx_sn_pair_refl/ qed.
+
+(* Basic inversion lemmas ***************************************************)
+
+lemma lpxs_inv_atom1: ∀h,o,G,L2. ⦃G, ⋆⦄ ⊢ ➡*[h, o] L2 → L2 = ⋆.
+/2 width=2 by TC_lpx_sn_inv_atom1/ qed-.
+
+lemma lpxs_inv_atom2: ∀h,o,G,L1. ⦃G, L1⦄ ⊢ ➡*[h, o] ⋆ → L1 = ⋆.
+/2 width=2 by TC_lpx_sn_inv_atom2/ qed-.
+
+(* Basic forward lemmas *****************************************************)
+
+lemma lpxs_fwd_length: ∀h,o,G,L1,L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → |L1| = |L2|.
+/2 width=2 by TC_lpx_sn_fwd_length/ qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/reduction/lpx_aaa.ma".
+include "basic_2/computation/lpxs.ma".
+
+(* SN EXTENDED PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS *******************)
+
+(* Properties about atomic arity assignment on terms ************************)
+
+lemma lpxs_aaa_conf: ∀h,o,G,L1,T,A. ⦃G, L1⦄ ⊢ T ⁝ A →
+ ∀L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → ⦃G, L2⦄ ⊢ T ⁝ A.
+#h #o #G #L1 #T #A #HT #L2 #HL12
+@(TC_Conf3 … (λL,A. ⦃G, L⦄ ⊢ T ⁝ A) … HT ? HL12) /2 width=5 by lpx_aaa_conf/
+qed-.
+
+lemma lprs_aaa_conf: ∀G,L1,T,A. ⦃G, L1⦄ ⊢ T ⁝ A →
+ ∀L2. ⦃G, L1⦄ ⊢ ➡* L2 → ⦃G, L2⦄ ⊢ T ⁝ A.
+/3 width=5 by lprs_lpxs, lpxs_aaa_conf/ qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/computation/cpxs_cpxs.ma".
+include "basic_2/computation/lpxs.ma".
+
+(* SN EXTENDED PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS *******************)
+
+(* Advanced properties ******************************************************)
+
+lemma lpxs_pair: ∀h,o,I,G,L1,L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 →
+ ∀V1,V2. ⦃G, L1⦄ ⊢ V1 ➡*[h, o] V2 →
+ ⦃G, L1.ⓑ{I}V1⦄ ⊢ ➡*[h, o] L2.ⓑ{I}V2.
+/2 width=1 by TC_lpx_sn_pair/ qed.
+
+(* Advanced inversion lemmas ************************************************)
+
+lemma lpxs_inv_pair1: ∀h,o,I,G,K1,L2,V1. ⦃G, K1.ⓑ{I}V1⦄ ⊢ ➡*[h, o] L2 →
+ ∃∃K2,V2. ⦃G, K1⦄ ⊢ ➡*[h, o] K2 & ⦃G, K1⦄ ⊢ V1 ➡*[h, o] V2 & L2 = K2.ⓑ{I}V2.
+/3 width=3 by TC_lpx_sn_inv_pair1, lpx_cpxs_trans/ qed-.
+
+lemma lpxs_inv_pair2: ∀h,o,I,G,L1,K2,V2. ⦃G, L1⦄ ⊢ ➡*[h, o] K2.ⓑ{I}V2 →
+ ∃∃K1,V1. ⦃G, K1⦄ ⊢ ➡*[h, o] K2 & ⦃G, K1⦄ ⊢ V1 ➡*[h, o] V2 & L1 = K1.ⓑ{I}V1.
+/3 width=3 by TC_lpx_sn_inv_pair2, lpx_cpxs_trans/ qed-.
+
+(* Advanced eliminators *****************************************************)
+
+lemma lpxs_ind_alt: ∀h,o,G. ∀R:relation lenv.
+ R (⋆) (⋆) → (
+ ∀I,K1,K2,V1,V2.
+ ⦃G, K1⦄ ⊢ ➡*[h, o] K2 → ⦃G, K1⦄ ⊢ V1 ➡*[h, o] V2 →
+ R K1 K2 → R (K1.ⓑ{I}V1) (K2.ⓑ{I}V2)
+ ) →
+ ∀L1,L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → R L1 L2.
+/3 width=4 by TC_lpx_sn_ind, lpx_cpxs_trans/ qed-.
+
+(* Properties on context-sensitive extended parallel computation for terms **)
+
+lemma lpxs_cpx_trans: ∀h,o,G. c_r_transitive … (cpx h o G) (λ_.lpxs h o G).
+/3 width=5 by c_r_trans_LTC2, lpx_cpxs_trans/ qed-.
+
+(* Note: alternative proof: /3 width=5 by s_r_trans_TC1, lpxs_cpx_trans/ *)
+lemma lpxs_cpxs_trans: ∀h,o,G. c_rs_transitive … (cpx h o G) (λ_.lpxs h o G).
+#h #o #G @c_r_to_c_rs_trans @c_r_trans_LTC2
+@c_rs_trans_TC1 /2 width=3 by lpx_cpxs_trans/ (**) (* full auto too slow *)
+qed-.
+
+lemma cpxs_bind2: ∀h,o,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 →
+ ∀I,T1,T2. ⦃G, L.ⓑ{I}V2⦄ ⊢ T1 ➡*[h, o] T2 →
+ ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡*[h, o] ⓑ{a,I}V2.T2.
+/4 width=5 by lpxs_cpxs_trans, lpxs_pair, cpxs_bind/ qed.
+
+(* Inversion lemmas on context-sensitive ext parallel computation for terms *)
+
+lemma cpxs_inv_abst1: ∀h,o,a,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓛ{a}V1.T1 ➡*[h, o] U2 →
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 & ⦃G, L.ⓛV1⦄ ⊢ T1 ➡*[h, o] T2 &
+ U2 = ⓛ{a}V2.T2.
+#h #o #a #G #L #V1 #T1 #U2 #H @(cpxs_ind … H) -U2 /2 width=5 by ex3_2_intro/
+#U0 #U2 #_ #HU02 * #V0 #T0 #HV10 #HT10 #H destruct
+elim (cpx_inv_abst1 … HU02) -HU02 #V2 #T2 #HV02 #HT02 #H destruct
+lapply (lpxs_cpx_trans … HT02 (L.ⓛV1) ?)
+/3 width=5 by lpxs_pair, cpxs_trans, cpxs_strap1, ex3_2_intro/
+qed-.
+
+lemma cpxs_inv_abbr1: ∀h,o,a,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓓ{a}V1.T1 ➡*[h, o] U2 → (
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 & ⦃G, L.ⓓV1⦄ ⊢ T1 ➡*[h, o] T2 &
+ U2 = ⓓ{a}V2.T2
+ ) ∨
+ ∃∃T2. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡*[h, o] T2 & ⬆[0, 1] U2 ≡ T2 & a = true.
+#h #o #a #G #L #V1 #T1 #U2 #H @(cpxs_ind … H) -U2 /3 width=5 by ex3_2_intro, or_introl/
+#U0 #U2 #_ #HU02 * *
+[ #V0 #T0 #HV10 #HT10 #H destruct
+ elim (cpx_inv_abbr1 … HU02) -HU02 *
+ [ #V2 #T2 #HV02 #HT02 #H destruct
+ lapply (lpxs_cpx_trans … HT02 (L.ⓓV1) ?)
+ /4 width=5 by lpxs_pair, cpxs_trans, cpxs_strap1, ex3_2_intro, or_introl/
+ | #T2 #HT02 #HUT2
+ lapply (lpxs_cpx_trans … HT02 (L.ⓓV1) ?) -HT02
+ /4 width=3 by lpxs_pair, cpxs_trans, ex3_intro, or_intror/
+ ]
+| #U1 #HTU1 #HU01
+ elim (lift_total U2 0 1) #U #HU2
+ /6 width=12 by cpxs_strap1, cpx_lift, drop_drop, ex3_intro, or_intror/
+]
+qed-.
+
+(* More advanced properties *************************************************)
+
+lemma lpxs_pair2: ∀h,o,I,G,L1,L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 →
+ ∀V1,V2. ⦃G, L2⦄ ⊢ V1 ➡*[h, o] V2 → ⦃G, L1.ⓑ{I}V1⦄ ⊢ ➡*[h, o] L2.ⓑ{I}V2.
+/3 width=3 by lpxs_pair, lpxs_cpxs_trans/ qed.
+
+(* Properties on supclosure *************************************************)
+
+lemma lpx_fqup_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
+ ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 →
+ ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡*[h, o] T & ⦃G1, K1, T⦄ ⊐+ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
+[ #G2 #L2 #T2 #H12 #K1 #HKL1 elim (lpx_fqu_trans … H12 … HKL1) -L1
+ /3 width=5 by cpx_cpxs, fqu_fqup, ex3_2_intro/
+| #G #G2 #L #L2 #T #T2 #_ #H2 #IH1 #K1 #HLK1 elim (IH1 … HLK1) -L1
+ #L0 #T0 #HT10 #HT0 #HL0 elim (lpx_fqu_trans … H2 … HL0) -L
+ #L #T3 #HT3 #HT32 #HL2 elim (fqup_cpx_trans … HT0 … HT3) -T
+ /3 width=7 by cpxs_strap1, fqup_strap1, ex3_2_intro/
+]
+qed-.
+
+lemma lpx_fqus_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
+ ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 →
+ ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡*[h, o] T & ⦃G1, K1, T⦄ ⊐* ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqus_ind … H) -G2 -L2 -T2 [ /2 width=5 by ex3_2_intro/ ]
+#G #G2 #L #L2 #T #T2 #_ #H2 #IH1 #K1 #HLK1 elim (IH1 … HLK1) -L1
+#L0 #T0 #HT10 #HT0 #HL0 elim (lpx_fquq_trans … H2 … HL0) -L
+#L #T3 #HT3 #HT32 #HL2 elim (fqus_cpx_trans … HT0 … HT3) -T
+/3 width=7 by cpxs_strap1, fqus_strap1, ex3_2_intro/
+qed-.
+
+lemma lpxs_fquq_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
+ ∀K1. ⦃G1, K1⦄ ⊢ ➡*[h, o] L1 →
+ ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡*[h, o] T & ⦃G1, K1, T⦄ ⊐⸮ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡*[h, o] L2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #HT12 #K1 #H @(lpxs_ind_dx … H) -K1
+[ /2 width=5 by ex3_2_intro/
+| #K1 #K #HK1 #_ * #L #T #HT1 #HT2 #HL2 -HT12
+ lapply (lpx_cpxs_trans … HT1 … HK1) -HT1
+ elim (lpx_fquq_trans … HT2 … HK1) -K
+ /3 width=7 by lpxs_strap2, cpxs_strap1, ex3_2_intro/
+]
+qed-.
+
+lemma lpxs_fqup_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
+ ∀K1. ⦃G1, K1⦄ ⊢ ➡*[h, o] L1 →
+ ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡*[h, o] T & ⦃G1, K1, T⦄ ⊐+ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡*[h, o] L2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #HT12 #K1 #H @(lpxs_ind_dx … H) -K1
+[ /2 width=5 by ex3_2_intro/
+| #K1 #K #HK1 #_ * #L #T #HT1 #HT2 #HL2 -HT12
+ lapply (lpx_cpxs_trans … HT1 … HK1) -HT1
+ elim (lpx_fqup_trans … HT2 … HK1) -K
+ /3 width=7 by lpxs_strap2, cpxs_trans, ex3_2_intro/
+]
+qed-.
+
+lemma lpxs_fqus_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
+ ∀K1. ⦃G1, K1⦄ ⊢ ➡*[h, o] L1 →
+ ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡*[h, o] T & ⦃G1, K1, T⦄ ⊐* ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡*[h, o] L2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqus_ind … H) -G2 -L2 -T2 /2 width=5 by ex3_2_intro/
+#G #G2 #L #L2 #T #T2 #_ #H2 #IH1 #K1 #HLK1 elim (IH1 … HLK1) -L1
+#L0 #T0 #HT10 #HT0 #HL0 elim (lpxs_fquq_trans … H2 … HL0) -L
+#L #T3 #HT3 #HT32 #HL2 elim (fqus_cpxs_trans … HT3 … HT0) -T
+/3 width=7 by cpxs_trans, fqus_strap1, ex3_2_intro/
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/reduction/lpx_drop.ma".
+include "basic_2/computation/lpxs.ma".
+
+(* SN EXTENDED PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS *******************)
+
+(* Properties on local environment slicing ***********************************)
+
+lemma lpxs_drop_conf: ∀h,o,G. dropable_sn (lpxs h o G).
+/3 width=3 by dropable_sn_TC, lpx_drop_conf/ qed-.
+
+lemma drop_lpxs_trans: ∀h,o,G. dedropable_sn (lpxs h o G).
+/3 width=3 by dedropable_sn_TC, drop_lpx_trans/ qed-.
+
+lemma lpxs_drop_trans_O1: ∀h,o,G. dropable_dx (lpxs h o G).
+/3 width=3 by dropable_dx_TC, lpx_drop_trans_O1/ qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/multiple/lleq_lleq.ma".
+include "basic_2/reduction/lpx_lleq.ma".
+include "basic_2/computation/cpxs_lreq.ma".
+include "basic_2/computation/lpxs_drop.ma".
+include "basic_2/computation/lpxs_cpxs.ma".
+
+(* SN EXTENDED PARALLEL COMPUTATION FOR LOCAL ENVIRONMENTS ******************)
+
+(* Properties on lazy equivalence for local environments ********************)
+
+lemma lleq_lpxs_trans: ∀h,o,G,L2,K2. ⦃G, L2⦄ ⊢ ➡*[h, o] K2 →
+ ∀L1,T,l. L1 ≡[T, l] L2 →
+ ∃∃K1. ⦃G, L1⦄ ⊢ ➡*[h, o] K1 & K1 ≡[T, l] K2.
+#h #o #G #L2 #K2 #H @(lpxs_ind … H) -K2 /2 width=3 by ex2_intro/
+#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,o,G,L1,L2,T,l. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → (L1 ≡[T, l] L2 → ⊥) →
+ ∃∃L,L0. ⦃G, L1⦄ ⊢ ➡[h, o] L & L1 ≡[T, l] L → ⊥ & ⦃G, L⦄ ⊢ ➡*[h, o] L0 & L0 ≡[T, l] L2.
+#h #o #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 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
+ #L #H1 #H lapply (nlleq_lleq_div … H … H2) -H2
+ #H2 elim (lleq_lpxs_trans … H3 … H) -L0
+ /3 width=8 by lleq_trans, ex4_2_intro/
+ | -H12 -IH2 /3 width=6 by ex4_2_intro/
+ ]
+]
+qed-.
+
+lemma lpxs_lleq_fqu_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+ ∀K1. ⦃G1, K1⦄ ⊢ ➡*[h, o] L1 → K1 ≡[T1, 0] L1 →
+ ∃∃K2. ⦃G1, K1, T1⦄ ⊐ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡*[h, o] L2 & K2 ≡[T2, 0] L2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
+[ #I #G1 #L1 #V1 #X #H1 #H2 elim (lpxs_inv_pair2 … H1) -H1
+ #K0 #V0 #H1KL1 #_ #H destruct
+ elim (lleq_inv_lref_ge_dx … H2 ? I L1 V1) -H2 //
+ #K1 #H #H2KL1 lapply (drop_inv_O2 … H) -H #H destruct
+ /2 width=4 by fqu_lref_O, ex3_intro/
+| * [ #a ] #I #G1 #L1 #V1 #T1 #K1 #HLK1 #H
+ [ elim (lleq_inv_bind … H)
+ | elim (lleq_inv_flat … H)
+ ] -H /2 width=4 by fqu_pair_sn, ex3_intro/
+| #a #I #G1 #L1 #V1 #T1 #K1 #HLK1 #H elim (lleq_inv_bind_O … H) -H
+ /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 #k #HL1 #HTU1 #K1 #H1KL1 #H2KL1
+ elim (drop_O1_le (Ⓕ) (k+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
+ /3 width=4 by fqu_drop, ex3_intro/
+ | lapply (drop_fwd_length_le2 … HL1) -L -T1 -o
+ lapply (lleq_fwd_length … H2KL1) //
+ ]
+]
+qed-.
+
+lemma lpxs_lleq_fquq_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
+ ∀K1. ⦃G1, K1⦄ ⊢ ➡*[h, o] L1 → K1 ≡[T1, 0] L1 →
+ ∃∃K2. ⦃G1, K1, T1⦄ ⊐⸮ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡*[h, o] L2 & K2 ≡[T2, 0] L2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #K1 #H1KL1 #H2KL1
+elim (fquq_inv_gen … H) -H
+[ #H elim (lpxs_lleq_fqu_trans … H … H1KL1 H2KL1) -L1
+ /3 width=4 by fqu_fquq, ex3_intro/
+| * #HG #HL #HT destruct /2 width=4 by ex3_intro/
+]
+qed-.
+
+lemma lpxs_lleq_fqup_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
+ ∀K1. ⦃G1, K1⦄ ⊢ ➡*[h, o] L1 → K1 ≡[T1, 0] L1 →
+ ∃∃K2. ⦃G1, K1, T1⦄ ⊐+ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡*[h, o] L2 & K2 ≡[T2, 0] L2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
+[ #G2 #L2 #T2 #H #K1 #H1KL1 #H2KL1 elim (lpxs_lleq_fqu_trans … H … H1KL1 H2KL1) -L1
+ /3 width=4 by fqu_fqup, ex3_intro/
+| #G #G2 #L #L2 #T #T2 #_ #HT2 #IHT1 #K1 #H1KL1 #H2KL1 elim (IHT1 … H2KL1) // -L1
+ #K #HT1 #H1KL #H2KL elim (lpxs_lleq_fqu_trans … HT2 … H1KL H2KL) -L
+ /3 width=5 by fqup_strap1, ex3_intro/
+]
+qed-.
+
+lemma lpxs_lleq_fqus_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
+ ∀K1. ⦃G1, K1⦄ ⊢ ➡*[h, o] L1 → K1 ≡[T1, 0] L1 →
+ ∃∃K2. ⦃G1, K1, T1⦄ ⊐* ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡*[h, o] L2 & K2 ≡[T2, 0] L2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #K1 #H1KL1 #H2KL1
+elim (fqus_inv_gen … H) -H
+[ #H elim (lpxs_lleq_fqup_trans … H … H1KL1 H2KL1) -L1
+ /3 width=4 by fqup_fqus, ex3_intro/
+| * #HG #HL #HT destruct /2 width=4 by ex3_intro/
+]
+qed-.
+
+fact lreq_lpxs_trans_lleq_aux: ∀h,o,G,L1,L0,l,k. L1 ⩬[l, k] L0 → k = ∞ →
+ ∀L2. ⦃G, L0⦄ ⊢ ➡*[h, o] L2 →
+ ∃∃L. L ⩬[l, k] L2 & ⦃G, L1⦄ ⊢ ➡*[h, o] L &
+ (∀T. L0 ≡[T, l] L2 ↔ L1 ≡[T, l] L).
+#h #o #G #L1 #L0 #l #k #H elim H -L1 -L0 -l -k
+[ #l #k #_ #L2 #H >(lpxs_inv_atom1 … H) -H
+ /3 width=5 by ex3_intro, conj/
+| #I1 #I0 #L1 #L0 #V1 #V0 #_ #_ #Hm destruct
+| #I #L1 #L0 #V1 #k #HL10 #IHL10 #Hm #Y #H
+ elim (lpxs_inv_pair1 … H) -H #L2 #V2 #HL02 #HV02 #H destruct
+ 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, lreq_cpxs_trans, lreq_pair/
+ #T elim (IH T) #HL0dx #HL0sn
+ @conj #H @(lleq_lreq_repl … H) -H /3 width=1 by lreq_sym, lreq_pair_O_Y/
+| #I1 #I0 #L1 #L0 #V1 #V0 #l #k #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, lreq_succ/
+ #T elim (IH T) #HL0dx #HL0sn
+ @conj #H @(lleq_lreq_repl … H) -H /3 width=1 by lreq_sym, lreq_succ/
+]
+qed-.
+
+lemma lreq_lpxs_trans_lleq: ∀h,o,G,L1,L0,l. L1 ⩬[l, ∞] L0 →
+ ∀L2. ⦃G, L0⦄ ⊢ ➡*[h, o] L2 →
+ ∃∃L. L ⩬[l, ∞] L2 & ⦃G, L1⦄ ⊢ ➡*[h, o] L &
+ (∀T. L0 ≡[T, l] L2 ↔ L1 ≡[T, l] L).
+/2 width=1 by lreq_lpxs_trans_lleq_aux/ qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/computation/lpxs.ma".
+
+(* SN EXTENDED PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS *******************)
+
+(* Main properties **********************************************************)
+
+theorem lpxs_trans: ∀h,o,G. Transitive … (lpxs h o G).
+/2 width=3 by trans_TC/ qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/lrsubeqc_4.ma".
+include "basic_2/static/lsubr.ma".
+include "basic_2/static/aaa.ma".
+include "basic_2/computation/gcp_cr.ma".
+
+(* LOCAL ENVIRONMENT REFINEMENT FOR GENERIC REDUCIBILITY ********************)
+
+inductive lsubc (RP) (G): relation lenv ≝
+| lsubc_atom: lsubc RP G (⋆) (⋆)
+| lsubc_pair: ∀I,L1,L2,V. lsubc RP G L1 L2 → lsubc RP G (L1.ⓑ{I}V) (L2.ⓑ{I}V)
+| lsubc_beta: ∀L1,L2,V,W,A. ⦃G, L1, V⦄ ϵ[RP] 〚A〛 → ⦃G, L1, W⦄ ϵ[RP] 〚A〛 → ⦃G, L2⦄ ⊢ W ⁝ A →
+ lsubc RP G L1 L2 → lsubc RP G (L1. ⓓⓝW.V) (L2.ⓛW)
+.
+
+interpretation
+ "local environment refinement (generic reducibility)"
+ 'LRSubEqC RP G L1 L2 = (lsubc RP G L1 L2).
+
+(* Basic inversion lemmas ***************************************************)
+
+fact lsubc_inv_atom1_aux: ∀RP,G,L1,L2. G ⊢ L1 ⫃[RP] L2 → L1 = ⋆ → L2 = ⋆.
+#RP #G #L1 #L2 * -L1 -L2
+[ //
+| #I #L1 #L2 #V #_ #H destruct
+| #L1 #L2 #V #W #A #_ #_ #_ #_ #H destruct
+]
+qed-.
+
+(* Basic_1: was just: csubc_gen_sort_r *)
+lemma lsubc_inv_atom1: ∀RP,G,L2. G ⊢ ⋆ ⫃[RP] L2 → L2 = ⋆.
+/2 width=5 by lsubc_inv_atom1_aux/ qed-.
+
+fact lsubc_inv_pair1_aux: ∀RP,G,L1,L2. G ⊢ L1 ⫃[RP] L2 → ∀I,K1,X. L1 = K1.ⓑ{I}X →
+ (∃∃K2. G ⊢ K1 ⫃[RP] K2 & L2 = K2.ⓑ{I}X) ∨
+ ∃∃K2,V,W,A. ⦃G, K1, V⦄ ϵ[RP] 〚A〛 & ⦃G, K1, W⦄ ϵ[RP] 〚A〛 & ⦃G, K2⦄ ⊢ W ⁝ A &
+ G ⊢ K1 ⫃[RP] K2 &
+ L2 = K2. ⓛW & X = ⓝW.V & I = Abbr.
+#RP #G #L1 #L2 * -L1 -L2
+[ #I #K1 #V #H destruct
+| #J #L1 #L2 #V #HL12 #I #K1 #W #H destruct /3 width=3 by ex2_intro, or_introl/
+| #L1 #L2 #V1 #W2 #A #HV1 #H1W2 #H2W2 #HL12 #I #K1 #V #H destruct /3 width=10 by ex7_4_intro, or_intror/
+]
+qed-.
+
+(* Basic_1: was: csubc_gen_head_r *)
+lemma lsubc_inv_pair1: ∀RP,I,G,K1,L2,X. G ⊢ K1.ⓑ{I}X ⫃[RP] L2 →
+ (∃∃K2. G ⊢ K1 ⫃[RP] K2 & L2 = K2.ⓑ{I}X) ∨
+ ∃∃K2,V,W,A. ⦃G, K1, V⦄ ϵ[RP] 〚A〛 & ⦃G, K1, W⦄ ϵ[RP] 〚A〛 & ⦃G, K2⦄ ⊢ W ⁝ A &
+ G ⊢ K1 ⫃[RP] K2 &
+ L2 = K2.ⓛW & X = ⓝW.V & I = Abbr.
+/2 width=3 by lsubc_inv_pair1_aux/ qed-.
+
+fact lsubc_inv_atom2_aux: ∀RP,G,L1,L2. G ⊢ L1 ⫃[RP] L2 → L2 = ⋆ → L1 = ⋆.
+#RP #G #L1 #L2 * -L1 -L2
+[ //
+| #I #L1 #L2 #V #_ #H destruct
+| #L1 #L2 #V #W #A #_ #_ #_ #_ #H destruct
+]
+qed-.
+
+(* Basic_1: was just: csubc_gen_sort_l *)
+lemma lsubc_inv_atom2: ∀RP,G,L1. G ⊢ L1 ⫃[RP] ⋆ → L1 = ⋆.
+/2 width=5 by lsubc_inv_atom2_aux/ qed-.
+
+fact lsubc_inv_pair2_aux: ∀RP,G,L1,L2. G ⊢ L1 ⫃[RP] L2 → ∀I,K2,W. L2 = K2.ⓑ{I} W →
+ (∃∃K1. G ⊢ K1 ⫃[RP] K2 & L1 = K1. ⓑ{I} W) ∨
+ ∃∃K1,V,A. ⦃G, K1, V⦄ ϵ[RP] 〚A〛 & ⦃G, K1, W⦄ ϵ[RP] 〚A〛 & ⦃G, K2⦄ ⊢ W ⁝ A &
+ G ⊢ K1 ⫃[RP] K2 &
+ L1 = K1.ⓓⓝW.V & I = Abst.
+#RP #G #L1 #L2 * -L1 -L2
+[ #I #K2 #W #H destruct
+| #J #L1 #L2 #V #HL12 #I #K2 #W #H destruct /3 width=3 by ex2_intro, or_introl/
+| #L1 #L2 #V1 #W2 #A #HV1 #H1W2 #H2W2 #HL12 #I #K2 #W #H destruct /3 width=8 by ex6_3_intro, or_intror/
+]
+qed-.
+
+(* Basic_1: was just: csubc_gen_head_l *)
+lemma lsubc_inv_pair2: ∀RP,I,G,L1,K2,W. G ⊢ L1 ⫃[RP] K2.ⓑ{I} W →
+ (∃∃K1. G ⊢ K1 ⫃[RP] K2 & L1 = K1.ⓑ{I} W) ∨
+ ∃∃K1,V,A. ⦃G, K1, V⦄ ϵ[RP] 〚A〛 & ⦃G, K1, W⦄ ϵ[RP] 〚A〛 & ⦃G, K2⦄ ⊢ W ⁝ A &
+ G ⊢ K1 ⫃[RP] K2 &
+ L1 = K1.ⓓⓝW.V & I = Abst.
+/2 width=3 by lsubc_inv_pair2_aux/ qed-.
+
+(* Basic forward lemmas *****************************************************)
+
+lemma lsubc_fwd_lsubr: ∀RP,G,L1,L2. G ⊢ L1 ⫃[RP] L2 → L1 ⫃ L2.
+#RP #G #L1 #L2 #H elim H -L1 -L2 /2 width=1 by lsubr_pair, lsubr_beta/
+qed-.
+
+(* Basic properties *********************************************************)
+
+(* Basic_1: was just: csubc_refl *)
+lemma lsubc_refl: ∀RP,G,L. G ⊢ L ⫃[RP] L.
+#RP #G #L elim L -L /2 width=1 by lsubc_pair/
+qed.
+
+(* Basic_1: removed theorems 3:
+ csubc_clear_conf csubc_getl_conf csubc_csuba
+*)
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/static/aaa_lift.ma".
+include "basic_2/computation/lsubc.ma".
+
+(* LOCAL ENVIRONMENT REFINEMENT FOR GENERIC REDUCIBILITY ********************)
+
+(* Properties concerning basic local environment slicing ********************)
+
+(* Basic_1: was: csubc_drop_conf_O *)
+(* Note: the constant 0 can not be generalized *)
+lemma lsubc_drop_O1_trans: ∀RP,G,L1,L2. G ⊢ L1 ⫃[RP] L2 → ∀K2,c,k. ⬇[c, 0, k] L2 ≡ K2 →
+ ∃∃K1. ⬇[c, 0, k] L1 ≡ K1 & G ⊢ K1 ⫃[RP] K2.
+#RP #G #L1 #L2 #H elim H -L1 -L2
+[ #X #c #k #H elim (drop_inv_atom1 … H) -H /4 width=3 by drop_atom, ex2_intro/
+| #I #L1 #L2 #V #_ #IHL12 #X #c #k #H
+ elim (drop_inv_O1_pair1 … H) -H * #Hm #H destruct
+ [ elim (IHL12 L2 c 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 #c #k #H
+ elim (drop_inv_O1_pair1 … H) -H * #Hm #H destruct
+ [ elim (IHL12 L2 c 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/
+ ]
+]
+qed-.
+
+(* Basic_1: was: csubc_drop_conf_rev *)
+lemma drop_lsubc_trans: ∀RR,RS,RP. gcp RR RS RP →
+ ∀G,L1,K1,l,k. ⬇[Ⓕ, l, k] L1 ≡ K1 → ∀K2. G ⊢ K1 ⫃[RP] K2 →
+ ∃∃L2. G ⊢ L1 ⫃[RP] L2 & ⬇[Ⓕ, l, k] L2 ≡ K2.
+#RR #RS #RP #Hgcp #G #L1 #K1 #l #k #H elim H -L1 -K1 -l -k
+[ #l #k #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 #k #_ #IHLK1 #K2 #HK12
+ elim (IHLK1 … HK12) -K1 /3 width=5 by lsubc_pair, drop_drop, ex2_intro/
+| #I #L1 #K1 #V1 #V2 #l #k #HLK1 #HV21 #IHLK1 #X #H
+ elim (lsubc_inv_pair1 … H) -H *
+ [ #K2 #HK12 #H destruct
+ elim (IHLK1 … HK12) -K1 /3 width=5 by lsubc_pair, drop_skip, ex2_intro/
+ | #K2 #V #W2 #A #HV2 #H1W2 #H2W2 #HK12 #H1 #H2 #H3 destruct
+ elim (lift_inv_flat1 … HV21) -HV21 #W3 #V3 #HW23 #HV3 #H destruct
+ elim (IHLK1 … HK12) #K #HL1K #HK2
+ lapply (gcr_lift … Hgcp … HV2 … HLK1 … HV3) -HV2
+ lapply (gcr_lift … Hgcp … H1W2 … HLK1 … HW23) -H1W2
+ /4 width=11 by lsubc_beta, aaa_lift, drop_skip, ex2_intro/
+ ]
+]
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/computation/lsubc_drop.ma".
+
+(* LOCAL ENVIRONMENT REFINEMENT FOR GENERIC REDUCIBILITY ********************)
+
+(* Properties concerning generic local environment slicing ******************)
+
+(* Basic_1: was: csubc_drop1_conf_rev *)
+lemma drops_lsubc_trans: ∀RR,RS,RP. gcp RR RS RP →
+ ∀G,L1,K1,cs. ⬇*[Ⓕ, cs] L1 ≡ K1 → ∀K2. G ⊢ K1 ⫃[RP] K2 →
+ ∃∃L2. G ⊢ L1 ⫃[RP] L2 & ⬇*[Ⓕ, cs] L2 ≡ K2.
+#RR #RS #RP #Hgcp #G #L1 #K1 #cs #H elim H -L1 -K1 -cs
+[ /2 width=3 by drops_nil, ex2_intro/
+| #L1 #L #K1 #cs #l #k #_ #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/
+]
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/static/lsuba.ma".
+include "basic_2/computation/gcp_aaa.ma".
+
+(* LOCAL ENVIRONMENT REFINEMENT FOR GENERIC REDUCIBILITY ********************)
+
+(* properties concerning lenv refinement for atomic arity assignment ********)
+
+lemma lsuba_lsubc: ∀RR,RS,RP. gcp RR RS RP → gcr RR RS RP RP →
+ ∀G,L1,L2. G ⊢ L1 ⫃⁝ L2 → G ⊢ L1 ⫃[RP] L2.
+#RR #RS #RP #H1RP #H2RP #G #L1 #L2 #H elim H -L1 -L2 /2 width=1 by lsubc_pair/
+#L1 #L2 #W #V #A #H elim (aaa_inv_cast … H) -H /3 width=4 by acr_aaa, lsubc_beta/
+qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/sn_6.ma".
+include "basic_2/multiple/lleq.ma".
+include "basic_2/reduction/lpx.ma".
+
+(* SN EXTENDED STRONGLY NORMALIZING LOCAL ENVIRONMENTS **********************)
+
+definition lsx: ∀h. sd h → relation4 ynat term genv lenv ≝
+ λh,o,l,T,G. SN … (lpx h o G) (lleq l T).
+
+interpretation
+ "extended strong normalization (local environment)"
+ 'SN h o l T G L = (lsx h o T l G L).
+
+(* Basic eliminators ********************************************************)
+
+lemma lsx_ind: ∀h,o,G,T,l. ∀R:predicate lenv.
+ (∀L1. G ⊢ ⬊*[h, o, T, l] L1 →
+ (∀L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 → (L1 ≡[T, l] L2 → ⊥) → R L2) →
+ R L1
+ ) →
+ ∀L. G ⊢ ⬊*[h, o, T, l] L → R L.
+#h #o #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,o,G,L1,T,l.
+ (∀L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 → (L1 ≡[T, l] L2 → ⊥) → G ⊢ ⬊*[h, o, T, l] L2) →
+ G ⊢ ⬊*[h, o, T, l] L1.
+/5 width=1 by lleq_sym, SN_intro/ qed.
+
+lemma lsx_atom: ∀h,o,G,T,l. G ⊢ ⬊*[h, o, T, l] ⋆.
+#h #o #G #T #l @lsx_intro
+#X #H #HT lapply (lpx_inv_atom1 … H) -H
+#H destruct elim HT -HT //
+qed.
+
+lemma lsx_sort: ∀h,o,G,L,l,s. G ⊢ ⬊*[h, o, ⋆s, l] L.
+#h #o #G #L1 #l #s @lsx_intro
+#L2 #HL12 #H elim H -H
+/3 width=4 by lpx_fwd_length, lleq_sort/
+qed.
+
+lemma lsx_gref: ∀h,o,G,L,l,p. G ⊢ ⬊*[h, o, §p, l] L.
+#h #o #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,o,G,L,T,U,lt,l,k. lt ≤ yinj l + yinj k →
+ ⬆[l, k] T ≡ U → G ⊢ ⬊*[h, o, U, lt] L → G ⊢ ⬊*[h, o, U, l] L.
+#h #o #G #L #T #U #lt #l #k #Hltlm #HTU #H @(lsx_ind … H) -L
+/5 width=7 by lsx_intro, lleq_ge_up/
+qed-.
+
+lemma lsx_ge: ∀h,o,G,L,T,l1,l2. l1 ≤ l2 →
+ G ⊢ ⬊*[h, o, T, l1] L → G ⊢ ⬊*[h, o, T, l2] L.
+#h #o #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,o,a,I,G,L,V,T,l. G ⊢ ⬊*[h, o, ⓑ{a,I}V.T, l] L →
+ G ⊢ ⬊*[h, o, V, l] L.
+#h #o #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,o,I,G,L,V,T,l. G ⊢ ⬊*[h, o, ⓕ{I}V.T, l] L →
+ G ⊢ ⬊*[h, o, V, l] L.
+#h #o #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,o,I,G,L,V,T,l. G ⊢ ⬊*[h, o, ⓕ{I}V.T, l] L →
+ G ⊢ ⬊*[h, o, T, l] L.
+#h #o #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,o,I,G,L,V,T,l. G ⊢ ⬊*[h, o, ②{I}V.T, l] L →
+ G ⊢ ⬊*[h, o, V, l] L.
+#h #o * /2 width=4 by lsx_fwd_bind_sn, lsx_fwd_flat_sn/
+qed-.
+
+(* Basic inversion lemmas ***************************************************)
+
+lemma lsx_inv_flat: ∀h,o,I,G,L,V,T,l. G ⊢ ⬊*[h, o, ⓕ{I}V.T, l] L →
+ G ⊢ ⬊*[h, o, V, l] L ∧ G ⊢ ⬊*[h, o, T, l] L.
+/3 width=3 by lsx_fwd_flat_sn, lsx_fwd_flat_dx, conj/ qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/snalt_6.ma".
+include "basic_2/computation/lpxs_lleq.ma".
+include "basic_2/computation/lsx.ma".
+
+(* SN EXTENDED STRONGLY NORMALIZING LOCAL ENVIRONMENTS **********************)
+
+(* alternative definition of lsx *)
+definition lsxa: ∀h. sd h → relation4 ynat term genv lenv ≝
+ λh,o,l,T,G. SN … (lpxs h o G) (lleq l T).
+
+interpretation
+ "extended strong normalization (local environment) alternative"
+ 'SNAlt h o l T G L = (lsxa h o T l G L).
+
+(* Basic eliminators ********************************************************)
+
+lemma lsxa_ind: ∀h,o,G,T,l. ∀R:predicate lenv.
+ (∀L1. G ⊢ ⬊⬊*[h, o, T, l] L1 →
+ (∀L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → (L1 ≡[T, l] L2 → ⊥) → R L2) →
+ R L1
+ ) →
+ ∀L. G ⊢ ⬊⬊*[h, o, T, l] L → R L.
+#h #o #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,o,G,L1,T,l.
+ (∀L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → (L1 ≡[T, l] L2 → ⊥) → G ⊢ ⬊⬊*[h, o, T, l] L2) →
+ G ⊢ ⬊⬊*[h, o, T, l] L1.
+/5 width=1 by lleq_sym, SN_intro/ qed.
+
+fact lsxa_intro_aux: ∀h,o,G,L1,T,l.
+ (∀L,L2. ⦃G, L⦄ ⊢ ➡*[h, o] L2 → L1 ≡[T, l] L → (L1 ≡[T, l] L2 → ⊥) → G ⊢ ⬊⬊*[h, o, T, l] L2) →
+ G ⊢ ⬊⬊*[h, o, T, l] L1.
+/4 width=3 by lsxa_intro/ qed-.
+
+lemma lsxa_lleq_trans: ∀h,o,T,G,L1,l. G ⊢ ⬊⬊*[h, o, T, l] L1 →
+ ∀L2. L1 ≡[T, l] L2 → G ⊢ ⬊⬊*[h, o, T, l] L2.
+#h #o #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,o,T,G,L1,l. G ⊢ ⬊⬊*[h, o, T, l] L1 →
+ ∀L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → G ⊢ ⬊⬊*[h, o, T, l] L2.
+#h #o #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,o,G,L1,T,l.
+ (∀L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 → (L1 ≡[T, l] L2 → ⊥) → G ⊢ ⬊⬊*[h, o, T, l] L2) →
+ G ⊢ ⬊⬊*[h, o, T, l] L1.
+#h #o #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 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/
+]
+qed-.
+
+(* Main properties **********************************************************)
+
+theorem lsx_lsxa: ∀h,o,G,L,T,l. G ⊢ ⬊*[h, o, T, l] L → G ⊢ ⬊⬊*[h, o, T, l] L.
+#h #o #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,o,G,L,T,l. G ⊢ ⬊⬊*[h, o, T, l] L → G ⊢ ⬊*[h, o, T, l] L.
+#h #o #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,o,G,L1,T,l.
+ (∀L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → (L1 ≡[T, l] L2 → ⊥) → G ⊢ ⬊*[h, o, T, l] L2) →
+ G ⊢ ⬊*[h, o, T, l] L1.
+/6 width=1 by lsxa_inv_lsx, lsx_lsxa, lsxa_intro/ qed.
+
+lemma lsx_lpxs_trans: ∀h,o,G,L1,T,l. G ⊢ ⬊*[h, o, T, l] L1 →
+ ∀L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → G ⊢ ⬊*[h, o, T, l] L2.
+/4 width=3 by lsxa_inv_lsx, lsx_lsxa, lsxa_lpxs_trans/ qed-.
+
+(* Advanced eliminators *****************************************************)
+
+lemma lsx_ind_alt: ∀h,o,G,T,l. ∀R:predicate lenv.
+ (∀L1. G ⊢ ⬊*[h, o, T, l] L1 →
+ (∀L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → (L1 ≡[T, l] L2 → ⊥) → R L2) →
+ R L1
+ ) →
+ ∀L. G ⊢ ⬊*[h, o, T, l] L → R L.
+#h #o #G #T #l #R #IH #L #H @(lsxa_ind h o G T l … L)
+/4 width=1 by lsxa_inv_lsx, lsx_lsxa/
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/computation/csx_lpxs.ma".
+include "basic_2/computation/lcosx_cpx.ma".
+
+(* SN EXTENDED STRONGLY NORMALIZING LOCAL ENVIRONMENTS **********************)
+
+(* Advanced properties ******************************************************)
+
+lemma lsx_lref_be_lpxs: ∀h,o,I,G,K1,V,i,l. l ≤ yinj i → ⦃G, K1⦄ ⊢ ⬊*[h, o] V →
+ ∀K2. G ⊢ ⬊*[h, o, V, 0] K2 → ⦃G, K1⦄ ⊢ ➡*[h, o] K2 →
+ ∀L2. ⬇[i] L2 ≡ K2.ⓑ{I}V → G ⊢ ⬊*[h, o, #i, l] L2.
+#h #o #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
+#Y #H #HLK2 elim (lpx_inv_pair1 … H) -H
+#K2 #V2 #HK02 #HV02 #H destruct
+elim (eq_term_dec V0 V2) #HnV02 destruct [ -IHV0 -HV02 -HK0 | -IHK0 -HnL02 -HLK0 ]
+[ /4 width=8 by lpxs_strap1, lleq_lref/
+| @(IHV0 … HnV02 … HLK2) -IHV0 -HnV02 -HLK2
+ /3 width=4 by lsx_cpx_trans_O, lsx_lpx_trans, lpxs_cpx_trans, lpxs_strap1/ (**) (* full auto too slow *)
+]
+qed.
+
+lemma lsx_lref_be: ∀h,o,I,G,K,V,i,l. l ≤ yinj i → ⦃G, K⦄ ⊢ ⬊*[h, o] V →
+ G ⊢ ⬊*[h, o, V, 0] K →
+ ∀L. ⬇[i] L ≡ K.ⓑ{I}V → G ⊢ ⬊*[h, o, #i, l] L.
+/2 width=8 by lsx_lref_be_lpxs/ qed.
+
+(* Main properties **********************************************************)
+
+theorem csx_lsx: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → ∀l. G ⊢ ⬊*[h, o, T, l] L.
+#h #o #G #L #T @(fqup_wf_ind_eq … G L T) -G -L -T
+#Z #Y #X #IH #G #L * * //
+[ #i #HG #HL #HT #H #l destruct
+ elim (lt_or_ge i (|L|)) /2 width=1 by lsx_lref_free/
+ 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 #l destruct
+ elim (csx_fwd_bind … H) -H /3 width=1 by lsx_bind/
+| #I #V #T #HG #HL #HT #H #l destruct
+ elim (csx_fwd_flat … H) -H /3 width=1 by lsx_flat/
+]
+qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/multiple/lleq_drop.ma".
+include "basic_2/reduction/lpx_drop.ma".
+include "basic_2/computation/lsx.ma".
+
+(* SN EXTENDED STRONGLY NORMALIZING LOCAL ENVIRONMENTS **********************)
+
+(* Advanced properties ******************************************************)
+
+lemma lsx_lref_free: ∀h,o,G,L,l,i. |L| ≤ i → G ⊢ ⬊*[h, o, #i, l] L.
+#h #o #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,o,G,L,l,i. yinj i < l → G ⊢ ⬊*[h, o, #i, l] L.
+#h #o #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,o,I,G,L,l,i. l ≤ yinj i → G ⊢ ⬊*[h, o, #i, l] L →
+ ∀K,V. ⬇[i] L ≡ K.ⓑ{I}V → G ⊢ ⬊*[h, o, V, 0] K.
+#h #o #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 (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/
+qed-.
+
+(* Properties on relocation *************************************************)
+
+lemma lsx_lift_le: ∀h,o,G,K,T,U,lt,l,k. lt ≤ yinj l →
+ ⬆[l, k] T ≡ U → G ⊢ ⬊*[h, o, T, lt] K →
+ ∀L. ⬇[Ⓕ, l, k] L ≡ K → G ⊢ ⬊*[h, o, U, lt] L.
+#h #o #G #K #T #U #lt #l #k #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,o,G,K,T,U,lt,l,k. yinj l ≤ lt →
+ ⬆[l, k] T ≡ U → G ⊢ ⬊*[h, o, T, lt] K →
+ ∀L. ⬇[Ⓕ, l, k] L ≡ K → G ⊢ ⬊*[h, o, U, lt + k] L.
+#h #o #G #K #T #U #lt #l #k #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/
+qed-.
+
+(* Inversion lemmas on relocation *******************************************)
+
+lemma lsx_inv_lift_le: ∀h,o,G,L,T,U,lt,l,k. lt ≤ yinj l →
+ ⬆[l, k] T ≡ U → G ⊢ ⬊*[h, o, U, lt] L →
+ ∀K. ⬇[Ⓕ, l, k] L ≡ K → G ⊢ ⬊*[h, o, T, lt] K.
+#h #o #G #L #T #U #lt #l #k #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,o,G,L,T,U,lt,l,k. yinj l ≤ lt → lt ≤ l + k →
+ ⬆[l, k] T ≡ U → G ⊢ ⬊*[h, o, U, lt] L →
+ ∀K. ⬇[Ⓕ, l, k] L ≡ K → G ⊢ ⬊*[h, o, T, l] K.
+#h #o #G #L #T #U #lt #l #k #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,o,G,L,T,U,lt,l,k. yinj l + yinj k ≤ lt →
+ ⬆[l, k] T ≡ U → G ⊢ ⬊*[h, o, U, lt] L →
+ ∀K. ⬇[Ⓕ, l, k] L ≡ K → G ⊢ ⬊*[h, o, T, lt-k] K.
+#h #o #G #L #T #U #lt #l #k #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/
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/multiple/lleq_lleq.ma".
+include "basic_2/reduction/lpx_lleq.ma".
+include "basic_2/computation/lsx.ma".
+
+(* SN EXTENDED STRONGLY NORMALIZING LOCAL ENVIRONMENTS **********************)
+
+(* Advanced properties ******************************************************)
+
+lemma lsx_lleq_trans: ∀h,o,T,G,L1,l. G ⊢ ⬊*[h, o, T, l] L1 →
+ ∀L2. L1 ≡[T, l] L2 → G ⊢ ⬊*[h, o, T, l] L2.
+#h #o #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,o,T,G,L1,l. G ⊢ ⬊*[h, o, T, l] L1 →
+ ∀L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 → G ⊢ ⬊*[h, o, T, l] L2.
+#h #o #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_lreq_conf: ∀h,o,G,L1,T,l. G ⊢ ⬊*[h, o, T, l] L1 →
+ ∀L2. L1 ⩬[l, ∞] L2 → G ⊢ ⬊*[h, o, T, l] L2.
+#h #o #G #L1 #T #l #H @(lsx_ind … H) -L1
+#L1 #_ #IHL1 #L2 #HL12 @lsx_intro
+#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,o,a,I,G,L,V,T,l. G ⊢ ⬊*[h, o, ⓑ{a,I}V.T, l] L →
+ G ⊢ ⬊*[h, o, T, ⫯l] L.ⓑ{I}V.
+#h #o #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_lreq_conf … (L2.ⓑ{I}V1)) /2 width=1 by lreq_succ/
+@IHL1 // #H @HT -IHL1 -HL12 -HT
+@(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,o,a,I,G,L,V,T,l. G ⊢ ⬊*[h, o, ⓑ{a, I}V.T, l] L →
+ G ⊢ ⬊*[h, o, V, l] L ∧ G ⊢ ⬊*[h, o, T, ⫯l] L.ⓑ{I}V.
+/3 width=4 by lsx_fwd_bind_sn, lsx_fwd_bind_dx, conj/ qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/computation/lpxs_lpxs.ma".
+include "basic_2/computation/lsx_alt.ma".
+
+(* SN EXTENDED STRONGLY NORMALIZING LOCAL ENVIRONMENTS **********************)
+
+(* Advanced properties ******************************************************)
+
+fact lsx_bind_lpxs_aux: ∀h,o,a,I,G,L1,V,l. G ⊢ ⬊*[h, o, V, l] L1 →
+ ∀Y,T. G ⊢ ⬊*[h, o, T, ⫯l] Y →
+ ∀L2. Y = L2.ⓑ{I}V → ⦃G, L1⦄ ⊢ ➡*[h, o] L2 →
+ G ⊢ ⬊*[h, o, ⓑ{a,I}V.T, l] L2.
+#h #o #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 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/
+ ]
+| /3 width=4 by lpxs_pair/
+]
+qed-.
+
+lemma lsx_bind: ∀h,o,a,I,G,L,V,l. G ⊢ ⬊*[h, o, V, l] L →
+ ∀T. G ⊢ ⬊*[h, o, T, ⫯l] L.ⓑ{I}V →
+ G ⊢ ⬊*[h, o, ⓑ{a,I}V.T, l] L.
+/2 width=3 by lsx_bind_lpxs_aux/ qed.
+
+lemma lsx_flat_lpxs: ∀h,o,I,G,L1,V,l. G ⊢ ⬊*[h, o, V, l] L1 →
+ ∀L2,T. G ⊢ ⬊*[h, o, T, l] L2 → ⦃G, L1⦄ ⊢ ➡*[h, o] L2 →
+ G ⊢ ⬊*[h, o, ⓕ{I}V.T, l] L2.
+#h #o #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 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/
+ ]
+| /3 width=1 by/
+]
+qed-.
+
+lemma lsx_flat: ∀h,o,I,G,L,V,l. G ⊢ ⬊*[h, o, V, l] L →
+ ∀T. G ⊢ ⬊*[h, o, T, l] L → G ⊢ ⬊*[h, o, ⓕ{I}V.T, l] L.
+/2 width=3 by lsx_flat_lpxs/ qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/dpredstar_7.ma".
+include "basic_2/static/da.ma".
+include "basic_2/computation/cprs.ma".
+
+(* STRATIFIED DECOMPOSED PARALLEL COMPUTATION ON TERMS **********************)
+
+definition scpds: ∀h. sd h → nat → relation4 genv lenv term term ≝
+ λh,o,d2,G,L,T1,T2.
+ ∃∃T,d1. d2 ≤ d1 & ⦃G, L⦄ ⊢ T1 ▪[h, o] d1 & ⦃G, L⦄ ⊢ T1 •*[h, d2] T & ⦃G, L⦄ ⊢ T ➡* T2.
+
+interpretation "stratified decomposed parallel computation (term)"
+ 'DPRedStar h o d G L T1 T2 = (scpds h o d G L T1 T2).
+
+(* Basic properties *********************************************************)
+
+lemma sta_cprs_scpds: ∀h,o,G,L,T1,T,T2,d. ⦃G, L⦄ ⊢ T1 ▪[h, o] d+1 → ⦃G, L⦄ ⊢ T1 •*[h, 1] T →
+ ⦃G, L⦄ ⊢ T ➡* T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, o, 1] T2.
+/2 width=6 by ex4_2_intro/ qed.
+
+lemma lstas_scpds: ∀h,o,G,L,T1,T2,d1. ⦃G, L⦄ ⊢ T1 ▪[h, o] d1 →
+ ∀d2. d2 ≤ d1 → ⦃G, L⦄ ⊢ T1 •*[h, d2] T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d2] T2.
+/2 width=6 by ex4_2_intro/ qed.
+
+lemma scpds_strap1: ∀h,o,G,L,T1,T,T2,d.
+ ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d] T → ⦃G, L⦄ ⊢ T ➡ T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d] T2.
+#h #o #G #L #T1 #T #T2 #d * /3 width=8 by cprs_strap1, ex4_2_intro/
+qed.
+
+(* Basic forward lemmas *****************************************************)
+
+lemma scpds_fwd_cprs: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 •*➡*[h, o, 0] T2 →
+ ⦃G, L⦄ ⊢ T1 ➡* T2.
+#h #o #G #L #T1 #T2 * /3 width=3 by cprs_strap2, lstas_cpr/
+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/unfold/lstas_aaa.ma".
+include "basic_2/computation/cpxs_aaa.ma".
+include "basic_2/computation/scpds.ma".
+
+(* STRATIFIED DECOMPOSED PARALLEL COMPUTATION ON TERMS **********************)
+
+(* Properties on atomic arity assignment for terms **************************)
+
+lemma scpds_aaa_conf: ∀h,o,G,L,d. Conf3 … (aaa G L) (scpds h o d G L).
+#h #o #G #L #d #A #T #HT #U * /3 width=6 by lstas_aaa_conf, cprs_aaa_conf/
+qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/static/da_lift.ma".
+include "basic_2/unfold/lstas_lift.ma".
+include "basic_2/computation/cprs_lift.ma".
+include "basic_2/computation/scpds.ma".
+
+(* STRATIFIED DECOMPOSED PARALLEL COMPUTATION ON TERMS **********************)
+
+(* Relocation properties ****************************************************)
+
+lemma scpds_lift: ∀h,o,G,d. d_liftable (scpds h o d G).
+#h #o #G #d2 #K #T1 #T2 * #T #d1 #Hd21 #Hd1 #HT1 #HT2 #L #c #l #k
+elim (lift_total T l k)
+/3 width=15 by cprs_lift, da_lift, lstas_lift, ex4_2_intro/
+qed.
+
+lemma scpds_inv_lift1: ∀h,o,G,d. d_deliftable_sn (scpds h o d G).
+#h #o #G #d2 #L #U1 #U2 * #U #d1 #Hd21 #Hd1 #HU1 #HU2 #K #c #l #k #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/
+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/unfold/lstas_da.ma".
+include "basic_2/computation/lprs_cprs.ma".
+include "basic_2/computation/cpxs_cpxs.ma".
+include "basic_2/computation/scpds.ma".
+
+(* STRATIFIED DECOMPOSED PARALLEL COMPUTATION ON TERMS **********************)
+
+(* Advanced properties ******************************************************)
+
+lemma scpds_strap2: ∀h,o,G,L,T1,T,T2,d1,d. ⦃G, L⦄ ⊢ T1 ▪[h, o] d1+1 →
+ ⦃G, L⦄ ⊢ T1 •*[h, 1] T → ⦃G, L⦄ ⊢ T •*➡*[h, o, d] T2 →
+ ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d+1] T2.
+#h #o #G #L #T1 #T #T2 #d1 #d #Hd1 #HT1 *
+#T0 #d0 #Hd0 #HTd0 #HT0 #HT02
+lapply (lstas_da_conf … HT1 … Hd1) <minus_plus_k_k #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,o,G,L,T1,T,T2,d.
+ ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d] T → ⦃G, L⦄ ⊢ T ➡* T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d] T2.
+#h #o #G #L #T1 #T #T2 #d * /3 width=8 by cprs_trans, ex4_2_intro/
+qed-.
+
+lemma lstas_scpds_trans: ∀h,o,G,L,T1,T,T2,d1,d2,d.
+ d2 ≤ d1 → ⦃G, L⦄ ⊢ T1 ▪[h, o] d1 →
+ ⦃G, L⦄ ⊢ T1 •*[h, d2] T → ⦃G, L⦄ ⊢ T •*➡*[h, o, d] T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d2+d] T2.
+#h #o #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,o,a,G,L,V1,T1,U2,d. ⦃G, L⦄ ⊢ ⓛ{a}V1.T1 •*➡*[h, o, d] U2 →
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡* V2 & ⦃G, L.ⓛV1⦄ ⊢ T1 •*➡*[h, o, d] T2 &
+ U2 = ⓛ{a}V2.T2.
+#h #o #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,o,a1,a2,G,L,V1,W2,T1,T2,d. ⦃G, L⦄ ⊢ ⓓ{a1}V1.T1 •*➡*[h, o, d] ⓛ{a2}W2.T2 →
+ ∃∃T. ⦃G, L.ⓓV1⦄ ⊢ T1 •*➡*[h, o, d] T & ⬆[0, 1] ⓛ{a2}W2.T2 ≡ T & a1 = true.
+#h #o #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
+| /3 width=6 by ex4_2_intro, ex3_intro/
+]
+qed-.
+
+lemma scpds_inv_lstas_eq: ∀h,o,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d] T2 →
+ ∀T. ⦃G, L⦄ ⊢ T1 •*[h, d] T → ⦃G, L⦄ ⊢ T ➡* T2.
+#h #o #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,o,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d] T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2.
+#h #o #G #L #T1 #T2 #d * /3 width=5 by cpxs_trans, lstas_cpxs, cprs_cpxs/
+qed-.
+
+(* Main properties **********************************************************)
+
+theorem scpds_conf_eq: ∀h,o,G,L,T0,T1,d. ⦃G, L⦄ ⊢ T0 •*➡*[h, o, d] T1 →
+ ∀T2. ⦃G, L⦄ ⊢ T0 •*➡*[h, o, d] T2 →
+ ∃∃T. ⦃G, L⦄ ⊢ T1 ➡* T & ⦃G, L⦄ ⊢ T2 ➡* T.
+#h #o #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-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/pred_4.ma".
+include "basic_2/static/lsubr.ma".
+include "basic_2/unfold/lstas.ma".
+
+(* CONTEXT-SENSITIVE PARALLEL REDUCTION FOR TERMS ***************************)
+
+(* activate genv *)
+(* Basic_1: includes: pr0_delta1 pr2_delta1 pr2_thin_dx *)
+(* Note: cpr_flat: does not hold in basic_1 *)
+inductive cpr: relation4 genv lenv term term ≝
+| cpr_atom : ∀I,G,L. cpr G L (⓪{I}) (⓪{I})
+| cpr_delta: ∀G,L,K,V,V2,W2,i.
+ ⬇[i] L ≡ K. ⓓV → cpr G K V V2 →
+ ⬆[0, i + 1] V2 ≡ W2 → cpr G L (#i) W2
+| cpr_bind : ∀a,I,G,L,V1,V2,T1,T2.
+ cpr G L V1 V2 → cpr G (L.ⓑ{I}V1) T1 T2 →
+ cpr G L (ⓑ{a,I}V1.T1) (ⓑ{a,I}V2.T2)
+| cpr_flat : ∀I,G,L,V1,V2,T1,T2.
+ cpr G L V1 V2 → cpr G L T1 T2 →
+ cpr G L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2)
+| cpr_zeta : ∀G,L,V,T1,T,T2. cpr G (L.ⓓV) T1 T →
+ ⬆[0, 1] T2 ≡ T → cpr G L (+ⓓV.T1) T2
+| cpr_eps : ∀G,L,V,T1,T2. cpr G L T1 T2 → cpr G L (ⓝV.T1) T2
+| cpr_beta : ∀a,G,L,V1,V2,W1,W2,T1,T2.
+ cpr G L V1 V2 → cpr G L W1 W2 → cpr G (L.ⓛW1) T1 T2 →
+ cpr G L (ⓐV1.ⓛ{a}W1.T1) (ⓓ{a}ⓝW2.V2.T2)
+| cpr_theta: ∀a,G,L,V1,V,V2,W1,W2,T1,T2.
+ cpr G L V1 V → ⬆[0, 1] V ≡ V2 → cpr G L W1 W2 → cpr G (L.ⓓW1) T1 T2 →
+ cpr G L (ⓐV1.ⓓ{a}W1.T1) (ⓓ{a}W2.ⓐV2.T2)
+.
+
+interpretation "context-sensitive parallel reduction (term)"
+ 'PRed G L T1 T2 = (cpr G L T1 T2).
+
+(* Basic properties *********************************************************)
+
+lemma lsubr_cpr_trans: ∀G. lsub_trans … (cpr G) lsubr.
+#G #L1 #T1 #T2 #H elim H -G -L1 -T1 -T2
+[ //
+| #G #L1 #K1 #V1 #V2 #W2 #i #HLK1 #_ #HVW2 #IHV12 #L2 #HL12
+ elim (lsubr_fwd_drop2_abbr … HL12 … HLK1) -L1 *
+ /3 width=6 by cpr_delta/
+|3,7: /4 width=1 by lsubr_pair, cpr_bind, cpr_beta/
+|4,6: /3 width=1 by cpr_flat, cpr_eps/
+|5,8: /4 width=3 by lsubr_pair, cpr_zeta, cpr_theta/
+]
+qed-.
+
+(* Basic_1: was by definition: pr2_free *)
+lemma tpr_cpr: ∀G,T1,T2. ⦃G, ⋆⦄ ⊢ T1 ➡ T2 → ∀L. ⦃G, L⦄ ⊢ T1 ➡ T2.
+#G #T1 #T2 #HT12 #L
+lapply (lsubr_cpr_trans … HT12 L ?) //
+qed.
+
+(* Basic_1: includes by definition: pr0_refl *)
+lemma cpr_refl: ∀G,T,L. ⦃G, L⦄ ⊢ T ➡ T.
+#G #T elim T -T // * /2 width=1 by cpr_bind, cpr_flat/
+qed.
+
+(* Basic_1: was: pr2_head_1 *)
+lemma cpr_pair_sn: ∀I,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡ V2 →
+ ∀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,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 #l #HLK elim (lt_or_eq_or_gt i l)
+ #Hil [1,3: /4 width=4 by lift_lref_ge_minus, lift_lref_lt, ylt_inj, yle_inj, 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 #l #HLK
+ elim (IHW1 … HLK) -IHW1 #W2 #W #HW12 #HW2
+ [ 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,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 #K #V1 #V2 #W2 #i #d #HLK #_ #HVW2 #IHV12 #H destruct
+ /3 width=6 by cpr_delta/
+| #G #L #K #V1 #V2 #W2 #i #d #_ #_ #_ #_ <plus_n_Sm #H destruct
+]
+qed-.
+
+lemma lstas_cpr: ∀h,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 •*[h, 0] T2 → ⦃G, L⦄ ⊢ T1 ➡ T2.
+/2 width=4 by lstas_cpr_aux/ qed.
+
+(* Basic inversion lemmas ***************************************************)
+
+fact cpr_inv_atom1_aux: ∀G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ∀I. T1 = ⓪{I} →
+ T2 = ⓪{I} ∨
+ ∃∃K,V,V2,i. ⬇[i] L ≡ K. ⓓV & ⦃G, K⦄ ⊢ V ➡ V2 &
+ ⬆[O, i + 1] V2 ≡ T2 & I = LRef i.
+#G #L #T1 #T2 * -G -L -T1 -T2
+[ #I #G #L #J #H destruct /2 width=1 by or_introl/
+| #L #G #K #V #V2 #T2 #i #HLK #HV2 #HVT2 #J #H destruct /3 width=8 by ex4_4_intro, or_intror/
+| #a #I #G #L #V1 #V2 #T1 #T2 #_ #_ #J #H destruct
+| #I #G #L #V1 #V2 #T1 #T2 #_ #_ #J #H destruct
+| #G #L #V #T1 #T #T2 #_ #_ #J #H destruct
+| #G #L #V #T1 #T2 #_ #J #H destruct
+| #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #J #H destruct
+| #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #J #H destruct
+]
+qed-.
+
+lemma cpr_inv_atom1: ∀I,G,L,T2. ⦃G, L⦄ ⊢ ⓪{I} ➡ T2 →
+ T2 = ⓪{I} ∨
+ ∃∃K,V,V2,i. ⬇[i] L ≡ K. ⓓV & ⦃G, K⦄ ⊢ V ➡ V2 &
+ ⬆[O, i + 1] V2 ≡ T2 & I = LRef i.
+/2 width=3 by cpr_inv_atom1_aux/ qed-.
+
+(* Basic_1: includes: pr0_gen_sort pr2_gen_sort *)
+lemma cpr_inv_sort1: ∀G,L,T2,s. ⦃G, L⦄ ⊢ ⋆s ➡ T2 → T2 = ⋆s.
+#G #L #T2 #s #H
+elim (cpr_inv_atom1 … H) -H //
+* #K #V #V2 #i #_ #_ #_ #H destruct
+qed-.
+
+(* Basic_1: includes: pr0_gen_lref pr2_gen_lref *)
+lemma cpr_inv_lref1: ∀G,L,T2,i. ⦃G, L⦄ ⊢ #i ➡ T2 →
+ T2 = #i ∨
+ ∃∃K,V,V2. ⬇[i] L ≡ K. ⓓV & ⦃G, K⦄ ⊢ V ➡ V2 &
+ ⬆[O, i + 1] V2 ≡ T2.
+#G #L #T2 #i #H
+elim (cpr_inv_atom1 … H) -H /2 width=1 by or_introl/
+* #K #V #V2 #j #HLK #HV2 #HVT2 #H destruct /3 width=6 by ex3_3_intro, or_intror/
+qed-.
+
+lemma cpr_inv_gref1: ∀G,L,T2,p. ⦃G, L⦄ ⊢ §p ➡ T2 → T2 = §p.
+#G #L #T2 #p #H
+elim (cpr_inv_atom1 … H) -H //
+* #K #V #V2 #i #_ #_ #_ #H destruct
+qed-.
+
+fact cpr_inv_bind1_aux: ∀G,L,U1,U2. ⦃G, L⦄ ⊢ U1 ➡ U2 →
+ ∀a,I,V1,T1. U1 = ⓑ{a,I}V1. T1 → (
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ➡ T2 &
+ U2 = ⓑ{a,I}V2.T2
+ ) ∨
+ ∃∃T. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡ T & ⬆[0, 1] U2 ≡ T &
+ a = true & I = Abbr.
+#G #L #U1 #U2 * -L -U1 -U2
+[ #I #G #L #b #J #W1 #U1 #H destruct
+| #L #G #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
+| #G #L #V #T1 #T #T2 #HT1 #HT2 #b #J #W #U1 #H destruct /3 width=3 by ex4_intro, or_intror/
+| #G #L #V #T1 #T2 #_ #b #J #W #U1 #H destruct
+| #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #b #J #W #U1 #H destruct
+| #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #b #J #W #U1 #H destruct
+]
+qed-.
+
+lemma cpr_inv_bind1: ∀a,I,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡ U2 → (
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ➡ T2 &
+ U2 = ⓑ{a,I}V2.T2
+ ) ∨
+ ∃∃T. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡ T & ⬆[0, 1] U2 ≡ T &
+ a = true & I = Abbr.
+/2 width=3 by cpr_inv_bind1_aux/ qed-.
+
+(* Basic_1: includes: pr0_gen_abbr pr2_gen_abbr *)
+lemma cpr_inv_abbr1: ∀a,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓓ{a}V1.T1 ➡ U2 → (
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L. ⓓV1⦄ ⊢ T1 ➡ T2 &
+ U2 = ⓓ{a}V2.T2
+ ) ∨
+ ∃∃T. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡ T & ⬆[0, 1] U2 ≡ T & a = true.
+#a #G #L #V1 #T1 #U2 #H
+elim (cpr_inv_bind1 … H) -H *
+/3 width=5 by ex3_2_intro, ex3_intro, or_introl, or_intror/
+qed-.
+
+(* Basic_1: includes: pr0_gen_abst pr2_gen_abst *)
+lemma cpr_inv_abst1: ∀a,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓛ{a}V1.T1 ➡ U2 →
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L.ⓛV1⦄ ⊢ T1 ➡ T2 &
+ U2 = ⓛ{a}V2.T2.
+#a #G #L #V1 #T1 #U2 #H
+elim (cpr_inv_bind1 … H) -H *
+[ /3 width=5 by ex3_2_intro/
+| #T #_ #_ #_ #H destruct
+]
+qed-.
+
+fact cpr_inv_flat1_aux: ∀G,L,U,U2. ⦃G, L⦄ ⊢ U ➡ U2 →
+ ∀I,V1,U1. U = ⓕ{I}V1.U1 →
+ ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L⦄ ⊢ U1 ➡ T2 &
+ U2 = ⓕ{I} V2. T2
+ | (⦃G, L⦄ ⊢ U1 ➡ U2 ∧ I = Cast)
+ | ∃∃a,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L⦄ ⊢ W1 ➡ W2 &
+ ⦃G, L.ⓛW1⦄ ⊢ T1 ➡ T2 & U1 = ⓛ{a}W1.T1 &
+ U2 = ⓓ{a}ⓝW2.V2.T2 & I = Appl
+ | ∃∃a,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡ V & ⬆[0,1] V ≡ V2 &
+ ⦃G, L⦄ ⊢ W1 ➡ W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ➡ T2 &
+ U1 = ⓓ{a}W1.T1 &
+ U2 = ⓓ{a}W2.ⓐV2.T2 & I = Appl.
+#G #L #U #U2 * -L -U -U2
+[ #I #G #L #J #W1 #U1 #H destruct
+| #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 or4_intro0, ex3_2_intro/
+| #G #L #V #T1 #T #T2 #_ #_ #J #W #U1 #H destruct
+| #G #L #V #T1 #T2 #HT12 #J #W #U1 #H destruct /3 width=1 by or4_intro1, conj/
+| #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #HV12 #HW12 #HT12 #J #W #U1 #H destruct /3 width=11 by or4_intro2, ex6_6_intro/
+| #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV1 #HV2 #HW12 #HT12 #J #W #U1 #H destruct /3 width=13 by or4_intro3, ex7_7_intro/
+]
+qed-.
+
+lemma cpr_inv_flat1: ∀I,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓕ{I}V1.U1 ➡ U2 →
+ ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L⦄ ⊢ U1 ➡ T2 &
+ U2 = ⓕ{I}V2.T2
+ | (⦃G, L⦄ ⊢ U1 ➡ U2 ∧ I = Cast)
+ | ∃∃a,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L⦄ ⊢ W1 ➡ W2 &
+ ⦃G, L.ⓛW1⦄ ⊢ T1 ➡ T2 & U1 = ⓛ{a}W1.T1 &
+ U2 = ⓓ{a}ⓝW2.V2.T2 & I = Appl
+ | ∃∃a,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡ V & ⬆[0,1] V ≡ V2 &
+ ⦃G, L⦄ ⊢ W1 ➡ W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ➡ T2 &
+ U1 = ⓓ{a}W1.T1 &
+ U2 = ⓓ{a}W2.ⓐV2.T2 & I = Appl.
+/2 width=3 by cpr_inv_flat1_aux/ qed-.
+
+(* Basic_1: includes: pr0_gen_appl pr2_gen_appl *)
+lemma cpr_inv_appl1: ∀G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓐV1.U1 ➡ U2 →
+ ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L⦄ ⊢ U1 ➡ T2 &
+ U2 = ⓐV2.T2
+ | ∃∃a,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L⦄ ⊢ W1 ➡ W2 &
+ ⦃G, L.ⓛW1⦄ ⊢ T1 ➡ T2 &
+ U1 = ⓛ{a}W1.T1 & U2 = ⓓ{a}ⓝW2.V2.T2
+ | ∃∃a,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡ V & ⬆[0,1] V ≡ V2 &
+ ⦃G, L⦄ ⊢ W1 ➡ W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ➡ T2 &
+ U1 = ⓓ{a}W1.T1 & U2 = ⓓ{a}W2.ⓐV2.T2.
+#G #L #V1 #U1 #U2 #H elim (cpr_inv_flat1 … H) -H *
+[ /3 width=5 by or3_intro0, ex3_2_intro/
+| #_ #H destruct
+| /3 width=11 by or3_intro1, ex5_6_intro/
+| /3 width=13 by or3_intro2, ex6_7_intro/
+]
+qed-.
+
+(* Note: the main property of simple terms *)
+lemma cpr_inv_appl1_simple: ∀G,L,V1,T1,U. ⦃G, L⦄ ⊢ ⓐV1. T1 ➡ U → 𝐒⦃T1⦄ →
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L⦄ ⊢ T1 ➡ T2 &
+ U = ⓐV2. T2.
+#G #L #V1 #T1 #U #H #HT1
+elim (cpr_inv_appl1 … H) -H *
+[ /2 width=5 by ex3_2_intro/
+| #a #V2 #W1 #W2 #U1 #U2 #_ #_ #_ #H #_ destruct
+ elim (simple_inv_bind … HT1)
+| #a #V #V2 #W1 #W2 #U1 #U2 #_ #_ #_ #_ #H #_ destruct
+ elim (simple_inv_bind … HT1)
+]
+qed-.
+
+(* Basic_1: includes: pr0_gen_cast pr2_gen_cast *)
+lemma cpr_inv_cast1: ∀G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓝ V1. U1 ➡ U2 → (
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L⦄ ⊢ U1 ➡ T2 &
+ U2 = ⓝ V2. T2
+ ) ∨ ⦃G, L⦄ ⊢ U1 ➡ U2.
+#G #L #V1 #U1 #U2 #H elim (cpr_inv_flat1 … H) -H *
+[ /3 width=5 by ex3_2_intro, or_introl/
+| /2 width=1 by or_intror/
+| #a #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #_ #H destruct
+| #a #V #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #_ #_ #H destruct
+]
+qed-.
+
+(* Basic forward lemmas *****************************************************)
+
+lemma cpr_fwd_bind1_minus: ∀I,G,L,V1,T1,T. ⦃G, L⦄ ⊢ -ⓑ{I}V1.T1 ➡ T → ∀b.
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ ⓑ{b,I}V1.T1 ➡ ⓑ{b,I}V2.T2 &
+ T = -ⓑ{I}V2.T2.
+#I #G #L #V1 #T1 #T #H #b
+elim (cpr_inv_bind1 … H) -H *
+[ #V2 #T2 #HV12 #HT12 #H destruct /3 width=4 by cpr_bind, ex2_2_intro/
+| #T2 #_ #_ #H destruct
+]
+qed-.
+
+(* Basic_1: removed theorems 11:
+ pr0_subst0_back pr0_subst0_fwd pr0_subst0
+ pr2_head_2 pr2_cflat clear_pr2_trans
+ pr2_gen_csort pr2_gen_cflat pr2_gen_cbind
+ pr2_gen_ctail pr2_ctail
+*)
+(* Basic_1: removed local theorems 4:
+ pr0_delta_eps pr0_cong_delta
+ pr2_free_free pr2_free_delta
+*)
--- /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/cir.ma".
+include "basic_2/reduction/cpr.ma".
+
+(* CONTEXT-SENSITIVE PARALLEL REDUCTION FOR TERMS ***************************)
+
+(* Advanced forward lemmas on irreducibility ********************************)
+
+lemma cpr_fwd_cir: ∀G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ⦃G, L⦄ ⊢ ➡ 𝐈⦃T1⦄ → T2 = T1.
+#G #L #T1 #T2 #H elim H -G -L -T1 -T2
+[ //
+| #G #L #K #V1 #V2 #W2 #i #HLK #_ #HVW2 #IHV12 #H
+ elim (cir_inv_delta … HLK) //
+| #a * #G #L #V1 #V2 #T1 #T2 #_ #_ #IHV1 #IHT1 #H
+ [ elim (cir_inv_bind … H) -H #HV1 #HT1 * #H destruct
+ lapply (IHV1 … HV1) -IHV1 -HV1 #H destruct
+ lapply (IHT1 … HT1) -IHT1 #H destruct //
+ | elim (cir_inv_ib2 … H) -H /3 width=2 by or_introl, eq_f2/
+ ]
+| * #G #L #V1 #V2 #T1 #T2 #_ #_ #IHV1 #IHT1 #H
+ [ elim (cir_inv_appl … H) -H #HV1 #HT1 #_
+ >IHV1 -IHV1 // -HV1 >IHT1 -IHT1 //
+ | elim (cir_inv_ri2 … H) /2 width=1 by/
+ ]
+| #G #L #V1 #T1 #T #T2 #_ #_ #_ #H
+ elim (cir_inv_ri2 … H) /2 width=1 by or_introl/
+| #G #L #V1 #T1 #T2 #_ #_ #H
+ elim (cir_inv_ri2 … H) /2 width=1 by/
+| #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #_ #_ #H
+ elim (cir_inv_appl … H) -H #_ #_ #H
+ elim (simple_inv_bind … H)
+| #a #G #L #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #_ #_ #_ #H
+ elim (cir_inv_appl … H) -H #_ #_ #H
+ elim (simple_inv_bind … H)
+]
+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_max.ma".
+include "basic_2/substitution/drop_drop.ma".
+include "basic_2/reduction/cpr.ma".
+
+(* CONTEXT-SENSITIVE PARALLEL REDUCTION FOR TERMS ***************************)
+
+(* Relocation properties ****************************************************)
+
+(* Basic_1: includes: pr0_lift pr2_lift *)
+lemma cpr_lift: ∀G. d_liftable (cpr G).
+#G #K #T1 #T2 #H elim H -G -K -T1 -T2
+[ #I #G #K #L #c #l #k #_ #U1 #H1 #U2 #H2
+ >(lift_mono … H1 … H2) -H1 -H2 //
+| #G #K #KV #V #V2 #W2 #i #HKV #HV2 #HVW2 #IHV2 #L #c #l #k #HLK #U1 #H #U2 #HWU2
+ elim (lift_inv_lref1 … H) * #Hil #H destruct
+ [ elim (lift_trans_ge … HVW2 … HWU2) -W2 /2 width=1 by ylt_fwd_le_succ1/ #W2 #HVW2 #HWU2
+ elim (drop_trans_le … HLK … HKV) -K /2 width=2 by ylt_fwd_le/ #X #HLK #H
+ elim (drop_inv_skip2 … H) -H /2 width=1 by ylt_to_minus/ -Hil
+ #K #Y #HKV #HVY #H destruct /3 width=9 by cpr_delta/
+ | lapply (lift_trans_be … HVW2 … HWU2 ? ?) -W2 /2 width=1 by yle_succ_dx/ >plus_plus_comm_23 #HVU2
+ 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 #c #l #k #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 #c #l #k #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 #c #l #k #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 #c #l #k #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 #c #l #k #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 #c #l #k #HLK #X1 #HX1 #X2 #HX2
+ elim (lift_inv_flat1 … HX1) -HX1 #V0 #X #HV10 #HX #HX1 destruct
+ elim (lift_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT10 #HX destruct
+ elim (lift_inv_bind1 … HX2) -HX2 #W3 #X #HW23 #HX #HX2 destruct
+ elim (lift_inv_flat1 … HX) -HX #V3 #T3 #HV3 #HT23 #HX destruct
+ elim (lift_trans_ge … HV2 … HV3) -V2 /4 width=6 by cpr_theta, drop_skip/
+]
+qed.
+
+(* Basic_1: includes: pr0_gen_lift pr2_gen_lift *)
+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 #c #l #k #_ #T1 #H
+ [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3 by 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 #c #l #k #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 // <yminus_succ2 <yplus_inj >yplus_SO2 >ymax_pre_sn /3 width=8 by cpr_delta, ylt_fwd_le_succ1, ex2_intro/
+ | elim (yle_inv_plus_inj2 … Hil) #Hlim #Hmi
+ lapply (yle_inv_inj … Hmi) -Hmi #Hmi
+ lapply (drop_conf_ge … HLK … HLV ?) -L // #HKLV
+ elim (lift_split … HVW2 l (i - k + 1)) -HVW2 [2,3,4: /2 width=1 by yle_succ_dx, 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=8 by cpr_delta, ex2_intro/
+ ]
+| #a #I #G #L #V1 #V2 #U1 #U2 #_ #_ #IHV12 #IHU12 #K #c #l #k #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 #c #l #k #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 #c #l #k #HLK #X #H
+ elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct
+ elim (IHU1 (K.ⓓW1) c … 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 #c #l #k #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 #c #l #k #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) c … 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 #c #l #k #HLK #X #HX
+ elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct
+ elim (lift_inv_bind2 … HY) -HY #W0 #T0 #HW01 #HT01 #HY destruct
+ elim (IHV1 … HLK … HV01) -V1 #V3 #HV3 #HV03
+ elim (IHT12 (K.ⓓW0) c … HT01) -T1 /2 width=1 by drop_skip/ #T3 #HT32 #HT03
+ elim (IHW12 … HLK … HW01) -W1
+ elim (lift_trans_le … HV3 … HV2) -V
+ /4 width=9 by cpr_theta, lift_flat, lift_bind, ex2_intro/
+]
+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_drop.ma".
+include "basic_2/reduction/cpr.ma".
+
+(* CONTEXT-SENSITIVE PARALLEL REDUCTION FOR TERMS ***************************)
+
+(* Properties on lazy sn pointwise extensions *******************************)
+
+lemma cpr_llpx_sn_conf: ∀R. (∀L. reflexive … (R L)) → d_liftable R → d_deliftable_sn R →
+ ∀G. c_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 #V1c #V2c #W2c #i #HLKs #_ #HVW2c #IHV12c #Ld #H elim (llpx_sn_inv_lref_ge_sn … H … HLKs) // -H
+ #Kd #V1l #HLKd #HV1c #HV1sd
+ lapply (drop_fwd_drop2 … HLKs) -HLKs #HLKs
+ lapply (drop_fwd_drop2 … HLKd) -HLKd #HLKd
+ @(llpx_sn_lift_le … HLKs HLKd … HVW2c) -HLKs -HLKd -HVW2c /2 width=1 by/ (**) (* full auto too slow *)
+| #a #I #G #Ls #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #Ld #H elim (llpx_sn_inv_bind_O … H) -H
+ /4 width=5 by llpx_sn_bind_repl_SO, llpx_sn_bind/
+| #I #G #Ls #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
+ /3 width=1 by llpx_sn_flat/
+| #G #Ls #V #T1 #T2 #T #_ #HT2 #IHT12 #Ld #H elim (llpx_sn_inv_bind_O … H) -H
+ /3 width=10 by llpx_sn_inv_lift_le, drop_drop/
+| #G #Ls #V #T1 #T2 #_ #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H /2 width=1 by/
+| #a #G #Ls #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
+ #HV1 #H elim (llpx_sn_inv_bind_O … H) -H
+ /4 width=5 by llpx_sn_bind_repl_SO, llpx_sn_flat, llpx_sn_bind/
+| #a #G #Ls #V1 #V2 #V #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
+ #HV1 #H elim (llpx_sn_inv_bind_O … H) -H //
+ #HW1 #HT1 @llpx_sn_bind_O /2 width=1 by/ @llpx_sn_flat (**) (* full auto too slow *)
+ [ /3 width=10 by llpx_sn_lift_le, drop_drop/
+ | /3 width=4 by llpx_sn_bind_repl_O/
+]
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/pred_6.ma".
+include "basic_2/static/sd.ma".
+include "basic_2/reduction/cpr.ma".
+
+(* CONTEXT-SENSITIVE EXTENDED PARALLEL REDUCTION FOR TERMS ******************)
+
+(* avtivate genv *)
+inductive cpx (h) (o): relation4 genv lenv term term ≝
+| cpx_atom : ∀I,G,L. cpx h o G L (⓪{I}) (⓪{I})
+| cpx_st : ∀G,L,s,d. deg h o s (d+1) → cpx h o G L (⋆s) (⋆(next h s))
+| cpx_delta: ∀I,G,L,K,V,V2,W2,i.
+ ⬇[i] L ≡ K.ⓑ{I}V → cpx h o G K V V2 →
+ ⬆[0, i+1] V2 ≡ W2 → cpx h o G L (#i) W2
+| cpx_bind : ∀a,I,G,L,V1,V2,T1,T2.
+ cpx h o G L V1 V2 → cpx h o G (L.ⓑ{I}V1) T1 T2 →
+ cpx h o G L (ⓑ{a,I}V1.T1) (ⓑ{a,I}V2.T2)
+| cpx_flat : ∀I,G,L,V1,V2,T1,T2.
+ cpx h o G L V1 V2 → cpx h o G L T1 T2 →
+ cpx h o G L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2)
+| cpx_zeta : ∀G,L,V,T1,T,T2. cpx h o G (L.ⓓV) T1 T →
+ ⬆[0, 1] T2 ≡ T → cpx h o G L (+ⓓV.T1) T2
+| cpx_eps : ∀G,L,V,T1,T2. cpx h o G L T1 T2 → cpx h o G L (ⓝV.T1) T2
+| cpx_ct : ∀G,L,V1,V2,T. cpx h o G L V1 V2 → cpx h o G L (ⓝV1.T) V2
+| cpx_beta : ∀a,G,L,V1,V2,W1,W2,T1,T2.
+ cpx h o G L V1 V2 → cpx h o G L W1 W2 → cpx h o G (L.ⓛW1) T1 T2 →
+ cpx h o G L (ⓐV1.ⓛ{a}W1.T1) (ⓓ{a}ⓝW2.V2.T2)
+| cpx_theta: ∀a,G,L,V1,V,V2,W1,W2,T1,T2.
+ cpx h o G L V1 V → ⬆[0, 1] V ≡ V2 → cpx h o G L W1 W2 →
+ cpx h o G (L.ⓓW1) T1 T2 →
+ cpx h o G L (ⓐV1.ⓓ{a}W1.T1) (ⓓ{a}W2.ⓐV2.T2)
+.
+
+interpretation
+ "context-sensitive extended parallel reduction (term)"
+ 'PRed h o G L T1 T2 = (cpx h o G L T1 T2).
+
+(* Basic properties *********************************************************)
+
+lemma lsubr_cpx_trans: ∀h,o,G. lsub_trans … (cpx h o G) lsubr.
+#h #o #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 (lsubr_fwd_drop2_pair … HL12 … HLK1) -HL12 -HLK1 *
+ /4 width=7 by cpx_delta, cpx_ct/
+|4,9: /4 width=1 by cpx_bind, cpx_beta, lsubr_pair/
+|5,7,8: /3 width=1 by cpx_flat, cpx_eps, cpx_ct/
+|6,10: /4 width=3 by cpx_zeta, cpx_theta, lsubr_pair/
+]
+qed-.
+
+(* Note: this is "∀h,g,L. reflexive … (cpx h g L)" *)
+lemma cpx_refl: ∀h,o,G,T,L. ⦃G, L⦄ ⊢ T ➡[h, o] T.
+#h #o #G #T elim T -T // * /2 width=1 by cpx_bind, cpx_flat/
+qed.
+
+lemma cpr_cpx: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ⦃G, L⦄ ⊢ T1 ➡[h, o] T2.
+#h #o #G #L #T1 #T2 #H elim H -L -T1 -T2
+/2 width=7 by cpx_delta, cpx_bind, cpx_flat, cpx_zeta, cpx_eps, cpx_beta, cpx_theta/
+qed.
+
+lemma cpx_pair_sn: ∀h,o,I,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 →
+ ∀T. ⦃G, L⦄ ⊢ ②{I}V1.T ➡[h, o] ②{I}V2.T.
+#h #o * /2 width=1 by cpx_bind, cpx_flat/
+qed.
+
+lemma cpx_delift: ∀h,o,I,G,K,V,T1,L,l. ⬇[l] L ≡ (K.ⓑ{I}V) →
+ ∃∃T2,T. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 & ⬆[l, 1] T ≡ T2.
+#h #o #I #G #K #V #T1 elim T1 -T1
+[ * #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: /4 width=4 by cpx_atom, lift_lref_ge_minus, lift_lref_lt, ylt_inj, yle_inj, 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 #l #HLK
+ elim (IHW1 … HLK) -IHW1 #W2 #W #HW12 #HW2
+ [ 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/
+ ]
+]
+qed-.
+
+(* Basic inversion lemmas ***************************************************)
+
+fact cpx_inv_atom1_aux: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → ∀J. T1 = ⓪{J} →
+ ∨∨ T2 = ⓪{J}
+ | ∃∃s,d. deg h o s (d+1) & T2 = ⋆(next h s) & J = Sort s
+ | ∃∃I,K,V,V2,i. ⬇[i] L ≡ K.ⓑ{I}V & ⦃G, K⦄ ⊢ V ➡[h, o] V2 &
+ ⬆[O, i+1] V2 ≡ T2 & J = LRef i.
+#G #h #o #L #T1 #T2 * -L -T1 -T2
+[ #I #G #L #J #H destruct /2 width=1 by or3_intro0/
+| #G #L #s #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
+| #G #L #V #T1 #T #T2 #_ #_ #J #H destruct
+| #G #L #V #T1 #T2 #_ #J #H destruct
+| #G #L #V1 #V2 #T #_ #J #H destruct
+| #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #J #H destruct
+| #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #J #H destruct
+]
+qed-.
+
+lemma cpx_inv_atom1: ∀h,o,J,G,L,T2. ⦃G, L⦄ ⊢ ⓪{J} ➡[h, o] T2 →
+ ∨∨ T2 = ⓪{J}
+ | ∃∃s,d. deg h o s (d+1) & T2 = ⋆(next h s) & J = Sort s
+ | ∃∃I,K,V,V2,i. ⬇[i] L ≡ K.ⓑ{I}V & ⦃G, K⦄ ⊢ V ➡[h, o] V2 &
+ ⬆[O, i+1] V2 ≡ T2 & J = LRef i.
+/2 width=3 by cpx_inv_atom1_aux/ qed-.
+
+lemma cpx_inv_sort1: ∀h,o,G,L,T2,s. ⦃G, L⦄ ⊢ ⋆s ➡[h, o] T2 → T2 = ⋆s ∨
+ ∃∃d. deg h o s (d+1) & T2 = ⋆(next h s).
+#h #o #G #L #T2 #s #H
+elim (cpx_inv_atom1 … H) -H /2 width=1 by or_introl/ *
+[ #s0 #d0 #Hkd0 #H1 #H2 destruct /3 width=4 by ex2_intro, or_intror/
+| #I #K #V #V2 #i #_ #_ #_ #H destruct
+]
+qed-.
+
+lemma cpx_inv_lref1: ∀h,o,G,L,T2,i. ⦃G, L⦄ ⊢ #i ➡[h, o] T2 →
+ T2 = #i ∨
+ ∃∃I,K,V,V2. ⬇[i] L ≡ K. ⓑ{I}V & ⦃G, K⦄ ⊢ V ➡[h, o] V2 &
+ ⬆[O, i+1] V2 ≡ T2.
+#h #o #G #L #T2 #i #H
+elim (cpx_inv_atom1 … H) -H /2 width=1 by or_introl/ *
+[ #s #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_lref1_ge: ∀h,o,G,L,T2,i. ⦃G, L⦄ ⊢ #i ➡[h, o] T2 → |L| ≤ i → T2 = #i.
+#h #o #G #L #T2 #i #H elim (cpx_inv_lref1 … H) -H // *
+#I #K #V1 #V2 #HLK #_ #_ #HL -h -G -V2 lapply (drop_fwd_length_lt2 … HLK) -K -I -V1
+#H elim (lt_refl_false i) /2 width=3 by lt_to_le_to_lt/
+qed-.
+
+lemma cpx_inv_gref1: ∀h,o,G,L,T2,p. ⦃G, L⦄ ⊢ §p ➡[h, o] T2 → T2 = §p.
+#h #o #G #L #T2 #p #H
+elim (cpx_inv_atom1 … H) -H // *
+[ #s #d #_ #_ #H destruct
+| #I #K #V #V2 #i #_ #_ #_ #H destruct
+]
+qed-.
+
+fact cpx_inv_bind1_aux: ∀h,o,G,L,U1,U2. ⦃G, L⦄ ⊢ U1 ➡[h, o] U2 →
+ ∀a,J,V1,T1. U1 = ⓑ{a,J}V1.T1 → (
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L.ⓑ{J}V1⦄ ⊢ T1 ➡[h, o] T2 &
+ U2 = ⓑ{a,J}V2.T2
+ ) ∨
+ ∃∃T. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡[h, o] T & ⬆[0, 1] U2 ≡ T &
+ a = true & J = Abbr.
+#h #o #G #L #U1 #U2 * -L -U1 -U2
+[ #I #G #L #b #J #W #U1 #H destruct
+| #G #L #s #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
+| #G #L #V #T1 #T #T2 #HT1 #HT2 #b #J #W #U1 #H destruct /3 width=3 by ex4_intro, or_intror/
+| #G #L #V #T1 #T2 #_ #b #J #W #U1 #H destruct
+| #G #L #V1 #V2 #T #_ #b #J #W #U1 #H destruct
+| #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #b #J #W #U1 #H destruct
+| #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #b #J #W #U1 #H destruct
+]
+qed-.
+
+lemma cpx_inv_bind1: ∀h,o,a,I,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡[h, o] U2 → (
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ➡[h, o] T2 &
+ U2 = ⓑ{a,I} V2. T2
+ ) ∨
+ ∃∃T. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡[h, o] T & ⬆[0, 1] U2 ≡ T &
+ a = true & I = Abbr.
+/2 width=3 by cpx_inv_bind1_aux/ qed-.
+
+lemma cpx_inv_abbr1: ∀h,o,a,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓓ{a}V1.T1 ➡[h, o] U2 → (
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L.ⓓV1⦄ ⊢ T1 ➡[h, o] T2 &
+ U2 = ⓓ{a} V2. T2
+ ) ∨
+ ∃∃T. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡[h, o] T & ⬆[0, 1] U2 ≡ T & a = true.
+#h #o #a #G #L #V1 #T1 #U2 #H
+elim (cpx_inv_bind1 … H) -H * /3 width=5 by ex3_2_intro, ex3_intro, or_introl, or_intror/
+qed-.
+
+lemma cpx_inv_abst1: ∀h,o,a,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓛ{a}V1.T1 ➡[h, o] U2 →
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L.ⓛV1⦄ ⊢ T1 ➡[h, o] T2 &
+ U2 = ⓛ{a} V2. T2.
+#h #o #a #G #L #V1 #T1 #U2 #H
+elim (cpx_inv_bind1 … H) -H *
+[ /3 width=5 by ex3_2_intro/
+| #T #_ #_ #_ #H destruct
+]
+qed-.
+
+fact cpx_inv_flat1_aux: ∀h,o,G,L,U,U2. ⦃G, L⦄ ⊢ U ➡[h, o] U2 →
+ ∀J,V1,U1. U = ⓕ{J}V1.U1 →
+ ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L⦄ ⊢ U1 ➡[h, o] T2 &
+ U2 = ⓕ{J}V2.T2
+ | (⦃G, L⦄ ⊢ U1 ➡[h, o] U2 ∧ J = Cast)
+ | (⦃G, L⦄ ⊢ V1 ➡[h, o] U2 ∧ J = Cast)
+ | ∃∃a,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L⦄ ⊢ W1 ➡[h, o] W2 &
+ ⦃G, L.ⓛW1⦄ ⊢ T1 ➡[h, o] T2 &
+ U1 = ⓛ{a}W1.T1 &
+ U2 = ⓓ{a}ⓝW2.V2.T2 & J = Appl
+ | ∃∃a,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V & ⬆[0,1] V ≡ V2 &
+ ⦃G, L⦄ ⊢ W1 ➡[h, o] W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ➡[h, o] T2 &
+ U1 = ⓓ{a}W1.T1 &
+ U2 = ⓓ{a}W2.ⓐV2.T2 & J = Appl.
+#h #o #G #L #U #U2 * -L -U -U2
+[ #I #G #L #J #W #U1 #H destruct
+| #G #L #s #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/
+| #G #L #V #T1 #T #T2 #_ #_ #J #W #U1 #H destruct
+| #G #L #V #T1 #T2 #HT12 #J #W #U1 #H destruct /3 width=1 by or5_intro1, conj/
+| #G #L #V1 #V2 #T #HV12 #J #W #U1 #H destruct /3 width=1 by or5_intro2, conj/
+| #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #HV12 #HW12 #HT12 #J #W #U1 #H destruct /3 width=11 by or5_intro3, ex6_6_intro/
+| #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV1 #HV2 #HW12 #HT12 #J #W #U1 #H destruct /3 width=13 by or5_intro4, ex7_7_intro/
+]
+qed-.
+
+lemma cpx_inv_flat1: ∀h,o,I,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓕ{I}V1.U1 ➡[h, o] U2 →
+ ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L⦄ ⊢ U1 ➡[h, o] T2 &
+ U2 = ⓕ{I} V2. T2
+ | (⦃G, L⦄ ⊢ U1 ➡[h, o] U2 ∧ I = Cast)
+ | (⦃G, L⦄ ⊢ V1 ➡[h, o] U2 ∧ I = Cast)
+ | ∃∃a,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L⦄ ⊢ W1 ➡[h, o] W2 &
+ ⦃G, L.ⓛW1⦄ ⊢ T1 ➡[h, o] T2 &
+ U1 = ⓛ{a}W1.T1 &
+ U2 = ⓓ{a}ⓝW2.V2.T2 & I = Appl
+ | ∃∃a,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V & ⬆[0,1] V ≡ V2 &
+ ⦃G, L⦄ ⊢ W1 ➡[h, o] W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ➡[h, o] T2 &
+ U1 = ⓓ{a}W1.T1 &
+ U2 = ⓓ{a}W2.ⓐV2.T2 & I = Appl.
+/2 width=3 by cpx_inv_flat1_aux/ qed-.
+
+lemma cpx_inv_appl1: ∀h,o,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓐ V1.U1 ➡[h, o] U2 →
+ ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L⦄ ⊢ U1 ➡[h, o] T2 &
+ U2 = ⓐ V2. T2
+ | ∃∃a,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L⦄ ⊢ W1 ➡[h, o] W2 &
+ ⦃G, L.ⓛW1⦄ ⊢ T1 ➡[h, o] T2 &
+ U1 = ⓛ{a}W1.T1 & U2 = ⓓ{a}ⓝW2.V2.T2
+ | ∃∃a,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V & ⬆[0,1] V ≡ V2 &
+ ⦃G, L⦄ ⊢ W1 ➡[h, o] W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ➡[h, o] T2 &
+ U1 = ⓓ{a}W1.T1 & U2 = ⓓ{a}W2. ⓐV2. T2.
+#h #o #G #L #V1 #U1 #U2 #H elim (cpx_inv_flat1 … H) -H *
+[ /3 width=5 by or3_intro0, ex3_2_intro/
+|2,3: #_ #H destruct
+| /3 width=11 by or3_intro1, ex5_6_intro/
+| /3 width=13 by or3_intro2, ex6_7_intro/
+]
+qed-.
+
+(* Note: the main property of simple terms *)
+lemma cpx_inv_appl1_simple: ∀h,o,G,L,V1,T1,U. ⦃G, L⦄ ⊢ ⓐV1.T1 ➡[h, o] U → 𝐒⦃T1⦄ →
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 &
+ U = ⓐV2.T2.
+#h #o #G #L #V1 #T1 #U #H #HT1
+elim (cpx_inv_appl1 … H) -H *
+[ /2 width=5 by ex3_2_intro/
+| #a #V2 #W1 #W2 #U1 #U2 #_ #_ #_ #H #_ destruct
+ elim (simple_inv_bind … HT1)
+| #a #V #V2 #W1 #W2 #U1 #U2 #_ #_ #_ #_ #H #_ destruct
+ elim (simple_inv_bind … HT1)
+]
+qed-.
+
+lemma cpx_inv_cast1: ∀h,o,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓝV1.U1 ➡[h, o] U2 →
+ ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L⦄ ⊢ U1 ➡[h, o] T2 &
+ U2 = ⓝ V2. T2
+ | ⦃G, L⦄ ⊢ U1 ➡[h, o] U2
+ | ⦃G, L⦄ ⊢ V1 ➡[h, o] U2.
+#h #o #G #L #V1 #U1 #U2 #H elim (cpx_inv_flat1 … H) -H *
+[ /3 width=5 by or3_intro0, ex3_2_intro/
+|2,3: /2 width=1 by or3_intro1, or3_intro2/
+| #a #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #_ #H destruct
+| #a #V #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #_ #_ #H destruct
+]
+qed-.
+
+(* Basic forward lemmas *****************************************************)
+
+lemma cpx_fwd_bind1_minus: ∀h,o,I,G,L,V1,T1,T. ⦃G, L⦄ ⊢ -ⓑ{I}V1.T1 ➡[h, o] T → ∀b.
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ ⓑ{b,I}V1.T1 ➡[h, o] ⓑ{b,I}V2.T2 &
+ T = -ⓑ{I}V2.T2.
+#h #o #I #G #L #V1 #T1 #T #H #b
+elim (cpx_inv_bind1 … H) -H *
+[ #V2 #T2 #HV12 #HT12 #H destruct /3 width=4 by cpx_bind, ex2_2_intro/
+| #T2 #_ #_ #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/reduction/cix.ma".
+include "basic_2/reduction/cpx.ma".
+
+(* CONTEXT-SENSITIVE EXTENDED PARALLEL REDUCTION FOR TERMS ******************)
+
+(* Advanced forward lemmas on irreducibility ********************************)
+
+lemma cpx_fwd_cix: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃T1⦄ → T2 = T1.
+#h #o #G #L #T1 #T2 #H elim H -G -L -T1 -T2
+[ //
+| #G #L #s #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
+ [ elim (cix_inv_bind … H) -H #HV1 #HT1 * #H destruct
+ lapply (IHV1 … HV1) -IHV1 -HV1 #H destruct
+ lapply (IHT1 … HT1) -IHT1 #H destruct //
+ | elim (cix_inv_ib2 … H) -H /3 width=2 by or_introl, eq_f2/
+ ]
+| * #G #L #V1 #V2 #T1 #T2 #_ #_ #IHV1 #IHT1 #H
+ [ elim (cix_inv_appl … H) -H #HV1 #HT1 #_
+ >IHV1 -IHV1 // -HV1 >IHT1 -IHT1 //
+ | elim (cix_inv_ri2 … H) /2 width=1 by/
+ ]
+| #G #L #V1 #T1 #T #T2 #_ #_ #_ #H
+ elim (cix_inv_ri2 … H) /2 width=1 by or_introl/
+| #G #L #V1 #T1 #T2 #_ #_ #H
+ elim (cix_inv_ri2 … H) /2 width=1 by/
+| #G #L #V1 #V2 #T #_ #_ #H
+ elim (cix_inv_ri2 … H) /2 width=1 by/
+| #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #_ #_ #H
+ elim (cix_inv_appl … H) -H #_ #_ #H
+ elim (simple_inv_bind … H)
+| #a #G #L #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #_ #_ #_ #H
+ elim (cix_inv_appl … H) -H #_ #_ #H
+ elim (simple_inv_bind … H)
+]
+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_max.ma".
+include "basic_2/substitution/drop_drop.ma".
+include "basic_2/multiple/fqus_alt.ma".
+include "basic_2/static/da.ma".
+include "basic_2/reduction/cpx.ma".
+
+(* CONTEXT-SENSITIVE EXTENDED PARALLEL REDUCTION FOR TERMS ******************)
+
+(* Advanced properties ******************************************************)
+
+fact sta_cpx_aux: ∀h,o,G,L,T1,T2,d2,d1. ⦃G, L⦄ ⊢ T1 •*[h, d2] T2 → d2 = 1 →
+ ⦃G, L⦄ ⊢ T1 ▪[h, o] d1+1 → ⦃G, L⦄ ⊢ T1 ➡[h, o] T2.
+#h #o #G #L #T1 #T2 #d2 #d1 #H elim H -G -L -T1 -T2 -d2
+[ #G #L #d2 #s #H0 destruct normalize
+ /3 width=4 by cpx_st, da_inv_sort/
+| #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 #d2 #HLK #HV12 #HVW2 #_ #H0 #H
+ lapply (discr_plus_xy_y … H0) -H0 #H0 destruct
+ 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/
+| /4 width=3 by cpx_flat, da_inv_flat/
+| /4 width=3 by cpx_eps, da_inv_flat/
+]
+qed-.
+
+lemma sta_cpx: ∀h,o,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 •*[h, 1] T2 →
+ ⦃G, L⦄ ⊢ T1 ▪[h, o] d+1 → ⦃G, L⦄ ⊢ T1 ➡[h, o] T2.
+/2 width=3 by sta_cpx_aux/ qed.
+
+(* Relocation properties ****************************************************)
+
+lemma cpx_lift: ∀h,o,G. d_liftable (cpx h o G).
+#h #o #G #K #T1 #T2 #H elim H -G -K -T1 -T2
+[ #I #G #K #L #c #l #k #_ #U1 #H1 #U2 #H2
+ >(lift_mono … H1 … H2) -H1 -H2 //
+| #G #K #s #d #Hkd #L #c #l #k #_ #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 #c #l #k #HLK #U1 #H #U2 #HWU2
+ elim (lift_inv_lref1 … H) * #Hil #H destruct
+ [ elim (lift_trans_ge … HVW2 … HWU2) -W2 /2 width=1 by ylt_fwd_le_succ1/ #W2 #HVW2 #HWU2
+ elim (drop_trans_le … HLK … HKV) -K /2 width=2 by ylt_fwd_le/ #X #HLK #H
+ elim (drop_inv_skip2 … H) -H /2 width=1 by ylt_to_minus/ -Hil
+ #K #Y #HKV #HVY #H destruct /3 width=10 by cpx_delta/
+ | lapply (lift_trans_be … HVW2 … HWU2 ? ?) -W2 /2 width=1 by yle_succ_dx/ >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 #c #l #k #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 #c #l #k #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 #c #l #k #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 #c #l #k #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 #c #l #k #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 #c #l #k #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 #c #l #k #HLK #X1 #HX1 #X2 #HX2
+ elim (lift_inv_flat1 … HX1) -HX1 #V0 #X #HV10 #HX #HX1 destruct
+ elim (lift_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT10 #HX destruct
+ elim (lift_inv_bind1 … HX2) -HX2 #W3 #X #HW23 #HX #HX2 destruct
+ elim (lift_inv_flat1 … HX) -HX #V3 #T3 #HV3 #HT23 #HX destruct
+ elim (lift_trans_ge … HV2 … HV3) -V2 /4 width=6 by cpx_theta, drop_skip/
+]
+qed.
+
+lemma cpx_inv_lift1: ∀h,o,G. d_deliftable_sn (cpx h o G).
+#h #o #G #L #U1 #U2 #H elim H -G -L -U1 -U2
+[ * #i #G #L #K #c #l #k #_ #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 * #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 #s #d #Hkd #K #c #l #k #_ #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 #c #l #k #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 // <yminus_succ2 <yplus_inj >yplus_SO2 >ymax_pre_sn /3 width=9 by cpx_delta, ylt_fwd_le_succ1, ex2_intro/
+ | elim (yle_inv_plus_inj2 … Hil) #Hlim #Hmi
+ lapply (yle_inv_inj … Hmi) -Hmi #Hmi
+ lapply (drop_conf_ge … HLK … HLV ?) -L // #HKLV
+ elim (lift_split … HVW2 l (i - k + 1)) -HVW2 /3 width=1 by yle_succ, yle_pred_sn, 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 #c #l #k #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 #c #l #k #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 #c #l #k #HLK #X #H
+ elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct
+ elim (IHU1 (K.ⓓW1) c … HTU1) /2 width=1 by drop_skip/ -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 #c #l #k #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 #c #l #k #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 #c #l #k #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) c … 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 #c #l #k #HLK #X #HX
+ elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct
+ elim (lift_inv_bind2 … HY) -HY #W0 #T0 #HW01 #HT01 #HY destruct
+ elim (IHV1 … HLK … HV01) -V1 #V3 #HV3 #HV03
+ elim (IHT12 (K.ⓓW0) c … HT01) -T1 /2 width=1 by drop_skip/ #T3 #HT32 #HT03
+ elim (IHW12 … HLK … HW01) -W1 #W3 #HW32 #HW03
+ elim (lift_trans_le … HV3 … HV2) -V
+ /4 width=9 by cpx_theta, lift_bind, lift_flat, ex2_intro/
+]
+qed-.
+
+(* Properties on supclosure *************************************************)
+
+lemma fqu_cpx_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡[h, o] U2 →
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] U1 & ⦃G1, L1, U1⦄ ⊐ ⦃G2, L2, U2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
+/3 width=3 by fqu_pair_sn, fqu_bind_dx, fqu_flat_dx, cpx_pair_sn, cpx_bind, cpx_flat, ex2_intro/
+[ #I #G #L #V2 #U2 #HVU2
+ elim (lift_total U2 0 1)
+ /4 width=7 by fqu_drop, cpx_delta, drop_pair, drop_drop, ex2_intro/
+| #G #L #K #T1 #U1 #k #HLK1 #HTU1 #T2 #HTU2
+ elim (lift_total T2 0 (k+1))
+ /3 width=11 by cpx_lift, fqu_drop, ex2_intro/
+]
+qed-.
+
+lemma fqu_sta_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 •*[h, 1] U2 →
+ ∀d. ⦃G2, L2⦄ ⊢ T2 ▪[h, o] d+1 →
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] U1 & ⦃G1, L1, U1⦄ ⊐ ⦃G2, L2, U2⦄.
+/3 width=5 by fqu_cpx_trans, sta_cpx/ qed-.
+
+lemma fquq_cpx_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡[h, o] U2 →
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] U1 & ⦃G1, L1, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #U2 #HTU2 elim (fquq_inv_gen … H) -H
+[ #HT12 elim (fqu_cpx_trans … HT12 … HTU2) /3 width=3 by fqu_fquq, ex2_intro/
+| * #H1 #H2 #H3 destruct /2 width=3 by ex2_intro/
+]
+qed-.
+
+lemma fquq_sta_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 •*[h, 1] U2 →
+ ∀d. ⦃G2, L2⦄ ⊢ T2 ▪[h, o] d+1 →
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] U1 & ⦃G1, L1, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
+/3 width=5 by fquq_cpx_trans, sta_cpx/ qed-.
+
+lemma fqup_cpx_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡[h, o] U2 →
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] U1 & ⦃G1, L1, U1⦄ ⊐+ ⦃G2, L2, U2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
+[ #G2 #L2 #T2 #H12 #U2 #HTU2 elim (fqu_cpx_trans … H12 … HTU2) -T2
+ /3 width=3 by fqu_fqup, ex2_intro/
+| #G #G2 #L #L2 #T #T2 #_ #HT2 #IHT1 #U2 #HTU2
+ elim (fqu_cpx_trans … HT2 … HTU2) -T2 #T2 #HT2 #HTU2
+ elim (IHT1 … HT2) -T /3 width=7 by fqup_strap1, ex2_intro/
+]
+qed-.
+
+lemma fqus_cpx_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡[h, o] U2 →
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] U1 & ⦃G1, L1, U1⦄ ⊐* ⦃G2, L2, U2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #U2 #HTU2 elim (fqus_inv_gen … H) -H
+[ #HT12 elim (fqup_cpx_trans … HT12 … HTU2) /3 width=3 by fqup_fqus, ex2_intro/
+| * #H1 #H2 #H3 destruct /2 width=3 by ex2_intro/
+]
+qed-.
+
+lemma fqu_cpx_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡[h, o] U2 → (T2 = U2 → ⊥) →
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐ ⦃G2, L2, U2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
+[ #I #G #L #V1 #V2 #HV12 #_ elim (lift_total V2 0 1)
+ #U2 #HVU2 @(ex3_intro … U2)
+ [1,3: /3 width=7 by fqu_drop, cpx_delta, drop_pair, drop_drop/
+ | #H destruct
+ lapply (lift_inv_lref2_be … HVU2 ? ?) -HVU2 //
+ ]
+| #I #G #L #V1 #T #V2 #HV12 #H @(ex3_intro … (②{I}V2.T))
+ [1,3: /2 width=4 by fqu_pair_sn, cpx_pair_sn/
+ | #H0 destruct /2 width=1 by/
+ ]
+| #a #I #G #L #V #T1 #T2 #HT12 #H @(ex3_intro … (ⓑ{a,I}V.T2))
+ [1,3: /2 width=4 by fqu_bind_dx, cpx_bind/
+ | #H0 destruct /2 width=1 by/
+ ]
+| #I #G #L #V #T1 #T2 #HT12 #H @(ex3_intro … (ⓕ{I}V.T2))
+ [1,3: /2 width=4 by fqu_flat_dx, cpx_flat/
+ | #H0 destruct /2 width=1 by/
+ ]
+| #G #L #K #T1 #U1 #k #HLK #HTU1 #T2 #HT12 #H elim (lift_total T2 0 (k+1))
+ #U2 #HTU2 @(ex3_intro … U2)
+ [1,3: /2 width=10 by cpx_lift, fqu_drop/
+ | #H0 destruct /3 width=5 by lift_inj/
+]
+qed-.
+
+lemma fquq_cpx_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡[h, o] U2 → (T2 = U2 → ⊥) →
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H12 #U2 #HTU2 #H elim (fquq_inv_gen … H12) -H12
+[ #H12 elim (fqu_cpx_trans_neq … H12 … HTU2 H) -T2
+ /3 width=4 by fqu_fquq, ex3_intro/
+| * #HG #HL #HT destruct /3 width=4 by ex3_intro/
+]
+qed-.
+
+lemma fqup_cpx_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡[h, o] U2 → (T2 = U2 → ⊥) →
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐+ ⦃G2, L2, U2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind_dx … H) -G1 -L1 -T1
+[ #G1 #L1 #T1 #H12 #U2 #HTU2 #H elim (fqu_cpx_trans_neq … H12 … HTU2 H) -T2
+ /3 width=4 by fqu_fqup, ex3_intro/
+| #G #G1 #L #L1 #T #T1 #H1 #_ #IH12 #U2 #HTU2 #H elim (IH12 … HTU2 H) -T2
+ #U1 #HTU1 #H #H12 elim (fqu_cpx_trans_neq … H1 … HTU1 H) -T1
+ /3 width=8 by fqup_strap2, ex3_intro/
+]
+qed-.
+
+lemma fqus_cpx_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡[h, o] U2 → (T2 = U2 → ⊥) →
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐* ⦃G2, L2, U2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H12 #U2 #HTU2 #H elim (fqus_inv_gen … H12) -H12
+[ #H12 elim (fqup_cpx_trans_neq … H12 … HTU2 H) -T2
+ /3 width=4 by fqup_fqus, ex3_intro/
+| * #HG #HL #HT destruct /3 width=4 by ex3_intro/
+]
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/multiple/lleq_drop.ma".
+include "basic_2/reduction/cpx_llpx_sn.ma".
+
+(* CONTEXT-SENSITIVE EXTENDED PARALLEL REDUCTION FOR TERMS ******************)
+
+(* Properties on lazy equivalence for local environments ********************)
+
+lemma lleq_cpx_trans: ∀h,o,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ➡[h, o] T2 →
+ ∀L1. L1 ≡[T1, 0] L2 → ⦃G, L1⦄ ⊢ T1 ➡[h, o] T2.
+#h #o #G #L2 #T1 #T2 #H elim H -G -L2 -T1 -T2 /2 width=2 by cpx_st/
+[ #I #G #L2 #K2 #V1 #V2 #W2 #i #HLK2 #_ #HVW2 #IHV12 #L1 #H elim (lleq_fwd_lref_dx … H … HLK2) -L2
+ [ #H elim (ylt_yle_false … H) //
+ | * /3 width=7 by cpx_delta/
+ ]
+| #a #I #G #L2 #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L1 #H elim (lleq_inv_bind_O … H) -H
+ /3 width=1 by cpx_bind/
+| #I #G #L2 #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L1 #H elim (lleq_inv_flat … H) -H
+ /3 width=1 by cpx_flat/
+| #G #L2 #V2 #T1 #T #T2 #_ #HT2 #IHT1 #L1 #H elim (lleq_inv_bind_O … H) -H
+ /3 width=3 by cpx_zeta/
+| #G #L2 #W1 #T1 #T2 #_ #IHT12 #L1 #H elim (lleq_inv_flat … H) -H
+ /3 width=1 by cpx_eps/
+| #G #L2 #W1 #W2 #T1 #_ #IHW12 #L1 #H elim (lleq_inv_flat … H) -H
+ /3 width=1 by cpx_ct/
+| #a #G #L1 #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #L1 #H elim (lleq_inv_flat … H) -H
+ #HV1 #H elim (lleq_inv_bind_O … H) -H /3 width=1 by cpx_beta/
+| #a #G #L1 #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV1 #IHW12 #IHT12 #L1 #H elim (lleq_inv_flat … H) -H
+ #HV1 #H elim (lleq_inv_bind_O … H) -H /3 width=3 by cpx_theta/
+]
+qed-.
+
+lemma cpx_lleq_conf: ∀h,o,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ➡[h, o] T2 →
+ ∀L1. L2 ≡[T1, 0] L1 → ⦃G, L1⦄ ⊢ T1 ➡[h, o] T2.
+/3 width=3 by lleq_cpx_trans, lleq_sym/ qed-.
+
+lemma cpx_lleq_conf_sn: ∀h,o,G. c_r_confluent1 … (cpx h o G) (lleq 0).
+/3 width=6 by cpx_llpx_sn_conf, lift_mono, ex2_intro/ qed-.
+
+lemma cpx_lleq_conf_dx: ∀h,o,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ➡[h, o] T2 →
+ ∀L1. L1 ≡[T1, 0] L2 → L1 ≡[T2, 0] L2.
+/4 width=6 by cpx_lleq_conf_sn, 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_drop.ma".
+include "basic_2/reduction/cpx.ma".
+
+(* CONTEXT-SENSITIVE EXTENDED PARALLEL REDUCTION FOR TERMS ******************)
+
+(* Properties on lazy sn pointwise extensions *******************************)
+
+(* Note: lemma 1000 *)
+lemma cpx_llpx_sn_conf: ∀R. (∀L. reflexive … (R L)) → d_liftable R → d_deliftable_sn R →
+ ∀h,o,G. c_r_confluent1 … (cpx h o G) (llpx_sn R 0).
+#R #H1R #H2R #H3R #h #o #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 #V1c #V2c #W2c #i #HLKs #_ #HVW2c #IHV12c #Ld #H elim (llpx_sn_inv_lref_ge_sn … H … HLKs) // -H
+ #Kd #V1l #HLKd #HV1c #HV1sd
+ lapply (drop_fwd_drop2 … HLKs) -HLKs #HLKs
+ lapply (drop_fwd_drop2 … HLKd) -HLKd #HLKd
+ @(llpx_sn_lift_le … HLKs HLKd … HVW2c) -HLKs -HLKd -HVW2c /2 width=1 by/ (**) (* full auto too slow *)
+| #a #I #G #Ls #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #Ld #H elim (llpx_sn_inv_bind_O … H) -H
+ /4 width=5 by llpx_sn_bind_repl_SO, llpx_sn_bind/
+| #I #G #Ls #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
+ /3 width=1 by llpx_sn_flat/
+| #G #Ls #V #T1 #T2 #T #_ #HT2 #IHT12 #Ld #H elim (llpx_sn_inv_bind_O … H) -H
+ /3 width=10 by llpx_sn_inv_lift_le, drop_drop/
+| #G #Ls #V #T1 #T2 #_ #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H /2 width=1 by/
+| #G #Ls #V1 #V2 #T #_ #IHV12 #Ld #H elim (llpx_sn_inv_flat … H) -H /2 width=1 by/
+| #a #G #Ls #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
+ #HV1 #H elim (llpx_sn_inv_bind_O … H) -H
+ /4 width=5 by llpx_sn_bind_repl_SO, llpx_sn_flat, llpx_sn_bind/
+| #a #G #Ls #V1 #V2 #V #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
+ #HV1 #H elim (llpx_sn_inv_bind_O … H) -H //
+ #HW1 #HT1 @llpx_sn_bind_O /2 width=1 by/ @llpx_sn_flat (**) (* full auto too slow *)
+ [ /3 width=10 by llpx_sn_lift_le, drop_drop/
+ | /3 width=4 by llpx_sn_bind_repl_O/
+]
+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_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,o,G. lsub_trans … (cpx h o G) (lreq 0 (∞)).
+#h #o #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-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/btpredproper_8.ma".
+include "basic_2/substitution/fqu.ma".
+include "basic_2/multiple/lleq.ma".
+include "basic_2/reduction/lpx.ma".
+
+(* "RST" PROPER PARALLEL COMPUTATION FOR CLOSURES ***************************)
+
+inductive fpb (h) (o) (G1) (L1) (T1): relation3 genv lenv term ≝
+| fpb_fqu: ∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → fpb h o G1 L1 T1 G2 L2 T2
+| fpb_cpx: ∀T2. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] T2 → (T1 = T2 → ⊥) → fpb h o G1 L1 T1 G1 L1 T2
+| fpb_lpx: ∀L2. ⦃G1, L1⦄ ⊢ ➡[h, o] L2 → (L1 ≡[T1, 0] L2 → ⊥) → fpb h o G1 L1 T1 G1 L2 T1
+.
+
+interpretation
+ "'rst' proper parallel reduction (closure)"
+ 'BTPRedProper h o G1 L1 T1 G2 L2 T2 = (fpb h o G1 L1 T1 G2 L2 T2).
+
+(* Basic properties *********************************************************)
+
+lemma cpr_fpb: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → (T1 = T2 → ⊥) →
+ ⦃G, L, T1⦄ ≻[h, o] ⦃G, L, T2⦄.
+/3 width=1 by fpb_cpx, cpr_cpx/ qed.
+
+lemma lpr_fpb: ∀h,o,G,L1,L2,T. ⦃G, L1⦄ ⊢ ➡ L2 → (L1 ≡[T, 0] L2 → ⊥) →
+ ⦃G, L1, T⦄ ≻[h, o] ⦃G, L2, T⦄.
+/3 width=1 by fpb_lpx, lpr_lpx/ 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/fleq.ma".
+include "basic_2/reduction/fpb_lleq.ma".
+
+(* "RST" PROPER PARALLEL COMPUTATION FOR CLOSURES ***************************)
+
+(* Properties on lazy equivalence for closures ******************************)
+
+lemma fleq_fpb_trans: ∀h,o,F1,F2,K1,K2,T1,T2. ⦃F1, K1, T1⦄ ≡[0] ⦃F2, K2, T2⦄ →
+ ∀G2,L2,U2. ⦃F2, K2, T2⦄ ≻[h, o] ⦃G2, L2, U2⦄ →
+ ∃∃G1,L1,U1. ⦃F1, K1, T1⦄ ≻[h, o] ⦃G1, L1, U1⦄ & ⦃G1, L1, U1⦄ ≡[0] ⦃G2, L2, U2⦄.
+#h #o #F1 #F2 #K1 #K2 #T1 #T2 * -F2 -K2 -T2
+#K2 #HK12 #G2 #L2 #U2 #H12 elim (lleq_fpb_trans … HK12 … H12) -K2
+/3 width=5 by fleq_intro, ex2_3_intro/
+qed-.
+
+(* Inversion lemmas on lazy equivalence for closures ************************)
+
+lemma fpb_inv_fleq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → ⊥.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2
+[ #G2 #L2 #T2 #H12 #H elim (fleq_inv_gen … H) -H
+ /3 width=8 by lleq_fwd_length, fqu_inv_eq/
+| #T2 #_ #HnT #H elim (fleq_inv_gen … H) -H /2 width=1 by/
+| #L2 #_ #HnL #H elim (fleq_inv_gen … H) -H /2 width=1 by/
+]
+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/unfold/lstas_da.ma".
+include "basic_2/reduction/cpx_lift.ma".
+include "basic_2/reduction/fpb.ma".
+
+(* "RST" PROPER PARALLEL COMPUTATION FOR CLOSURES ***************************)
+
+(* Advanced properties ******************************************************)
+
+lemma sta_fpb: ∀h,o,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 ▪[h, o] d+1 →
+ ⦃G, L⦄ ⊢ T1 •*[h, 1] T2 → ⦃G, L, T1⦄ ≻[h, o] ⦃G, L, T2⦄.
+#h #o #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.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/multiple/lleq_fqus.ma".
+include "basic_2/multiple/lleq_lleq.ma".
+include "basic_2/reduction/lpx_lleq.ma".
+include "basic_2/reduction/fpb.ma".
+
+(* "RST" PROPER PARALLEL COMPUTATION FOR CLOSURES ***************************)
+
+(* Properties on lazy equivalence for local environments ********************)
+
+lemma lleq_fpb_trans: ∀h,o,F,K1,K2,T. K1 ≡[T, 0] K2 →
+ ∀G,L2,U. ⦃F, K2, T⦄ ≻[h, o] ⦃G, L2, U⦄ →
+ ∃∃L1. ⦃F, K1, T⦄ ≻[h, o] ⦃G, L1, U⦄ & L1 ≡[U, 0] L2.
+#h #o #F #K1 #K2 #T #HT #G #L2 #U * -G -L2 -U
+[ #G #L2 #U #H2 elim (lleq_fqu_trans … H2 … HT) -K2
+ /3 width=3 by fpb_fqu, ex2_intro/
+| /4 width=10 by fpb_cpx, cpx_lleq_conf_sn, lleq_cpx_trans, ex2_intro/
+| #L2 #HKL2 #HnKL2 elim (lleq_lpx_trans … HKL2 … HT) -HKL2
+ /6 width=3 by fpb_lpx, lleq_canc_sn, ex2_intro/ (* 2 lleq_canc_sn *)
+]
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/btpred_8.ma".
+include "basic_2/substitution/fquq.ma".
+include "basic_2/multiple/lleq.ma".
+include "basic_2/reduction/lpx.ma".
+
+(* "QRST" PARALLEL REDUCTION FOR CLOSURES ***********************************)
+
+inductive fpbq (h) (o) (G1) (L1) (T1): relation3 genv lenv term ≝
+| fpbq_fquq: ∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → fpbq h o G1 L1 T1 G2 L2 T2
+| fpbq_cpx : ∀T2. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] T2 → fpbq h o G1 L1 T1 G1 L1 T2
+| fpbq_lpx : ∀L2. ⦃G1, L1⦄ ⊢ ➡[h, o] L2 → fpbq h o G1 L1 T1 G1 L2 T1
+| fpbq_lleq: ∀L2. L1 ≡[T1, 0] L2 → fpbq h o G1 L1 T1 G1 L2 T1
+.
+
+interpretation
+ "'qrst' parallel reduction (closure)"
+ 'BTPRed h o G1 L1 T1 G2 L2 T2 = (fpbq h o G1 L1 T1 G2 L2 T2).
+
+(* Basic properties *********************************************************)
+
+lemma fpbq_refl: ∀h,o. tri_reflexive … (fpbq h o).
+/2 width=1 by fpbq_cpx/ qed.
+
+lemma cpr_fpbq: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ⦃G, L, T1⦄ ≽[h, o] ⦃G, L, T2⦄.
+/3 width=1 by fpbq_cpx, cpr_cpx/ qed.
+
+lemma lpr_fpbq: ∀h,o,G,L1,L2,T. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L1, T⦄ ≽[h, o] ⦃G, L2, T⦄.
+/3 width=1 by fpbq_lpx, lpr_lpx/ qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/static/aaa_fqus.ma".
+include "basic_2/static/aaa_lleq.ma".
+include "basic_2/reduction/lpx_aaa.ma".
+include "basic_2/reduction/fpbq.ma".
+
+(* "QRST" PARALLEL REDUCTION FOR CLOSURES ***********************************)
+
+(* Properties on atomic arity assignment for terms **************************)
+
+lemma fpbq_aaa_conf: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ →
+ ∀A1. ⦃G1, L1⦄ ⊢ T1 ⁝ A1 → ∃A2. ⦃G2, L2⦄ ⊢ T2 ⁝ A2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2
+/3 width=6 by aaa_lleq_conf, lpx_aaa_conf, cpx_aaa_conf, aaa_fquq_conf, ex_intro/
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/btpredalt_8.ma".
+include "basic_2/reduction/fpb_fleq.ma".
+include "basic_2/reduction/fpbq.ma".
+
+(* "QRST" PARALLEL REDUCTION FOR CLOSURES ***********************************)
+
+(* alternative definition of fpbq *)
+definition fpbqa: ∀h. sd h → tri_relation genv lenv term ≝
+ λh,o,G1,L1,T1,G2,L2,T2.
+ ⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ ∨ ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄.
+
+interpretation
+ "'qrst' parallel reduction (closure) alternative"
+ 'BTPRedAlt h o G1 L1 T1 G2 L2 T2 = (fpbqa h o G1 L1 T1 G2 L2 T2).
+
+(* Basic properties *********************************************************)
+
+lemma fleq_fpbq: ∀h,o,G1,G2,L1,L2,T1,T2.
+ ⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 * /2 width=1 by fpbq_lleq/
+qed.
+
+lemma fpb_fpbq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2
+/3 width=1 by fpbq_fquq, fpbq_cpx, fpbq_lpx, fqu_fquq/
+qed.
+
+(* Main properties **********************************************************)
+
+theorem fpbq_fpbqa: ∀h,o,G1,G2,L1,L2,T1,T2.
+ ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ≽≽[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2
+[ #G2 #L2 #T2 #H elim (fquq_inv_gen … H) -H
+ [ /3 width=1 by fpb_fqu, or_intror/
+ | * #H1 #H2 #H3 destruct /2 width=1 by or_introl/
+ ]
+| #T2 #HT12 elim (eq_term_dec T1 T2)
+ #HnT12 destruct /4 width=1 by fpb_cpx, or_intror, or_introl/
+| #L2 #HL12 elim (lleq_dec … T1 L1 L2 0)
+ /4 width=1 by fpb_lpx, fleq_intro, or_intror, or_introl/
+| /3 width=1 by fleq_intro, or_introl/
+]
+qed.
+
+theorem fpbqa_inv_fpbq: ∀h,o,G1,G2,L1,L2,T1,T2.
+ ⦃G1, L1, T1⦄ ≽≽[h, o] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 * /2 width=1 by fleq_fpbq, fpb_fpbq/
+qed-.
+
+(* Advanced eliminators *****************************************************)
+
+lemma fpbq_ind_alt: ∀h,o,G1,G2,L1,L2,T1,T2. ∀R:Prop.
+ (⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → R) →
+ (⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ → R) →
+ ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ → R.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #R #HI1 #HI2 #H elim (fpbq_fpbqa … H) /2 width=1 by/
+qed-.
+
+(* aternative definition of fpb *********************************************)
+
+lemma fpb_fpbq_alt: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ ∧ (⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → ⊥).
+/3 width=10 by fpb_fpbq, fpb_inv_fleq, conj/ qed.
+
+lemma fpbq_inv_fpb_alt: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ →
+ (⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → ⊥) → ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #H0 @(fpbq_ind_alt … H) -H //
+#H elim H0 -H0 //
+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_lift.ma".
+include "basic_2/reduction/fpbq.ma".
+
+(* "QRST" PARALLEL REDUCTION FOR CLOSURES ***********************************)
+
+(* Advanced properties ******************************************************)
+
+lemma sta_fpbq: ∀h,o,G,L,T1,T2,d.
+ ⦃G, L⦄ ⊢ T1 ▪[h, o] d+1 → ⦃G, L⦄ ⊢ T1 •*[h, 1] T2 →
+ ⦃G, L, T1⦄ ≽[h, o] ⦃G, L, T2⦄.
+/3 width=4 by fpbq_cpx, sta_cpx/ qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/predsn_3.ma".
+include "basic_2/substitution/lpx_sn.ma".
+include "basic_2/reduction/cpr.ma".
+
+(* SN PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS *****************************)
+
+definition lpr: relation3 genv lenv lenv ≝ λG. lpx_sn (cpr G).
+
+interpretation "parallel reduction (local environment, sn variant)"
+ 'PRedSn G L1 L2 = (lpr G L1 L2).
+
+(* Basic inversion lemmas ***************************************************)
+
+(* Basic_1: includes: wcpr0_gen_sort *)
+lemma lpr_inv_atom1: ∀G,L2. ⦃G, ⋆⦄ ⊢ ➡ L2 → L2 = ⋆.
+/2 width=4 by lpx_sn_inv_atom1_aux/ qed-.
+
+(* Basic_1: includes: wcpr0_gen_head *)
+lemma lpr_inv_pair1: ∀I,G,K1,V1,L2. ⦃G, K1.ⓑ{I}V1⦄ ⊢ ➡ L2 →
+ ∃∃K2,V2. ⦃G, K1⦄ ⊢ ➡ K2 & ⦃G, K1⦄ ⊢ V1 ➡ V2 & L2 = K2.ⓑ{I}V2.
+/2 width=3 by lpx_sn_inv_pair1_aux/ qed-.
+
+lemma lpr_inv_atom2: ∀G,L1. ⦃G, L1⦄ ⊢ ➡ ⋆ → L1 = ⋆.
+/2 width=4 by lpx_sn_inv_atom2_aux/ qed-.
+
+lemma lpr_inv_pair2: ∀I,G,L1,K2,V2. ⦃G, L1⦄ ⊢ ➡ K2.ⓑ{I}V2 →
+ ∃∃K1,V1. ⦃G, K1⦄ ⊢ ➡ K2 & ⦃G, K1⦄ ⊢ V1 ➡ V2 & L1 = K1. ⓑ{I} V1.
+/2 width=3 by lpx_sn_inv_pair2_aux/ qed-.
+
+(* Basic properties *********************************************************)
+
+(* Note: lemma 250 *)
+lemma lpr_refl: ∀G,L. ⦃G, L⦄ ⊢ ➡ L.
+/2 width=1 by lpx_sn_refl/ qed.
+
+lemma lpr_pair: ∀I,G,K1,K2,V1,V2. ⦃G, K1⦄ ⊢ ➡ K2 → ⦃G, K1⦄ ⊢ V1 ➡ V2 →
+ ⦃G, K1.ⓑ{I}V1⦄ ⊢ ➡ K2.ⓑ{I}V2.
+/2 width=1 by lpx_sn_pair/ qed.
+
+(* Basic forward lemmas *****************************************************)
+
+lemma lpr_fwd_length: ∀G,L1,L2. ⦃G, L1⦄ ⊢ ➡ L2 → |L1| = |L2|.
+/2 width=2 by lpx_sn_fwd_length/ qed-.
+
+(* Basic_1: removed theorems 3: wcpr0_getl wcpr0_getl_back
+ pr0_subst1_back
+*)
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/substitution/lpx_sn_drop.ma".
+include "basic_2/substitution/fquq_alt.ma".
+include "basic_2/reduction/cpr_lift.ma".
+include "basic_2/reduction/lpr.ma".
+
+(* SN PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS *****************************)
+
+(* Properties on local environment slicing ***********************************)
+
+(* Basic_1: includes: wcpr0_drop *)
+lemma lpr_drop_conf: ∀G. dropable_sn (lpr G).
+/3 width=6 by lpx_sn_deliftable_dropable, cpr_inv_lift1/ qed-.
+
+(* Basic_1: includes: wcpr0_drop_back *)
+lemma drop_lpr_trans: ∀G. dedropable_sn (lpr G).
+/3 width=10 by lpx_sn_liftable_dedropable, cpr_lift/ qed-.
+
+lemma lpr_drop_trans_O1: ∀G. dropable_dx (lpr G).
+/2 width=3 by lpx_sn_dropable/ qed-.
+
+(* Properties on context-sensitive parallel reduction for terms *************)
+
+lemma fqu_cpr_trans_dx: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡ U2 →
+ ∃∃L,U1. ⦃G1, L1⦄ ⊢ ➡ L & ⦃G1, L⦄ ⊢ T1 ➡ U1 & ⦃G1, L, U1⦄ ⊐ ⦃G2, L2, U2⦄.
+#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
+/3 width=5 by fqu_lref_O, fqu_pair_sn, fqu_bind_dx, fqu_flat_dx, lpr_pair, cpr_pair_sn, cpr_atom, cpr_bind, cpr_flat, ex3_2_intro/
+#G #L #K #U #T #k #HLK #HUT #U2 #HU2
+elim (lift_total U2 0 (k+1)) #T2 #HUT2
+lapply (cpr_lift … HU2 … HLK … HUT … HUT2) -HU2 -HUT /3 width=9 by fqu_drop, ex3_2_intro/
+qed-.
+
+lemma fquq_cpr_trans_dx: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡ U2 →
+ ∃∃L,U1. ⦃G1, L1⦄ ⊢ ➡ L & ⦃G1, L⦄ ⊢ T1 ➡ U1 & ⦃G1, L, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
+#G1 #G2 #L1 #L2 #T1 #T2 #H #U2 #HTU2 elim (fquq_inv_gen … H) -H
+[ #HT12 elim (fqu_cpr_trans_dx … HT12 … HTU2) /3 width=5 by fqu_fquq, ex3_2_intro/
+| * #H1 #H2 #H3 destruct /2 width=5 by ex3_2_intro/
+]
+qed-.
+
+lemma fqu_cpr_trans_sn: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡ U2 →
+ ∃∃L,U1. ⦃G1, L1⦄ ⊢ ➡ L & ⦃G1, L1⦄ ⊢ T1 ➡ U1 & ⦃G1, L, U1⦄ ⊐ ⦃G2, L2, U2⦄.
+#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
+/3 width=5 by fqu_lref_O, fqu_pair_sn, fqu_bind_dx, fqu_flat_dx, lpr_pair, cpr_pair_sn, cpr_atom, cpr_bind, cpr_flat, ex3_2_intro/
+#G #L #K #U #T #k #HLK #HUT #U2 #HU2
+elim (lift_total U2 0 (k+1)) #T2 #HUT2
+lapply (cpr_lift … HU2 … HLK … HUT … HUT2) -HU2 -HUT /3 width=9 by fqu_drop, ex3_2_intro/
+qed-.
+
+lemma fquq_cpr_trans_sn: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡ U2 →
+ ∃∃L,U1. ⦃G1, L1⦄ ⊢ ➡ L & ⦃G1, L1⦄ ⊢ T1 ➡ U1 & ⦃G1, L, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
+#G1 #G2 #L1 #L2 #T1 #T2 #H #U2 #HTU2 elim (fquq_inv_gen … H) -H
+[ #HT12 elim (fqu_cpr_trans_sn … HT12 … HTU2) /3 width=5 by fqu_fquq, ex3_2_intro/
+| * #H1 #H2 #H3 destruct /2 width=5 by ex3_2_intro/
+]
+qed-.
+
+lemma fqu_lpr_trans: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+ ∀K2. ⦃G2, L2⦄ ⊢ ➡ K2 →
+ ∃∃K1,T. ⦃G1, L1⦄ ⊢ ➡ K1 & ⦃G1, L1⦄ ⊢ T1 ➡ T & ⦃G1, K1, T⦄ ⊐ ⦃G2, K2, T2⦄.
+#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
+/3 width=5 by fqu_lref_O, fqu_pair_sn, fqu_flat_dx, lpr_pair, ex3_2_intro/
+[ #a #I #G2 #L2 #V2 #T2 #X #H elim (lpr_inv_pair1 … H) -H
+ #K2 #W2 #HLK2 #HVW2 #H destruct
+ /3 width=5 by fqu_fquq, cpr_pair_sn, fqu_bind_dx, ex3_2_intro/
+| #G #L1 #K1 #T1 #U1 #k #HLK1 #HTU1 #K2 #HK12
+ elim (drop_lpr_trans … HLK1 … HK12) -HK12
+ /3 width=7 by fqu_drop, ex3_2_intro/
+]
+qed-.
+
+lemma fquq_lpr_trans: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
+ ∀K2. ⦃G2, L2⦄ ⊢ ➡ K2 →
+ ∃∃K1,T. ⦃G1, L1⦄ ⊢ ➡ K1 & ⦃G1, L1⦄ ⊢ T1 ➡ T & ⦃G1, K1, T⦄ ⊐⸮ ⦃G2, K2, T2⦄.
+#G1 #G2 #L1 #L2 #T1 #T2 #H #K2 #HLK2 elim (fquq_inv_gen … H) -H
+[ #HT12 elim (fqu_lpr_trans … HT12 … HLK2) /3 width=5 by fqu_fquq, ex3_2_intro/
+| * #H1 #H2 #H3 destruct /2 width=5 by ex3_2_intro/
+]
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/substitution/lpx_sn_lpx_sn.ma".
+include "basic_2/multiple/fqup.ma".
+include "basic_2/reduction/lpr_drop.ma".
+
+(* SN PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS *****************************)
+
+(* Main properties on context-sensitive parallel reduction for terms ********)
+
+fact cpr_conf_lpr_atom_atom:
+ ∀I,G,L1,L2. ∃∃T. ⦃G, L1⦄ ⊢ ⓪{I} ➡ T & ⦃G, L2⦄ ⊢ ⓪{I} ➡ T.
+/2 width=3 by cpr_atom, ex2_intro/ qed-.
+
+fact cpr_conf_lpr_atom_delta:
+ ∀G,L0,i. (
+ ∀L,T. ⦃G, L0, #i⦄ ⊐+ ⦃G, L, T⦄ →
+ ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
+ ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
+ ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
+ ) →
+ ∀K0,V0. ⬇[i] L0 ≡ K0.ⓓV0 →
+ ∀V2. ⦃G, K0⦄ ⊢ V0 ➡ V2 → ∀T2. ⬆[O, i + 1] V2 ≡ T2 →
+ ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
+ ∃∃T. ⦃G, L1⦄ ⊢ #i ➡ T & ⦃G, L2⦄ ⊢ T2 ➡ T.
+#G #L0 #i #IH #K0 #V0 #HLK0 #V2 #HV02 #T2 #HVT2 #L1 #HL01 #L2 #HL02
+elim (lpr_drop_conf … HLK0 … HL01) -HL01 #X1 #H1 #HLK1
+elim (lpr_inv_pair1 … H1) -H1 #K1 #V1 #HK01 #HV01 #H destruct
+elim (lpr_drop_conf … HLK0 … HL02) -HL02 #X2 #H2 #HLK2
+elim (lpr_inv_pair1 … H2) -H2 #K2 #W2 #HK02 #_ #H destruct
+lapply (drop_fwd_drop2 … HLK2) -W2 #HLK2
+lapply (fqup_lref … G … HLK0) -HLK0 #HLK0
+elim (IH … HLK0 … HV01 … HV02 … HK01 … HK02) -L0 -K0 -V0 #V #HV1 #HV2
+elim (lift_total V 0 (i+1))
+/3 width=12 by cpr_lift, cpr_delta, ex2_intro/
+qed-.
+
+(* Basic_1: includes: pr0_delta_delta pr2_delta_delta *)
+fact cpr_conf_lpr_delta_delta:
+ ∀G,L0,i. (
+ ∀L,T. ⦃G, L0, #i⦄ ⊐+ ⦃G, L, T⦄ →
+ ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
+ ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
+ ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
+ ) →
+ ∀K0,V0. ⬇[i] L0 ≡ K0.ⓓV0 →
+ ∀V1. ⦃G, K0⦄ ⊢ V0 ➡ V1 → ∀T1. ⬆[O, i + 1] V1 ≡ T1 →
+ ∀KX,VX. ⬇[i] L0 ≡ KX.ⓓVX →
+ ∀V2. ⦃G, KX⦄ ⊢ VX ➡ V2 → ∀T2. ⬆[O, i + 1] V2 ≡ T2 →
+ ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
+ ∃∃T. ⦃G, L1⦄ ⊢ T1 ➡ T & ⦃G, L2⦄ ⊢ T2 ➡ T.
+#G #L0 #i #IH #K0 #V0 #HLK0 #V1 #HV01 #T1 #HVT1
+#KX #VX #H #V2 #HV02 #T2 #HVT2 #L1 #HL01 #L2 #HL02
+lapply (drop_mono … H … HLK0) -H #H destruct
+elim (lpr_drop_conf … HLK0 … HL01) -HL01 #X1 #H1 #HLK1
+elim (lpr_inv_pair1 … H1) -H1 #K1 #W1 #HK01 #_ #H destruct
+lapply (drop_fwd_drop2 … HLK1) -W1 #HLK1
+elim (lpr_drop_conf … HLK0 … HL02) -HL02 #X2 #H2 #HLK2
+elim (lpr_inv_pair1 … H2) -H2 #K2 #W2 #HK02 #_ #H destruct
+lapply (drop_fwd_drop2 … HLK2) -W2 #HLK2
+lapply (fqup_lref … G … HLK0) -HLK0 #HLK0
+elim (IH … HLK0 … HV01 … HV02 … HK01 … HK02) -L0 -K0 -V0 #V #HV1 #HV2
+elim (lift_total V 0 (i+1)) /3 width=12 by cpr_lift, ex2_intro/
+qed-.
+
+fact cpr_conf_lpr_bind_bind:
+ ∀a,I,G,L0,V0,T0. (
+ ∀L,T. ⦃G, L0, ⓑ{a,I}V0.T0⦄ ⊐+ ⦃G, L, T⦄ →
+ ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
+ ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
+ ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
+ ) →
+ ∀V1. ⦃G, L0⦄ ⊢ V0 ➡ V1 → ∀T1. ⦃G, L0.ⓑ{I}V0⦄ ⊢ T0 ➡ T1 →
+ ∀V2. ⦃G, L0⦄ ⊢ V0 ➡ V2 → ∀T2. ⦃G, L0.ⓑ{I}V0⦄ ⊢ T0 ➡ T2 →
+ ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
+ ∃∃T. ⦃G, L1⦄ ⊢ ⓑ{a,I}V1.T1 ➡ T & ⦃G, L2⦄ ⊢ ⓑ{a,I}V2.T2 ➡ T.
+#a #I #G #L0 #V0 #T0 #IH #V1 #HV01 #T1 #HT01
+#V2 #HV02 #T2 #HT02 #L1 #HL01 #L2 #HL02
+elim (IH … HV01 … HV02 … HL01 … HL02) //
+elim (IH … HT01 … HT02 (L1.ⓑ{I}V1) … (L2.ⓑ{I}V2)) -IH
+/3 width=5 by lpr_pair, cpr_bind, ex2_intro/
+qed-.
+
+fact cpr_conf_lpr_bind_zeta:
+ ∀G,L0,V0,T0. (
+ ∀L,T. ⦃G, L0, +ⓓV0.T0⦄ ⊐+ ⦃G, L, T⦄ →
+ ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
+ ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
+ ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
+ ) →
+ ∀V1. ⦃G, L0⦄ ⊢ V0 ➡ V1 → ∀T1. ⦃G, L0.ⓓV0⦄ ⊢ T0 ➡ T1 →
+ ∀T2. ⦃G, L0.ⓓV0⦄ ⊢ T0 ➡ T2 → ∀X2. ⬆[O, 1] X2 ≡ T2 →
+ ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
+ ∃∃T. ⦃G, L1⦄ ⊢ +ⓓV1.T1 ➡ T & ⦃G, L2⦄ ⊢ X2 ➡ T.
+#G #L0 #V0 #T0 #IH #V1 #HV01 #T1 #HT01
+#T2 #HT02 #X2 #HXT2 #L1 #HL01 #L2 #HL02
+elim (IH … HT01 … HT02 (L1.ⓓV1) … (L2.ⓓV1)) -IH -HT01 -HT02 /2 width=1 by lpr_pair/ -L0 -V0 -T0 #T #HT1 #HT2
+elim (cpr_inv_lift1 … HT2 L2 … HXT2) -T2 /3 width=3 by cpr_zeta, drop_drop, ex2_intro/
+qed-.
+
+fact cpr_conf_lpr_zeta_zeta:
+ ∀G,L0,V0,T0. (
+ ∀L,T. ⦃G, L0, +ⓓV0.T0⦄ ⊐+ ⦃G, L, T⦄ →
+ ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
+ ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
+ ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
+ ) →
+ ∀T1. ⦃G, L0.ⓓV0⦄ ⊢ T0 ➡ T1 → ∀X1. ⬆[O, 1] X1 ≡ T1 →
+ ∀T2. ⦃G, L0.ⓓV0⦄ ⊢ T0 ➡ T2 → ∀X2. ⬆[O, 1] X2 ≡ T2 →
+ ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
+ ∃∃T. ⦃G, L1⦄ ⊢ X1 ➡ T & ⦃G, L2⦄ ⊢ X2 ➡ T.
+#G #L0 #V0 #T0 #IH #T1 #HT01 #X1 #HXT1
+#T2 #HT02 #X2 #HXT2 #L1 #HL01 #L2 #HL02
+elim (IH … HT01 … HT02 (L1.ⓓV0) … (L2.ⓓV0)) -IH -HT01 -HT02 /2 width=1 by lpr_pair/ -L0 -T0 #T #HT1 #HT2
+elim (cpr_inv_lift1 … HT1 L1 … HXT1) -T1 /2 width=2 by drop_drop/ #T1 #HT1 #HXT1
+elim (cpr_inv_lift1 … HT2 L2 … HXT2) -T2 /2 width=2 by drop_drop/ #T2 #HT2 #HXT2
+lapply (lift_inj … HT2 … HT1) -T #H destruct /2 width=3 by ex2_intro/
+qed-.
+
+fact cpr_conf_lpr_flat_flat:
+ ∀I,G,L0,V0,T0. (
+ ∀L,T. ⦃G, L0, ⓕ{I}V0.T0⦄ ⊐+ ⦃G, L, T⦄ →
+ ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
+ ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
+ ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
+ ) →
+ ∀V1. ⦃G, L0⦄ ⊢ V0 ➡ V1 → ∀T1. ⦃G, L0⦄ ⊢ T0 ➡ T1 →
+ ∀V2. ⦃G, L0⦄ ⊢ V0 ➡ V2 → ∀T2. ⦃G, L0⦄ ⊢ T0 ➡ T2 →
+ ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
+ ∃∃T. ⦃G, L1⦄ ⊢ ⓕ{I}V1.T1 ➡ T & ⦃G, L2⦄ ⊢ ⓕ{I}V2.T2 ➡ T.
+#I #G #L0 #V0 #T0 #IH #V1 #HV01 #T1 #HT01
+#V2 #HV02 #T2 #HT02 #L1 #HL01 #L2 #HL02
+elim (IH … HV01 … HV02 … HL01 … HL02) //
+elim (IH … HT01 … HT02 … HL01 … HL02) /3 width=5 by cpr_flat, ex2_intro/
+qed-.
+
+fact cpr_conf_lpr_flat_eps:
+ ∀G,L0,V0,T0. (
+ ∀L,T. ⦃G, L0, ⓝV0.T0⦄ ⊐+ ⦃G, L, T⦄ →
+ ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
+ ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
+ ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
+ ) →
+ ∀V1,T1. ⦃G, L0⦄ ⊢ T0 ➡ T1 → ∀T2. ⦃G, L0⦄ ⊢ T0 ➡ T2 →
+ ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
+ ∃∃T. ⦃G, L1⦄ ⊢ ⓝV1.T1 ➡ T & ⦃G, L2⦄ ⊢ T2 ➡ T.
+#G #L0 #V0 #T0 #IH #V1 #T1 #HT01
+#T2 #HT02 #L1 #HL01 #L2 #HL02
+elim (IH … HT01 … HT02 … HL01 … HL02) // -L0 -V0 -T0 /3 width=3 by cpr_eps, ex2_intro/
+qed-.
+
+fact cpr_conf_lpr_eps_eps:
+ ∀G,L0,V0,T0. (
+ ∀L,T. ⦃G, L0, ⓝV0.T0⦄ ⊐+ ⦃G, L, T⦄ →
+ ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
+ ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
+ ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
+ ) →
+ ∀T1. ⦃G, L0⦄ ⊢ T0 ➡ T1 → ∀T2. ⦃G, L0⦄ ⊢ T0 ➡ T2 →
+ ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
+ ∃∃T. ⦃G, L1⦄ ⊢ T1 ➡ T & ⦃G, L2⦄ ⊢ T2 ➡ T.
+#G #L0 #V0 #T0 #IH #T1 #HT01
+#T2 #HT02 #L1 #HL01 #L2 #HL02
+elim (IH … HT01 … HT02 … HL01 … HL02) // -L0 -V0 -T0 /2 width=3 by ex2_intro/
+qed-.
+
+fact cpr_conf_lpr_flat_beta:
+ ∀a,G,L0,V0,W0,T0. (
+ ∀L,T. ⦃G, L0, ⓐV0.ⓛ{a}W0.T0⦄ ⊐+ ⦃G, L, T⦄ →
+ ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
+ ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
+ ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
+ ) →
+ ∀V1. ⦃G, L0⦄ ⊢ V0 ➡ V1 → ∀T1. ⦃G, L0⦄ ⊢ ⓛ{a}W0.T0 ➡ T1 →
+ ∀V2. ⦃G, L0⦄ ⊢ V0 ➡ V2 → ∀W2. ⦃G, L0⦄ ⊢ W0 ➡ W2 → ∀T2. ⦃G, L0.ⓛW0⦄ ⊢ T0 ➡ T2 →
+ ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
+ ∃∃T. ⦃G, L1⦄ ⊢ ⓐV1.T1 ➡ T & ⦃G, L2⦄ ⊢ ⓓ{a}ⓝW2.V2.T2 ➡ T.
+#a #G #L0 #V0 #W0 #T0 #IH #V1 #HV01 #X #H
+#V2 #HV02 #W2 #HW02 #T2 #HT02 #L1 #HL01 #L2 #HL02
+elim (cpr_inv_abst1 … H) -H #W1 #T1 #HW01 #HT01 #H destruct
+elim (IH … HV01 … HV02 … HL01 … HL02) -HV01 -HV02 /2 width=1 by/ #V #HV1 #HV2
+elim (IH … HW01 … HW02 … HL01 … HL02) /2 width=1 by/ #W #HW1 #HW2
+elim (IH … HT01 … HT02 (L1.ⓛW1) … (L2.ⓛW2)) /2 width=1 by lpr_pair/ -L0 -V0 -W0 -T0 #T #HT1 #HT2
+lapply (lsubr_cpr_trans … HT2 (L2.ⓓⓝW2.V2) ?) -HT2 /2 width=1 by lsubr_beta/ (**) (* full auto not tried *)
+/4 width=5 by cpr_bind, cpr_flat, cpr_beta, ex2_intro/
+qed-.
+
+(* Basic-1: includes:
+ pr0_cong_upsilon_refl pr0_cong_upsilon_zeta
+ pr0_cong_upsilon_cong pr0_cong_upsilon_delta
+*)
+fact cpr_conf_lpr_flat_theta:
+ ∀a,G,L0,V0,W0,T0. (
+ ∀L,T. ⦃G, L0, ⓐV0.ⓓ{a}W0.T0⦄ ⊐+ ⦃G, L, T⦄ →
+ ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
+ ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
+ ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
+ ) →
+ ∀V1. ⦃G, L0⦄ ⊢ V0 ➡ V1 → ∀T1. ⦃G, L0⦄ ⊢ ⓓ{a}W0.T0 ➡ T1 →
+ ∀V2. ⦃G, L0⦄ ⊢ V0 ➡ V2 → ∀U2. ⬆[O, 1] V2 ≡ U2 →
+ ∀W2. ⦃G, L0⦄ ⊢ W0 ➡ W2 → ∀T2. ⦃G, L0.ⓓW0⦄ ⊢ T0 ➡ T2 →
+ ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
+ ∃∃T. ⦃G, L1⦄ ⊢ ⓐV1.T1 ➡ T & ⦃G, L2⦄ ⊢ ⓓ{a}W2.ⓐU2.T2 ➡ T.
+#a #G #L0 #V0 #W0 #T0 #IH #V1 #HV01 #X #H
+#V2 #HV02 #U2 #HVU2 #W2 #HW02 #T2 #HT02 #L1 #HL01 #L2 #HL02
+elim (IH … HV01 … HV02 … HL01 … HL02) -HV01 -HV02 /2 width=1 by/ #V #HV1 #HV2
+elim (lift_total V 0 1) #U #HVU
+lapply (cpr_lift … HV2 (L2.ⓓW2) … HVU2 … HVU) -HVU2 /2 width=2 by drop_drop/ #HU2
+elim (cpr_inv_abbr1 … H) -H *
+[ #W1 #T1 #HW01 #HT01 #H destruct
+ elim (IH … HW01 … HW02 … HL01 … HL02) /2 width=1 by/
+ elim (IH … HT01 … HT02 (L1.ⓓW1) … (L2.ⓓW2)) /2 width=1 by lpr_pair/ -L0 -V0 -W0 -T0
+ /4 width=7 by cpr_bind, cpr_flat, cpr_theta, ex2_intro/
+| #T1 #HT01 #HXT1 #H destruct
+ elim (IH … HT01 … HT02 (L1.ⓓW2) … (L2.ⓓW2)) /2 width=1 by lpr_pair/ -L0 -V0 -W0 -T0 #T #HT1 #HT2
+ elim (cpr_inv_lift1 … HT1 L1 … HXT1) -HXT1
+ /4 width=9 by cpr_flat, cpr_zeta, drop_drop, lift_flat, ex2_intro/
+]
+qed-.
+
+fact cpr_conf_lpr_beta_beta:
+ ∀a,G,L0,V0,W0,T0. (
+ ∀L,T. ⦃G, L0, ⓐV0.ⓛ{a}W0.T0⦄ ⊐+ ⦃G, L, T⦄ →
+ ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
+ ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
+ ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
+ ) →
+ ∀V1. ⦃G, L0⦄ ⊢ V0 ➡ V1 → ∀W1. ⦃G, L0⦄ ⊢ W0 ➡ W1 → ∀T1. ⦃G, L0.ⓛW0⦄ ⊢ T0 ➡ T1 →
+ ∀V2. ⦃G, L0⦄ ⊢ V0 ➡ V2 → ∀W2. ⦃G, L0⦄ ⊢ W0 ➡ W2 → ∀T2. ⦃G, L0.ⓛW0⦄ ⊢ T0 ➡ T2 →
+ ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
+ ∃∃T. ⦃G, L1⦄ ⊢ ⓓ{a}ⓝW1.V1.T1 ➡ T & ⦃G, L2⦄ ⊢ ⓓ{a}ⓝW2.V2.T2 ➡ T.
+#a #G #L0 #V0 #W0 #T0 #IH #V1 #HV01 #W1 #HW01 #T1 #HT01
+#V2 #HV02 #W2 #HW02 #T2 #HT02 #L1 #HL01 #L2 #HL02
+elim (IH … HV01 … HV02 … HL01 … HL02) -HV01 -HV02 /2 width=1 by/ #V #HV1 #HV2
+elim (IH … HW01 … HW02 … HL01 … HL02) /2 width=1 by/ #W #HW1 #HW2
+elim (IH … HT01 … HT02 (L1.ⓛW1) … (L2.ⓛW2)) /2 width=1 by lpr_pair/ -L0 -V0 -W0 -T0 #T #HT1 #HT2
+lapply (lsubr_cpr_trans … HT1 (L1.ⓓⓝW1.V1) ?) -HT1 /2 width=1 by lsubr_beta/
+lapply (lsubr_cpr_trans … HT2 (L2.ⓓⓝW2.V2) ?) -HT2 /2 width=1 by lsubr_beta/
+/4 width=5 by cpr_bind, cpr_flat, ex2_intro/ (**) (* full auto not tried *)
+qed-.
+
+(* Basic_1: was: pr0_upsilon_upsilon *)
+fact cpr_conf_lpr_theta_theta:
+ ∀a,G,L0,V0,W0,T0. (
+ ∀L,T. ⦃G, L0, ⓐV0.ⓓ{a}W0.T0⦄ ⊐+ ⦃G, L, T⦄ →
+ ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
+ ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
+ ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
+ ) →
+ ∀V1. ⦃G, L0⦄ ⊢ V0 ➡ V1 → ∀U1. ⬆[O, 1] V1 ≡ U1 →
+ ∀W1. ⦃G, L0⦄ ⊢ W0 ➡ W1 → ∀T1. ⦃G, L0.ⓓW0⦄ ⊢ T0 ➡ T1 →
+ ∀V2. ⦃G, L0⦄ ⊢ V0 ➡ V2 → ∀U2. ⬆[O, 1] V2 ≡ U2 →
+ ∀W2. ⦃G, L0⦄ ⊢ W0 ➡ W2 → ∀T2. ⦃G, L0.ⓓW0⦄ ⊢ T0 ➡ T2 →
+ ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
+ ∃∃T. ⦃G, L1⦄ ⊢ ⓓ{a}W1.ⓐU1.T1 ➡ T & ⦃G, L2⦄ ⊢ ⓓ{a}W2.ⓐU2.T2 ➡ T.
+#a #G #L0 #V0 #W0 #T0 #IH #V1 #HV01 #U1 #HVU1 #W1 #HW01 #T1 #HT01
+#V2 #HV02 #U2 #HVU2 #W2 #HW02 #T2 #HT02 #L1 #HL01 #L2 #HL02
+elim (IH … HV01 … HV02 … HL01 … HL02) -HV01 -HV02 /2 width=1 by/ #V #HV1 #HV2
+elim (IH … HW01 … HW02 … HL01 … HL02) /2 width=1 by/
+elim (IH … HT01 … HT02 (L1.ⓓW1) … (L2.ⓓW2)) /2 width=1 by lpr_pair/ -L0 -V0 -W0 -T0
+elim (lift_total V 0 1) #U #HVU
+lapply (cpr_lift … HV1 (L1.ⓓW1) … HVU1 … HVU) -HVU1 /2 width=2 by drop_drop/
+lapply (cpr_lift … HV2 (L2.ⓓW2) … HVU2 … HVU) -HVU2 /2 width=2 by drop_drop/
+/4 width=7 by cpr_bind, cpr_flat, ex2_intro/ (**) (* full auto not tried *)
+qed-.
+
+theorem cpr_conf_lpr: ∀G. lpx_sn_confluent (cpr G) (cpr G).
+#G #L0 #T0 @(fqup_wf_ind_eq … G L0 T0) -G -L0 -T0 #G #L #T #IH #G0 #L0 * [| * ]
+[ #I0 #HG #HL #HT #T1 #H1 #T2 #H2 #L1 #HL01 #L2 #HL02 destruct
+ elim (cpr_inv_atom1 … H1) -H1
+ elim (cpr_inv_atom1 … H2) -H2
+ [ #H2 #H1 destruct
+ /2 width=1 by cpr_conf_lpr_atom_atom/
+ | * #K0 #V0 #V2 #i2 #HLK0 #HV02 #HVT2 #H2 #H1 destruct
+ /3 width=10 by cpr_conf_lpr_atom_delta/
+ | #H2 * #K0 #V0 #V1 #i1 #HLK0 #HV01 #HVT1 #H1 destruct
+ /4 width=10 by ex2_commute, cpr_conf_lpr_atom_delta/
+ | * #X #Y #V2 #z #H #HV02 #HVT2 #H2
+ * #K0 #V0 #V1 #i #HLK0 #HV01 #HVT1 #H1 destruct
+ /3 width=17 by cpr_conf_lpr_delta_delta/
+ ]
+| #a #I #V0 #T0 #HG #HL #HT #X1 #H1 #X2 #H2 #L1 #HL01 #L2 #HL02 destruct
+ elim (cpr_inv_bind1 … H1) -H1 *
+ [ #V1 #T1 #HV01 #HT01 #H1
+ | #T1 #HT01 #HXT1 #H11 #H12
+ ]
+ elim (cpr_inv_bind1 … H2) -H2 *
+ [1,3: #V2 #T2 #HV02 #HT02 #H2
+ |2,4: #T2 #HT02 #HXT2 #H21 #H22
+ ] destruct
+ [ /3 width=10 by cpr_conf_lpr_bind_bind/
+ | /4 width=11 by ex2_commute, cpr_conf_lpr_bind_zeta/
+ | /3 width=11 by cpr_conf_lpr_bind_zeta/
+ | /3 width=12 by cpr_conf_lpr_zeta_zeta/
+ ]
+| #I #V0 #T0 #HG #HL #HT #X1 #H1 #X2 #H2 #L1 #HL01 #L2 #HL02 destruct
+ elim (cpr_inv_flat1 … H1) -H1 *
+ [ #V1 #T1 #HV01 #HT01 #H1
+ | #HX1 #H1
+ | #a1 #V1 #Y1 #W1 #Z1 #T1 #HV01 #HYW1 #HZT1 #H11 #H12 #H13
+ | #a1 #V1 #U1 #Y1 #W1 #Z1 #T1 #HV01 #HVU1 #HYW1 #HZT1 #H11 #H12 #H13
+ ]
+ elim (cpr_inv_flat1 … H2) -H2 *
+ [1,5,9,13: #V2 #T2 #HV02 #HT02 #H2
+ |2,6,10,14: #HX2 #H2
+ |3,7,11,15: #a2 #V2 #Y2 #W2 #Z2 #T2 #HV02 #HYW2 #HZT2 #H21 #H22 #H23
+ |4,8,12,16: #a2 #V2 #U2 #Y2 #W2 #Z2 #T2 #HV02 #HVU2 #HYW2 #HZT2 #H21 #H22 #H23
+ ] destruct
+ [ /3 width=10 by cpr_conf_lpr_flat_flat/
+ | /4 width=8 by ex2_commute, cpr_conf_lpr_flat_eps/
+ | /4 width=12 by ex2_commute, cpr_conf_lpr_flat_beta/
+ | /4 width=14 by ex2_commute, cpr_conf_lpr_flat_theta/
+ | /3 width=8 by cpr_conf_lpr_flat_eps/
+ | /3 width=7 by cpr_conf_lpr_eps_eps/
+ | /3 width=12 by cpr_conf_lpr_flat_beta/
+ | /3 width=13 by cpr_conf_lpr_beta_beta/
+ | /3 width=14 by cpr_conf_lpr_flat_theta/
+ | /3 width=17 by cpr_conf_lpr_theta_theta/
+ ]
+]
+qed-.
+
+(* Basic_1: includes: pr0_confluence pr2_confluence *)
+theorem cpr_conf: ∀G,L. confluent … (cpr G L).
+/2 width=6 by cpr_conf_lpr/ qed-.
+
+(* Properties on context-sensitive parallel reduction for terms *************)
+
+lemma lpr_cpr_conf_dx: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡ T1 → ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 →
+ ∃∃T. ⦃G, L1⦄ ⊢ T0 ➡ T & ⦃G, L1⦄ ⊢ T1 ➡ T.
+#G #L0 #T0 #T1 #HT01 #L1 #HL01
+elim (cpr_conf_lpr … HT01 T0 … HL01 … HL01) /2 width=3 by ex2_intro/
+qed-.
+
+lemma lpr_cpr_conf_sn: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡ T1 → ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 →
+ ∃∃T. ⦃G, L1⦄ ⊢ T0 ➡ T & ⦃G, L0⦄ ⊢ T1 ➡ T.
+#G #L0 #T0 #T1 #HT01 #L1 #HL01
+elim (cpr_conf_lpr … HT01 T0 … L0 … HL01) /2 width=3 by ex2_intro/
+qed-.
+
+(* Main properties **********************************************************)
+
+theorem lpr_conf: ∀G. confluent … (lpr G).
+/3 width=6 by lpx_sn_conf, cpr_conf_lpr/
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/predsn_5.ma".
+include "basic_2/reduction/lpr.ma".
+include "basic_2/reduction/cpx.ma".
+
+(* SN EXTENDED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS ********************)
+
+definition lpx: ∀h. sd h → relation3 genv lenv lenv ≝
+ λh,o,G. lpx_sn (cpx h o G).
+
+interpretation "extended parallel reduction (local environment, sn variant)"
+ 'PRedSn h o G L1 L2 = (lpx h o G L1 L2).
+
+(* Basic inversion lemmas ***************************************************)
+
+lemma lpx_inv_atom1: ∀h,o,G,L2. ⦃G, ⋆⦄ ⊢ ➡[h, o] L2 → L2 = ⋆.
+/2 width=4 by lpx_sn_inv_atom1_aux/ qed-.
+
+lemma lpx_inv_pair1: ∀h,o,I,G,K1,V1,L2. ⦃G, K1.ⓑ{I}V1⦄ ⊢ ➡[h, o] L2 →
+ ∃∃K2,V2. ⦃G, K1⦄ ⊢ ➡[h, o] K2 & ⦃G, K1⦄ ⊢ V1 ➡[h, o] V2 &
+ L2 = K2. ⓑ{I} V2.
+/2 width=3 by lpx_sn_inv_pair1_aux/ qed-.
+
+lemma lpx_inv_atom2: ∀h,o,G,L1. ⦃G, L1⦄ ⊢ ➡[h, o] ⋆ → L1 = ⋆.
+/2 width=4 by lpx_sn_inv_atom2_aux/ qed-.
+
+lemma lpx_inv_pair2: ∀h,o,I,G,L1,K2,V2. ⦃G, L1⦄ ⊢ ➡[h, o] K2.ⓑ{I}V2 →
+ ∃∃K1,V1. ⦃G, K1⦄ ⊢ ➡[h, o] K2 & ⦃G, K1⦄ ⊢ V1 ➡[h, o] V2 &
+ L1 = K1. ⓑ{I} V1.
+/2 width=3 by lpx_sn_inv_pair2_aux/ qed-.
+
+lemma lpx_inv_pair: ∀h,o,I1,I2,G,L1,L2,V1,V2. ⦃G, L1.ⓑ{I1}V1⦄ ⊢ ➡[h, o] L2.ⓑ{I2}V2 →
+ ∧∧ ⦃G, L1⦄ ⊢ ➡[h, o] L2 & ⦃G, L1⦄ ⊢ V1 ➡[h, o] V2 & I1 = I2.
+/2 width=1 by lpx_sn_inv_pair/ qed-.
+
+(* Basic properties *********************************************************)
+
+lemma lpx_refl: ∀h,o,G,L. ⦃G, L⦄ ⊢ ➡[h, o] L.
+/2 width=1 by lpx_sn_refl/ qed.
+
+lemma lpx_pair: ∀h,o,I,G,K1,K2,V1,V2. ⦃G, K1⦄ ⊢ ➡[h, o] K2 → ⦃G, K1⦄ ⊢ V1 ➡[h, o] V2 →
+ ⦃G, K1.ⓑ{I}V1⦄ ⊢ ➡[h, o] K2.ⓑ{I}V2.
+/2 width=1 by lpx_sn_pair/ qed.
+
+lemma lpr_lpx: ∀h,o,G,L1,L2. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L1⦄ ⊢ ➡[h, o] L2.
+#h #o #G #L1 #L2 #H elim H -L1 -L2 /3 width=1 by lpx_pair, cpr_cpx/
+qed.
+
+(* Basic forward lemmas *****************************************************)
+
+lemma lpx_fwd_length: ∀h,o,G,L1,L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 → |L1| = |L2|.
+/2 width=2 by lpx_sn_fwd_length/ qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/static/aaa_lift.ma".
+include "basic_2/static/lsuba_aaa.ma".
+include "basic_2/reduction/lpx_drop.ma".
+
+(* SN EXTENDED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS ********************)
+
+(* Properties on atomic arity assignment for terms **************************)
+
+(* Note: lemma 500 *)
+lemma cpx_lpx_aaa_conf: ∀h,o,G,L1,T1,A. ⦃G, L1⦄ ⊢ T1 ⁝ A →
+ ∀T2. ⦃G, L1⦄ ⊢ T1 ➡[h, o] T2 →
+ ∀L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 → ⦃G, L2⦄ ⊢ T2 ⁝ A.
+#h #o #G #L1 #T1 #A #H elim H -G -L1 -T1 -A
+[ #o #L1 #s #X #H
+ elim (cpx_inv_sort1 … H) -H // * //
+| #I #G #L1 #K1 #V1 #B #i #HLK1 #_ #IHV1 #X #H #L2 #HL12
+ elim (cpx_inv_lref1 … H) -H
+ [ #H destruct
+ elim (lpx_drop_conf … HLK1 … HL12) -L1 #X #H #HLK2
+ elim (lpx_inv_pair1 … H) -H
+ #K2 #V2 #HK12 #HV12 #H destruct /3 width=6 by aaa_lref/
+ | * #J #Y #Z #V2 #H #HV12 #HV2
+ lapply (drop_mono … H … HLK1) -H #H destruct
+ elim (lpx_drop_conf … HLK1 … HL12) -L1 #Z #H #HLK2
+ elim (lpx_inv_pair1 … H) -H #K2 #V0 #HK12 #_ #H destruct
+ /3 width=8 by aaa_lift, drop_fwd_drop2/
+ ]
+| #a #G #L1 #V1 #T1 #B #A #_ #_ #IHV1 #IHT1 #X #H #L2 #HL12
+ elim (cpx_inv_abbr1 … H) -H *
+ [ #V2 #T2 #HV12 #HT12 #H destruct /4 width=2 by lpx_pair, aaa_abbr/
+ | #T2 #HT12 #HT2 #H destruct -IHV1
+ /4 width=8 by lpx_pair, aaa_inv_lift, drop_drop/
+ ]
+| #a #G #L1 #V1 #T1 #B #A #_ #_ #IHV1 #IHT1 #X #H #L2 #HL12
+ elim (cpx_inv_abst1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct
+ /4 width=1 by lpx_pair, aaa_abst/
+| #G #L1 #V1 #T1 #B #A #_ #_ #IHV1 #IHT1 #X #H #L2 #HL12
+ elim (cpx_inv_appl1 … H) -H *
+ [ #V2 #T2 #HV12 #HT12 #H destruct /3 width=3 by aaa_appl/
+ | #b #V2 #W1 #W2 #U1 #U2 #HV12 #HW12 #HU12 #H1 #H2 destruct
+ lapply (IHV1 … HV12 … HL12) -IHV1 -HV12 #HV2
+ lapply (IHT1 (ⓛ{b}W2.U2) … HL12) -IHT1 /2 width=1 by cpx_bind/ -L1 #H
+ elim (aaa_inv_abst … H) -H #B0 #A0 #HW1 #HU2 #H destruct
+ /5 width=6 by lsuba_aaa_trans, lsuba_beta, aaa_abbr, aaa_cast/
+ | #b #V #V2 #W1 #W2 #U1 #U2 #HV1 #HV2 #HW12 #HU12 #H1 #H2 destruct
+ lapply (aaa_lift G L2 … B … (L2.ⓓW2) … HV2) -HV2 /2 width=2 by drop_drop/ #HV2
+ lapply (IHT1 (ⓓ{b}W2.U2) … HL12) -IHT1 /2 width=1 by cpx_bind/ -L1 #H
+ elim (aaa_inv_abbr … H) -H /3 width=3 by aaa_abbr, aaa_appl/
+ ]
+| #G #L1 #V1 #T1 #A #_ #_ #IHV1 #IHT1 #X #H #L2 #HL12
+ elim (cpx_inv_cast1 … H) -H
+ [ * #V2 #T2 #HV12 #HT12 #H destruct /3 width=1 by aaa_cast/
+ | -IHV1 /2 width=1 by/
+ | -IHT1 /2 width=1 by/
+ ]
+]
+qed-.
+
+lemma cpx_aaa_conf: ∀h,o,G,L,T1,A. ⦃G, L⦄ ⊢ T1 ⁝ A → ∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → ⦃G, L⦄ ⊢ T2 ⁝ A.
+/2 width=7 by cpx_lpx_aaa_conf/ qed-.
+
+lemma lpx_aaa_conf: ∀h,o,G,L1,T,A. ⦃G, L1⦄ ⊢ T ⁝ A → ∀L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 → ⦃G, L2⦄ ⊢ T ⁝ A.
+/2 width=7 by cpx_lpx_aaa_conf/ qed-.
+
+lemma cpr_aaa_conf: ∀G,L,T1,A. ⦃G, L⦄ ⊢ T1 ⁝ A → ∀T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ⦃G, L⦄ ⊢ T2 ⁝ A.
+/3 width=5 by cpx_aaa_conf, cpr_cpx/ qed-.
+
+lemma lpr_aaa_conf: ∀G,L1,T,A. ⦃G, L1⦄ ⊢ T ⁝ A → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L2⦄ ⊢ T ⁝ A.
+/3 width=5 by lpx_aaa_conf, lpr_lpx/ qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/substitution/lpx_sn_drop.ma".
+include "basic_2/reduction/cpx_lift.ma".
+include "basic_2/reduction/lpx.ma".
+
+(* SN EXTENDED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS ********************)
+
+(* Properties on local environment slicing ***********************************)
+
+lemma lpx_drop_conf: ∀h,o,G. dropable_sn (lpx h o G).
+/3 width=6 by lpx_sn_deliftable_dropable, cpx_inv_lift1/ qed-.
+
+lemma drop_lpx_trans: ∀h,o,G. dedropable_sn (lpx h o G).
+/3 width=10 by lpx_sn_liftable_dedropable, cpx_lift/ qed-.
+
+lemma lpx_drop_trans_O1: ∀h,o,G. dropable_dx (lpx h o G).
+/2 width=3 by lpx_sn_dropable/ qed-.
+
+(* Properties on supclosure *************************************************)
+
+lemma fqu_lpx_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+ ∀K2. ⦃G2, L2⦄ ⊢ ➡[h, o] K2 →
+ ∃∃K1,T. ⦃G1, L1⦄ ⊢ ➡[h, o] K1 & ⦃G1, L1⦄ ⊢ T1 ➡[h, o] T & ⦃G1, K1, T⦄ ⊐ ⦃G2, K2, T2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
+/3 width=5 by fqu_lref_O, fqu_pair_sn, fqu_flat_dx, lpx_pair, ex3_2_intro/
+[ #a #I #G2 #L2 #V2 #T2 #X #H elim (lpx_inv_pair1 … H) -H
+ #K2 #W2 #HLK2 #HVW2 #H destruct
+ /3 width=5 by cpx_pair_sn, fqu_bind_dx, ex3_2_intro/
+| #G #L1 #K1 #T1 #U1 #k #HLK1 #HTU1 #K2 #HK12
+ elim (drop_lpx_trans … HLK1 … HK12) -HK12
+ /3 width=7 by fqu_drop, ex3_2_intro/
+]
+qed-.
+
+lemma fquq_lpx_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
+ ∀K2. ⦃G2, L2⦄ ⊢ ➡[h, o] K2 →
+ ∃∃K1,T. ⦃G1, L1⦄ ⊢ ➡[h, o] K1 & ⦃G1, L1⦄ ⊢ T1 ➡[h, o] T & ⦃G1, K1, T⦄ ⊐⸮ ⦃G2, K2, T2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #K2 #HLK2 elim (fquq_inv_gen … H) -H
+[ #HT12 elim (fqu_lpx_trans … HT12 … HLK2) /3 width=5 by fqu_fquq, ex3_2_intro/
+| * #H1 #H2 #H3 destruct /2 width=5 by ex3_2_intro/
+]
+qed-.
+
+lemma lpx_fqu_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+ ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 →
+ ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡[h, o] T & ⦃G1, K1, T⦄ ⊐ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
+/3 width=7 by fqu_pair_sn, fqu_bind_dx, fqu_flat_dx, lpx_pair, ex3_2_intro/
+[ #I #G1 #L1 #V1 #X #H elim (lpx_inv_pair2 … H) -H
+ #K1 #W1 #HKL1 #HWV1 #H destruct elim (lift_total V1 0 1)
+ /4 width=7 by cpx_delta, fqu_drop, drop_drop, ex3_2_intro/
+| #G #L1 #K1 #T1 #U1 #k #HLK1 #HTU1 #L0 #HL01
+ elim (lpx_drop_trans_O1 … HL01 … HLK1) -L1
+ /3 width=5 by fqu_drop, ex3_2_intro/
+]
+qed-.
+
+lemma lpx_fquq_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
+ ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 →
+ ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡[h, o] T & ⦃G1, K1, T⦄ ⊐⸮ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #K1 #HKL1 elim (fquq_inv_gen … H) -H
+[ #HT12 elim (lpx_fqu_trans … HT12 … HKL1) /3 width=5 by fqu_fquq, ex3_2_intro/
+| * #H1 #H2 #H3 destruct /2 width=5 by ex3_2_intro/
+]
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/multiple/frees_lreq.ma".
+include "basic_2/multiple/frees_lift.ma".
+include "basic_2/reduction/lpx_drop.ma".
+
+(* SN EXTENDED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS ********************)
+
+(* Properties on context-sensitive free variables ***************************)
+
+lemma lpx_cpx_frees_trans: ∀h,o,G,L1,U1,U2. ⦃G, L1⦄ ⊢ U1 ➡[h, o] U2 →
+ ∀L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 →
+ ∀i. L2 ⊢ i ϵ 𝐅*[0]⦃U2⦄ → L1 ⊢ i ϵ 𝐅*[0]⦃U1⦄.
+#h #o #G #L1 #U1 @(fqup_wf_ind_eq … G L1 U1) -G -L1 -U1
+#G0 #L0 #U0 #IH #G #L1 * *
+[ -IH #s #HG #HL #HU #U2 #H1 #L2 #_ #i #H2 elim (cpx_inv_sort1 … H1) -H1
+ [| * #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 (lpx_drop_trans_O1 … HL12 … HLK2) -HL12 #Y #HLK1 #H
+ elim (lpx_inv_pair2 … H) -H #K1 #W1 #HK12 #HW12 #H destruct
+ /4 width=11 by frees_lref_be, fqup_lref/
+ | * #I #K1 #W1 #W0 #HLK1 #HW10 #HW0U2
+ lapply (drop_fwd_drop2 … HLK1) #H0
+ elim (lpx_drop_conf … H0 … HL12) -H0 -HL12 #K2 #HK12 #HLK2
+ elim (ylt_split i (j+1)) >yplus_SO2 #Hji
+ [ -IH elim (frees_inv_lift_be … H2 … HLK2 … HW0U2) /2 width=1 by ylt_fwd_succ2/
+ | lapply (frees_inv_lift_ge … H2 … HLK2 … HW0U2 ?) -L2 -U2 // destruct
+ /4 width=11 by frees_lref_be, fqup_lref, yle_succ1_inj/
+ ]
+ ]
+| -IH #p #HG #HL #HU #U2 #H1 >(cpx_inv_gref1 … H1) -H1 destruct
+ #L2 #_ #i #H2 elim (frees_inv_gref … H2)
+| #a #I #W1 #U1 #HG #HL #HU #X #HX #L2 #HL12 #i #Hi destruct
+ elim (cpx_inv_bind1 … HX) -HX *
+ [ #W2 #U2 #HW12 #HU12 #H destruct
+ elim (frees_inv_bind_O … Hi) -Hi
+ /4 width=7 by frees_bind_dx_O, frees_bind_sn, lpx_pair/
+ | #U2 #HU12 #HXU2 #H1 #H2 destruct
+ lapply (frees_lift_ge … Hi (L2.ⓓW1) (Ⓕ) … HXU2 ?)
+ /4 width=7 by frees_bind_dx_O, lpx_pair, drop_drop/
+ ]
+| #I #W1 #X1 #HG #HL #HU #X2 #HX2 #L2 #HL12 #i #Hi destruct
+ elim (cpx_inv_flat1 … HX2) -HX2 *
+ [ #W2 #U2 #HW12 #HU12 #H destruct
+ elim (frees_inv_flat … Hi) -Hi /3 width=7 by frees_flat_dx, frees_flat_sn/
+ | #HU12 #H destruct /3 width=7 by frees_flat_dx/
+ | #HW12 #H destruct /3 width=7 by frees_flat_sn/
+ | #b #W2 #V1 #V2 #U1 #U2 #HW12 #HV12 #HU12 #H1 #H2 #H3 destruct
+ 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 (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/
+ ]
+ | #b #W2 #W0 #V1 #V2 #U1 #U2 #HW12 #HW20 #HV12 #HU12 #H1 #H2 #H3 destruct
+ elim (frees_inv_bind_O … Hi) -Hi #Hi
+ [ /4 width=7 by frees_flat_dx, frees_bind_sn/
+ | elim (frees_inv_flat … Hi) -Hi
+ [ #HW0 lapply (frees_inv_lift_ge … HW0 L2 (Ⓕ) … HW20 ?) -W0
+ /3 width=7 by frees_flat_sn, drop_drop/
+ | /5 width=7 by frees_bind_dx_O, frees_flat_dx, lpx_pair/
+ ]
+ ]
+ ]
+]
+qed-.
+
+lemma cpx_frees_trans: ∀h,o,G. frees_trans (cpx h o G).
+/2 width=8 by lpx_cpx_frees_trans/ qed-.
+
+lemma lpx_frees_trans: ∀h,o,G,L1,L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 →
+ ∀U,i. L2 ⊢ i ϵ 𝐅*[0]⦃U⦄ → L1 ⊢ i ϵ 𝐅*[0]⦃U⦄.
+/2 width=8 by lpx_cpx_frees_trans/ 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/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_lreq.ma".
+include "basic_2/multiple/lleq_llor.ma".
+include "basic_2/reduction/cpx_lreq.ma".
+include "basic_2/reduction/cpx_lleq.ma".
+include "basic_2/reduction/lpx_frees.ma".
+
+(* SN EXTENDED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS ********************)
+
+(* Properties on lazy equivalence for local environments ********************)
+
+(* Note: contains a proof of llpx_cpx_conf *)
+lemma lleq_lpx_trans: ∀h,o,G,L2,K2. ⦃G, L2⦄ ⊢ ➡[h, o] K2 →
+ ∀L1,T,l. L1 ≡[T, l] L2 →
+ ∃∃K1. ⦃G, L1⦄ ⊢ ➡[h, o] K1 & K1 ≡[T, l] K2.
+#h #o #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 … 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 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=5 by llpx_sn_llor_fwd_sn, ex2_intro/
+]
+qed-.
+
+lemma lpx_lleq_fqu_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+ ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 → K1 ≡[T1, 0] L1 →
+ ∃∃K2. ⦃G1, K1, T1⦄ ⊐ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2 & K2 ≡[T2, 0] L2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
+[ #I #G1 #L1 #V1 #X #H1 #H2 elim (lpx_inv_pair2 … H1) -H1
+ #K0 #V0 #H1KL1 #_ #H destruct
+ elim (lleq_inv_lref_ge_dx … H2 ? I L1 V1) -H2 //
+ #K1 #H #H2KL1 lapply (drop_inv_O2 … H) -H #H destruct
+ /2 width=4 by fqu_lref_O, ex3_intro/
+| * [ #a ] #I #G1 #L1 #V1 #T1 #K1 #HLK1 #H
+ [ elim (lleq_inv_bind … H)
+ | elim (lleq_inv_flat … H)
+ ] -H /2 width=4 by fqu_pair_sn, ex3_intro/
+| #a #I #G1 #L1 #V1 #T1 #K1 #HLK1 #H elim (lleq_inv_bind_O … H) -H
+ /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 #k #HL1 #HTU1 #K1 #H1KL1 #H2KL1
+ elim (drop_O1_le (Ⓕ) (k+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
+ /3 width=4 by fqu_drop, ex3_intro/
+ | lapply (drop_fwd_length_le2 … HL1) -L -T1 -o
+ lapply (lleq_fwd_length … H2KL1) //
+ ]
+]
+qed-.
+
+lemma lpx_lleq_fquq_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
+ ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 → K1 ≡[T1, 0] L1 →
+ ∃∃K2. ⦃G1, K1, T1⦄ ⊐⸮ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2 & K2 ≡[T2, 0] L2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #K1 #H1KL1 #H2KL1
+elim (fquq_inv_gen … H) -H
+[ #H elim (lpx_lleq_fqu_trans … H … H1KL1 H2KL1) -L1
+ /3 width=4 by fqu_fquq, ex3_intro/
+| * #HG #HL #HT destruct /2 width=4 by ex3_intro/
+]
+qed-.
+
+lemma lpx_lleq_fqup_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
+ ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 → K1 ≡[T1, 0] L1 →
+ ∃∃K2. ⦃G1, K1, T1⦄ ⊐+ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2 & K2 ≡[T2, 0] L2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
+[ #G2 #L2 #T2 #H #K1 #H1KL1 #H2KL1 elim (lpx_lleq_fqu_trans … H … H1KL1 H2KL1) -L1
+ /3 width=4 by fqu_fqup, ex3_intro/
+| #G #G2 #L #L2 #T #T2 #_ #HT2 #IHT1 #K1 #H1KL1 #H2KL1 elim (IHT1 … H2KL1) // -L1
+ #K #HT1 #H1KL #H2KL elim (lpx_lleq_fqu_trans … HT2 … H1KL H2KL) -L
+ /3 width=5 by fqup_strap1, ex3_intro/
+]
+qed-.
+
+lemma lpx_lleq_fqus_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
+ ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 → K1 ≡[T1, 0] L1 →
+ ∃∃K2. ⦃G1, K1, T1⦄ ⊐* ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2 & K2 ≡[T2, 0] L2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #K1 #H1KL1 #H2KL1
+elim (fqus_inv_gen … H) -H
+[ #H elim (lpx_lleq_fqup_trans … H … H1KL1 H2KL1) -L1
+ /3 width=4 by fqup_fqus, ex3_intro/
+| * #HG #HL #HT destruct /2 width=4 by ex3_intro/
+]
+qed-.
+
+fact lreq_lpx_trans_lleq_aux: ∀h,o,G,L1,L0,l,k. L1 ⩬[l, k] L0 → k = ∞ →
+ ∀L2. ⦃G, L0⦄ ⊢ ➡[h, o] L2 →
+ ∃∃L. L ⩬[l, k] L2 & ⦃G, L1⦄ ⊢ ➡[h, o] L &
+ (∀T. L0 ≡[T, l] L2 ↔ L1 ≡[T, l] L).
+#h #o #G #L1 #L0 #l #k #H elim H -L1 -L0 -l -k
+[ #l #k #_ #L2 #H >(lpx_inv_atom1 … H) -H
+ /3 width=5 by ex3_intro, conj/
+| #I1 #I0 #L1 #L0 #V1 #V0 #_ #_ #Hm destruct
+| #I #L1 #L0 #V1 #k #HL10 #IHL10 #Hm #Y #H
+ elim (lpx_inv_pair1 … H) -H #L2 #V2 #HL02 #HV02 #H destruct
+ 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, lreq_cpx_trans, lreq_pair/
+ #T elim (IH T) #HL0dx #HL0sn
+ @conj #H @(lleq_lreq_repl … H) -H /3 width=1 by lreq_sym, lreq_pair_O_Y/
+| #I1 #I0 #L1 #L0 #V1 #V0 #l #k #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, lreq_succ/
+ #T elim (IH T) #HL0dx #HL0sn
+ @conj #H @(lleq_lreq_repl … H) -H /3 width=1 by lreq_sym, lreq_succ/
+]
+qed-.
+
+lemma lreq_lpx_trans_lleq: ∀h,o,G,L1,L0,l. L1 ⩬[l, ∞] L0 →
+ ∀L2. ⦃G, L0⦄ ⊢ ➡[h, o] L2 →
+ ∃∃L. L ⩬[l, ∞] L2 & ⦃G, L1⦄ ⊢ ➡[h, o] L &
+ (∀T. L0 ≡[T, l] L2 ↔ L1 ≡[T, l] L).
+/2 width=1 by lreq_lpx_trans_lleq_aux/ qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/statictypestar_6.ma".
-include "basic_2/grammar/genv.ma".
-include "basic_2/substitution/drop.ma".
-include "basic_2/static/sh.ma".
-
-(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************)
-
-(* activate genv *)
-inductive lstas (h): nat → relation4 genv lenv term term ≝
-| 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,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 d T U = (lstas h d G L T U).
-
-(* Basic inversion lemmas ***************************************************)
-
-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 #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,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,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 & d = 0
- ) |
- (∃∃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 #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 #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,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 & d = 0
- ) |
- (∃∃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-.
-
-lemma lstas_inv_lref1_O: ∀h,G,L,X,i. ⦃G, L⦄ ⊢ #i •*[h, 0] X →
- (∃∃K,V,W. ⬇[i] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V •*[h, 0] W &
- ⬆[0, i+1] W ≡ X
- ) ∨
- (∃∃K,W,V. ⬇[i] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W •*[h, 0] V &
- 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 #d #_ #_ #_ <plus_n_Sm #H destruct
-qed-.
-
-(* Basic_1: was just: sty0_gen_lref *)
-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, d] V &
- ⬆[0, i+1] V ≡ X
- ).
-#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,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 #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,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,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 #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,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,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 #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,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,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 #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,d. ⦃G, L⦄ ⊢ ⓝW.T •*[h, d] U → ⦃G, L⦄ ⊢ T •*[h, d] U.
-/2 width=4 by lstas_inv_cast1_aux/
-qed-.
-
-(* Basic_1: removed theorems 7:
- sty1_abbr sty1_appl sty1_bind sty1_cast2
- sty1_correct sty1_lift sty1_trans
-*)
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/static/aaa_lift.ma".
-include "basic_2/unfold/lstas_lstas.ma".
-
-(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************)
-
-(* Properties on atomic arity assignment for terms **************************)
-
-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 #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 … d) -d
- [ elim (IHV 0) -IHV /3 width=7 by lstas_zero, aaa_lref, ex2_intro/
- | #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 #d elim (IHT d) -IHT
- /3 width=7 by lstas_bind, aaa_abbr, ex2_intro/
-| #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 #d elim (IHT d) -IHT
- /3 width=6 by lstas_appl, aaa_appl, ex2_intro/
-| #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,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-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/static/da_da.ma".
-include "basic_2/unfold/lstas_lstas.ma".
-
-(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************)
-
-(* Properties on degree assignment for terms ********************************)
-
-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 #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 #d1 #HLK #HW #IHW #d2 @(nat_ind_plus … d2) -d2
- [ elim (IHW 0) -IHW /3 width=6 by lstas_zero, da_ldec, ex2_intro/
- | #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 #d1 #_ #IHT #d2 elim (IHT … d2) -IHT
- /3 width=6 by lstas_bind, da_bind, ex2_intro/
-| * #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,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,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 #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 #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,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 #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 #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 #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,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/
-| -d1 <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 "ground_2/ynat/ynat_max.ma".
-include "basic_2/substitution/drop_drop.ma".
-include "basic_2/unfold/lstas.ma".
-
-(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************)
-
-(* Properties on relocation *************************************************)
-
-(* Basic_1: was just: sty0_lift *)
-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 #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 /2 width=1 by ylt_fwd_le_succ1/ #W2 #HW12 #HWU2
- elim (drop_trans_le … HL21 … HLK1) -L1 /2 width=2 by ylt_fwd_le/ #X #HLK2 #H
- elim (drop_inv_skip2 … H) -H /2 width=1 by ylt_to_minus/ -Hil #K2 #V2 #HK21 #HV12 #H destruct
- /3 width=9 by lstas_ldef/
- | lapply (lift_trans_be … HW1 … HWU2 ? ?) -W /2 width=1 by yle_succ_dx/ #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 #l #m #HL21 #X #H #U2 #HWU2
- >(lift_mono … HWU2 … H) -U2
- 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 ylt_fwd_le/ #X #HLK2 #H
- elim (drop_inv_skip2 … H) -H /2 width=1 by ylt_to_minus/ -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 #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 /2 width=1 by ylt_fwd_le_succ1/ #W #HW1 #HWU2
- elim (drop_trans_le … HL21 … HLK1) -L1 /2 width=2 by ylt_fwd_le/ #X #HLK2 #H
- elim (drop_inv_skip2 … H) -H /2 width=1 by ylt_to_minus/ -Hil #K2 #W2 #HK21 #HW12 #H destruct
- /3 width=9 by lstas_succ/
- | lapply (lift_trans_be … HW1 … HWU2 ? ?) -W /2 width=1 by yle_succ_dx/ #HW1U2
- lapply (drop_trans_ge … HL21 … HLK1 ?) -L1 /3 width=9 by lstas_succ, drop_inv_gen/
- ]
-| #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 #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 #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,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 #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 // <yminus_succ2 <yplus_inj >yplus_SO2 >ymax_pre_sn /3 width=8 by lstas_ldef, ylt_fwd_le_succ1, ex2_intro/
- | lapply (drop_conf_ge … HL21 … HLK2 ?) -L2 // #HL1K2
- elim (yle_inv_plus_inj2 … Hil) -Hil #Hlim #mi
- elim (lift_split … HW2 l (i-m+1)) -HW2 /2 width=1 by yle_succ_dx, le_S_S/
- #W0 #HW20 <le_plus_minus_comm /2 width=1 by yle_inv_inj/ >minus_minus_m_m /3 width=8 by lstas_ldef, yle_inv_inj, le_S, ex2_intro/
- ]
-| #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 #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 // <yminus_succ2 <yplus_inj >yplus_SO2 >ymax_pre_sn /3 width=8 by lstas_succ, ylt_fwd_le_succ1, ex2_intro/
- | lapply (drop_conf_ge … HL21 … HLK2 ?) -L2 // #HL1K2
- elim (yle_inv_plus_inj2 … Hil) -Hil #Hlim #mi
- elim (lift_split … HW2 l (i-m+1)) -HW2 /2 width=1 by yle_succ_dx, le_S_S/
- #W0 #HW20 <le_plus_minus_comm /2 width=1 by yle_inv_inj/ >minus_minus_m_m /3 width=8 by lstas_succ, yle_inv_inj, le_S, ex2_intro/
- ]
-| #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 #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 #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/
-]
-qed-.
-
-(* Advanced inversion lemmas ************************************************)
-
-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 #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 #_ #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 #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/
- | #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 #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,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,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 #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 #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 #d #HLK #_ #_ #IHW1 #d2
- @(nat_ind_plus … d2) -d2
- [ elim (IHW1 0) -IHW1 /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 ex_intro, lstas_succ/
- ]
-| #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-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/multiple/llpx_sn_drop.ma".
-include "basic_2/unfold/lstas.ma".
-
-(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************)
-
-(* Properties on lazy sn pointwise extensions *******************************)
-
-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 #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 #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 #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 #d #_ #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
- /3 width=1 by llpx_sn_flat/
-| #G #Ls #V #T1 #T2 #d #_ #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
- /3 width=1 by llpx_sn_flat/
-]
-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/unfold/lstas_lift.ma".
-
-(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************)
-
-(* Main properties **********************************************************)
-
-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 #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 #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 #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 #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,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: #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/
-| #G #L #K #W #W1 #i #HLK #_ #_ #X #H
- 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 #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 #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 #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 #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,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 #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 #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 #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 #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,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,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/
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/unfold_4.ma".
-include "basic_2/grammar/lenv_append.ma".
-include "basic_2/grammar/genv.ma".
-include "basic_2/substitution/drop.ma".
-
-(* CONTEXT-SENSITIVE UNFOLD FOR TERMS ***************************************)
-
-(* activate genv *)
-inductive unfold: relation4 genv lenv term lenv ≝
-| unfold_sort: ∀G,L,k. unfold G L (⋆k) L
-| unfold_lref: ∀I,G,L1,L2,K1,K2,V,i. ⬇[i] L1 ≡ K1. ⓑ{I}V →
- unfold G K1 V K2 → ⬇[Ⓣ, |L2|, i] L2 ≡ K2 →
- unfold G L1 (#i) (L1@@L2)
-| unfold_bind: ∀a,I,G,L1,L2,V,T.
- unfold G (L1.ⓑ{I}V) T L2 → unfold G L1 (ⓑ{a,I}V.T) L2
-| unfold_flat: ∀I,G,L1,L2,V,T.
- unfold G L1 T L2 → unfold G L1 (ⓕ{I}V.T) L2
-.
-
-interpretation "context-sensitive unfold (term)"
- 'Unfold G L1 T L2 = (unfold G L1 T L2).