X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fdama%2Fexcedence.ma;h=c2a5ffd4fde20049391060de83b85f2cd22cf310;hb=952ced6c96e98fa678c59729d18975f9a376623e;hp=84a033c3a8ebb071bf2b56abca7b4fa1bbf8a873;hpb=4c1356a13d67f07ef2524a8c1612accd878b810a;p=helm.git diff --git a/matita/dama/excedence.ma b/matita/dama/excedence.ma index 84a033c3a..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 → Prop; + exc_relation: exc_carr → exc_carr → Type; exc_coreflexive: coreflexive ? exc_relation; exc_cotransitive: cotransitive ? exc_relation }. @@ -63,7 +63,7 @@ intros (E); apply (mk_apartness E (apart E)); intros (H1); apply (H x); cases H1; assumption; |2: unfold; intros(x y H); cases H; clear H; [right|left] assumption; |3: intros (E); unfold; cases E (T f _ cTf); simplify; intros (x y z Axy); - cases Axy (H); lapply (cTf ? ? z H) as H1; cases H1; clear Axy H1; + cases Axy (H H); lapply (cTf ? ? z H) as H1; cases H1; clear Axy H1; [left; left|right; left|right; right|left; right] assumption] qed. @@ -79,21 +79,23 @@ lemma eq_reflexive:∀E. reflexive ? (eq E). intros (E); unfold; intros (x); apply ap_coreflexive; qed. -lemma eq_symmetric:∀E.symmetric ? (eq E). +lemma eq_sym_:∀E.symmetric ? (eq E). intros (E); unfold; intros (x y Exy); unfold; unfold; intros (Ayx); apply Exy; apply ap_symmetric; assumption; qed. -lemma eq_symmetric_:∀E:apartness.∀x,y:E.x ≈ y → y ≈ x := eq_symmetric. +lemma eq_sym:∀E:apartness.∀x,y:E.x ≈ y → y ≈ x := eq_sym_. -coercion cic:/matita/excedence/eq_symmetric_.con. +coercion cic:/matita/excedence/eq_sym.con. -lemma eq_transitive: ∀E.transitive ? (eq E). +lemma eq_trans_: ∀E.transitive ? (eq E). (* bug. intros k deve fare whd quanto basta *) intros 6 (E x y z Exy Eyz); intro Axy; cases (ap_cotransitive ???y Axy); [apply Exy|apply Eyz] assumption. qed. +lemma eq_trans:∀E:apartness.∀x,y,z:E.x ≈ y → y ≈ z → x ≈ z ≝ eq_trans_. + (* BUG: vedere se ricapita *) lemma le_antisymmetric: ∀E.antisymmetric ? (le E) (eq ?). intros 5 (E x y Lxy Lyx); intro H; @@ -125,12 +127,8 @@ intros (E a b Lab); cases Lab (LEab Aab); cases Aab (H H); [cases (LEab H)] fold normalize (b ≰ a); assumption; (* BUG *) qed. -theorem le_le_to_eq: ∀E:excedence.∀x,y:E. x ≤ y → y ≤ x → x ≈ y. -intros 6 (E x y L1 L2 H); cases H; [apply (L1 H1)|apply (L2 H1)] -qed. - lemma unfold_apart: ∀E:excedence. ∀x,y:E. x ≰ y ∨ y ≰ x → x # y. -unfold apart_of_excedence; unfold apart; simplify; intros; assumption; +intros; assumption; qed. lemma le_rewl: ∀E:excedence.∀z,y,x:E. x ≈ y → x ≤ z → y ≤ z. @@ -152,3 +150,13 @@ 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. + +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.