]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/dama/dama/ordered_uniform.ma
...
[helm.git] / helm / software / matita / contribs / dama / dama / ordered_uniform.ma
index 0bc8a32255cde1804c708a588c4f25f772215ed6..63c966056db54a2f07ad19786ba1329b2722e1c6 100644 (file)
@@ -37,19 +37,6 @@ record ordered_uniform_space : Type ≝ {
   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));
-qed.
-
-definition Type_of_ordered_uniform_space_dual : ordered_uniform_space → Type.
-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.
-
 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.
@@ -62,9 +49,17 @@ interpretation "relation invertion" 'invert a = (invert_os_relation _ a).
 interpretation "relation invertion" 'invert_symbol = (invert_os_relation _).
 interpretation "relation invertion" 'invert_appl a x = (invert_os_relation _ a x).
 
+lemma hint_segment: ∀O.
+ segment (Type_of_ordered_set O) →
+ segment (hos_carr (os_l O)).
+intros; assumption;
+qed.
+
+coercion hint_segment nocomposites.
+
 lemma segment_square_of_ordered_set_square: 
-  ∀O:ordered_set.∀u,v:O.∀x:O squareO.
-   \fst x ∈ [u,v] → \snd x ∈ [u,v] → {[u,v]} squareO.
+  ∀O:ordered_set.∀s:‡O.∀x:O squareO.
+   \fst x ∈ s → \snd x ∈ s → {[s]} squareO.
 intros; split; exists; [1: apply (\fst x) |3: apply (\snd x)] assumption;
 qed.
 
@@ -73,109 +68,102 @@ coercion segment_square_of_ordered_set_square with 0 2 nocomposites.
 alias symbol "pi1" (instance 4) = "exT \fst".
 alias symbol "pi1" (instance 2) = "exT \fst".
 lemma ordered_set_square_of_segment_square : 
- ∀O:ordered_set.∀u,v:O.{[u,v]} squareO → O squareO ≝ 
-  λO:ordered_set.λu,v:O.λb:{[u,v]} squareO.〈\fst(\fst b),\fst(\snd b)〉.
+ ∀O:ordered_set.∀s:‡O.{[s]} squareO → O squareO ≝ 
+  λO:ordered_set.λs:‡O.λb:{[s]} squareO.〈\fst(\fst b),\fst(\snd b)〉.
 
 coercion ordered_set_square_of_segment_square nocomposites.
 
 lemma restriction_agreement : 
