X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fdama%2Fdama%2Fproperty_exhaustivity.ma;h=48551f85e70ab95c2867e65d4aea9b1e93fcd919;hb=6abb01e8b00db927e16aa790354d1da57af7875b;hp=96170fa262a6d25f553560e90bb76b5b3a11e5cb;hpb=6d27950e804ea499909ae0fabceea99f35d118e9;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 96170fa26..48551f85e 100644 --- a/helm/software/matita/contribs/dama/dama/property_exhaustivity.ma +++ b/helm/software/matita/contribs/dama/dama/property_exhaustivity.ma @@ -15,73 +15,110 @@ 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: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