Skip to content

Commit 4267362

Browse files
This forward ports the changes introduced by leanprover-community/mathlib3#17483 No change is needed to `Mathlib.Data.EReal` as the proofs have been golfed in a different way. Co-authored-by: Eric Wieser <wieser.eric@gmail.com>
1 parent 19a7e60 commit 4267362

File tree

8 files changed

+22
-34
lines changed

8 files changed

+22
-34
lines changed

Mathlib/Algebra/Group/Basic.lean

Lines changed: 6 additions & 17 deletions
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Jeremy Avigad, Leonardo de Moura, Simon Hudon, Mario Carneiro
55
66
! This file was ported from Lean 3 source module algebra.group.basic
7-
! leanprover-community/mathlib commit 966e0cf0685c9cedf8a3283ac69eef4d5f2eaca2
7+
! leanprover-community/mathlib commit 2196ab363eb097c008d4497125e0dde23fb36db2
88
! Please do not edit these lines, except to modify the commit id
99
! if you have ported upstream changes.
1010
-/
@@ -241,21 +241,10 @@ theorem inv_inj {a b : G} : a⁻¹ = b⁻¹ ↔ a = b :=
241241
#align neg_inj neg_inj
242242

243243
@[to_additive]
244-
theorem eq_inv_of_eq_inv (h : a = b⁻¹) : b = a⁻¹ := by simp [h]
245-
#align eq_inv_of_eq_inv eq_inv_of_eq_inv
246-
#align eq_neg_of_eq_neg eq_neg_of_eq_neg
247-
248-
@[to_additive]
249-
theorem eq_inv_iff_eq_inv : a = b⁻¹ ↔ b = a⁻¹ :=
250-
⟨eq_inv_of_eq_inv, eq_inv_of_eq_inv⟩
251-
#align eq_inv_iff_eq_inv eq_inv_iff_eq_inv
252-
#align eq_neg_iff_eq_neg eq_neg_iff_eq_neg
253-
254-
@[to_additive]
255-
theorem inv_eq_iff_inv_eq : a⁻¹ = b ↔ b⁻¹ = a :=
256-
eq_comm.trans <| eq_inv_iff_eq_inv.trans eq_comm
257-
#align inv_eq_iff_inv_eq inv_eq_iff_inv_eq
258-
#align neg_eq_iff_neg_eq neg_eq_iff_neg_eq
244+
theorem inv_eq_iff_eq_inv : a⁻¹ = b ↔ a = b⁻¹ :=
245+
fun h => h ▸ (inv_inv a).symm, fun h => h.symm ▸ inv_inv b⟩
246+
#align inv_eq_iff_eq_inv inv_eq_iff_eq_inv
247+
#align neg_eq_iff_eq_neg neg_eq_iff_eq_neg
259248

260249
variable (G)
261250

@@ -663,7 +652,7 @@ theorem mul_eq_one_iff_eq_inv : a * b = 1 ↔ a = b⁻¹ :=
663652

664653
@[to_additive]
665654
theorem mul_eq_one_iff_inv_eq : a * b = 1 ↔ a⁻¹ = b :=
666-
by rw [mul_eq_one_iff_eq_inv, eq_inv_iff_eq_inv, eq_comm]
655+
by rw [mul_eq_one_iff_eq_inv, inv_eq_iff_eq_inv]
667656
#align mul_eq_one_iff_inv_eq mul_eq_one_iff_inv_eq
668657
#align add_eq_zero_iff_neg_eq add_eq_zero_iff_neg_eq
669658

Mathlib/Algebra/GroupWithZero/Basic.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Johan Commelin
55
66
! This file was ported from Lean 3 source module algebra.group_with_zero.basic
7-
! leanprover-community/mathlib commit 2f3994e1b117b1e1da49bcfb67334f33460c3ce4
7+
! leanprover-community/mathlib commit 2196ab363eb097c008d4497125e0dde23fb36db2
88
! Please do not edit these lines, except to modify the commit id
99
! if you have ported upstream changes.
1010
-/
@@ -371,7 +371,7 @@ theorem one_div_ne_zero {a : G₀} (h : a ≠ 0) : 1 / a ≠ 0 := by
371371
#align one_div_ne_zero one_div_ne_zero
372372

