]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/dama/metric_lattice.ma
the is_structure is a bad idea if you don't have canonical structures.
[helm.git] / helm / software / matita / dama / metric_lattice.ma
index 0e92f0a89c42c8b18215b3459befa2e17c40c1e3..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,24 +33,20 @@ 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
 }.
 
 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.
 
@@ -76,7 +72,7 @@ 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.
 
@@ -84,83 +80,24 @@ 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) ??)); 
+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)); [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≈ (δ(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))); [apply feq_plusr; apply (meq_l ????? (join_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))); [apply (meq_l ????? (meet_comm ???));]
+apply (Eq≈ (δ(y∧x) (y∧z)) (meq_l ????? (meet_comm ?x y)));
 apply eq_reflexive;   
 qed.  
 
-include "sequence.ma".
-include "nat/plus.ma".
-
-definition d2s ≝ 
-  λR.λml:mlattice R.λs:sequence (pordered_set_of_excedence 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).    
-*)
-
-alias symbol "leq" = "ordered sets less or equal than".
-alias symbol "and" = "constructive and".
-alias symbol "exists" = "constructive exists (Type)".
-lemma carabinieri: (* non trova la coercion .... *)
-  ∀R.∀ml:mlattice R.∀an,bn,xn:sequence (pordered_set_of_excedence ml).
-    (∀n. (an n ≤ xn n) ∧ (xn n ≤ bn n)) →
-    ∀x:ml. tends0 ? (d2s ? ml an x) → tends0 ? (d2s ? ml bn x) →
-    tends0 ? (d2s ? ml xn x).
-intros (R ml an bn xn H x Ha Hb); unfold tends0 in Ha Hb ⊢ %. unfold d2s in Ha Hb ⊢ %.
-intros (e He);
-elim (Ha ? He) (n1 H1); clear Ha; elim (Hb e He) (n2 H2); clear Hb;
-constructor 1; [apply (n1 + n2);] intros (n3 Hn3);
-cut (n1<n3) [2: 
-  generalize in match Hn3; elim n2; [rewrite > sym_plus in H3; assumption]
-  apply H3; rewrite > sym_plus in H4; simplify in H4; apply lt_S_to_lt; 
-  rewrite > sym_plus in H4; assumption;]
-elim (H1 ? Hcut) (H3 H4); clear Hcut;
-cut (n2<n3) [2: 
-  generalize in match Hn3; elim n1; [assumption]
-  apply H5; simplify in H6; apply lt_S_to_lt; assumption] 
-elim (H2 ? Hcut) (H5 H6); clear Hcut;
-elim (H n3) (H7 H8); clear H H1 H2; 
-lapply (le_to_eqm ??? H7) as Dxm; lapply (le_to_eqj ??? H7) as Dym;
-cut (δ (xn n3) x ≤ δ (bn n3) x + δ (an n3) x + δ (an n3) x); [2:
-  apply (le_transitive ???? (mtineq ???? (an n3)));
-  lapply (le_mtri ????? H7 H8);
-  lapply (feq_plusr ? (- δ(xn n3) (bn n3)) ?? Hletin);
-  cut ( δ(an n3) (bn n3)+- δ(xn n3) (bn n3)≈( δ(an n3) (xn n3))); [2:
-    apply (Eq≈  (0 + δ(an n3) (xn n3)) ? (zero_neutral ??));
-    apply (Eq≈  (δ(an n3) (xn n3) + 0) ? (plus_comm ???));
-    apply (Eq≈  (δ(an n3) (xn n3) +  (-δ(xn n3) (bn n3) +  δ(xn n3) (bn n3))) ? (opp_inverse ??));
-    apply (Eq≈  (δ(an n3) (xn n3) +  (δ(xn n3) (bn n3) + -δ(xn n3) (bn n3))) ? (plus_comm ?? (δ(xn n3) (bn n3))));
-    apply (Eq≈  ? ? (eq_sym ??? (plus_assoc ????))); assumption;] clear Hletin1;
-  apply (le_rewl ???  ( δ(an n3) (xn n3)+ δ(an n3) x));[
-    apply feq_plusr; apply msymmetric;]
-  apply (le_rewl ???  (δ(an n3) (bn n3)+- δ(xn n3) (bn n3)+ δ(an n3) x));[
-    apply feq_plusr; assumption;]
-  ]
-[2: split; [
-  apply (lt_le_transitive ????? (mpositive ????));
-  split; elim He; [apply  le_zero_x_to_le_opp_x_zero|
-  cases t1; 
-    [left; apply exc_zero_opp_x_to_exc_x_zero;
-     apply (Ex≫ ? (eq_opp_opp_x_x ??));
-    |right;  apply exc_opp_x_zero_to_exc_zero_x;
-     apply (Ex≪ ? (eq_opp_opp_x_x ??));]] assumption;]
-clear Dxm Dym H7 H8 Hn3 H5 H3 n1 n2;
 
-  
 (* 3.17 conclusione: δ x y ≈ 0 *)
 (* 3.20 conclusione: δ x y ≈ 0 *)
 (* 3.21 sup forte