From: Enrico Tassi Date: Mon, 14 Jan 2008 14:58:26 +0000 (+0000) Subject: fixed a pulback and proved 3.17 X-Git-Tag: make_still_working~5670 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=954ed2eaf305a60d7e046206472bc0397a421ad2;p=helm.git fixed a pulback and proved 3.17 --- diff --git a/helm/software/matita/dama/metric_lattice.ma b/helm/software/matita/dama/metric_lattice.ma index 36742163b..40f3d41eb 100644 --- a/helm/software/matita/dama/metric_lattice.ma +++ b/helm/software/matita/dama/metric_lattice.ma @@ -19,12 +19,26 @@ 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_lattice_: lattice; + ml_with_: ms_carr ? ml_mspace_ = l_carr ml_lattice_; + ml_with2_: l_carr ml_lattice_ = apart_of_metric_space ? ml_mspace_ }. +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); unfold Type_OF_mlattice_; +intros (R ml); apply (mk_metric_space R 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)); @@ -107,4 +121,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 *) diff --git a/helm/software/matita/dama/metric_space.ma b/helm/software/matita/dama/metric_space.ma index d36e90a68..595407de8 100644 --- a/helm/software/matita/dama/metric_space.ma +++ b/helm/software/matita/dama/metric_space.ma @@ -23,7 +23,6 @@ record metric_space (R : todgroup) : Type ≝ { mreflexive: ∀a. metric a a ≈ 0; msymmetric: ∀a,b. metric a b ≈ metric b a; mtineq: ∀a,b,c:ms_carr. metric a b ≤ metric a c + metric c b - (* manca qualcosa per essere una metrica e non una semimetrica *) }. notation < "\nbsp \delta a \nbsp b" non associative with precedence 80 for @{ 'delta2 $a $b}. @@ -49,3 +48,6 @@ intros (R ms); apply (mk_apartness ? (apart_metric ? ms)); unfold apart_metric; qed. coercion cic:/matita/metric_space/apart_of_metric_space.con. + +lemma ap2delta: ∀R.∀m:metric_space R.∀x,y:m.x#y → 0 < δ x y. +intros 2 (R m); cases m 0; simplify; intros; assumption; qed. diff --git a/helm/software/matita/dama/sandwich_corollary.ma b/helm/software/matita/dama/sandwich_corollary.ma new file mode 100644 index 000000000..f41b65028 --- /dev/null +++ b/helm/software/matita/dama/sandwich_corollary.ma @@ -0,0 +1,45 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "sandwich.ma". + +include "metric_lattice.ma". + +(* 3.17 *) +lemma tends_uniq: + ∀R.∀ml:mlattice R.∀xn:sequence ml. + ∀x,y:apart_of_metric_space ? ml. + (* BUG: it inserts a compesoed coercion instead of an hand made one, + what to do? prefer the human made one or allow to kill a coercion? + *) + xn ⇝ x → xn ⇝ y → x ≈ y. +intros (R ml xn x y H1 H2); unfold tends0 in H1 H2; unfold d2s in H1 H2; +intro Axy; lapply (ap2delta R ml x y Axy) as ge0; +cases (H1 (δ x y/1) (divide_preserves_lt ??? ge0)) (n1 Hn1); clear H1; +cases (H2 (δ x y/1) (divide_preserves_lt ??? ge0)) (n2 Hn2); clear H2; +letin N ≝ (S (n2 + n1)); +cases (Hn1 N ?) (H1 H2); [apply (ltwr ? n2); rewrite < sym_plus; apply le_n;] +cases (Hn2 N ?) (H3 H4); [apply (ltwl ? n1); rewrite < sym_plus; apply le_n;] +clear H1 H3 Hn2 Hn1 N ge0 Axy; lapply (mtineq ?? x y (xn (S (n2+n1)))) as H5; +cut ( δx (xn (S (n2+n1)))+ δ(xn (S (n2+n1))) y <   δx y/1 + δ(xn (S (n2+n1))) y) as H6;[2: + apply flt_plusr; apply (Lt≪ ? (msymmetric ????)); assumption] +lapply (le_lt_transitive ???? H5 H6) as H7; clear H6; +cut (δx y/1+ δ(xn (S (n2+n1))) y < δx y/1+ δx y/1) as H6; [2:apply flt_plusl; assumption] +lapply (lt_transitive ???? H7 H6) as ABS; clear H6 H7 H4 H5 H2 n1 n2 xn; +lapply (divpow ? (δ x y) 1) as D; lapply (Lt≪ ? (eq_sym ??? D) ABS) as H; +change in H with ( δx y/1+ δx y/1< δx y/1+ δx y/1); +apply (lt_coreflexive ?? H); +qed. + +