pc1/props pc1_t
pc1/props pc1_pr0_u2
pc1/props pc1_head
+
pc3/dec pc3_dec
pc3/dec pc3_abst_dec
pc3/fsubst0 pc3_pr2_fsubst0
pc3/fsubst0 pc3_pr2_fsubst0_back
pc3/fsubst0 pc3_fsubst0
-pc3/fwd pc3_gen_sort
-pc3/fwd pc3_gen_abst
-pc3/fwd pc3_gen_abst_shift
-pc3/fwd pc3_gen_lift
pc3/fwd pc3_gen_not_abst
pc3/fwd pc3_gen_lift_abst
-pc3/fwd pc3_gen_sort_abst
-pc3/left pc3_ind_left__pc3_left_pr3
-pc3/left pc3_ind_left__pc3_left_trans
-pc3/left pc3_ind_left__pc3_left_sym
-pc3/left pc3_ind_left__pc3_left_pc3
-pc3/left pc3_ind_left__pc3_pc3_left
-pc3/left pc3_ind_left
pc3/nf2 pc3_nf2
pc3/nf2 pc3_nf2_unfold
pc3/pc1 pc3_pc1
-pc3/props pc3_pr2_r
-pc3/props pc3_pr2_x
-pc3/props pc3_pr3_r
-pc3/props pc3_pr3_x
-pc3/props pc3_pr3_t
-pc3/props pc3_s
-pc3/props pc3_thin_dx
-pc3/props pc3_head_1
-pc3/props pc3_head_2
-pc3/props pc3_pr2_u
-pc3/props pc3_t
-pc3/props pc3_pr2_u2
-pc3/props pc3_pr3_conf
-pc3/props pc3_head_12
-pc3/props pc3_head_21
pc3/props pc3_pr0_pr2_t
pc3/props pc3_pr2_pr2_t
pc3/props pc3_pr2_pr3_t
pc3/props pc3_pr3_pc3_t
-pc3/props pc3_lift
pc3/props pc3_eta
pc3/subst1 pc3_gen_cabbr
pc3/wcpr0 pc3_wcpr0__pc3_wcpr0_t_aux
pc3/wcpr0 pc3_wcpr0_t
pc3/wcpr0 pc3_wcpr0
+
pr0/fwd pr0_gen_void
pr0/dec nf0_dec
pr0/subst1 pr0_subst1_back
(* Advanced properties ******************************************************)
-lemma cprs_abbr_dx: ∀L,V1,V2. L ⊢ V1 ➡ V2 → ∀T1,T2. L. ⓓV1 ⊢ T1 ➡* T2 →
- L ⊢ ⓓV1. T1 ➡* ⓓV2. T2.
+lemma cprs_abbr1_dx: ∀L,V1,V2. L ⊢ V1 ➡ V2 → ∀T1,T2. L. ⓓV1 ⊢ T1 ➡* T2 →
+ L ⊢ ⓓV1. T1 ➡* ⓓV2. T2.
#L #V1 #V2 #HV12 #T1 #T2 #HT12 @(cprs_ind_dx … HT12) -T1
[ /3 width=5/
| #T1 #T #HT1 #_ #IHT1
]
qed.
-lemma cpr_abbr: ∀L,V1,V2. L ⊢ V1 ➡ V2 → ∀T1,T2. L. ⓓV1 ⊢ T1 ➡ T2 →
- L ⊢ ⓓV1. T1 ➡* ⓓV2. T2.
+lemma cpr_abbr1: ∀L,V1,V2. L ⊢ V1 ➡ V2 → ∀T1,T2. L. ⓓV1 ⊢ T1 ➡ T2 →
+ L ⊢ ⓓV1. T1 ➡* ⓓV2. T2.
/3 width=1/ qed.
lemma cpr_abbr2: ∀L,V1,V2. L ⊢ V1 ➡ V2 → ∀T1,T2. L. ⓓV2 ⊢ T1 ➡ T2 →
@(cprs_trans … IHV1) -IHV1 /2 width=1/
qed.
-lemma cprs_abbr: ∀L,V1,T1,T2. L. ⓓV1 ⊢ T1 ➡* T2 → ∀V2. L ⊢ V1 ➡* V2 →
- L ⊢ ⓓV1. T1 ➡* ⓓV2. T2.
+lemma cprs_abbr1: ∀L,V1,T1,T2. L. ⓓV1 ⊢ T1 ➡* T2 → ∀V2. L ⊢ V1 ➡* V2 →
+ L ⊢ ⓓV1. T1 ➡* ⓓV2. T2.
#L #V1 #T1 #T2 #HT12 #V2 #HV12 @(cprs_ind … HV12) -V2 /2 width=1/
#V #V2 #_ #HV2 #IHV1
@(cprs_trans … IHV1) -IHV1 /2 width=1/
qed.
+lemma cprs_abbr2_dx: ∀L,V1,V2. L ⊢ V1 ➡ V2 → ∀T1,T2. L. ⓓV2 ⊢ T1 ➡* T2 →
+ L ⊢ ⓓV1. T1 ➡* ⓓV2. T2.
+#L #V1 #V2 #HV12 #T1 #T2 #HT12 @(cprs_ind_dx … HT12) -T1
+[ /2 width=1/
+| #T1 #T #HT1 #_ #IHT1
+ lapply (lcpr_cpr_trans (L. ⓓV1) … HT1) -HT1 /2 width=1/ #HT1
+ @(cprs_trans … IHT1) -IHT1 /2 width=1/
+]
+qed.
+
+lemma cprs_abbr2: ∀L,V1,V2. L ⊢ V1 ➡* V2 → ∀T1,T2. L. ⓓV2 ⊢ T1 ➡* T2 →
+ L ⊢ ⓓV1. T1 ➡* ⓓV2. T2.
+#L #V1 #V2 #HV12 @(cprs_ind … HV12) -V2 /2 width=1/
+#V #V2 #_ #HV2 #IHV1 #T1 #T2 #HT12
+lapply (IHV1 T1 T1 ?) -IHV1 // #HV1
+@(cprs_trans … HV1) -HV1 /2 width=1/
+qed.
+
(* Basic_1: was only: pr3_pr2_pr3_t pr3_wcpr0_t *)
lemma lcpr_cprs_trans: ∀L1,L2. L1 ⊢ ➡ L2 →
∀T1,T2. L2 ⊢ T1 ➡* T2 → L1 ⊢ T1 ➡* T2.
-HTU1 #U #U2 #_ #HU2 * #T #HTU #HT1
elim (cpr_inv_lift … HLK … HTU … HU2) -U -HLK /3 width=5/
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/reducibility/lcpr_aaa.ma".
+include "basic_2/computation/cprs.ma".
+include "basic_2/computation/lcprs.ma".
+
+(* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS *************)
+
+(* Properties about atomic arity assignment on terms ************************)
+
+lemma aaa_lcprs_conf: ∀L1,T,A. L1 ⊢ T ÷ A → ∀L2. L1 ⊢ ➡* L2 → L2 ⊢ T ÷ A.
+#L1 #T #A #HT #L2 #HL12
+@(TC_Conf3 … (λL,A. L ⊢ T ÷ A) … HT ? HL12) /2 width=3/
+qed.
+
+lemma aaa_cprs_conf: ∀L,T1,A. L ⊢ T1 ÷ A → ∀T2. L ⊢ T1 ➡* T2 → L ⊢ T2 ÷ A.
+#L #T1 #A #HT1 #T2 #HT12
+@(TC_Conf3 … HT1 ? HT12) /2 width=3/
+qed.
+
+lemma aaa_lcprs_cprs_conf: ∀L1,T1,A. L1 ⊢ T1 ÷ A → ∀L2. L1 ⊢ ➡* L2 →
+ ∀T2. L2 ⊢ T1 ➡* T2 → L2 ⊢ T2 ÷ A.
+/3 width=3/ qed.
#L #T1 #R #HT1 #IHT1 #T2 #HT12 @(TC_star_ind … HT1 IHT1 … HT12) //
qed-.
+lemma cpcs_ind_dx: ∀L,T2. ∀R:predicate term. R T2 →
+ (∀T1,T. L ⊢ T1 ⬌ T → L ⊢ T ⬌* T2 → R T → R T1) →
+ ∀T1. L ⊢ T1 ⬌* T2 → R T1.
+#L #T2 #R #HT2 #IHT2 #T1 #HT12
+@(TC_star_ind_dx … HT2 IHT2 … HT12) //
+qed-.
+
(* Basic properties *********************************************************)
(* Basic_1: was: pc3_refl *)
lemma cpcs_refl: ∀L,T. L ⊢ T ⬌* T.
/2 width=1/ qed.
-lemma cpcs_strap1: ∀L,T1,T,T2.
- L ⊢ T1 ⬌* T → L ⊢ T ⬌ T2 → L ⊢ T1 ⬌* T2.
+lemma cpcs_strap1: ∀L,T1,T,T2. L ⊢ T1 ⬌* T → L ⊢ T ⬌ T2 → L ⊢ T1 ⬌* T2.
/2 width=3/ qed.
-lemma cpcs_strap2: ∀L,T1,T,T2.
- L ⊢ T1 ⬌ T → L ⊢ T ⬌* T2 → L ⊢ T1 ⬌* T2.
+lemma cpcs_strap2: ∀L,T1,T,T2. L ⊢ T1 ⬌ T → L ⊢ T ⬌* T2 → L ⊢ T1 ⬌* T2.
/2 width=3/ qed.
-(* Basic_1: removed theorems 1: clear_pc3_trans *)
+(* Basic_1: was: pc3_pr2_r *)
+lemma cpcs_cpr_dx: ∀L,T1,T2. L ⊢ T1 ➡ T2 → L ⊢ T1 ⬌* T2.
+/3 width=1/ qed.
+
+(* Basic_1: was: pc3_pr2_x *)
+lemma cpcs_cpr_sn: ∀L,T1,T2. L ⊢ T2 ➡ T1 → L ⊢ T1 ⬌* T2.
+/3 width=1/ qed.
+
+lemma cpcs_cpr_strap1: ∀L,T1,T. L ⊢ T1 ⬌* T → ∀T2. L ⊢ T ➡ T2 → L ⊢ T1 ⬌* T2.
+/3 width=3/ qed.
+
+(* Basic_1: was: pc3_pr2_u *)
+lemma cpcs_cpr_strap2: ∀L,T1,T. L ⊢ T1 ➡ T → ∀T2. L ⊢ T ⬌* T2 → L ⊢ T1 ⬌* T2.
+/3 width=3/ qed.
+
+lemma cpcs_cpr_div: ∀L,T1,T. L ⊢ T1 ⬌* T → ∀T2. L ⊢ T2 ➡ T → L ⊢ T1 ⬌* T2.
+/3 width=3/ qed.
+
+(* Basic_1: was: pc3_pr2_u2 *)
+lemma cpcs_cpr_conf: ∀L,T1,T. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ⬌* T2 → L ⊢ T1 ⬌* T2.
+/3 width=3/ qed.
+
+(* Basic_1: was: pc3_s *)
+lemma cprs_comm: ∀L,T1,T2. L ⊢ T1 ⬌* T2 → L ⊢ T2 ⬌* T1.
+#L #T1 #T2 #H @(cpcs_ind … H) -T2 // /3 width=3/
+qed.
+
+(* Basic_1: removed theorems 6:
+ clear_pc3_trans pc3_ind_left
+ pc3_head_1 pc3_head_2 pc3_head_12 pc3_head_21
+ Basic_1: removed local theorems 5:
+ pc3_left_pr3 pc3_left_trans pc3_left_sym pc3_left_pc3 pc3_pc3_left
+*)
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| 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_lift.ma".
+include "basic_2/computation/cprs_cprs.ma".
+include "basic_2/conversion/cpc_cpc.ma".
+include "basic_2/equivalence/cpcs_cprs.ma".
+
+(* CONTEXT-SENSITIVE PARALLEL EQUIVALENCE ON TERMS **************************)
+
+(* Advanced inversion lemmas ************************************************)
+
+lemma cpcs_inv_cprs: ∀L,T1,T2. L ⊢ T1 ⬌* T2 →
+ ∃∃T. L ⊢ T1 ➡* T & L ⊢ T2 ➡* T.
+#L #T1 #T2 #H @(cpcs_ind … H) -T2
+[ /3 width=3/
+| #T #T2 #_ #HT2 * #T0 #HT10 elim HT2 -HT2 #HT2 #HT0
+ [ elim (cprs_strip … HT0 … HT2) -T #T #HT0 #HT2
+ lapply (cprs_strap1 … HT10 … HT0) -T0 /2 width=3/
+ | lapply (cprs_strap2 … HT2 … HT0) -T /2 width=3/
+ ]
+]
+qed-.
+
+(* Basic_1: was: pc3_gen_sort *)
+lemma cpcs_inv_sort: ∀L,k1,k2. L ⊢ ⋆k1 ⬌* ⋆k2 → k1 = k2.
+#L #k1 #k2 #H
+elim (cpcs_inv_cprs … H) -H #T #H1
+>(cprs_inv_sort1 … H1) -T #H2
+lapply (cprs_inv_sort1 … H2) -L #H destruct //
+qed-.
+
+(* Basic_1: was: pc3_gen_sort_abst *)
+lemma cpcs_inv_sort_abst: ∀L,W,T,k. L ⊢ ⋆k ⬌* ⓛW.T → False.
+#L #W #T #k #H
+elim (cpcs_inv_cprs … H) -H #X #H1
+>(cprs_inv_sort1 … H1) -X #H2
+elim (cprs_inv_abst1 Abst W … H2) -H2 #W0 #T0 #_ #_ #H destruct
+qed-.
+
+(* Basic_1: was: pc3_gen_abst *)
+lemma cpcs_inv_abst: ∀L,W1,W2,T1,T2. L ⊢ ⓛW1.T1 ⬌* ⓛW2.T2 → ∀I,V.
+ L ⊢ W1 ⬌* W2 ∧ L. ②{I}V ⊢ T1 ⬌* T2.
+#L #W1 #W2 #T1 #T2 #H #I #V
+elim (cpcs_inv_cprs … H) -H #T #H1 #H2
+elim (cprs_inv_abst1 I V … H1) -H1 #W0 #T0 #HW10 #HT10 #H destruct
+elim (cprs_inv_abst1 I V … H2) -H2 #W #T #HW2 #HT2 #H destruct /3 width=3/
+qed-.
+
+(* Basic_1: was: pc3_gen_abst_shift *)
+lemma cpcs_inv_abst_shift: ∀L,W1,W2,T1,T2. L ⊢ ⓛW1.T1 ⬌* ⓛW2.T2 → ∀W.
+ L ⊢ W1 ⬌* W2 ∧ L. ⓛW ⊢ T1 ⬌* T2.
+#L #W1 #W2 #T1 #T2 #H #W
+lapply (cpcs_inv_abst … H Abst W) -H //
+qed.
+
+(* Basic_1: was: pc3_gen_lift *)
+lemma cpcs_inv_lift: ∀L,K,d,e. ⇩[d, e] L ≡ K →
+ ∀T1,U1. ⇧[d, e] T1 ≡ U1 → ∀T2,U2. ⇧[d, e] T2 ≡ U2 →
+ L ⊢ U1 ⬌* U2 → K ⊢ T1 ⬌* T2.
+#L #K #d #e #HLK #T1 #U1 #HTU1 #T2 #U2 #HTU2 #HU12
+elim (cpcs_inv_cprs … HU12) -HU12 #U #HU1 #HU2
+elim (cprs_inv_lift … HLK … HTU1 … HU1) -U1 #T #HTU #HT1
+elim (cprs_inv_lift … HLK … HTU2 … HU2) -L -U2 #X #HXU
+>(lift_inj … HXU … HTU) -X -U -d -e /2 width=3/
+qed-.
+
+(* advanced properties ******************************************************)
+
+(* Basic_1: was only: pc3_thin_dx *)
+lemma cpcs_flat: ∀L,V1,V2. L ⊢ V1 ⬌* V2 → ∀T1,T2. L ⊢ T1 ⬌* T2 →
+ ∀I. L ⊢ ⓕ{I}V1. T1 ⬌* ⓕ{I}V2. T2.
+#L #V1 #V2 #HV12 #T1 #T2 #HT12 #I
+elim (cpcs_inv_cprs … HV12) -HV12 #V #HV1 #HV2
+elim (cpcs_inv_cprs … HT12) -HT12 /3 width=5 by cprs_flat, cprs_div/ (**) (* /3 width=5/ is too slow *)
+qed.
+
+lemma cpcs_abbr_dx: ∀L,V,T1,T2. L.ⓓV ⊢ T1 ⬌* T2 → L ⊢ ⓓV. T1 ⬌* ⓓV. T2.
+#L #V #T1 #T2 #HT12
+elim (cpcs_inv_cprs … HT12) -HT12 /3 width=5 by cprs_div, cprs_abbr1/ (**) (* /3 width=5/ is a bit slow *)
+qed.
+
+lemma cpcs_abbr_sn: ∀L,V1,V2,T. L ⊢ V1 ⬌* V2 → L ⊢ ⓓV1. T ⬌* ⓓV2. T.
+#L #V1 #V2 #T #HV12
+elim (cpcs_inv_cprs … HV12) -HV12 /3 width=5 by cprs_div, cprs_abbr1/ (**) (* /3 width=5/ is a bit slow *)
+qed.
+
+(* Basic_1: was: pc3_gen_lift *)
+lemma cpcs_lift: ∀L,K,d,e. ⇩[d, e] L ≡ K →
+ ∀T1,U1. ⇧[d, e] T1 ≡ U1 → ∀T2,U2. ⇧[d, e] T2 ≡ U2 →
+ K ⊢ T1 ⬌* T2 → L ⊢ U1 ⬌* U2.
+#L #K #d #e #HLK #T1 #U1 #HTU1 #T2 #U2 #HTU2 #HT12
+elim (cpcs_inv_cprs … HT12) -HT12 #T #HT1 #HT2
+elim (lift_total T d e) #U #HTU
+lapply (cprs_lift … HLK … HTU1 … HT1 … HTU) -T1 #HU1
+lapply (cprs_lift … HLK … HTU2 … HT2 … HTU) -K -T2 -T -d -e /2 width=3/
+qed.
+
+lemma cpcs_strip: ∀L,T1,T. L ⊢ T ⬌* T1 → ∀T2. L ⊢ T ⬌ T2 →
+ ∃∃T0. L ⊢ T1 ⬌ T0 & L ⊢ T2 ⬌* T0.
+/3 width=3/ qed.
+
+(* Main properties **********************************************************)
+
+(* Basic_1: was pc3_t *)
+theorem cpcs_trans: ∀L,T1,T. L ⊢ T1 ⬌* T → ∀T2. L ⊢ T ⬌* T2 → L ⊢ T1 ⬌* T2.
+/2 width=3/ qed.
+
+theorem cpcs_canc_sn: ∀L,T,T1,T2. L ⊢ T ⬌* T1 → L ⊢ T ⬌* T2 → L ⊢ T1 ⬌* T2.
+/3 width=3 by cpcs_trans, cprs_comm/ qed. (**) (* /3 width=3/ is too slow *)
+
+theorem cpcs_canc_dx: ∀L,T,T1,T2. L ⊢ T1 ⬌* T → L ⊢ T2 ⬌* T → L ⊢ T1 ⬌* T2.
+/3 width=3 by cpcs_trans, cprs_comm/ qed. (**) (* /3 width=3/ is too slow *)
+
+lemma cpcs_abbr1: ∀L,V1,V2. L ⊢ V1 ⬌* V2 → ∀T1,T2. L.ⓓV1 ⊢ T1 ⬌* T2 →
+ L ⊢ ⓓV1. T1 ⬌* ⓓV2. T2.
+#L #V1 #V2 #HV12 #T1 #T2 #HT12
+@(cpcs_trans … (ⓓV1.T2)) /2 width=1/
+qed.
+
+lemma cpcs_abbr2: ∀L,V1,V2. L ⊢ V1 ⬌* V2 → ∀T1,T2. L.ⓓV2 ⊢ T1 ⬌* T2 →
+ L ⊢ ⓓV1. T1 ⬌* ⓓV2. T2.
+#L #V1 #V2 #HV12 #T1 #T2 #HT12
+@(cpcs_trans … (ⓓV2.T1)) /2 width=1/
+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.ma".
+include "basic_2/equivalence/cpcs.ma".
+
+(* CONTEXT-SENSITIVE PARALLEL EQUIVALENCE ON TERMS **************************)
+
+(* Properties about context sensitive computation on terms ******************)
+
+(* Basic_1: was: pc3_pr3_r *)
+lemma cpcs_cprs_dx: ∀L,T1,T2. L ⊢ T1 ➡* T2 → L ⊢ T1 ⬌* T2.
+#L #T1 #T2 #H @(cprs_ind … H) -T2 /width=1/ /3 width=3/
+qed.
+
+(* Basic_1: was: pc3_pr3_x *)
+lemma cpcs_cprs_sn: ∀L,T1,T2. L ⊢ T2 ➡* T1 → L ⊢ T1 ⬌* T2.
+#L #T1 #T2 #H @(cprs_ind_dx … H) -T2 /width=1/ /3 width=3/
+qed.
+
+lemma cpcs_cprs_strap1: ∀L,T1,T. L ⊢ T1 ⬌* T → ∀T2. L ⊢ T ➡* T2 → L ⊢ T1 ⬌* T2.
+#L #T1 #T #HT1 #T2 #H @(cprs_ind … H) -T2 /width=1/ /2 width=3/
+qed.
+
+lemma cpcs_cprs_strap2: ∀L,T1,T. L ⊢ T1 ➡* T → ∀T2. L ⊢ T ⬌* T2 → L ⊢ T1 ⬌* T2.
+#L #T1 #T #H #T2 #HT2 @(cprs_ind_dx … H) -T1 /width=1/ /2 width=3/
+qed.
+
+lemma cpcs_cprs_div: ∀L,T1,T. L ⊢ T1 ⬌* T → ∀T2. L ⊢ T2 ➡* T → L ⊢ T1 ⬌* T2.
+#L #T1 #T #HT1 #T2 #H @(cprs_ind_dx … H) -T2 /width=1/ /2 width=3/
+qed.
+
+(* Basic_1: was: pc3_pr3_conf *)
+lemma cpcs_cprs_conf: ∀L,T1,T. L ⊢ T ➡* T1 → ∀T2. L ⊢ T ⬌* T2 → L ⊢ T1 ⬌* T2.
+#L #T1 #T #H #T2 #HT2 @(cprs_ind … H) -T1 /width=1/ /2 width=3/
+qed.
+
+(* Basic_1: was: pc3_pr3_t *)
+(* Basic_1: note: pc3_pr3_t should be renamed *)
+lemma cprs_div: ∀L,T1,T. L ⊢ T1 ➡* T → ∀T2. L ⊢ T2 ➡* T → L ⊢ T1 ⬌* T2.
+#L #T1 #T #HT1 #T2 #H @(cprs_ind_dx … H) -T2 /2 width=1/ /2 width=3/
+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_ltpss.ma".
+include "basic_2/reducibility/ltpr_aaa.ma".
+include "basic_2/reducibility/cpr.ma".
+include "basic_2/reducibility/lcpr.ma".
+
+(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS *************)
+
+(* Properties about atomic arity assignment on terms ************************)
+
+lemma aaa_lcpr_conf: ∀L1,T,A. L1 ⊢ T ÷ A → ∀L2. L1 ⊢ ➡ L2 → L2 ⊢ T ÷ A.
+#L1 #T #A #HT #L2 * /3 width=5/
+qed.
+
+lemma aaa_cpr_conf: ∀L,T1,A. L ⊢ T1 ÷ A → ∀T2. L ⊢ T1 ➡ T2 → L ⊢ T2 ÷ A.
+#L #T1 #A #HT1 #T2 * /3 width=5/
+qed.
+
+lemma aaa_lcpr_cpr_conf: ∀L1,T1,A. L1 ⊢ T1 ÷ A → ∀L2. L1 ⊢ ➡ L2 →
+ ∀T2. L2 ⊢ T1 ➡ T2 → L2 ⊢ T2 ÷ A.
+/3 width=3/ qed.
(* *)
(**************************************************************************)
-include "basic_2/reducibility/ltpr_ldrop.ma".
include "basic_2/static/aaa_ltps.ma".
include "basic_2/static/lsuba_aaa.ma".
+include "basic_2/reducibility/ltpr_ldrop.ma".
(* CONTEXT-FREE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ********************)