X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fdama%2Fmetric_lattice.ma;h=968ae8f3b65515e9e1411fac2530981e949a52a9;hb=55dc87ec418b5b8afe10164cfc61b5ad80a88e6c;hp=0bfc3db678179c3ceedaab49b34ff762961f6f75;hpb=515b66b082bf6e1553d1aa75ba632b99a4d88e27;p=helm.git diff --git a/helm/software/matita/dama/metric_lattice.ma b/helm/software/matita/dama/metric_lattice.ma index 0bfc3db67..968ae8f3b 100644 --- a/helm/software/matita/dama/metric_lattice.ma +++ b/helm/software/matita/dama/metric_lattice.ma @@ -12,27 +12,27 @@ (* *) (**************************************************************************) -set "baseuri" "cic:/matita/metric_lattice/". - 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_ = apart_of_excess (pl_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_; +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; @@ -76,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. @@ -83,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. @@ -107,4 +138,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 *)