]> matita.cs.unibo.it Git - helm.git/blob - matita/dama/ordered_sets.ma
b8c6952c7912bcdf0a1439b23fc58ad08e174404
[helm.git] / matita / dama / ordered_sets.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/ordered_sets/".
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 excedence : Type ≝ {
23   exc_carr:> Type;
24   exc_relation: exc_carr → exc_carr → Prop;
25   exc_coreflexive: coreflexive ? exc_relation;
26   exc_cotransitive: cotransitive ? exc_relation 
27 }.
28
29 interpretation "excedence" 'nleq a b =
30  (cic:/matita/ordered_sets/exc_relation.con _ a b). 
31
32 definition le ≝ λE:excedence.λa,b:E. ¬ (a ≰ b).
33
34 interpretation "ordered sets less or equal than" 'leq a b = 
35  (cic:/matita/ordered_sets/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 definition apart ≝ λE:excedence.λa,b:E. a ≰ b ∨ b ≰ a.
47
48 notation "a # b" non associative with precedence 50 for @{ 'apart $a $b}.
49 interpretation "apartness" 'apart a b = (cic:/matita/ordered_sets/apart.con _ a b). 
50
51 lemma apart_coreflexive: ∀E.coreflexive ? (apart E).
52 intros (E); unfold; cases E; simplify; clear E; intros (x); unfold;
53 intros (H1); apply (H x); cases H1; assumption;
54 qed.
55
56 lemma apart_symmetric: ∀E.symmetric ? (apart E).
57 intros (E); unfold; intros(x y H); cases H; clear H; [right|left] assumption; 
58 qed.
59
60 lemma apart_cotrans: ∀E. cotransitive ? (apart E).
61 intros (E); unfold; cases E (T f _ cTf); simplify; intros (x y z Axy);
62 cases Axy (H); lapply (cTf ? ? z H) as H1; cases H1; clear Axy H1;
63 [left; left|right; left|right; right|left; right] assumption.
64 qed.
65
66 definition eq ≝ λE:excedence.λa,b:E. ¬ (a # b).
67
68 notation "a ≈ b" non associative with precedence 50 for @{ 'napart $a $b}.    
69 interpretation "alikeness" 'napart a b =
70   (cic:/matita/ordered_sets/eq.con _ a b). 
71
72 lemma eq_reflexive:∀E. reflexive ? (eq E).
73 intros (E); unfold; cases E (T f cRf _); simplify; unfold Not; intros (x H);
74 apply (cRf x); cases H; assumption;
75 qed.
76
77 lemma eq_symmetric:∀E.symmetric ? (eq E).
78 intros (E); unfold; unfold eq; unfold Not;
79 intros (x y H1 H2); apply H1; cases H2; [right|left] assumption; 
80 qed.
81
82 lemma eq_transitive: ∀E.transitive ? (eq E).
83 intros (E); unfold; cases E (T f _ cTf); simplify; unfold Not; 
84 intros (x y z H1 H2 H3); cases H3 (H4 H4); clear E H3; lapply (cTf ? ? y H4) as H5;
85 cases H5; clear H5 H4 cTf; [1,4: apply H1|*:apply H2] clear H1 H2;
86 [1,3:left|*:right] assumption;
87 qed.
88
89 lemma le_antisymmetric: ∀E.antisymmetric ? (le E) (eq E).
90 intros (E); unfold; intros (x y Lxy Lyx); unfold; unfold; intros (H);
91 cases H; [apply Lxy;|apply Lyx] assumption;
92 qed.
93
94 definition lt ≝ λE:excedence.λa,b:E. a ≤ b ∧ a # b.
95
96 interpretation "ordered sets less than" 'lt a b =
97  (cic:/matita/ordered_sets/lt.con _ a b).
98
99 lemma lt_coreflexive: ∀E.coreflexive ? (lt E).
100 intros (E); unfold; unfold Not; intros (x H); cases H (_ ABS); 
101 apply (apart_coreflexive ? x ABS);
102 qed.
103
104 lemma lt_transitive: ∀E.transitive ? (lt E).
105 intros (E); unfold; intros (x y z H1 H2); cases H1 (Lxy Axy); cases H2 (Lyz Ayz); 
106 split; [apply (le_transitive ???? Lxy Lyz)] clear H1 H2;
107 cases Axy (H1 H1); cases Ayz (H2 H2); [1:cases (Lxy H1)|3:cases (Lyz H2)]
108 clear Axy Ayz;lapply (exc_cotransitive E) as c; unfold cotransitive in c;
109 lapply (exc_coreflexive E) as r; unfold coreflexive in r;
110 [1: lapply (c ?? y H1) as H3; cases H3 (H4 H4); [cases (Lxy H4)|cases (r ? H4)]
111 |2: lapply (c ?? x H2) as H3; cases H3 (H4 H4); [right; assumption|cases (Lxy H4)]]
112 qed.
113
114 theorem mah: ∀E:excedence.∀a,b:E. (a < b) → (b ≰ a).
115 intros (E a b Lab); cases Lab (LEab Aab);
116 cases Aab (H H); [cases (LEab H)] fold normalize (b ≰ a); assumption; (* BUG *)  
117 qed.
118
119 -- altro file
120 opposto TH è assioma per ordine totale.
121
122 -- 
123
124
125
126
127
128
129
130 record is_order_relation (C:Type) (le:C→C→Prop) (eq:C→C→Prop) : Type ≝ { 
131   or_reflexive: reflexive ? le;
132   or_transitive: transitive ? le;
133   or_antisimmetric: antisymmetric ? le eq
134 }.
135
136 record ordered_set: Type ≝ { 
137   os_carr:> excedence;
138   os_order_relation_properties:> is_order_relation ? (le os_carr) (apart os_carr)
139 }.
140
141 ordered_set.
142
143 E
144
145 E
146
147 theorem antisimmetric_to_cotransitive_to_transitive:
148  ∀C.∀le:relation C. antisimmetric ? le → cotransitive ? le →
149   transitive ? le.
150  intros;
151  unfold transitive;
152  intros;
153  elim (c ? ? z H1);
154   [ assumption
155   | rewrite < (H ? ? H2 t);
156     assumption
157   ].
158 qed.
159
160 definition is_increasing ≝ λO:ordered_set.λa:nat→O.∀n:nat.a n ≤ a (S n).
161 definition is_decreasing ≝ λO:ordered_set.λa:nat→O.∀n:nat.a (S n) ≤ a n.
162
163 definition is_upper_bound ≝ λO:ordered_set.λa:nat→O.λu:O.∀n:nat.a n ≤ u.
164 definition is_lower_bound ≝ λO:ordered_set.λa:nat→O.λu:O.∀n:nat.u ≤ a n.
165
166 record is_sup (O:ordered_set) (a:nat→O) (u:O) : Prop ≝
167  { sup_upper_bound: is_upper_bound O a u; 
168    sup_least_upper_bound: ∀v:O. is_upper_bound O a v → u≤v
169  }.
170
171 record is_inf (O:ordered_set) (a:nat→O) (u:O) : Prop ≝
172  { inf_lower_bound: is_lower_bound O a u; 
173    inf_greatest_lower_bound: ∀v:O. is_lower_bound O a v → v≤u
174  }.
175
176 record is_bounded_below (O:ordered_set) (a:nat→O) : Type ≝
177  { ib_lower_bound: O;
178    ib_lower_bound_is_lower_bound: is_lower_bound ? a ib_lower_bound
179  }.
180
181 record is_bounded_above (O:ordered_set) (a:nat→O) : Type ≝
182  { ib_upper_bound: O;
183    ib_upper_bound_is_upper_bound: is_upper_bound ? a ib_upper_bound
184  }.
185
186 record is_bounded (O:ordered_set) (a:nat→O) : Type ≝
187  { ib_bounded_below:> is_bounded_below ? a;
188    ib_bounded_above:> is_bounded_above ? a
189  }.
190
191 record bounded_below_sequence (O:ordered_set) : Type ≝
192  { bbs_seq:1> nat→O;
193    bbs_is_bounded_below:> is_bounded_below ? bbs_seq
194  }.
195
196 record bounded_above_sequence (O:ordered_set) : Type ≝
197  { bas_seq:1> nat→O;
198    bas_is_bounded_above:> is_bounded_above ? bas_seq
199  }.
200
201 record bounded_sequence (O:ordered_set) : Type ≝
202  { bs_seq:1> nat → O;
203    bs_is_bounded_below: is_bounded_below ? bs_seq;
204    bs_is_bounded_above: is_bounded_above ? bs_seq
205  }.
206
207 definition bounded_below_sequence_of_bounded_sequence ≝
208  λO:ordered_set.λb:bounded_sequence O.
209   mk_bounded_below_sequence ? b (bs_is_bounded_below ? b).
210
211 coercion cic:/matita/ordered_sets/bounded_below_sequence_of_bounded_sequence.con.
212
213 definition bounded_above_sequence_of_bounded_sequence ≝
214  λO:ordered_set.λb:bounded_sequence O.
215   mk_bounded_above_sequence ? b (bs_is_bounded_above ? b).
216
217 coercion cic:/matita/ordered_sets/bounded_above_sequence_of_bounded_sequence.con.
218
219 definition lower_bound ≝
220  λO:ordered_set.λb:bounded_below_sequence O.
221   ib_lower_bound ? b (bbs_is_bounded_below ? b).
222
223 lemma lower_bound_is_lower_bound:
224  ∀O:ordered_set.∀b:bounded_below_sequence O.
225   is_lower_bound ? b (lower_bound ? b).
226  intros;
227  unfold lower_bound;
228  apply ib_lower_bound_is_lower_bound.
229 qed.
230
231 definition upper_bound ≝
232  λO:ordered_set.λb:bounded_above_sequence O.
233   ib_upper_bound ? b (bas_is_bounded_above ? b).
234
235 lemma upper_bound_is_upper_bound:
236  ∀O:ordered_set.∀b:bounded_above_sequence O.
237   is_upper_bound ? b (upper_bound ? b).
238  intros;
239  unfold upper_bound;
240  apply ib_upper_bound_is_upper_bound.
241 qed.
242
243 definition lt ≝ λO:ordered_set.λa,b:O.a ≤ b ∧ a ≠ b.
244
245 interpretation "Ordered set lt" 'lt a b =
246  (cic:/matita/ordered_sets/lt.con _ a b).
247
248 definition reverse_ordered_set: ordered_set → ordered_set.
249  intros;
250  apply mk_ordered_set;
251   [2:apply (λx,y:o.y ≤ x)
252   | skip
253   | apply mk_is_order_relation;
254      [ simplify;
255        intros;
256        apply (or_reflexive ? ? o)
257      | simplify;
258        intros;
259        apply (or_transitive ? ? o);
260         [2: apply H1
261         | skip
262         | assumption
263         ] 
264      | simplify;
265        intros;
266        apply (or_antisimmetric ? ? o);
267        assumption
268      ]
269   ].
270 qed.
271  
272 interpretation "Ordered set ge" 'geq a b =
273  (cic:/matita/ordered_sets/os_le.con _
274   (cic:/matita/ordered_sets/os_pre_ordered_set.con _
275    (cic:/matita/ordered_sets/reverse_ordered_set.con _ _)) a b).
276
277 lemma is_lower_bound_reverse_is_upper_bound:
278  ∀O:ordered_set.∀a:nat→O.∀l:O.
279   is_lower_bound O a l → is_upper_bound (reverse_ordered_set O) a l.
280  intros;
281  unfold;
282  intro;
283  unfold;
284  unfold reverse_ordered_set;
285  simplify;
286  apply H.
287 qed.
288
289 lemma is_upper_bound_reverse_is_lower_bound:
290  ∀O:ordered_set.∀a:nat→O.∀l:O.
291   is_upper_bound O a l → is_lower_bound (reverse_ordered_set O) a l.
292  intros;
293  unfold;
294  intro;
295  unfold;
296  unfold reverse_ordered_set;
297  simplify;
298  apply H.
299 qed.
300
301 lemma reverse_is_lower_bound_is_upper_bound:
302  ∀O:ordered_set.∀a:nat→O.∀l:O.
303   is_lower_bound (reverse_ordered_set O) a l → is_upper_bound O a l.
304  intros;
305  unfold in H;
306  unfold reverse_ordered_set in H;
307  apply H.
308 qed.
309
310 lemma reverse_is_upper_bound_is_lower_bound:
311  ∀O:ordered_set.∀a:nat→O.∀l:O.
312   is_upper_bound (reverse_ordered_set O) a l → is_lower_bound O a l.
313  intros;
314  unfold in H;
315  unfold reverse_ordered_set in H;
316  apply H.
317 qed.
318
319
320 lemma is_inf_to_reverse_is_sup:
321  ∀O:ordered_set.∀a:bounded_below_sequence O.∀l:O.
322   is_inf O a l → is_sup (reverse_ordered_set O) a l.
323  intros;
324  apply (mk_is_sup (reverse_ordered_set O));
325   [ apply is_lower_bound_reverse_is_upper_bound;
326     apply inf_lower_bound;
327     assumption
328   | intros;
329     change in v with (os_carrier O);
330     change with (v ≤ l);
331     apply (inf_greatest_lower_bound ? ? ? H);
332     apply reverse_is_upper_bound_is_lower_bound;
333     assumption
334   ].
335 qed.
336  
337 lemma is_sup_to_reverse_is_inf:
338  ∀O:ordered_set.∀a:bounded_above_sequence O.∀l:O.
339   is_sup O a l → is_inf (reverse_ordered_set O) a l.
340  intros;
341  apply (mk_is_inf (reverse_ordered_set O));
342   [ apply is_upper_bound_reverse_is_lower_bound;
343     apply sup_upper_bound;
344     assumption
345   | intros;
346     change in v with (os_carrier O);
347     change with (l ≤ v);
348     apply (sup_least_upper_bound ? ? ? H);
349     apply reverse_is_lower_bound_is_upper_bound;
350     assumption
351   ].
352 qed.
353
354 lemma reverse_is_sup_to_is_inf:
355  ∀O:ordered_set.∀a:bounded_above_sequence O.∀l:O.
356   is_sup (reverse_ordered_set O) a l → is_inf O a l.
357  intros;
358  apply mk_is_inf;
359   [ apply reverse_is_upper_bound_is_lower_bound;
360     change in l with (os_carrier (reverse_ordered_set O));
361     apply sup_upper_bound;
362     assumption
363   | intros;
364     change in l with (os_carrier (reverse_ordered_set O));
365     change in v with (os_carrier (reverse_ordered_set O));
366     change with (os_le (reverse_ordered_set O) l v);
367     apply (sup_least_upper_bound ? ? ? H);
368     change in v with (os_carrier O);
369     apply is_lower_bound_reverse_is_upper_bound;
370     assumption
371   ].
372 qed.
373
374 lemma reverse_is_inf_to_is_sup:
375  ∀O:ordered_set.∀a:bounded_above_sequence O.∀l:O.
376   is_inf (reverse_ordered_set O) a l → is_sup O a l.
377  intros;
378  apply mk_is_sup;
379   [ apply reverse_is_lower_bound_is_upper_bound;
380     change in l with (os_carrier (reverse_ordered_set O));
381     apply (inf_lower_bound ? ? ? H)
382   | intros;
383     change in l with (os_carrier (reverse_ordered_set O));
384     change in v with (os_carrier (reverse_ordered_set O));
385     change with (os_le (reverse_ordered_set O) v l);
386     apply (inf_greatest_lower_bound ? ? ? H);
387     change in v with (os_carrier O);
388     apply is_upper_bound_reverse_is_lower_bound;
389     assumption
390   ].
391 qed.
392
393 record cotransitively_ordered_set: Type :=
394  { cos_ordered_set :> ordered_set;
395    cos_cotransitive: cotransitive ? (os_le cos_ordered_set)
396  }.