X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fdama%2Fsandwich.ma;h=47709bd24c78f5bbfd27b0d528decf4e6741dd4c;hb=156e9d172f3fd2fd7f2d5f0f88c8662c2c0fd796;hp=95dd0d3cbc0c5add6a783457b962f96e791509ce;hpb=6067115471521e8b9ea805531cb94a0a80774314;p=helm.git diff --git a/helm/software/matita/dama/sandwich.ma b/helm/software/matita/dama/sandwich.ma index 95dd0d3cb..47709bd24 100644 --- a/helm/software/matita/dama/sandwich.ma +++ b/helm/software/matita/dama/sandwich.ma @@ -12,7 +12,7 @@ (* *) (**************************************************************************) -set "baseuri" "cic:/matita/sandwich/". + include "nat/plus.ma". include "nat/orders.ma". @@ -36,11 +36,12 @@ notation "s ⇝ x" non associative with precedence 50 for @{'tends $s $x}. interpretation "tends to" 'tends s x = (cic:/matita/sequence/tends0.con _ (cic:/matita/sandwich/d2s.con _ _ s x)). - -alias symbol "and" = "constructive and". + theorem sandwich: +let ugo ≝ excess_OF_lattice1 in ∀R.∀ml:mlattice R.∀an,bn,xn:sequence ml.∀x:ml. - (∀n. (an n ≤ xn n) ∧ (xn n ≤ bn n)) → + (∀n. (le (excess_OF_lattice1 ml) (xn n) (an n)) ∧ + (le (excess_OF_lattice1 ?) (bn n) (xn n))) → True. an ⇝ x → bn ⇝ x → xn ⇝ x. intros (R ml an bn xn x H Ha Hb); unfold tends0 in Ha Hb ⊢ %; unfold d2s in Ha Hb ⊢ %; intros (e He); @@ -54,16 +55,16 @@ cases (H n3) (H7 H8); clear Lt_n1n3 Lt_n2n3 Lt_n1n2_n3 c H1 H2 H n1 n2; (* the main inequality *) cut (δ (xn n3) x ≤ δ (bn n3) x + δ (an n3) x + δ (an n3) x) as main_ineq; [2: apply (le_transitive ???? (mtineq ???? (an n3))); - lapply (le_mtri ????? H7 H8) as H9; clear H7 H8; - lapply (feq_plusr ? (- δ(xn n3) (bn n3)) ?? H9) as H10; clear H9; cut ( δ(an n3) (bn n3)+- δ(xn n3) (bn n3)≈( δ(an n3) (xn n3))) as H11; [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 H10; - apply (le_rewl ???  ( δ(an n3) (xn n3)+ δ(an n3) x) (msymmetric ??(an n3)(xn n3))); - apply (le_rewl ???  (δ(an n3) (bn n3)+- δ(xn n3) (bn n3)+ δ(an n3) x) H11); + lapply (le_mtri ?? ??? (ge_to_le ??? H7) (ge_to_le ??? H8)) as H9; clear H7 H8; + lapply (feq_plusr ? (- δ(xn n3) (bn n3)) ?? H9) as H10; clear H9; + 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;] + apply (Le≪ (δ(an n3) (xn n3)+ δ(an n3) x) (msymmetric ??(an n3)(xn n3))); + apply (Le≪ (δ(an n3) (bn n3)+- δ(xn n3) (bn n3)+ δ(an n3) x) H11); apply (Le≪ (- δ(xn n3) (bn n3)+ δ(an n3) (bn n3)+δ(an n3) x) (plus_comm ??(δ(an n3) (bn n3)))); apply (Le≪ (- δ(xn n3) (bn n3)+ (δ(an n3) (bn n3)+δ(an n3) x)) (plus_assoc ????)); apply (Le≪ ((δ(an n3) (bn n3)+δ(an n3) x)+- δ(xn n3) (bn n3)) (plus_comm ???));