X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fdama%2Fexcedence.ma;h=c2a5ffd4fde20049391060de83b85f2cd22cf310;hb=952ced6c96e98fa678c59729d18975f9a376623e;hp=3b2afd5dc982b86c8a8308018b72bc599bc04fc2;hpb=8e7751f97e50bdc18537aac7478d0621d45ab956;p=helm.git diff --git a/matita/dama/excedence.ma b/matita/dama/excedence.ma index 3b2afd5dc..c2a5ffd4f 100644 --- a/matita/dama/excedence.ma +++ b/matita/dama/excedence.ma @@ -21,7 +21,7 @@ include "constructive_higher_order_relations.ma". record excedence : Type ≝ { exc_carr:> Type; - exc_relation: exc_carr → exc_carr → Type; (* Big bug: era in Prop!!! *) + exc_relation: exc_carr → exc_carr → Type; exc_coreflexive: coreflexive ? exc_relation; exc_cotransitive: cotransitive ? exc_relation }. @@ -127,12 +127,6 @@ intros (E a b Lab); cases Lab (LEab Aab); cases Aab (H H); [cases (LEab H)] fold normalize (b ≰ a); assumption; (* BUG *) qed. -(* CSC: lo avevi gia' dimostrato; ho messo apply! *) -theorem le_le_to_eq: ∀E:excedence.∀x,y:E. x ≤ y → y ≤ x → x ≈ y. -apply le_antisymmetric; -qed. - -(* CSC: perche' quel casino: bastava intros; assumption! *) lemma unfold_apart: ∀E:excedence. ∀x,y:E. x ≰ y ∨ y ≰ x → x # y. intros; assumption; qed. @@ -155,4 +149,14 @@ qed. lemma ap_rewr: ∀A:apartness.∀x,z,y:A. x ≈ y → z # y → z # x. intros (A x z y Exy Azy); apply ap_symmetric; apply (ap_rewl ???? Exy); apply ap_symmetric; assumption; -qed. \ No newline at end of file +qed. + +lemma exc_rewl: ∀A:excedence.∀x,z,y:A. x ≈ y → y ≰ z → x ≰ z. +intros (A x z y Exy Ayz); elim (exc_cotransitive ???x Ayz); [2:assumption] +cases Exy; right; assumption; +qed. + +lemma exc_rewr: ∀A:excedence.∀x,z,y:A. x ≈ y → z ≰ y → z ≰ x. +intros (A x z y Exy Azy); elim (exc_cotransitive ???x Azy); [assumption] +elim (Exy); left; assumption; +qed.