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 notation "hvbox( ⦃ term 46 L1 ⦄ ➡ break ⦃ term 46 L2 ⦄ )"
16 non associative with precedence 45
17 for @{ 'FocalizedPRed $L1 $L2 }.
19 include "basic_2/unfold/ltpss_sn.ma".
20 include "basic_2/reducibility/ltpr.ma".
22 (* FOCALIZED PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ***********************)
24 definition lfpr: relation lenv ≝
25 λL1,L2. ∃∃L. L1 ➡ L & L ⊢ ▶* [0, |L|] L2
29 "focalized parallel reduction (environment)"
30 'FocalizedPRed L1 L2 = (lfpr L1 L2).
32 (* Basic properties *********************************************************)
35 lemma lfpr_refl: ∀L. ⦃L⦄ ➡ ⦃L⦄.
38 lemma ltpss_sn_lfpr: ∀L1,L2,d,e. L1 ⊢ ▶* [d, e] L2 → ⦃L1⦄ ➡ ⦃L2⦄.
41 lemma ltpr_lfpr: ∀L1,L2. L1 ➡ L2 → ⦃L1⦄ ➡ ⦃L2⦄.
44 (* Basic inversion lemmas ***************************************************)
46 lemma lfpr_inv_atom1: ∀L2. ⦃⋆⦄ ➡ ⦃L2⦄ → L2 = ⋆.
47 #L2 * #L #HL >(ltpr_inv_atom1 … HL) -HL #HL2 >(ltpss_sn_inv_atom1 … HL2) -HL2 //