X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fdama%2Fdama%2Fproperty_exhaustivity.ma;h=713588e7066f0feebcccdbe62f2cedc7179804bd;hb=62571dd402d272b1632b7739607d25df3552cc04;hp=cf3a5b0574cfe099fd04f79cb29f62ef087cb416;hpb=5070f476ff80ee53fe444d284f9e7587a37022f4;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 cf3a5b057..713588e70 100644 --- a/helm/software/matita/contribs/dama/dama/property_exhaustivity.ma +++ b/helm/software/matita/contribs/dama/dama/property_exhaustivity.ma @@ -22,85 +22,129 @@ definition exhaustive ≝ (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 ↑ (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. +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 ↓ (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; +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 _)). + +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: - ∀C:ordered_uniform_space.∀l,u:C.∀a:sequence {[l,u]}. - a is_cauchy → (λn:nat.fst (a n)) is_cauchy. -intros 7; + ∀C:ordered_uniform_space.∀s:‡C.∀a:sequence {[s]}. + a is_cauchy → ⌊n,\fst (a n)⌋ is_cauchy. +intros 6; alias symbol "pi1" (instance 3) = "pair pi1". -apply (H (λx:{[l,u]} square.U 〈fst (fst x),fst (snd x)〉)); +alias symbol "pi2" = "pair pi2". +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. -(* Lemma 3.8 *) +(* Lemma 3.8 NON DUALIZZATO *) lemma restrict_uniform_convergence_uparrow: ∀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). + ∀a:sequence {[l,u]}.∀x:C. ⌊n,\fst (a n)⌋ ↑ x → + x∈[l,u] ∧ ∀h:x ∈ [l,u].a uniform_converges ≪x,h≫. intros; cases H2 (Ha Hx); clear H2; cases Hx; split; [1: split; - [1: apply (supremum_is_upper_bound C ?? Hx u); + [1: apply (supremum_is_upper_bound ? x Hx u); apply (segment_upperbound ? l); - |2: apply (le_transitive ?? (fst (a 0))); [2: apply H2;] - apply (segment_lowerbound ?l u);] + |2: apply (le_transitive l ? x ? (H2 O)); + apply (segment_lowerbound ? l u a 0);] |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; + lapply (uparrow_upperlocated a ≪x,h≫) as Ha1; + [2: apply (segment_preserves_uparrow C l u);split; assumption;] + lapply (segment_preserves_supremum C l u a ≪?,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); [left; split; assumption] apply restric_uniform_convergence; assumption;] qed. - + +lemma hint_mah1: + ∀C. Type_OF_ordered_uniform_space1 C → hos_carr (os_r C). + intros; assumption; qed. + +coercion hint_mah1 nocomposites. + +lemma hint_mah2: + ∀C. sequence (hos_carr (os_l C)) → sequence (hos_carr (os_r C)). + intros; assumption; qed. + +coercion hint_mah2 nocomposites. + +lemma hint_mah3: + ∀C. Type_OF_ordered_uniform_space C → hos_carr (os_r C). + intros; assumption; qed. + +coercion hint_mah3 nocomposites. + +lemma hint_mah4: + ∀C. sequence (hos_carr (os_r C)) → sequence (hos_carr (os_l C)). + intros; assumption; qed. + +coercion hint_mah4 nocomposites. + lemma restrict_uniform_convergence_downarrow: ∀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). + ∀a:sequence {[l,u]}.∀x: C. ⌊n,\fst (a n)⌋ ↓ x → + x∈[l,u] ∧ ∀h:x ∈ [l,u].a uniform_converges ≪x,h≫. intros; cases H2 (Ha Hx); clear H2; cases Hx; split; [1: split; - [2: apply (infimum_is_lower_bound C ?? Hx l); + [2: apply (infimum_is_lower_bound ? x Hx l); apply (segment_lowerbound ? l u); - |1: apply (le_transitive ?? (fst (a 0))); [apply H2;] - apply (segment_upperbound ? l u);] + |1: lapply (ge_transitive ? ? x ? (H2 O)); [apply u||assumption] + apply (segment_upperbound ? l u a 0);] |2: intros; - lapply (downarrow_lowerlocated ? a (sig_in ?? x h)) as Ha1; - [2: apply segment_preserves_downarrow;split; assumption;] - lapply (segment_preserves_infimum ?l u a (sig_in ??? h)) as Ha2; + lapply (downarrow_lowerlocated a ≪x,h≫) as Ha1; + [2: apply (segment_preserves_downarrow ? l u);split; assumption;] + lapply (segment_preserves_infimum C l u a ≪x,h≫) as Ha2; [2:split; assumption]; cases Ha2; clear Ha2; cases (H1 a a); lapply (H7 H4 Ha1) as HaC; lapply (segment_cauchy ? l u ? HaC) as Ha; lapply (sigma_cauchy ? H ? x ? Ha); [right; split; assumption] apply restric_uniform_convergence; assumption;] -qed. \ No newline at end of file +qed.