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=66155917bd62757397bc324029d0baade2cc281f;hp=63c966056db54a2f07ad19786ba1329b2722e1c6;hpb=f71d9d74c33e4ab3d86dc8244a7bdd3576bb0ac0;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 63c966056..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 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. @@ -109,13 +110,24 @@ lemma bs2_of_bss2: coercion bs2_of_bss2 nocomposites. -(* -lemma xxx : - ∀O,s,x.bs2_of_bss2 (ordered_set_OF_ordered_uniform_space O) s x - = - x. -intros; reflexivity; -*) + +lemma a2sa : + ∀O:ordered_uniform_space.∀s:‡(ordered_set_OF_ordered_uniform_space O). + ∀x: + bs_carr + (square_bishop_set + (bishop_set_of_ordered_set + (segment_ordered_set + (ordered_set_OF_ordered_uniform_space O) s))). + (\fst x) ≈ (\snd x) → + (\fst (bs2_of_bss2 (ordered_set_OF_ordered_uniform_space O) s x)) + ≈ + (\snd (bs2_of_bss2 (ordered_set_OF_ordered_uniform_space O) s x)). +intros 3; cases x (a b); clear x; cases a (x Hx); cases b (y Hy); clear a b; +simplify; intros 2 (K H); apply K; clear K; whd; whd in H; cases H; +[left|right] apply x2sx; assumption; +qed. + lemma segment_ordered_uniform_space: ∀O:ordered_uniform_space.∀s:‡O.ordered_uniform_space. @@ -127,59 +139,83 @@ intros (O s); apply mk_ordered_uniform_space; apply (mk_uniform_space (bishop_set_of_ordered_set {[s]}) f); [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) as IH;[2:intro;apply Hm;cases H; clear H; - [left;apply (x2sx ? s (\fst x) (\snd x) H1); - |right;apply (x2sx ? s ?? H1);] - + lapply (us_phi1 O w Gw x (a2sa ??? Hm)) as IH; apply (restrict ? s ??? 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 O u v Gu Gv) (w HW); cases HW (Gw Hw); clear HW; - exists; [apply (λb:{[l,r]} squareB.w b)] split; + exists; [apply (λb:{[s]} 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 O l r ??? HuU H)|apply (restrict O l r ??? HvV H1);]] + [apply (restrict O s ??? HuU H)|apply (restrict O s ??? HvV H1);]] |3: intros (U Hu); cases Hu (u HU); cases HU (Gu HuU); clear Hu HU; cases (us_phi3 O u Gu) (w HW); cases HW (Gw Hwu); clear HW; - exists; [apply (λx:{[l,r]} squareB.w x)] split; + exists; [apply (λx:{[s]} 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 O l r ??? HuU); apply Hwu; + |2: intros 2 (x Hx); apply (restrict O s ??? 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 O u Gu x) (Hul Hur); split; intros; - [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);]] + [1: lapply (invert_restriction_agreement O s ?? HuU) as Ra; + apply (restrict O s ?? x Ra); + apply Hul; apply (unrestrict O s ??? HuU H); + |2: apply (restrict O s ??? HuU); apply Hur; + apply (unrestrict O s ??? (invert_restriction_agreement O s ?? 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 p ? H2 y H3) as Cu; - [1: apply (unrestrict O l r ??? HuU); apply H1; - |2: apply (restrict O l r ??? HuU Cu);]] +|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_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; + apply (x2sx (os_l O) s (\fst y) h1 H); + |5: change with (\fst h ≤ \fst (\fst y)); intro; apply H4; + apply (x2sx (os_l O) s h (\fst y) H); + |6: change with (\fst (\snd y) ≤ \fst h1); intro; apply H5; + apply (x2sx (os_l O) s (\snd y) h1 H); + |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);] +|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". lemma restric_uniform_convergence: - ∀O:ordered_uniform_space.∀l,u:O. - ∀x:{[l,u]}. - ∀a:sequence {[l,u]}. + ∀O:ordered_uniform_space.∀s:‡O. + ∀x:{[s]}. + ∀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 ≝