-
Notifications
You must be signed in to change notification settings - Fork 0
/
Copy pathSuma_constante_es_suprayectiva.lean
83 lines (68 loc) · 1.98 KB
/
Suma_constante_es_suprayectiva.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
79
80
81
82
83
-- Suma_constante_es_suprayectiva.lean
-- La función (x ↦ x + c) es suprayectiva.
-- José A. Alonso Jiménez <https://jaalonso.github.io>
-- Sevilla, 8-noviembre-2023
-- ---------------------------------------------------------------------
-- ---------------------------------------------------------------------
-- Demostrar que para todo número real c, la función
-- f(x) = x + c
-- es suprayectiva.
-- ----------------------------------------------------------------------
-- Demostración en lenguaje natural
-- ================================
-- Tenemos que demostrar que
-- (∀ x ∈ ℝ)(∃ y ∈ ℝ)[y+c = x]
-- Sea x ∈ ℝ. Entonces, y = x-c ∈ ℝ y
-- y + c = (x - c) + c
-- = x
-- Demostraciones con Lean4
-- ========================
import Mathlib.Data.Real.Basic
import Mathlib.Tactic
variable {c : ℝ}
open Function
-- 1ª demostración
example : Surjective (fun x ↦ x + c) :=
by
intro x
-- x : ℝ
-- ⊢ ∃ a, (fun x => x + c) a = x
use x - c
-- ⊢ (fun x => x + c) (x - c) = x
dsimp
-- ⊢ (x - c) + c = x
exact sub_add_cancel x c
-- 2ª demostración
example : Surjective (fun x ↦ x + c) :=
by
intro x
-- x : ℝ
-- ⊢ ∃ a, (fun x => x + c) a = x
use x - c
-- ⊢ (fun x => x + c) (x - c) = x
change (x - c) + c = x
-- ⊢ (x - c) + c = x
exact sub_add_cancel x c
-- 3ª demostración
example : Surjective (fun x ↦ x + c) :=
by
intro x
-- x : ℝ
-- ⊢ ∃ a, (fun x => x + c) a = x
use x - c
-- ⊢ (fun x => x + c) (x - c) = x
exact sub_add_cancel x c
-- 4ª demostración
example : Surjective (fun x ↦ x + c) :=
fun x ↦ ⟨x - c, sub_add_cancel x c⟩
-- 5ª demostración
example : Surjective (fun x ↦ x + c) :=
fun x ↦ ⟨x - c, by ring⟩
-- 6ª demostración
example : Surjective (fun x ↦ x + c) :=
add_right_surjective c
-- Lemas usados
-- ============
-- variable (a b : ℝ)
-- #check (sub_add_cancel a b : (a - b) + b = a)
-- #check (add_right_surjective c : Surjective (fun x ↦ x + c))