1 lemma lpx_sn_append: ∀R. l_appendable_sn R →
2 ∀K1,K2. lpx_sn R K1 K2 → ∀L1,L2. lpx_sn R L1 L2 →
3 lpx_sn R (L1 @@ K1) (L2 @@ K2).
4 #R #HR #K1 #K2 #H elim H -K1 -K2 /3 width=1 by lpx_sn_pair/
7 (* Advanced forward lemmas **************************************************)
9 lemma lpx_sn_fwd_append1: ∀R,L1,K1,L. lpx_sn R (K1 @@ L1) L →
10 ∃∃K2,L2. lpx_sn R K1 K2 & L = K2 @@ L2.
13 @(ex2_2_intro … K2 (⋆)) // (* explicit constructor, /2 width=4/ does not work *)
14 | #L1 #I #V1 #IH #K1 #X #H
15 elim (lpx_sn_inv_pair1 … H) -H #L #V2 #H1 #HV12 #H destruct
16 elim (IH … H1) -IH -H1 #K2 #L2 #HK12 #H destruct
17 @(ex2_2_intro … (L2.ⓑ{I}V2) HK12) // (* explicit constructor, /2 width=4/ does not work *)
21 lemma lpx_sn_fwd_append2: ∀R,L2,K2,L. lpx_sn R L (K2 @@ L2) →
22 ∃∃K1,L1. lpx_sn R K1 K2 & L = K1 @@ L1.
25 @(ex2_2_intro … K1 (⋆)) // (**) (* explicit constructor, /2 width=4/ does not work *)
26 | #L2 #I #V2 #IH #K2 #X #H
27 elim (lpx_sn_inv_pair2 … H) -H #L #V1 #H1 #HV12 #H destruct
28 elim (IH … H1) -IH -H1 #K1 #L1 #HK12 #H destruct
29 @(ex2_2_intro … (L1.ⓑ{I}V1) HK12) // (* explicit constructor, /2 width=4/ does not work *)