-
Notifications
You must be signed in to change notification settings - Fork 0
/
Copy pathComposicion_de_funciones_monotonas.lean
78 lines (68 loc) · 2.06 KB
/
Composicion_de_funciones_monotonas.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
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
-- Composicion_de_funciones_monotonas.lean
-- La composición de dos funciones monótonas es monótona.
-- José A. Alonso Jiménez <https://jaalonso.github.io>
-- Sevilla, 12-octubre-2023
-- ---------------------------------------------------------------------
-- ---------------------------------------------------------------------
-- Demostrar que la composición de dos funciones monótonas es monótona.
-- ----------------------------------------------------------------------
-- Demostración en lenguaje natural
-- ================================
-- Sean f y g dos funciones monótonas de ℝ en ℝ. Tenemos que demostrar
-- que f ∘ g es monótona; es decir, que
-- (∀ a, b ∈ ℝ) [a ≤ b → (f ∘ g)(a) ≤ (f ∘ g)(b)]
-- Sean a, b ∈ ℝ tales que a ≤ b. Por ser g monótona, se tiene
-- g(a) ≤ g(b)
-- y, por ser f monótona, se tiene
-- f(g(a)) ≤ f(g(b))
-- Finalmente, por la definición de composición,
-- (f ∘ g)(a) ≤ (f ∘ g)(b)
-- que es lo que había que demostrar.
-- Demostraciones con Lean4
-- ========================
import Mathlib.Data.Real.Basic
variable (f g : ℝ → ℝ)
-- 1ª demostración
example
(mf : Monotone f)
(mg : Monotone g)
: Monotone (f ∘ g) :=
by
have h1 : ∀ a b, a ≤ b → (f ∘ g) a ≤ (f ∘ g) b := by
{ intros a b hab
have h1 : g a ≤ g b := mg hab
show (f ∘ g) a ≤ (f ∘ g) b
exact mf h1 }
show Monotone (f ∘ g)
exact h1
-- 2ª demostración
example
(mf : Monotone f)
(mg : Monotone g)
: Monotone (f ∘ g) :=
by
have h1 : ∀ a b, a ≤ b → (f ∘ g) a ≤ (f ∘ g) b := by
{ intros a b hab
show (f ∘ g) a ≤ (f ∘ g) b
exact mf (mg hab) }
show Monotone (f ∘ g)
exact h1
-- 3ª demostración
example
(mf : Monotone f)
(mg : Monotone g)
: Monotone (f ∘ g) :=
by
-- a b : ℝ
-- hab : a ≤ b
intros a b hab
-- (f ∘ g) a ≤ (f ∘ g) b
apply mf
-- g a ≤ g b
apply mg
-- a ≤ b
apply hab
-- 4ª demostración
example (mf : Monotone f) (mg : Monotone g) :
Monotone (f ∘ g) :=
λ _ _ hab ↦ mf (mg hab)