X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fdama%2Fmetric_lattice.ma;h=ae55b5551dae807bd16ab72fae962914c0e70b4d;hb=3a5b9647787e1402f6886d41824f664db290963f;hp=6028ada59748c3659533fb48b5fd7c02ec8b694a;hpb=194a530bca425408e72a6ea86afc2947b857da30;p=helm.git diff --git a/helm/software/matita/dama/metric_lattice.ma b/helm/software/matita/dama/metric_lattice.ma index 6028ada59..ae55b5551 100644 --- a/helm/software/matita/dama/metric_lattice.ma +++ b/helm/software/matita/dama/metric_lattice.ma @@ -17,7 +17,7 @@ set "baseuri" "cic:/matita/metric_lattice/". include "metric_space.ma". include "lattice.ma". -record mlattice_ (R : ogroup) : Type ≝ { +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) @@ -33,41 +33,167 @@ qed. coercion cic:/matita/metric_lattice/ml_mspace.con. -record is_mlattice (R : ogroup) (ml: mlattice_ R) : Type ≝ { +record is_mlattice (R : todgroup) (ml: mlattice_ R) : Type ≝ { ml_prop1: ∀a,b:ml. 0 < δ a b → a # b; ml_prop2: ∀a,b,c:ml. δ (a∨b) (a∨c) + δ (a∧b) (a∧c) ≤ δ b c }. -record mlattice (R : ogroup) : Type ≝ { +record mlattice (R : todgroup) : Type ≝ { ml_carr :> mlattice_ R; ml_props:> is_mlattice R ml_carr }. -lemma eq_to_zero: ∀R.∀ml:mlattice R.∀x,y:ml.x ≈ y → δ x y ≈ 0. +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 ?? ml); +assumption; +qed. + +lemma eq_to_dzero: ∀R.∀ml:mlattice R.∀x,y:ml.x ≈ y → δ x y ≈ 0. intros (R ml x y H); intro H1; apply H; clear H; apply (ml_prop1 ?? ml); split [apply mpositive] apply ap_symmetric; assumption; qed. -lemma meq_joinl: ∀R.∀ml:mlattice R.∀x,y,z:ml. x≈z → δx y ≈ δz y. +lemma meq_l: ∀R.∀ml:mlattice R.∀x,y,z:ml. x≈z → δx y ≈ δz y. intros (R ml x y z); apply le_le_eq; [ apply (le_transitive ???? (mtineq ???y z)); - apply (le_rewl ??? (0+δz y) (eq_to_zero ???? H)); + apply (le_rewl ??? (0+δz y) (eq_to_dzero ???? H)); apply (le_rewl ??? (δz y) (zero_neutral ??)); apply le_reflexive; | apply (le_transitive ???? (mtineq ???y x)); - apply (le_rewl ??? (0+δx y) (eq_to_zero ??z x H)); + apply (le_rewl ??? (0+δx y) (eq_to_dzero ??z x H)); apply (le_rewl ??? (δx y) (zero_neutral ??)); apply le_reflexive;] qed. -lemma meq_joinr: ∀R.∀ml:mlattice R.∀x,y,z:ml. x≈z → δy x ≈ δy z. +(* 3.3 *) +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_joinl; assumption; -qed. +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; +qed. + +lemma dap_to_ap: ∀R.∀ml:mlattice R.∀x,y:ml. δ x y # 0 → x # y. +intros (R ml x y H); apply (ml_prop1 ?? ml); split; [apply mpositive;] +apply ap_symmetric; assumption; +qed. (* 3.11 *) lemma le_mtri: ∀R.∀ml:mlattice R.∀x,y,z:ml. x ≤ y → y ≤ z → δ x z ≈ δ x y + δ y z. intros (R ml x y z Lxy Lyz); apply le_le_eq; [apply mtineq] apply (le_transitive ????? (ml_prop2 ?? ml (y) ??)); -(* auto type. assert failure *) -whd; +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; +apply (Eq≈ (δ(x∧y) y + δy z)); [apply feq_plusr; apply (meq_l ????? Dxm);] +apply (Eq≈ (δ(x∧y) (y∧z) + δy z)); [apply feq_plusr; apply (meq_r ????? Dym);] +apply (Eq≈ (δ(x∧y) (y∧z) + δ(x∨y) z)); [apply feq_plusl; apply (meq_l ????? Dxj);] +apply (Eq≈ (δ(x∧y) (y∧z) + δ(x∨y) (y∨z))); [apply feq_plusl; apply (meq_r ????? Dyj);] +apply (Eq≈ ? (plus_comm ???)); +apply (Eq≈ (δ(y∨x) (y∨z)+ δ(x∧y) (y∧z))); [apply feq_plusr; apply (meq_l ????? (join_comm ???));] +apply feq_plusl; +apply (Eq≈ (δ(y∧x) (y∧z))); [apply (meq_l ????? (meet_comm ???));] +apply eq_reflexive; +qed. + +include "sequence.ma". +include "nat/plus.ma". + +lemma ltwl: ∀a,b,c:nat. b + a < c → a < c. +intros 3 (x y z); elim y (H z IH H); [apply H] +apply IH; apply lt_S_to_lt; apply H; +qed. + +lemma ltwr: ∀a,b,c:nat. a + b < c → a < c. +intros 3 (x y z); rewrite > sym_plus; apply ltwl; +qed. + + +definition d2s ≝ + λR.λml:mlattice R.λs:sequence ml.λk.λn. δ (s n) k. +(* +notation "s ⇝ 'Zero'" non associative with precedence 50 for @{'tends0 $s }. + +interpretation "tends to" 'tends s x = + (cic:/matita/sequence/tends0.con _ s). +*) + +axiom core1: ∀G:todgroup.∀e:G.0