X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fdama%2Fdama%2Fbishop_set.ma;h=9fee3d9fd3c0cdd3e640ee9054a84b89bf4d59ce;hb=7e33e23e18dc5d008b3b3dc0052aa4d7b236415e;hp=7ac1b83f01a42a8a4acab41fa2b06ff6f77f6c4e;hpb=9eabe046c1182960de8cfdba96c5414224e3a61e;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 7ac1b83f0..9fee3d9fd 100644 --- a/helm/software/matita/contribs/dama/dama/bishop_set.ma +++ b/helm/software/matita/contribs/dama/dama/bishop_set.ma @@ -23,10 +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 = (bs_apart _ x y). +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)); @@ -41,7 +41,7 @@ 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 \approx b)" non associative with precedence 45 for @{ 'napart $a $b}. interpretation "Bishop set alikeness" 'napart a b = (eq _ a b). @@ -57,7 +57,6 @@ qed. lemma eq_sym:∀E:bishop_set.∀x,y:E.x ≈ y → y ≈ x ≝ eq_sym_. lemma eq_trans_: ∀E:bishop_set.transitive ? (eq E). -(* bug. intros k deve fare whd quanto basta *) intros 6 (E x y z Exy Eyz); intro Axy; cases (bs_cotransitive ???y Axy); [apply Exy|apply Eyz] assumption. qed. @@ -95,3 +94,22 @@ 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" 'subset a b = (bs_subset _ a b). + +definition square_bishop_set : bishop_set → bishop_set. +intro S; apply (mk_bishop_set (S × S)); +[1: intros (x y); apply ((fst x # fst y) ∨ (snd x # snd y)); +|2: intro x; simplify; intro; cases H (X X); clear H; apply (bs_coreflexive ?? X); +|3: intros 2 (x y); simplify; intro H; cases H (X X); clear H; [left|right] apply (bs_symmetric ??? X); +|4: intros 3 (x y z); simplify; intro H; cases H (X X); clear H; + [1: cases (bs_cotransitive ??? (fst z) X); [left;left|right;left]assumption; + |2: cases (bs_cotransitive ??? (snd z) X); [left;right|right;right]assumption;]] +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). +interpretation "bishop set square" 'square_bs x = (square_bishop_set x). \ No newline at end of file