intro Xyz; apply Exy; apply unfold_apart; right; assumption;
qed.
-notation > "'Ex'≪" non associative with precedence 50 for
- @{'excessrewritel}.
-
-interpretation "exc_rewl" 'excessrewritel =
- (cic:/matita/excess/exc_rewl.con _ _ _).
-
lemma le_rewr: ∀E:excess.∀z,y,x:E. x ≈ y → z ≤ x → z ≤ y.
intros (E z y x Exy Lxz); apply (le_transitive ???? Lxz);
intro Xyz; apply Exy; apply unfold_apart; left; assumption;
qed.
-notation > "'Ex'≫" non associative with precedence 50 for
- @{'excessrewriter}.
+notation > "'Le'≪" non associative with precedence 50 for
+ @{'lerewritel}.
-interpretation "exc_rewr" 'excessrewriter =
- (cic:/matita/excess/exc_rewr.con _ _ _).
+interpretation "le_rewl" 'lerewritel =
+ (cic:/matita/excess/le_rewl.con _ _ _).
+
+notation > "'Le'≫" non associative with precedence 50 for
+ @{'lerewriter}.
+
+interpretation "le_rewr" 'lerewriter =
+ (cic:/matita/excess/le_rewr.con _ _ _).
lemma ap_rewl: ∀A:apartness.∀x,z,y:A. x ≈ y → y # z → x # z.
intros (A x z y Exy Ayz); cases (ap_cotransitive ???x Ayz); [2:assumption]
elim (Exy); left; assumption;
qed.
+notation > "'Ex'≪" non associative with precedence 50 for
+ @{'excessrewritel}.
+
+interpretation "exc_rewl" 'excessrewritel =
+ (cic:/matita/excess/exc_rewl.con _ _ _).
+
+notation > "'Ex'≫" non associative with precedence 50 for
+ @{'excessrewriter}.
+
+interpretation "exc_rewr" 'excessrewriter =
+ (cic:/matita/excess/exc_rewr.con _ _ _).
+
lemma lt_rewr: ∀A:excess.∀x,z,y:A. x ≈ y → z < y → z < x.
intros (A x y z E H); split; elim H;
[apply (le_rewr ???? (eq_sym ??? E));|apply (ap_rewr ???? E)] assumption;
[apply (le_rewl ???? (eq_sym ??? E));| apply (ap_rewl ???? E);] assumption;
qed.
+notation > "'Lt'≪" non associative with precedence 50 for
+ @{'ltrewritel}.
+
+interpretation "lt_rewl" 'ltrewritel =
+ (cic:/matita/excess/lt_rewl.con _ _ _).
+
+notation > "'Lt'≫" non associative with precedence 50 for
+ @{'ltrewriter}.
+
+interpretation "lt_rewr" 'ltrewriter =
+ (cic:/matita/excess/lt_rewr.con _ _ _).
+
lemma lt_le_transitive: ∀A:excess.∀x,y,z:A.x < y → y ≤ z → x < z.
intros (A x y z LT LE); cases LT (LEx APx); split; [apply (le_transitive ???? LEx LE)]
whd in LE LEx APx; cases APx (EXx EXx); [cases (LEx EXx)]
include "sequence.ma".
include "metric_lattice.ma".
-lemma sandwich_ineq: ∀G:todgroup.∀e:G.0<e → e/3 + e/2 + e/2 < e.
-intro G; cases G; unfold divide; intro e;
-alias num (instance 0) = "natural number".
-cases (dg_prop (mk_todgroup todg_order todg_division_ H) e 3) 0;
-cases (dg_prop (mk_todgroup todg_order todg_division_ H) e 2) 0; simplify;
-intro H3;
-cut (0<w1) as H4; [2:
- apply (ltmul_lt ??? 2); apply (lt_rewr ???? H2);
- apply (lt_rewl ???? (mulzero ? 3)); assumption]
-cut (0<w) as H5; [2:
- apply (ltmul_lt ??? 3); apply (lt_rewr ???? H1);
- apply (lt_rewl ???? (mulzero ? 4)); assumption]
-cut (w<w1) as H6; [2:
- apply (muleqplus_lt ??? 2 O); try assumption; apply (Eq≈ ? H2 H1);]
-apply (plus_cancr_lt ??? w1);
-apply (lt_rewl ??? (w+e)); [
- apply (Eq≈ (w+3*w1) ? H2);
- apply (Eq≈ (w+w1+(w1+w1)) (plus_assoc ??w1 w1));
- apply (Eq≈ (w+(w1+(w1+w1))) (plus_assoc ?w w1 ?));
- simplify; repeat apply feq_plusl; apply eq_sym;
- apply (Eq≈ ? (plus_comm ???)); apply zero_neutral;]
-apply (lt_rewl ???? (plus_comm ???));
-apply flt_plusl; assumption;
-qed.
-
definition d2s ≝
λR.λml:mlattice R.λs:sequence ml.λk.λn. δ (s n) k.
theorem sandwich:
∀R.∀ml:mlattice R.∀an,bn,xn:sequence ml.∀x:ml.
(∀n. (an n ≤ xn n) ∧ (xn n ≤ bn n)) →
- 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);
-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;
-elim (H n3) (H7 H8); clear H H1 H2;
-lapply (le_to_eqm ??? H7) as Dxm; lapply (le_to_eqj ??? H7) as Dym;
-(* the main step *)
+ 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);
+alias num (instance 0) = "natural number".
+cases (Ha (e/2) (divide_preserves_lt ??? He)) (n1 H1); clear Ha;
+cases (Hb (e/2) (divide_preserves_lt ??? He)) (n2 H2); clear Hb;
+apply (ex_introT ?? (n1+n2)); intros (n3 Lt_n1n2_n3);
+lapply (ltwr ??? Lt_n1n2_n3) as Lt_n1n3; lapply (ltwl ??? Lt_n1n2_n3) as Lt_n2n3;
+cases (H1 ? Lt_n1n3) (c daxe); cases (H2 ? Lt_n2n3) (c dbxe);
+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;
+ 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≈ ? ? (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);
- clear Dym Dxm H8 H7 H6 H5 H4 H3;
- apply (le_rewl ??? (- δ(xn n3) (bn n3)+ δ(an n3) (bn n3)+δ(an n3) x) (plus_comm ??(δ(an n3) (bn n3))));
- apply (le_rewl ??? (- δ(xn n3) (bn n3)+ (δ(an n3) (bn n3)+δ(an n3) x)) (plus_assoc ????));
- apply (le_rewl ??? ((δ(an n3) (bn n3)+δ(an n3) x)+- δ(xn n3) (bn n3)) (plus_comm ???));
+ 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 ???));
apply lew_opp; [apply mpositive] apply fle_plusr;
- apply (le_rewr ???? (plus_comm ???));
- apply (le_rewr ??? ( δ(an n3) x+ δx (bn n3)) (msymmetric ????));
+ apply (Le≫ ? (plus_comm ???));
+ apply (Le≫ ( δ(an n3) x+ δx (bn n3)) (msymmetric ????));
apply mtineq;]
-split; [ (* first the trivial case *)
+split; [ (* first the trivial case: -e< δ(xn n3) x *)
apply (lt_le_transitive ????? (mpositive ????));
- apply lt_zero_x_to_lt_opp_x_zero; assumption;]
-clear Dxm Dym H7 H8 Hn3 H5 H3 n1 n2; cases H4 (H7 H8); clear H4;
-apply (le_lt_transitive ???? ? (sandwich_ineq ?? He));
-apply (le_transitive ???? main_ineq);
-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; assumption]
-apply (le_transitive ?? (e/3+ e/2 + e/2)); [apply fle_plusl; assumption;]
-apply le_reflexive;
+ apply lt_zero_x_to_lt_opp_x_zero; assumption;]
+(* the main goal: δ(xn n3) x<e *)
+apply (le_lt_transitive ???? main_ineq);
+apply (Lt≫ (e/2+e/2+e/2)); [apply (divpow ?e 2)]
+repeat (apply ltplus; try assumption);
qed.