1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 set "baseuri" "cic:/matita/ordered_sets/".
17 include "higher_order_defs/relations.ma".
18 include "nat/plus.ma".
19 include "constructive_connectives.ma".
21 definition cotransitive ≝
22 λC:Type.λle:C→C→Prop.∀x,y,z:C. le x y → le x z ∨ le z y.
24 definition antisimmetric ≝
25 λC:Type.λle:C→C→Prop.∀x,y:C. le x y → le y x → x=y.
27 record is_order_relation (C:Type) (le:C→C→Prop) : Type ≝
28 { or_reflexive: reflexive ? le;
29 or_transitive: transitive ? le;
30 or_antisimmetric: antisimmetric ? le
33 record ordered_set: Type ≝
35 os_le: os_carrier → os_carrier → Prop;
36 os_order_relation_properties:> is_order_relation ? os_le
39 interpretation "Ordered Sets le" 'leq a b =
40 (cic:/matita/ordered_sets/os_le.con _ a b).
42 theorem antisimmetric_to_cotransitive_to_transitive:
43 ∀C.∀le:relation C. antisimmetric ? le → cotransitive ? le →
50 | rewrite < (H ? ? H2 t);
55 definition is_increasing ≝ λO:ordered_set.λa:nat→O.∀n:nat.a n ≤ a (S n).
56 definition is_decreasing ≝ λO:ordered_set.λa:nat→O.∀n:nat.a (S n) ≤ a n.
58 definition is_upper_bound ≝ λO:ordered_set.λa:nat→O.λu:O.∀n:nat.a n ≤ u.
59 definition is_lower_bound ≝ λO:ordered_set.λa:nat→O.λu:O.∀n:nat.u ≤ a n.
61 record is_sup (O:ordered_set) (a:nat→O) (u:O) : Prop ≝
62 { sup_upper_bound: is_upper_bound O a u;
63 sup_least_upper_bound: ∀v:O. is_upper_bound O a v → u≤v
66 record is_inf (O:ordered_set) (a:nat→O) (u:O) : Prop ≝
67 { inf_lower_bound: is_lower_bound O a u;
68 inf_greatest_lower_bound: ∀v:O. is_lower_bound O a v → v≤u
71 record is_bounded_below (O:ordered_set) (a:nat→O) : Type ≝
73 ib_lower_bound_is_lower_bound: is_lower_bound ? a ib_lower_bound
76 record is_bounded_above (O:ordered_set) (a:nat→O) : Type ≝
78 ib_upper_bound_is_upper_bound: is_upper_bound ? a ib_upper_bound
81 record is_bounded (O:ordered_set) (a:nat→O) : Type ≝
82 { ib_bounded_below:> is_bounded_below ? a;
83 ib_bounded_above:> is_bounded_above ? a
86 record bounded_below_sequence (O:ordered_set) : Type ≝
88 bbs_is_bounded_below:> is_bounded_below ? bbs_seq
91 record bounded_above_sequence (O:ordered_set) : Type ≝
93 bas_is_bounded_above:> is_bounded_above ? bas_seq
96 record bounded_sequence (O:ordered_set) : Type ≝
98 bs_is_bounded_below: is_bounded_below ? bs_seq;
99 bs_is_bounded_above: is_bounded_above ? bs_seq
102 definition bounded_below_sequence_of_bounded_sequence ≝
103 λO:ordered_set.λb:bounded_sequence O.
104 mk_bounded_below_sequence ? b (bs_is_bounded_below ? b).
106 coercion cic:/matita/ordered_sets/bounded_below_sequence_of_bounded_sequence.con.
108 definition bounded_above_sequence_of_bounded_sequence ≝
109 λO:ordered_set.λb:bounded_sequence O.
110 mk_bounded_above_sequence ? b (bs_is_bounded_above ? b).
112 coercion cic:/matita/ordered_sets/bounded_above_sequence_of_bounded_sequence.con.
114 definition lower_bound ≝
115 λO:ordered_set.λb:bounded_below_sequence O.
116 ib_lower_bound ? b (bbs_is_bounded_below ? b).
118 lemma lower_bound_is_lower_bound:
119 ∀O:ordered_set.∀b:bounded_below_sequence O.
120 is_lower_bound ? b (lower_bound ? b).
123 apply ib_lower_bound_is_lower_bound.
126 definition upper_bound ≝
127 λO:ordered_set.λb:bounded_above_sequence O.
128 ib_upper_bound ? b (bas_is_bounded_above ? b).
130 lemma upper_bound_is_upper_bound:
131 ∀O:ordered_set.∀b:bounded_above_sequence O.
132 is_upper_bound ? b (upper_bound ? b).
135 apply ib_upper_bound_is_upper_bound.
138 record is_dedekind_sigma_complete (O:ordered_set) : Type ≝
139 { dsc_inf: ∀a:nat→O.∀m:O. is_lower_bound ? a m → ex ? (λs:O.is_inf O a s);
140 dsc_inf_proof_irrelevant:
141 ∀a:nat→O.∀m,m':O.∀p:is_lower_bound ? a m.∀p':is_lower_bound ? a m'.
142 (match dsc_inf a m p with [ ex_intro s _ ⇒ s ]) =
143 (match dsc_inf a m' p' with [ ex_intro s' _ ⇒ s' ]);
144 dsc_sup: ∀a:nat→O.∀m:O. is_upper_bound ? a m → ex ? (λs:O.is_sup O a s);
145 dsc_sup_proof_irrelevant:
146 ∀a:nat→O.∀m,m':O.∀p:is_upper_bound ? a m.∀p':is_upper_bound ? a m'.
147 (match dsc_sup a m p with [ ex_intro s _ ⇒ s ]) =
148 (match dsc_sup a m' p' with [ ex_intro s' _ ⇒ s' ])
151 record dedekind_sigma_complete_ordered_set : Type ≝
152 { dscos_ordered_set:> ordered_set;
153 dscos_dedekind_sigma_complete_properties:>
154 is_dedekind_sigma_complete dscos_ordered_set
158 ∀O:dedekind_sigma_complete_ordered_set.
159 bounded_below_sequence O → O.
162 (dsc_inf O (dscos_dedekind_sigma_complete_properties O) b);
164 | apply (lower_bound ? b)
165 | apply lower_bound_is_lower_bound
170 ∀O:dedekind_sigma_complete_ordered_set.
171 ∀a:bounded_below_sequence O.
172 is_inf ? a (inf ? a).
176 elim (dsc_inf O (dscos_dedekind_sigma_complete_properties O) a
177 (lower_bound O a) (lower_bound_is_lower_bound O a));
182 lemma inf_proof_irrelevant:
183 ∀O:dedekind_sigma_complete_ordered_set.
184 ∀a,a':bounded_below_sequence O.
185 bbs_seq ? a = bbs_seq ? a' →
191 generalize in match i1;
196 rewrite < (dsc_inf_proof_irrelevant O O f (ib_lower_bound ? f i)
197 (ib_lower_bound ? f i2) (ib_lower_bound_is_lower_bound ? f i)
198 (ib_lower_bound_is_lower_bound ? f i2));
203 ∀O:dedekind_sigma_complete_ordered_set.
204 bounded_above_sequence O → O.
207 (dsc_sup O (dscos_dedekind_sigma_complete_properties O) b);
209 | apply (upper_bound ? b)
210 | apply upper_bound_is_upper_bound
215 ∀O:dedekind_sigma_complete_ordered_set.
216 ∀a:bounded_above_sequence O.
217 is_sup ? a (sup ? a).
221 elim (dsc_sup O (dscos_dedekind_sigma_complete_properties O) a
222 (upper_bound O a) (upper_bound_is_upper_bound O a));
227 lemma sup_proof_irrelevant:
228 ∀O:dedekind_sigma_complete_ordered_set.
229 ∀a,a':bounded_above_sequence O.
230 bas_seq ? a = bas_seq ? a' →
236 generalize in match i1;
241 rewrite < (dsc_sup_proof_irrelevant O O f (ib_upper_bound ? f i2)
242 (ib_upper_bound ? f i) (ib_upper_bound_is_upper_bound ? f i2)
243 (ib_upper_bound_is_upper_bound ? f i));
250 ∀O:dedekind_sigma_complete_ordered_set.
251 ∀a:bounded_sequence O. inf ? a ≤ sup ? a.
253 apply (or_transitive ? ? O' ? (a O));
254 [ elim daemon (*apply (inf_lower_bound ? ? ? ? (inf_is_inf ? ? a))*)
255 | elim daemon (*apply (sup_upper_bound ? ? ? ? (sup_is_sup ? ? a))*)
259 lemma inf_respects_le:
260 ∀O:dedekind_sigma_complete_ordered_set.
261 ∀a:bounded_below_sequence O.∀m:O.
262 is_upper_bound ? a m → inf ? a ≤ m.
264 apply (or_transitive ? ? O' ? (sup ? (mk_bounded_sequence ? a ? ?)));
265 [ apply (bbs_is_bounded_below ? a)
266 | apply (mk_is_bounded_above ? ? m H)
269 (sup_least_upper_bound ? ? ?
270 (sup_is_sup ? (mk_bounded_sequence O' a a
271 (mk_is_bounded_above O' a m H))));
276 definition is_sequentially_monotone ≝
277 λO:ordered_set.λf:O→O.
278 ∀a:nat→O.∀p:is_increasing ? a.
279 is_increasing ? (λi.f (a i)).
281 record is_order_continuous
282 (O:dedekind_sigma_complete_ordered_set) (f:O→O) : Prop
284 { ioc_is_sequentially_monotone: is_sequentially_monotone ? f;
285 ioc_is_upper_bound_f_sup:
286 ∀a:bounded_above_sequence O.
287 is_upper_bound ? (λi.f (a i)) (f (sup ? a));
289 ∀a:bounded_above_sequence O.
292 sup ? (mk_bounded_above_sequence ? (λi.f (a i))
293 (mk_is_bounded_above ? ? (f (sup ? a))
294 (ioc_is_upper_bound_f_sup a)));
295 ioc_is_lower_bound_f_inf:
296 ∀a:bounded_below_sequence O.
297 is_lower_bound ? (λi.f (a i)) (f (inf ? a));
299 ∀a:bounded_below_sequence O.
302 inf ? (mk_bounded_below_sequence ? (λi.f (a i))
303 (mk_is_bounded_below ? ? (f (inf ? a))
304 (ioc_is_lower_bound_f_inf a)))
307 theorem tail_inf_increasing:
308 ∀O:dedekind_sigma_complete_ordered_set.
309 ∀a:bounded_below_sequence O.
310 let y ≝ λi.mk_bounded_below_sequence ? (λj.a (i+j)) ? in
311 let x ≝ λi.inf ? (y i) in
313 [ apply (mk_is_bounded_below ? ? (ib_lower_bound ? a a));
316 apply (ib_lower_bound_is_lower_bound ? a a)
318 unfold is_increasing;
320 unfold x in ⊢ (? ? ? %);
321 apply (inf_greatest_lower_bound ? ? ? (inf_is_inf ? (y (S n))));
322 change with (is_lower_bound ? (y (S n)) (inf ? (y n)));
323 unfold is_lower_bound;
325 generalize in match (inf_lower_bound ? ? ? (inf_is_inf ? (y n)) (S n1));
326 (*CSC: coercion per FunClass inserita a mano*)
327 suppose (inf ? (y n) ≤ bbs_seq ? (y n) (S n1)) (H);
328 cut (bbs_seq ? (y n) (S n1) = bbs_seq ? (y (S n)) n1);
333 autobatch paramodulation library
338 definition is_liminf:
339 ∀O:dedekind_sigma_complete_ordered_set.
340 bounded_below_sequence O → O → Prop.
343 (is_sup ? (λi.inf ? (mk_bounded_below_sequence ? (λj.b (i+j)) ?)) t);
344 apply (mk_is_bounded_below ? ? (ib_lower_bound ? b b));
347 apply (ib_lower_bound_is_lower_bound ? b b).
351 ∀O:dedekind_sigma_complete_ordered_set.
352 bounded_sequence O → O.
356 (mk_bounded_above_sequence ?
358 (mk_bounded_below_sequence ?
359 (λj.b (i+j)) ?)) ?));
360 [ apply (mk_is_bounded_below ? ? (ib_lower_bound ? b b));
363 apply (ib_lower_bound_is_lower_bound ? b b)
364 | apply (mk_is_bounded_above ? ? (ib_upper_bound ? b b));
365 unfold is_upper_bound;
369 (mk_bounded_below_sequence O (\lambda j:nat.b (n+j))
370 (mk_is_bounded_below O (\lambda j:nat.b (n+j)) (ib_lower_bound O b b)
371 (\lambda j:nat.ib_lower_bound_is_lower_bound O b b (n+j))))
372 \leq ib_upper_bound O b b);
373 apply (inf_respects_le O);
376 apply (ib_upper_bound_is_upper_bound ? b b)
380 definition reverse_ordered_set: ordered_set → ordered_set.
382 apply mk_ordered_set;
383 [2:apply (λx,y:o.y ≤ x)
385 | apply mk_is_order_relation;
388 apply (or_reflexive ? ? o)
391 apply (or_transitive ? ? o);
398 apply (or_antisimmetric ? ? o);
404 interpretation "Ordered set ge" 'geq a b =
405 (cic:/matita/ordered_sets/os_le.con _
406 (cic:/matita/ordered_sets/os_pre_ordered_set.con _
407 (cic:/matita/ordered_sets/reverse_ordered_set.con _ _)) a b).
409 lemma is_lower_bound_reverse_is_upper_bound:
410 ∀O:ordered_set.∀a:nat→O.∀l:O.
411 is_lower_bound O a l → is_upper_bound (reverse_ordered_set O) a l.
416 unfold reverse_ordered_set;
421 lemma is_upper_bound_reverse_is_lower_bound:
422 ∀O:ordered_set.∀a:nat→O.∀l:O.
423 is_upper_bound O a l → is_lower_bound (reverse_ordered_set O) a l.
428 unfold reverse_ordered_set;
433 lemma reverse_is_lower_bound_is_upper_bound:
434 ∀O:ordered_set.∀a:nat→O.∀l:O.
435 is_lower_bound (reverse_ordered_set O) a l → is_upper_bound O a l.
438 unfold reverse_ordered_set in H;
442 lemma reverse_is_upper_bound_is_lower_bound:
443 ∀O:ordered_set.∀a:nat→O.∀l:O.
444 is_upper_bound (reverse_ordered_set O) a l → is_lower_bound O a l.
447 unfold reverse_ordered_set in H;
452 lemma is_inf_to_reverse_is_sup:
453 ∀O:ordered_set.∀a:bounded_below_sequence O.∀l:O.
454 is_inf O a l → is_sup (reverse_ordered_set O) a l.
456 apply (mk_is_sup (reverse_ordered_set O));
457 [ apply is_lower_bound_reverse_is_upper_bound;
458 apply inf_lower_bound;
461 change in v with (os_carrier O);
463 apply (inf_greatest_lower_bound ? ? ? H);
464 apply reverse_is_upper_bound_is_lower_bound;
469 lemma is_sup_to_reverse_is_inf:
470 ∀O:ordered_set.∀a:bounded_above_sequence O.∀l:O.
471 is_sup O a l → is_inf (reverse_ordered_set O) a l.
473 apply (mk_is_inf (reverse_ordered_set O));
474 [ apply is_upper_bound_reverse_is_lower_bound;
475 apply sup_upper_bound;
478 change in v with (os_carrier O);
480 apply (sup_least_upper_bound ? ? ? H);
481 apply reverse_is_lower_bound_is_upper_bound;
486 lemma reverse_is_sup_to_is_inf:
487 ∀O:ordered_set.∀a:bounded_above_sequence O.∀l:O.
488 is_sup (reverse_ordered_set O) a l → is_inf O a l.
491 [ apply reverse_is_upper_bound_is_lower_bound;
492 change in l with (os_carrier (reverse_ordered_set O));
493 apply sup_upper_bound;
496 change in l with (os_carrier (reverse_ordered_set O));
497 change in v with (os_carrier (reverse_ordered_set O));
498 change with (os_le (reverse_ordered_set O) l v);
499 apply (sup_least_upper_bound ? ? ? H);
500 change in v with (os_carrier O);
501 apply is_lower_bound_reverse_is_upper_bound;
506 lemma reverse_is_inf_to_is_sup:
507 ∀O:ordered_set.∀a:bounded_above_sequence O.∀l:O.
508 is_inf (reverse_ordered_set O) a l → is_sup O a l.
511 [ apply reverse_is_lower_bound_is_upper_bound;
512 change in l with (os_carrier (reverse_ordered_set O));
513 apply (inf_lower_bound ? ? ? H)
515 change in l with (os_carrier (reverse_ordered_set O));
516 change in v with (os_carrier (reverse_ordered_set O));
517 change with (os_le (reverse_ordered_set O) v l);
518 apply (inf_greatest_lower_bound ? ? ? H);
519 change in v with (os_carrier O);
520 apply is_upper_bound_reverse_is_lower_bound;
526 definition reverse_dedekind_sigma_complete_ordered_set:
527 dedekind_sigma_complete_ordered_set → dedekind_sigma_complete_ordered_set.
529 apply mk_dedekind_sigma_complete_ordered_set;
530 [ apply (reverse_ordered_set d)
532 (*apply mk_is_dedekind_sigma_complete;
534 elim (dsc_sup ? ? d a m) 0;
535 [ generalize in match H; clear H;
536 generalize in match m; clear m;
540 change in a1 with (Type_OF_ordered_set ? (reverse_ordered_set ? o));
541 apply (ex_intro ? ? a1);
543 change in m with (Type_OF_ordered_set ? o);
544 apply (is_sup_to_reverse_is_inf ? ? ? ? H1)
545 | generalize in match H; clear H;
546 generalize in match m; clear m;
549 change in t with (Type_OF_ordered_set ? o);
551 apply reverse_is_lower_bound_is_upper_bound;
554 | apply is_sup_reverse_is_inf;
561 definition reverse_bounded_sequence:
562 ∀O:dedekind_sigma_complete_ordered_set.
564 bounded_sequence (reverse_dedekind_sigma_complete_ordered_set O).
566 apply mk_bounded_sequence;
568 unfold reverse_dedekind_sigma_complete_ordered_set;
577 λO:dedekind_sigma_complete_ordered_set.
578 λa:bounded_sequence O.
579 liminf (reverse_dedekind_sigma_complete_ordered_set O)
580 (reverse_bounded_sequence O a).
582 notation "hvbox(〈a〉)"
583 non associative with precedence 45
584 for @{ 'hide_everything_but $a }.
586 interpretation "mk_bounded_above_sequence" 'hide_everything_but a
587 = (cic:/matita/ordered_sets/bounded_above_sequence.ind#xpointer(1/1/1) _ _ a _).
589 interpretation "mk_bounded_below_sequence" 'hide_everything_but a
590 = (cic:/matita/ordered_sets/bounded_below_sequence.ind#xpointer(1/1/1) _ _ a _).
592 theorem eq_f_sup_sup_f:
593 ∀O':dedekind_sigma_complete_ordered_set.
594 ∀f:O'→O'. ∀H:is_order_continuous ? f.
595 ∀a:bounded_above_sequence O'.
596 ∀p:is_increasing ? a.
597 f (sup ? a) = sup ? (mk_bounded_above_sequence ? (λi.f (a i)) ?).
598 [ apply (mk_is_bounded_above ? ? (f (sup ? a)));
599 apply ioc_is_upper_bound_f_sup;
602 apply ioc_respects_sup;
607 theorem eq_f_sup_sup_f':
608 ∀O':dedekind_sigma_complete_ordered_set.
609 ∀f:O'→O'. ∀H:is_order_continuous ? f.
610 ∀a:bounded_above_sequence O'.
611 ∀p:is_increasing ? a.
612 ∀p':is_bounded_above ? (λi.f (a i)).
613 f (sup ? a) = sup ? (mk_bounded_above_sequence ? (λi.f (a i)) p').
615 rewrite > (eq_f_sup_sup_f ? f H a H1);
616 apply sup_proof_irrelevant;
620 theorem eq_f_liminf_sup_f_inf:
621 ∀O':dedekind_sigma_complete_ordered_set.
622 ∀f:O'→O'. ∀H:is_order_continuous ? f.
623 ∀a:bounded_sequence O'.
627 (mk_bounded_above_sequence ?
629 (mk_bounded_below_sequence ?
633 [ apply (mk_is_bounded_below ? ? (ib_lower_bound ? a a));
636 apply (ib_lower_bound_is_lower_bound ? a a)
637 | apply (mk_is_bounded_above ? ? (f (sup ? a)));
638 unfold is_upper_bound;
640 apply (or_transitive ? ? O' ? (f (a n)));
641 [ generalize in match (ioc_is_lower_bound_f_inf ? ? H);
644 rewrite > (plus_n_O n) in ⊢ (? ? ? (? (? ? ? %)));
645 apply (H1 (mk_bounded_below_sequence O' (\lambda j:nat.a (n+j))
646 (mk_is_bounded_below O' (\lambda j:nat.a (n+j)) (ib_lower_bound O' a a)
647 (\lambda j:nat.ib_lower_bound_is_lower_bound O' a a (n+j)))) O);
648 | elim daemon (*apply (ioc_is_upper_bound_f_sup ? ? ? H)*)
653 generalize in match (\lambda n:nat
655 (mk_bounded_below_sequence O' (\lambda j:nat.a (plus n j))
656 (mk_is_bounded_below O' (\lambda j:nat.a (plus n j))
657 (ib_lower_bound O' a a)
658 (\lambda j:nat.ib_lower_bound_is_lower_bound O' a a (plus n j))))
659 (ib_upper_bound O' a a)
660 (\lambda n1:nat.ib_upper_bound_is_upper_bound O' a a (plus n n1)));
662 apply (eq_f_sup_sup_f' ? f H (mk_bounded_above_sequence O'
665 (mk_bounded_below_sequence O' (\lambda j:nat.a (plus i j))
666 (mk_is_bounded_below O' (\lambda j:nat.a (plus i j))
667 (ib_lower_bound O' a a)
668 (\lambda n:nat.ib_lower_bound_is_lower_bound O' a a (plus i n)))))
669 (mk_is_bounded_above O'
672 (mk_bounded_below_sequence O' (\lambda j:nat.a (plus i j))
673 (mk_is_bounded_below O' (\lambda j:nat.a (plus i j))
674 (ib_lower_bound O' a a)
675 (\lambda n:nat.ib_lower_bound_is_lower_bound O' a a (plus i n)))))
676 (ib_upper_bound O' a a) p2)));
679 (is_increasing ? (\lambda i:nat
681 (mk_bounded_below_sequence O' (\lambda j:nat.a (plus i j))
682 (mk_is_bounded_below O' (\lambda j:nat.a (plus i j))
683 (ib_lower_bound O' a a)
684 (\lambda n:nat.ib_lower_bound_is_lower_bound O' a a (plus i n))))));
685 apply tail_inf_increasing
692 definition lt ≝ λO:ordered_set.λa,b:O.a ≤ b ∧ a ≠ b.
694 interpretation "Ordered set lt" 'lt a b =
695 (cic:/matita/ordered_sets/lt.con _ a b).