]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/dama/metric_lattice.ma
snopshot before isabellization
[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 include "metric_space.ma".
16 include "lattice.ma".
17
18 record mlattice_ (R : todgroup) : Type ≝ {
19   ml_mspace_: metric_space R;
20   ml_lattice:> lattice;
21   ml_with_: ms_carr ? ml_mspace_ = apart_of_excess (pl_carr ml_lattice) 
22 }.
23
24 lemma ml_mspace: ∀R.mlattice_ R → metric_space R.
25 intros (R ml); apply (mk_metric_space R (apart_of_excess ml)); 
26 unfold apartness_OF_mlattice_; 
27 [rewrite < (ml_with_ ? ml); apply (metric ? (ml_mspace_ ? ml))]
28 cases (ml_with_ ? ml); simplify;
29 [apply (mpositive ? (ml_mspace_ ? ml));|apply (mreflexive ? (ml_mspace_ ? ml));
30 |apply (msymmetric ? (ml_mspace_ ? ml));|apply (mtineq ? (ml_mspace_ ? ml))]
31 qed.
32
33 coercion cic:/matita/metric_lattice/ml_mspace.con.
34
35 alias symbol "plus" = "Abelian group plus".
36 record mlattice (R : todgroup) : Type ≝ {
37   ml_carr :> mlattice_ R;
38   ml_prop1: ∀a,b:ml_carr. 0 < δ a b → a # b;
39   ml_prop2: ∀a,b,c:ml_carr. δ (a∨b) (a∨c) + δ (a∧b) (a∧c) ≤ δ b c
40 }.
41
42 lemma eq_to_ndlt0: ∀R.∀ml:mlattice R.∀a,b:ml. a ≈ b → ¬ 0 < δ a b.
43 intros (R ml a b E); intro H; apply E; apply ml_prop1;
44 assumption;
45 qed.
46
47 lemma eq_to_dzero: ∀R.∀ml:mlattice R.∀x,y:ml.x ≈ y → δ x y ≈ 0.
48 intros (R ml x y H); intro H1; apply H; clear H; 
49 apply ml_prop1; split [apply mpositive] apply ap_symmetric;
50 assumption;
51 qed.
52
53 lemma meq_l: ∀R.∀ml:mlattice R.∀x,y,z:ml. x≈z → δx y ≈ δz y.
54 intros (R ml x y z); apply le_le_eq;
55 [ apply (le_transitive ???? (mtineq ???y z)); 
56   apply (le_rewl ??? (0+δz y) (eq_to_dzero ???? H));
57   apply (le_rewl ??? (δz y) (zero_neutral ??)); apply le_reflexive;
58 | apply (le_transitive ???? (mtineq ???y x));
59   apply (le_rewl ??? (0+δx y) (eq_to_dzero ??z x H));
60   apply (le_rewl ??? (δx y) (zero_neutral ??)); apply le_reflexive;]
61 qed.
62
63 (* 3.3 *)
64 lemma meq_r: ∀R.∀ml:mlattice R.∀x,y,z:ml. x≈z → δy x ≈ δy z.
65 intros; apply (eq_trans ???? (msymmetric ??y x));
66 apply (eq_trans ????? (msymmetric ??z y)); apply meq_l; assumption;
67 qed.
68  
69
70 lemma dap_to_lt: ∀R.∀ml:mlattice R.∀x,y:ml. δ x y # 0 → 0 < δ x y.
71 intros; split [apply mpositive] apply ap_symmetric; assumption;
72 qed.
73
74 lemma dap_to_ap: ∀R.∀ml:mlattice R.∀x,y:ml. δ x y # 0 → x # y.
75 intros (R ml x y H); apply ml_prop1; split; [apply mpositive;]
76 apply ap_symmetric; assumption;
77 qed.
78
79 interpretation "Lattive meet le" 'leq a b =
80  (cic:/matita/excess/le.con (cic:/matita/lattice/excess_OF_lattice1.con _) a b).
81
82 interpretation "Lattive join le (aka ge)" 'geq a b =
83  (cic:/matita/excess/le.con (cic:/matita/lattice/excess_OF_lattice.con _) a b).
84
85 lemma le_to_ge: ∀l:lattice.∀a,b:l.a ≤ b → b ≥ a.
86 intros(l a b H); apply H;
87 qed.
88
89 lemma ge_to_le: ∀l:lattice.∀a,b:l.b ≥ a → a ≤ b.
90 intros(l a b H); apply H;
91 qed.
92
93 lemma eq_to_eq:∀l:lattice.∀a,b:l.
94   (eq (apart_of_excess (pl_carr (latt_jcarr l))) a b) →
95   (eq (apart_of_excess (pl_carr (latt_mcarr l))) a b).
96 intros 3; unfold eq; unfold apartness_OF_lattice;
97 unfold apartness_OF_lattice_1; unfold latt_jcarr; simplify;
98 unfold dual_exc; simplify; intros 2 (H H1); apply H;
99 cases H1;[right|left]assumption;
100 qed. 
101
102 coercion cic:/matita/metric_lattice/eq_to_eq.con nocomposites.
103
104 (* 3.11 *)
105 lemma le_mtri: 
106   ∀R.∀ml:mlattice R.∀x,y,z:ml. x ≤ y → y ≤ z → δ x z ≈ δ x y + δ y z.
107 intros (R ml x y z Lxy Lyz); apply le_le_eq; [apply mtineq]
108 apply (le_transitive ????? (ml_prop2 ?? (y) ??)); 
109 cut ( δx y+ δy z ≈ δ(y∨x) (y∨z)+ δ(y∧x) (y∧z)); [
110   apply (le_rewr ??? (δx y+ δy z)); [assumption] apply le_reflexive]
111 lapply (le_to_eqm ?? Lxy) as Dxm; lapply (le_to_eqm ?? Lyz) as Dym;
112 lapply (le_to_eqj ?? (le_to_ge ??? Lxy)) as Dxj; lapply (le_to_eqj ?? (le_to_ge ??? Lyz)) as Dyj; clear Lxy Lyz;
113 apply (Eq≈ (δ(x∧y) y + δy z) (meq_l ????? Dxm));
114 apply (Eq≈ (δ(x∧y) (y∧z) + δy z) (meq_r ????? Dym));
115 apply (Eq≈ (δ(x∧y) (y∧z) + δ(y∨x) z));[
116   apply feq_plusl; apply meq_l; clear Dyj Dxm Dym;
117   unfold apartness_OF_mlattice1;
118   exact (eq_to_eq ??? Dxj);]
119 apply (Eq≈ (δ(x∧y) (y∧z) + δ(y∨x) (z∨y))); [
120   apply (feq_plusl ? (δ(x∧y) (y∧z)) ?? (meq_r ??? (y∨x) ? Dyj));]
121 apply (Eq≈ ? (plus_comm ???));
122 apply (Eq≈ (δ(y∨x) (y∨z)+ δ(x∧y) (y∧z)));[
123   apply feq_plusr;
124   apply meq_r; 
125   apply (join_comm y z);]
126 apply feq_plusl;
127 apply (Eq≈ (δ(y∧x) (y∧z)) (meq_l ????? (meet_comm x y)));
128 apply eq_reflexive;   
129 qed.  
130
131
132 (* 3.17 conclusione: δ x y ≈ 0 *)
133 (* 3.20 conclusione: δ x y ≈ 0 *)
134 (* 3.21 sup forte
135    strong_sup x ≝ ∀n. s n ≤ x ∧ ∀y x ≰ y → ∃n. s n ≰ y
136    strong_sup_zoli x ≝  ∀n. s n ≤ x ∧ ∄y. y#x ∧ y ≤ x
137 *)
138 (* 3.22 sup debole (più piccolo dei maggioranti) *)
139 (* 3.23 conclusion: δ x sup(...) ≈ 0 *)
140 (* 3.25 vero nel reticolo e basta (niente δ) *)
141 (* 3.36 conclusion: δ x y ≈ 0 *)