--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "sandwich_corollary.ma".
+
+(* 3.19 *)
+definition supremum ≝
+ λR.λml:mlattice R.λxn:sequence ml.λx:ml.
+ increasing ? xn → upper_bound ? xn x ∧ xn ⇝ x.
+
+definition infimum ≝
+ λR.λml:mlattice R.λxn:sequence ml.λx:ml.
+ decreasing ? xn → lower_bound ? xn x ∧ xn ⇝ x.
+
+(* 3.20 *)
+lemma supremum_uniq:
+ ∀R.∀ml:mlattice R.∀xn:sequence ml.increasing ? xn → (* BUG again the wrong coercion is chosen *)
+ ∀x,y:apart_of_metric_space ? ml.supremum ?? xn x → supremum ?? xn y → x ≈ y.
+intros (R ml xn Hxn x y Sx Sy);
+elim (Sx Hxn) (_ Hx); elim (Sy Hxn) (_ Hy);
+apply (tends_uniq ?? xn ?? Hx Hy);
+qed.
+
+(* 3.21 *)