-
Notifications
You must be signed in to change notification settings - Fork 0
/
Copy pathOpuesto_del_opuesto.lean
56 lines (44 loc) · 1.38 KB
/
Opuesto_del_opuesto.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
-- Opuesto_del_opuesto.lean
-- Si R es un anillo y a ∈ R, entonces -(-a) = a.
-- José A. Alonso Jiménez <https://jaalonso.github.io>
-- Sevilla, 10-agosto-2023
-- ---------------------------------------------------------------------
-- ---------------------------------------------------------------------
-- Demostrar que si R es un anillo y a ∈ R, entonces
-- -(-a) = a
-- ----------------------------------------------------------------------
-- Demostración en lenguaje natural
-- ================================
-- Es consecuencia de las siguiente propiedades demostradas en
-- ejercicios anteriores:
-- ∀ a b ∈ R, a + b = 0 → -a = b
-- ∀ a ∈ R, -a + a = 0
-- Demostraciones con Lean4
-- ========================
import Mathlib.Algebra.Ring.Defs
import Mathlib.Tactic
variable {R : Type _} [Ring R]
variable {a : R}
-- 1ª demostración
example : -(-a) = a :=
by
have h1 : -a + a = 0 := neg_add_cancel a
show -(-a) = a
exact neg_eq_of_add_eq_zero_right h1
-- 2ª demostración
example : -(-a) = a :=
by
apply neg_eq_of_add_eq_zero_right
rw [neg_add_cancel]
-- 3ª demostración
example : -(-a) = a :=
neg_neg a
-- 4ª demostración
example : -(-a) = a :=
by simp
-- Lemas usados
-- ============
-- variable (b : R)
-- #check (neg_add_cancel a : -a + a = 0)
-- #check (neg_eq_of_add_eq_zero_right : a + b = 0 → -a = b)
-- #check (neg_neg a : -(-a) = a)