]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/dama/metric_lattice.ma
snopshot before isabellization
[helm.git] / helm / software / matita / dama / metric_lattice.ma
index 40f3d41eb8f32fa057cf46faf463fe2f045f5381..968ae8f3b65515e9e1411fac2530981e949a52a9 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-
-
 include "metric_space.ma".
 include "lattice.ma".
 
 record mlattice_ (R : todgroup) : Type ≝ {
   ml_mspace_: metric_space R;
-  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_ 
+  ml_lattice:> lattice;
+  ml_with_: ms_carr ? ml_mspace_ = apart_of_excess (pl_carr ml_lattice) 
 }.
 
-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);
+intros (R ml); apply (mk_metric_space R (apart_of_excess ml)); 
+unfold apartness_OF_mlattice_; 
+[rewrite < (ml_with_ ? ml); apply (metric ? (ml_mspace_ ? 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));
-|apply (mtineq ? (ml_mspace_ ? ml))]
+[apply (mpositive ? (ml_mspace_ ? ml));|apply (mreflexive ? (ml_mspace_ ? ml));
+|apply (msymmetric ? (ml_mspace_ ? ml));|apply (mtineq ? (ml_mspace_ ? ml))]
 qed.
 
 coercion cic:/matita/metric_lattice/ml_mspace.con.
 
+alias symbol "plus" = "Abelian group plus".
 record mlattice (R : todgroup) : Type ≝ {
   ml_carr :> mlattice_ R;
   ml_prop1: ∀a,b:ml_carr. 0 < δ a b → a # b;
@@ -90,6 +76,31 @@ intros (R ml x y H); apply ml_prop1; split; [apply mpositive;]
 apply ap_symmetric; assumption;
 qed.
 
+interpretation "Lattive meet le" 'leq a b =
+ (cic:/matita/excess/le.con (cic:/matita/lattice/excess_OF_lattice1.con _) a b).
+
+interpretation "Lattive join le (aka ge)" 'geq a b =
+ (cic:/matita/excess/le.con (cic:/matita/lattice/excess_OF_lattice.con _) a b).
+
+lemma le_to_ge: ∀l:lattice.∀a,b:l.a ≤ b → b ≥ a.
+intros(l a b H); apply H;
+qed.
+
+lemma ge_to_le: ∀l:lattice.∀a,b:l.b ≥ a → a ≤ b.
+intros(l a b H); apply H;
+qed.
+
+lemma eq_to_eq:∀l:lattice.∀a,b:l.
+  (eq (apart_of_excess (pl_carr (latt_jcarr l))) a b) →
+  (eq (apart_of_excess (pl_carr (latt_mcarr l))) a b).
+intros 3; unfold eq; unfold apartness_OF_lattice;
+unfold apartness_OF_lattice_1; unfold latt_jcarr; simplify;
+unfold dual_exc; simplify; intros 2 (H H1); apply H;
+cases H1;[right|left]assumption;
+qed. 
+
+coercion cic:/matita/metric_lattice/eq_to_eq.con nocomposites.
+
 (* 3.11 *)
 lemma le_mtri: 
   ∀R.∀ml:mlattice R.∀x,y,z:ml. x ≤ y → y ≤ z → δ x z ≈ δ x y + δ y z.
@@ -97,17 +108,23 @@ intros (R ml x y z Lxy Lyz); apply le_le_eq; [apply mtineq]
 apply (le_transitive ????? (ml_prop2 ?? (y) ??)); 
 cut ( δx y+ δy z ≈ δ(y∨x) (y∨z)+ δ(y∧x) (y∧z)); [
   apply (le_rewr ??? (δx y+ δy z)); [assumption] apply le_reflexive]
-lapply (le_to_eqm ??? Lxy) as Dxm; lapply (le_to_eqm ??? Lyz) as Dym;
-lapply (le_to_eqj ??? Lxy) as Dxj; lapply (le_to_eqj ??? Lyz) as Dyj; clear Lxy Lyz;
+lapply (le_to_eqm ?? Lxy) as Dxm; lapply (le_to_eqm ?? Lyz) as Dym;
+lapply (le_to_eqj ?? (le_to_ge ??? Lxy)) as Dxj; lapply (le_to_eqj ?? (le_to_ge ??? Lyz)) as Dyj; clear Lxy Lyz;
 apply (Eq≈ (δ(x∧y) y + δy z) (meq_l ????? Dxm));
 apply (Eq≈ (δ(x∧y) (y∧z) + δy z) (meq_r ????? Dym));
-apply (Eq≈ (δ(x∧y) (y∧z) + δ(x∨y) z) (meq_l ????? Dxj));
-apply (Eq≈ (δ(x∧y) (y∧z) + δ(x∨y) (y∨z))); [
-  apply (feq_plusl ? (δ(x∧y) (y∧z)) ?? (meq_r ??? (x∨y) ? Dyj));]
+apply (Eq≈ (δ(x∧y) (y∧z) + δ(y∨x) z));[
+  apply feq_plusl; apply meq_l; clear Dyj Dxm Dym;
+  unfold apartness_OF_mlattice1;
+  exact (eq_to_eq ??? Dxj);]
+apply (Eq≈ (δ(x∧y) (y∧z) + δ(y∨x) (z∨y))); [
+  apply (feq_plusl ? (δ(x∧y) (y∧z)) ?? (meq_r ??? (y∨x) ? Dyj));]
 apply (Eq≈ ? (plus_comm ???));
-apply (Eq≈ (δ(y∨x) (y∨z)+ δ(x∧y) (y∧z)) (meq_l ????? (join_comm ?x y)));
+apply (Eq≈ (δ(y∨x) (y∨z)+ δ(x∧y) (y∧z)));[
+  apply feq_plusr;
+  apply meq_r; 
+  apply (join_comm y z);]
 apply feq_plusl;
-apply (Eq≈ (δ(y∧x) (y∧z)) (meq_l ????? (meet_comm ?x y)));
+apply (Eq≈ (δ(y∧x) (y∧z)) (meq_l ????? (meet_comm x y)));
 apply eq_reflexive;   
 qed.