373373
@[simp]
374-
theorem inv_eq_zero {a : G₀} : a⁻¹ = 0 ↔ a = 0 := by rw [inv_eq_iff_inv_eq, inv_zero, eq_comm]
374+
theorem inv_eq_zero {a : G₀} : a⁻¹ = 0 ↔ a = 0 := by rw [inv_eq_iff_eq_inv, inv_zero]
375375
#align inv_eq_zero inv_eq_zero
376376

377377
@[simp]

Mathlib/Algebra/Order/Group/Abs.lean

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl
55
66
! This file was ported from Lean 3 source module algebra.order.group.abs
7-
! leanprover-community/mathlib commit a95b16cbade0f938fc24abd05412bde1e84bab9b
7+
! leanprover-community/mathlib commit 2196ab363eb097c008d4497125e0dde23fb36db2
88
! Please do not edit these lines, except to modify the commit id
99
! if you have ported upstream changes.
1010
-/
@@ -88,13 +88,13 @@ theorem abs_neg (a : α) : |(-a)| = |a| := by rw [abs_eq_max_neg, max_comm, neg_
8888
#align abs_neg abs_neg
8989

9090
theorem eq_or_eq_neg_of_abs_eq {a b : α} (h : |a| = b) : a = b ∨ a = -b := by
91-
simpa only [← h, eq_comm, eq_neg_iff_eq_neg] using abs_choice a
91+
simpa only [← h, eq_comm (a := |a|), neg_eq_iff_eq_neg] using abs_choice a
9292
#align eq_or_eq_neg_of_abs_eq eq_or_eq_neg_of_abs_eq
9393

9494
theorem abs_eq_abs {a b : α} : |a| = |b| ↔ a = b ∨ a = -b := by
9595
refine' ⟨fun h => _, fun h => _⟩
9696
· obtain rfl | rfl := eq_or_eq_neg_of_abs_eq h <;>
97-
simpa only [neg_eq_iff_neg_eq, neg_inj, or_comm, @eq_comm _ (-b)] using abs_choice b
97+
simpa only [neg_eq_iff_eq_neg (a := |b|), neg_inj, or_comm] using abs_choice b
9898
· cases' h with h h <;>
9999
simp [h, abs_neg]
100100
#align abs_eq_abs abs_eq_abs

Mathlib/Algebra/Periodic.lean

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Benjamin Davidson
55
66
! This file was ported from Lean 3 source module algebra.periodic
7-
! leanprover-community/mathlib commit 4c19a16e4b705bf135cf9a80ac18fcc99c438514
7+
! leanprover-community/mathlib commit 2196ab363eb097c008d4497125e0dde23fb36db2
88
! Please do not edit these lines, except to modify the commit id
99
! if you have ported upstream changes.
1010
-/
@@ -371,7 +371,7 @@ protected theorem Antiperiodic.funext [Add α] [Neg β] (h : Antiperiodic f c) :
371371

372372
protected theorem Antiperiodic.funext' [Add α] [InvolutiveNeg β] (h : Antiperiodic f c) :
373373
(fun x => -f (x + c)) = f :=
374-
(eq_neg_iff_eq_neg.mp h.funext).symm
374+
neg_eq_iff_eq_neg.mpr h.funext
375375
#align function.antiperiodic.funext' Function.Antiperiodic.funext'
376376

377377
/-- If a function is `antiperiodic` with antiperiod `c`, then it is also `periodic` with period
@@ -405,7 +405,7 @@ theorem Antiperiodic.int_odd_mul_antiperiodic [Ring α] [InvolutiveNeg β] (h :
405405
#align function.antiperiodic.int_odd_mul_antiperiodic Function.Antiperiodic.int_odd_mul_antiperiodic
406406

407407
theorem Antiperiodic.sub_eq [AddGroup α] [InvolutiveNeg β] (h : Antiperiodic f c) (x : α) :
408-
f (x - c) = -f x := by simp only [eq_neg_iff_eq_neg.mp (h (x - c)), sub_add_cancel]
408+
f (x - c) = -f x := by simp only [← neg_eq_iff_eq_neg, ← h (x - c), sub_add_cancel]
409409
#align function.antiperiodic.sub_eq Function.Antiperiodic.sub_eq
410410

411411
theorem Antiperiodic.sub_eq' [AddCommGroup α] [Neg β] (h : Antiperiodic f c) :

Mathlib/Data/Int/Basic.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Jeremy Avigad
55
66
! This file was ported from Lean 3 source module data.int.basic
7-
! leanprover-community/mathlib commit e1bccd6e40ae78370f01659715d3c948716e3b7e
7+
! leanprover-community/mathlib commit 2196ab363eb097c008d4497125e0dde23fb36db2
88
! Please do not edit these lines, except to modify the commit id
99
! if you have ported upstream changes.
1010
-/
@@ -153,7 +153,7 @@ theorem succ_neg_succ (a : ℤ) : succ (-succ a) = -a := by rw [neg_succ, succ_p
153153
#align int.succ_neg_succ Int.succ_neg_succ
154154

155155
theorem neg_pred (a : ℤ) : -pred a = succ (-a) := by
156-
rw [eq_neg_of_eq_neg (neg_succ (-a)).symm, neg_neg]
156+
rw [neg_eq_iff_eq_neg.mp (neg_succ (-a)), neg_neg]
157157
#align int.neg_pred Int.neg_pred
158158

159159
theorem pred_neg_pred (a : ℤ) : pred (-pred a) = -a := by rw [neg_pred, pred_succ]

Mathlib/Data/Real/ConjugateExponents.lean

Lines changed: 3 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Sébastien Gouëzel, Yury Kudryashov
55
66
! This file was ported from Lean 3 source module data.real.conjugate_exponents
7-
! leanprover-community/mathlib commit bd9851ca476957ea4549eb19b40e7b5ade9428cc
7+
! leanprover-community/mathlib commit 2196ab363eb097c008d4497125e0dde23fb36db2
88
! Please do not edit these lines, except to modify the commit id
99
! if you have ported upstream changes.
1010
-/
@@ -73,8 +73,8 @@ theorem one_div_ne_zero : 1 / p ≠ 0 := ne_of_gt h.one_div_pos
7373

7474
theorem conj_eq : q = p / (p - 1) := by
7575
have := h.inv_add_inv_conj
76-
rw [← eq_sub_iff_add_eq', one_div, inv_eq_iff_inv_eq] at this
77-
field_simp [this, h.ne_zero]
76+
rw [← eq_sub_iff_add_eq', one_div, inv_eq_iff_eq_inv] at this
77+
field_simp [this, h.ne_zero]
7878
#align real.is_conjugate_exponent.conj_eq Real.IsConjugateExponent.conj_eq
7979

8080
theorem conjugate_eq : conjugateExponent p = q := h.conj_eq.symm
@@ -134,4 +134,3 @@ theorem isConjugateExponent_one_div {a b : ℝ} (ha : 0 < a) (hb : 0 < b) (hab :
134134
#align real.is_conjugate_exponent_one_div Real.isConjugateExponent_one_div
135135

136136
end Real
137-

Mathlib/Data/Real/EReal.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Kevin Buzzard
55
66
! This file was ported from Lean 3 source module data.real.ereal
7-
! leanprover-community/mathlib commit afdb4fa3b32d41106a4a09b371ce549ad7958abd
7+
! leanprover-community/mathlib commit 2196ab363eb097c008d4497125e0dde23fb36db2
88
! Please do not edit these lines, except to modify the commit id
99
! if you have ported upstream changes.
1010
-/

Mathlib/Data/Set/Pointwise/Interval.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Yury G. Kudryashov, Patrick Massot
55
66
! This file was ported from Lean 3 source module data.set.pointwise.interval
7-
! leanprover-community/mathlib commit 59694bd07f0a39c5beccba34bd9f413a160782bf
7+
! leanprover-community/mathlib commit 2196ab363eb097c008d4497125e0dde23fb36db2
88
! Please do not edit these lines, except to modify the commit id
99
! if you have ported upstream changes.
1010
-/
@@ -770,7 +770,7 @@ theorem inv_Ioo_0_left {a : α} (ha : 0 < a) : (Ioo 0 a)⁻¹ = Ioi a⁻¹ := by
770770
#align set.inv_Ioo_0_left Set.inv_Ioo_0_left
771771

772772
theorem inv_Ioi {a : α} (ha : 0 < a) : (Ioi a)⁻¹ = Ioo 0 a⁻¹ := by
773-
rw [inv_eq_iff_inv_eq, inv_Ioo_0_left (inv_pos.2 ha), inv_inv]
773+
rw [inv_eq_iff_eq_inv, inv_Ioo_0_left (inv_pos.2 ha), inv_inv]
774774
#align set.inv_Ioi Set.inv_Ioi
775775

776776
theorem image_const_mul_Ioi_zero {k : Type _} [LinearOrderedField k] {x : k} (hx : 0 < x) :

0 commit comments

Comments
 (0)