Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[Merged by Bors] - feat(GroupTheory/GroupAction/Group): invOf smul lemmas #9972

Closed
wants to merge 10 commits into from
22 changes: 22 additions & 0 deletions Mathlib/GroupTheory/GroupAction/Group.lean
Original file line number Diff line number Diff line change
Expand Up @@ -4,6 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Chris Hughes
-/
import Mathlib.Algebra.Group.Aut
import Mathlib.Algebra.Invertible.Defs
import Mathlib.Algebra.Invertible.Basic
mans0954 marked this conversation as resolved.
Show resolved Hide resolved
import Mathlib.GroupTheory.GroupAction.Units

#align_import group_theory.group_action.group from "leanprover-community/mathlib"@"3b52265189f3fb43aa631edffce5d060fafaf82f"
Expand Down Expand Up @@ -178,6 +180,26 @@ theorem smul_eq_iff_eq_inv_smul (g : α) {x y : β} : g • x = y ↔ x = g⁻¹

end Group

section Monoid

variable [Monoid α] [MulAction α β]

variable {c : α} {x y : β} [Invertible c]

@[simp]
theorem invOf_smul_smul : ⅟c • c • x = x := inv_smul_smul (unitOfInvertible c) _

@[simp]
theorem smul_invOf_smul : c • (⅟ c • x) = x := smul_inv_smul (unitOfInvertible c) _
Copy link
Member

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

c and x should be explicit for these lemmas.

Copy link
Collaborator Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Okay, done. Please can you explain why though so I know for next time.

Copy link
Member

@eric-wieser eric-wieser Jan 26, 2024

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

The simple answer here is that this lemma is almost identical to smul_inv_smul, and so you should try to match the argument explicitness too.

The slightly more general answer is that lemmas with = in their goal should ensure that every variable in their goal is captured in an an explicit variable (either as (x : X) or {x : X} (hx : P x)).


theorem invOf_smul_eq_iff : ⅟c • x = y ↔ x = c • y :=
inv_smul_eq_iff (a := unitOfInvertible c)

theorem smul_eq_iff_eq_invOf_smul : c • x = y ↔ x = ⅟c • y :=
smul_eq_iff_eq_inv_smul (g := unitOfInvertible c)

end Monoid

/-- `Monoid.toMulAction` is faithful on nontrivial cancellative monoids with zero. -/
instance CancelMonoidWithZero.faithfulSMul [CancelMonoidWithZero α] [Nontrivial α] :
FaithfulSMul α α :=
Expand Down
Loading