Menu Close

Etiqueta: L.clear

Si a es un punto de acumulación de u, entonces ∀ε>0, ∀ N, ∃k≥N, |u(k)−a| < ε

Para extraer una subsucesión se aplica una función de extracción que conserva el orden; por ejemplo, la subsucesión

   uₒ, u₂, u₄, u₆, ...

se ha obtenido con la función de extracción φ tal que φ(n) = 2*n.

En Lean, se puede definir que φ es una función de extracción por

   def extraccion (φ : ℕ → ℕ) :=
     ∀ n m, n < m → φ n < φ m

También se puede definir que a es un límite de u por

   def limite (u : ℕ → ℝ) (a : ℝ) :=
     ∀ ε > 0, ∃ N, ∀ k ≥ N, |u k - a| < ε

Los puntos de acumulación de una sucesión son los límites de sus subsucesiones. En Lean se puede definir por

   def punto_acumulacion (u : ℕ → ℝ) (a : ℝ) :=
     ∃ φ, extraccion φ ∧ limite (u ∘ φ) a

Demostrar que si a es un punto de acumulación de u, entonces

   ∀ ε > 0, ∀ N, ∃ k ≥ N, |u k - a| < ε

Para ello, completar la siguiente teoría de Lean:

import data.real.basic
open nat
 
variable  {u :   }
variables {a : }
variable  {φ :   }
 
def extraccion (φ :   ):=
   n m, n < m  φ n < φ m
 
notation `|`x`|` := abs x
 
def limite (u :   ) (a : ) :=
   ε > 0,  N,  k  N, |u k - a| < ε
 
def punto_acumulacion (u :   ) (a : ) :=
   φ, extraccion φ  limite (u  φ) a
 
example
  (h : punto_acumulacion u a)
  :  ε > 0,  N,  k  N, |u k - a| < ε :=
sorry

[expand title=»Soluciones con Lean»]

import data.real.basic
open nat
 
variable  {u :   }
variables {a : }
variable  {φ :   }
 
def extraccion (φ :   ):=
   n m, n < m  φ n < φ m
 
notation `|`x`|` := abs x
 
def limite (u :   ) (a : ) :=
   ε > 0,  N,  k  N, |u k - a| < ε
 
def punto_acumulacion (u :   ) (a : ) :=
   φ, extraccion φ  limite (u  φ) a
 
-- En la demostración se usarán los siguientes lemas.
 
lemma aux1
  (h : extraccion φ)
  :  n, n  φ n :=
begin
  intro n,
  induction n with m HI,
  { exact nat.zero_le (φ 0), },
  { apply nat.succ_le_of_lt,
    calc m  φ m        : HI
       ... < φ (succ m) : h m (m+1) (lt_add_one m), },
end
 
lemma aux2
  (h : extraccion φ)
  :  N N',  n  N', φ n  N :=
