]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/dama/dama/property_exhaustivity.ma
lebesge works
[helm.git] / helm / software / matita / contribs / dama / dama / property_exhaustivity.ma
index 0e0f013e988ec9758afba613b1a8d6493928e80c..48551f85e70ab95c2867e65d4aea9b1e93fcd919 100644 (file)
@@ -69,22 +69,37 @@ definition exhaustive ≝
      (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;
+ ∀O:half_ordered_set.∀s:segment O.∀x:O.
+   seg_l O s (λl.l ≤≤ x) → seg_u O s (λu.x ≤≤ u) → x ∈ s.
+intros; unfold; cases (wloss_prop 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 h_uparrow_to_in_segment:
+  ∀C:half_ordered_set.
+   ∀s:segment C.
+     ∀a:sequence C.
+      (∀i.a i ∈ s) →
+       ∀x:C. uparrow C a x → 
+         in_segment C s x.
+intros (C H a H1 x H2); unfold in H2; cases H2; clear H2;unfold in H3 H4; cases H4; clear H4; unfold in H2;
+cases (wloss_prop C) (W W); apply prove_in_segment; unfold; rewrite <W;simplify;
+[ apply (hle_transitive ??? x ? (H2 O)); lapply (H1 O) as K; unfold in K; rewrite <W in K;
+  cases K; unfold in H4 H6; rewrite <W in H6 H4; simplify in H4 H6; assumption;
+| intro; cases (H5 ? H4); clear H5 H4;lapply(H1 w) as K; unfold in K; rewrite<W in K;
+  cases K; unfold in H5 H4; rewrite<W in H4 H5; simplify in H4 H5; apply (H5 H6);    
+| apply (hle_transitive ??? x ?  (H2 O)); lapply (H1 0) as K; unfold in K; rewrite <W in K;
+  cases K; unfold in H4 H6; rewrite <W in H4 H6; simplify in H4 H6; assumption;
+| intro; cases (H5 ? H4); clear H5 H4;lapply(H1 w) as K; unfold in K; rewrite<W in K;
+  cases K; unfold in H5 H4; rewrite<W in H4 H5; simplify in H4 H5; apply (H4 H6);]    
+qed.
 
+notation "'uparrow_to_in_segment'" non associative with precedence 90 for @{'uparrow_to_in_segment}.
+notation "'downarrow_to_in_segment'" non associative with precedence 90 for @{'downarrow_to_in_segment}.
 
+interpretation "uparrow_to_in_segment" 'uparrow_to_in_segment = (h_uparrow_to_in_segment (os_l _)).
+interpretation "downarrow_to_in_segment" 'downarrow_to_in_segment = (h_uparrow_to_in_segment (os_r _)).
 (* Lemma 3.8 NON DUALIZZATO *)
 lemma restrict_uniform_convergence_uparrow:
   ∀C:ordered_uniform_space.property_sigma C →
@@ -93,72 +108,17 @@ lemma restrict_uniform_convergence_uparrow:
       ∀x:C. ⌊n,\fst (a n)⌋ ↑ x → 
        in_segment (os_l C) s x ∧ ∀h:x ∈ s.a uniform_converges ≪x,h≫.
 intros; split;
-[1: unfold in H2; cases H2; clear H2;unfold in H3 H4; cases H4; clear H4; unfold in H2;
-    cases (wloss_prop (os_l C)) (W W); apply prove_in_segment; unfold; rewrite <W;
-    simplify;
-    [ apply (le_transitive ?? x ? (H2 O));  
-      lapply (under_wloss_upperbound (os_r C) ?? (h_segment_upperbound (os_r C) s a) O) as K;
-      unfold in K; whd in K:(?%????); simplify in K; rewrite <W in K; apply K;
-    | intro; cases (H5 ? H4); clear H5 H4;    
-      lapply (under_wloss_upperbound (os_l C) ?? (segment_upperbound s a) w) as K;
-      unfold in K; whd in K:(?%????); simplify in K; rewrite <W in K;
-      apply K; apply H6;
-    | intro; unfold in H4; rewrite <W in H4;
-      lapply depth=0 (H5 (seg_u_ (os_l C) s)) as k; unfold in k:(%???→?);
-      simplify in k; rewrite <W in k; lapply (k
-     simplify;intro; cases (H5 ? H4); clear H5 H4;    
-      lapply (under_wloss_upperbound (os_l C) ?? (segment_upperbound s a) w) as K;
-      unfold in K; whd in K:(?%????); simplify in K; rewrite <W in K;
-      apply K; apply H6;    
-      
-      
-      
-    cases H2 (Ha Hx); clear H2; cases Hx; split; 
-    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));
-        apply (segment_lowerbound ? l u a 0);]
+[1: apply (uparrow_to_in_segment s ⌊n,\fst (a \sub n)⌋ ? x H2); 
+    simplify; intros; cases (a i); assumption;
 |2: intros;
     lapply (uparrow_upperlocated a ≪x,h≫) as Ha1;
-      [2: apply (segment_preserves_uparrow C l u);split; assumption;] 
-    lapply (segment_preserves_supremum C l u a ≪?,h≫) as Ha2; 
-      [2:split; assumption]; cases Ha2; clear Ha2;
-    cases (H1 a a); lapply (H6 H4 Ha1) as HaC;
-    lapply (segment_cauchy ? l u ? HaC) as Ha;
-    lapply (sigma_cauchy ? H  ? x ? Ha); [left; split; assumption]
-    apply restric_uniform_convergence; assumption;]
+      [2: apply (segment_preserves_uparrow s); assumption;] 
+    lapply (segment_preserves_supremum s a ≪?,h≫ H2) as Ha2; 
+    cases Ha2; clear Ha2;
+    cases (H1 a a); lapply (H5 H3 Ha1) as HaC;
+    lapply (segment_cauchy C s ? HaC) as Ha;
+    lapply (sigma_cauchy ? H  ? x ? Ha); [left; assumption]
+    apply (restric_uniform_convergence C s ≪x,h≫ a Hletin)]
 qed.
 
 lemma hint_mah1:
@@ -185,24 +145,34 @@ lemma hint_mah4:
   
 coercion hint_mah4 nocomposites.
 
+lemma hint_mah5:
+  ∀C. segment (hos_carr (os_r C)) → segment (hos_carr (os_l C)).
+  intros; assumption; qed.
+  
+coercion hint_mah5 nocomposites.
+
+lemma hint_mah6:
+  ∀C. segment (hos_carr (os_l C)) → segment (hos_carr (os_r C)).
+  intros; assumption; qed.
+
+coercion hint_mah6 nocomposites.
+
 lemma restrict_uniform_convergence_downarrow:
   ∀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≫.
-intros; cases H2 (Ha Hx); clear H2; cases Hx; split;
-[1: split;
-    [2: apply (infimum_is_lower_bound ? x Hx l); 
-        apply (segment_lowerbound ? l u);
-    |1: lapply (ge_transitive ? ? x ? (H2 O)); [apply u||assumption]
-        apply (segment_upperbound ? l u a 0);]
+    ∀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; split;       
+[1: apply (downarrow_to_in_segment s ⌊n,\fst (a n)⌋ ? x); [2: apply H2]; 
+    simplify; intros; cases (a i); assumption;
 |2: intros;
     lapply (downarrow_lowerlocated a ≪x,h≫) as Ha1;
-      [2: apply (segment_preserves_downarrow ? l u);split; assumption;]
-          lapply (segment_preserves_infimum C l u a ≪x,h≫) as Ha2; 
-      [2:split; assumption]; cases Ha2; clear Ha2;
-    cases (H1 a a); lapply (H7 H4 Ha1) as HaC;
-    lapply (segment_cauchy ? l u ? HaC) as Ha;
-    lapply (sigma_cauchy ? H  ? x ? Ha); [right; split; assumption]
-    apply restric_uniform_convergence; assumption;]
+      [2: apply (segment_preserves_downarrow s a x h H2);] 
+    lapply (segment_preserves_infimum s a ≪?,h≫ H2) as Ha2; 
+    cases Ha2; clear Ha2;
+    cases (H1 a a); lapply (H6 H3 Ha1) as HaC;
+    lapply (segment_cauchy C s ? HaC) as Ha;
+    lapply (sigma_cauchy ? H  ? x ? Ha); [right; assumption]
+    apply (restric_uniform_convergence C s ≪x,h≫ a Hletin)]
 qed.