X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fdama%2Fmetric_lattice.ma;h=f0242da284896e612f5774aa801bc881a9606ff8;hb=fb9237a1eb706f8d7e6ed0fea9e6f6a65fa5d7fe;hp=9f5811b6d710c0c896f3b74c04bb8d29ab757fa1;hpb=9483f7e6c85ec11f88bdb219b2cebad2039b1a74;p=helm.git diff --git a/helm/software/matita/dama/metric_lattice.ma b/helm/software/matita/dama/metric_lattice.ma index 9f5811b6d..f0242da28 100644 --- a/helm/software/matita/dama/metric_lattice.ma +++ b/helm/software/matita/dama/metric_lattice.ma @@ -32,13 +32,18 @@ qed. coercion cic:/matita/metric_lattice/ml_mspace.con. alias symbol "plus" = "Abelian group plus". -alias symbol "leq" = "ordered sets less or equal than". +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) }. +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; @@ -82,8 +87,9 @@ 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) + δ(y∨x) z));[