From ad22323241bc9d60798ca7004f189caf7a924890 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 25 Jan 2024 00:17:40 +0000 Subject: [PATCH] feat: two lemmas about division (#9966) --- Mathlib/Algebra/Group/Basic.lean | 6 +++++- Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean | 7 +++++++ 2 files changed, 12 insertions(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Group/Basic.lean b/Mathlib/Algebra/Group/Basic.lean index 9cc7b2c117dbc..3aa7be269ac5b 100644 --- a/Mathlib/Algebra/Group/Basic.lean +++ b/Mathlib/Algebra/Group/Basic.lean @@ -410,7 +410,7 @@ end DivInvOneMonoid section DivisionMonoid -variable [DivisionMonoid α] {a b c : α} +variable [DivisionMonoid α] {a b c d : α} attribute [local simp] mul_assoc div_eq_mul_inv @@ -471,6 +471,10 @@ theorem one_div_one_div : 1 / (1 / a) = a := by simp #align one_div_one_div one_div_one_div #align zero_sub_zero_sub zero_sub_zero_sub +@[to_additive] +theorem div_eq_div_iff_comm : a / b = c / d ↔ b / a = d / c := + inv_inj.symm.trans <| by simp only [inv_div] + @[to_additive SubtractionMonoid.toSubNegZeroMonoid] instance (priority := 100) DivisionMonoid.toDivInvOneMonoid : DivInvOneMonoid α := { DivisionMonoid.toDivInvMonoid with diff --git a/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean b/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean index e4b934971710d..78e9e0bd0ceba 100644 --- a/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean +++ b/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean @@ -150,6 +150,7 @@ theorem divp_mk0 (a : G₀) {b : G₀} (hb : b ≠ 0) : a /ₚ Units.mk0 b hb = namespace Commute +/-- The `MonoidWithZero` version of `div_eq_div_iff_mul_eq_mul`. -/ protected lemma div_eq_div_iff (hbd : Commute b d) (hb : b ≠ 0) (hd : d ≠ 0) : a / b = c / d ↔ a * d = c * b := hbd.div_eq_div_iff_of_isUnit hb.isUnit hd.isUnit @@ -195,6 +196,12 @@ theorem div_eq_div_iff (hb : b ≠ 0) (hd : d ≠ 0) : a / b = c / d ↔ a * d = IsUnit.div_eq_div_iff hb.isUnit hd.isUnit #align div_eq_div_iff div_eq_div_iff +/-- The `CommGroupWithZero` version of `div_eq_div_iff_div_eq_div`. -/ +theorem div_eq_div_iff_div_eq_div' (hb : b ≠ 0) (hc : c ≠ 0) : a / b = c / d ↔ a / c = b / d := by + conv_lhs => rw [← mul_left_inj' hb, div_mul_cancel _ hb] + conv_rhs => rw [← mul_left_inj' hc, div_mul_cancel _ hc] + rw [mul_comm _ c, div_mul_eq_mul_div, mul_div_assoc] + theorem div_div_cancel' (ha : a ≠ 0) : a / (a / b) = b := IsUnit.div_div_cancel ha.isUnit #align div_div_cancel' div_div_cancel'