]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/lib/basics/star.ma
lambda: some refactoring + support for subsets of subterms started
[helm.git] / matita / matita / lib / basics / star.ma
index aafb3fa6be9174cde3f2914240066503666fc4ee..36aee8fe45465c016a6ca3709046a0c6fefd11c7 100644 (file)
@@ -122,15 +122,25 @@ lemma starl_to_star: ∀A,R,a,b.starl A R a b → star A R a b.
 #a #b #c #Rab #sbc #sbc @(star_compl … Rab) //
 qed.
 
-lemma star_ind_l : 
-  ∀A:Type[0].∀R:relation A.∀Q:A → A → Prop.
-  (∀a.Q a a) → 
-  (∀a,b,c.R a b → star A R b c → Q b c → Q a c) → 
-  ∀a,b.star A R a b → Q a b.
-#A #R #Q #H1 #H2 #a #b #H0 
-elim (star_to_starl ???? H0) // -H0 -b -a 
-#a #b #c #Rab #slbc @H2 // @starl_to_star //
-qed.  
+fact star_ind_l_aux: ∀A,R,a2. ∀P:predicate A.
+                     P a2 →
+                     (∀a1,a. R a1 a → star … R a a2 → P a → P a1) →
+                     ∀a1,a. star … R a1 a → a = a2 → P a1.
+#A #R #a2 #P #H1 #H2 #a1 #a #Ha1
+elim (star_to_starl ???? Ha1) -a1 -a
+[ #a #b #c #Hab #Hbc #IH #H destruct /3 width=4/
+| #a #H destruct /2 width=1/
+]
+qed-.
+
+(* imporeved version of star_ind_l with "left_parameter" *)
+lemma star_ind_l: ∀A,R,a2. ∀P:predicate A.
+                  P a2 →
+                  (∀a1,a. R a1 a → star … R a a2 → P a → P a1) →
+                  ∀a1. star … R a1 a2 → P a1.
+#A #R #a2 #P #H1 #H2 #a1 #Ha12
+@(star_ind_l_aux … H1 H2 … Ha12) //
+qed.
 
 (* RC and star *)
 
@@ -263,7 +273,7 @@ lemma bi_TC_strap: ∀A,B. ∀R:bi_relation A B. ∀a1,a,a2,b1,b,b2.
                    R a1 b1 a b → bi_TC … R a b a2 b2 → bi_TC … R a1 b1 a2 b2.
 #A #B #R #a1 #a #a2 #b1 #b #b2 #HR #H elim H -a2 -b2 /2 width=4/ /3 width=4/
 qed.
-
+                       
 lemma bi_TC_reflexive: ∀A,B,R. bi_reflexive A B R →
                        bi_reflexive A B (bi_TC … R).
 /2 width=1/ qed.
@@ -411,3 +421,21 @@ lemma bi_star_ind_dx: ∀A,B,R,a2,b2. ∀P:relation2 A B. P a2 b2 →
 | #H12 @(bi_TC_ind_dx ?????????? H12) -a1 -b1 /2 width=5/ -H /3 width=5/
 ]
 qed-.
+
+(* ************ confluence of star *****************)
+
+lemma star_strip: ∀A,R. confluent A R →
+                  ∀a0,a1. star … R a0 a1 → ∀a2. R a0 a2 →
+                  ∃∃a. R a1 a & star … R a2 a.
+#A #R #HR #a0 #a1 #H elim H -a1 /2 width=3/
+#a #a1 #_ #Ha1 #IHa0 #a2 #Ha02
+elim (IHa0 … Ha02) -a0 #a0 #Ha0 #Ha20
+elim (HR … Ha1 … Ha0) -a /3 width=5/
+qed-.
+
+lemma star_confluent: ∀A,R. confluent A R → confluent A (star … R).
+#A #R #HR #a0 #a1 #H elim H -a1 /2 width=3/
+#a #a1 #_ #Ha1 #IHa0 #a2 #Ha02
+elim (IHa0 … Ha02) -a0 #a0 #Ha0 #Ha20
+elim (star_strip … HR … Ha0 … Ha1) -a /3 width=5/
+qed-.