]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/dama/dama/ordered_uniform.ma
removed <_,_> notation second interpretation for dependent pair, since it used
[helm.git] / helm / software / matita / contribs / dama / dama / ordered_uniform.ma
index 974405214746fa6acf4ccf4866d65fd2056f391d..bf0260a510ac775fede207a1760f26e371e36daf 100644 (file)
@@ -34,29 +34,143 @@ coercion cic:/matita/dama/ordered_uniform/ous_unifspace.con.
 
 record ordered_uniform_space : Type ≝ {
   ous_stuff :> ordered_uniform_space_;
-  ous_prop1: ∀U.us_unifbase ous_stuff U → convex ous_stuff U
-}.
+  ous_convex: ∀U.us_unifbase ous_stuff U → convex ous_stuff U
+}.   
+
+definition invert_os_relation ≝
+  λC:ordered_set.λU:C square → Prop.
+    λx:C square. U 〈\snd x,\fst x〉.
+
+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 segment_square_of_ordered_set_square: 
+  ∀O:ordered_set.∀u,v:O.∀x:O square.
+   \fst x ∈ [u,v] → \snd x ∈ [u,v] → {[u,v]} square.
+intros; split; exists; [1: apply (\fst x) |3: apply (\snd x)] assumption;
+qed.
 
+coercion cic:/matita/dama/ordered_uniform/segment_square_of_ordered_set_square.con 0 2.
+
+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]} square → O square ≝ 
+  λO:ordered_set.λu,v:O.λb:{[u,v]} square.〈\fst(\fst b),\fst(\snd b)〉.
+
+coercion cic:/matita/dama/ordered_uniform/ordered_set_square_of_segment_square.con.
+
+lemma restriction_agreement : 
+  ∀O:ordered_uniform_space.∀l,r:O.∀P:{[l,r]} square → Prop.∀OP:O square → Prop.Prop.
+apply(λO:ordered_uniform_space.λl,r:O.
+       λP:{[l,r]} square → Prop.λOP:O square → Prop.
+          ∀b:O square.∀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]} square.
+  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; 
+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]} square.
+  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;
+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]} square → Prop.∀u:O square → 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);]
+qed. 
+    
+alias symbol "square" (instance 8) = "bishop set square".
+lemma bs_of_ss: 
+ ∀O:ordered_set.∀u,v:O.{[u,v]} square → (bishop_set_of_ordered_set O) square ≝ 
+  λO:ordered_set.λu,v:O.λb:{[u,v]} square.〈\fst(\fst b),\fst(\snd b)〉.
+
+notation < "x \sub \neq" with precedence 91 for @{'bsss $x}.
+interpretation "bs_of_ss" 'bsss x = (bs_of_ss _ _ _ x).
+
+alias symbol "square" (instance 7) = "ordered set square".
+lemma ss_of_bs: 
+ ∀O:ordered_set.∀u,v:O.
+  ∀b:O square.\fst b ∈ [u,v] → \snd b ∈ [u,v] → {[u,v]} square ≝ 
+ λO:ordered_set.λu,v:O.
+  λb:(O:bishop_set) square.λ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 segment_ordered_uniform_space: 
   ∀O:ordered_uniform_space.∀u,v:O.ordered_uniform_space.
-intros (O u v); apply mk_ordered_uniform_space;
-[1: apply mk_ordered_uniform_space_;
-    [1: apply (mk_ordered_set (sigma ? (λx.x ∈ [u,v])));
-        [1: intros (x y); apply (fst x ≰ fst y);
-        |2: intros 1; cases x; simplify; apply os_coreflexive;
-        |3: intros 3; cases x; cases y; cases z; simplify; apply os_cotransitive]
-    |2: apply (mk_uniform_space (bishop_set_of_ordered_set (mk_ordered_set (sigma ? (λx.x ∈ [u,v])) ???)));
-    |3: apply refl_eq;
+intros (O l r); apply mk_ordered_uniform_space;
+[1: apply (mk_ordered_uniform_space_ {[l,r]});
+    [1: alias symbol "and" = "constructive and".
+        letin f ≝ (λP:{[l,r]} square → Prop. ∃OP:O square → Prop.
+                    (us_unifbase O OP) ∧ restriction_agreement ??? P OP);
+        apply (mk_uniform_space (bishop_set_of_ordered_set {[l,r]}) f);
+        [1: intros (U H); intro x; simplify; 
+            cases H (w Hw); cases Hw (Gw Hwp); clear H Hw; intro Hm;
+            lapply (us_phi1 ?? Gw x Hm) as IH;
+            apply (restrict ?????? 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;
+            [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);]]
+        |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;
+            [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; 
+                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);
+            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);]]
+    |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);]]
 qed.
 
+interpretation "Ordered uniform space segment" 'segment_set a b = 
+ (segment_ordered_uniform_space _ a b).
 
 (* Lemma 3.2 *)
-lemma foo:
+alias symbol "pi1" = "exT \fst".
+lemma restric_uniform_convergence:
  ∀O:ordered_uniform_space.∀l,u:O.
-  ∀x:(segment_ordered_uniform_space O l u).
-   ∀a:sequence (segment_ordered_uniform_space O l u).
-     (* (λn.fst (a n)) uniform_converges (fst x) → *) 
+  ∀x:{[l,u]}.
+   ∀a:sequence {[l,u]}.
+     ⌊n,\fst (a n)⌋ uniform_converges (\fst x) → 
       a uniform_converges x.
-     
-      
\ No newline at end of file
+intros 8; cases H1; cases H2; clear H2 H1;
+cases (H ? H3) (m Hm); exists [apply m]; intros; 
+apply (restrict ? l u ??? H4); apply (Hm ? H1);
+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).