X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=sidebyside;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fdama%2Fdama%2Fproperty_exhaustivity.ma;h=48551f85e70ab95c2867e65d4aea9b1e93fcd919;hb=f20f1ac4aea15d81599bd2283c5440fce8d4cf6a;hp=713588e7066f0feebcccdbe62f2cedc7179804bd;hpb=62571dd402d272b1632b7739607d25df3552cc04;p=helm.git diff --git a/helm/software/matita/contribs/dama/dama/property_exhaustivity.ma b/helm/software/matita/contribs/dama/dama/property_exhaustivity.ma index 713588e70..48551f85e 100644 --- a/helm/software/matita/contribs/dama/dama/property_exhaustivity.ma +++ b/helm/software/matita/contribs/dama/dama/property_exhaustivity.ma @@ -15,13 +15,6 @@ 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. @@ -54,18 +47,6 @@ notation "'segment_preserves_downarrow'" non associative with precedence 90 for 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: @@ -80,27 +61,64 @@ 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: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 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