]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/dama/dama/bishop_set.ma
1c68695d3038dbc8806519257c1ae3d867330f02
[helm.git] / helm / software / matita / contribs / dama / dama / bishop_set.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 include "ordered_set.ma".
16
17 (* Definition 2.2, setoid *)
18 record bishop_set: Type ≝ {
19   bs_carr:> Type;
20   bs_apart: bs_carr → bs_carr → CProp;
21   bs_coreflexive: coreflexive ? bs_apart;
22   bs_symmetric: symmetric ? bs_apart;
23   bs_cotransitive: cotransitive ? bs_apart
24 }.
25
26 notation "hvbox(a break # b)" non associative with precedence 50 
27   for @{ 'apart $a $b}.
28   
29 interpretation "bishop_setapartness" 'apart x y = 
30   (cic:/matita/dama/bishop_set/bs_apart.con _ x y).
31
32 definition bishop_set_of_ordered_set: ordered_set → bishop_set.
33 intros (E); apply (mk_bishop_set E (λa,b:E. a ≰ b ∨ b ≰ a));  
34 [1: unfold; cases E; simplify; clear E; intros (x); unfold;
35     intros (H1); apply (H x); cases H1; assumption;
36 |2: unfold; intros(x y H); cases H; clear H; [right|left] assumption;
37 |3: intros (E); unfold; cases E (T f _ cTf); simplify; intros (x y z Axy);
38     cases Axy (H H); lapply (cTf ? ? z H) as H1; cases H1; clear Axy H1;
39     [left; left|right; left|right; right|left; right] assumption]
40 qed.
41
42 (* Definition 2.2 (2) *)
43 definition eq ≝ λA:bishop_set.λa,b:A. ¬ (a # b).
44
45 notation "hvbox(a break ≈ b)" non associative with precedence 50 
46   for @{ 'napart $a $b}.
47       
48 interpretation "Bishop set alikeness" 'napart a b = 
49   (cic:/matita/dama/bishop_set/eq.con _ a b). 
50
51 lemma eq_reflexive:∀E:bishop_set. reflexive ? (eq E).
52 intros (E); unfold; intros (x); apply bs_coreflexive; 
53 qed.
54
55 lemma eq_sym_:∀E:bishop_set.symmetric ? (eq E).
56 intros 4 (E x y H); intro T; cases (H (bs_symmetric ??? T)); 
57 qed.
58
59 lemma eq_sym:∀E:bishop_set.∀x,y:E.x ≈ y → y ≈ x ≝ eq_sym_.
60
61 lemma eq_trans_: ∀E:bishop_set.transitive ? (eq E).
62 (* bug. intros k deve fare whd quanto basta *)
63 intros 6 (E x y z Exy Eyz); intro Axy; cases (bs_cotransitive ???y Axy); 
64 [apply Exy|apply Eyz] assumption.
65 qed.
66
67 coercion cic:/matita/dama/bishop_set/bishop_set_of_ordered_set.con.
68
69 lemma le_antisymmetric: 
70   ∀E:ordered_set.antisymmetric E (le E) (eq E).
71 intros 5 (E x y Lxy Lyx); intro H; 
72 cases H; [apply Lxy;|apply Lyx] assumption;
73 qed.
74
75 lemma le_le_eq: ∀E:ordered_set.∀a,b:E. a ≤ b → b ≤ a → a ≈ b.
76 intros (E x y L1 L2); intro H; cases H; [apply L1|apply L2] assumption;
77 qed.