]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_reqx.ma
λδ-2B is released
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_computation / cpxs_reqx.ma
index eb1444b26f69ccddd91a3e0e7d9f6203c4762bea..64be00b99ccc3150af51ecf8aa0a16bb103b3964 100644 (file)
@@ -19,7 +19,7 @@ include "basic_2/rt_computation/cpxs_teqx.ma".
 
 (* Properties with sort-irrelevant equivalence for local environments *******)
 
-(* Basic_2A1: was just: lleq_cpxs_trans *) 
+(* Basic_2A1: was just: lleq_cpxs_trans *)
 lemma reqx_cpxs_trans: ∀h,G,L0,T0,T1. ⦃G,L0⦄ ⊢ T0 ⬈*[h] T1 →
                        ∀L2. L2 ≛[T0] L0 →
                        ∃∃T. ⦃G,L2⦄ ⊢ T0 ⬈*[h] T & T ≛ T1.
@@ -31,19 +31,19 @@ elim (teqx_cpxs_trans … H2 … H3) -T #U0 #H2 #H3
 /3 width=5 by cpxs_strap2, teqx_trans, ex2_intro/
 qed-.
 
-(* Basic_2A1: was just: cpxs_lleq_conf *) 
+(* Basic_2A1: was just: cpxs_lleq_conf *)
 lemma cpxs_reqx_conf: ∀h,G,L0,T0,T1. ⦃G,L0⦄ ⊢ T0 ⬈*[h] T1 →
                       ∀L2. L0 ≛[T0] L2 →
                       ∃∃T. ⦃G,L2⦄ ⊢ T0 ⬈*[h] T & T ≛ T1.
 /3 width=3 by reqx_cpxs_trans, reqx_sym/ qed-.
 
-(* Basic_2A1: was just: cpxs_lleq_conf_dx *) 
+(* Basic_2A1: was just: cpxs_lleq_conf_dx *)
 lemma cpxs_reqx_conf_dx: ∀h,G,L2,T1,T2. ⦃G,L2⦄ ⊢ T1 ⬈*[h] T2 →
                          ∀L1. L1 ≛[T1] L2 → L1 ≛[T2] L2.
 #h #G #L2 #T1 #T2 #H @(cpxs_ind … H) -T2 /3 width=6 by cpx_reqx_conf_dx/
 qed-.
 
-(* Basic_2A1: was just: lleq_conf_sn *) 
+(* Basic_2A1: was just: lleq_conf_sn *)
 lemma cpxs_reqx_conf_sn: ∀h,G,L1,T1,T2. ⦃G,L1⦄ ⊢ T1 ⬈*[h] T2 →
                          ∀L2. L1 ≛[T1] L2 → L1 ≛[T2] L2.
 /4 width=6 by cpxs_reqx_conf_dx, reqx_sym/ qed-.