X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fdama%2Fdama%2Fproperty_exhaustivity.ma;h=7145b42886c802cc4d75574a77e3c64906e357ea;hb=bb7af347df386afcd3ea2adea8e7e982e3a5a253;hp=fc7121349d89c6917a7384710b211dc29a4a8e01;hpb=3f80c07b790d3abdb3aafcd1da33f14fa90ada25;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 fc7121349..7145b4288 100644 --- a/helm/software/matita/contribs/dama/dama/property_exhaustivity.ma +++ b/helm/software/matita/contribs/dama/dama/property_exhaustivity.ma @@ -19,10 +19,10 @@ 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; + 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}. @@ -68,23 +68,31 @@ definition exhaustive ≝ (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; +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