]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/dama/metric_lattice.ma
Added summation formula for the power of a binomial.
[helm.git] / helm / software / matita / dama / metric_lattice.ma
index 56120525d8cf3d794f6889000b7cbc1e12f8a995..0bfc3db678179c3ceedaab49b34ff762961f6f75 100644 (file)
@@ -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,33 +33,78 @@ 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 ≝ {
+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
 }.
 
-axiom meq_joinl: ∀R.∀ml:mlattice R.∀x,y,z:ml. x≈z → δx y ≈ δz y.
+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;
+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; split [apply mpositive] apply ap_symmetric;
+assumption;
+qed.
+
+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_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_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;
+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 foo: ∀R.∀ml:mlattice R.∀x,y,z:ml. δx y ≈ δ(y∨x) (y∨z).
-intros; apply le_le_eq; [
-  apply (le_rewr ???? (meq_joinl ????? (join_comm ???)));
-  apply (le_rewr ???? (meq_joinr ????? (join_comm ???)));
-*) 
+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; 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∧x) ??)); 
-(* auto type. assert failure *)
+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;
+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≈ ? (plus_comm ???));
+apply (Eq≈ (δ(y∨x) (y∨z)+ δ(x∧y) (y∧z)) (meq_l ????? (join_comm ?x y)));
+apply feq_plusl;
+apply (Eq≈ (δ(y∧x) (y∧z)) (meq_l ????? (meet_comm ?x y)));
+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 *)
\ No newline at end of file