-
Notifications
You must be signed in to change notification settings - Fork 0
/
Copy pathMenor_por_intermedio.lean
53 lines (40 loc) · 1.44 KB
/
Menor_por_intermedio.lean
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
-- Menor_por_intermedio.lean
-- Si (∃z ∈ ℝ)[x < z < y], entonces x < y.
-- José A. Alonso Jiménez <https://jaalonso.github.io>
-- Sevilla, 15-diciembre-2023
-- ---------------------------------------------------------------------
-- ---------------------------------------------------------------------
-- Demostrar que si (∃z ∈ ℝ)[x < z < y], entonces x < y.
-- ---------------------------------------------------------------------
-- Demostración en lenguaje natural
-- ================================
-- Sea z tal que verifica las siguientes relaciones:
-- x < z (1)
-- z < y (2)
-- Aplicando la propiedad transitiva a (1) y (2) se tiene que
-- x < y.
-- Demostraciones con Lean4
-- ========================
import Mathlib.Data.Real.Basic
variable (x y : ℝ)
-- 1ª demostración
-- ===============
example : (∃ z : ℝ, x < z ∧ z < y) → x < y :=
by
rintro ⟨z, h1 : x < z, h2 : z < y⟩
show x < y
exact lt_trans h1 h2
-- 2ª demostración
-- ===============
example : (∃ z : ℝ, x < z ∧ z < y) → x < y :=
by
rintro ⟨z, h1, h2⟩
exact lt_trans h1 h2
-- 3ª demostración
-- ===============
example : (∃ z : ℝ, x < z ∧ z < y) → x < y :=
fun ⟨_, h1, h2⟩ ↦ lt_trans h1 h2
-- Lemas usados
-- ============
-- variable (a b c : ℝ)
-- #check (lt_trans : a < b → b < c → a < c)