--- /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/notation/relations/btpredalt_8.ma".
+include "basic_2/reduction/fpb_fleq.ma".
+include "basic_2/reduction/fpbq.ma".
+
+(* "QRST" PARALLEL REDUCTION FOR CLOSURES ***********************************)
+
+(* Basic properties *********************************************************)
+
+lemma fleq_fpbq: ∀h,o,G1,G2,L1,L2,T1,T2.
+ ⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 * /2 width=1 by fpbq_lleq/
+qed.
+
+lemma fpb_fpbq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2
+/3 width=1 by fpbq_fquq, fpbq_cpx, fpbq_lpx, fqu_fquq/
+qed.
+
+(* Advanced eliminators *****************************************************)
+
+lemma fpbq_ind_alt: ∀h,o,G1,G2,L1,L2,T1,T2. ∀R:Prop.
+ (⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → R) →
+ (⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ → R) →
+ ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ → R.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #R #HI1 #HI2 #H elim (fpbq_fpbqa … H) /2 width=1 by/
+qed-.
+
+(* aternative definition of fpb *********************************************)
+
+lemma fpb_fpbq_alt: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ ∧ (⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → ⊥).
+/3 width=10 by fpb_fpbq, fpb_inv_fleq, conj/ qed.
+
+lemma fpbq_inv_fpb_alt: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ →
+ (⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → ⊥) → ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #H0 @(fpbq_ind_alt … H) -H //
+#H elim H0 -H0 //
+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/multiple/llor_drop.ma".
+include "basic_2/multiple/llpx_sn_llor.ma".
+include "basic_2/multiple/llpx_sn_lpx_sn.ma".
+include "basic_2/multiple/lleq_lreq.ma".
+include "basic_2/multiple/lleq_llor.ma".
+include "basic_2/reduction/cpx_lreq.ma".
+include "basic_2/reduction/cpx_lleq.ma".
+include "basic_2/reduction/lpx_frees.ma".
+
+(* SN EXTENDED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS ********************)
+
+
+fact lreq_lpx_trans_lleq_aux: ∀h,o,G,L1,L0,l,k. L1 ⩬[l, k] L0 → k = ∞ →
+ ∀L2. ⦃G, L0⦄ ⊢ ➡[h, o] L2 →
+ ∃∃L. L ⩬[l, k] L2 & ⦃G, L1⦄ ⊢ ➡[h, o] L &
+ (∀T. L0 ≡[T, l] L2 ↔ L1 ≡[T, l] L).
+#h #o #G #L1 #L0 #l #k #H elim H -L1 -L0 -l -k
+[ #l #k #_ #L2 #H >(lpx_inv_atom1 … H) -H
+ /3 width=5 by ex3_intro, conj/
+| #I1 #I0 #L1 #L0 #V1 #V0 #_ #_ #Hm destruct
+| #I #L1 #L0 #V1 #k #HL10 #IHL10 #Hm #Y #H
+ elim (lpx_inv_pair1 … H) -H #L2 #V2 #HL02 #HV02 #H destruct
+ lapply (ysucc_inv_Y_dx … Hm) -Hm #Hm
+ elim (IHL10 … HL02) // -IHL10 -HL02 #L #HL2 #HL1 #IH
+ @(ex3_intro … (L.ⓑ{I}V2)) /3 width=3 by lpx_pair, lreq_cpx_trans, lreq_pair/
+ #T elim (IH T) #HL0dx #HL0sn
+ @conj #H @(lleq_lreq_repl … H) -H /3 width=1 by lreq_sym, lreq_pair_O_Y/
+| #I1 #I0 #L1 #L0 #V1 #V0 #l #k #HL10 #IHL10 #Hm #Y #H
+ elim (lpx_inv_pair1 … H) -H #L2 #V2 #HL02 #HV02 #H destruct
+ elim (IHL10 … HL02) // -IHL10 -HL02 #L #HL2 #HL1 #IH
+ @(ex3_intro … (L.ⓑ{I1}V1)) /3 width=1 by lpx_pair, lreq_succ/
+ #T elim (IH T) #HL0dx #HL0sn
+ @conj #H @(lleq_lreq_repl … H) -H /3 width=1 by lreq_sym, lreq_succ/
+]
+qed-.
+
+lemma lreq_lpx_trans_lleq: ∀h,o,G,L1,L0,l. L1 ⩬[l, ∞] L0 →
+ ∀L2. ⦃G, L0⦄ ⊢ ➡[h, o] L2 →
+ ∃∃L. L ⩬[l, ∞] L2 & ⦃G, L1⦄ ⊢ ➡[h, o] L &
+ (∀T. L0 ≡[T, l] L2 ↔ L1 ≡[T, l] L).
+/2 width=1 by lreq_lpx_trans_lleq_aux/ 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 *)
+(* *)
+(**************************************************************************)
+
+(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
+
+notation "hvbox( ⦃ term 46 G1, break term 46 L1, break term 46 T1 ⦄ ≽ [ break term 46 h ] ⦃ break term 46 G2, break term 46 L2 , break term 46 T2 ⦄ )"
+ non associative with precedence 45
+ for @{ 'BTPRed $h $G1 $L1 $T1 $G2 $L2 $T2 }.
+++ /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 *)
-(* *)
-(**************************************************************************)
-
-(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-
-notation "hvbox( ⦃ term 46 G1, break term 46 L1, break term 46 T1 ⦄ ≽ [ break term 46 h, break term 46 o ] ⦃ break term 46 G2, break term 46 L2 , break term 46 T2 ⦄ )"
- non associative with precedence 45
- for @{ 'BTPRed $h $o $G1 $L1 $T1 $G2 $L2 $T2 }.
+++ /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 *)
-(* *)
-(**************************************************************************)
-
-(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-
-notation "hvbox( ⦃ term 46 G1, break term 46 L1, break term 46 T1 ⦄ ≽ ≽ [ break term 46 h, break term 46 o ] ⦃ break term 46 G2, break term 46 L2 , break term 46 T2 ⦄ )"
- non associative with precedence 45
- for @{ 'BTPRedAlt $h $o $G1 $L1 $T1 $G2 $L2 $T2 }.
--- /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/lfdeq.ma".
+include "basic_2/rt_transition/cpx_lfxs.ma".
+
+(* UNCOUNTED CONTEXT-SENSITIVE PARALLEL RT-TRANSITION FOR TERMS *************)
+
+(* Properties with degree-based equivalence for local environments **********)
+
+(* Basic_2A1: was just: cpx_lleq_conf_sn *)
+lemma cpx_lfdeq_conf_sn: ∀h,o,G. s_r_confluent1 … (cpx h G) (lfdeq h o).
+/3 width=6 by cpx_lfxs_conf/ qed-.
+
+(* Basic_2A1: was just: cpx_lleq_conf_dx *)
+lemma cpx_lfdeq_conf_dx: ∀h,o,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ⬈[h] T2 →
+ ∀L1. L1 ≡[h, o, T1] L2 → L1 ≡[h, o, T2] L2.
+/4 width=4 by cpx_lfdeq_conf_sn, lfdeq_sym/ qed-.
(* PROPER PARALLEL RST-TRANSITION FOR CLOSURES ******************************)
inductive fpb (h) (o) (G1) (L1) (T1): relation3 genv lenv term ≝
-| fpb_fqu: ∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → fpb h o G1 L1 T1 G2 L2 T2
-| fpb_cpx: ∀T2. ⦃G1, L1⦄ ⊢ T1 ⬈[h] T2 → (T1 ≡[h, o] T2 → ⊥) → fpb h o G1 L1 T1 G1 L1 T2
-| fpb_lpx: ∀L2. ⦃G1, L1⦄ ⊢ ⬈[h, T1] L2 → (L1 ≡[h, o, T1] L2 → ⊥) → fpb h o G1 L1 T1 G1 L2 T1
+| fpb_fqu : ∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → fpb h o G1 L1 T1 G2 L2 T2
+| fpb_cpx : ∀T2. ⦃G1, L1⦄ ⊢ T1 ⬈[h] T2 → (T1 ≡[h, o] T2 → ⊥) → fpb h o G1 L1 T1 G1 L1 T2
+| fpb_lfpx: ∀L2. ⦃G1, L1⦄ ⊢ ⬈[h, T1] L2 → (L1 ≡[h, o, T1] L2 → ⊥) → fpb h o G1 L1 T1 G1 L2 T1
.
interpretation
⦃G, L, T1⦄ ≻[h, o] ⦃G, L, T2⦄.
/3 width=2 by fpb_cpx, cpm_fwd_cpx/ qed.
-lemma lpr_fpb: ∀h,o,G,L1,L2,T. ⦃G, L1⦄ ⊢ ➡[h, T] L2 → (L1 ≡[h, o, T] L2 → ⊥) →
- ⦃G, L1, T⦄ ≻[h, o] ⦃G, L2, T⦄.
-/3 width=1 by fpb_lpx, lfpr_fwd_lfpx/ qed.
+(* Basic_2A1: includes: lpr_fpb *)
+lemma lfpr_fpb: ∀h,o,G,L1,L2,T. ⦃G, L1⦄ ⊢ ➡[h, T] L2 → (L1 ≡[h, o, T] L2 → ⊥) →
+ ⦃G, L1, T⦄ ≻[h, o] ⦃G, L2, T⦄.
+/3 width=1 by fpb_lfpx, lfpr_fwd_lfpx/ qed.
(* *)
(**************************************************************************)
-include "basic_2/notation/relations/btpred_8.ma".
-include "basic_2/substitution/fquq.ma".
-include "basic_2/multiple/lleq.ma".
-include "basic_2/reduction/lpx.ma".
-
-(* "QRST" PARALLEL REDUCTION FOR CLOSURES ***********************************)
-
-inductive fpbq (h) (o) (G1) (L1) (T1): relation3 genv lenv term ≝
-| fpbq_fquq: ∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → fpbq h o G1 L1 T1 G2 L2 T2
-| fpbq_cpx : ∀T2. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] T2 → fpbq h o G1 L1 T1 G1 L1 T2
-| fpbq_lpx : ∀L2. ⦃G1, L1⦄ ⊢ ➡[h, o] L2 → fpbq h o G1 L1 T1 G1 L2 T1
-| fpbq_lleq: ∀L2. L1 ≡[T1, 0] L2 → fpbq h o G1 L1 T1 G1 L2 T1
+include "basic_2/notation/relations/btpred_7.ma".
+include "basic_2/s_transition/fquq.ma".
+include "basic_2/rt_transition/lfpr_lfpx.ma".
+
+(* PARALLEL RST-TRANSITION FOR CLOSURES *************************************)
+
+(* Basic_2A1: includes: fpbq_lleq *)
+inductive fpbq (h) (G1) (L1) (T1): relation3 genv lenv term ≝
+| fpbq_fquq: ∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → fpbq h G1 L1 T1 G2 L2 T2
+| fpbq_cpx : ∀T2. ⦃G1, L1⦄ ⊢ T1 ⬈[h] T2 → fpbq h G1 L1 T1 G1 L1 T2
+| fpbq_lfpx: ∀L2. ⦃G1, L1⦄ ⊢ ⬈[h, T1] L2 → fpbq h G1 L1 T1 G1 L2 T1
.
interpretation
- "'qrst' parallel reduction (closure)"
- 'BTPRed h o G1 L1 T1 G2 L2 T2 = (fpbq h o G1 L1 T1 G2 L2 T2).
+ "parallel rst-transition (closure)"
+ 'BTPRed h G1 L1 T1 G2 L2 T2 = (fpbq h G1 L1 T1 G2 L2 T2).
(* Basic properties *********************************************************)
-lemma fpbq_refl: ∀h,o. tri_reflexive … (fpbq h o).
+lemma fpbq_refl: ∀h. tri_reflexive … (fpbq h).
/2 width=1 by fpbq_cpx/ qed.
-lemma cpr_fpbq: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ⦃G, L, T1⦄ ≽[h, o] ⦃G, L, T2⦄.
-/3 width=1 by fpbq_cpx, cpr_cpx/ qed.
+(* Basic_2A1: includes: cpr_fpbq *)
+lemma cpm_fpbq: ∀n,h,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[n, h] T2 → ⦃G, L, T1⦄ ≽[h] ⦃G, L, T2⦄.
+/3 width=2 by fpbq_cpx, cpm_fwd_cpx/ qed.
+
+lemma lfpr_fpbq: ∀h,G,L1,L2,T. ⦃G, L1⦄ ⊢ ➡[h, T] L2 → ⦃G, L1, T⦄ ≽[h] ⦃G, L2, T⦄.
+/3 width=1 by fpbq_lfpx, lfpr_fwd_lfpx/ qed.
-lemma lpr_fpbq: ∀h,o,G,L1,L2,T. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L1, T⦄ ≽[h, o] ⦃G, L2, T⦄.
-/3 width=1 by fpbq_lpx, lpr_lpx/ qed.
+(* Basic_2A1: removed theorems 2:
+ fpbq_fpbqa fpbqa_inv_fpbq
+*)
(**************************************************************************)
include "basic_2/static/aaa_fqus.ma".
-include "basic_2/static/aaa_lleq.ma".
-include "basic_2/reduction/lpx_aaa.ma".
-include "basic_2/reduction/fpbq.ma".
+include "basic_2/rt_transition/lfpx_aaa.ma".
+include "basic_2/rt_transition/fpbq.ma".
-(* "QRST" PARALLEL REDUCTION FOR CLOSURES ***********************************)
+(* PARALLEL RST-TRANSITION FOR CLOSURES *************************************)
-(* Properties on atomic arity assignment for terms **************************)
+(* Properties with atomic arity assignment for terms ************************)
-lemma fpbq_aaa_conf: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ →
+lemma fpbq_aaa_conf: ∀h,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≽[h] ⦃G2, L2, T2⦄ →
∀A1. ⦃G1, L1⦄ ⊢ T1 ⁝ A1 → ∃A2. ⦃G2, L2⦄ ⊢ T2 ⁝ A2.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2
-/3 width=6 by aaa_lleq_conf, lpx_aaa_conf, cpx_aaa_conf, aaa_fquq_conf, ex_intro/
+#h #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2
+/3 width=6 by lfpx_aaa_conf, cpx_aaa_conf, aaa_fquq_conf, ex_intro/
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/notation/relations/btpredalt_8.ma".
-include "basic_2/reduction/fpb_fleq.ma".
-include "basic_2/reduction/fpbq.ma".
-
-(* "QRST" PARALLEL REDUCTION FOR CLOSURES ***********************************)
-
-(* alternative definition of fpbq *)
-definition fpbqa: ∀h. sd h → tri_relation genv lenv term ≝
- λh,o,G1,L1,T1,G2,L2,T2.
- ⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ ∨ ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄.
-
-interpretation
- "'qrst' parallel reduction (closure) alternative"
- 'BTPRedAlt h o G1 L1 T1 G2 L2 T2 = (fpbqa h o G1 L1 T1 G2 L2 T2).
-
-(* Basic properties *********************************************************)
-
-lemma fleq_fpbq: ∀h,o,G1,G2,L1,L2,T1,T2.
- ⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 * /2 width=1 by fpbq_lleq/
-qed.
-
-lemma fpb_fpbq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2
-/3 width=1 by fpbq_fquq, fpbq_cpx, fpbq_lpx, fqu_fquq/
-qed.
-
-(* Main properties **********************************************************)
-
-theorem fpbq_fpbqa: ∀h,o,G1,G2,L1,L2,T1,T2.
- ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ≽≽[h, o] ⦃G2, L2, T2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2
-[ #G2 #L2 #T2 #H elim (fquq_inv_gen … H) -H
- [ /3 width=1 by fpb_fqu, or_intror/
- | * #H1 #H2 #H3 destruct /2 width=1 by or_introl/
- ]
-| #T2 #HT12 elim (eq_term_dec T1 T2)
- #HnT12 destruct /4 width=1 by fpb_cpx, or_intror, or_introl/
-| #L2 #HL12 elim (lleq_dec … T1 L1 L2 0)
- /4 width=1 by fpb_lpx, fleq_intro, or_intror, or_introl/
-| /3 width=1 by fleq_intro, or_introl/
-]
-qed.
-
-theorem fpbqa_inv_fpbq: ∀h,o,G1,G2,L1,L2,T1,T2.
- ⦃G1, L1, T1⦄ ≽≽[h, o] ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 * /2 width=1 by fleq_fpbq, fpb_fpbq/
-qed-.
-
-(* Advanced eliminators *****************************************************)
-
-lemma fpbq_ind_alt: ∀h,o,G1,G2,L1,L2,T1,T2. ∀R:Prop.
- (⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → R) →
- (⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ → R) →
- ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ → R.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #R #HI1 #HI2 #H elim (fpbq_fpbqa … H) /2 width=1 by/
-qed-.
-
-(* aternative definition of fpb *********************************************)
-
-lemma fpb_fpbq_alt: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ ∧ (⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → ⊥).
-/3 width=10 by fpb_fpbq, fpb_inv_fleq, conj/ qed.
-
-lemma fpbq_inv_fpb_alt: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ →
- (⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → ⊥) → ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #H0 @(fpbq_ind_alt … H) -H //
-#H elim H0 -H0 //
-qed-.
include "basic_2/static/lfdeq_fqup.ma".
include "basic_2/rt_transition/lfpx.ma".
-(* UNCOUNTED CONTEXT-SENSITIVE PARALLEL RT-TRANSITION FOR TERMS *************)
+(* UNCOUNTED PARALLEL RT-TRANSITION FOR LOCAL ENV.S ON REFERRED ENTRIES *****)
-(* Properties with degree-based equivalence for terms ***********************)
+(* Properties with degree-based equivalence for local environments **********)
-(* Basic_2A1: was just: cpx_lleq_conf *)
lemma cpx_tdeq_conf_lexs: ∀h,o,G. R_confluent2_lfxs … (cpx h G) (cdeq h o) (cpx h G) (cdeq h o).
#h #o #G #L0 #T0 #T1 #H @(cpx_ind … H) -G -L0 -T0 -T1 /2 width=3 by ex2_intro/
[ #G #L0 #s0 #X0 #H0 #L1 #HL01 #L2 #HL02
elim (cpx_tdeq_conf … HT01 T2) -HT01 /3 width=3 by tdeq_sym, ex2_intro/
qed-.
+(* Basic_2A1: was just: cpx_lleq_conf *)
+lemma cpx_lfdeq_conf: ∀h,o,G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ⬈[h] T1 →
+ ∀L2. L0 ≡[h, o, T0] L2 →
+ ∃∃T. ⦃G, L2⦄ ⊢ T0 ⬈[h] T & T1 ≡[h, o] T.
+#h #o #G #L0 #T0 #T1 #HT01 #L2 #HL02
+elim (cpx_tdeq_conf_lexs … HT01 T0 … L0 … HL02) -HT01 -HL02
+/2 width=3 by lfxs_refl, ex2_intro/
+qed-.
+
+(* Basic_2A1: was just: lleq_cpx_trans *)
+lemma lfdeq_cpx_trans: ∀h,o,G,L2,L0,T0. L2 ≡[h, o, T0] L0 →
+ ∀T1. ⦃G, L0⦄ ⊢ T0 ⬈[h] T1 →
+ ∃∃T. ⦃G, L2⦄ ⊢ T0 ⬈[h] T & T ≡[h, o] T1.
+#h #o #G #L2 #L0 #T0 #HL20 #T1 #HT01
+elim (cpx_lfdeq_conf … o … HT01 L2) -HT01
+/3 width=3 by lfdeq_sym, tdeq_sym, ex2_intro/
+qed-.
+
+include "basic_2/static/lfxs_lfxs.ma".
+
+axiom lfpx_lfdeq_conf: ∀h,o,G,T. confluent2 … (lfpx h G T) (lfdeq h o T).
+(*
+#H1 #H2 #H3 #H4 #H5 #H6 #H7 #H8 #H9
+@lfxs_conf
+*)
+(* Basic_2A1: was just: lleq_lpx_trans *)
+lemma lfdeq_lfpx_trans: ∀h,o,G,T,L2,K2. ⦃G, L2⦄ ⊢ ⬈[h, T] K2 →
+ ∀L1. L1 ≡[h, o, T] L2 →
+ ∃∃K1. ⦃G, L1⦄ ⊢ ⬈[h, T] K1 & K1 ≡[h, o, T] K2.
+#h #o #G #T #L2 #K2 #HLK2 #L1 #HL12
+elim (lfpx_lfdeq_conf … o … HLK2 L1)
+/3 width=3 by lfdeq_sym, ex2_intro/
+qed-.
(*
-lemma cpx_lleq_conf: ∀h,o,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ⬈[h, o] T2 →
- ∀L1. L2 ≡[T1, 0] L1 → ⦃G, L1⦄ ⊢ T1 ⬈[h, o] T2.
-/3 width=3 by lleq_cpx_trans, lleq_sym/ qed-.
+(* Properties with supclosure ***********************************************)
-lemma cpx_lleq_conf_sn: ∀h,o,G. s_r_confluent1 … (cpx h o G) (lleq 0).
-/3 width=6 by cpx_llpx_sn_conf, lift_mono, ex2_intro/ qed-.
+lemma lpx_lleq_fqu_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+ ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 → K1 ≡[T1, 0] L1 →
+ ∃∃K2. ⦃G1, K1, T1⦄ ⊐ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2 & K2 ≡[T2, 0] L2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
+[ #I #G1 #L1 #V1 #X #H1 #H2 elim (lpx_inv_pair2 … H1) -H1
+ #K0 #V0 #H1KL1 #_ #H destruct
+ elim (lleq_inv_lref_ge_dx … H2 ? I L1 V1) -H2 //
+ #K1 #H #H2KL1 lapply (drop_inv_O2 … H) -H #H destruct
+ /2 width=4 by fqu_lref_O, ex3_intro/
+| * [ #a ] #I #G1 #L1 #V1 #T1 #K1 #HLK1 #H
+ [ elim (lleq_inv_bind … H)
+ | elim (lleq_inv_flat … H)
+ ] -H /2 width=4 by fqu_pair_sn, ex3_intro/
+| #a #I #G1 #L1 #V1 #T1 #K1 #HLK1 #H elim (lleq_inv_bind_O … H) -H
+ /3 width=4 by lpx_pair, fqu_bind_dx, ex3_intro/
+| #I #G1 #L1 #V1 #T1 #K1 #HLK1 #H elim (lleq_inv_flat … H) -H
+ /2 width=4 by fqu_flat_dx, ex3_intro/
+| #G1 #L1 #L #T1 #U1 #k #HL1 #HTU1 #K1 #H1KL1 #H2KL1
+ elim (drop_O1_le (Ⓕ) (k+1) K1)
+ [ #K #HK1 lapply (lleq_inv_lift_le … H2KL1 … HK1 HL1 … HTU1 ?) -H2KL1 //
+ #H2KL elim (lpx_drop_trans_O1 … H1KL1 … HL1) -L1
+ #K0 #HK10 #H1KL lapply (drop_mono … HK10 … HK1) -HK10 #H destruct
+ /3 width=4 by fqu_drop, ex3_intro/
+ | lapply (drop_fwd_length_le2 … HL1) -L -T1 -o
+ lapply (lleq_fwd_length … H2KL1) //
+ ]
+]
+qed-.
+
+lemma lpx_lleq_fquq_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
+ ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 → K1 ≡[T1, 0] L1 →
+ ∃∃K2. ⦃G1, K1, T1⦄ ⊐⸮ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2 & K2 ≡[T2, 0] L2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #K1 #H1KL1 #H2KL1
+elim (fquq_inv_gen … H) -H
+[ #H elim (lpx_lleq_fqu_trans … H … H1KL1 H2KL1) -L1
+ /3 width=4 by fqu_fquq, ex3_intro/
+| * #HG #HL #HT destruct /2 width=4 by ex3_intro/
+]
+qed-.
-lemma cpx_lleq_conf_dx: ∀h,o,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ⬈[h, o] T2 →
- ∀L1. L1 ≡[T1, 0] L2 → L1 ≡[T2, 0] L2.
-/4 width=6 by cpx_lleq_conf_sn, lleq_sym/ qed-.
-*)
\ No newline at end of file
+lemma lpx_lleq_fqup_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
+ ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 → K1 ≡[T1, 0] L1 →
+ ∃∃K2. ⦃G1, K1, T1⦄ ⊐+ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2 & K2 ≡[T2, 0] L2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
+[ #G2 #L2 #T2 #H #K1 #H1KL1 #H2KL1 elim (lpx_lleq_fqu_trans … H … H1KL1 H2KL1) -L1
+ /3 width=4 by fqu_fqup, ex3_intro/
+| #G #G2 #L #L2 #T #T2 #_ #HT2 #IHT1 #K1 #H1KL1 #H2KL1 elim (IHT1 … H2KL1) // -L1
+ #K #HT1 #H1KL #H2KL elim (lpx_lleq_fqu_trans … HT2 … H1KL H2KL) -L
+ /3 width=5 by fqup_strap1, ex3_intro/
+]
+qed-.
+
+lemma lpx_lleq_fqus_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
+ ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 → K1 ≡[T1, 0] L1 →
+ ∃∃K2. ⦃G1, K1, T1⦄ ⊐* ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2 & K2 ≡[T2, 0] L2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #K1 #H1KL1 #H2KL1
+elim (fqus_inv_gen … H) -H
+[ #H elim (lpx_lleq_fqup_trans … H … H1KL1 H2KL1) -L1
+ /3 width=4 by fqup_fqus, ex3_intro/
+| * #HG #HL #HT destruct /2 width=4 by ex3_intro/
+]
+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/multiple/llor_drop.ma".
-include "basic_2/multiple/llpx_sn_llor.ma".
-include "basic_2/multiple/llpx_sn_lpx_sn.ma".
-include "basic_2/multiple/lleq_lreq.ma".
-include "basic_2/multiple/lleq_llor.ma".
-include "basic_2/reduction/cpx_lreq.ma".
-include "basic_2/reduction/cpx_lleq.ma".
-include "basic_2/reduction/lpx_frees.ma".
-
-(* SN EXTENDED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS ********************)
-
-(* Properties on lazy equivalence for local environments ********************)
-
-(* Note: contains a proof of llpx_cpx_conf *)
-lemma lleq_lpx_trans: ∀h,o,G,L2,K2. ⦃G, L2⦄ ⊢ ➡[h, o] K2 →
- ∀L1,T,l. L1 ≡[T, l] L2 →
- ∃∃K1. ⦃G, L1⦄ ⊢ ➡[h, o] K1 & K1 ≡[T, l] K2.
-#h #o #G #L2 #K2 #HLK2 #L1 #T #l #HL12
-lapply (lpx_fwd_length … HLK2) #H1
-lapply (lleq_fwd_length … HL12) #H2
-lapply (lpx_sn_llpx_sn … T … l HLK2) // -HLK2 #H
-lapply (lleq_llpx_sn_trans … HL12 … H) /2 width=3 by lleq_cpx_trans/ -HL12 -H #H
-elim (llor_total L1 K2 T l) // -H1 -H2 #K1 #HLK1
-lapply (llpx_sn_llor_dx_sym … H … HLK1)
-[ /2 width=6 by cpx_frees_trans/
-| /3 width=10 by cpx_llpx_sn_conf, cpx_inv_lift1, cpx_lift/
-| /3 width=5 by llpx_sn_llor_fwd_sn, ex2_intro/
-]
-qed-.
-
-lemma lpx_lleq_fqu_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
- ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 → K1 ≡[T1, 0] L1 →
- ∃∃K2. ⦃G1, K1, T1⦄ ⊐ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2 & K2 ≡[T2, 0] L2.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
-[ #I #G1 #L1 #V1 #X #H1 #H2 elim (lpx_inv_pair2 … H1) -H1
- #K0 #V0 #H1KL1 #_ #H destruct
- elim (lleq_inv_lref_ge_dx … H2 ? I L1 V1) -H2 //
- #K1 #H #H2KL1 lapply (drop_inv_O2 … H) -H #H destruct
- /2 width=4 by fqu_lref_O, ex3_intro/
-| * [ #a ] #I #G1 #L1 #V1 #T1 #K1 #HLK1 #H
- [ elim (lleq_inv_bind … H)
- | elim (lleq_inv_flat … H)
- ] -H /2 width=4 by fqu_pair_sn, ex3_intro/
-| #a #I #G1 #L1 #V1 #T1 #K1 #HLK1 #H elim (lleq_inv_bind_O … H) -H
- /3 width=4 by lpx_pair, fqu_bind_dx, ex3_intro/
-| #I #G1 #L1 #V1 #T1 #K1 #HLK1 #H elim (lleq_inv_flat … H) -H
- /2 width=4 by fqu_flat_dx, ex3_intro/
-| #G1 #L1 #L #T1 #U1 #k #HL1 #HTU1 #K1 #H1KL1 #H2KL1
- elim (drop_O1_le (Ⓕ) (k+1) K1)
- [ #K #HK1 lapply (lleq_inv_lift_le … H2KL1 … HK1 HL1 … HTU1 ?) -H2KL1 //
- #H2KL elim (lpx_drop_trans_O1 … H1KL1 … HL1) -L1
- #K0 #HK10 #H1KL lapply (drop_mono … HK10 … HK1) -HK10 #H destruct
- /3 width=4 by fqu_drop, ex3_intro/
- | lapply (drop_fwd_length_le2 … HL1) -L -T1 -o
- lapply (lleq_fwd_length … H2KL1) //
- ]
-]
-qed-.
-
-lemma lpx_lleq_fquq_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 → K1 ≡[T1, 0] L1 →
- ∃∃K2. ⦃G1, K1, T1⦄ ⊐⸮ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2 & K2 ≡[T2, 0] L2.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #K1 #H1KL1 #H2KL1
-elim (fquq_inv_gen … H) -H
-[ #H elim (lpx_lleq_fqu_trans … H … H1KL1 H2KL1) -L1
- /3 width=4 by fqu_fquq, ex3_intro/
-| * #HG #HL #HT destruct /2 width=4 by ex3_intro/
-]
-qed-.
-
-lemma lpx_lleq_fqup_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
- ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 → K1 ≡[T1, 0] L1 →
- ∃∃K2. ⦃G1, K1, T1⦄ ⊐+ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2 & K2 ≡[T2, 0] L2.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
-[ #G2 #L2 #T2 #H #K1 #H1KL1 #H2KL1 elim (lpx_lleq_fqu_trans … H … H1KL1 H2KL1) -L1
- /3 width=4 by fqu_fqup, ex3_intro/
-| #G #G2 #L #L2 #T #T2 #_ #HT2 #IHT1 #K1 #H1KL1 #H2KL1 elim (IHT1 … H2KL1) // -L1
- #K #HT1 #H1KL #H2KL elim (lpx_lleq_fqu_trans … HT2 … H1KL H2KL) -L
- /3 width=5 by fqup_strap1, ex3_intro/
-]
-qed-.
-
-lemma lpx_lleq_fqus_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
- ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 → K1 ≡[T1, 0] L1 →
- ∃∃K2. ⦃G1, K1, T1⦄ ⊐* ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2 & K2 ≡[T2, 0] L2.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #K1 #H1KL1 #H2KL1
-elim (fqus_inv_gen … H) -H
-[ #H elim (lpx_lleq_fqup_trans … H … H1KL1 H2KL1) -L1
- /3 width=4 by fqup_fqus, ex3_intro/
-| * #HG #HL #HT destruct /2 width=4 by ex3_intro/
-]
-qed-.
-
-fact lreq_lpx_trans_lleq_aux: ∀h,o,G,L1,L0,l,k. L1 ⩬[l, k] L0 → k = ∞ →
- ∀L2. ⦃G, L0⦄ ⊢ ➡[h, o] L2 →
- ∃∃L. L ⩬[l, k] L2 & ⦃G, L1⦄ ⊢ ➡[h, o] L &
- (∀T. L0 ≡[T, l] L2 ↔ L1 ≡[T, l] L).
-#h #o #G #L1 #L0 #l #k #H elim H -L1 -L0 -l -k
-[ #l #k #_ #L2 #H >(lpx_inv_atom1 … H) -H
- /3 width=5 by ex3_intro, conj/
-| #I1 #I0 #L1 #L0 #V1 #V0 #_ #_ #Hm destruct
-| #I #L1 #L0 #V1 #k #HL10 #IHL10 #Hm #Y #H
- elim (lpx_inv_pair1 … H) -H #L2 #V2 #HL02 #HV02 #H destruct
- lapply (ysucc_inv_Y_dx … Hm) -Hm #Hm
- elim (IHL10 … HL02) // -IHL10 -HL02 #L #HL2 #HL1 #IH
- @(ex3_intro … (L.ⓑ{I}V2)) /3 width=3 by lpx_pair, lreq_cpx_trans, lreq_pair/
- #T elim (IH T) #HL0dx #HL0sn
- @conj #H @(lleq_lreq_repl … H) -H /3 width=1 by lreq_sym, lreq_pair_O_Y/
-| #I1 #I0 #L1 #L0 #V1 #V0 #l #k #HL10 #IHL10 #Hm #Y #H
- elim (lpx_inv_pair1 … H) -H #L2 #V2 #HL02 #HV02 #H destruct
- elim (IHL10 … HL02) // -IHL10 -HL02 #L #HL2 #HL1 #IH
- @(ex3_intro … (L.ⓑ{I1}V1)) /3 width=1 by lpx_pair, lreq_succ/
- #T elim (IH T) #HL0dx #HL0sn
- @conj #H @(lleq_lreq_repl … H) -H /3 width=1 by lreq_sym, lreq_succ/
-]
-qed-.
-
-lemma lreq_lpx_trans_lleq: ∀h,o,G,L1,L0,l. L1 ⩬[l, ∞] L0 →
- ∀L2. ⦃G, L0⦄ ⊢ ➡[h, o] L2 →
- ∃∃L. L ⩬[l, ∞] L2 & ⦃G, L1⦄ ⊢ ➡[h, o] L &
- (∀T. L0 ≡[T, l] L2 ↔ L1 ≡[T, l] L).
-/2 width=1 by lreq_lpx_trans_lleq_aux/ qed-.
cpg.ma cpg_simple.ma cpg_drops.ma cpg_lsubr.ma
-cpx.ma cpx_simple.ma cpx_drops.ma cpx_fqus.ma cpx_lsubr.ma cpx_lfxs.ma
+cpx.ma cpx_simple.ma cpx_drops.ma cpx_fqus.ma cpx_lsubr.ma cpx_lfxs.ma cpx_lfdeq.ma
lfpx.ma lfpx_length.ma lfpx_drops.ma lfpx_fqup.ma lfpx_frees.ma lfpx_lfdeq.ma lfpx_aaa.ma lfpx_lfpx.ma
cnx.ma cnx_simple.ma cnx_drops.ma cnx_cnx.ma
cpm.ma cpm_simple.ma cpm_drops.ma cpm_lsubr.ma cpm_lfxs.ma cpm_cpx.ma
cpr.ma cpr_drops.ma
lfpr.ma lfpr_length.ma lfpr_drops.ma lfpr_fqup.ma lfpr_frees.ma lfpr_aaa.ma lfpr_lfpx.ma lfpr_lfpr.ma
fpb.ma
+fpbq.ma fpbq_aaa.ma
]
class "cyan"
[ { "rt-transition" * } {
- [ { "parallel qrst-transition" * } {
-(* [ "fpbq ( ⦃?,?,?⦄ ≽[?,?] ⦃?,?,?⦄ )" "fpbq_alt ( ⦃?,?,?⦄ ≽≽[?,?] ⦃?,?,?⦄ )" + "fpbq_aaa" * ] *)
+ [ { "parallel rst-transition" * } {
+ [ "fpbq ( ⦃?,?,?⦄ ≽[?] ⦃?,?,?⦄ )" "fpbq_aaa" * ]
[ "fpb ( ⦃?,?,?⦄ ≻[?,?] ⦃?,?,?⦄ )" (* "fpb_lleq" + "fpb_fleq" *) * ]
}
]
-(*
- [ { "context-sensitive rt-reduction" * } {
- [ "lpx_lleq" * ]
- [ "cpx_lleq" * ]
- }
- ]
-*)
[ { "t-bound context-sensitive rt-transition" * } {
[ "lfpr ( ⦃?,?⦄ ⊢ ➡[?,?] ? )" "lfpr_length" + "lfpr_drops" + "lfpr_fqup" + "lfpr_frees" + "lfpr_aaa" + "lfpr_lfpx" + "lfpr_lfpr" * ]
[ "cpr ( ⦃?,?⦄ ⊢ ? ➡[?] ? )" "cpr_drops" * ]
}
]
[ { "uncounted context-sensitive rt-transition" * } {
- [ "cnx ( ⦃?,?⦄ ⊢ ⬈[?,?] 𝐍⦃?⦄ )" "cnx_simple" + "cnx_drops" *+ "cnx_cnx" ]
+ [ "cnx ( ⦃?,?⦄ ⊢ ⬈[?,?] 𝐍⦃?⦄ )" "cnx_simple" + "cnx_drops" + "cnx_cnx" * ]
[ "lfpx ( ⦃?,?⦄ ⊢ ⬈[?,?] ? )" "lfpx_length" + "lfpx_drops" + "lfpx_fqup" + "lfpx_frees" + "lfpx_lfdeq" + "lfpx_aaa" + "lfpx_lfpx" * ]
[ "cpx ( ⦃?,?⦄ ⊢ ? ⬈[?] ? )" "cpx_simple" + "cpx_drops" + "cpx_fqus" + "cpx_lsubr" + "cpx_lfxs" * ]
}