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=917e2a199037d55157ccae646a30bd5da2c66eac;hpb=0e93f77172427eed198b974e32c7f3e80d2c0251;p=helm.git diff --git a/helm/software/matita/dama/metric_lattice.ma b/helm/software/matita/dama/metric_lattice.ma index 917e2a199..f0242da28 100644 --- a/helm/software/matita/dama/metric_lattice.ma +++ b/helm/software/matita/dama/metric_lattice.ma @@ -12,20 +12,18 @@ (* *) (**************************************************************************) -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) + 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,34 +31,30 @@ qed. coercion cic:/matita/metric_lattice/ml_mspace.con. -record is_mlattice (R : ogroup) (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 ≝ { +alias symbol "plus" = "Abelian group plus". +alias symbol "leq" = "Excess less or equal than". +record mlattice (R : todgroup) : Type ≝ { ml_carr :> mlattice_ R; - ml_props:> is_mlattice R ml_carr + 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 ?? ml); +intros (R ml a b E); intro H; apply E; apply ml_prop1; 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; +apply ml_prop1; split [apply mpositive] apply ap_symmetric; assumption; qed. -(* -lemma lt_to_dpos: ∀R.∀ml:mlattice R.∀x,y:ml.x < y → 0 < δ x y. -intros 4; repeat (unfold in ⊢ (? % ? ?→?)); simplify; unfold excl; -intro H; elim H (H1 H2); elim H2 (H3 H3); [cases (H1 H3)] -split; [apply mpositive] -*) - 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)); @@ -77,16 +71,12 @@ intros; apply (eq_trans ???? (msymmetric ??y x)); apply (eq_trans ????? (msymmetric ??z y)); apply meq_l; assumption; qed. -lemma ap_le_to_lt: ∀O:ogroup.∀a,c:O.c # a → c ≤ a → c < a. -intros (R a c A L); split; 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;] +intros (R ml x y H); apply ml_prop1; split; [apply mpositive;] apply ap_symmetric; assumption; qed. @@ -94,23 +84,34 @@ qed. 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) ??)); - 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_trans ?? (δ(x∧y) y + δy z)); [apply feq_plusr; apply (meq_l ????? Dxm);] - apply (eq_trans ?? (δ(x∧y) (y∧z) + δy z)); [apply feq_plusr; apply (meq_r ????? Dym);] - apply (eq_trans ?? (δ(x∧y) (y∧z) + δ(x∨y) z)); [apply feq_plusl; apply (meq_l ????? Dxj);] - apply (eq_trans ?? (δ(x∧y) (y∧z) + δ(x∨y) (y∨z))); [apply feq_plusl; apply (meq_r ????? Dyj);] - apply (eq_trans ?? ? ? (plus_comm ???)); - apply (eq_trans ?? (δ(y∨x) (y∨z)+ δ(x∧y) (y∧z))); [apply feq_plusr; apply (meq_l ????? (join_comm ???));] - apply feq_plusl; - apply (eq_trans ?? (δ(y∧x) (y∧z))); [apply (meq_l ????? (meet_comm ???));] - apply eq_reflexive; +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 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));[ + 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)));[ + apply feq_plusr; apply meq_r; apply (join_comm ??);] +apply feq_plusl; +apply (Eq≈ (δ(y∧x) (y∧z)) (meq_l ????? (meet_comm ??))); +apply eq_reflexive; qed. - +(* 3.17 conclusione: δ x y ≈ 0 *) +(* 3.20 conclusione: δ x y ≈ 0 *) +(* 3.21 sup forte + strong_sup x ≝ ∀n. s n ≤ x ∧ ∀y x ≰ y → ∃n. s n ≰ y + strong_sup_zoli x ≝ ∀n. s n ≤ x ∧ ∄y. y#x ∧ y ≤ x +*) +(* 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 *)