λ N N', ⟨max N N', ⟨le_max_right N N',
                    le_trans (le_max_left N N')
                             (aux1 h (max N N'))⟩⟩
 
-- 1ª demostración
example
  (h : punto_acumulacion u a)
  :  ε > 0,  N,  k  N, |u k - a| < ε :=
begin
  intros ε hε N,
  unfold punto_acumulacion at h,
  rcases h with ⟨φ, hφ1, hφ2⟩,
  unfold limite at2,
  cases hφ2 ε hε with N' hN',
  rcases aux2 hφ1 N N' with ⟨m, hm, hm'⟩,
  clear hφ12,
  use φ m,
  split,
  { exact hm', },
  { exact hN' m hm, },
end
 
-- 2ª demostración
example
  (h : punto_acumulacion u a)
  :  ε > 0,  N,  n  N, |u n - a| < ε :=
begin
  intros ε hε N,
  rcases h with ⟨φ, hφ1, hφ2⟩,
  cases hφ2 ε hε with N' hN',
  rcases aux2 hφ1 N N' with ⟨m, hm, hm'⟩,
  use φ m,
  exact ⟨hm', hN' m hm⟩,
end
 
-- 3ª demostración
example
  (h : punto_acumulacion u a)
  :  ε > 0,  N,  n  N, |u n - a| < ε :=
begin
  intros ε hε N,
  rcases h with ⟨φ, hφ1, hφ2⟩,
  cases hφ2 ε hε with N' hN',
  rcases aux2 hφ1 N N' with ⟨m, hm, hm'⟩,
  exact ⟨φ m, hm', hN' _ hm⟩,
end
 
-- 4ª demostración
example
  (h : punto_acumulacion u a)
  :  ε > 0,  N,  n  N, |u n - a| < ε :=
begin
  intros ε hε N,
  rcases h with ⟨φ, hφ1, hφ2⟩,
  cases hφ2 ε hε with N' hN',
  rcases aux2 hφ1 N N' with ⟨m, hm, hm'⟩,
  use φ m ; finish,
end
 
-- 5ª demostración
example
  (h : punto_acumulacion u a)
  :  ε > 0,  N,  n  N, |u n - a| < ε :=
assume ε,
assume hε : ε > 0,
assume N,
exists.elim h
  ( assume φ,
    assume hφ : extraccion φ  limite (u  φ) a,
    exists.elim (hφ.2 ε hε)
      ( assume N',
        assume hN' :  (n : ), n  N'  |(u  φ) n - a| < ε,
        have h1 :  n  N', φ n  N,
          from aux2 hφ.1 N N',
        exists.elim h1
          ( assume m,
            assume hm :  (H : m  N'), φ m  N,
            exists.elim hm
              ( assume hm1 : m  N',
                assume hm2 : φ m  N,
                have h2 : |u (φ m) - a| < ε,
                  from hN' m hm1,
                show  n  N, |u n - a| < ε,
                  from exists.intro (φ m) (exists.intro hm2 h2)))))
 
-- 6ª demostración
example
  (h : punto_acumulacion u a)
  :  ε > 0,  N,  n  N, |u n - a| < ε :=
assume ε,
assume hε : ε > 0,
assume N,
exists.elim h
  ( assume φ,
    assume hφ : extraccion φ  limite (u  φ) a,
    exists.elim (hφ.2 ε hε)
      ( assume N',
        assume hN' :  (n : ), n  N'  |(u  φ) n - a| < ε,
        have h1 :  n  N', φ n  N,
          from aux2 hφ.1 N N',
        exists.elim h1
          ( assume m,
            assume hm :  (H : m  N'), φ m  N,
            exists.elim hm
              ( assume hm1 : m  N',
                assume hm2 : φ m  N,
                have h2 : |u (φ m) - a| < ε,
                  from hN' m hm1,
                show  n  N, |u n - a| < ε,
                  from ⟨φ m, hm2, h2⟩))))
 
-- 7ª demostración
example
  (h : punto_acumulacion u a)
  :  ε > 0,  N,  n  N, |u n - a| < ε :=
assume ε,
assume hε : ε > 0,
assume N,
exists.elim h
  ( assume φ,
    assume hφ : extraccion φ  limite (u  φ) a,
    exists.elim (hφ.2 ε hε)
      ( assume N',
        assume hN' :  (n : ), n  N'  |(u  φ) n - a| < ε,
        have h1 :  n  N', φ n  N,
          from aux2 hφ.1 N N',
        exists.elim h1
          ( assume m,
            assume hm :  (H : m  N'), φ m  N,
            exists.elim hm
              ( assume hm1 : m  N',
                assume hm2 : φ m  N,
                have h2 : |u (φ m) - a| < ε,
                  from hN' m hm1,
                ⟨φ m, hm2, h2⟩))))
 
-- 8ª demostración
example
  (h : punto_acumulacion u a)
  :  ε > 0,  N,  n  N, |u n - a| < ε :=
assume ε,
assume hε : ε > 0,
assume N,
exists.elim h
  ( assume φ,
    assume hφ : extraccion φ  limite (u  φ) a,
    exists.elim (hφ.2 ε hε)
      ( assume N',
        assume hN' :  (n : ), n  N'  |(u  φ) n - a| < ε,
        have h1 :  n  N', φ n  N,
          from aux2 hφ.1 N N',
        exists.elim h1
          ( assume m,
            assume hm :  (H : m  N'), φ m  N,
            exists.elim hm
              ( assume hm1 : m  N',
                assume hm2 : φ m  N,
                ⟨φ m, hm2, hN' m hm1⟩))))
 
-- 9ª demostración
example
  (h : punto_acumulacion u a)
  :  ε > 0,  N,  n  N, |u n - a| < ε :=
assume ε,
assume hε : ε > 0,
assume N,
exists.elim h
  ( assume φ,
    assume hφ : extraccion φ  limite (u  φ) a,
    exists.elim (hφ.2 ε hε)
      ( assume N',
        assume hN' :  (n : ), n  N'  |(u  φ) n - a| < ε,
        have h1 :  n  N', φ n  N,
          from aux2 hφ.1 N N',
        exists.elim h1
          ( assume m,
            assume hm :  (H : m  N'), φ m  N,
            exists.elim hm
              (λ hm1 hm2, ⟨φ m, hm2, hN' m hm1⟩))))
 
-- 10ª demostración
example
  (h : punto_acumulacion u a)
  :  ε > 0,  N,  n  N, |u n - a| < ε :=
assume ε,
assume hε : ε > 0,
assume N,
exists.elim h
  ( assume φ,
    assume hφ : extraccion φ  limite (u  φ) a,
    exists.elim (hφ.2 ε hε)
      ( assume N',
        assume hN' :  (n : ), n  N'  |(u  φ) n - a| < ε,
        have h1 :  n  N', φ n  N,
          from aux2 hφ.1 N N',
        exists.elim h1
          (λ m hm, exists.elim hm (λ hm1 hm2, ⟨φ m, hm2, hN' m hm1⟩))))
 
-- 11ª demostración
example
  (h : punto_acumulacion u a)
  :  ε > 0,  N,  n  N, |u n - a| < ε :=
assume ε,
assume hε : ε > 0,
assume N,
exists.elim h
  ( assume φ,
    assume hφ : extraccion φ  limite (u  φ) a,
    exists.elim (hφ.2 ε hε)
      ( assume N',
        assume hN' :  (n : ), n  N'  |(u  φ) n - a| < ε,
        exists.elim (aux2 hφ.1 N N')
          (λ m hm, exists.elim hm (λ hm1 hm2, ⟨φ m, hm2, hN' m hm1⟩))))
 
-- 12ª demostración
example
  (h : punto_acumulacion u a)
  :  ε > 0,  N,  n  N, |u n - a| < ε :=
assume ε,
assume hε : ε > 0,
assume N,
exists.elim h
  ( assume φ,
    assume hφ : extraccion φ  limite (u  φ) a,
    exists.elim (hφ.2 ε hε)
      (λ N' hN', exists.elim (aux2 hφ.1 N N')
        (λ m hm, exists.elim hm
          (λ hm1 hm2, ⟨φ m, hm2, hN' m hm1⟩))))
 
-- 13ª demostración
example
  (h : punto_acumulacion u a)
  :  ε > 0,  N,  n  N, |u n - a| < ε :=
assume ε,
assume hε : ε > 0,
assume N,
exists.elim h
  (λ φ hφ, exists.elim (hφ.2 ε hε)
    (λ N' hN', exists.elim (aux2 hφ.1 N N')
      (λ m hm, exists.elim hm
        (λ hm1 hm2, ⟨φ m, hm2, hN' m hm1⟩))))
 
-- 14ª demostración
example
  (h : punto_acumulacion u a)
  :  ε > 0,  N,  n  N, |u n - a| < ε :=
λ ε hε N, exists.elim h
  (λ φ hφ, exists.elim (hφ.2 ε hε)
    (λ N' hN', exists.elim (aux2 hφ.1 N N')
      (λ m hm, exists.elim hm
        (λ hm1 hm2, ⟨φ m, hm2, hN' m hm1⟩))))

Se puede interactuar con la prueba anterior en esta sesión con Lean.

En los comentarios se pueden escribir otras soluciones, escribiendo el código entre una línea con <pre lang="lean"> y otra con </pre>
[/expand]

[expand title=»Soluciones con Isabelle/HOL»]

theory "Si_a_es_un_punto_de_acumulacion_de_u,_entonces_a_tiene_puntos_cercanos"
imports Main HOL.Real
begin
 
definition extraccion :: "(nat ⇒ nat) ⇒ bool" where
  "extraccion φ ⟷ (∀ n m. n < m ⟶ φ n < φ m)"
 
definition limite :: "(nat ⇒ real) ⇒ real ⇒ bool"
  where "limite u a ⟷ (∀ε>0. ∃N. ∀k≥N. ¦u k - a¦ < ε)"
 
definition punto_acumulacion :: "(nat ⇒ real) ⇒ real ⇒ bool"
  where "punto_acumulacion u a ⟷ (∃φ. extraccion φ ∧ limite (u ∘ φ) a)"
 
(* En la demostración se usarán los siguientes lemas *)
lemma aux1 :
  assumes "extraccion φ"
  shows   "n ≤ φ n"
proof (induct n)
  show "0 ≤ φ 0" by simp
next
  fix n assume HI : "n ≤ φ n"
  then show "Suc n ≤ φ (Suc n)"
    using assms extraccion_def
    by (metis Suc_leI lessI order_le_less_subst1)
qed
 
lemma aux2 :
  assumes "extraccion φ"
  shows   "∀ N N'. ∃ k ≥ N'. φ k ≥ N"
proof (intro allI)
  fix N N' :: nat
  have "max N N' ≥ N' ∧ φ (max N N') ≥ N"
    by (meson assms aux1 max.bounded_iff max.cobounded2)
  then show "∃k ≥ N'. φ k ≥ N"
    by blast
qed
 
(* 1ª demostración *)
lemma
  assumes "punto_acumulacion u a"
  shows   "∀ε>0. ∀ N. ∃k≥N. ¦u k - a¦ < ε"
proof (intro allI impI)
  fix ε :: real and N :: nat
  assume "ε > 0"
  obtain φ where1 : "extraccion φ"
             and hφ2 : "limite (u ∘ φ) a"
    using assms punto_acumulacion_def by blast
  obtain N' where hN' : "∀k≥N'. ¦(u ∘ φ) k - a¦ < ε"
    using2 limite_def ‹ε > 0by auto
  obtain m where hm1 : "m ≥ N'" and hm2 : "φ m ≥ N"
    using aux2 hφ1 by blast
  have "φ m ≥ N ∧ ¦u (φ m) - a¦ < ε"
    using hN' hm1 hm2 by force
  then show "∃k≥N. ¦u k - a¦ < ε"
    by auto
qed
 
(* 2ª demostración *)
lemma
  assumes "punto_acumulacion u a"
  shows   "∀ε>0. ∀ N. ∃k≥N. ¦u k - a¦ < ε"
proof (intro allI impI)
  fix ε :: real and N :: nat
  assume "ε > 0"
  obtain φ where1 : "extraccion φ"
             and hφ2 : "limite (u ∘ φ) a"
    using assms punto_acumulacion_def by blast
  obtain N' where hN' : "∀k≥N'. ¦(u ∘ φ) k - a¦ < ε"
    using2 limite_def ‹ε > 0by auto
  obtain m where "m ≥ N' ∧ φ m ≥ N"
    using aux2 hφ1 by blast
  then show "∃k≥N. ¦u k - a¦ < ε"
    using hN' by auto
qed
 
end

En los comentarios se pueden escribir otras soluciones, escribiendo el código entre una línea con <pre lang="isar"> y otra con </pre>
[/expand]

Las sucesiones convergentes son sucesiones de Cauchy

Nota: El problema de hoy lo ha escrito Sara Díaz Real y es uno de los que se encuentran en su Trabajo Fin de Máster Formalización en Lean de problemas de las Olimpiadas Internacionales de Matemáticas (IMO). Concretamente, el problema se encuentra en la página 52 junto con la demostración en lenguaje natural.


En Lean, una sucesión u₀, u₁, u₂, … se puede representar mediante una función (u : ℕ → ℝ) de forma que u(n) es uₙ.

Se define

  • el valor absoluto de x por
     notation `|`x`|` := abs x
  • a es un límite de la sucesión u, por
     def limite (u : ℕ → ℝ) (a : ℝ) :=
       ∀ ε > 0, ∃ N, ∀ n ≥ N, |u n - a| ≤ ε
  • la sucesión u es convergente por
     def suc_convergente (u : ℕ → ℝ) :=
       ∃ l, limite u l
  • la sucesión u es de Cauchy por
     def suc_cauchy (u : ℕ → ℝ) :=
       ∀ ε > 0, ∃ N, ∀ p ≥ N, ∀ q ≥ N, |u p - u q| ≤ ε

Demostrar que las sucesiones convergentes son de Cauchy.

Para ello, completar la siguiente teoría de Lean:

import data.real.basic
 
variable {u :    }
 
notation `|`x`|` := abs x
 
def limite (u :   ) (l : ) : Prop :=
   ε > 0,  N,  n  N, |u n - l|  ε
 
def suc_convergente (u :   ) :=
   l, limite u l
 
def suc_cauchy (u :   ) :=
   ε > 0,  N,  p  N,  q  N, |u p - u q|  ε
 
example
  (h : suc_convergente u)
  : suc_cauchy u :=
sorry

[expand title=»Soluciones con Lean»]

import data.real.basic
 
variable {u :    }
 
notation `|`x`|` := abs x
 
def limite (u :   ) (l : ) : Prop :=
   ε > 0,  N,  n  N, |u n - l|  ε
 
def suc_convergente (u :   ) :=
   l, limite u l
 
def suc_cauchy (u :   ) :=
   ε > 0,  N,  p  N,  q  N, |u p - u q|  ε
 
-- 1ª demostración
example
  (h : suc_convergente u)
  : suc_cauchy u :=
begin
  unfold suc_cauchy,
  intros ε hε,
  have2 : 0 < ε/2 := half_pos hε,
  cases h with l hl,
  cases hl (ε/2)2 with N hN,
  clear hε hl hε2,
  use N,
  intros p hp q hq,
  calc |u p - u q|
       = |(u p - l) + (l - u q)| : by ring_nf
   ...  |u p - l|  + |l - u q|  : abs_add (u p - l) (l - u q)
   ... = |u p - l|  + |u q - l|  : congr_arg2 (+) rfl (abs_sub_comm l (u q))
   ...  ε/2 + ε/2               : add_le_add (hN p hp) (hN q hq)
   ... = ε                       : add_halves ε,
end
 
-- 2ª demostración
example
  (h : suc_convergente u)
  : suc_cauchy u :=
begin
  intros ε hε,
  cases h with l hl,
  cases hl (ε/2) (half_pos hε) with N hN,
  clear hε hl,
  use N,
  intros p hp q hq,
  calc |u p - u q|
       = |(u p - l) + (l - u q)| : by ring_nf
   ...  |u p - l|  + |l - u q|  : abs_add (u p - l) (l - u q)
   ... = |u p - l|  + |u q - l|  : congr_arg2 (+) rfl (abs_sub_comm l (u q))
   ...  ε/2 + ε/2               : add_le_add (hN p hp) (hN q hq)
   ... = ε                       : add_halves ε,
end
 
-- 3ª demostración
example
  (h : suc_convergente u)
  : suc_cauchy u :=
begin
  intros ε hε,
  cases h with l hl,
  cases hl (ε/2) (half_pos hε) with N hN,
  clear hε hl,
  use N,
  intros p hp q hq,
  have cota1 : |u p - l|  ε / 2 := hN p hp,
  have cota2 : |u q - l|  ε / 2 := hN q hq,
  clear hN hp hq,
  calc |u p - u q|
       = |(u p - l) + (l - u q)| : by ring_nf
   ...  |u p - l|  + |l - u q|  : abs_add (u p - l) (l - u q)
   ... = |u p - l|  + |u q - l|  : by rw abs_sub_comm l (u q)
   ...  ε                       : by linarith,
end
 
-- 4ª demostración
example
  (h : suc_convergente u)
  : suc_cauchy u :=
begin
  intros ε hε,
  cases h with l hl,
  cases hl (ε/2) (half_pos hε) with N hN,
  clear hε hl,
  use N,
  intros p hp q hq,
  calc |u p - u q|
       = |(u p - l) + (l - u q)| : by ring_nf
   ...  |u p - l|  + |l - u q|  : abs_add (u p - l) (l - u q)
   ... = |u p - l|  + |u q - l|  : by rw abs_sub_comm l (u q)
   ...  ε                       : by linarith [hN p hp, hN q hq],
end
 
-- 5ª demostración
example
  (h : suc_convergente u)
  : suc_cauchy u :=
begin
  intros ε hε,
  cases h with l hl,
  cases hl (ε/2) (by linarith) with N hN,
  use N,
  intros p hp q hq,
  calc |u p - u q|
       = |(u p - l) + (l - u q)| : by ring_nf
   ...  |u p - l|  + |l - u q|  : by simp [abs_add]
   ... = |u p - l|  + |u q - l|  : by simp [abs_sub_comm]
   ...  ε                       : by linarith [hN p hp, hN q hq],
end

Se puede interactuar con la prueba anterior en esta sesión con Lean.

En los comentarios se pueden escribir otras soluciones, escribiendo el código entre una línea con <pre lang="lean"> y otra con </pre>
[/expand]

[expand title=»Soluciones con Isabelle/HOL»]

theory Las_sucesiones_convergentes_son_sucesiones_de_Cauchy
imports Main HOL.Real
begin
 
definition limite :: "(nat ⇒ real) ⇒ real ⇒ bool"
  where "limite u c ⟷ (∀ε>0. ∃k::nat. ∀n≥k. ¦u n - c¦ < ε)"
 
definition suc_convergente :: "(nat ⇒ real) ⇒ bool"
  where "suc_convergente u ⟷ (∃ l. limite u l)"
 
definition suc_cauchy :: "(nat ⇒ real) ⇒ bool"
  where "suc_cauchy u ⟷ (∀ε>0. ∃k. ∀m≥k. ∀n≥k. ¦u m - u n¦ < ε)"
 
(* 1ª demostración *)
lemma
  assumes "suc_convergente u"
  shows   "suc_cauchy u"
proof (unfold suc_cauchy_def; intro allI impI)
  fix ε :: real
  assume "0 < ε"
  then have "0 < ε/2"
    by simp
  obtain a where "limite u a"
    using assms suc_convergente_def by blast
  then obtain k where hk : "∀n≥k. ¦u n - a¦ < ε/2"
    using0 < ε / 2› limite_def by blast
  have "∀m≥k. ∀n≥k. ¦u m - u n¦ < ε"
  proof (intro allI impI)
    fix p q
    assume hp : "p ≥ k" and hq : "q ≥ k"
    then have hp' : "¦u p - a¦ < ε/2"
      using hk by blast
    have hq' : "¦u q - a¦ < ε/2"
      using hk hq by blast
    have "¦u p - u q¦ = ¦(u p - a) + (a - u q)¦"
      by simp
    also have "… ≤ ¦u p - a¦  + ¦a - u q¦"
      by simp
    also have "… = ¦u p - a¦  + ¦u q - a¦"
      by simp
    also have "… < ε/2 + ε/2"
      using hp' hq' by simp
    also have "… = ε"
      by simp
    finally show "¦u p - u q¦ < ε"
      by this
  qed
  then show "∃k. ∀m≥k. ∀n≥k. ¦u m - u n¦ < ε"
    by (rule exI)
qed
 
(* 2ª demostración *)
lemma
  assumes "suc_convergente u"
  shows   "suc_cauchy u"
proof (unfold suc_cauchy_def; intro allI impI)
  fix ε :: real
  assume "0 < ε"
  then have "0 < ε/2"
    by simp
  obtain a where "limite u a"
    using assms suc_convergente_def by blast
  then obtain k where hk : "∀n≥k. ¦u n - a¦ < ε/2"
    using0 < ε / 2› limite_def by blast
  have "∀m≥k. ∀n≥k. ¦u m - u n¦ < ε"
  proof (intro allI impI)
    fix p q
    assume hp : "p ≥ k" and hq : "q ≥ k"
    then have hp' : "¦u p - a¦ < ε/2"
      using hk by blast
    have hq' : "¦u q - a¦ < ε/2"
      using hk hq by blast
    show "¦u p - u q¦ < ε"
      using hp' hq' by argo
  qed
  then show "∃k. ∀m≥k. ∀n≥k. ¦u m - u n¦ < ε"
    by (rule exI)
qed
 
(* 3ª demostración *)
lemma
  assumes "suc_convergente u"
  shows   "suc_cauchy u"
proof (unfold suc_cauchy_def; intro allI impI)
  fix ε :: real
  assume "0 < ε"
  then have "0 < ε/2"
    by simp
  obtain a where "limite u a"
    using assms suc_convergente_def by blast
  then obtain k where hk : "∀n≥k. ¦u n - a¦ < ε/2"
    using0 < ε / 2› limite_def by blast
  have "∀m≥k. ∀n≥k. ¦u m - u n¦ < ε"
    using hk by (smt (z3) field_sum_of_halves)
  then show "∃k. ∀m≥k. ∀n≥k. ¦u m - u n¦ < ε"
    by (rule exI)
qed
 
(* 3ª demostración *)
lemma
  assumes "suc_convergente u"
  shows   "suc_cauchy u"
proof (unfold suc_cauchy_def; intro allI impI)
  fix ε :: real
  assume "0 < ε"
  then have "0 < ε/2"
    by simp
  obtain a where "limite u a"
    using assms suc_convergente_def by blast
  then obtain k where hk : "∀n≥k. ¦u n - a¦ < ε/2"
    using0 < ε / 2› limite_def by blast
  then show "∃k. ∀m≥k. ∀n≥k. ¦u m - u n¦ < ε"
    by (smt (z3) field_sum_of_halves)
qed
 
end

En los comentarios se pueden escribir otras soluciones, escribiendo el código entre una línea con <pre lang="isar"> y otra con </pre>
[/expand]