X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fdama%2Fdama%2Fbishop_set.ma;h=d69bb2732b3f1a4a023aa21bdd54fc4635c3b7b7;hb=c231702a57076acf0c161cdb4799bf83158175f0;hp=64ae4495d14c11fb1d215ea8d709f8b19c2db2d2;hpb=80ea6f314e89d9d280338c41860cb04949319629;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 64ae4495d..d69bb2732 100644 --- a/helm/software/matita/contribs/dama/dama/bishop_set.ma +++ b/helm/software/matita/contribs/dama/dama/bishop_set.ma @@ -27,12 +27,12 @@ interpretation "bishop set apartness" '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)); -[1: unfold; cases E; simplify; clear E; intros (x); unfold; - 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 H); lapply (cTf ? ? z H) as H1; cases H1; clear Axy H1; - [left; left|right; left|right; right|left; right] assumption] +[1: intro x; simplify; intro H; cases H; clear H; + apply (exc_coreflexive x H1); +|2: intros 3 (x y H); simplify in H ⊢ %; cases H; [right|left]assumption; +|3: intros 4 (x y z H); simplify in H ⊢ %; cases H; clear H; + [ cases (exc_cotransitive x y z H1); [left;left|right;left] assumption; + | cases (exc_cotransitive y x z H1); [right;right|left;right] assumption;]] qed. (* Definition 2.2 (2) *) @@ -55,10 +55,10 @@ intros 6 (E x y z Exy Eyz); intro Axy; cases (bs_cotransitive ???y Axy); [apply Exy|apply Eyz] assumption. qed. -coercion cic:/matita/dama/bishop_set/bishop_set_of_ordered_set.con. +coercion bishop_set_of_ordered_set. lemma le_antisymmetric: - ∀E:ordered_set.antisymmetric E (le E) (eq E). + ∀E:ordered_set.antisymmetric E (le (os_l E)) (eq E). intros 5 (E x y Lxy Lyx); intro H; cases H; [apply Lxy;|apply Lyx] assumption; qed. @@ -67,28 +67,6 @@ 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. - definition bs_subset ≝ λO:bishop_set.λP,Q:O→Prop.∀x:O.P x → Q x. interpretation "bishop set subset" 'subseteq a b = (bs_subset _ a b). @@ -105,5 +83,7 @@ qed. notation "s 2 \atop \neq" non associative with precedence 90 for @{ 'square_bs $s }. -interpretation "bishop set square" 'square x = (square_bishop_set x). +notation > "s 'squareB'" non associative with precedence 90 + for @{ 'squareB $s }. +interpretation "bishop set square" 'squareB x = (square_bishop_set x). interpretation "bishop set square" 'square_bs x = (square_bishop_set x). \ No newline at end of file