λG,L,T1,T2. ∃c. ⦃G, L⦄ ⊢ T1 ⬈[c, h] T2.
interpretation
- "uncounted context-sensitive parallel reduction (term)"
+ "uncounted context-sensitive parallel rt-transition (term)"
'PRedTy h G L T1 T2 = (cpx h G L T1 T2).
(* Basic properties *********************************************************)
lemma lpr_fwd_length: ∀G,L1,L2. ⦃G, L1⦄ ⊢ ➡ L2 → |L1| = |L2|.
/2 width=2 by lpx_sn_fwd_length/ qed-.
+lemma lpr_lpx: ∀h,o,G,L1,L2. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L1⦄ ⊢ ➡[h, o] L2.
+#h #o #G #L1 #L2 #H elim H -L1 -L2 /3 width=1 by lpx_pair, cpr_cpx/
+qed.
+
(* Basic_1: removed theorems 3: wcpr0_getl wcpr0_getl_back
pr0_subst1_back
*)
(Y1 = ⋆ ∧ Y2 = ⋆) ∨
∃∃I,L1,L2,V1,V2. L1 ≡[#i] L2 &
Y1 = L1.ⓑ{I}V1 & Y2 = L2.ⓑ{I}V2.
-#Y1 #Y2 #i #H elim (lfxs_inv_lref … H) -H *
-/3 width=8 by ex3_5_intro, or_introl, or_intror, conj/
-qed-.
+/2 width=1 by lfxs_inv_lref/ qed-.
lemma lfeq_inv_bind: ∀p,I,L1,L2,V,T. L1 ≡[ⓑ{p,I}V.T] L2 →
L1 ≡[V] L2 ∧ L1.ⓑ{I}V ≡[T] L2.ⓑ{I}V.
-#p #I #L1 #L2 #V #T #H elim (lfxs_inv_bind … H) -H /2 width=3 by conj/
-qed-.
+/2 width=2 by lfxs_inv_bind/ qed-.
lemma lfeq_inv_flat: ∀I,L1,L2,V,T. L1 ≡[ⓕ{I}V.T] L2 →
L1 ≡[V] L2 ∧ L1 ≡[T] L2.
-#I #L1 #L2 #V #T #H elim (lfxs_inv_flat … H) -H /2 width=3 by conj/
-qed-.
+/2 width=2 by lfxs_inv_flat/ qed-.
(* Advanced inversion lemmas ************************************************)