X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Fbasics%2Fstar.ma;h=04d2d7317e4416b7031d18f5125f2ce78eeb72bc;hb=2405e80ecd3a66780ef1d27066a648330aacf1b0;hp=8fd32c2e1f3092ade0b5a40208b3afb2641c2f54;hpb=1185204d6a1e634a107fac71a45c9f87f49ccc31;p=helm.git diff --git a/matita/matita/lib/basics/star.ma b/matita/matita/lib/basics/star.ma index 8fd32c2e1..04d2d7317 100644 --- a/matita/matita/lib/basics/star.ma +++ b/matita/matita/lib/basics/star.ma @@ -88,6 +88,56 @@ theorem star_inv: ∀A,R. #H (elim H) /2/ normalize #c #d #H1 #H2 #H3 @(trans_star … H3) /2/ qed. +lemma star_decomp_l : + ∀A,R,x,y.star A R x y → x = y ∨ ∃z.R x z ∧ star A R z y. +#A #R #x #y #Hstar elim Hstar +[ #b #c #Hleft #Hright * + [ #H1 %2 @(ex_intro ?? c) % // + | * #x0 * #H1 #H2 %2 @(ex_intro ?? x0) % /2/ ] +| /2/ ] +qed. + +(* right associative version of star *) +inductive starl (A:Type[0]) (R:relation A) : A → A → Prop ≝ + |injl: ∀a,b,c.R a b → starl A R b c → starl A R a c + |refll: ∀a.starl A R a a. + +lemma starl_comp : ∀A,R,a,b,c. + starl A R a b → R b c → starl A R a c. +#A #R #a #b #c #Hstar elim Hstar + [#a1 #b1 #c1 #Rab #sbc #Hind #a1 @(injl … Rab) @Hind // + |#a1 #Rac @(injl … Rac) // + ] +qed. + +lemma star_compl : ∀A,R,a,b,c. + R a b → star A R b c → star A R a c. +#A #R #a #b #c #Rab #Hstar elim Hstar + [#b1 #c1 #sbb1 #Rb1c1 #Hind @(inj … Rb1c1) @Hind + |@(inj … Rab) // + ] +qed. + +lemma star_to_starl: ∀A,R,a,b.star A R a b → starl A R a b. +#A #R #a #b #Hs elim Hs // +#d #c #sad #Rdc #sad @(starl_comp … Rdc) // +qed. + +lemma starl_to_star: ∀A,R,a,b.starl A R a b → star A R a b. +#A #R #a #b #Hs elim Hs // -Hs -b -a +#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. + (* RC and star *) lemma TC_to_star: ∀A,R,a,b.TC A R a b → star A R a b.