(* *)
(**************************************************************************)
+include "basic_2/static/lfxs_lfxs.ma".
+include "basic_2/rt_transition/lfpx_frees.ma".
include "basic_2/rt_transition/cpm_lsubr.ma".
include "basic_2/rt_transition/cpr.ma".
include "basic_2/rt_transition/cpr_drops.ma".
include "basic_2/rt_transition/lfpr_drops.ma".
include "basic_2/rt_transition/lfpr_fqup.ma".
+include "basic_2/rt_transition/lfpr_lfpx.ma".
(* PARALLEL R-TRANSITION FOR LOCAL ENV.S ON REFERRED ENTRIES ****************)
#h #G #L0 #T0 #T1 #HT01 #L1 #HL01
elim (cpr_conf_lfpr … HT01 T0 … L0 … HL01) /2 width=3 by ex2_intro/
qed-.
+
+(* Main properties **********************************************************)
+
+theorem lfpr_conf: ∀h,G,T. confluent … (lfpr h G T).
+/4 width=6 by cpr_conf_lfpr, lfpx_frees_conf_fwd_lfpr, lfpx_frees_conf, lfxs_conf/ qed-.
(* Fwd. lemmas with unc. rt-transition for local env.s on referred entries **)
+lemma lfpx_frees_conf_fwd_lfpr: ∀h,G. lexs_frees_confluent (cpx h G) cfull →
+ lexs_frees_confluent (cpm 0 h G) cfull.
+/4 width=7 by cpm_fwd_cpx, lexs_co/ qed-.
+
lemma lfpr_fwd_lfpx: ∀h,T,G,L1,L2. ⦃G, L1⦄ ⊢ ➡[h, T] L2 → ⦃G, L1⦄ ⊢ ⬈[h, T] L2.
/3 width=3 by cpm_fwd_cpx, lfxs_co/ 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/lfxs_lfxs.ma".
-include "basic_2/rt_transition/lfpr_lfpr.ma".
-
-theorem lfpr_conf: ∀h,G,T. confluent … (lfpr h G T).
-#h #G @lfxs_conf [ | @cpr_conf_lfpr ]
-qed-.
-
(* Properties with context-sensitive free variables *************************)
-axiom frees_pair_flat: ∀L,T,f1,I1,V1. L.ⓑ{I1}V1 ⊢ 𝐅*⦃T⦄ ≡ f1 →
- ∀f2,I2,V2. L.ⓑ{I2}V2 ⊢ 𝐅*⦃T⦄ ≡ f2 →
- ∀f0. f1 ⋓ f2 ≡ f0 →
- ∀I0,I. L.ⓑ{I0}ⓕ{I}V1.V2 ⊢ 𝐅*⦃T⦄ ≡ f0.
-
(* Basic_2A1: was: lpx_cpx_frees_trans *)
lemma cpx_frees_conf_lfpx: ∀h,G,L1,T1,f1. L1 ⊢ 𝐅*⦃T1⦄ ≡ f1 →
∀L2. L1 ⦻*[cpx h G, cfull, f1] L2 →
cpg.ma cpg_simple.ma cpg_drops.ma cpg_lsubr.ma
cpx.ma cpx_simple.ma cpx_drops.ma cpx_lsubr.ma
-lfpx.ma lfpx_length.ma lfpx_drops.ma lfpx_fqup.ma
+lfpx.ma lfpx_length.ma lfpx_drops.ma lfpx_fqup.ma lfpx_frees.ma
cpm.ma cpm_simple.ma cpm_drops.ma cpm_lsubr.ma cpm_cpx.ma
cpr.ma cpr_drops.ma
lfpr.ma lfpr_length.ma lfpr_drops.ma lfpr_fqup.ma lfpr_lfpx.ma lfpr_lfpr.ma
<subsection name="A2">Stage "A2": "Extending the Applicability Condition"</subsection>
+ <news class="alpha" date="2017 January 17.">
+ Confluence for parallel r-transition on referred entries of local environments.
+ </news>
<news class="alpha" date="2016 September 15.">
Confluence for context-sensitive parallel r-transition on terms.
</news>
}
]
[ { "context-sensitive rt-reduction" * } {
- [ "lpx_drop" + "lpx_frees" + "lpx_lleq" + "lpx_aaa" * ]
+ [ "lpx_lleq" + "lpx_aaa" * ]
[ "cpx_lreq" + "cpx_fqus" + "cpx_llpx_sn" + "cpx_lleq" * ]
}
]
}
]
[ { "uncounted context-sensitive rt-transition" * } {
- [ "lfpx ( ⦃?,?⦄ ⊢ ⬈[?,?] ? )" "lfpx_length" + "lfpx_drops" + "lfpx_fqup" * ]
+ [ "lfpx ( ⦃?,?⦄ ⊢ ⬈[?,?] ? )" "lfpx_length" + "lfpx_drops" + "lfpx_fqup" + "lfpx_frees" * ]
[ "cpx ( ⦃?,?⦄ ⊢ ? ⬈[?] ? )" "cpx_simple" + "cpx_drops" + "cpx_lsubr" * ]
}
]