]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/dama/dama/property_exhaustivity.ma
WIP
[helm.git] / helm / software / matita / contribs / dama / dama / property_exhaustivity.ma
index 96170fa262a6d25f553560e90bb76b5b3a11e5cb..0e0f013e988ec9758afba613b1a8d6493928e80c 100644 (file)
 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 segment_upperbound:
-  ∀C:ordered_set.∀l,u:C.∀a:sequence {[l,u]}.u is_upper_bound ⌊n,\fst (a n)⌋.
-intros 5; change with (\fst (a n) ≤ u); cases (a n); cases H; assumption;
+lemma h_segment_upperbound:
+  ∀C:half_ordered_set.
+   ∀s:segment C.
+    ∀a:sequence (half_segment_ordered_set C s).
+     (seg_u C s) (upper_bound ? ⌊n,\fst (a n)⌋). 
+intros; cases (wloss_prop C); unfold; rewrite < H; simplify; intro n;
+cases (a n); simplify; unfold in H1; rewrite < H in H1; cases H1; 
+simplify in H2 H3; rewrite < H in H2 H3; assumption;
 qed.
 
-lemma segment_lowerbound:
-  ∀C:ordered_set.∀l,u:C.∀a:sequence {[l,u]}.l is_lower_bound ⌊n,\fst (a n)⌋.
-intros 5; change with (l ≤ \fst (a n)); cases (a n); cases H; assumption;
-qed.
+notation "'segment_upperbound'" non associative with precedence 90 for @{'segment_upperbound}.
+notation "'segment_lowerbound'" non associative with precedence 90 for @{'segment_lowerbound}.
 
-lemma segment_preserves_uparrow:
-  ∀C:ordered_set.∀l,u:C.∀a:sequence {[l,u]}.∀x,h. 
-    ⌊n,\fst (a n)⌋ ↑ x → a ↑ ≪x,h≫.
-intros; cases H (Ha Hx); split [apply Ha] cases Hx; 
-split; [apply H1] intros;
-cases (H2 (\fst y)); [2: apply H3;] exists [apply w] assumption;
-qed.
+interpretation "segment_upperbound" 'segment_upperbound = (h_segment_upperbound (os_l _)).
+interpretation "segment_lowerbound" 'segment_lowerbound = (h_segment_upperbound (os_r _)).
 
-lemma segment_preserves_downarrow:
-  ∀C:ordered_set.∀l,u:C.∀a:sequence {[l,u]}.∀x,h. 
-    ⌊n,\fst (a n)⌋ ↓ x → a ↓ ≪x,h≫.
-intros; cases H (Ha Hx); split [apply Ha] cases Hx; 
-split; [apply H1] intros;
-cases (H2 (\fst y));[2:apply H3]; exists [apply w] assumption;
+lemma h_segment_preserves_uparrow:
+  ∀C:half_ordered_set.∀s:segment C.∀a:sequence (half_segment_ordered_set C s).
+   ∀x,h. uparrow C ⌊n,\fst (a n)⌋ x → uparrow (half_segment_ordered_set C s) a ≪x,h≫.
+intros; cases H (Ha Hx); split;
+[ intro n; intro H; apply (Ha n); apply (sx2x ???? H);
+| cases Hx; split; 
+  [ intro n; intro H; apply (H1 n);apply (sx2x ???? H); 
+  | intros; cases (H2 (\fst y)); [2: apply (sx2x ???? H3);] 
+    exists [apply w] apply (x2sx ?? (a w) y H4);]]
 qed.
-    
+
+notation "'segment_preserves_uparrow'" non associative with precedence 90 for @{'segment_preserves_uparrow}.
+notation "'segment_preserves_downarrow'" non associative with precedence 90 for @{'segment_preserves_downarrow}.
+
+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 _)).
 (* Fact 2.18 *)
 lemma segment_cauchy:
-  ∀C:ordered_uniform_space.∀l,u:C.∀a:sequence {[l,u]}.
+  ∀C:ordered_uniform_space.∀s:‡C.∀a:sequence {[s]}.
     a is_cauchy → ⌊n,\fst (a n)⌋ is_cauchy.
-intros 7
+intros 6
 alias symbol "pi1" (instance 3) = "pair pi1".
 alias symbol "pi2" = "pair pi2".
-apply (H (λx:{[l,u]} squareB.U 〈\fst (\fst x),\fst (\snd x)〉));
+apply (H (λx:{[s]} squareB.U 〈\fst (\fst x),\fst (\snd x)〉));
 (unfold segment_ordered_uniform_space; simplify);
 exists [apply U] split; [assumption;]
 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≫.
-intros; cases H2 (Ha Hx); clear H2; cases Hx; split;
-[1: split;
+    ∀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: 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));
@@ -76,7 +153,7 @@ intros; cases H2 (Ha Hx); clear H2; cases Hx; split;
 |2: intros;
     lapply (uparrow_upperlocated a ≪x,h≫) as Ha1;
       [2: apply (segment_preserves_uparrow C l u);split; assumption;] 
-    lapply (segment_preserves_supremum l u a ≪?,h≫) as Ha2; 
+    lapply (segment_preserves_supremum 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;
@@ -108,12 +185,11 @@ lemma hint_mah4:
   
 coercion hint_mah4 nocomposites.
 
-axiom restrict_uniform_convergence_downarrow:
+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); 
@@ -122,13 +198,11 @@ intros; cases H2 (Ha Hx); clear H2; cases Hx; split;
         apply (segment_upperbound ? l u a 0);]
 |2: intros;
     lapply (downarrow_lowerlocated a ≪x,h≫) as Ha1;
-      [2: apply (segment_preserves_downarrow ? l u);split; assumption;] 
-    lapply (segment_preserves_infimum l u); 
-      [2: apply a; ≪?,h≫) as Ha2; 
+      [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;]
 qed.
-*)
\ No newline at end of file