]> matita.cs.unibo.it Git - helm.git/blob - matita/dama/excess.ma
experimental branch with no set baseuri command and no developments
[helm.git] / matita / dama / excess.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 set "baseuri" "cic:/matita/excess/".
16
17 include "higher_order_defs/relations.ma".
18 include "nat/plus.ma".
19 include "constructive_connectives.ma".
20 include "constructive_higher_order_relations.ma".
21
22 record excess : Type ≝ {
23   exc_carr:> Type;
24   exc_relation: exc_carr → exc_carr → Type;
25   exc_coreflexive: coreflexive ? exc_relation;
26   exc_cotransitive: cotransitive ? exc_relation 
27 }.
28
29 interpretation "excess" 'nleq a b =
30  (cic:/matita/excess/exc_relation.con _ a b). 
31
32 definition le ≝ λE:excess.λa,b:E. ¬ (a ≰ b).
33
34 interpretation "ordered sets less or equal than" 'leq a b = 
35  (cic:/matita/excess/le.con _ a b).
36
37 lemma le_reflexive: ∀E.reflexive ? (le E).
38 intros (E); unfold; cases E; simplify; intros (x); apply (H x);
39 qed.
40
41 lemma le_transitive: ∀E.transitive ? (le E).
42 intros (E); unfold; cases E; simplify; unfold Not; intros (x y z Rxy Ryz H2); 
43 cases (c x z y H2) (H4 H5); clear H2; [exact (Rxy H4)|exact (Ryz H5)] 
44 qed.
45
46 record apartness : Type ≝ {
47   ap_carr:> Type;
48   ap_apart: ap_carr → ap_carr → Type;
49   ap_coreflexive: coreflexive ? ap_apart;
50   ap_symmetric: symmetric ? ap_apart;
51   ap_cotransitive: cotransitive ? ap_apart
52 }.
53
54 notation "a break # b" non associative with precedence 50 for @{ 'apart $a $b}.
55 interpretation "apartness" 'apart x y = 
56   (cic:/matita/excess/ap_apart.con _ x y).
57
58 definition strong_ext ≝ λA:apartness.λop:A→A.∀x,y. op x # op y → x # y.
59
60 definition apart ≝ λE:excess.λa,b:E. a ≰ b ∨ b ≰ a.
61
62 definition apart_of_excess: excess → apartness.
63 intros (E); apply (mk_apartness E (apart E));  
64 [1: unfold; cases E; simplify; clear E; intros (x); unfold;
65     intros (H1); apply (H x); cases H1; assumption;
66 |2: unfold; intros(x y H); cases H; clear H; [right|left] assumption;
67 |3: intros (E); unfold; cases E (T f _ cTf); simplify; intros (x y z Axy);
68     cases Axy (H H); lapply (cTf ? ? z H) as H1; cases H1; clear Axy H1;
69     [left; left|right; left|right; right|left; right] assumption]
70 qed.
71
72 coercion cic:/matita/excess/apart_of_excess.con.
73
74 definition eq ≝ λA:apartness.λa,b:A. ¬ (a # b).
75
76 notation "a break ≈ b" non associative with precedence 50 for @{ 'napart $a $b}.    
77 interpretation "alikeness" 'napart a b =
78   (cic:/matita/excess/eq.con _ a b). 
79
80 lemma eq_reflexive:∀E. reflexive ? (eq E).
81 intros (E); unfold; intros (x); apply ap_coreflexive; 
82 qed.
83
84 lemma eq_sym_:∀E.symmetric ? (eq E).
85 intros (E); unfold; intros (x y Exy); unfold; unfold; intros (Ayx); apply Exy;
86 apply ap_symmetric; assumption; 
87 qed.
88
89 lemma eq_sym:∀E:apartness.∀x,y:E.x ≈ y → y ≈ x ≝ eq_sym_.
90
91 coercion cic:/matita/excess/eq_sym.con.
92
93 lemma eq_trans_: ∀E.transitive ? (eq E).
94 (* bug. intros k deve fare whd quanto basta *)
95 intros 6 (E x y z Exy Eyz); intro Axy; cases (ap_cotransitive ???y Axy); 
96 [apply Exy|apply Eyz] assumption.
97 qed.
98
99 lemma eq_trans:∀E:apartness.∀x,z,y:E.x ≈ y → y ≈ z → x ≈ z ≝ 
100   λE,x,y,z.eq_trans_ E x z y.
101
102 notation > "'Eq'≈" non associative with precedence 50 for
103  @{'eqrewrite}.
104  
105 interpretation "eq_rew" 'eqrewrite = 
106  (cic:/matita/excess/eq_trans.con _ _ _).
107
108 (* BUG: vedere se ricapita *)
109 alias id "antisymmetric" = "cic:/matita/constructive_higher_order_relations/antisymmetric.con".
110 lemma le_antisymmetric: ∀E.antisymmetric ? (le E) (eq ?).
111 intros 5 (E x y Lxy Lyx); intro H;
112 cases H; [apply Lxy;|apply Lyx] assumption;
113 qed.
114
115 definition lt ≝ λE:excess.λa,b:E. a ≤ b ∧ a # b.
116
117 interpretation "ordered sets less than" 'lt a b =
118  (cic:/matita/excess/lt.con _ a b).
119
120 lemma lt_coreflexive: ∀E.coreflexive ? (lt E).
121 intros 2 (E x); intro H; cases H (_ ABS); 
122 apply (ap_coreflexive ? x ABS);
123 qed.
124  
125 lemma lt_transitive: ∀E.transitive ? (lt E).
126 intros (E); unfold; intros (x y z H1 H2); cases H1 (Lxy Axy); cases H2 (Lyz Ayz); 
127 split; [apply (le_transitive ???? Lxy Lyz)] clear H1 H2;
128 cases Axy (H1 H1); cases Ayz (H2 H2); [1:cases (Lxy H1)|3:cases (Lyz H2)]
129 clear Axy Ayz;lapply (exc_cotransitive E) as c; unfold cotransitive in c;
130 lapply (exc_coreflexive E) as r; unfold coreflexive in r;
131 [1: lapply (c ?? y H1) as H3; cases H3 (H4 H4); [cases (Lxy H4)|cases (r ? H4)]
132 |2: lapply (c ?? x H2) as H3; cases H3 (H4 H4); [right; assumption|cases (Lxy H4)]]
133 qed.
134
135 theorem lt_to_excess: ∀E:excess.∀a,b:E. (a < b) → (b ≰ a).
136 intros (E a b Lab); cases Lab (LEab Aab);
137 cases Aab (H H); [cases (LEab H)] fold normalize (b ≰ a); assumption; (* BUG *)  
138 qed.
139
140 lemma unfold_apart: ∀E:excess. ∀x,y:E. x ≰ y ∨ y ≰ x → x # y.
141 intros; assumption;
142 qed.
143
144 lemma le_rewl: ∀E:excess.∀z,y,x:E. x ≈ y → x ≤ z → y ≤ z.
145 intros (E z y x Exy Lxz); apply (le_transitive ???? ? Lxz);
146 intro Xyz; apply Exy; apply unfold_apart; right; assumption;
147 qed.
148
149 lemma le_rewr: ∀E:excess.∀z,y,x:E. x ≈ y → z ≤ x → z ≤ y.
150 intros (E z y x Exy Lxz); apply (le_transitive ???? Lxz);
151 intro Xyz; apply Exy; apply unfold_apart; left; assumption;
152 qed.
153
154 notation > "'Le'≪" non associative with precedence 50 for
155  @{'lerewritel}.
156  
157 interpretation "le_rewl" 'lerewritel = 
158  (cic:/matita/excess/le_rewl.con _ _ _).
159
160 notation > "'Le'≫" non associative with precedence 50 for
161  @{'lerewriter}.
162  
163 interpretation "le_rewr" 'lerewriter = 
164  (cic:/matita/excess/le_rewr.con _ _ _).
165
166 lemma ap_rewl: ∀A:apartness.∀x,z,y:A. x ≈ y → y # z → x # z.
167 intros (A x z y Exy Ayz); cases (ap_cotransitive ???x Ayz); [2:assumption]
168 cases (Exy (ap_symmetric ??? a));
169 qed.
170   
171 lemma ap_rewr: ∀A:apartness.∀x,z,y:A. x ≈ y → z # y → z # x.
172 intros (A x z y Exy Azy); apply ap_symmetric; apply (ap_rewl ???? Exy);
173 apply ap_symmetric; assumption;
174 qed.
175
176 lemma exc_rewl: ∀A:excess.∀x,z,y:A. x ≈ y → y ≰ z → x ≰ z.
177 intros (A x z y Exy Ayz); elim (exc_cotransitive ???x Ayz); [2:assumption]
178 cases Exy; right; assumption;
179 qed.
180   
181 lemma exc_rewr: ∀A:excess.∀x,z,y:A. x ≈ y → z ≰ y → z ≰ x.
182 intros (A x z y Exy Azy); elim (exc_cotransitive ???x Azy); [assumption]
183 elim (Exy); left; assumption;
184 qed.
185
186 notation > "'Ex'≪" non associative with precedence 50 for
187  @{'excessrewritel}.
188  
189 interpretation "exc_rewl" 'excessrewritel = 
190  (cic:/matita/excess/exc_rewl.con _ _ _).
191
192 notation > "'Ex'≫" non associative with precedence 50 for
193  @{'excessrewriter}.
194  
195 interpretation "exc_rewr" 'excessrewriter = 
196  (cic:/matita/excess/exc_rewr.con _ _ _).
197
198 lemma lt_rewr: ∀A:excess.∀x,z,y:A. x ≈ y → z < y → z < x.
199 intros (A x y z E H); split; elim H; 
200 [apply (le_rewr ???? (eq_sym ??? E));|apply (ap_rewr ???? E)] assumption;
201 qed.
202
203 lemma lt_rewl: ∀A:excess.∀x,z,y:A. x ≈ y → y < z → x < z.
204 intros (A x y z E H); split; elim H; 
205 [apply (le_rewl ???? (eq_sym ??? E));| apply (ap_rewl ???? E);] assumption;
206 qed.
207
208 notation > "'Lt'≪" non associative with precedence 50 for
209  @{'ltrewritel}.
210  
211 interpretation "lt_rewl" 'ltrewritel = 
212  (cic:/matita/excess/lt_rewl.con _ _ _).
213
214 notation > "'Lt'≫" non associative with precedence 50 for
215  @{'ltrewriter}.
216  
217 interpretation "lt_rewr" 'ltrewriter = 
218  (cic:/matita/excess/lt_rewr.con _ _ _).
219
220 lemma lt_le_transitive: ∀A:excess.∀x,y,z:A.x < y → y ≤ z → x < z.
221 intros (A x y z LT LE); cases LT (LEx APx); split; [apply (le_transitive ???? LEx LE)]
222 whd in LE LEx APx; cases APx (EXx EXx); [cases (LEx EXx)]
223 cases (exc_cotransitive ??? z EXx) (EXz EXz); [cases (LE EXz)]
224 right; assumption;
225 qed.
226
227 lemma le_lt_transitive: ∀A:excess.∀x,y,z:A.x ≤ y → y < z → x < z.
228 intros (A x y z LE LT); cases LT (LEx APx); split; [apply (le_transitive ???? LE LEx)]
229 whd in LE LEx APx; cases APx (EXx EXx); [cases (LEx EXx)]
230 cases (exc_cotransitive ??? x EXx) (EXz EXz); [right; assumption]
231 cases LE; assumption;
232 qed.
233     
234 lemma le_le_eq: ∀E:excess.∀a,b:E. a ≤ b → b ≤ a → a ≈ b.
235 intros (E x y L1 L2); intro H; cases H; [apply L1|apply L2] assumption;
236 qed.
237
238 lemma eq_le_le: ∀E:excess.∀a,b:E. a ≈ b → a ≤ b ∧ b ≤ a.
239 intros (E x y H); unfold apart_of_excess in H; unfold apart in H;
240 simplify in H; split; intro; apply H; [left|right] assumption.
241 qed.
242
243 lemma ap_le_to_lt: ∀E:excess.∀a,c:E.c # a → c ≤ a → c < a.
244 intros; split; assumption;
245 qed.
246
247 definition total_order_property : ∀E:excess. Type ≝
248   λE:excess. ∀a,b:E. a ≰ b → b < a.
249