X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Fbasics%2Frelations.ma;h=ed5130955886e557a1e11b8c83b09f7ad649cb48;hb=eae50cc815292d335df1c488a00b39ef98fa5870;hp=afb5d5aae05515a1c9801c45abb9064f2a55a91f;hpb=f6464ba2cffc9936b4d8285604786cd91531e0d0;p=helm.git diff --git a/matita/matita/lib/basics/relations.ma b/matita/matita/lib/basics/relations.ma index afb5d5aae..ed5130955 100644 --- a/matita/matita/lib/basics/relations.ma +++ b/matita/matita/lib/basics/relations.ma @@ -23,6 +23,9 @@ definition relation : Type[0] → Type[0] definition relation2 : Type[0] → Type[0] → Type[0] ≝ λA,B.A→B→Prop. +definition relation3 : Type[0] → Type[0] → Type[0] → Type[0] +≝ λA,B,C.A→B→C→Prop. + definition reflexive: ∀A.∀R :relation A.Prop ≝ λA.λR.∀x:A.R x x. @@ -45,6 +48,14 @@ definition tight_apart: ∀A.∀eq,ap:relation A.Prop definition antisymmetric: ∀A.∀R:relation A.Prop ≝ λA.λR.∀x,y:A. R x y → ¬(R y x). +(* Reflexive closure ************) + +definition RC: ∀A:Type[0]. relation A → relation A ≝ + λA,R,x,y. R … x y ∨ x = y. + +lemma RC_reflexive: ∀A,R. reflexive A (RC … R). +/2 width=1/ qed. + (********** operations **********) definition Rcomp ≝ λA.λR1,R2:relation A.λa1,a2. ∃am.R1 a1 am ∧ R2 am a2. @@ -161,5 +172,12 @@ interpretation "functional extentional equality" definition bi_relation: Type[0] → Type[0] → Type[0] ≝ λA,B.A→B→A→B→Prop. -definition bi_reflexive: ∀A,B. ∀R :bi_relation A B. Prop +definition bi_reflexive: ∀A,B. ∀R:bi_relation A B. Prop ≝ λA,B,R. ∀x,y. R x y x y. + +definition bi_symmetric: ∀A,B. ∀R: bi_relation A B. Prop ≝ λA,B,R. + ∀a1,a2,b1,b2. R a2 b2 a1 b1 → R a1 b1 a2 b2. + +definition bi_transitive: ∀A,B. ∀R: bi_relation A B. Prop ≝ λA,B,R. + ∀a1,a,b1,b. R a1 b1 a b → + ∀a2,b2. R a b a2 b2 → R a1 b1 a2 b2.