X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=inline;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fdama%2Fdama%2Fbishop_set.ma;h=d572820799b26ed2080451ef81083c0064784a86;hb=9f064ffe2834ae91f306b44f79bfbfb68a4631c5;hp=b877e094cd63d26dfc68045622e576a6c31b0d52;hpb=648db01678fac09ddfb3cce900bc72e8a1c420da;p=helm.git diff --git a/helm/software/matita/contribs/dama/dama/bishop_set.ma b/helm/software/matita/contribs/dama/dama/bishop_set.ma index b877e094c..d57282079 100644 --- a/helm/software/matita/contribs/dama/dama/bishop_set.ma +++ b/helm/software/matita/contribs/dama/dama/bishop_set.ma @@ -23,11 +23,10 @@ record bishop_set: Type ≝ { bs_cotransitive: cotransitive ? bs_apart }. -notation "hvbox(a break # b)" non associative with precedence 50 +notation "hvbox(a break # b)" non associative with precedence 45 for @{ 'apart $a $b}. -interpretation "bishop_setapartness" 'apart x y = - (cic:/matita/dama/bishop_set/bs_apart.con _ x y). +interpretation "bishop_setapartness" 'apart x y = (bs_apart _ x y). definition bishop_set_of_ordered_set: ordered_set → bishop_set. intros (E); apply (mk_bishop_set E (λa,b:E. a ≰ b ∨ b ≰ a)); @@ -42,11 +41,10 @@ qed. (* Definition 2.2 (2) *) definition eq ≝ λA:bishop_set.λa,b:A. ¬ (a # b). -notation "hvbox(a break ≈ b)" non associative with precedence 50 +notation "hvbox(a break ≈ b)" non associative with precedence 45 for @{ 'napart $a $b}. -interpretation "Bishop set alikeness" 'napart a b = - (cic:/matita/dama/bishop_set/eq.con _ a b). +interpretation "Bishop set alikeness" 'napart a b = (eq _ a b). lemma eq_reflexive:∀E:bishop_set. reflexive ? (eq E). intros (E); unfold; intros (x); apply bs_coreflexive; @@ -72,3 +70,28 @@ intros 5 (E x y Lxy Lyx); intro H; cases H; [apply Lxy;|apply Lyx] assumption; qed. +lemma le_le_eq: ∀E:ordered_set.∀a,b:E. a ≤ b → b ≤ a → a ≈ b. +intros (E x y L1 L2); intro H; cases H; [apply L1|apply L2] assumption; +qed. + +definition lt ≝ λE:ordered_set.λa,b:E. a ≤ b ∧ a # b. + +interpretation "ordered sets less than" 'lt a b = (lt _ a b). + +lemma lt_coreflexive: ∀E.coreflexive ? (lt E). +intros 2 (E x); intro H; cases H (_ ABS); +apply (bs_coreflexive ? x ABS); +qed. + +lemma lt_transitive: ∀E.transitive ? (lt E). +intros (E); unfold; intros (x y z H1 H2); cases H1 (Lxy Axy); cases H2 (Lyz Ayz); +split; [apply (le_transitive ???? Lxy Lyz)] clear H1 H2; +cases Axy (H1 H1); cases Ayz (H2 H2); [1:cases (Lxy H1)|3:cases (Lyz H2)]clear Axy Ayz; +[1: cases (os_cotransitive ??? y H1) (X X); [cases (Lxy X)|cases (os_coreflexive ?? X)] +|2: cases (os_cotransitive ??? x H2) (X X); [right;assumption|cases (Lxy X)]] +qed. + +theorem lt_to_excess: ∀E:ordered_set.∀a,b:E. (a < b) → (b ≰ a). +intros (E a b Lab); cases Lab (LEab Aab); cases Aab (H H);[cases (LEab H)] +assumption; +qed.