X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fdama%2Fdama%2Fbishop_set.ma;h=64ae4495d14c11fb1d215ea8d709f8b19c2db2d2;hb=80ea6f314e89d9d280338c41860cb04949319629;hp=8cb77ec24289c633e6adcb535394156e0f053224;hpb=ca41435a6021292ccba239aa173651c0be705b45;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 8cb77ec24..64ae4495d 100644 --- a/helm/software/matita/contribs/dama/dama/bishop_set.ma +++ b/helm/software/matita/contribs/dama/dama/bishop_set.ma @@ -23,9 +23,6 @@ record bishop_set: Type ≝ { bs_cotransitive: cotransitive ? bs_apart }. -notation "hvbox(a break # b)" non associative with precedence 45 - for @{ 'apart $a $b}. - interpretation "bishop set apartness" 'apart x y = (bs_apart _ x y). definition bishop_set_of_ordered_set: ordered_set → bishop_set. @@ -41,9 +38,6 @@ qed. (* Definition 2.2 (2) *) definition eq ≝ λA:bishop_set.λa,b:A. ¬ (a # b). -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). lemma eq_reflexive:∀E:bishop_set. reflexive ? (eq E). @@ -97,7 +91,7 @@ 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). +interpretation "bishop set subset" 'subseteq a b = (bs_subset _ a b). definition square_bishop_set : bishop_set → bishop_set. intro S; apply (mk_bishop_set (S × S));