1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "basic_2/rt_transition/fpbc_lpx.ma".
16 include "basic_2/rt_computation/fpbs_lpxs.ma".
17 include "basic_2/rt_computation/fpbg_fqup.ma".
19 (* PROPER PARALLEL RST-COMPUTATION FOR CLOSURES *****************************)
21 (* Properties with extended rt-computation on full local environments *******)
23 (* Basic_2A1: uses: lpxs_fpbg *)
24 lemma lpxs_rneqx_fpbg:
25 ∀G,L1,L2,T. ❨G,L1❩ ⊢ ⬈* L2 →
26 (L1 ≅[T] L2 → ⊥) → ❨G,L1,T❩ > ❨G,L2,T❩.
28 elim (lpxs_rneqg_inv_step_sn … H … H0) -H -H0
29 /4 width=10 by fpbc_fpbs_fpbg, lpxs_feqg_fpbs, lpx_fpbc, feqg_intro_sn, sfull_dec/