]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/dama/dama/bishop_set.ma
more notation
[helm.git] / helm / software / matita / contribs / dama / dama / bishop_set.ma
index 7ac1b83f01a42a8a4acab41fa2b06ff6f77f6c4e..78176d776816c85c513348d356568924e90c77f0 100644 (file)
@@ -23,7 +23,7 @@ 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).
@@ -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.