From 6abb01e8b00db927e16aa790354d1da57af7875b Mon Sep 17 00:00:00 2001 From: Enrico Tassi Date: Tue, 28 Oct 2008 15:35:05 +0000 Subject: [PATCH] lebesge works --- .../matita/contribs/dama/dama/lebesgue.ma | 176 +++++++++++------- .../matita/contribs/dama/dama/ordered_set.ma | 4 +- .../dama/dama/property_exhaustivity.ma | 152 ++++++--------- .../matita/contribs/dama/dama/sandwich.ma | 9 +- 4 files changed, 180 insertions(+), 161 deletions(-) diff --git a/helm/software/matita/contribs/dama/dama/lebesgue.ma b/helm/software/matita/contribs/dama/dama/lebesgue.ma index 2e6b0a1e4..5957235cc 100644 --- a/helm/software/matita/contribs/dama/dama/lebesgue.ma +++ b/helm/software/matita/contribs/dama/dama/lebesgue.ma @@ -16,102 +16,152 @@ include "sandwich.ma". include "property_exhaustivity.ma". +(* NOT DUALIZED *) +alias symbol "low" = "lower". +alias id "le" = "cic:/matita/dama/ordered_set/le.con". lemma order_converges_bigger_lowsegment: ∀C:ordered_set. - ∀a:sequence C.∀l,u:C.∀H:∀i:nat.a i ∈ [l,u]. - ∀x:C.∀p:a order_converges x. - ∀j.l ≤ (pi1exT23 ???? p) j. + ∀a:sequence (os_l C).∀s:segment C.∀H:∀i:nat.a i ∈ s. + ∀x:C.∀p:order_converge C a x. + ∀j. seg_l (os_l C) s (λl.le (os_l C) l (pi1exT23 ???? p j)). intros; cases p (xi yi Ux Dy Hxy); clear p; simplify; cases Ux (Ixi Sxi); clear Ux; cases Dy (Dyi Iyi); clear Dy; cases (Hxy j) (Ia Sa); clear Hxy; cases Ia (Da SSa); cases Sa (Inca SIa); clear Ia Sa; -intro H2; cases (SSa l H2) (w Hw); simplify in Hw; -cases (H (w+j)) (Hal Hau); apply (Hau Hw); +cases (wloss_prop (os_l C))(W W);rewrite (H1 ? (hos_excess_ O)); lapply (hos_cotransitive O ?? z H); [assumption] cases Hletin;[right|left]assumption;] qed. +*) definition dual_hos : half_ordered_set → half_ordered_set. intro; constructor 1; @@ -75,8 +77,6 @@ intro hos; constructor 1; [apply hos; | apply (dual_hos hos); | reflexivity] qed. -(* coercion half2full. *) - definition Type_of_ordered_set : ordered_set → Type. intro o; apply (hos_carr (os_l o)); qed. diff --git a/helm/software/matita/contribs/dama/dama/property_exhaustivity.ma b/helm/software/matita/contribs/dama/dama/property_exhaustivity.ma index 0e0f013e9..48551f85e 100644 --- a/helm/software/matita/contribs/dama/dama/property_exhaustivity.ma +++ b/helm/software/matita/contribs/dama/dama/property_exhaustivity.ma @@ -69,22 +69,37 @@ definition exhaustive ≝ (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; + ∀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 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

sym_plus in H1; apply (le_w_plus mb); assumption; |2: apply Hmb; apply (le_w_plus ma); assumption] - |2: simplify; apply (le_transitive (a i) ?? Lax Lxb); - |3: simplify; repeat split; try assumption; - [1: apply (le_transitive ?? (b i) Lax Lxb); - |2: intro X; cases (exc_coreflexive (a i) X)]] + |2,3: simplify; apply (le_transitive (a i) ?? Lax Lxb); + |4: apply (le_reflexive); + |5,6: assumption;] |1: apply HW; exists[apply l] simplify; split; [1: apply (us_phi1 ?? Gw); unfold; apply eq_reflexive; |2: apply Hma; rewrite > sym_plus in H1; apply (le_w_plus mb); assumption;]] -- 2.39.2