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 @{ 'FocalizedPConvAlt $L1 $L2 }.
19 include "basic_2/reducibility/lfpr.ma".
21 (* FOCALIZED PARALLEL CONVERSION ON LOCAL ENVIRONMENTS **********************)
23 definition lfpc: relation lenv ≝
24 λL1,L2. ⦃L1⦄ ➡ ⦃L2⦄ ∨ ⦃L2⦄ ➡ ⦃L1⦄.
27 "focalized parallel conversion (local environment)"
28 'FocalizedPConv L1 L2 = (lfpc L1 L2).
30 (* Basic properties *********************************************************)
32 lemma lfpc_refl: ∀L. ⦃L⦄ ⬌ ⦃L⦄.
35 lemma lfpc_sym: ∀L1,L2. ⦃L1⦄ ⬌ ⦃L2⦄ → ⦃L2⦄ ⬌ ⦃L1⦄.
39 lemma lfpc_lfpr: ∀L1,L2. ⦃L1⦄ ⬌ ⦃L2⦄ → ∃∃L. ⦃L1⦄ ➡ ⦃L⦄ & ⦃L2⦄ ➡ ⦃L⦄.