]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/ground_2/lib/star.ma
- former llpx_sn an lleq reactivated as lfxs and lfeq
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / lib / star.ma
index 50fda90652eb9f0c4773e230c7860aed86ae53a0..8849d759beeaf774379ecc33015d548353b6e913 100644 (file)
@@ -17,10 +17,22 @@ include "ground_2/xoa/xoa_props.ma".
 
 (* PROPERTIES OF RELATIONS **************************************************)
 
+definition relation5 : Type[0] → Type[0] → Type[0] → Type[0] → Type[0] → Type[0]
+≝ λA,B,C,D,E.A→B→C→D→E→Prop.
+
+definition relation6 : Type[0] → Type[0] → Type[0] → Type[0] → Type[0] → Type[0] → Type[0]
+≝ λA,B,C,D,E,F.A→B→C→D→E→F→Prop.
+
 definition Decidable: Prop → Prop ≝ λR. R ∨ (R → ⊥).
 
 definition Transitive: ∀A. ∀R: relation A. Prop ≝ λA,R.
                        ∀a1,a0. R a1 a0 → ∀a2. R a0 a2 → R a1 a2.
+                       
+definition left_cancellable: ∀A. ∀R: relation A. Prop ≝ λA,R.
+                             ∀a0,a1. R a0 a1 → ∀a2. R a0 a2 → R a1 a2.
+
+definition right_cancellable: ∀A. ∀R: relation A. Prop ≝ λA,R.
+                              ∀a1,a0. R a1 a0 → ∀a2. R a2 a0 → R a1 a2.
 
 definition confluent2: ∀A. ∀R1,R2: relation A. Prop ≝ λA,R1,R2.
                        ∀a0,a1. R1 a0 a1 → ∀a2. R2 a0 a2 →
@@ -178,6 +190,28 @@ lemma s_r_trans_LTC2: ∀A,B,S,R. s_rs_transitive A B S R → s_r_transitive A B
 #A #B #S #R #HSR #L2 #T1 #T2 #HT12 #L1 #H @(TC_ind_dx … L1 H) -L1 /3 width=3 by inj/
 qed-.
 
+lemma s_r_to_s_rs_trans: ∀A,B,S,R. s_r_transitive A B (LTC … S) R →
+                         s_rs_transitive A B S R.
+#A #B #S #R #HSR #L2 #T1 #T2 #HL2 #L1 #HT1
+elim (TC_idem … (S L1) …  T1 T2)
+#_ #H @H @HSR //
+qed-.
+
+lemma s_rs_to_s_r_trans: ∀A,B,S,R. s_rs_transitive A B S R →
+                         s_r_transitive A B (LTC … S) R.
+#A #B #S #R #HSR #L2 #T1 #T2 #HL2 #L1 #HT1
+elim (TC_idem … (S L1) …  T1 T2)
+#H #_ @H @HSR //
+qed-.
+
+lemma s_rs_trans_TC1: ∀A,B,S,R. s_rs_transitive A B S R →
+                      s_rs_transitive A B (LTC … S) R.
+#A #B #S #R #HSR #L2 #T1 #T2 #HL2 #L1 #HT1
+elim (TC_idem … (S L1) …  T1 T2)
+elim (TC_idem … (S L2) …  T1 T2)
+#_ #H1 #H2 #_ @H2 @HSR /3 width=3 by/
+qed-.
+
 (* relations on unboxed pairs ***********************************************)
 
 lemma bi_TC_strip: ∀A,B,R. bi_confluent A B R →