]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2A/etc/append/lpx_sn.etc
milestone update in ground_2 and basic_2A
[helm.git] / matita / matita / contribs / lambdadelta / basic_2A / etc / append / lpx_sn.etc
diff --git a/matita/matita/contribs/lambdadelta/basic_2A/etc/append/lpx_sn.etc b/matita/matita/contribs/lambdadelta/basic_2A/etc/append/lpx_sn.etc
new file mode 100644 (file)
index 0000000..d399e0a
--- /dev/null
@@ -0,0 +1,31 @@
+lemma lpx_sn_append: ∀R. l_appendable_sn R →
+                     ∀K1,K2. lpx_sn R K1 K2 → ∀L1,L2. lpx_sn R L1 L2 →
+                     lpx_sn R (L1 @@ K1) (L2 @@ K2).
+#R #HR #K1 #K2 #H elim H -K1 -K2 /3 width=1 by lpx_sn_pair/
+qed-.
+
+(* Advanced forward lemmas **************************************************)
+
+lemma lpx_sn_fwd_append1: ∀R,L1,K1,L. lpx_sn R (K1 @@ L1) L →
+                          ∃∃K2,L2. lpx_sn R K1 K2 &  L = K2 @@ L2.
+#R #L1 elim L1 -L1
+[ #K1 #K2 #HK12
+  @(ex2_2_intro … K2 (⋆)) // (* explicit constructor, /2 width=4/ does not work *)
+| #L1 #I #V1 #IH #K1 #X #H
+  elim (lpx_sn_inv_pair1 … H) -H #L #V2 #H1 #HV12 #H destruct
+  elim (IH … H1) -IH -H1 #K2 #L2 #HK12 #H destruct
+  @(ex2_2_intro … (L2.ⓑ{I}V2) HK12) // (* explicit constructor, /2 width=4/ does not work *)
+]
+qed-.
+
+lemma lpx_sn_fwd_append2: ∀R,L2,K2,L. lpx_sn R L (K2 @@ L2) →
+                          ∃∃K1,L1. lpx_sn R K1 K2 & L = K1 @@ L1.
+#R #L2 elim L2 -L2
+[ #K2 #K1 #HK12
+  @(ex2_2_intro … K1 (⋆)) // (**) (* explicit constructor, /2 width=4/ does not work *)
+| #L2 #I #V2 #IH #K2 #X #H
+  elim (lpx_sn_inv_pair2 … H) -H #L #V1 #H1 #HV12 #H destruct
+  elim (IH … H1) -IH -H1 #K1 #L1 #HK12 #H destruct
+  @(ex2_2_intro … (L1.ⓑ{I}V1) HK12) // (* explicit constructor, /2 width=4/ does not work *)
+]
+qed-.