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_sets2".
17 include "ordered_sets.ma".
19 theorem le_f_inf_inf_f:
20 ∀O':dedekind_sigma_complete_ordered_set.
21 ∀f:O'→O'. ∀H:is_order_continuous ? f.
22 ∀a:bounded_below_sequence O'.
24 f (inf ? a) ≤ inf ? (mk_bounded_below_sequence ? (λi. f (a i)) p).
25 [ apply mk_is_bounded_below;
26 [2: apply ioc_is_lower_bound_f_inf;
32 apply (inf_greatest_lower_bound ? ? ? (inf_is_inf ? ?));
35 letin b := (λi.match i with [ O ⇒ inf ? a | S _ ⇒ a n]);
36 change with (f (b O) ≤ f (b (S O)));
37 apply (ioc_is_sequentially_monotone ? ? H);
42 [ apply (inf_lower_bound ? ? ? (inf_is_inf ? ?));
43 | apply (or_reflexive O' ? (dscos_ordered_set O'))
48 theorem le_to_le_sup_sup:
49 ∀O':dedekind_sigma_complete_ordered_set.
50 ∀a,b:bounded_above_sequence O'.
51 (∀i.a i ≤ b i) → sup ? a ≤ sup ? b.
53 apply (sup_least_upper_bound ? ? ? (sup_is_sup ? a));
56 apply (or_transitive ? ? O');
59 | apply (sup_upper_bound ? ? ? (sup_is_sup ? b))
63 interpretation "mk_bounded_sequence" 'hide_everything_but a
64 = (cic:/matita/ordered_sets/bounded_sequence.ind#xpointer(1/1/1) _ _ a _ _).
67 ∀O:ordered_set.∀a:nat→O.∀p.∀i.
68 bas_seq ? (mk_bounded_above_sequence ? a p) i = a i.
73 (*lemma reduce_bbs_seq:
74 ∀C.∀O:ordered_set C.∀a:nat→O.∀p.∀i.
75 bbs_seq ? ? (mk_bounded_below_sequence ? ? a p) i = a i.
80 axiom inf_extensional:
81 ∀O:dedekind_sigma_complete_ordered_set.
82 ∀a,b:bounded_below_sequence O.
83 (∀i.a i = b i) → inf ? a = inf O b.
85 lemma eq_to_le: ∀O:ordered_set.∀x,y:O.x=y → x ≤ y.
88 apply (or_reflexive ? ? O).
92 ∀O':dedekind_sigma_complete_ordered_set.
93 ∀f:O'→O'. ∀H:is_order_continuous ? f.
94 ∀a:bounded_sequence O'.
97 f (liminf ? a) ≤ liminf ? (mk_bounded_sequence ? (λi. f (a i)) pb pa).
98 [ letin bas ≝ (bounded_above_sequence_of_bounded_sequence ? a);
99 apply mk_is_bounded_above;
100 [2: apply (ioc_is_upper_bound_f_sup ? ? H bas)
103 | letin bbs ≝ (bounded_below_sequence_of_bounded_sequence ? a);
104 apply mk_is_bounded_below;
105 [2: apply (ioc_is_lower_bound_f_inf ? ? H bbs)
109 rewrite > eq_f_liminf_sup_f_inf in ⊢ (? ? % ?);
111 apply le_to_le_sup_sup;
113 rewrite > reduce_bas_seq;
114 rewrite > reduce_bas_seq;
115 apply (or_transitive ? ? O');
116 [2: apply le_f_inf_inf_f;
120 apply inf_extensional;
129 record cotransitively_ordered_set: Type :=
130 { cos_ordered_set :> ordered_set;
131 cos_cotransitive: cotransitive ? (os_le cos_ordered_set)