X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fdama%2Fdama%2Fordered_uniform.ma;h=36ba2d287503736707f4911fe6407d6fcee08d27;hb=c33fae30b4ce40198b8e1889ea1c1b58697cd567;hp=1f4f6f6867a772e8b37e21d91e1c02b7446c47a4;hpb=858e703335b47065529d481891863d380a2156d7;p=helm.git diff --git a/helm/software/matita/contribs/dama/dama/ordered_uniform.ma b/helm/software/matita/contribs/dama/dama/ordered_uniform.ma index 1f4f6f686..36ba2d287 100644 --- a/helm/software/matita/contribs/dama/dama/ordered_uniform.ma +++ b/helm/software/matita/contribs/dama/dama/ordered_uniform.ma @@ -34,7 +34,8 @@ coercion ous_unifspace. record ordered_uniform_space : Type ≝ { ous_stuff :> ordered_uniform_space_; - ous_convex: ∀U.us_unifbase ous_stuff U → convex (os_l ous_stuff) U + ous_convex_l: ∀U.us_unifbase ous_stuff U → convex (os_l ous_stuff) U; + ous_convex_r: ∀U.us_unifbase ous_stuff U → convex (os_r ous_stuff) U }. definition half_ordered_set_OF_ordered_uniform_space : ordered_uniform_space → half_ordered_set. @@ -170,7 +171,7 @@ intros (O s); apply mk_ordered_uniform_space; |2: simplify; unfold convex; intros 3; cases s1; intros; (* TODO: x2sx is for ≰, we need one for ≤ *) cases H (u HU); cases HU (Gu HuU); clear HU H; - lapply depth=0 (ous_convex ?? Gu 〈\fst h,\fst h1〉 ???????) as K3; + lapply depth=0 (ous_convex_l ?? Gu 〈\fst h,\fst h1〉 ???????) as K3; [2: intro; apply H2; apply (x2sx (os_l O) s h h1 H); |3: apply 〈\fst (\fst y),\fst (\snd y)〉; |4: intro; change in H with (\fst (\fst y) ≰ \fst h1); apply H3; @@ -182,11 +183,27 @@ intros (O s); apply mk_ordered_uniform_space; |7: change with (\fst (\fst y) ≤ \fst (\snd y)); intro; apply H6; apply (x2sx (os_l O) s (\fst y) (\snd y) H); |8: apply (restrict O s U u y HuU K3); - |1: apply (unrestrict O s ?? 〈h,h1〉 HuU H1);]] + |1: apply (unrestrict O s ?? 〈h,h1〉 HuU H1);] +|3: simplify; unfold convex; intros 3; cases s1; intros; (* TODO *) + cases H (u HU); cases HU (Gu HuU); clear HU H; + lapply depth=0 (ous_convex_r ?? Gu 〈\fst h,\fst h1〉 ???????) as K3; + [2: intro; apply H2; apply (x2sx (os_r O) s h h1 H); + |3: apply 〈\fst (\fst y),\fst (\snd y)〉; + |4: intro; (*change in H with (\fst (\fst y) ≱ \fst h1);*) apply H3; + apply (x2sx (os_r O) s (\fst y) h1 H); + |5: (*change with (\fst h ≥ \fst (\fst y));*) intro; apply H4; + apply (x2sx (os_r O) s h (\fst y) H); + |6: (*change with (\fst (\snd y) ≤ \fst h1);*) intro; apply H5; + apply (x2sx (os_r O) s (\snd y) h1 H); + |7: (*change with (\fst (\fst y) ≤ \fst (\snd y));*) intro; apply H6; + apply (x2sx (os_r O) s (\fst y) (\snd y) H); + |8: apply (restrict O s U u y HuU K3); + |1: apply (unrestrict O s ?? 〈h,h1〉 HuU H1);] +] qed. -interpretation "Ordered uniform space segment" 'segment_set a b = - (segment_ordered_uniform_space _ a b). +interpretation "Ordered uniform space segment" 'segset a = + (segment_ordered_uniform_space _ a). (* Lemma 3.2 *) alias symbol "pi1" = "exT \fst". @@ -196,9 +213,9 @@ lemma restric_uniform_convergence: ∀a:sequence {[s]}. (⌊n, \fst (a n)⌋ : sequence O) uniform_converges (\fst x) → a uniform_converges x. -intros 8; cases H1; cases H2; clear H2 H1; +intros 7; cases H1; cases H2; clear H2 H1; cases (H ? H3) (m Hm); exists [apply m]; intros; -apply (restrict ? l u ??? H4); apply (Hm ? H1); +apply (restrict ? s ??? H4); apply (Hm ? H1); qed. definition order_continuity ≝