]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/dama/sandwich.ma
snapshot with more duality, almost where we left without duality
[helm.git] / helm / software / matita / dama / sandwich.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15
16
17 include "nat/plus.ma".
18 include "nat/orders.ma".
19
20 lemma ltwl: ∀a,b,c:nat. b + a < c → a < c.
21 intros 3 (x y z); elim y (H z IH H); [apply H]
22 apply IH; apply lt_S_to_lt; apply H;
23 qed.
24
25 lemma ltwr: ∀a,b,c:nat. a + b < c → a < c.
26 intros 3 (x y z); rewrite > sym_plus; apply ltwl;
27 qed. 
28
29 include "tend.ma".
30
31 alias symbol "leq" = "ordered sets less or equal than".
32 alias symbol "and" = "constructive and".
33 theorem sandwich:
34 let ugo ≝ excess_OF_lattice1 in
35   ∀R.∀ml:mlattice R.∀an,bn,xn:sequence ml.∀x:ml.
36     (∀n. (xn n ≤ an n) ∧ (bn n ≤ xn n)) →  
37     an ⇝ x → bn ⇝ x → xn ⇝ x.
38 intros (R ml an bn xn x H Ha Hb); 
39 unfold tends0 in Ha Hb ⊢ %; unfold d2s in Ha Hb ⊢ %; intros (e He);
40 alias num (instance 0) = "natural number".
41 cases (Ha (e/2) (divide_preserves_lt ??? He)) (n1 H1); clear Ha; 
42 cases (Hb (e/2) (divide_preserves_lt ??? He)) (n2 H2); clear Hb;
43 apply (ex_introT ?? (n1+n2)); intros (n3 Lt_n1n2_n3);
44 lapply (ltwr ??? Lt_n1n2_n3) as Lt_n1n3; lapply (ltwl ??? Lt_n1n2_n3) as Lt_n2n3;
45 cases (H1 ? Lt_n1n3) (c daxe); cases (H2 ? Lt_n2n3) (c dbxe); 
46 cases (H n3) (H7 H8); clear Lt_n1n3 Lt_n2n3 Lt_n1n2_n3 c H1 H2 H n1 n2;     
47 (* the main inequality *)
48 cut (δ (xn n3) x ≤ δ (bn n3) x + δ (an n3) x + δ (an n3) x) as main_ineq; [2:
49   apply (le_transitive ???? (mtineq ???? (an n3)));
50   cut ( δ(an n3) (bn n3)+- δ(xn n3) (bn n3)≈( δ(an n3) (xn n3))) as H11; [2:
51     lapply (le_mtri ?? ??? (ge_to_le ??? H7) (ge_to_le ??? H8)) as H9; clear H7 H8;
52     lapply (feq_plusr ? (- δ(xn n3) (bn n3)) ?? H9) as H10; clear H9;
53     apply (Eq≈ (0 + δ(an n3) (xn n3)) ? (zero_neutral ??));
54     apply (Eq≈ (δ(an n3) (xn n3) + 0) ? (plus_comm ???));
55     apply (Eq≈ (δ(an n3) (xn n3) + (-δ(xn n3) (bn n3) +  δ(xn n3) (bn n3))) ? (opp_inverse ??));
56     apply (Eq≈ (δ(an n3) (xn n3) + (δ(xn n3) (bn n3) + -δ(xn n3) (bn n3))) ? (plus_comm ?? (δ(xn n3) (bn n3))));
57     apply (Eq≈ ?? (eq_sym ??? (plus_assoc ????))); assumption;]
58   apply (Le≪ (δ(an n3) (xn n3)+ δ(an n3) x) (msymmetric ??(an n3)(xn n3)));    
59   apply (Le≪ (δ(an n3) (bn n3)+- δ(xn n3) (bn n3)+ δ(an n3) x) H11);
60   apply (Le≪ (- δ(xn n3) (bn n3)+ δ(an n3) (bn n3)+δ(an n3) x) (plus_comm ??(δ(an n3) (bn n3))));
61   apply (Le≪ (- δ(xn n3) (bn n3)+ (δ(an n3) (bn n3)+δ(an n3) x)) (plus_assoc ????));
62   apply (Le≪ ((δ(an n3) (bn n3)+δ(an n3) x)+- δ(xn n3) (bn n3)) (plus_comm ???));
63   apply lew_opp; [apply mpositive] apply fle_plusr;
64   apply (Le≫ ? (plus_comm ???));
65   apply (Le≫ ( δ(an n3) x+ δx (bn n3)) (msymmetric ????));
66   apply mtineq;]
67 split; [ (* first the trivial case: -e< δ(xn n3) x *)
68   apply (lt_le_transitive ????? (mpositive ????));
69   apply lt_zero_x_to_lt_opp_x_zero; assumption;]
70 (* the main goal: δ(xn n3) x<e *)
71 apply (le_lt_transitive ???? main_ineq);
72 apply (Lt≫ (e/2+e/2+e/2)); [apply (divpow ?e 2)]
73 repeat (apply ltplus; try assumption);
74 qed.