]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/dama/metric_lattice.ma
56120525d8cf3d794f6889000b7cbc1e12f8a995
[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 axiom meq_joinl: ∀R.∀ml:mlattice R.∀x,y,z:ml. x≈z → δx y ≈ δz y.
47
48 lemma meq_joinr: ∀R.∀ml:mlattice R.∀x,y,z:ml. x≈z → δy x ≈ δy z.
49 intros; apply (eq_trans ???? (msymmetric ??y x));
50 apply (eq_trans ????? (msymmetric ??z y)); apply meq_joinl; assumption;
51 qed.
52
53 (*
54 lemma foo: ∀R.∀ml:mlattice R.∀x,y,z:ml. δx y ≈ δ(y∨x) (y∨z).
55 intros; apply le_le_eq; [
56   apply (le_rewr ???? (meq_joinl ????? (join_comm ???)));
57   apply (le_rewr ???? (meq_joinr ????? (join_comm ???)));
58 *) 
59
60 (* 3.11 *)
61 lemma le_mtri: 
62   ∀R.∀ml:mlattice R.∀x,y,z:ml. x ≤ y → y ≤ z → δ x z ≈ δ x y + δ y z.
63 intros (R ml x y z Lxy Lyz); apply le_le_eq; [apply mtineq]
64 apply (le_transitive ????? (ml_prop2 ?? ml (y∧x) ??)); 
65 (* auto type. assert failure *)