X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fdama%2Fdama%2Fproperty_exhaustivity.ma;h=96170fa262a6d25f553560e90bb76b5b3a11e5cb;hb=6d27950e804ea499909ae0fabceea99f35d118e9;hp=f7250f9e2e167a7a428b1d604c14638142580772;hpb=6e2dfd0a82ab76d3c0aeec5f6149e7ee5992d687;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 f7250f9e2..96170fa26 100644 --- a/helm/software/matita/contribs/dama/dama/property_exhaustivity.ma +++ b/helm/software/matita/contribs/dama/dama/property_exhaustivity.ma @@ -55,13 +55,13 @@ lemma segment_cauchy: intros 7; alias symbol "pi1" (instance 3) = "pair pi1". alias symbol "pi2" = "pair pi2". -apply (H (λx:{[l,u]} square.U 〈\fst (\fst x),\fst (\snd x)〉)); +apply (H (λx:{[l,u]} 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]} → @@ -69,39 +69,66 @@ lemma restrict_uniform_convergence_uparrow: 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 ? ??? ? (H2 O)); - 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 ≪x,h≫) as Ha1; - [2: apply segment_preserves_uparrow;split; assumption;] - lapply (segment_preserves_supremum ? l u a ≪?,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 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 restrict_uniform_convergence_downarrow: + +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. + +axiom 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 → + ∀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 ???? (H2 O)); - 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 ≪x,h≫) as Ha1; - [2: apply segment_preserves_downarrow;split; assumption;] - lapply (segment_preserves_infimum ?l u a ≪?,h≫) as Ha2; + lapply (downarrow_lowerlocated a ≪x,h≫) as Ha1; + [2: apply (segment_preserves_downarrow ? l u);split; assumption;] + lapply (segment_preserves_infimum l u); + [2: apply a; ≪?,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. +*) \ No newline at end of file