]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/dama/metric_lattice.ma
renaming
[helm.git] / helm / software / matita / dama / metric_lattice.ma
index 36742163bc541dab30cb60408921d7e16f9b8977..f0242da284896e612f5774aa801bc881a9606ff8 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_ = ap_carr (l_carr ml_lattice)
+  ml_with: ms_carr ? ml_mspace_ = Type_OF_lattice ml_lattice 
 }.
 
 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;
+intros  (R ml); apply (mk_metric_space R (Type_OF_mlattice_ ? 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));
 |apply (mtineq ? (ml_mspace_ ? ml))]
@@ -33,12 +31,19 @@ qed.
 
 coercion cic:/matita/metric_lattice/ml_mspace.con.
 
+alias symbol "plus" = "Abelian group plus".
+alias symbol "leq" = "Excess less or equal than".
 record mlattice (R : todgroup) : Type ≝ {
   ml_carr :> mlattice_ R;
   ml_prop1: ∀a,b:ml_carr. 0 < δ a b → a # b;
-  ml_prop2: ∀a,b,c:ml_carr. δ (a∨b) (a∨c) + δ (a∧b) (a∧c) ≤ δ b c
+  ml_prop2: ∀a,b,c:ml_carr. δ (a∨b) (a∨c) + δ (a∧b) (a∧c) ≤ (δ b c)
 }.
 
+interpretation "Metric lattice leq" 'leq a b = 
+  (cic:/matita/excess/le.con (cic:/matita/metric_lattice/excess_OF_mlattice1.con _ _) a b). 
+interpretation "Metric lattice geq" 'geq a b = 
+  (cic:/matita/excess/le.con (cic:/matita/metric_lattice/excess_OF_mlattice.con _ _) a b). 
+
 lemma eq_to_ndlt0: ∀R.∀ml:mlattice R.∀a,b:ml. a ≈ b → ¬ 0 < δ a b.
 intros (R ml a b E); intro H; apply E; apply ml_prop1;
 assumption;
@@ -65,7 +70,6 @@ lemma meq_r: ∀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_l; assumption;
 qed.
 
 lemma dap_to_lt: ∀R.∀ml:mlattice R.∀x,y:ml. δ x y # 0 → 0 < δ x y.
 intros; split [apply mpositive] apply ap_symmetric; assumption;
@@ -83,17 +87,20 @@ 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 y x Lxy) as Dxm; lapply (le_to_eqm z y Lyz) as Dym;
+lapply (le_to_eqj x y Lxy) as Dxj; lapply (le_to_eqj y z Lyz) as Dyj; clear Lxy Lyz;
+STOP
 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; assumption]
+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 ??);]
 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 ??)));
 apply eq_reflexive;   
 qed.  
 
@@ -107,4 +114,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 *)