]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/dama/metric_lattice.ma
ok
[helm.git] / helm / software / matita / dama / metric_lattice.ma
index 81eab3137e9f121d28f84a96e4f29c9b4b62b129..ae55b5551dae807bd16ab72fae962914c0e70b4d 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,12 +33,12 @@ 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
 }.
@@ -103,40 +103,6 @@ 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).    
-*)
-
-lemma lew_opp : ∀O:ogroup.∀a,b,c:O.0 ≤ b → a ≤ c → a + -b ≤ c.
-intros (O a b c L0 L);
-apply (le_transitive ????? L);
-apply (plus_cancl_le ??? (-a));
-apply (le_rewr ??? 0 (opp_inverse ??));
-apply (le_rewl ??? (-a+a+-b) (plus_assoc ????));
-apply (le_rewl ??? (0+-b) (opp_inverse ??));
-apply (le_rewl ??? (-b) (zero_neutral ?(-b)));
-apply le_zero_x_to_le_opp_x_zero;
-assumption;
-qed.
-
-
-lemma ltw_opp : ∀O:ogroup.∀a,b,c:O.0 < b → a < c → a + -b < c.
-intros (O a b c P L);
-apply (lt_transitive ????? L);
-apply (plus_cancl_lt ??? (-a));
-apply (lt_rewr ??? 0 (opp_inverse ??));
-apply (lt_rewl ??? (-a+a+-b) (plus_assoc ????));
-apply (lt_rewl ??? (0+-b) (opp_inverse ??));
-apply (lt_rewl ??? ? (zero_neutral ??));
-apply lt_zero_x_to_lt_opp_x_zero;
-assumption;
-qed.
-
 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;
@@ -146,17 +112,31 @@ 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<e → e/3 + e/2 + e/2 < e.
+
 alias symbol "leq" = "ordered sets less or equal than".
 alias symbol "and" = "constructive and".
 alias symbol "exists" = "constructive exists (Type)".
-theorem carabinieri: (* non trova la coercion .... *)
-  ∀R.∀ml:mlattice R.∀an,bn,xn:sequence (pordered_set_of_excedence ml).
+theorem carabinieri:
+  ∀R.∀ml:mlattice R.∀an,bn,xn:sequence 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;
+alias num (instance 0) = "natural number".
+elim (Ha (e/2) (divide_preserves_lt ??? He)) (n1 H1); clear Ha; 
+elim (Hb (e/3) (divide_preserves_lt ??? He)) (n2 H2); clear Hb;
 constructor 1; [apply (n1 + n2);] intros (n3 Hn3);
 lapply (ltwr ??? Hn3) as Hn1n3; lapply (ltwl ??? Hn3) as Hn2n3;
 elim (H1 ? Hn1n3) (H3 H4); elim (H2 ? Hn2n3) (H5 H6); clear Hn1n3 Hn2n3;
@@ -195,6 +175,16 @@ split; [
   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;
+apply (le_lt_transitive ???? ? (core1 ?? He));
+apply (le_transitive ???? Hcut);
+apply (le_transitive ??  (e/3+ δ(an n3) x+ δ(an n3) x)); [
+  repeat apply fle_plusr; cases H6; assumption;]
+apply (le_transitive ??  (e/3+ e/2 + δ(an n3) x)); [
+  apply fle_plusr; apply fle_plusl; cases H4; assumption;]
+apply (le_transitive ??  (e/3+ e/2 + e/2)); [
+  apply fle_plusl; cases H4; assumption;]
+apply le_reflexive;
+qed.
 
   
 (* 3.17 conclusione: δ x y ≈ 0 *)