X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Fdama%2Fbishop_set.ma;h=a532b22014f23f6b0282240759b1152634ea029c;hb=456a157eba1428fd4ec02713e60ac2b653a0e0b0;hp=d69bb2732b3f1a4a023aa21bdd54fc4635c3b7b7;hpb=6fbeff97e37927fd95b3aee3eb23b4309fc465c4;p=helm.git diff --git a/helm/software/matita/library/dama/bishop_set.ma b/helm/software/matita/library/dama/bishop_set.ma index d69bb2732..a532b2201 100644 --- a/helm/software/matita/library/dama/bishop_set.ma +++ b/helm/software/matita/library/dama/bishop_set.ma @@ -12,7 +12,7 @@ (* *) (**************************************************************************) -include "ordered_set.ma". +include "dama/ordered_set.ma". (* Definition 2.2, setoid *) record bishop_set: Type ≝ { @@ -23,7 +23,7 @@ record bishop_set: Type ≝ { bs_cotransitive: cotransitive ? bs_apart }. -interpretation "bishop set apartness" '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)); @@ -38,7 +38,7 @@ qed. (* Definition 2.2 (2) *) definition eq ≝ λA:bishop_set.λa,b:A. ¬ (a # b). -interpretation "Bishop set alikeness" 'napart a b = (eq _ 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; @@ -69,7 +69,7 @@ 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). +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)); @@ -86,4 +86,4 @@ notation "s 2 \atop \neq" non associative with precedence 90 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 +interpretation "bishop set square" 'square_bs x = (square_bishop_set x).