X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fcontribs%2Fdama%2Fdama%2Fconstructive_higher_order_relations.ma;fp=matita%2Fcontribs%2Fdama%2Fdama%2Fconstructive_higher_order_relations.ma;h=8d195396cc455efc20ec7c0875b7e4264d2e6d5e;hp=0000000000000000000000000000000000000000;hb=f61af501fb4608cc4fb062a0864c774e677f0d76;hpb=58ae1809c352e71e7b5530dc41e2bfc834e1aef1 diff --git a/matita/contribs/dama/dama/constructive_higher_order_relations.ma b/matita/contribs/dama/dama/constructive_higher_order_relations.ma new file mode 100644 index 000000000..8d195396c --- /dev/null +++ b/matita/contribs/dama/dama/constructive_higher_order_relations.ma @@ -0,0 +1,51 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + + + +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. + +definition coreflexive ≝ λC:Type.λlt:C→C→Type. ∀x:C. ¬ (lt x x). + +definition antisymmetric ≝ + λC:Type.λle:C→C→Type.λeq:C→C→Type.∀x,y:C.le x y → le y x → eq x y. + +definition symmetric ≝ + λC:Type.λle:C→C→Type.∀x,y:C.le x y → le y x. + +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. + +