X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fdama%2Fconstructive_higher_order_relations.ma;h=8d195396cc455efc20ec7c0875b7e4264d2e6d5e;hb=954ed2eaf305a60d7e046206472bc0397a421ad2;hp=b66ba684309c47c2cb606f2607e4db506b374f31;hpb=2ab6644dd2dff227ac1bf335df7ff0d244ebe8dc;p=helm.git diff --git a/helm/software/matita/dama/constructive_higher_order_relations.ma b/helm/software/matita/dama/constructive_higher_order_relations.ma index b66ba6843..8d195396c 100644 --- a/helm/software/matita/dama/constructive_higher_order_relations.ma +++ b/helm/software/matita/dama/constructive_higher_order_relations.ma @@ -12,9 +12,10 @@ (* *) (**************************************************************************) -set "baseuri" "cic:/matita/constructive_higher_order_relations". + include "constructive_connectives.ma". +include "higher_order_defs/relations.ma". definition cotransitive ≝ λC:Type.λlt:C→C→Type.∀x,y,z:C. lt x y → lt x z ∨ lt z y. @@ -29,3 +30,22 @@ definition symmetric ≝ definition transitive ≝ λC:Type.λle:C→C→Type.∀x,y,z:C.le x y → le y z → le x z. + +definition associative ≝ + λC:Type.λop:C→C→C.λeq:C→C→Type.∀x,y,z. eq (op x (op y z)) (op (op x y) z). + +definition commutative ≝ + λC:Type.λop:C→C→C.λeq:C→C→Type.∀x,y. eq (op x y) (op y x). + +alias id "antisymmetric" = "cic:/matita/higher_order_defs/relations/antisymmetric.con". +theorem antisimmetric_to_cotransitive_to_transitive: + ∀C:Type.∀le:C→C→Prop. antisymmetric ? le → cotransitive ? le → transitive ? le. +intros (T f Af cT); unfold transitive; intros (x y z fxy fyz); +lapply (cT ??z fxy) as H; cases H; [assumption] cases (Af ? ? fyz H1); +qed. + +lemma Or_symmetric: symmetric ? Or. +unfold; intros (x y H); cases H; [right|left] assumption; +qed. + +