]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/grammar/lpx_sn.ma
- main proposition on lsx finally proved!
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / grammar / lpx_sn.ma
index 6001f40fa962427299acecb860891db8a0e5ff55..fddde2de5bdd33df53e458ea82a49f0ca7172eba 100644 (file)
@@ -12,7 +12,7 @@
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/grammar/lenv_append.ma".
+include "basic_2/grammar/lenv_length.ma".
 
 (* SN POINTWISE EXTENSION OF A CONTEXT-SENSITIVE REALTION FOR TERMS *********)
 
@@ -23,6 +23,12 @@ inductive lpx_sn (R:lenv→relation term): relation lenv ≝
                lpx_sn R (K1. ⓑ{I} V1) (K2. ⓑ{I} V2)
 .
 
+(* Basic properties *********************************************************)
+
+lemma lpx_sn_refl: ∀R. (∀L. reflexive ? (R L)) → reflexive … (lpx_sn R).
+#R #HR #L elim L -L /2 width=1 by lpx_sn_atom, lpx_sn_pair/
+qed-.
+
 (* Basic inversion lemmas ***************************************************)
 
 fact lpx_sn_inv_atom1_aux: ∀R,L1,L2. lpx_sn R L1 L2 → L1 = ⋆ → L2 = ⋆.
@@ -81,41 +87,3 @@ qed-.
 lemma lpx_sn_fwd_length: ∀R,L1,L2. lpx_sn R L1 L2 → |L1| = |L2|.
 #R #L1 #L2 #H elim H -L1 -L2 normalize //
 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-.
-
-(* Basic properties *********************************************************)
-
-lemma lpx_sn_refl: ∀R. (∀L. reflexive ? (R L)) → reflexive … (lpx_sn R).
-#R #HR #L elim L -L /2 width=1 by lpx_sn_atom, lpx_sn_pair/
-qed-.
-
-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-.