record ordered_uniform_space : Type ≝ {
ous_stuff :> ordered_uniform_space_;
- ous_convex: ∀U.us_unifbase ous_stuff U → convex ous_stuff U
+ ous_convex: ∀U.us_unifbase ous_stuff U → convex (os_l ous_stuff) U
}.
definition half_ordered_set_OF_ordered_uniform_space : ordered_uniform_space → half_ordered_set.
|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;
-
-lemma ls2l:
- ∀O:ordered_set.∀s:‡O.∀x,y:os_l (square_ordered_set (segment_ordered_set O s)).
- le (os_l (square_ordered_set (segment_ordered_set O s))) x y →
- \fst x ≤ \fst y.
-intros 4; cases x (a1 a2); cases y (b1 b2); clear x y;
-intros 2 (H K); apply H; clear H;
-simplify in K:(? ? ? %);
-simplify in K:(? ? % ?);
-whd in ⊢ (? (? (% ?)) ? ?);
-whd in ⊢ (? (? ((λ_:?.? ? ? (? ? (? (% ?)))) ?)) ? ?);
-simplify;
-whd in K:(% ? ? ?);
-simplify in K:(%);
-whd in ⊢ (? (% ?) ? ? ? ?);
-simplify in ⊢ (? (% ?) ? ? ? ?);
-simplify in ⊢ (? % ? ? ? ?);
-simplify in ⊢ (%);
-cases (wloss_prop (os_l (segment_ordered_set O s)));
-rewrite <H in K ⊢ %;
-whd in ⊢ (? % ? ?);
-simplify in ⊢ (%);
-unfold hos_excess; do 2 rewrite <H;
-
-SQUARE SEEMS NOT DUAL!
-
-cases a1 (a _); cases a2 (b _);
-cases b1 (c _); cases b2 (d _); clear a1 a2 b1 b2; simplify;
-intros 2 (H K); apply H; clear H; apply K;
-
- lapply (ous_convex ?? Gu p ? H2 y H3) as Cu;
- [1: apply (unrestrict O s ??? HuU); apply H1;
- |2: apply (restrict O s ??? 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 ?? 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);]]
qed.
interpretation "Ordered uniform space segment" 'segment_set a b =
(* 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;
alias symbol "pi2" = "pair pi2".
alias symbol "pi1" = "pair pi1".
definition square_segment ≝
- λO:ordered_set.λa,b:O.λx: O squareO.
- And4 (\fst x ≤ b) (a ≤ \fst x) (\snd x ≤ b) (a ≤ \snd x).
+ λO:half_ordered_set.λs:segment O.λx: square_half_ordered_set O.
+ in_segment ? s (\fst x) ∧ in_segment ? s (\snd x).
definition convex ≝
- λO:ordered_set.λU:O squareO → Prop.
- ∀p.U p → \fst p ≤ \snd p → ∀y.
- square_segment O (\fst p) (\snd p) y → U y.
+ λO:half_ordered_set.λU:square_half_ordered_set O → Prop.
+ ∀s.U s → le O (\fst s) (\snd s) →
+ ∀y.
+ le O (\fst y) (\snd s) →
+ le O (\fst s) (\fst y) →
+ le O (\snd y) (\snd s) →
+ le O (\fst y) (\snd y) →
+ U y.
(* Definition 2.11 *)
definition upper_located ≝