]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/grammar/lenv_px_sn.ma
- we commit just the components before "reducibility"
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / grammar / lenv_px_sn.ma
index e3f661cfc50d0ea751bc7dccbef6f81c00374fe0..e746586aef863d8c143d3b72f57d4e3b89dfdf23 100644 (file)
@@ -23,15 +23,14 @@ inductive lpx_sn (R:lenv→relation term): relation lenv ≝
                lpx_sn R (K1. ⓑ{I} V1) (K2. ⓑ{I} V2)
 .
 
-definition lpx_sn_confluent: predicate (lenv→relation term) ≝ λR.
-                             ∀L0,T0,T1. R L0 T0 T1 → ∀T2. R L0 T0 T2 →
-                             ∀L1. lpx_sn R L0 L1 → ∀L2. lpx_sn R L0 L2 →
-                             ∃∃T. R L1 T1 T & R L2 T2 T.
-
-definition lpx_sn_transitive: predicate (lenv→relation term) ≝ λR.
-                              ∀L1,T1,T. R L1 T1 T → ∀L2. lpx_sn R L1 L2 →
-                              ∀T2. R L2 T T2 → R L1 T1 T2.
+definition lpx_sn_confluent: relation (lenv→relation term) ≝ λR1,R2.
+                             ∀L0,T0,T1. R1 L0 T0 T1 → ∀T2. R2 L0 T0 T2 →
+                             ∀L1. lpx_sn R1 L0 L1 → ∀L2. lpx_sn R2 L0 L2 →
+                             ∃∃T. R1 L1 T1 T & R2 L2 T2 T.
 
+definition lpx_sn_transitive: relation (lenv→relation term) ≝ λR1,R2.
+                              ∀L1,T1,T. R1 L1 T1 T → ∀L2. lpx_sn R1 L1 L2 →
+                              ∀T2. R2 L2 T T2 → R1 L1 T1 T2.
 
 (* Basic inversion lemmas ***************************************************)
 
@@ -85,6 +84,32 @@ 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).
@@ -97,13 +122,15 @@ lemma lpx_sn_append: ∀R. l_appendable_sn R →
 #R #HR #K1 #K2 #H elim H -K1 -K2 // /3 width=1/
 qed-.
 
-lemma lpx_sn_trans: ∀R. lpx_sn_transitive R → Transitive … (lpx_sn R).
+(* Advanced properties ******************************************************)
+
+lemma lpx_sn_trans: ∀R. lpx_sn_transitive R R → Transitive … (lpx_sn R).
 #R #HR #L1 #L #H elim H -L1 -L //
 #I #L1 #L #V1 #V #HL1 #HV1 #IHL1 #X #H
 elim (lpx_sn_inv_pair1 … H) -H #L2 #V2 #HL2 #HV2 #H destruct /3 width=5/
 qed-.
 
-lemma lpx_sn_conf: ∀R. lpx_sn_confluent R → confluent … (lpx_sn R).
+lemma lpx_sn_conf: ∀R. lpx_sn_confluent R → confluent … (lpx_sn R).
 #R #HR #L0 @(f_ind … length … L0) -L0 #n #IH *
 [ #_ #X1 #H1 #X2 #H2 -n
   >(lpx_sn_inv_atom1 … H1) -X1