Skip to content

Commit ee38ee9

Browse files
committed
feat(Algebra/GroupWithZero/Units/Lemmas): div_div_div_cancel_left' (#6606)
Co-authored-by: Moritz Firsching <firsching@google.com>
1 parent 38c0722 commit ee38ee9

File tree

2 files changed

+5
-4
lines changed

2 files changed

+5
-4
lines changed

Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean

Lines changed: 3 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -199,6 +199,9 @@ theorem div_helper (b : G₀) (h : a ≠ 0) : 1 / (a * b) * a = 1 / b := by
199199
rw [div_mul_eq_mul_div, one_mul, div_mul_right _ h]
200200
#align div_helper div_helper
201201

202+
theorem div_div_div_cancel_left' (a b : G₀) (hc : c ≠ 0) : c / a / (c / b) = b / a := by
203+
rw [div_div_div_eq, mul_comm, mul_div_mul_right _ _ hc]
204+
202205
end CommGroupWithZero
203206

204207
section MonoidWithZero

Mathlib/Analysis/SpecialFunctions/Log/Base.lean

Lines changed: 2 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -102,10 +102,8 @@ theorem mul_logb {a b c : ℝ} (h₁ : b ≠ 0) (h₂ : b ≠ 1) (h₃ : b ≠ -
102102
#align real.mul_logb Real.mul_logb
103103

104104
theorem div_logb {a b c : ℝ} (h₁ : c ≠ 0) (h₂ : c ≠ 1) (h₃ : c ≠ -1) :
105-
logb a c / logb b c = logb a b := by
106-
unfold logb
107-
-- TODO: div_div_div_cancel_left is missing for `group_with_zero`,
108-
rw [div_div_div_eq, mul_comm, mul_div_mul_right _ _ (log_ne_zero.mpr ⟨h₁, h₂, h₃⟩)]
105+
logb a c / logb b c = logb a b :=
106+
div_div_div_cancel_left' _ _ <| log_ne_zero.mpr ⟨h₁, h₂, h₃⟩
109107
#align real.div_logb Real.div_logb
110108

111109
section BPosAndNeOne

0 commit comments

Comments
 (0)