-
Notifications
You must be signed in to change notification settings - Fork 0
/
Copy pathSi_bc_eq_ef_entonces_((ab)c)d_eq_((ae)f)d.lean
61 lines (51 loc) · 1.58 KB
/
Si_bc_eq_ef_entonces_((ab)c)d_eq_((ae)f)d.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
-- Si_bc_eq_ef_entonces_((ab)c)d_eq_((ae)f)d.lean
-- Si bc = ef, entonces ((ab)c)d = ((ae)f)d
-- José A. Alonso Jiménez <https://jaalonso.github.io>
-- Sevilla, 18-julio-2023
-- ---------------------------------------------------------------------
-- ---------------------------------------------------------------------
-- Demostrar que si a, b, c, d, e y f son números reales tales que
-- b * c = e * f
-- entonces
-- ((a * b) * c) * d = ((a * e) * f) * d
-- ---------------------------------------------------------------------
-- Demostración en lenguaje natural
-- ================================
-- Por la siguiente cadena de igualdades
-- ((ab)c)d
-- = (a(bc))d [por la asociativa]
-- = (a(ef))d [por la hipótesis]
-- = ((ae)f)d [por la asociativa]
-- Demostraciones con Lean4
-- ========================
import Mathlib.Data.Real.Basic
import Mathlib.Tactic
-- 1ª demostración
example
(a b c d e f : ℝ)
(h : b * c = e * f)
: ((a * b) * c) * d = ((a * e) * f) * d :=
calc
((a * b) * c) * d
= (a * (b * c)) * d := by rw [mul_assoc a]
_ = (a * (e * f)) * d := by rw [h]
_ = ((a * e) * f) * d := by rw [←mul_assoc a]
-- 2ª demostración
example
(a b c d e f : ℝ)
(h : b * c = e * f)
: ((a * b) * c) * d = ((a * e) * f) * d :=
by
rw [mul_assoc a]
rw [h]
rw [←mul_assoc a]
-- 3ª demostración
example
(a b c d e f : ℝ)
(h : b * c = e * f)
: ((a * b) * c) * d = ((a * e) * f) * d :=
by
rw [mul_assoc a, h, ←mul_assoc a]
-- Lemas usados
-- ============
-- #check (mul_assoc : ∀ (a b c : ℝ), (a * b) * c = a * (b * c))