]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/dama/metric_lattice.ma
fixed a pulback and proved 3.17
[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
16
17 include "metric_space.ma".
18 include "lattice.ma".
19
20 record mlattice_ (R : todgroup) : Type ≝ {
21   ml_mspace_: metric_space R;
22   ml_lattice_: lattice;
23   ml_with_: ms_carr ? ml_mspace_ = l_carr ml_lattice_;
24   ml_with2_: l_carr ml_lattice_ = apart_of_metric_space ? ml_mspace_ 
25 }.
26
27 lemma ml_lattice: ∀R.mlattice_ R → lattice.
28 intros (R ml); apply (mk_lattice (apart_of_metric_space ? (ml_mspace_ ? ml))); try unfold eq;
29 cases (ml_with2_ ? ml);
30 [apply (join (ml_lattice_ ? ml));|apply (meet (ml_lattice_ ? ml));
31 |apply (join_refl (ml_lattice_ R ml));| apply (meet_refl (ml_lattice_ ? ml));
32 |apply (join_comm (ml_lattice_ ? ml));| apply (meet_comm (ml_lattice_ ? ml));
33 |apply (join_assoc (ml_lattice_ ? ml));|apply (meet_assoc (ml_lattice_ ? ml));
34 |apply (absorbjm (ml_lattice_ ? ml)); |apply (absorbmj (ml_lattice_ ? ml));
35 |apply (strong_extj (ml_lattice_ ? ml));|apply (strong_extm (ml_lattice_ ? ml));]
36 qed.
37
38 coercion cic:/matita/metric_lattice/ml_lattice.con.
39
40 lemma ml_mspace: ∀R.mlattice_ R → metric_space R.
41 intros (R ml); apply (mk_metric_space R ml);
42 cases (ml_with_ ? ml); simplify;
43 [apply (metric ? (ml_mspace_ ? ml));|apply (mpositive ? (ml_mspace_ ? ml));
44 |apply (mreflexive ? (ml_mspace_ ? ml));|apply (msymmetric ? (ml_mspace_ ? ml));
45 |apply (mtineq ? (ml_mspace_ ? ml))]
46 qed.
47
48 coercion cic:/matita/metric_lattice/ml_mspace.con.
49
50 record mlattice (R : todgroup) : Type ≝ {
51   ml_carr :> mlattice_ R;
52   ml_prop1: ∀a,b:ml_carr. 0 < δ a b → a # b;
53   ml_prop2: ∀a,b,c:ml_carr. δ (a∨b) (a∨c) + δ (a∧b) (a∧c) ≤ δ b c
54 }.
55
56 lemma eq_to_ndlt0: ∀R.∀ml:mlattice R.∀a,b:ml. a ≈ b → ¬ 0 < δ a b.
57 intros (R ml a b E); intro H; apply E; apply ml_prop1;
58 assumption;
59 qed.
60
61 lemma eq_to_dzero: ∀R.∀ml:mlattice R.∀x,y:ml.x ≈ y → δ x y ≈ 0.
62 intros (R ml x y H); intro H1; apply H; clear H; 
63 apply ml_prop1; split [apply mpositive] apply ap_symmetric;
64 assumption;
65 qed.
66
67 lemma meq_l: ∀R.∀ml:mlattice R.∀x,y,z:ml. x≈z → δx y ≈ δz y.
68 intros (R ml x y z); apply le_le_eq;
69 [ apply (le_transitive ???? (mtineq ???y z)); 
70   apply (le_rewl ??? (0+δz y) (eq_to_dzero ???? H));
71   apply (le_rewl ??? (δz y) (zero_neutral ??)); apply le_reflexive;
72 | apply (le_transitive ???? (mtineq ???y x));
73   apply (le_rewl ??? (0+δx y) (eq_to_dzero ??z x H));
74   apply (le_rewl ??? (δx y) (zero_neutral ??)); apply le_reflexive;]
75 qed.
76
77 (* 3.3 *)
78 lemma meq_r: ∀R.∀ml:mlattice R.∀x,y,z:ml. x≈z → δy x ≈ δy z.
79 intros; apply (eq_trans ???? (msymmetric ??y x));
80 apply (eq_trans ????? (msymmetric ??z y)); apply meq_l; assumption;
81 qed.
82  
83
84 lemma dap_to_lt: ∀R.∀ml:mlattice R.∀x,y:ml. δ x y # 0 → 0 < δ x y.
85 intros; split [apply mpositive] apply ap_symmetric; assumption;
86 qed.
87
88 lemma dap_to_ap: ∀R.∀ml:mlattice R.∀x,y:ml. δ x y # 0 → x # y.
89 intros (R ml x y H); apply ml_prop1; split; [apply mpositive;]
90 apply ap_symmetric; assumption;
91 qed.
92
93 (* 3.11 *)
94 lemma le_mtri: 
95   ∀R.∀ml:mlattice R.∀x,y,z:ml. x ≤ y → y ≤ z → δ x z ≈ δ x y + δ y z.
96 intros (R ml x y z Lxy Lyz); apply le_le_eq; [apply mtineq]
97 apply (le_transitive ????? (ml_prop2 ?? (y) ??)); 
98 cut ( δx y+ δy z ≈ δ(y∨x) (y∨z)+ δ(y∧x) (y∧z)); [
99   apply (le_rewr ??? (δx y+ δy z)); [assumption] apply le_reflexive]
100 lapply (le_to_eqm ??? Lxy) as Dxm; lapply (le_to_eqm ??? Lyz) as Dym;
101 lapply (le_to_eqj ??? Lxy) as Dxj; lapply (le_to_eqj ??? Lyz) as Dyj; clear Lxy Lyz;
102 apply (Eq≈ (δ(x∧y) y + δy z) (meq_l ????? Dxm));
103 apply (Eq≈ (δ(x∧y) (y∧z) + δy z) (meq_r ????? Dym));
104 apply (Eq≈ (δ(x∧y) (y∧z) + δ(x∨y) z) (meq_l ????? Dxj));
105 apply (Eq≈ (δ(x∧y) (y∧z) + δ(x∨y) (y∨z))); [
106   apply (feq_plusl ? (δ(x∧y) (y∧z)) ?? (meq_r ??? (x∨y) ? Dyj));]
107 apply (Eq≈ ? (plus_comm ???));
108 apply (Eq≈ (δ(y∨x) (y∨z)+ δ(x∧y) (y∧z)) (meq_l ????? (join_comm ?x y)));
109 apply feq_plusl;
110 apply (Eq≈ (δ(y∧x) (y∧z)) (meq_l ????? (meet_comm ?x y)));
111 apply eq_reflexive;   
112 qed.  
113
114
115 (* 3.17 conclusione: δ x y ≈ 0 *)
116 (* 3.20 conclusione: δ x y ≈ 0 *)
117 (* 3.21 sup forte
118    strong_sup x ≝ ∀n. s n ≤ x ∧ ∀y x ≰ y → ∃n. s n ≰ y
119    strong_sup_zoli x ≝  ∀n. s n ≤ x ∧ ∄y. y#x ∧ y ≤ x
120 *)
121 (* 3.22 sup debole (più piccolo dei maggioranti) *)
122 (* 3.23 conclusion: δ x sup(...) ≈ 0 *)
123 (* 3.25 vero nel reticolo e basta (niente δ) *)
124 (* 3.36 conclusion: δ x y ≈ 0 *)