Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 9c3cb72

Browse files
committed
feat(data/int/basic): Add unit lemmas (#13565)
This PR adds a few more unit lemmas, and cleans up some of the proofs.
1 parent 695e0b7 commit 9c3cb72

File tree

2 files changed

+12
-6
lines changed

2 files changed

+12
-6
lines changed

src/algebra/group_power/lemmas.lean

Lines changed: 0 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -520,9 +520,6 @@ theorem nat.cast_le_pow_div_sub {K : Type*} [linear_ordered_field K] {a : K} (H
520520

521521
namespace int
522522

523-
lemma units_sq (u : ℤˣ) : u ^ 2 = 1 :=
524-
(sq u).symm ▸ units_mul_self u
525-
526523
alias int.units_sq ← int.units_pow_two
527524

528525
lemma units_pow_eq_pow_mod_two (u : ℤˣ) (n : ℕ) : u ^ n = u ^ (n % 2) :=

src/data/int/basic.lean

Lines changed: 12 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -1317,11 +1317,20 @@ by rw [is_unit_iff_nat_abs_eq, abs_eq_nat_abs, ←int.coe_nat_one, coe_nat_inj']
13171317
lemma of_nat_is_unit {n : ℕ} : is_unit (n : ℤ) ↔ is_unit n :=
13181318
by rw [nat.is_unit_iff, is_unit_iff_nat_abs_eq, nat_abs_of_nat]
13191319

1320-
lemma units_inv_eq_self (u : ℤˣ) : u⁻¹ = u :=
1321-
(units_eq_one_or u).elim (λ h, h.symm ▸ rfl) (λ h, h.symm ▸ rfl)
1320+
lemma is_unit_mul_self {a : ℤ} (ha : is_unit a) : a * a = 1 :=
1321+
(is_unit_eq_one_or ha).elim (λ h, h.symm ▸ rfl) (λ h, h.symm ▸ rfl)
1322+
1323+
lemma is_unit_sq {a : ℤ} (ha : is_unit a) : a ^ 2 = 1 :=
1324+
by rw [sq, is_unit_mul_self ha]
1325+
1326+
@[simp] lemma units_sq (u : ℤˣ) : u ^ 2 = 1 :=
1327+
by rw [units.ext_iff, units.coe_pow, units.coe_one, is_unit_sq u.is_unit]
13221328

13231329
@[simp] lemma units_mul_self (u : ℤˣ) : u * u = 1 :=
1324-
(units_eq_one_or u).elim (λ h, h.symm ▸ rfl) (λ h, h.symm ▸ rfl)
1330+
by rw [←sq, units_sq]
1331+
1332+
@[simp] lemma units_inv_eq_self (u : ℤˣ) : u⁻¹ = u :=
1333+
by rw [inv_eq_iff_mul_eq_one, units_mul_self]
13251334

13261335
-- `units.coe_mul` is a "wrong turn" for the simplifier, this undoes it and simplifies further
13271336
@[simp] lemma units_coe_mul_self (u : ℤˣ) : (u * u : ℤ) = 1 :=

0 commit comments

Comments
 (0)