-  ∀O:ordered_uniform_space.∀l,r:O.∀P:{[l,r]} squareO → Prop.∀OP:O squareO → Prop.Prop.
-apply(λO:ordered_uniform_space.λl,r:O.
-       λP:{[l,r]} squareO → Prop. λOP:O squareO → Prop.
+  ∀O:ordered_uniform_space.∀s:‡O.∀P:{[s]} squareO → Prop.∀OP:O squareO → Prop.Prop.
+apply(λO:ordered_uniform_space.λs:‡O.
+       λP:{[s]} squareO → Prop. λOP:O squareO → Prop.
           ∀b:O squareO.∀H1,H2.(P b → OP b) ∧ (OP b → P b));
 [5,7: apply H1|6,8:apply H2]skip;
 qed.
 
-lemma unrestrict: ∀O:ordered_uniform_space.∀l,r:O.∀U,u.∀x:{[l,r]} squareO.
-  restriction_agreement ? l r U u → U x → u x.
-intros 7; cases x (b b1); cases b (w1 H1); cases b1 (w2 H2); clear b b1 x; 
+lemma unrestrict: ∀O:ordered_uniform_space.∀s:‡O.∀U,u.∀x:{[s]} squareO.
+  restriction_agreement ? s U u → U x → u x.
+intros 6; cases x (b b1); cases b (w1 H1); cases b1 (w2 H2); clear b b1 x; 
 cases (H 〈w1,w2〉 H1 H2) (L _); intro Uw; apply L; apply Uw;
 qed.
 
-lemma restrict: ∀O:ordered_uniform_space.∀l,r:O.∀U,u.∀x:{[l,r]} squareO.
-  restriction_agreement ? l r U u → u x → U x.
-intros 6; cases x (b b1); cases b (w1 H1); cases b1 (w2 H2); clear b1 b x;
+lemma restrict: ∀O:ordered_uniform_space.∀s:‡O.∀U,u.∀x:{[s]} squareO.
+  restriction_agreement ? s U u → u x → U x.
+intros 5; cases x (b b1); cases b (w1 H1); cases b1 (w2 H2); clear b1 b x;
 intros (Ra uw); cases (Ra 〈w1,w2〉 H1 H2) (_ R); apply R; apply uw;
 qed.
 
 lemma invert_restriction_agreement: 
-  ∀O:ordered_uniform_space.∀l,r:O.
-   ∀U:{[l,r]} squareO → Prop.∀u:O squareO → Prop.
-    restriction_agreement ? l r U u →
-    restriction_agreement ? l r (\inv U) (\inv u).
-intros 9; split; intro;
-[1: apply (unrestrict ????? (segment_square_of_ordered_set_square ??? 〈\snd b,\fst b〉 H2 H1) H H3);
-|2: apply (restrict ????? (segment_square_of_ordered_set_square ??? 〈\snd b,\fst b〉 H2 H1) H H3);]
+  ∀O:ordered_uniform_space.∀s:‡O.
+   ∀U:{[s]} squareO → Prop.∀u:O squareO → Prop.
+    restriction_agreement ? s U u →
+    restriction_agreement ? s (\inv U) (\inv u).
+intros 8; split; intro;
+[1: apply (unrestrict ???? (segment_square_of_ordered_set_square ?? 〈\snd b,\fst b〉 H2 H1) H H3);
+|2: apply (restrict ???? (segment_square_of_ordered_set_square ?? 〈\snd b,\fst b〉 H2 H1) H H3);]
 qed. 
 
 lemma bs2_of_bss2: 
- ∀O:ordered_set.∀u,v:O.(bishop_set_of_ordered_set {[u,v]}) squareB → (bishop_set_of_ordered_set O) squareB ≝ 
-  λO:ordered_set.λu,v:O.λb:{[u,v]} squareO.〈\fst(\fst b),\fst(\snd b)〉.
+ ∀O:ordered_set.∀s:‡O.(bishop_set_of_ordered_set {[s]}) squareB → (bishop_set_of_ordered_set O) squareB ≝ 
+  λO:ordered_set.λs:‡O.λb:{[s]} squareO.〈\fst(\fst b),\fst(\snd b)〉.
 
 coercion bs2_of_bss2 nocomposites.
 
 (*
-notation < "x \sub \neq" with precedence 91 for @{'bsss $x}.
-interpretation "bs_of_ss" 'bsss x = (bs_of_ss _ _ _ x).
-*)
-
-(*
-lemma ss_of_bs: 
- ∀O:ordered_set.∀u,v:O.
-  ∀b:O squareO.\fst b ∈ [u,v] → \snd b ∈ [u,v] → {[u,v]} squareO ≝ 
- λO:ordered_set.λu,v:O.
-  λb:O squareB.λH1,H2.〈≪\fst b,H1≫,≪\snd b,H2≫〉.
-*)
-
-(*
-notation < "x \sub \nleq" with precedence 91 for @{'ssbs $x}.
-interpretation "ss_of_bs" 'ssbs x = (ss_of_bs _ _ _ x _ _).
+lemma xxx :
+ ∀O,s,x.bs2_of_bss2 (ordered_set_OF_ordered_uniform_space O) s x
+              =
+              x.
+intros; reflexivity;
 *)
 
 lemma segment_ordered_uniform_space: 
-  ∀O:ordered_uniform_space.∀u,v:O.ordered_uniform_space.
-intros (O l r); apply mk_ordered_uniform_space;
-[1: apply (mk_ordered_uniform_space_ {[l,r]});
+  ∀O:ordered_uniform_space.∀s:‡O.ordered_uniform_space.
+intros (O s); apply mk_ordered_uniform_space;
+[1: apply (mk_ordered_uniform_space_ {[s]});
     [1: alias symbol "and" = "constructive and".
-        letin f ≝ (λP:{[l,r]} squareO → Prop. ∃OP:O squareO → Prop.
-                    (us_unifbase O OP) ∧ restriction_agreement ??? P OP);
-        apply (mk_uniform_space (bishop_set_of_ordered_set {[l,r]}) f);
+        letin f ≝ (λP:{[s]} squareO → Prop. ∃OP:O squareO → Prop.
+                    (us_unifbase O OP) ∧ restriction_agreement ?? P OP);
+        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 Hm) as IH;
-            apply (restrict ? l r ??? Hwp IH); STOP
+            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);]
+
+            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 ??? 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 = 
@@ -187,7 +175,7 @@ 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) → 
+    (⌊n, \fst (a n)⌋ : sequence O) uniform_converges (\fst x) → 
       a uniform_converges x.
 intros 8; cases H1; cases H2; clear H2 H1;
 cases (H ? H3) (m Hm); exists [apply m]; intros; 
@@ -197,3 +185,16 @@ qed.
 definition order_continuity ≝
   λC:ordered_uniform_space.∀a:sequence C.∀x:C.
     (a ↑ x → a uniform_converges x) ∧ (a ↓ x → a uniform_converges x).
+
+lemma hint_boh1: ∀O. Type_OF_ordered_uniform_space O → hos_carr (os_l O).
+intros; assumption;
+qed.
+
+coercion hint_boh1 nocomposites. 
+
+lemma hint_boh2: ∀O:ordered_uniform_space. hos_carr (os_l O) → Type_OF_ordered_uniform_space O.
+intros; assumption;
+qed.
+
+coercion hint_boh2 nocomposites. 
+