X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fdama%2Fdama%2Fordered_uniform.ma;h=e724dc2e6ba4b7613b011973f483f89109096b63;hb=6e2dfd0a82ab76d3c0aeec5f6149e7ee5992d687;hp=0bc8a32255cde1804c708a588c4f25f772215ed6;hpb=f36588e673e67f0758fdbec52baa515a28fd9a7a;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 0bc8a3225..e724dc2e6 100644 --- a/helm/software/matita/contribs/dama/dama/ordered_uniform.ma +++ b/helm/software/matita/contribs/dama/dama/ordered_uniform.ma @@ -36,7 +36,7 @@ record ordered_uniform_space : Type ≝ { ous_stuff :> ordered_uniform_space_; ous_convex: ∀U.us_unifbase ous_stuff U → convex ous_stuff U }. - +(* definition Type_of_ordered_uniform_space : ordered_uniform_space → Type. intro; compose ordered_set_OF_ordered_uniform_space with os_l. apply (hos_carr (f o)); @@ -47,9 +47,9 @@ intro; compose ordered_set_OF_ordered_uniform_space with os_r. apply (hos_carr (f o)); qed. -coercion Type_of_ordered_uniform_space. coercion Type_of_ordered_uniform_space_dual. - +coercion Type_of_ordered_uniform_space. +*) definition half_ordered_set_OF_ordered_uniform_space : ordered_uniform_space → half_ordered_set. intro; compose ordered_set_OF_ordered_uniform_space with os_l. apply (f o); qed. @@ -143,39 +143,39 @@ intros (O l r); apply mk_ordered_uniform_space; [1: intros (U H); intro x; simplify; cases H (w Hw); cases Hw (Gw Hwp); clear H Hw; intro Hm; lapply (us_phi1 O w Gw x Hm) as IH; - apply (restrict ? l r ??? Hwp IH); STOP + apply (restrict ? l r ??? Hwp IH); |2: intros (U V HU HV); cases HU (u Hu); cases HV (v Hv); clear HU HV; cases Hu (Gu HuU); cases Hv (Gv HvV); clear Hu Hv; - cases (us_phi2 ??? Gu Gv) (w HW); cases HW (Gw Hw); clear HW; - exists; [apply (λb:{[l,r]} square.w b)] split; + cases (us_phi2 O u v Gu Gv) (w HW); cases HW (Gw Hw); clear HW; + exists; [apply (λb:{[l,r]} squareB.w b)] split; [1: unfold f; simplify; clearbody f; exists; [apply w]; split; [assumption] intro b; simplify; unfold segment_square_of_ordered_set_square; cases b; intros; split; intros; assumption; |2: intros 2 (x Hx); cases (Hw ? Hx); split; - [apply (restrict ?????? HuU H)|apply (restrict ?????? HvV H1);]] + [apply (restrict O l r ??? HuU H)|apply (restrict O l r ??? HvV H1);]] |3: intros (U Hu); cases Hu (u HU); cases HU (Gu HuU); clear Hu HU; - cases (us_phi3 ?? Gu) (w HW); cases HW (Gw Hwu); clear HW; - exists; [apply (λx:{[l,r]} square.w x)] split; + cases (us_phi3 O u Gu) (w HW); cases HW (Gw Hwu); clear HW; + exists; [apply (λx:{[l,r]} squareB.w x)] split; [1: exists;[apply w];split;[assumption] intros; simplify; intro; unfold segment_square_of_ordered_set_square; cases b; intros; split; intro; assumption; - |2: intros 2 (x Hx); apply (restrict ?????? HuU); apply Hwu; + |2: intros 2 (x Hx); apply (restrict O l r ??? HuU); apply Hwu; cases Hx (m Hm); exists[apply (\fst m)] apply Hm;] |4: intros (U HU x); cases HU (u Hu); cases Hu (Gu HuU); clear HU Hu; - cases (us_phi4 ?? Gu x) (Hul Hur); + cases (us_phi4 O u Gu x) (Hul Hur); split; intros; - [1: lapply (invert_restriction_agreement ????? HuU) as Ra; - apply (restrict ????? x Ra); - apply Hul; apply (unrestrict ?????? HuU H); - |2: apply (restrict ?????? HuU); apply Hur; - apply (unrestrict ?????? (invert_restriction_agreement ????? HuU) H);]] + [1: lapply (invert_restriction_agreement O l r ?? HuU) as Ra; + apply (restrict O l r ?? x Ra); + apply Hul; apply (unrestrict O l r ??? HuU H); + |2: apply (restrict O l r ??? HuU); apply Hur; + apply (unrestrict O l r ??? (invert_restriction_agreement O l r ?? HuU) H);]] |2: simplify; reflexivity;] |2: simplify; unfold convex; intros; cases H (u HU); cases HU (Gu HuU); clear HU H; - lapply (ous_convex ?? Gu (bs_of_ss ? l r p) ? H2 (bs_of_ss ? l r y) H3) as Cu; - [1: apply (unrestrict ?????? HuU); apply H1; - |2: apply (restrict ?????? HuU Cu);]] + lapply (ous_convex ?? Gu p ? H2 y H3) as Cu; + [1: apply (unrestrict O l r ??? HuU); apply H1; + |2: apply (restrict O l r ??? HuU Cu);]] qed. interpretation "Ordered uniform space segment" 'segment_set a b = @@ -185,9 +185,10 @@ interpretation "Ordered uniform space segment" 'segment_set a b = alias symbol "pi1" = "exT \fst". lemma restric_uniform_convergence: ∀O:ordered_uniform_space.∀l,u:O. - ∀x:{[l,u]}. - ∀a:sequence {[l,u]}. - ⌊n,\fst (a n)⌋ uniform_converges (\fst x) → + ∀x:(segment_ordered_uniform_space O l u). + ∀a:sequence (segment_ordered_uniform_space O l u). + uniform_converge (segment_ordered_uniform_space O l u) + (mk_seq O (λn:nat.\fst (a n))) (\fst x) → True. a uniform_converges x. intros 8; cases H1; cases H2; clear H2 H1; cases (H ? H3) (m Hm); exists [apply m]; intros;