]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/dama/metric_lattice.ma
c8a9c895fc4be18204bd2dc3e851b186085ada06
[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.