]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/dama/dama/ordered_uniform.ma
exhaustivity completed
[helm.git] / helm / software / matita / contribs / dama / dama / ordered_uniform.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 include "uniform.ma".
16
17 record ordered_uniform_space_ : Type ≝ {
18   ous_os:> ordered_set;
19   ous_us_: uniform_space;
20   with_ : us_carr ous_us_ = bishop_set_of_ordered_set ous_os
21 }.
22
23 lemma ous_unifspace: ordered_uniform_space_ → uniform_space.
24 intro X; apply (mk_uniform_space (bishop_set_of_ordered_set X));
25 unfold bishop_set_OF_ordered_uniform_space_;
26 [1: rewrite < (with_ X); simplify; apply (us_unifbase (ous_us_ X));
27 |2: cases (with_ X); simplify; apply (us_phi1 (ous_us_ X));
28 |3: cases (with_ X); simplify; apply (us_phi2 (ous_us_ X));
29 |4: cases (with_ X); simplify; apply (us_phi3 (ous_us_ X));
30 |5: cases (with_ X); simplify; apply (us_phi4 (ous_us_ X))]
31 qed.
32
33 coercion cic:/matita/dama/ordered_uniform/ous_unifspace.con.
34
35 record ordered_uniform_space : Type ≝ {
36   ous_stuff :> ordered_uniform_space_;
37   ous_convex: ∀U.us_unifbase ous_stuff U → convex ous_stuff U
38 }.   
39
40 lemma segment_square_of_O_square: 
41   ∀O:ordered_set.∀u,v:O.∀x:O square.fst x ∈ [u,v] → snd x ∈ [u,v] → {[u,v]} square.
42 intros; split; exists; [1: apply (fst x) |3: apply (snd x)] assumption;
43 qed.
44
45 coercion cic:/matita/dama/ordered_uniform/segment_square_of_O_square.con 0 2.
46
47 alias symbol "pi1" (instance 4) = "sigma pi1".
48 alias symbol "pi1" (instance 2) = "sigma pi1".
49 lemma O_square_of_segment_square : 
50  ∀O:ordered_set.∀u,v:O.{[u,v]} square → O square ≝ 
51  λO:ordered_set.λu,v:O.λb:{[u,v]} square.〈fst(fst b),fst(snd b)〉.
52
53 coercion cic:/matita/dama/ordered_uniform/O_square_of_segment_square.con.
54
55 lemma restriction_agreement : 
56   ∀O:ordered_uniform_space.∀l,r:O.∀P:{[l,r]} square → Prop.∀OP:O square → Prop.Prop.
57 apply(λO:ordered_uniform_space.λl,r:O.λP:{[l,r]} square → Prop.λOP:O square → Prop.
58     ∀b:O square.∀H1,H2.
59       (P b → OP b) ∧ (OP b → P b));
60 [5,7: apply H1|6,8:apply H2]skip;
61 qed.
62
63 lemma unrestrict: ∀O:ordered_uniform_space.∀l,r:O.∀U,u.∀x:{[l,r]} square.
64   restriction_agreement ? l r U u → U x → u x.
65 intros 7; cases x (b b1); cases b; cases b1; 
66 cases (H 〈x1,x2〉 H1 H2) (L _); intros; apply L; assumption;
67 qed.
68
69 lemma restrict: ∀O:ordered_uniform_space.∀l,r:O.∀U,u.∀x:{[l,r]} square.
70   restriction_agreement ? l r U u → u x → U x.
71 intros 6; cases x (b b1); cases b; cases b1; intros (X); 
72 cases (X 〈x1,x2〉 H H1) (_ R); apply R; assumption;
73 qed.
74
75 lemma invert_restriction_agreement: 
76   ∀O:ordered_uniform_space.∀l,r:O.∀U:{[l,r]} square → Prop.∀u.
77     restriction_agreement ? l r U u →
78     restriction_agreement ? l r (inv U) (inv u).
79 intros 9; split; intro;
80 [1: apply (unrestrict ????? (segment_square_of_O_square ??? 〈snd b,fst b〉 H2 H1) H H3);
81 |2: apply (restrict ????? (segment_square_of_O_square ??? 〈snd b,fst b〉 H2 H1) H H3);]
82 qed. 
83     
84 lemma bs_of_ss: 
85  ∀O:ordered_set.∀u,v:O.{[u,v]} square → (bishop_set_of_ordered_set O) square ≝ 
86  λO:ordered_set.λu,v:O.λb:{[u,v]} square.〈fst(fst b),fst(snd b)〉.
87
88 notation < "x \sub \neq" non associative with precedence 91 for @{'bsss $x}.
89 interpretation "bs_of_ss" 'bsss x = 
90   (cic:/matita/dama/ordered_uniform/bs_of_ss.con _ _ _ x).
91
92 lemma segment_ordered_uniform_space: 
93   ∀O:ordered_uniform_space.∀u,v:O.ordered_uniform_space.
94 intros (O l r); apply mk_ordered_uniform_space;
95 [1: apply (mk_ordered_uniform_space_ {[l,r]});
96     [1: alias symbol "and" = "constructive and".
97         letin f ≝ (λP:{[l,r]} square → Prop. ∃OP:O square → Prop.
98                     (us_unifbase O OP) ∧ restriction_agreement ??? P OP);
99         apply (mk_uniform_space (bishop_set_of_ordered_set {[l,r]}) f);
100         [1: intros (U H); intro x; unfold mk_set; simplify; 
101             cases H (w Hw); cases Hw (Gw Hwp); clear H Hw; intro Hm;
102             lapply (us_phi1 ?? Gw x Hm) as IH;
103             apply (restrict ?????? Hwp IH);
104         |2: intros (U V HU HV); cases HU (u Hu); cases HV (v Hv); clear HU HV;
105             cases Hu (Gu HuU); cases Hv (Gv HvV); clear Hu Hv;
106             cases (us_phi2 ??? Gu Gv) (w HW); cases HW (Gw Hw); clear HW;
107             exists; [apply (λb:{[l,r]} square.w b)] split;
108             [1: unfold f; simplify; clearbody f;
109                 exists; [apply w]; split; [assumption] intro b; simplify;
110                 unfold segment_square_of_O_square; (* ??? *)
111                 cases b; intros; split; intros; assumption;
112             |2: intros 2 (x Hx); unfold mk_set; cases (Hw ? Hx); split;
113                 [apply (restrict ?????? HuU H)|apply (restrict ?????? HvV H1);]]
114         |3: intros (U Hu); cases Hu (u HU); cases HU (Gu HuU); clear Hu HU;
115             cases (us_phi3 ?? Gu) (w HW); cases HW (Gw Hwu); clear HW;
116             exists; [apply (λx:{[l,r]} square.w x)] split;
117             [1: exists;[apply w];split;[assumption] intros; simplify; intro;
118                 unfold segment_square_of_O_square; (* ??? *)
119                 cases b; intros; split; intro; assumption;
120             |2: intros 2 (x Hx); apply (restrict ?????? HuU); apply Hwu; 
121                 cases Hx (m Hm); exists[apply (fst m)] apply Hm;]
122         |4: intros (U HU x); cases HU (u Hu); cases Hu (Gu HuU); clear HU Hu;
123             cases (us_phi4 ?? Gu x) (Hul Hur);
124             split; intros; 
125             [1: apply (restrict ?????? (invert_restriction_agreement ????? HuU));
126                 apply Hul; apply (unrestrict ?????? HuU H);
127             |2: apply (restrict ?????? HuU); apply Hur; 
128                 apply (unrestrict ?????? (invert_restriction_agreement ????? HuU) H);]]
129     |2: simplify; reflexivity;]
130 |2: simplify; unfold convex; intros;
131     cases H (u HU); cases HU (Gu HuU); clear HU H; 
132     lapply (ous_convex ?? Gu (bs_of_ss ? l r p) ? H2 (bs_of_ss ? l r y) H3) as Cu;
133     [1: apply (unrestrict ?????? HuU); apply H1;
134     |2: apply (restrict ?????? HuU Cu);]]
135 qed.
136
137 interpretation "Ordered uniform space segment" 'segment_set a b = 
138  (cic:/matita/dama/ordered_uniform/segment_ordered_uniform_space.con _ a b).
139
140 (* Lemma 3.2 *)
141 alias symbol "pi1" = "sigma pi1".
142 lemma restric_uniform_convergence:
143  ∀O:ordered_uniform_space.∀l,u:O.
144   ∀x:{[l,u]}.
145    ∀a:sequence {[l,u]}.
146      (λn.fst (a n)) uniform_converges (fst x) → 
147       a uniform_converges x.
148 intros 8; cases H1; cases H2; clear H2 H1;
149 cases (H ? H3) (m Hm); exists [apply m]; intros; 
150 apply (restrict ? l u ??? H4); apply (Hm ? H1);
151 qed.
152
153 definition order_continuity ≝
154   λC:ordered_uniform_space.∀a:sequence C.∀x:C.
155     (a ↑ x → a uniform_converges x) ∧ (a ↓ x → a uniform_converges x).