X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flibrary%2Fdama%2Fproperty_exhaustivity.ma;fp=matita%2Fmatita%2Flibrary%2Fdama%2Fproperty_exhaustivity.ma;h=8605105bf8e597c021761b0e45d3b7854ac268d3;hb=2c01ff6094173915e7023076ea48b5804dca7778;hp=0000000000000000000000000000000000000000;hpb=a050e3f80d7ea084ce0184279af98e8251c7d2a6;p=helm.git diff --git a/matita/matita/library/dama/property_exhaustivity.ma b/matita/matita/library/dama/property_exhaustivity.ma new file mode 100644 index 000000000..8605105bf --- /dev/null +++ b/matita/matita/library/dama/property_exhaustivity.ma @@ -0,0 +1,172 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "dama/ordered_uniform.ma". +include "dama/property_sigma.ma". + +lemma h_segment_upperbound: + ∀C:half_ordered_set. + ∀s:segment C. + ∀a:sequence (half_segment_ordered_set C s). + upper_bound ? ⌊n,\fst (a n)⌋ (seg_u C s). +intros 4; simplify; cases (a n); simplify; unfold in H; +cases (wloss_prop C); rewrite < H1 in H; simplify; 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}. + +interpretation "segment_upperbound" 'segment_upperbound = (h_segment_upperbound (os_l ?)). +interpretation "segment_lowerbound" 'segment_lowerbound = (h_segment_upperbound (os_r ?)). + +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 rule H; +| cases Hx; split; + [ intro n; intro H; apply (H1 n);apply rule H; + | intros; cases (H2 (\fst y)); [2: apply rule 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.∀s:‡C.∀a:sequence {[s]}. + a is_cauchy → ⌊n,\fst (a n)⌋ is_cauchy. +intros 6; +alias symbol "pi1" (instance 3) = "pair pi1". +alias symbol "pi2" = "pair pi2". +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 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; +[ apply (hle_transitive ??? x ? (H2 O)); lapply (H1 O) as K; unfold in K; rewrite