From c3b8ed2e554cbdb677729747c5b5a96112ae5169 Mon Sep 17 00:00:00 2001 From: Enrico Tassi Date: Mon, 9 Jun 2008 10:24:09 +0000 Subject: [PATCH] exhaustivity completed --- .../contribs/dama/dama/ordered_uniform.ma | 1 - .../dama/dama/property_exhaustivity.ma | 53 +++++++++++++++++-- .../matita/contribs/dama/dama/uniform.ma | 2 - 3 files changed, 48 insertions(+), 8 deletions(-) diff --git a/helm/software/matita/contribs/dama/dama/ordered_uniform.ma b/helm/software/matita/contribs/dama/dama/ordered_uniform.ma index 63f263d27..3bdfbe978 100644 --- a/helm/software/matita/contribs/dama/dama/ordered_uniform.ma +++ b/helm/software/matita/contribs/dama/dama/ordered_uniform.ma @@ -153,4 +153,3 @@ qed. definition order_continuity ≝ λC:ordered_uniform_space.∀a:sequence C.∀x:C. (a ↑ x → a uniform_converges x) ∧ (a ↓ x → a uniform_converges x). - diff --git a/helm/software/matita/contribs/dama/dama/property_exhaustivity.ma b/helm/software/matita/contribs/dama/dama/property_exhaustivity.ma index 173739f27..b9bb51e3b 100644 --- a/helm/software/matita/contribs/dama/dama/property_exhaustivity.ma +++ b/helm/software/matita/contribs/dama/dama/property_exhaustivity.ma @@ -21,14 +21,57 @@ definition exhaustive ≝ ∀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; +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. + +lemma segment_preserves_uparrow: + ∀C:ordered_set.∀l,u:C.∀a:sequence {[l,u]}.∀x,h. + (λn.fst (a n)) ↑ x → a ↑ (sig_in ?? x h). +intros; cases H (Ha Hx); split [apply Ha] cases Hx; +split; [apply H1] intros; +cases (H2 (fst y) H3); exists [apply w] assumption; +qed. + +(* Fact 2.18 *) +lemma segment_cauchy: + ∀C:ordered_uniform_space.∀l,u:C.∀a:sequence {[l,u]}. + a is_cauchy → (λn:nat.fst (a n)) is_cauchy. +intros 7; +alias symbol "pi1" (instance 3) = "pair pi1". +apply (H (λx:{[l,u]} square.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. + (* Lemma 3.8 *) -lemma xxx: +lemma restrict_uniform_convergence: ∀C:ordered_uniform_space.property_sigma C → ∀l,u:C.exhaustive {[l,u]} → ∀a:sequence {[l,u]}.∀x:C. (λn.fst (a n)) ↑ x → x∈[l,u] ∧ ∀h:x ∈ [l,u].a uniform_converges (sig_in ?? x h). -intros; split; -[1: (* manca fact 2.5 *) - +intros; cases H2 (Ha Hx); clear H2; cases Hx; split; +[1: split; + [1: apply (supremum_is_upper_bound C ?? Hx u); + apply (segment_upperbound ? l); + |2: apply (le_transitive ?? (fst (a 0))); [2: apply H2;] + apply (segment_lowerbound ?l u);] +|2: intros; + lapply (uparrow_upperlocated ? a (sig_in ?? x h)) as Ha1; + [2: apply segment_preserves_uparrow;split; assumption;] + lapply (segment_preserves_supremum ?l u a (sig_in ??? h)) as Ha2; + [2:split; assumption]; cases Ha2; clear Ha2; + cases (H1 a a); lapply (H6 H4 Ha1) as HaC; + lapply (segment_cauchy ? l u ? HaC) as Ha; + lapply (sigma_cauchy ? H ? x ? Ha); [split; assumption] + apply restric_uniform_convergence; assumption;] +qed. \ No newline at end of file diff --git a/helm/software/matita/contribs/dama/dama/uniform.ma b/helm/software/matita/contribs/dama/dama/uniform.ma index d5d714726..41b593ae4 100644 --- a/helm/software/matita/contribs/dama/dama/uniform.ma +++ b/helm/software/matita/contribs/dama/dama/uniform.ma @@ -139,5 +139,3 @@ interpretation "explicit big set" 'explicitset t = definition restrict_uniformity ≝ λC:uniform_space.λX:C→Prop. {U:C square → Prop| (U ⊆ {x:C square|X (fst x) ∧ X(snd x)}) ∧ us_unifbase C U}. - - -- 2.39.2