X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Fbasics%2Fstar.ma;h=5ac7d91fb624e3b4a872d9c20a8c770a506fd38e;hb=fd0e4ca5e7994d02b0cf923b14969a79f1287dcc;hp=aafb3fa6be9174cde3f2914240066503666fc4ee;hpb=53f874fba5b9c39a788085515a4fefe5d29281da;p=helm.git diff --git a/matita/matita/lib/basics/star.ma b/matita/matita/lib/basics/star.ma index aafb3fa6b..5ac7d91fb 100644 --- a/matita/matita/lib/basics/star.ma +++ b/matita/matita/lib/basics/star.ma @@ -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 *) @@ -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-.