include "ordered_uniform.ma".
include "property_sigma.ma".
-(* Definition 3.7 *)
-definition exhaustive ≝
- λC:ordered_uniform_space.
- ∀a,b:sequence C.
- (a is_increasing → a is_upper_located → a is_cauchy) ∧
- (b is_decreasing → b is_lower_located → b is_cauchy).
-
lemma h_segment_upperbound:
∀C:half_ordered_set.
∀s:segment C.
interpretation "segment_preserves_uparrow" 'segment_preserves_uparrow = (h_segment_preserves_uparrow (os_l _)).
interpretation "segment_preserves_downarrow" 'segment_preserves_downarrow = (h_segment_preserves_uparrow (os_r _)).
-
-lemma hint_pippo:
- ∀C,s.
- sequence
- (Type_of_ordered_set
- (segment_ordered_set
- (ordered_set_OF_ordered_uniform_space C) s))
- →
- sequence (Type_OF_uniform_space (segment_ordered_uniform_space C s)). intros; assumption;
-qed.
-
-coercion hint_pippo nocomposites.
(* Fact 2.18 *)
lemma segment_cauchy:
intro; cases b; intros; simplify; split; intros; assumption;
qed.
+(* Definition 3.7 *)
+definition exhaustive ≝
+ λC:ordered_uniform_space.
+ ∀a,b:sequence C.
+ (a is_increasing → a is_upper_located → a is_cauchy) ∧
+ (b is_decreasing → b is_lower_located → b is_cauchy).
+
+lemma prove_in_segment:
+ ∀O:ordered_set.∀s:segment (os_l O).∀x:O.
+ 𝕝_s (λl.l ≤ x) → 𝕦_s (λu.x ≤ u) → x ∈ s.
+intros; unfold; cases (wloss_prop (os_l O)); rewrite < H2;
+split; assumption;
+qed.
+
+lemma under_wloss_upperbound:
+ ∀C:half_ordered_set.∀s:segment C.∀a:sequence C.
+ seg_u C s (upper_bound C a) →
+ ∀i.seg_u C s (λu.a i ≤≤ u).
+intros; unfold in H; unfold;
+cases (wloss_prop C); rewrite <H1 in H ⊢ %;
+apply (H i);
+qed.
+
+
(* Lemma 3.8 NON DUALIZZATO *)
lemma restrict_uniform_convergence_uparrow:
∀C:ordered_uniform_space.property_sigma C →
- ∀l,u:C.exhaustive {[l,u]} →
- ∀a:sequence {[l,u]}.∀x:C. ⌊n,\fst (a n)⌋ ↑ x →
- x∈[l,u] ∧ ∀h:x ∈ [l,u].a uniform_converges ≪x,h≫.
+ ∀s:segment (os_l C).exhaustive (segment_ordered_uniform_space C s) →
+ ∀a:sequence (segment_ordered_uniform_space C s).
+ ∀x:C. ⌊n,\fst (a n)⌋ ↑ x →
+ in_segment (os_l C) s x ∧ ∀h:x ∈ s.a uniform_converges ≪x,h≫.
intros; cases H2 (Ha Hx); clear H2; cases Hx; split;
-[1: split;
+[1: apply prove_in_segment;
+ lapply depth=0 (under_wloss_upperbound (os_l C) ?? (segment_upperbound s a) O) as W1;
+ lapply depth=0 (under_wloss_upperbound (os_r C) ?? (h_segment_upperbound (os_r C) s a) O) as W2;
+ lapply (H2 O); simplify in Hletin; simplify in W2 W1;
+ cases a in Hletin W2 W1; simplify; cases (f O); simplify; intros;
+ whd in H6:(? % ? ? ? ?);
+ simplify in H6:(%);
+ cases (wloss_prop (os_l C)); rewrite <H8 in H5 H6 ⊢ %;
+ [ change in H6 with (le (os_l C) (seg_l_ (os_l C) s) w);
+ apply (le_transitive ??? H6 H7);
+ | apply (le_transitive (seg_u_ (os_l C) s) w x H6 H7);
+ |
+ lapply depth=0 (supremum_is_upper_bound ? x Hx (seg_u_ (os_l C) s)) as K;
+ lapply depth=0 (under_wloss_upperbound (os_l C) ?? (segment_upperbound s a));
+ apply K; intro; lapply (Hletin n); unfold; unfold in Hletin1;
+ rewrite < H8 in Hletin1; intro; apply Hletin1; clear Hletin1; apply H9;
+ | lapply depth=0 (h_supremum_is_upper_bound (os_r C) ⌊n,\fst (a n)⌋ x Hx (seg_l_ (os_r C) s)) as K;
+ lapply depth=0 (under_wloss_upperbound (os_r C) ?? (h_segment_upperbound (os_r C) s a));
+ apply K; intro; lapply (Hletin n); unfold; unfold in Hletin1;
+whd in Hletin1:(? % ? ? ? ?);
+simplify in Hletin1:(%);
+ rewrite < H8 in Hletin1; intro; apply Hletin1; clear Hletin1; apply H9;
+
+
+ apply (segment_upperbound ? l);
+ generalize in match (H2 O); generalize in match Hx; unfold supremum;
+ unfold upper_bound; whd in ⊢ (?→%→?); rewrite < H4;
+ split; unfold; rewrite < H4; simplify;
+ [1: lapply (infimum_is_lower_bound ? ? Hx u);
+
+
+
+split;
[1: apply (supremum_is_upper_bound ? x Hx u);
apply (segment_upperbound ? l);
|2: apply (le_transitive l ? x ? (H2 O));