]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/dama/metric_lattice.ma
...
[helm.git] / helm / software / matita / dama / metric_lattice.ma
index c8a9c895fc4be18204bd2dc3e851b186085ada06..56120525d8cf3d794f6889000b7cbc1e12f8a995 100644 (file)
@@ -17,14 +17,14 @@ set "baseuri" "cic:/matita/metric_lattice/".
 include "metric_space.ma".
 include "lattice.ma".
 
-record mlattice (R : ogroup) : Type ≝ {
+record mlattice_ (R : ogroup) : Type ≝ {
   ml_mspace_: metric_space R;
   ml_lattice:> lattice;
   ml_with_: ms_carr ? ml_mspace_ = ap_carr (l_carr ml_lattice)
 }.
 
-lemma ml_mspace: ∀R.mlattice R → metric_space R.
-intros (R ml); apply (mk_metric_space R ml); unfold Type_OF_mlattice;
+lemma ml_mspace: ∀R.mlattice_ R → metric_space R.
+intros (R ml); apply (mk_metric_space R ml); unfold Type_OF_mlattice_;
 cases (ml_with_ ? ml); simplify;
 [apply (metric ? (ml_mspace_ ? ml));|apply (mpositive ? (ml_mspace_ ? ml));
 |apply (mreflexive ? (ml_mspace_ ? ml));|apply (msymmetric ? (ml_mspace_ ? ml));
@@ -32,3 +32,34 @@ cases (ml_with_ ? ml); simplify;
 qed.
 
 coercion cic:/matita/metric_lattice/ml_mspace.con.
+
+record is_mlattice (R : ogroup) (ml: mlattice_ R) : Type ≝ {
+  ml_prop1: ∀a,b:ml. 0 < δ a b → a # b;
+  ml_prop2: ∀a,b,c:ml. δ (a∨b) (a∨c) + δ (a∧b) (a∧c) ≤ δ b c
+}.
+
+record mlattice (R : ogroup) : Type ≝ {
+  ml_carr :> mlattice_ R;
+  ml_props:> is_mlattice R ml_carr
+}.
+
+axiom meq_joinl: ∀R.∀ml:mlattice R.∀x,y,z:ml. x≈z → δx y ≈ δz y.
+
+lemma meq_joinr: ∀R.∀ml:mlattice R.∀x,y,z:ml. x≈z → δy x ≈ δy z.
+intros; apply (eq_trans ???? (msymmetric ??y x));
+apply (eq_trans ????? (msymmetric ??z y)); apply meq_joinl; assumption;
+qed.
+
+(*
+lemma foo: ∀R.∀ml:mlattice R.∀x,y,z:ml. δx y ≈ δ(y∨x) (y∨z).
+intros; apply le_le_eq; [
+  apply (le_rewr ???? (meq_joinl ????? (join_comm ???)));
+  apply (le_rewr ???? (meq_joinr ????? (join_comm ???)));
+*) 
+
+(* 3.11 *)
+lemma le_mtri: 
+  ∀R.∀ml:mlattice R.∀x,y,z:ml. x ≤ y → y ≤ z → δ x z ≈ δ x y + δ y z.
+intros (R ml x y z Lxy Lyz); apply le_le_eq; [apply mtineq]
+apply (le_transitive ????? (ml_prop2 ?? ml (y∧x) ??)); 
+(* auto type. assert failure *)