]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/dama/metric_lattice.ma
carabinieri almost done
[helm.git] / helm / software / matita / dama / metric_lattice.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 set "baseuri" "cic:/matita/metric_lattice/".
16
17 include "metric_space.ma".
18 include "lattice.ma".
19
20 record mlattice_ (R : ogroup) : Type ≝ {
21   ml_mspace_: metric_space R;
22   ml_lattice:> lattice;
23   ml_with_: ms_carr ? ml_mspace_ = ap_carr (l_carr ml_lattice)
24 }.
25
26 lemma ml_mspace: ∀R.mlattice_ R → metric_space R.
27 intros (R ml); apply (mk_metric_space R ml); unfold Type_OF_mlattice_;
28 cases (ml_with_ ? ml); simplify;
29 [apply (metric ? (ml_mspace_ ? ml));|apply (mpositive ? (ml_mspace_ ? ml));
30 |apply (mreflexive ? (ml_mspace_ ? ml));|apply (msymmetric ? (ml_mspace_ ? ml));
31 |apply (mtineq ? (ml_mspace_ ? ml))]
32 qed.
33
34 coercion cic:/matita/metric_lattice/ml_mspace.con.
35
36 record is_mlattice (R : ogroup) (ml: mlattice_ R) : Type ≝ {
37   ml_prop1: ∀a,b:ml. 0 < δ a b → a # b;
38   ml_prop2: ∀a,b,c:ml. δ (a∨b) (a∨c) + δ (a∧b) (a∧c) ≤ δ b c
39 }.
40
41 record mlattice (R : ogroup) : Type ≝ {
42   ml_carr :> mlattice_ R;
43   ml_props:> is_mlattice R ml_carr
44 }.
45
46 lemma eq_to_ndlt0: ∀R.∀ml:mlattice R.∀a,b:ml. a ≈ b → ¬ 0 < δ a b.
47 intros (R ml a b E); intro H; apply E; apply (ml_prop1 ?? ml);
48 assumption;
49 qed.
50
51 lemma eq_to_dzero: ∀R.∀ml:mlattice R.∀x,y:ml.x ≈ y → δ x y ≈ 0.
52 intros (R ml x y H); intro H1; apply H; clear H; 
53 apply (ml_prop1 ?? ml); split [apply mpositive] apply ap_symmetric;
54 assumption;
55 qed.
56
57 lemma meq_l: ∀R.∀ml:mlattice R.∀x,y,z:ml. x≈z → δx y ≈ δz y.
58 intros (R ml x y z); apply le_le_eq;
59 [ apply (le_transitive ???? (mtineq ???y z)); 
60   apply (le_rewl ??? (0+δz y) (eq_to_dzero ???? H));
61   apply (le_rewl ??? (δz y) (zero_neutral ??)); apply le_reflexive;
62 | apply (le_transitive ???? (mtineq ???y x));
63   apply (le_rewl ??? (0+δx y) (eq_to_dzero ??z x H));
64   apply (le_rewl ??? (δx y) (zero_neutral ??)); apply le_reflexive;]
65 qed.
66
67 (* 3.3 *)
68 lemma meq_r: ∀R.∀ml:mlattice R.∀x,y,z:ml. x≈z → δy x ≈ δy z.
69 intros; apply (eq_trans ???? (msymmetric ??y x));
70 apply (eq_trans ????? (msymmetric ??z y)); apply meq_l; assumption;
71 qed.
72  
73
74 lemma dap_to_lt: ∀R.∀ml:mlattice R.∀x,y:ml. δ x y # 0 → 0 < δ x y.
75 intros; split [apply mpositive] apply ap_symmetric; assumption;
76 qed.
77
78 lemma dap_to_ap: ∀R.∀ml:mlattice R.∀x,y:ml. δ x y # 0 → x # y.
79 intros (R ml x y H); apply (ml_prop1 ?? ml); split; [apply mpositive;]
80 apply ap_symmetric; assumption;
81 qed.
82
83 (* 3.11 *)
84 lemma le_mtri: 
85   ∀R.∀ml:mlattice R.∀x,y,z:ml. x ≤ y → y ≤ z → δ x z ≈ δ x y + δ y z.
86 intros (R ml x y z Lxy Lyz); apply le_le_eq; [apply mtineq]
87 apply (le_transitive ????? (ml_prop2 ?? ml (y) ??)); 
88 cut ( δx y+ δy z ≈ δ(y∨x) (y∨z)+ δ(y∧x) (y∧z)); [
89   apply (le_rewr ??? (δx y+ δy z)); [assumption] apply le_reflexive]
90 lapply (le_to_eqm ??? Lxy) as Dxm; lapply (le_to_eqm ??? Lyz) as Dym;
91 lapply (le_to_eqj ??? Lxy) as Dxj; lapply (le_to_eqj ??? Lyz) as Dyj; clear Lxy Lyz;
92 apply (Eq≈ (δ(x∧y) y + δy z)); [apply feq_plusr; apply (meq_l ????? Dxm);]
93 apply (Eq≈ (δ(x∧y) (y∧z) + δy z)); [apply feq_plusr; apply (meq_r ????? Dym);]
94 apply (Eq≈ (δ(x∧y) (y∧z) + δ(x∨y) z)); [apply feq_plusl; apply (meq_l ????? Dxj);]
95 apply (Eq≈ (δ(x∧y) (y∧z) + δ(x∨y) (y∨z))); [apply feq_plusl; apply (meq_r ????? Dyj);]
96 apply (Eq≈ ? (plus_comm ???));
97 apply (Eq≈ (δ(y∨x) (y∨z)+ δ(x∧y) (y∧z))); [apply feq_plusr; apply (meq_l ????? (join_comm ???));]
98 apply feq_plusl;
99 apply (Eq≈ (δ(y∧x) (y∧z))); [apply (meq_l ????? (meet_comm ???));]
100 apply eq_reflexive;   
101 qed.  
102
103 include "sequence.ma".
104 include "nat/plus.ma".
105
106 definition d2s ≝ 
107   λR.λml:mlattice R.λs:sequence (pordered_set_of_excedence ml).λk.λn. δ (s n) k.
108 (*
109 notation "s ⇝ 'Zero'" non associative with precedence 50 for @{'tends0 $s }.
110   
111 interpretation "tends to" 'tends s x = 
112   (cic:/matita/sequence/tends0.con _ s).    
113 *)
114
115 lemma lew_opp : ∀O:ogroup.∀a,b,c:O.0 ≤ b → a ≤ c → a + -b ≤ c.
116 intros (O a b c L0 L);
117 apply (le_transitive ????? L);
118 apply (plus_cancl_le ??? (-a));
119 apply (le_rewr ??? 0 (opp_inverse ??));
120 apply (le_rewl ??? (-a+a+-b) (plus_assoc ????));
121 apply (le_rewl ??? (0+-b) (opp_inverse ??));
122 apply (le_rewl ??? (-b) (zero_neutral ?(-b)));
123 apply le_zero_x_to_le_opp_x_zero;
124 assumption;
125 qed.
126
127
128 lemma ltw_opp : ∀O:ogroup.∀a,b,c:O.0 < b → a < c → a + -b < c.
129 intros (O a b c P L);
130 apply (lt_transitive ????? L);
131 apply (plus_cancl_lt ??? (-a));
132 apply (lt_rewr ??? 0 (opp_inverse ??));
133 apply (lt_rewl ??? (-a+a+-b) (plus_assoc ????));
134 apply (lt_rewl ??? (0+-b) (opp_inverse ??));
135 apply (lt_rewl ??? ? (zero_neutral ??));
136 apply lt_zero_x_to_lt_opp_x_zero;
137 assumption;
138 qed.
139
140 lemma ltwl: ∀a,b,c:nat. b + a < c → a < c.
141 intros 3 (x y z); elim y (H z IH H); [apply H]
142 apply IH; apply lt_S_to_lt; apply H;
143 qed.
144
145 lemma ltwr: ∀a,b,c:nat. a + b < c → a < c.
146 intros 3 (x y z); rewrite > sym_plus; apply ltwl;
147 qed. 
148
149 alias symbol "leq" = "ordered sets less or equal than".
150 alias symbol "and" = "constructive and".
151 alias symbol "exists" = "constructive exists (Type)".
152 theorem carabinieri: (* non trova la coercion .... *)
153   ∀R.∀ml:mlattice R.∀an,bn,xn:sequence (pordered_set_of_excedence ml).
154     (∀n. (an n ≤ xn n) ∧ (xn n ≤ bn n)) →
155     ∀x:ml. tends0 ? (d2s ? ml an x) → tends0 ? (d2s ? ml bn x) →
156     tends0 ? (d2s ? ml xn x).
157 intros (R ml an bn xn H x Ha Hb); unfold tends0 in Ha Hb ⊢ %. unfold d2s in Ha Hb ⊢ %.
158 intros (e He);
159 elim (Ha ? He) (n1 H1); clear Ha; elim (Hb e He) (n2 H2); clear Hb;
160 constructor 1; [apply (n1 + n2);] intros (n3 Hn3);
161 lapply (ltwr ??? Hn3) as Hn1n3; lapply (ltwl ??? Hn3) as Hn2n3;
162 elim (H1 ? Hn1n3) (H3 H4); elim (H2 ? Hn2n3) (H5 H6); clear Hn1n3 Hn2n3;
163 elim (H n3) (H7 H8); clear H H1 H2; 
164 lapply (le_to_eqm ??? H7) as Dxm; lapply (le_to_eqj ??? H7) as Dym;
165 (* the main step *)
166 cut (δ (xn n3) x ≤ δ (bn n3) x + δ (an n3) x + δ (an n3) x); [2:
167   apply (le_transitive ???? (mtineq ???? (an n3)));
168   lapply (le_mtri ????? H7 H8);
169   lapply (feq_plusr ? (- δ(xn n3) (bn n3)) ?? Hletin);
170   cut ( δ(an n3) (bn n3)+- δ(xn n3) (bn n3)≈( δ(an n3) (xn n3))); [2:
171     apply (Eq≈  (0 + δ(an n3) (xn n3)) ? (zero_neutral ??));
172     apply (Eq≈  (δ(an n3) (xn n3) + 0) ? (plus_comm ???));
173     apply (Eq≈  (δ(an n3) (xn n3) +  (-δ(xn n3) (bn n3) +  δ(xn n3) (bn n3))) ? (opp_inverse ??));
174     apply (Eq≈  (δ(an n3) (xn n3) +  (δ(xn n3) (bn n3) + -δ(xn n3) (bn n3))) ? (plus_comm ?? (δ(xn n3) (bn n3))));
175     apply (Eq≈  ? ? (eq_sym ??? (plus_assoc ????))); assumption;] clear Hletin1;
176   apply (le_rewl ???  ( δ(an n3) (xn n3)+ δ(an n3) x));[
177     apply feq_plusr; apply msymmetric;]
178   apply (le_rewl ???  (δ(an n3) (bn n3)+- δ(xn n3) (bn n3)+ δ(an n3) x));[
179     apply feq_plusr; assumption;]
180   clear Hcut Hletin Dym Dxm H8 H7 H6 H5 H4 H3;
181   apply (le_rewl ??? (- δ(xn n3) (bn n3)+ δ(an n3) (bn n3)+δ(an n3) x));[
182     apply feq_plusr; apply plus_comm;]
183   apply (le_rewl ??? (- δ(xn n3) (bn n3)+ (δ(an n3) (bn n3)+δ(an n3) x)) (plus_assoc ????));
184   apply (le_rewl ??? ((δ(an n3) (bn n3)+δ(an n3) x)+- δ(xn n3) (bn n3)) (plus_comm ???));
185   apply lew_opp; [apply mpositive] apply fle_plusr;
186   apply (le_rewr ???? (plus_comm ???));
187   apply (le_rewr ??? ( δ(an n3) x+ δx (bn n3)) (msymmetric ????));
188   apply mtineq;]
189 split; [
190   apply (lt_le_transitive ????? (mpositive ????));
191   split; elim He; [apply  le_zero_x_to_le_opp_x_zero; assumption;]
192   cases t1; [ 
193     left; apply exc_zero_opp_x_to_exc_x_zero;
194     apply (Ex≫ ? (eq_opp_opp_x_x ??));assumption;]
195   right;  apply exc_opp_x_zero_to_exc_zero_x;
196   apply (Ex≪ ? (eq_opp_opp_x_x ??)); assumption;]
197 clear Dxm Dym H7 H8 Hn3 H5 H3 n1 n2;
198
199   
200 (* 3.17 conclusione: δ x y ≈ 0 *)
201 (* 3.20 conclusione: δ x y ≈ 0 *)
202 (* 3.21 sup forte
203    strong_sup x ≝ ∀n. s n ≤ x ∧ ∀y x ≰ y → ∃n. s n ≰ y
204    strong_sup_zoli x ≝  ∀n. s n ≤ x ∧ ∄y. y#x ∧ y ≤ x
205 *)
206 (* 3.22 sup debole (più piccolo dei maggioranti) *)
207 (* 3.23 conclusion: δ x sup(...) ≈ 0 *)
208 (* 3.25 vero nel reticolo e basta (niente δ) *)
209 (* 3.36 conclusion: δ x y ≈ 0 *)