]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/dama/metric_lattice.ma
fixed a pulback and proved 3.17
[helm.git] / helm / software / matita / dama / metric_lattice.ma
index 36742163bc541dab30cb60408921d7e16f9b8977..40f3d41eb8f32fa057cf46faf463fe2f045f5381 100644 (file)
@@ -19,12 +19,26 @@ include "lattice.ma".
 
 record mlattice_ (R : todgroup) : Type ≝ {
   ml_mspace_: metric_space R;
-  ml_lattice:> lattice;
-  ml_with_: ms_carr ? ml_mspace_ = ap_carr (l_carr ml_lattice)
+  ml_lattice_: lattice;
+  ml_with_: ms_carr ? ml_mspace_ = l_carr ml_lattice_;
+  ml_with2_: l_carr ml_lattice_ = apart_of_metric_space ? ml_mspace_ 
 }.
 
+lemma ml_lattice: ∀R.mlattice_ R → lattice.
+intros (R ml); apply (mk_lattice (apart_of_metric_space ? (ml_mspace_ ? ml))); try unfold eq;
+cases (ml_with2_ ? ml);
+[apply (join (ml_lattice_ ? ml));|apply (meet (ml_lattice_ ? ml));
+|apply (join_refl (ml_lattice_ R ml));| apply (meet_refl (ml_lattice_ ? ml));
+|apply (join_comm (ml_lattice_ ? ml));| apply (meet_comm (ml_lattice_ ? ml));
+|apply (join_assoc (ml_lattice_ ? ml));|apply (meet_assoc (ml_lattice_ ? ml));
+|apply (absorbjm (ml_lattice_ ? ml)); |apply (absorbmj (ml_lattice_ ? ml));
+|apply (strong_extj (ml_lattice_ ? ml));|apply (strong_extm (ml_lattice_ ? ml));]
+qed.
+
+coercion cic:/matita/metric_lattice/ml_lattice.con.
+
 lemma ml_mspace: ∀R.mlattice_ R → metric_space R.
-intros (R ml); apply (mk_metric_space R ml); unfold Type_OF_mlattice_;
+intros (R ml); apply (mk_metric_space R ml);
 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));
@@ -107,4 +121,4 @@ qed.
 (* 3.22 sup debole (più piccolo dei maggioranti) *)
 (* 3.23 conclusion: δ x sup(...) ≈ 0 *)
 (* 3.25 vero nel reticolo e basta (niente δ) *)
-(* 3.36 conclusion: δ x y ≈ 0 *)
\ No newline at end of file
+(* 3.36 conclusion: δ x y ≈ 0 *)