]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/dama/dama/bishop_set.ma
more notation moved to core notation, unification of duplicated CProp connectives
[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 interpretation "bishop set apartness" 'apart x y = (bs_apart _ x y).
27
28 definition bishop_set_of_ordered_set: ordered_set → bishop_set.
29 intros (E); apply (mk_bishop_set E (λa,b:E. a ≰ b ∨ b ≰ a));  
30 [1: unfold; cases E; simplify; clear E; intros (x); unfold;
31     intros (H1); apply (H x); cases H1; assumption;
32 |2: unfold; intros(x y H); cases H; clear H; [right|left] assumption;
33 |3: intros (E); unfold; cases E (T f _ cTf); simplify; intros (x y z Axy);
34     cases Axy (H H); lapply (cTf ? ? z H) as H1; cases H1; clear Axy H1;
35     [left; left|right; left|right; right|left; right] assumption]
36 qed.
37
38 (* Definition 2.2 (2) *)
39 definition eq ≝ λA:bishop_set.λa,b:A. ¬ (a # b).
40
41 interpretation "Bishop set alikeness" 'napart a b = (eq _ a b). 
42
43 lemma eq_reflexive:∀E:bishop_set. reflexive ? (eq E).
44 intros (E); unfold; intros (x); apply bs_coreflexive; 
45 qed.
46
47 lemma eq_sym_:∀E:bishop_set.symmetric ? (eq E).
48 intros 4 (E x y H); intro T; cases (H (bs_symmetric ??? T)); 
49 qed.
50
51 lemma eq_sym:∀E:bishop_set.∀x,y:E.x ≈ y → y ≈ x ≝ eq_sym_.
52
53 lemma eq_trans_: ∀E:bishop_set.transitive ? (eq E).
54 intros 6 (E x y z Exy Eyz); intro Axy; cases (bs_cotransitive ???y Axy); 
55 [apply Exy|apply Eyz] assumption.
56 qed.
57
58 coercion cic:/matita/dama/bishop_set/bishop_set_of_ordered_set.con.
59
60 lemma le_antisymmetric: 
61   ∀E:ordered_set.antisymmetric E (le E) (eq E).
62 intros 5 (E x y Lxy Lyx); intro H; 
63 cases H; [apply Lxy;|apply Lyx] assumption;
64 qed.
65
66 lemma le_le_eq: ∀E:ordered_set.∀a,b:E. a ≤ b → b ≤ a → a ≈ b.
67 intros (E x y L1 L2); intro H; cases H; [apply L1|apply L2] assumption;
68 qed.
69
70 definition lt ≝ λE:ordered_set.λa,b:E. a ≤ b ∧ a # b.
71
72 interpretation "ordered sets less than" 'lt a b = (lt _ a b).
73
74 lemma lt_coreflexive: ∀E.coreflexive ? (lt E).
75 intros 2 (E x); intro H; cases H (_ ABS);
76 apply (bs_coreflexive ? x ABS);
77 qed.
78
79 lemma lt_transitive: ∀E.transitive ? (lt E).
80 intros (E); unfold; intros (x y z H1 H2); cases H1 (Lxy Axy); cases H2 (Lyz Ayz);
81 split; [apply (le_transitive ???? Lxy Lyz)] clear H1 H2;
82 cases Axy (H1 H1); cases Ayz (H2 H2); [1:cases (Lxy H1)|3:cases (Lyz H2)]clear Axy Ayz;
83 [1: cases (os_cotransitive ??? y H1) (X X); [cases (Lxy X)|cases (os_coreflexive ?? X)]
84 |2: cases (os_cotransitive ??? x H2) (X X); [right;assumption|cases (Lxy X)]]
85 qed.
86
87 theorem lt_to_excess: ∀E:ordered_set.∀a,b:E. (a < b) → (b ≰ a).
88 intros (E a b Lab); cases Lab (LEab Aab); cases Aab (H H);[cases (LEab H)]
89 assumption;
90 qed.
91
92 definition bs_subset ≝ λO:bishop_set.λP,Q:O→Prop.∀x:O.P x → Q x.
93
94 interpretation "bishop set subset" 'subseteq a b = (bs_subset _ a b). 
95
96 definition square_bishop_set : bishop_set → bishop_set.
97 intro S; apply (mk_bishop_set (S × S));
98 [1: intros (x y); apply ((\fst x # \fst y) ∨ (\snd x # \snd y));
99 |2: intro x; simplify; intro; cases H (X X); clear H; apply (bs_coreflexive ?? X);   
100 |3: intros 2 (x y); simplify; intro H; cases H (X X); clear H; [left|right] apply (bs_symmetric ??? X); 
101 |4: intros 3 (x y z); simplify; intro H; cases H (X X); clear H;
102     [1: cases (bs_cotransitive ??? (\fst z) X); [left;left|right;left]assumption;
103     |2: cases (bs_cotransitive ??? (\snd z) X); [left;right|right;right]assumption;]]
104 qed.
105
106 notation "s 2 \atop \neq" non associative with precedence 90
107   for @{ 'square_bs $s }.
108 interpretation "bishop set square" 'square x = (square_bishop_set x).
109 interpretation "bishop set square" 'square_bs x = (square_bishop_set x).