X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fdama%2Fdama%2Fbishop_set.ma;h=78176d776816c85c513348d356568924e90c77f0;hb=3c1ca5620048ad842144fba291f8bc5f0dca7061;hp=d572820799b26ed2080451ef81083c0064784a86;hpb=b1087303f8e1601aaa3723c42a32626b31e5b10e;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 d57282079..78176d776 100644 --- a/helm/software/matita/contribs/dama/dama/bishop_set.ma +++ b/helm/software/matita/contribs/dama/dama/bishop_set.ma @@ -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 45 +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.