From c529711fd42c23d223715b8150ac81936f834463 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Wed, 27 Oct 2021 17:47:27 +0000 Subject: [PATCH] refactor(*): rename fpow and gpow to zpow (#9989) Historically, we had two notions of integer power, one on groups called `gpow` and the other one on fields called `fpow`. Since the introduction of `div_inv_monoid` and `group_with_zero`, these definitions have been merged into one, and the boundaries are not clear any more. We rename both of them to `zpow`, adding a subscript `0` to disambiguate lemma names when there is a collision, where the subscripted version is for groups with zeroes (as is already done for inv lemmas). To limit the scope of the change. this PR does not rename `gsmul` to `zsmul` or `gmultiples` to `zmultiples`. --- src/algebra/field_power.lean | 178 ++++++++--------- src/algebra/group/conj.lean | 4 +- src/algebra/group/defs.lean | 48 ++--- src/algebra/group/hom_instances.lean | 10 +- src/algebra/group/pi.lean | 6 +- src/algebra/group/prod.lean | 8 +- src/algebra/group/to_additive.lean | 2 +- src/algebra/group/type_tags.lean | 16 +- src/algebra/group/ulift.lean | 4 +- src/algebra/group_power/basic.lean | 94 ++++----- src/algebra/group_power/lemmas.lean | 138 ++++++------- src/algebra/group_power/order.lean | 4 +- src/algebra/group_with_zero/power.lean | 174 ++++++++-------- src/algebra/iterate_hom.lean | 6 +- src/algebra/opposites.lean | 12 +- src/algebra/order/archimedean.lean | 12 +- src/algebra/punit_instances.lean | 2 +- src/algebra/star/basic.lean | 8 +- .../asymptotics/specific_asymptotics.lean | 12 +- .../asymptotics/superpolynomial_decay.lean | 20 +- src/analysis/calculus/deriv.lean | 68 +++---- src/analysis/convex/specific_functions.lean | 12 +- src/analysis/fourier.lean | 6 +- src/analysis/mean_inequalities.lean | 6 +- src/analysis/normed_space/basic.lean | 24 +-- src/analysis/special_functions/exp.lean | 2 +- .../special_functions/polynomials.lean | 10 +- src/analysis/special_functions/pow.lean | 8 +- src/analysis/specific_limits.lean | 12 +- src/category_theory/conj.lean | 4 +- src/category_theory/endomorphism.lean | 2 +- src/data/complex/basic.lean | 12 +- src/data/complex/is_R_or_C.lean | 4 +- src/data/equiv/mul_add_aut.lean | 4 +- src/data/equiv/ring_aut.lean | 2 +- src/data/int/gcd.lean | 4 +- src/data/real/irrational.lean | 4 +- src/deprecated/subfield.lean | 4 +- .../rotation_number/translation_number.lean | 12 +- src/field_theory/finite/basic.lean | 4 +- src/field_theory/intermediate_field.lean | 4 +- src/group_theory/free_group.lean | 6 +- src/group_theory/group_action/group.lean | 2 +- src/group_theory/order_of_element.lean | 136 ++++++------- src/group_theory/perm/basic.lean | 4 +- src/group_theory/perm/concrete_cycle.lean | 2 +- src/group_theory/perm/cycles.lean | 164 +++++++-------- src/group_theory/perm/fin.lean | 2 +- src/group_theory/perm/list.lean | 6 +- src/group_theory/perm/support.lean | 30 +-- src/group_theory/quotient_group.lean | 4 +- src/group_theory/specific_groups/cyclic.lean | 88 ++++----- src/group_theory/subgroup/basic.lean | 98 ++++----- src/group_theory/sylow.lean | 12 +- .../matrix/{fpow.lean => zpow.lean} | 186 +++++++++--------- src/measure_theory/group/arithmetic.lean | 10 +- src/number_theory/arithmetic_function.lean | 2 +- src/number_theory/padics/padic_integers.lean | 24 +-- src/number_theory/padics/padic_norm.lean | 18 +- src/number_theory/padics/padic_numbers.lean | 20 +- src/number_theory/padics/ring_homs.lean | 4 +- src/number_theory/quadratic_reciprocity.lean | 4 +- src/ring_theory/roots_of_unity.lean | 118 +++++------ src/tactic/group.lean | 46 ++--- src/topology/algebra/group_with_zero.lean | 34 ++-- src/topology/algebra/ordered/basic.lean | 20 +- test/refine_struct.lean | 28 +-- 67 files changed, 1017 insertions(+), 1017 deletions(-) rename src/linear_algebra/matrix/{fpow.lean => zpow.lean} (56%) diff --git a/src/algebra/field_power.lean b/src/algebra/field_power.lean index cb94312989cac..bb35ad0880d84 100644 --- a/src/algebra/field_power.lean +++ b/src/algebra/field_power.lean @@ -16,62 +16,62 @@ integer power. More specialised results are provided in the case of a linearly o universe u -@[simp] lemma ring_hom.map_fpow {K L : Type*} [division_ring K] [division_ring L] (f : K →+* L) : +@[simp] lemma ring_hom.map_zpow {K L : Type*} [division_ring K] [division_ring L] (f : K →+* L) : ∀ (a : K) (n : ℤ), f (a ^ n) = f a ^ n := -f.to_monoid_with_zero_hom.map_fpow +f.to_monoid_with_zero_hom.map_zpow -@[simp] lemma ring_equiv.map_fpow {K L : Type*} [division_ring K] [division_ring L] (f : K ≃+* L) : +@[simp] lemma ring_equiv.map_zpow {K L : Type*} [division_ring K] [division_ring L] (f : K ≃+* L) : ∀ (a : K) (n : ℤ), f (a ^ n) = f a ^ n := -f.to_ring_hom.map_fpow +f.to_ring_hom.map_zpow -@[simp] lemma fpow_bit0_neg {K : Type*} [division_ring K] (x : K) (n : ℤ) : +@[simp] lemma zpow_bit0_neg {K : Type*} [division_ring K] (x : K) (n : ℤ) : (-x) ^ (bit0 n) = x ^ bit0 n := -by rw [fpow_bit0', fpow_bit0', neg_mul_neg] +by rw [zpow_bit0', zpow_bit0', neg_mul_neg] -lemma even.fpow_neg {K : Type*} [division_ring K] {n : ℤ} (h : even n) (a : K) : +lemma even.zpow_neg {K : Type*} [division_ring K] {n : ℤ} (h : even n) (a : K) : (-a) ^ n = a ^ n := begin obtain ⟨k, rfl⟩ := h, - rw [←bit0_eq_two_mul, fpow_bit0_neg], + rw [←bit0_eq_two_mul, zpow_bit0_neg], end -@[simp] lemma fpow_bit1_neg {K : Type*} [division_ring K] (x : K) (n : ℤ) : +@[simp] lemma zpow_bit1_neg {K : Type*} [division_ring K] (x : K) (n : ℤ) : (-x) ^ (bit1 n) = - x ^ bit1 n := -by rw [fpow_bit1', fpow_bit1', neg_mul_neg, neg_mul_eq_mul_neg] +by rw [zpow_bit1', zpow_bit1', neg_mul_neg, neg_mul_eq_mul_neg] section ordered_field_power open int variables {K : Type u} [linear_ordered_field K] {a : K} {n : ℤ} -lemma fpow_eq_zero_iff (hn : 0 < n) : +lemma zpow_eq_zero_iff (hn : 0 < n) : a ^ n = 0 ↔ a = 0 := begin - refine ⟨fpow_eq_zero, _⟩, + refine ⟨zpow_eq_zero, _⟩, rintros rfl, - exact zero_fpow _ hn.ne' + exact zero_zpow _ hn.ne' end -lemma fpow_nonneg {a : K} (ha : 0 ≤ a) : ∀ (z : ℤ), 0 ≤ a ^ z -| (n : ℕ) := by { rw gpow_coe_nat, exact pow_nonneg ha _ } -| -[1+n] := by { rw gpow_neg_succ_of_nat, exact inv_nonneg.2 (pow_nonneg ha _) } +lemma zpow_nonneg {a : K} (ha : 0 ≤ a) : ∀ (z : ℤ), 0 ≤ a ^ z +| (n : ℕ) := by { rw zpow_coe_nat, exact pow_nonneg ha _ } +| -[1+n] := by { rw zpow_neg_succ_of_nat, exact inv_nonneg.2 (pow_nonneg ha _) } -lemma fpow_pos_of_pos {a : K} (ha : 0 < a) : ∀ (z : ℤ), 0 < a ^ z -| (n : ℕ) := by { rw gpow_coe_nat, exact pow_pos ha _ } -| -[1+n] := by { rw gpow_neg_succ_of_nat, exact inv_pos.2 (pow_pos ha _) } +lemma zpow_pos_of_pos {a : K} (ha : 0 < a) : ∀ (z : ℤ), 0 < a ^ z +| (n : ℕ) := by { rw zpow_coe_nat, exact pow_pos ha _ } +| -[1+n] := by { rw zpow_neg_succ_of_nat, exact inv_pos.2 (pow_pos ha _) } -lemma fpow_le_of_le {x : K} (hx : 1 ≤ x) {a b : ℤ} (h : a ≤ b) : x ^ a ≤ x ^ b := +lemma zpow_le_of_le {x : K} (hx : 1 ≤ x) {a b : ℤ} (h : a ≤ b) : x ^ a ≤ x ^ b := begin induction a with a a; induction b with b b, - { simp only [of_nat_eq_coe, gpow_coe_nat], + { simp only [of_nat_eq_coe, zpow_coe_nat], apply pow_le_pow hx, apply le_of_coe_nat_le_coe_nat h }, { apply absurd h, apply not_le_of_gt, exact lt_of_lt_of_le (neg_succ_lt_zero _) (of_nat_nonneg _) }, - { simp only [gpow_neg_succ_of_nat, one_div, of_nat_eq_coe, gpow_coe_nat], + { simp only [zpow_neg_succ_of_nat, one_div, of_nat_eq_coe, zpow_coe_nat], apply le_trans (inv_le_one _); apply one_le_pow_of_one_le hx }, - { simp only [gpow_neg_succ_of_nat], + { simp only [zpow_neg_succ_of_nat], apply (inv_le_inv _ _).2, { apply pow_le_pow hx, have : -(↑(a+1) : ℤ) ≤ -(↑(b+1) : ℤ), from h, @@ -85,133 +85,133 @@ lemma pow_le_max_of_min_le {x : K} (hx : 1 ≤ x) {a b c : ℤ} (h : min a b ≤ begin wlog hle : a ≤ b, have hnle : -b ≤ -a, from neg_le_neg hle, - have hfle : x ^ (-b) ≤ x ^ (-a), from fpow_le_of_le hx hnle, + have hfle : x ^ (-b) ≤ x ^ (-a), from zpow_le_of_le hx hnle, have : x ^ (-c) ≤ x ^ (-a), - { apply fpow_le_of_le hx, + { apply zpow_le_of_le hx, simpa only [min_eq_left hle, neg_le_neg_iff] using h }, simpa only [max_eq_left hfle] end -lemma fpow_le_one_of_nonpos {p : K} (hp : 1 ≤ p) {z : ℤ} (hz : z ≤ 0) : p ^ z ≤ 1 := -calc p ^ z ≤ p ^ 0 : fpow_le_of_le hp hz +lemma zpow_le_one_of_nonpos {p : K} (hp : 1 ≤ p) {z : ℤ} (hz : z ≤ 0) : p ^ z ≤ 1 := +calc p ^ z ≤ p ^ 0 : zpow_le_of_le hp hz ... = 1 : by simp -lemma one_le_fpow_of_nonneg {p : K} (hp : 1 ≤ p) {z : ℤ} (hz : 0 ≤ z) : 1 ≤ p ^ z := -calc p ^ z ≥ p ^ 0 : fpow_le_of_le hp hz +lemma one_le_zpow_of_nonneg {p : K} (hp : 1 ≤ p) {z : ℤ} (hz : 0 ≤ z) : 1 ≤ p ^ z := +calc p ^ z ≥ p ^ 0 : zpow_le_of_le hp hz ... = 1 : by simp -theorem fpow_bit0_nonneg (a : K) (n : ℤ) : 0 ≤ a ^ bit0 n := -by { rw fpow_bit0, exact mul_self_nonneg _ } +theorem zpow_bit0_nonneg (a : K) (n : ℤ) : 0 ≤ a ^ bit0 n := +by { rw zpow_bit0₀, exact mul_self_nonneg _ } -theorem fpow_two_nonneg (a : K) : 0 ≤ a ^ 2 := +theorem zpow_two_nonneg (a : K) : 0 ≤ a ^ 2 := pow_bit0_nonneg a 1 -theorem fpow_bit0_pos {a : K} (h : a ≠ 0) (n : ℤ) : 0 < a ^ bit0 n := -(fpow_bit0_nonneg a n).lt_of_ne (fpow_ne_zero _ h).symm +theorem zpow_bit0_pos {a : K} (h : a ≠ 0) (n : ℤ) : 0 < a ^ bit0 n := +(zpow_bit0_nonneg a n).lt_of_ne (zpow_ne_zero _ h).symm -theorem fpow_two_pos_of_ne_zero (a : K) (h : a ≠ 0) : 0 < a ^ 2 := +theorem zpow_two_pos_of_ne_zero (a : K) (h : a ≠ 0) : 0 < a ^ 2 := pow_bit0_pos h 1 -@[simp] theorem fpow_bit1_neg_iff : a ^ bit1 n < 0 ↔ a < 0 := -⟨λ h, not_le.1 $ λ h', not_le.2 h $ fpow_nonneg h' _, - λ h, by rw [bit1, fpow_add_one h.ne]; exact mul_neg_of_pos_of_neg (fpow_bit0_pos h.ne _) h⟩ +@[simp] theorem zpow_bit1_neg_iff : a ^ bit1 n < 0 ↔ a < 0 := +⟨λ h, not_le.1 $ λ h', not_le.2 h $ zpow_nonneg h' _, + λ h, by rw [bit1, zpow_add_one₀ h.ne]; exact mul_neg_of_pos_of_neg (zpow_bit0_pos h.ne _) h⟩ -@[simp] theorem fpow_bit1_nonneg_iff : 0 ≤ a ^ bit1 n ↔ 0 ≤ a := -le_iff_le_iff_lt_iff_lt.2 fpow_bit1_neg_iff +@[simp] theorem zpow_bit1_nonneg_iff : 0 ≤ a ^ bit1 n ↔ 0 ≤ a := +le_iff_le_iff_lt_iff_lt.2 zpow_bit1_neg_iff -@[simp] theorem fpow_bit1_nonpos_iff : a ^ bit1 n ≤ 0 ↔ a ≤ 0 := +@[simp] theorem zpow_bit1_nonpos_iff : a ^ bit1 n ≤ 0 ↔ a ≤ 0 := begin - rw [le_iff_lt_or_eq, fpow_bit1_neg_iff], + rw [le_iff_lt_or_eq, zpow_bit1_neg_iff], split, { rintro (h | h), { exact h.le }, - { exact (fpow_eq_zero h).le } }, + { exact (zpow_eq_zero h).le } }, { intro h, rcases eq_or_lt_of_le h with rfl|h, - { exact or.inr (zero_fpow _ (bit1_ne_zero n)) }, + { exact or.inr (zero_zpow _ (bit1_ne_zero n)) }, { exact or.inl h } } end -@[simp] theorem fpow_bit1_pos_iff : 0 < a ^ bit1 n ↔ 0 < a := -lt_iff_lt_of_le_iff_le fpow_bit1_nonpos_iff +@[simp] theorem zpow_bit1_pos_iff : 0 < a ^ bit1 n ↔ 0 < a := +lt_iff_lt_of_le_iff_le zpow_bit1_nonpos_iff -lemma even.fpow_nonneg {n : ℤ} (hn : even n) (a : K) : +lemma even.zpow_nonneg {n : ℤ} (hn : even n) (a : K) : 0 ≤ a ^ n := begin cases le_or_lt 0 a with h h, - { exact fpow_nonneg h _ }, - { exact (hn.fpow_neg a).subst (fpow_nonneg (neg_nonneg_of_nonpos h.le) _) } + { exact zpow_nonneg h _ }, + { exact (hn.zpow_neg a).subst (zpow_nonneg (neg_nonneg_of_nonpos h.le) _) } end -theorem even.fpow_pos (hn : even n) (ha : a ≠ 0) : 0 < a ^ n := -by cases hn with k hk; simpa only [hk, two_mul] using fpow_bit0_pos ha k +theorem even.zpow_pos (hn : even n) (ha : a ≠ 0) : 0 < a ^ n := +by cases hn with k hk; simpa only [hk, two_mul] using zpow_bit0_pos ha k -theorem odd.fpow_nonneg (hn : odd n) (ha : 0 ≤ a) : 0 ≤ a ^ n := -by cases hn with k hk; simpa only [hk, two_mul] using fpow_bit1_nonneg_iff.mpr ha +theorem odd.zpow_nonneg (hn : odd n) (ha : 0 ≤ a) : 0 ≤ a ^ n := +by cases hn with k hk; simpa only [hk, two_mul] using zpow_bit1_nonneg_iff.mpr ha -theorem odd.fpow_pos (hn : odd n) (ha : 0 < a) : 0 < a ^ n := -by cases hn with k hk; simpa only [hk, two_mul] using fpow_bit1_pos_iff.mpr ha +theorem odd.zpow_pos (hn : odd n) (ha : 0 < a) : 0 < a ^ n := +by cases hn with k hk; simpa only [hk, two_mul] using zpow_bit1_pos_iff.mpr ha -theorem odd.fpow_nonpos (hn : odd n) (ha : a ≤ 0) : a ^ n ≤ 0:= -by cases hn with k hk; simpa only [hk, two_mul] using fpow_bit1_nonpos_iff.mpr ha +theorem odd.zpow_nonpos (hn : odd n) (ha : a ≤ 0) : a ^ n ≤ 0:= +by cases hn with k hk; simpa only [hk, two_mul] using zpow_bit1_nonpos_iff.mpr ha -theorem odd.fpow_neg (hn : odd n) (ha : a < 0) : a ^ n < 0:= -by cases hn with k hk; simpa only [hk, two_mul] using fpow_bit1_neg_iff.mpr ha +theorem odd.zpow_neg (hn : odd n) (ha : a < 0) : a ^ n < 0:= +by cases hn with k hk; simpa only [hk, two_mul] using zpow_bit1_neg_iff.mpr ha -lemma even.fpow_abs {p : ℤ} (hp : even p) (a : K) : |a| ^ p = a ^ p := +lemma even.zpow_abs {p : ℤ} (hp : even p) (a : K) : |a| ^ p = a ^ p := begin cases abs_choice a with h h; - simp only [h, hp.fpow_neg _], + simp only [h, hp.zpow_neg _], end -@[simp] lemma fpow_bit0_abs (a : K) (p : ℤ) : |a| ^ bit0 p = a ^ bit0 p := -(even_bit0 _).fpow_abs _ +@[simp] lemma zpow_bit0_abs (a : K) (p : ℤ) : |a| ^ bit0 p = a ^ bit0 p := +(even_bit0 _).zpow_abs _ -lemma even.abs_fpow {p : ℤ} (hp : even p) (a : K) : |a ^ p| = a ^ p := +lemma even.abs_zpow {p : ℤ} (hp : even p) (a : K) : |a ^ p| = a ^ p := begin rw [abs_eq_self], - exact hp.fpow_nonneg _ + exact hp.zpow_nonneg _ end -@[simp] lemma abs_fpow_bit0 (a : K) (p : ℤ) : +@[simp] lemma abs_zpow_bit0 (a : K) (p : ℤ) : |a ^ bit0 p| = a ^ bit0 p := -(even_bit0 _).abs_fpow _ +(even_bit0 _).abs_zpow _ end ordered_field_power -lemma one_lt_fpow {K} [linear_ordered_field K] {p : K} (hp : 1 < p) : +lemma one_lt_zpow {K} [linear_ordered_field K] {p : K} (hp : 1 < p) : ∀ z : ℤ, 0 < z → 1 < p ^ z -| (n : ℕ) h := (gpow_coe_nat p n).symm.subst (one_lt_pow hp $ int.coe_nat_ne_zero.mp h.ne') +| (n : ℕ) h := (zpow_coe_nat p n).symm.subst (one_lt_pow hp $ int.coe_nat_ne_zero.mp h.ne') | -[1+ n] h := ((int.neg_succ_not_pos _).mp h).elim section ordered variables {K : Type*} [linear_ordered_field K] -lemma nat.fpow_pos_of_pos {p : ℕ} (h : 0 < p) (n:ℤ) : 0 < (p:K)^n := -by { apply fpow_pos_of_pos, exact_mod_cast h } +lemma nat.zpow_pos_of_pos {p : ℕ} (h : 0 < p) (n:ℤ) : 0 < (p:K)^n := +by { apply zpow_pos_of_pos, exact_mod_cast h } -lemma nat.fpow_ne_zero_of_pos {p : ℕ} (h : 0 < p) (n:ℤ) : (p:K)^n ≠ 0 := -ne_of_gt (nat.fpow_pos_of_pos h n) +lemma nat.zpow_ne_zero_of_pos {p : ℕ} (h : 0 < p) (n:ℤ) : (p:K)^n ≠ 0 := +ne_of_gt (nat.zpow_pos_of_pos h n) -lemma fpow_strict_mono {x : K} (hx : 1 < x) : +lemma zpow_strict_mono {x : K} (hx : 1 < x) : strict_mono (λ n:ℤ, x ^ n) := λ m n h, show x ^ m < x ^ n, begin have xpos : 0 < x := zero_lt_one.trans hx, have h₀ : x ≠ 0 := xpos.ne', - have hxm : 0 < x^m := fpow_pos_of_pos xpos m, - have h : 1 < x ^ (n - m) := one_lt_fpow hx _ (sub_pos_of_lt h), + have hxm : 0 < x^m := zpow_pos_of_pos xpos m, + have h : 1 < x ^ (n - m) := one_lt_zpow hx _ (sub_pos_of_lt h), replace h := mul_lt_mul_of_pos_right h hxm, - rwa [sub_eq_add_neg, fpow_add h₀, mul_assoc, fpow_neg_mul_fpow_self _ h₀, one_mul, mul_one] at h, + rwa [sub_eq_add_neg, zpow_add₀ h₀, mul_assoc, zpow_neg_mul_zpow_self _ h₀, one_mul, mul_one] at h, end -@[simp] lemma fpow_lt_iff_lt {x : K} (hx : 1 < x) {m n : ℤ} : +@[simp] lemma zpow_lt_iff_lt {x : K} (hx : 1 < x) {m n : ℤ} : x ^ m < x ^ n ↔ m < n := -(fpow_strict_mono hx).lt_iff_lt +(zpow_strict_mono hx).lt_iff_lt -@[simp] lemma fpow_le_iff_le {x : K} (hx : 1 < x) {m n : ℤ} : +@[simp] lemma zpow_le_iff_le {x : K} (hx : 1 < x) {m n : ℤ} : x ^ m ≤ x ^ n ↔ m ≤ n := -(fpow_strict_mono hx).le_iff_le +(zpow_strict_mono hx).le_iff_le @[simp] lemma pos_div_pow_pos {a b : K} (ha : 0 < a) (hb : 0 < b) (k : ℕ) : 0 < a/b^k := div_pos ha (pow_pos hb k) @@ -221,29 +221,29 @@ div_pos ha (pow_pos hb k) (calc a = a * 1 : (mul_one a).symm ... ≤ a*b^k : (mul_le_mul_left ha).mpr $ one_le_pow_of_one_le hb _) -lemma fpow_injective {x : K} (h₀ : 0 < x) (h₁ : x ≠ 1) : +lemma zpow_injective {x : K} (h₀ : 0 < x) (h₁ : x ≠ 1) : function.injective ((^) x : ℤ → K) := begin intros m n h, rcases h₁.lt_or_lt with H|H, - { apply (fpow_strict_mono (one_lt_inv h₀ H)).injective, + { apply (zpow_strict_mono (one_lt_inv h₀ H)).injective, show x⁻¹ ^ m = x⁻¹ ^ n, - rw [← fpow_neg_one, ← fpow_mul, ← fpow_mul, mul_comm _ m, mul_comm _ n, fpow_mul, fpow_mul, + rw [← zpow_neg_one₀, ← zpow_mul₀, ← zpow_mul₀, mul_comm _ m, mul_comm _ n, zpow_mul₀, zpow_mul₀, h], }, - { exact (fpow_strict_mono H).injective h, }, + { exact (zpow_strict_mono H).injective h, }, end -@[simp] lemma fpow_inj {x : K} (h₀ : 0 < x) (h₁ : x ≠ 1) {m n : ℤ} : +@[simp] lemma zpow_inj {x : K} (h₀ : 0 < x) (h₁ : x ≠ 1) {m n : ℤ} : x ^ m = x ^ n ↔ m = n := -(fpow_injective h₀ h₁).eq_iff +(zpow_injective h₀ h₁).eq_iff end ordered section variables {K : Type*} [field K] -@[simp, norm_cast] theorem rat.cast_fpow [char_zero K] (q : ℚ) (n : ℤ) : +@[simp, norm_cast] theorem rat.cast_zpow [char_zero K] (q : ℚ) (n : ℤ) : ((q ^ n : ℚ) : K) = q ^ n := -(rat.cast_hom K).map_fpow q n +(rat.cast_hom K).map_zpow q n end diff --git a/src/algebra/group/conj.lean b/src/algebra/group/conj.lean index 977d04f0a68cb..39efac7b61797 100644 --- a/src/algebra/group/conj.lean +++ b/src/algebra/group/conj.lean @@ -74,11 +74,11 @@ begin { simp [pow_succ, hi] } end -@[simp] lemma conj_gpow {i : ℤ} {a b : α} : (a * b * a⁻¹) ^ i = a * (b ^ i) * a⁻¹ := +@[simp] lemma conj_zpow {i : ℤ} {a b : α} : (a * b * a⁻¹) ^ i = a * (b ^ i) * a⁻¹ := begin induction i, { simp }, - { simp [gpow_neg_succ_of_nat, conj_pow] } + { simp [zpow_neg_succ_of_nat, conj_pow] } end lemma conj_injective {x : α} : function.injective (λ (g : α), x * g * x⁻¹) := diff --git a/src/algebra/group/defs.lean b/src/algebra/group/defs.lean index 290e128b09307..bec75e867fb61 100644 --- a/src/algebra/group/defs.lean +++ b/src/algebra/group/defs.lean @@ -532,9 +532,9 @@ instance cancel_comm_monoid.to_cancel_monoid (M : Type u) [cancel_comm_monoid M] end cancel_monoid -/-- The fundamental power operation in a group. `gpow_rec n a = a*a*...*a` n times, for integer `n`. +/-- The fundamental power operation in a group. `zpow_rec n a = a*a*...*a` n times, for integer `n`. Use instead `a ^ n`, which has better definitional behavior. -/ -def gpow_rec {M : Type*} [has_one M] [has_mul M] [has_inv M] : ℤ → M → M +def zpow_rec {M : Type*} [has_one M] [has_mul M] [has_inv M] : ℤ → M → M | (int.of_nat n) a := npow_rec n a | -[1+ n] a := (npow_rec n.succ a) ⁻¹ @@ -544,7 +544,7 @@ def gsmul_rec {M : Type*} [has_zero M] [has_add M] [has_neg M]: ℤ → M → M | (int.of_nat n) a := nsmul_rec n a | -[1+ n] a := - (nsmul_rec n.succ a) -attribute [to_additive] gpow_rec +attribute [to_additive] zpow_rec /-- A `div_inv_monoid` is a `monoid` with operations `/` and `⁻¹` satisfying `div_eq_mul_inv : ∀ a b, a / b = a * b⁻¹`. @@ -561,7 +561,7 @@ also have an instance `∀ X [cromulent X], group_with_zero (foo X)`. Then the `(/)` coming from `group_with_zero_has_div` cannot be definitionally equal to the `(/)` coming from `foo.has_div`. -In the same way, adding a `gpow` field makes it possible to avoid definitional failures +In the same way, adding a `zpow` field makes it possible to avoid definitional failures in diamonds. See the definition of `monoid` and Note [forgetful inheritance] for more explanations on this. -/ @@ -569,14 +569,14 @@ explanations on this. class div_inv_monoid (G : Type u) extends monoid G, has_inv G, has_div G := (div := λ a b, a * b⁻¹) (div_eq_mul_inv : ∀ a b : G, a / b = a * b⁻¹ . try_refl_tac) -(gpow : ℤ → G → G := gpow_rec) -(gpow_zero' : ∀ (a : G), gpow 0 a = 1 . try_refl_tac) -(gpow_succ' : - ∀ (n : ℕ) (a : G), gpow (int.of_nat n.succ) a = a * gpow (int.of_nat n) a . try_refl_tac) -(gpow_neg' : - ∀ (n : ℕ) (a : G), gpow (-[1+ n]) a = (gpow n.succ a)⁻¹ . try_refl_tac) +(zpow : ℤ → G → G := zpow_rec) +(zpow_zero' : ∀ (a : G), zpow 0 a = 1 . try_refl_tac) +(zpow_succ' : + ∀ (n : ℕ) (a : G), zpow (int.of_nat n.succ) a = a * zpow (int.of_nat n) a . try_refl_tac) +(zpow_neg' : + ∀ (n : ℕ) (a : G), zpow (-[1+ n]) a = (zpow n.succ a)⁻¹ . try_refl_tac) -export div_inv_monoid (gpow) +export div_inv_monoid (zpow) /-- A `sub_neg_monoid` is an `add_monoid` with unary `-` and binary `-` operations satisfying `sub_eq_add_neg : ∀ a b, a - b = a + -b`. @@ -617,9 +617,9 @@ begin let iM : div_inv_monoid M := m₁, unfreezingI { cases m₁ with mul₁ _ one₁ one_mul₁ mul_one₁ npow₁ npow_zero₁ npow_succ₁ inv₁ div₁ - div_eq_mul_inv₁ gpow₁ gpow_zero'₁ gpow_succ'₁ gpow_neg'₁, + div_eq_mul_inv₁ zpow₁ zpow_zero'₁ zpow_succ'₁ zpow_neg'₁, cases m₂ with mul₂ _ one₂ one_mul₂ mul_one₂ npow₂ npow_zero₂ npow_succ₂ inv₂ div₂ - div_eq_mul_inv₂ gpow₂ gpow_zero'₂ gpow_succ'₂ gpow_neg'₂ }, + div_eq_mul_inv₂ zpow₂ zpow_zero'₂ zpow_succ'₂ zpow_neg'₂ }, change mul₁ = mul₂ at h_mul, subst h_mul, have h_one : one₁ = one₂, @@ -640,21 +640,21 @@ begin { exact div_eq_mul_inv₁ a b }, { exact div_eq_mul_inv₂ a b } }, subst h_div, - have h_gpow_aux : ∀ n g, gpow₁ (int.of_nat n) g = gpow₂ (int.of_nat n) g, + have h_zpow_aux : ∀ n g, zpow₁ (int.of_nat n) g = zpow₂ (int.of_nat n) g, { intros n g, induction n with n IH, { convert (rfl : (1 : M) = 1), - { exact gpow_zero'₁ g }, - { exact gpow_zero'₂ g } }, - { rw [gpow_succ'₁, gpow_succ'₂, IH] } }, - have h_gpow : gpow₁ = gpow₂, + { exact zpow_zero'₁ g }, + { exact zpow_zero'₂ g } }, + { rw [zpow_succ'₁, zpow_succ'₂, IH] } }, + have h_zpow : zpow₁ = zpow₂, { ext z, cases z with z z, - { exact h_gpow_aux z x }, - { rw [gpow_neg'₁, gpow_neg'₂], + { exact h_zpow_aux z x }, + { rw [zpow_neg'₁, zpow_neg'₂], congr', - exact h_gpow_aux _ _ } }, - subst h_gpow, + exact h_zpow_aux _ _ } }, + subst h_zpow, end @[to_additive] @@ -754,9 +754,9 @@ begin let iG : group G := g₁, unfreezingI { cases g₁ with mul₁ _ one₁ one_mul₁ mul_one₁ npow₁ npow_zero'₁ npow_succ'₁ inv₁ div₁ - div_eq_mul_inv₁ gpow₁ gpow_zero'₁ gpow_succ'₁ gpow_neg'₁ mul_left_inv₁, + div_eq_mul_inv₁ zpow₁ zpow_zero'₁ zpow_succ'₁ zpow_neg'₁ mul_left_inv₁, cases g₂ with mul₂ _ one₂ one_mul₂ mul_one₂ npow₂ npow_zero'₂ npow_succ'₂ inv₂ div₂ - div_eq_mul_inv₂ gpow₂ gpow_zero'₂ gpow_succ'₂ gpow_neg'₂ mul_left_inv₂, }, + div_eq_mul_inv₂ zpow₂ zpow_zero'₂ zpow_succ'₂ zpow_neg'₂ mul_left_inv₂, }, change mul₁ = mul₂ at h_mul, subst h_mul, have h_one : one₁ = one₂, diff --git a/src/algebra/group/hom_instances.lean b/src/algebra/group/hom_instances.lean index edb9f6f1c2d25..6bf2d1435195d 100644 --- a/src/algebra/group/hom_instances.lean +++ b/src/algebra/group/hom_instances.lean @@ -61,12 +61,12 @@ instance {M G} [mul_one_class M] [comm_group G] : comm_group (M →* G) := div := has_div.div, div_eq_mul_inv := by { intros, ext, apply div_eq_mul_inv }, mul_left_inv := by intros; ext; apply mul_left_inv, - gpow := λ n f, { to_fun := λ x, gpow n (f x), + zpow := λ n f, { to_fun := λ x, zpow n (f x), map_one' := by simp, - map_mul' := λ x y, by simp [mul_gpow] }, - gpow_zero' := λ f, by { ext x, simp }, - gpow_succ' := λ n f, by { ext x, simp [gpow_of_nat, pow_succ] }, - gpow_neg' := λ n f, by { ext x, simp }, + map_mul' := λ x y, by simp [mul_zpow] }, + zpow_zero' := λ f, by { ext x, simp }, + zpow_succ' := λ n f, by { ext x, simp [zpow_of_nat, pow_succ] }, + zpow_neg' := λ n f, by { ext x, simp }, ..monoid_hom.comm_monoid } /-- If `G` is an additive commutative group, then `M →+ G` is an additive commutative group too. -/ diff --git a/src/algebra/group/pi.lean b/src/algebra/group/pi.lean index c63306f818a4d..dfafad5aa6807 100644 --- a/src/algebra/group/pi.lean +++ b/src/algebra/group/pi.lean @@ -59,17 +59,17 @@ tactic.pi_instance_derive_field instance div_inv_monoid [∀ i, div_inv_monoid $ f i] : div_inv_monoid (Π i : I, f i) := by refine_struct { one := (1 : Π i, f i), mul := (*), inv := has_inv.inv, div := has_div.div, - npow := npow, gpow := λ z x i, gpow z (x i) }; tactic.pi_instance_derive_field + npow := npow, zpow := λ z x i, zpow z (x i) }; tactic.pi_instance_derive_field @[to_additive] instance group [∀ i, group $ f i] : group (Π i : I, f i) := by refine_struct { one := (1 : Π i, f i), mul := (*), inv := has_inv.inv, div := has_div.div, - npow := npow, gpow := gpow }; tactic.pi_instance_derive_field + npow := npow, zpow := zpow }; tactic.pi_instance_derive_field @[to_additive] instance comm_group [∀ i, comm_group $ f i] : comm_group (Π i : I, f i) := by refine_struct { one := (1 : Π i, f i), mul := (*), inv := has_inv.inv, div := has_div.div, - npow := npow, gpow := gpow }; tactic.pi_instance_derive_field + npow := npow, zpow := zpow }; tactic.pi_instance_derive_field @[to_additive add_left_cancel_semigroup] instance left_cancel_semigroup [∀ i, left_cancel_semigroup $ f i] : diff --git a/src/algebra/group/prod.lean b/src/algebra/group/prod.lean index 8c1d82f2818b3..9d5f674c68af6 100644 --- a/src/algebra/group/prod.lean +++ b/src/algebra/group/prod.lean @@ -103,10 +103,10 @@ instance [monoid M] [monoid N] : monoid (M × N) := @[to_additive] instance [div_inv_monoid G] [div_inv_monoid H] : div_inv_monoid (G × H) := { div_eq_mul_inv := λ a b, mk.inj_iff.mpr ⟨div_eq_mul_inv _ _, div_eq_mul_inv _ _⟩, - gpow := λ z a, ⟨div_inv_monoid.gpow z a.1, div_inv_monoid.gpow z a.2⟩, - gpow_zero' := λ z, ext (div_inv_monoid.gpow_zero' _) (div_inv_monoid.gpow_zero' _), - gpow_succ' := λ z a, ext (div_inv_monoid.gpow_succ' _ _) (div_inv_monoid.gpow_succ' _ _), - gpow_neg' := λ z a, ext (div_inv_monoid.gpow_neg' _ _) (div_inv_monoid.gpow_neg' _ _), + zpow := λ z a, ⟨div_inv_monoid.zpow z a.1, div_inv_monoid.zpow z a.2⟩, + zpow_zero' := λ z, ext (div_inv_monoid.zpow_zero' _) (div_inv_monoid.zpow_zero' _), + zpow_succ' := λ z a, ext (div_inv_monoid.zpow_succ' _ _) (div_inv_monoid.zpow_succ' _ _), + zpow_neg' := λ z a, ext (div_inv_monoid.zpow_neg' _ _) (div_inv_monoid.zpow_neg' _ _), .. prod.monoid, .. prod.has_inv, .. prod.has_div } @[to_additive] diff --git a/src/algebra/group/to_additive.lean b/src/algebra/group/to_additive.lean index 3291e73e0b0bd..fc096820bb3e3 100644 --- a/src/algebra/group/to_additive.lean +++ b/src/algebra/group/to_additive.lean @@ -212,7 +212,7 @@ meta def tr : bool → list string → list string | is_comm ("prod" :: s) := add_comm_prefix is_comm "sum" :: tr ff s | is_comm ("finprod" :: s) := add_comm_prefix is_comm "finsum" :: tr ff s | is_comm ("npow" :: s) := add_comm_prefix is_comm "nsmul" :: tr ff s -| is_comm ("gpow" :: s) := add_comm_prefix is_comm "gsmul" :: tr ff s +| is_comm ("zpow" :: s) := add_comm_prefix is_comm "gsmul" :: tr ff s | is_comm ("monoid" :: s) := ("add_" ++ add_comm_prefix is_comm "monoid") :: tr ff s | is_comm ("submonoid" :: s) := ("add_" ++ add_comm_prefix is_comm "submonoid") :: tr ff s | is_comm ("group" :: s) := ("add_" ++ add_comm_prefix is_comm "group") :: tr ff s diff --git a/src/algebra/group/type_tags.lean b/src/algebra/group/type_tags.lean index 4c4108e6be764..25e5603af9394 100644 --- a/src/algebra/group/type_tags.lean +++ b/src/algebra/group/type_tags.lean @@ -229,18 +229,18 @@ rfl instance [div_inv_monoid α] : sub_neg_monoid (additive α) := { sub_eq_add_neg := @div_eq_mul_inv α _, - gsmul := @gpow α _, - gsmul_zero' := div_inv_monoid.gpow_zero', - gsmul_succ' := div_inv_monoid.gpow_succ', - gsmul_neg' := div_inv_monoid.gpow_neg', + gsmul := @zpow α _, + gsmul_zero' := div_inv_monoid.zpow_zero', + gsmul_succ' := div_inv_monoid.zpow_succ', + gsmul_neg' := div_inv_monoid.zpow_neg', .. additive.has_neg, .. additive.has_sub, .. additive.add_monoid } instance [sub_neg_monoid α] : div_inv_monoid (multiplicative α) := { div_eq_mul_inv := @sub_eq_add_neg α _, - gpow := @gsmul α _, - gpow_zero' := sub_neg_monoid.gsmul_zero', - gpow_succ' := sub_neg_monoid.gsmul_succ', - gpow_neg' := sub_neg_monoid.gsmul_neg', + zpow := @gsmul α _, + zpow_zero' := sub_neg_monoid.gsmul_zero', + zpow_succ' := sub_neg_monoid.gsmul_succ', + zpow_neg' := sub_neg_monoid.gsmul_neg', .. multiplicative.has_inv, .. multiplicative.has_div, .. multiplicative.monoid } instance [group α] : add_group (additive α) := diff --git a/src/algebra/group/ulift.lean b/src/algebra/group/ulift.lean index 1e790e3fd147d..00eb69c63d9b1 100644 --- a/src/algebra/group/ulift.lean +++ b/src/algebra/group/ulift.lean @@ -77,12 +77,12 @@ tactic.pi_instance_derive_field @[to_additive] instance group [group α] : group (ulift α) := by refine_struct { one := (1 : ulift α), mul := (*), inv := has_inv.inv, div := has_div.div, - npow := λ n f, ⟨npow n f.down⟩, gpow := λ n f, ⟨gpow n f.down⟩ }; tactic.pi_instance_derive_field + npow := λ n f, ⟨npow n f.down⟩, zpow := λ n f, ⟨zpow n f.down⟩ }; tactic.pi_instance_derive_field @[to_additive] instance comm_group [comm_group α] : comm_group (ulift α) := by refine_struct { one := (1 : ulift α), mul := (*), inv := has_inv.inv, div := has_div.div, - npow := λ n f, ⟨npow n f.down⟩, gpow := λ n f, ⟨gpow n f.down⟩ }; tactic.pi_instance_derive_field + npow := λ n f, ⟨npow n f.down⟩, zpow := λ n f, ⟨zpow n f.down⟩ }; tactic.pi_instance_derive_field @[to_additive add_left_cancel_semigroup] instance left_cancel_semigroup [left_cancel_semigroup α] : diff --git a/src/algebra/group_power/basic.lean b/src/algebra/group_power/basic.lean index 0a03c45def955..8c8c43ef4c867 100644 --- a/src/algebra/group_power/basic.lean +++ b/src/algebra/group_power/basic.lean @@ -47,7 +47,7 @@ instance add_monoid.has_scalar_nat [add_monoid M] : has_scalar ℕ M := ⟨nsmul attribute [to_additive add_monoid.has_scalar_nat] monoid.has_pow -instance div_inv_monoid.has_pow [div_inv_monoid M] : has_pow M ℤ := ⟨λ x n, gpow n x⟩ +instance div_inv_monoid.has_pow [div_inv_monoid M] : has_pow M ℤ := ⟨λ x n, zpow n x⟩ instance sub_neg_monoid.has_scalar_int [sub_neg_monoid M] : has_scalar ℤ M := ⟨gsmul⟩ @@ -57,7 +57,7 @@ attribute [to_additive sub_neg_monoid.has_scalar_int] div_inv_monoid.has_pow lemma npow_eq_pow {M : Type*} [monoid M] (n : ℕ) (x : M) : npow n x = x^n := rfl @[simp, to_additive gsmul_eq_smul] -lemma gpow_eq_pow {M : Type*} [div_inv_monoid M] (n : ℤ) (x : M) : gpow n x = x^n := rfl +lemma zpow_eq_pow {M : Type*} [div_inv_monoid M] (n : ℤ) (x : M) : zpow n x = x^n := rfl /-! ### Commutativity @@ -260,30 +260,30 @@ variable [div_inv_monoid G] open int @[simp, norm_cast, to_additive] -theorem gpow_coe_nat (a : G) (n : ℕ) : a ^ (n:ℤ) = a ^ n := +theorem zpow_coe_nat (a : G) (n : ℕ) : a ^ (n:ℤ) = a ^ n := begin induction n with n ih, - { change gpow 0 a = a ^ 0, rw [div_inv_monoid.gpow_zero', pow_zero] }, - { change gpow (of_nat n) a = a ^ n at ih, - change gpow (of_nat n.succ) a = a ^ n.succ, - rw [div_inv_monoid.gpow_succ', pow_succ, ih] } + { change zpow 0 a = a ^ 0, rw [div_inv_monoid.zpow_zero', pow_zero] }, + { change zpow (of_nat n) a = a ^ n at ih, + change zpow (of_nat n.succ) a = a ^ n.succ, + rw [div_inv_monoid.zpow_succ', pow_succ, ih] } end @[to_additive] -theorem gpow_of_nat (a : G) (n : ℕ) : a ^ of_nat n = a ^ n := -gpow_coe_nat _ _ +theorem zpow_of_nat (a : G) (n : ℕ) : a ^ of_nat n = a ^ n := +zpow_coe_nat _ _ @[simp, to_additive] -theorem gpow_neg_succ_of_nat (a : G) (n : ℕ) : a ^ -[1+n] = (a ^ n.succ)⁻¹ := -by { rw ← gpow_coe_nat, exact div_inv_monoid.gpow_neg' n a } +theorem zpow_neg_succ_of_nat (a : G) (n : ℕ) : a ^ -[1+n] = (a ^ n.succ)⁻¹ := +by { rw ← zpow_coe_nat, exact div_inv_monoid.zpow_neg' n a } @[simp, to_additive zero_gsmul] -theorem gpow_zero (a : G) : a ^ (0:ℤ) = 1 := -by { convert pow_zero a using 1, exact gpow_coe_nat a 0 } +theorem zpow_zero (a : G) : a ^ (0:ℤ) = 1 := +by { convert pow_zero a using 1, exact zpow_coe_nat a 0 } @[simp, to_additive one_gsmul] -theorem gpow_one (a : G) : a ^ (1:ℤ) = a := -by { convert pow_one a using 1, exact gpow_coe_nat a 1 } +theorem zpow_one (a : G) : a ^ (1:ℤ) = a := +by { convert pow_one a using 1, exact zpow_coe_nat a 1 } end div_inv_monoid @@ -314,31 +314,31 @@ theorem pow_inv_comm (a : G) (m n : ℕ) : (a⁻¹)^m * a^n = a^n * (a⁻¹)^m : end nat @[simp, to_additive gsmul_zero] -theorem one_gpow : ∀ (n : ℤ), (1 : G) ^ n = 1 -| (n : ℕ) := by rw [gpow_coe_nat, one_pow] -| -[1+ n] := by rw [gpow_neg_succ_of_nat, one_pow, one_inv] +theorem one_zpow : ∀ (n : ℤ), (1 : G) ^ n = 1 +| (n : ℕ) := by rw [zpow_coe_nat, one_pow] +| -[1+ n] := by rw [zpow_neg_succ_of_nat, one_pow, one_inv] @[simp, to_additive neg_gsmul] -theorem gpow_neg (a : G) : ∀ (n : ℤ), a ^ -n = (a ^ n)⁻¹ -| (n+1:ℕ) := div_inv_monoid.gpow_neg' _ _ +theorem zpow_neg (a : G) : ∀ (n : ℤ), a ^ -n = (a ^ n)⁻¹ +| (n+1:ℕ) := div_inv_monoid.zpow_neg' _ _ | 0 := by { change a ^ (0 : ℤ) = (a ^ (0 : ℤ))⁻¹, simp } -| -[1+ n] := by { rw [gpow_neg_succ_of_nat, inv_inv, ← gpow_coe_nat], refl } +| -[1+ n] := by { rw [zpow_neg_succ_of_nat, inv_inv, ← zpow_coe_nat], refl } -lemma mul_gpow_neg_one (a b : G) : (a*b)^(-(1:ℤ)) = b^(-(1:ℤ))*a^(-(1:ℤ)) := -by simp only [mul_inv_rev, gpow_one, gpow_neg] +lemma mul_zpow_neg_one (a b : G) : (a*b)^(-(1:ℤ)) = b^(-(1:ℤ))*a^(-(1:ℤ)) := +by simp only [mul_inv_rev, zpow_one, zpow_neg] @[to_additive neg_one_gsmul] -theorem gpow_neg_one (x : G) : x ^ (-1:ℤ) = x⁻¹ := -by { rw [← congr_arg has_inv.inv (pow_one x), gpow_neg, ← gpow_coe_nat], refl } +theorem zpow_neg_one (x : G) : x ^ (-1:ℤ) = x⁻¹ := +by { rw [← congr_arg has_inv.inv (pow_one x), zpow_neg, ← zpow_coe_nat], refl } @[to_additive gsmul_neg] -theorem inv_gpow (a : G) : ∀n:ℤ, a⁻¹ ^ n = (a ^ n)⁻¹ -| (n : ℕ) := by rw [gpow_coe_nat, gpow_coe_nat, inv_pow] -| -[1+ n] := by rw [gpow_neg_succ_of_nat, gpow_neg_succ_of_nat, inv_pow] +theorem inv_zpow (a : G) : ∀n:ℤ, a⁻¹ ^ n = (a ^ n)⁻¹ +| (n : ℕ) := by rw [zpow_coe_nat, zpow_coe_nat, inv_pow] +| -[1+ n] := by rw [zpow_neg_succ_of_nat, zpow_neg_succ_of_nat, inv_pow] @[to_additive add_commute.gsmul_add] -theorem commute.mul_gpow {a b : G} (h : commute a b) : ∀ n : ℤ, (a * b) ^ n = a ^ n * b ^ n -| (n : ℕ) := by simp [gpow_coe_nat, h.mul_pow n] +theorem commute.mul_zpow {a b : G} (h : commute a b) : ∀ n : ℤ, (a * b) ^ n = a ^ n * b ^ n +| (n : ℕ) := by simp [zpow_coe_nat, h.mul_pow n] | -[1+n] := by simp [h.mul_pow, (h.pow_pow n.succ n.succ).inv_inv.symm.eq] end group @@ -347,20 +347,20 @@ section comm_group variables [comm_group G] [add_comm_group A] @[to_additive gsmul_add] -theorem mul_gpow (a b : G) (n : ℤ) : (a * b)^n = a^n * b^n := (commute.all a b).mul_gpow n +theorem mul_zpow (a b : G) (n : ℤ) : (a * b)^n = a^n * b^n := (commute.all a b).mul_zpow n @[to_additive gsmul_sub] -theorem div_gpow (a b : G) (n : ℤ) : (a / b) ^ n = a ^ n / b ^ n := -by rw [div_eq_mul_inv, div_eq_mul_inv, mul_gpow, inv_gpow] +theorem div_zpow (a b : G) (n : ℤ) : (a / b) ^ n = a ^ n / b ^ n := +by rw [div_eq_mul_inv, div_eq_mul_inv, mul_zpow, inv_zpow] /-- The `n`th power map (`n` an integer) on a commutative group, considered as a group homomorphism. -/ @[to_additive "Multiplication by an integer `n` on a commutative additive group, considered as an additive group homomorphism.", simps] -def gpow_group_hom (n : ℤ) : G →* G := +def zpow_group_hom (n : ℤ) : G →* G := { to_fun := (^ n), - map_one' := one_gpow n, - map_mul' := λ a b, mul_gpow a b n } + map_one' := one_zpow n, + map_mul' := λ a b, mul_zpow a b n } end comm_group @@ -496,30 +496,30 @@ lemma of_add_gsmul [add_group A] (x : A) (n : ℤ) : lemma of_mul_pow {A : Type*} [monoid A] (x : A) (n : ℕ) : additive.of_mul (x ^ n) = n • (additive.of_mul x) := rfl -lemma of_mul_gpow [group G] (x : G) (n : ℤ) : additive.of_mul (x ^ n) = n • additive.of_mul x := +lemma of_mul_zpow [group G] (x : G) (n : ℤ) : additive.of_mul (x ^ n) = n • additive.of_mul x := rfl -@[simp] lemma semiconj_by.gpow_right [group G] {a x y : G} (h : semiconj_by a x y) : +@[simp] lemma semiconj_by.zpow_right [group G] {a x y : G} (h : semiconj_by a x y) : ∀ m : ℤ, semiconj_by a (x^m) (y^m) -| (n : ℕ) := by simp [gpow_coe_nat, h.pow_right n] +| (n : ℕ) := by simp [zpow_coe_nat, h.pow_right n] | -[1+n] := by simp [(h.pow_right n.succ).inv_right] namespace commute variables [group G] {a b : G} -@[simp] lemma gpow_right (h : commute a b) (m : ℤ) : commute a (b^m) := -h.gpow_right m +@[simp] lemma zpow_right (h : commute a b) (m : ℤ) : commute a (b^m) := +h.zpow_right m -@[simp] lemma gpow_left (h : commute a b) (m : ℤ) : commute (a^m) b := -(h.symm.gpow_right m).symm +@[simp] lemma zpow_left (h : commute a b) (m : ℤ) : commute (a^m) b := +(h.symm.zpow_right m).symm -lemma gpow_gpow (h : commute a b) (m n : ℤ) : commute (a^m) (b^n) := (h.gpow_left m).gpow_right n +lemma zpow_zpow (h : commute a b) (m n : ℤ) : commute (a^m) (b^n) := (h.zpow_left m).zpow_right n variables (a) (m n : ℤ) -@[simp] theorem self_gpow : commute a (a ^ n) := (commute.refl a).gpow_right n -@[simp] theorem gpow_self : commute (a ^ n) a := (commute.refl a).gpow_left n -@[simp] theorem gpow_gpow_self : commute (a ^ m) (a ^ n) := (commute.refl a).gpow_gpow m n +@[simp] theorem self_zpow : commute a (a ^ n) := (commute.refl a).zpow_right n +@[simp] theorem zpow_self : commute (a ^ n) a := (commute.refl a).zpow_left n +@[simp] theorem zpow_zpow_self : commute (a ^ m) (a ^ n) := (commute.refl a).zpow_zpow m n end commute diff --git a/src/algebra/group_power/lemmas.lean b/src/algebra/group_power/lemmas.lean index 3d179cab974d0..d3af132484d12 100644 --- a/src/algebra/group_power/lemmas.lean +++ b/src/algebra/group_power/lemmas.lean @@ -113,87 +113,87 @@ open nat theorem gsmul_one [has_one A] (n : ℤ) : n • (1 : A) = n := by cases n; simp -lemma gpow_add_one (a : G) : ∀ n : ℤ, a ^ (n + 1) = a ^ n * a +lemma zpow_add_one (a : G) : ∀ n : ℤ, a ^ (n + 1) = a ^ n * a | (of_nat n) := by simp [← int.coe_nat_succ, pow_succ'] | -[1+0] := by simp [int.neg_succ_of_nat_eq] -| -[1+(n+1)] := by rw [int.neg_succ_of_nat_eq, gpow_neg, neg_add, neg_add_cancel_right, gpow_neg, - ← int.coe_nat_succ, gpow_coe_nat, gpow_coe_nat, pow_succ _ (n + 1), mul_inv_rev, +| -[1+(n+1)] := by rw [int.neg_succ_of_nat_eq, zpow_neg, neg_add, neg_add_cancel_right, zpow_neg, + ← int.coe_nat_succ, zpow_coe_nat, zpow_coe_nat, pow_succ _ (n + 1), mul_inv_rev, inv_mul_cancel_right] theorem add_one_gsmul : ∀ (a : A) (i : ℤ), (i + 1) • a = i • a + a := -@gpow_add_one (multiplicative A) _ +@zpow_add_one (multiplicative A) _ -lemma gpow_sub_one (a : G) (n : ℤ) : a ^ (n - 1) = a ^ n * a⁻¹ := +lemma zpow_sub_one (a : G) (n : ℤ) : a ^ (n - 1) = a ^ n * a⁻¹ := calc a ^ (n - 1) = a ^ (n - 1) * a * a⁻¹ : (mul_inv_cancel_right _ _).symm - ... = a^n * a⁻¹ : by rw [← gpow_add_one, sub_add_cancel] + ... = a^n * a⁻¹ : by rw [← zpow_add_one, sub_add_cancel] -lemma gpow_add (a : G) (m n : ℤ) : a ^ (m + n) = a ^ m * a ^ n := +lemma zpow_add (a : G) (m n : ℤ) : a ^ (m + n) = a ^ m * a ^ n := begin induction n using int.induction_on with n ihn n ihn, case hz : { simp }, - { simp only [← add_assoc, gpow_add_one, ihn, mul_assoc] }, - { rw [gpow_sub_one, ← mul_assoc, ← ihn, ← gpow_sub_one, add_sub_assoc] } + { simp only [← add_assoc, zpow_add_one, ihn, mul_assoc] }, + { rw [zpow_sub_one, ← mul_assoc, ← ihn, ← zpow_sub_one, add_sub_assoc] } end -lemma mul_self_gpow (b : G) (m : ℤ) : b*b^m = b^(m+1) := -by { conv_lhs {congr, rw ← gpow_one b }, rw [← gpow_add, add_comm] } +lemma mul_self_zpow (b : G) (m : ℤ) : b*b^m = b^(m+1) := +by { conv_lhs {congr, rw ← zpow_one b }, rw [← zpow_add, add_comm] } -lemma mul_gpow_self (b : G) (m : ℤ) : b^m*b = b^(m+1) := -by { conv_lhs {congr, skip, rw ← gpow_one b }, rw [← gpow_add, add_comm] } +lemma mul_zpow_self (b : G) (m : ℤ) : b^m*b = b^(m+1) := +by { conv_lhs {congr, skip, rw ← zpow_one b }, rw [← zpow_add, add_comm] } theorem add_gsmul : ∀ (a : A) (i j : ℤ), (i + j) • a = i • a + j • a := -@gpow_add (multiplicative A) _ +@zpow_add (multiplicative A) _ -lemma gpow_sub (a : G) (m n : ℤ) : a ^ (m - n) = a ^ m * (a ^ n)⁻¹ := -by rw [sub_eq_add_neg, gpow_add, gpow_neg] +lemma zpow_sub (a : G) (m n : ℤ) : a ^ (m - n) = a ^ m * (a ^ n)⁻¹ := +by rw [sub_eq_add_neg, zpow_add, zpow_neg] lemma sub_gsmul (m n : ℤ) (a : A) : (m - n) • a = m • a - n • a := -by simpa only [sub_eq_add_neg] using @gpow_sub (multiplicative A) _ _ _ _ +by simpa only [sub_eq_add_neg] using @zpow_sub (multiplicative A) _ _ _ _ -theorem gpow_one_add (a : G) (i : ℤ) : a ^ (1 + i) = a * a ^ i := -by rw [gpow_add, gpow_one] +theorem zpow_one_add (a : G) (i : ℤ) : a ^ (1 + i) = a * a ^ i := +by rw [zpow_add, zpow_one] theorem one_add_gsmul : ∀ (a : A) (i : ℤ), (1 + i) • a = a + i • a := -@gpow_one_add (multiplicative A) _ +@zpow_one_add (multiplicative A) _ -theorem gpow_mul_comm (a : G) (i j : ℤ) : a ^ i * a ^ j = a ^ j * a ^ i := -by rw [← gpow_add, ← gpow_add, add_comm] +theorem zpow_mul_comm (a : G) (i j : ℤ) : a ^ i * a ^ j = a ^ j * a ^ i := +by rw [← zpow_add, ← zpow_add, add_comm] theorem gsmul_add_comm : ∀ (a : A) (i j : ℤ), i • a + j • a = j • a + i • a := -@gpow_mul_comm (multiplicative A) _ +@zpow_mul_comm (multiplicative A) _ -theorem gpow_mul (a : G) (m n : ℤ) : a ^ (m * n) = (a ^ m) ^ n := -int.induction_on n (by simp) (λ n ihn, by simp [mul_add, gpow_add, ihn]) - (λ n ihn, by simp only [mul_sub, gpow_sub, ihn, mul_one, gpow_one]) +theorem zpow_mul (a : G) (m n : ℤ) : a ^ (m * n) = (a ^ m) ^ n := +int.induction_on n (by simp) (λ n ihn, by simp [mul_add, zpow_add, ihn]) + (λ n ihn, by simp only [mul_sub, zpow_sub, ihn, mul_one, zpow_one]) theorem gsmul_mul' : ∀ (a : A) (m n : ℤ), (m * n) • a = n • (m • a) := -@gpow_mul (multiplicative A) _ +@zpow_mul (multiplicative A) _ -theorem gpow_mul' (a : G) (m n : ℤ) : a ^ (m * n) = (a ^ n) ^ m := -by rw [mul_comm, gpow_mul] +theorem zpow_mul' (a : G) (m n : ℤ) : a ^ (m * n) = (a ^ n) ^ m := +by rw [mul_comm, zpow_mul] theorem mul_gsmul (a : A) (m n : ℤ) : (m * n) • a = m • (n • a) := by rw [mul_comm, gsmul_mul'] -theorem gpow_bit0 (a : G) (n : ℤ) : a ^ bit0 n = a ^ n * a ^ n := gpow_add _ _ _ +theorem zpow_bit0 (a : G) (n : ℤ) : a ^ bit0 n = a ^ n * a ^ n := zpow_add _ _ _ theorem bit0_gsmul (a : A) (n : ℤ) : bit0 n • a = n • a + n • a := -@gpow_bit0 (multiplicative A) _ _ _ +@zpow_bit0 (multiplicative A) _ _ _ -theorem gpow_bit1 (a : G) (n : ℤ) : a ^ bit1 n = a ^ n * a ^ n * a := -by rw [bit1, gpow_add, gpow_bit0, gpow_one] +theorem zpow_bit1 (a : G) (n : ℤ) : a ^ bit1 n = a ^ n * a ^ n * a := +by rw [bit1, zpow_add, zpow_bit0, zpow_one] theorem bit1_gsmul : ∀ (a : A) (n : ℤ), bit1 n • a = n • a + n • a + a := -@gpow_bit1 (multiplicative A) _ +@zpow_bit1 (multiplicative A) _ -@[simp] theorem monoid_hom.map_gpow (f : G →* H) (a : G) (n : ℤ) : f (a ^ n) = f a ^ n := +@[simp] theorem monoid_hom.map_zpow (f : G →* H) (a : G) (n : ℤ) : f (a ^ n) = f a ^ n := by cases n; simp @[simp] theorem add_monoid_hom.map_gsmul (f : A →+ B) (a : A) (n : ℤ) : f (n • a) = n • f a := -f.to_multiplicative.map_gpow a n +f.to_multiplicative.map_zpow a n -@[simp, norm_cast] lemma units.coe_gpow (u : units G) (n : ℤ) : ((u ^ n : units G) : G) = u ^ n := -(units.coe_hom G).map_gpow u n +@[simp, norm_cast] lemma units.coe_zpow (u : units G) (n : ℤ) : ((u ^ n : units G) : G) = u ^ n := +(units.coe_hom G).map_zpow u n end group @@ -652,11 +652,11 @@ def powers_hom [monoid M] : M ≃ (multiplicative ℕ →* M) := /-- Monoid homomorphisms from `multiplicative ℤ` are defined by the image of `multiplicative.of_add 1`. -/ -def gpowers_hom [group G] : G ≃ (multiplicative ℤ →* G) := -{ to_fun := λ x, ⟨λ n, x ^ n.to_add, gpow_zero x, λ m n, gpow_add x m n⟩, +def zpowers_hom [group G] : G ≃ (multiplicative ℤ →* G) := +{ to_fun := λ x, ⟨λ n, x ^ n.to_add, zpow_zero x, λ m n, zpow_add x m n⟩, inv_fun := λ f, f (multiplicative.of_add 1), - left_inv := gpow_one, - right_inv := λ f, monoid_hom.ext $ λ n, by { simp [← f.map_gpow, ← of_add_gsmul ] } } + left_inv := zpow_one, + right_inv := λ f, monoid_hom.ext $ λ n, by { simp [← f.map_zpow, ← of_add_gsmul ] } } /-- Additive homomorphisms from `ℕ` are defined by the image of `1`. -/ def multiples_hom [add_monoid A] : A ≃ (ℕ →+ A) := @@ -680,11 +680,11 @@ variables {M G A} @[simp] lemma powers_hom_symm_apply [monoid M] (f : multiplicative ℕ →* M) : (powers_hom M).symm f = f (multiplicative.of_add 1) := rfl -@[simp] lemma gpowers_hom_apply [group G] (x : G) (n : multiplicative ℤ) : - gpowers_hom G x n = x ^ n.to_add := rfl +@[simp] lemma zpowers_hom_apply [group G] (x : G) (n : multiplicative ℤ) : + zpowers_hom G x n = x ^ n.to_add := rfl -@[simp] lemma gpowers_hom_symm_apply [group G] (f : multiplicative ℤ →* G) : - (gpowers_hom G).symm f = f (multiplicative.of_add 1) := rfl +@[simp] lemma zpowers_hom_symm_apply [group G] (f : multiplicative ℤ →* G) : + (zpowers_hom G).symm f = f (multiplicative.of_add 1) := rfl @[simp] lemma multiples_hom_apply [add_monoid A] (x : A) (n : ℕ) : multiples_hom A x n = n • x := rfl @@ -708,7 +708,7 @@ monoid_hom.ext $ λ n, by rw [f.apply_mnat, g.apply_mnat, h] lemma monoid_hom.apply_mint [group M] (f : multiplicative ℤ →* M) (n : multiplicative ℤ) : f n = (f (multiplicative.of_add 1)) ^ n.to_add := -by rw [← gpowers_hom_symm_apply, ← gpowers_hom_apply, equiv.apply_symm_apply] +by rw [← zpowers_hom_symm_apply, ← zpowers_hom_apply, equiv.apply_symm_apply] /-! `monoid_hom.ext_mint` is defined in `data.int.cast` -/ @@ -731,10 +731,10 @@ def powers_mul_hom [comm_monoid M] : M ≃* (multiplicative ℕ →* M) := { map_mul' := λ a b, monoid_hom.ext $ by simp [mul_pow], ..powers_hom M} -/-- If `M` is commutative, `gpowers_hom` is a multiplicative equivalence. -/ -def gpowers_mul_hom [comm_group G] : G ≃* (multiplicative ℤ →* G) := -{ map_mul' := λ a b, monoid_hom.ext $ by simp [mul_gpow], - ..gpowers_hom G} +/-- If `M` is commutative, `zpowers_hom` is a multiplicative equivalence. -/ +def zpowers_mul_hom [comm_group G] : G ≃* (multiplicative ℤ →* G) := +{ map_mul' := λ a b, monoid_hom.ext $ by simp [mul_zpow], + ..zpowers_hom G} /-- If `M` is commutative, `multiples_hom` is an additive equivalence. -/ def multiples_add_hom [add_comm_monoid A] : A ≃+ (ℕ →+ A) := @@ -754,11 +754,11 @@ variables {M G A} @[simp] lemma powers_mul_hom_symm_apply [comm_monoid M] (f : multiplicative ℕ →* M) : (powers_mul_hom M).symm f = f (multiplicative.of_add 1) := rfl -@[simp] lemma gpowers_mul_hom_apply [comm_group G] (x : G) (n : multiplicative ℤ) : - gpowers_mul_hom G x n = x ^ n.to_add := rfl +@[simp] lemma zpowers_mul_hom_apply [comm_group G] (x : G) (n : multiplicative ℤ) : + zpowers_mul_hom G x n = x ^ n.to_add := rfl -@[simp] lemma gpowers_mul_hom_symm_apply [comm_group G] (f : multiplicative ℤ →* G) : - (gpowers_mul_hom G).symm f = f (multiplicative.of_add 1) := rfl +@[simp] lemma zpowers_mul_hom_symm_apply [comm_group G] (f : multiplicative ℤ →* G) : + (zpowers_mul_hom G).symm f = f (multiplicative.of_add 1) := rfl @[simp] lemma multiples_add_hom_apply [add_comm_monoid A] (x : A) (n : ℕ) : multiples_add_hom A x n = n • x := rfl @@ -775,7 +775,7 @@ variables {M G A} /-! ### Commutativity (again) -Facts about `semiconj_by` and `commute` that require `gpow` or `gsmul`, or the fact that integer +Facts about `semiconj_by` and `commute` that require `zpow` or `gsmul`, or the fact that integer multiplication equals semiring multiplication. -/ @@ -800,10 +800,10 @@ end variables [monoid M] [group G] [ring R] -@[simp] lemma units_gpow_right {a : M} {x y : units M} (h : semiconj_by a x y) : +@[simp] lemma units_zpow_right {a : M} {x y : units M} (h : semiconj_by a x y) : ∀ m : ℤ, semiconj_by a (↑(x^m)) (↑(y^m)) -| (n : ℕ) := by simp only [gpow_coe_nat, units.coe_pow, h, pow_right] -| -[1+n] := by simp only [gpow_neg_succ_of_nat, units.coe_pow, units_inv_right, h, pow_right] +| (n : ℕ) := by simp only [zpow_coe_nat, units.coe_pow, h, pow_right] +| -[1+n] := by simp only [zpow_neg_succ_of_nat, units.coe_pow, units_inv_right, h, pow_right] variables {a b x y x' y' : R} @@ -849,13 +849,13 @@ end variables [monoid M] [group G] [ring R] -@[simp] lemma units_gpow_right {a : M} {u : units M} (h : commute a u) (m : ℤ) : +@[simp] lemma units_zpow_right {a : M} {u : units M} (h : commute a u) (m : ℤ) : commute a (↑(u^m)) := -h.units_gpow_right m +h.units_zpow_right m -@[simp] lemma units_gpow_left {u : units M} {a : M} (h : commute ↑u a) (m : ℤ) : +@[simp] lemma units_zpow_left {u : units M} {a : M} (h : commute ↑u a) (m : ℤ) : commute (↑(u^m)) a := -(h.symm.units_gpow_right m).symm +(h.symm.units_zpow_right m).symm variables {a b : R} @@ -902,13 +902,13 @@ end @[simp] lemma int.to_add_pow (a : multiplicative ℤ) (b : ℕ) : to_add (a ^ b) = to_add a * b := by induction b; simp [*, mul_add, pow_succ, add_comm] -@[simp] lemma int.to_add_gpow (a : multiplicative ℤ) (b : ℤ) : to_add (a ^ b) = to_add a * b := +@[simp] lemma int.to_add_zpow (a : multiplicative ℤ) (b : ℤ) : to_add (a ^ b) = to_add a * b := int.induction_on b (by simp) - (by simp [gpow_add, mul_add] {contextual := tt}) - (by simp [gpow_add, mul_add, sub_eq_add_neg, -int.add_neg_one] {contextual := tt}) + (by simp [zpow_add, mul_add] {contextual := tt}) + (by simp [zpow_add, mul_add, sub_eq_add_neg, -int.add_neg_one] {contextual := tt}) @[simp] lemma int.of_add_mul (a b : ℤ) : of_add (a * b) = of_add a ^ b := -(int.to_add_gpow _ _).symm +(int.to_add_zpow _ _).symm end multiplicative @@ -931,7 +931,7 @@ namespace opposite @[simp] lemma unop_pow [monoid M] (x : Mᵒᵖ) (n : ℕ) : unop (x ^ n) = (unop x) ^ n := rfl /-- Moving to the opposite group or group_with_zero commutes with taking powers. -/ -@[simp] lemma op_gpow [div_inv_monoid M] (x : M) (z : ℤ) : op (x ^ z) = (op x) ^ z := rfl -@[simp] lemma unop_gpow [div_inv_monoid M] (x : Mᵒᵖ) (z : ℤ) : unop (x ^ z) = (unop x) ^ z := rfl +@[simp] lemma op_zpow [div_inv_monoid M] (x : M) (z : ℤ) : op (x ^ z) = (op x) ^ z := rfl +@[simp] lemma unop_zpow [div_inv_monoid M] (x : Mᵒᵖ) (z : ℤ) : unop (x ^ z) = (unop x) ^ z := rfl end opposite diff --git a/src/algebra/group_power/order.lean b/src/algebra/group_power/order.lean index c2fe4946e3e03..cc5401e8f2442 100644 --- a/src/algebra/group_power/order.lean +++ b/src/algebra/group_power/order.lean @@ -104,11 +104,11 @@ section group variables [group G] [preorder G] [covariant_class G G (*) (≤)] @[to_additive gsmul_nonneg] -theorem one_le_gpow {x : G} (H : 1 ≤ x) {n : ℤ} (hn : 0 ≤ n) : +theorem one_le_zpow {x : G} (H : 1 ≤ x) {n : ℤ} (hn : 0 ≤ n) : 1 ≤ x ^ n := begin lift n to ℕ using hn, - rw gpow_coe_nat, + rw zpow_coe_nat, apply one_le_pow_of_one_le' H, end diff --git a/src/algebra/group_with_zero/power.lean b/src/algebra/group_with_zero/power.lean index e6d46601c63a3..d0b51d1a39609 100644 --- a/src/algebra/group_with_zero/power.lean +++ b/src/algebra/group_with_zero/power.lean @@ -61,53 +61,53 @@ variables {G₀ : Type*} [group_with_zero G₀] local attribute [ematch] le_of_lt -@[simp] theorem one_fpow : ∀ (n : ℤ), (1 : G₀) ^ n = 1 -| (n : ℕ) := by rw [gpow_coe_nat, one_pow] -| -[1+ n] := by rw [gpow_neg_succ_of_nat, one_pow, inv_one] +@[simp] theorem one_zpow₀ : ∀ (n : ℤ), (1 : G₀) ^ n = 1 +| (n : ℕ) := by rw [zpow_coe_nat, one_pow] +| -[1+ n] := by rw [zpow_neg_succ_of_nat, one_pow, inv_one] -lemma zero_fpow : ∀ z : ℤ, z ≠ 0 → (0 : G₀) ^ z = 0 -| (n : ℕ) h := by { rw [gpow_coe_nat, zero_pow'], simpa using h } +lemma zero_zpow : ∀ z : ℤ, z ≠ 0 → (0 : G₀) ^ z = 0 +| (n : ℕ) h := by { rw [zpow_coe_nat, zero_pow'], simpa using h } | -[1+n] h := by simp lemma fzero_pow_eq (n : ℤ) : (0 : G₀) ^ n = if n = 0 then 1 else 0 := begin split_ifs with h, - { rw [h, gpow_zero] }, - { rw [zero_fpow _ h] } + { rw [h, zpow_zero] }, + { rw [zero_zpow _ h] } end -@[simp] theorem fpow_neg (a : G₀) : ∀ (n : ℤ), a ^ -n = (a ^ n)⁻¹ -| (n+1:ℕ) := div_inv_monoid.gpow_neg' _ _ +@[simp] theorem zpow_neg₀ (a : G₀) : ∀ (n : ℤ), a ^ -n = (a ^ n)⁻¹ +| (n+1:ℕ) := div_inv_monoid.zpow_neg' _ _ | 0 := by { change a ^ (0 : ℤ) = (a ^ (0 : ℤ))⁻¹, simp } -| -[1+ n] := by { rw [gpow_neg_succ_of_nat, inv_inv₀, ← gpow_coe_nat], refl } +| -[1+ n] := by { rw [zpow_neg_succ_of_nat, inv_inv₀, ← zpow_coe_nat], refl } -theorem fpow_neg_one (x : G₀) : x ^ (-1:ℤ) = x⁻¹ := -by { rw [← congr_arg has_inv.inv (pow_one x), fpow_neg, ← gpow_coe_nat], refl } +theorem zpow_neg_one₀ (x : G₀) : x ^ (-1:ℤ) = x⁻¹ := +by { rw [← congr_arg has_inv.inv (pow_one x), zpow_neg₀, ← zpow_coe_nat], refl } -theorem inv_fpow (a : G₀) : ∀n:ℤ, a⁻¹ ^ n = (a ^ n)⁻¹ -| (n : ℕ) := by rw [gpow_coe_nat, gpow_coe_nat, inv_pow₀] -| -[1+ n] := by rw [gpow_neg_succ_of_nat, gpow_neg_succ_of_nat, inv_pow₀] +theorem inv_zpow₀ (a : G₀) : ∀n:ℤ, a⁻¹ ^ n = (a ^ n)⁻¹ +| (n : ℕ) := by rw [zpow_coe_nat, zpow_coe_nat, inv_pow₀] +| -[1+ n] := by rw [zpow_neg_succ_of_nat, zpow_neg_succ_of_nat, inv_pow₀] -lemma fpow_add_one {a : G₀} (ha : a ≠ 0) : ∀ n : ℤ, a ^ (n + 1) = a ^ n * a +lemma zpow_add_one₀ {a : G₀} (ha : a ≠ 0) : ∀ n : ℤ, a ^ (n + 1) = a ^ n * a | (n : ℕ) := by simp [← int.coe_nat_succ, pow_succ'] | -[1+0] := by simp [int.neg_succ_of_nat_eq, ha] -| -[1+(n+1)] := by rw [int.neg_succ_of_nat_eq, fpow_neg, neg_add, neg_add_cancel_right, fpow_neg, - ← int.coe_nat_succ, gpow_coe_nat, gpow_coe_nat, pow_succ _ (n + 1), mul_inv_rev₀, mul_assoc, +| -[1+(n+1)] := by rw [int.neg_succ_of_nat_eq, zpow_neg₀, neg_add, neg_add_cancel_right, zpow_neg₀, + ← int.coe_nat_succ, zpow_coe_nat, zpow_coe_nat, pow_succ _ (n + 1), mul_inv_rev₀, mul_assoc, inv_mul_cancel ha, mul_one] -lemma fpow_sub_one {a : G₀} (ha : a ≠ 0) (n : ℤ) : a ^ (n - 1) = a ^ n * a⁻¹ := +lemma zpow_sub_one₀ {a : G₀} (ha : a ≠ 0) (n : ℤ) : a ^ (n - 1) = a ^ n * a⁻¹ := calc a ^ (n - 1) = a ^ (n - 1) * a * a⁻¹ : by rw [mul_assoc, mul_inv_cancel ha, mul_one] - ... = a^n * a⁻¹ : by rw [← fpow_add_one ha, sub_add_cancel] + ... = a^n * a⁻¹ : by rw [← zpow_add_one₀ ha, sub_add_cancel] -lemma fpow_add {a : G₀} (ha : a ≠ 0) (m n : ℤ) : a ^ (m + n) = a ^ m * a ^ n := +lemma zpow_add₀ {a : G₀} (ha : a ≠ 0) (m n : ℤ) : a ^ (m + n) = a ^ m * a ^ n := begin induction n using int.induction_on with n ihn n ihn, case hz : { simp }, - { simp only [← add_assoc, fpow_add_one ha, ihn, mul_assoc] }, - { rw [fpow_sub_one ha, ← mul_assoc, ← ihn, ← fpow_sub_one ha, add_sub_assoc] } + { simp only [← add_assoc, zpow_add_one₀ ha, ihn, mul_assoc] }, + { rw [zpow_sub_one₀ ha, ← mul_assoc, ← ihn, ← zpow_sub_one₀ ha, add_sub_assoc] } end -lemma fpow_add' {a : G₀} {m n : ℤ} (h : a ≠ 0 ∨ m + n ≠ 0 ∨ m = 0 ∧ n = 0) : +lemma zpow_add' {a : G₀} {m n : ℤ} (h : a ≠ 0 ∨ m + n ≠ 0 ∨ m = 0 ∧ n = 0) : a ^ (m + n) = a ^ m * a ^ n := begin by_cases hm : m = 0, { simp [hm] }, @@ -115,112 +115,112 @@ begin by_cases ha : a = 0, { subst a, simp only [false_or, eq_self_iff_true, not_true, ne.def, hm, hn, false_and, or_false] at h, - rw [zero_fpow _ h, zero_fpow _ hm, zero_mul] }, - { exact fpow_add ha m n } + rw [zero_zpow _ h, zero_zpow _ hm, zero_mul] }, + { exact zpow_add₀ ha m n } end -theorem fpow_one_add {a : G₀} (h : a ≠ 0) (i : ℤ) : a ^ (1 + i) = a * a ^ i := -by rw [fpow_add h, gpow_one] +theorem zpow_one_add₀ {a : G₀} (h : a ≠ 0) (i : ℤ) : a ^ (1 + i) = a * a ^ i := +by rw [zpow_add₀ h, zpow_one] -theorem semiconj_by.fpow_right {a x y : G₀} (h : semiconj_by a x y) : +theorem semiconj_by.zpow_right₀ {a x y : G₀} (h : semiconj_by a x y) : ∀ m : ℤ, semiconj_by a (x^m) (y^m) | (n : ℕ) := by simp [h.pow_right n] | -[1+n] := by simp [(h.pow_right (n + 1)).inv_right₀] -theorem commute.fpow_right {a b : G₀} (h : commute a b) : ∀ m : ℤ, commute a (b^m) := -h.fpow_right +theorem commute.zpow_right₀ {a b : G₀} (h : commute a b) : ∀ m : ℤ, commute a (b^m) := +h.zpow_right₀ -theorem commute.fpow_left {a b : G₀} (h : commute a b) (m : ℤ) : commute (a^m) b := -(h.symm.fpow_right m).symm +theorem commute.zpow_left₀ {a b : G₀} (h : commute a b) (m : ℤ) : commute (a^m) b := +(h.symm.zpow_right₀ m).symm -theorem commute.fpow_fpow {a b : G₀} (h : commute a b) (m n : ℤ) : commute (a^m) (b^n) := -(h.fpow_left m).fpow_right n +theorem commute.zpow_zpow₀ {a b : G₀} (h : commute a b) (m n : ℤ) : commute (a^m) (b^n) := +(h.zpow_left₀ m).zpow_right₀ n -theorem commute.fpow_self (a : G₀) (n : ℤ) : commute (a^n) a := (commute.refl a).fpow_left n +theorem commute.zpow_self₀ (a : G₀) (n : ℤ) : commute (a^n) a := (commute.refl a).zpow_left₀ n -theorem commute.self_fpow (a : G₀) (n : ℤ) : commute a (a^n) := (commute.refl a).fpow_right n +theorem commute.self_zpow₀ (a : G₀) (n : ℤ) : commute a (a^n) := (commute.refl a).zpow_right₀ n -theorem commute.fpow_fpow_self (a : G₀) (m n : ℤ) : commute (a^m) (a^n) := -(commute.refl a).fpow_fpow m n +theorem commute.zpow_zpow_self₀ (a : G₀) (m n : ℤ) : commute (a^m) (a^n) := +(commute.refl a).zpow_zpow₀ m n -theorem fpow_bit0 (a : G₀) (n : ℤ) : a ^ bit0 n = a ^ n * a ^ n := +theorem zpow_bit0₀ (a : G₀) (n : ℤ) : a ^ bit0 n = a ^ n * a ^ n := begin - apply fpow_add', right, + apply zpow_add', right, by_cases hn : n = 0, { simp [hn] }, { simp [← two_mul, hn, two_ne_zero] } end -theorem fpow_bit1 (a : G₀) (n : ℤ) : a ^ bit1 n = a ^ n * a ^ n * a := +theorem zpow_bit1₀ (a : G₀) (n : ℤ) : a ^ bit1 n = a ^ n * a ^ n * a := begin - rw [← fpow_bit0, bit1, fpow_add', gpow_one], + rw [← zpow_bit0₀, bit1, zpow_add', zpow_one], right, left, apply bit1_ne_zero end -theorem fpow_mul (a : G₀) : ∀ m n : ℤ, a ^ (m * n) = (a ^ m) ^ n -| (m : ℕ) (n : ℕ) := by { rw [gpow_coe_nat, gpow_coe_nat, ← pow_mul, ← gpow_coe_nat], refl } -| (m : ℕ) -[1+ n] := by { rw [gpow_coe_nat, gpow_neg_succ_of_nat, ← pow_mul, coe_nat_mul_neg_succ, - fpow_neg, inv_inj₀, ← gpow_coe_nat], refl } -| -[1+ m] (n : ℕ) := by { rw [gpow_coe_nat, gpow_neg_succ_of_nat, ← inv_pow₀, ← pow_mul, - neg_succ_mul_coe_nat, fpow_neg, inv_pow₀, inv_inj₀, ← gpow_coe_nat], refl } -| -[1+ m] -[1+ n] := by { rw [gpow_neg_succ_of_nat, gpow_neg_succ_of_nat, neg_succ_mul_neg_succ, - inv_pow₀, inv_inv₀, ← pow_mul, ← gpow_coe_nat], refl } +theorem zpow_mul₀ (a : G₀) : ∀ m n : ℤ, a ^ (m * n) = (a ^ m) ^ n +| (m : ℕ) (n : ℕ) := by { rw [zpow_coe_nat, zpow_coe_nat, ← pow_mul, ← zpow_coe_nat], refl } +| (m : ℕ) -[1+ n] := by { rw [zpow_coe_nat, zpow_neg_succ_of_nat, ← pow_mul, coe_nat_mul_neg_succ, + zpow_neg₀, inv_inj₀, ← zpow_coe_nat], refl } +| -[1+ m] (n : ℕ) := by { rw [zpow_coe_nat, zpow_neg_succ_of_nat, ← inv_pow₀, ← pow_mul, + neg_succ_mul_coe_nat, zpow_neg₀, inv_pow₀, inv_inj₀, ← zpow_coe_nat], refl } +| -[1+ m] -[1+ n] := by { rw [zpow_neg_succ_of_nat, zpow_neg_succ_of_nat, neg_succ_mul_neg_succ, + inv_pow₀, inv_inv₀, ← pow_mul, ← zpow_coe_nat], refl } -theorem fpow_mul' (a : G₀) (m n : ℤ) : a ^ (m * n) = (a ^ n) ^ m := -by rw [mul_comm, fpow_mul] +theorem zpow_mul₀' (a : G₀) (m n : ℤ) : a ^ (m * n) = (a ^ n) ^ m := +by rw [mul_comm, zpow_mul₀] -@[simp, norm_cast] lemma units.coe_gpow₀ (u : units G₀) : +@[simp, norm_cast] lemma units.coe_zpow₀ (u : units G₀) : ∀ (n : ℤ), ((u ^ n : units G₀) : G₀) = u ^ n -| (n : ℕ) := by { rw [gpow_coe_nat, gpow_coe_nat], exact u.coe_pow n } -| -[1+k] := by rw [gpow_neg_succ_of_nat, gpow_neg_succ_of_nat, units.coe_inv', u.coe_pow] +| (n : ℕ) := by { rw [zpow_coe_nat, zpow_coe_nat], exact u.coe_pow n } +| -[1+k] := by rw [zpow_neg_succ_of_nat, zpow_neg_succ_of_nat, units.coe_inv', u.coe_pow] -lemma fpow_ne_zero_of_ne_zero {a : G₀} (ha : a ≠ 0) : ∀ (z : ℤ), a ^ z ≠ 0 -| (n : ℕ) := by { rw gpow_coe_nat, exact pow_ne_zero _ ha } -| -[1+n] := by { rw gpow_neg_succ_of_nat, exact inv_ne_zero (pow_ne_zero _ ha) } +lemma zpow_ne_zero_of_ne_zero {a : G₀} (ha : a ≠ 0) : ∀ (z : ℤ), a ^ z ≠ 0 +| (n : ℕ) := by { rw zpow_coe_nat, exact pow_ne_zero _ ha } +| -[1+n] := by { rw zpow_neg_succ_of_nat, exact inv_ne_zero (pow_ne_zero _ ha) } -lemma fpow_sub {a : G₀} (ha : a ≠ 0) (z1 z2 : ℤ) : a ^ (z1 - z2) = a ^ z1 / a ^ z2 := -by rw [sub_eq_add_neg, fpow_add ha, fpow_neg, div_eq_mul_inv] +lemma zpow_sub₀ {a : G₀} (ha : a ≠ 0) (z1 z2 : ℤ) : a ^ (z1 - z2) = a ^ z1 / a ^ z2 := +by rw [sub_eq_add_neg, zpow_add₀ ha, zpow_neg₀, div_eq_mul_inv] -lemma commute.mul_fpow {a b : G₀} (h : commute a b) : +lemma commute.mul_zpow₀ {a b : G₀} (h : commute a b) : ∀ (i : ℤ), (a * b) ^ i = (a ^ i) * (b ^ i) | (n : ℕ) := by simp [h.mul_pow n] | -[1+n] := by simp [h.mul_pow, (h.pow_pow _ _).eq, mul_inv_rev₀] -lemma mul_fpow {G₀ : Type*} [comm_group_with_zero G₀] (a b : G₀) (m : ℤ): +lemma mul_zpow₀ {G₀ : Type*} [comm_group_with_zero G₀] (a b : G₀) (m : ℤ): (a * b) ^ m = (a ^ m) * (b ^ m) := -(commute.all a b).mul_fpow m +(commute.all a b).mul_zpow₀ m -theorem fpow_bit0' (a : G₀) (n : ℤ) : a ^ bit0 n = (a * a) ^ n := -(fpow_bit0 a n).trans ((commute.refl a).mul_fpow n).symm +theorem zpow_bit0' (a : G₀) (n : ℤ) : a ^ bit0 n = (a * a) ^ n := +(zpow_bit0₀ a n).trans ((commute.refl a).mul_zpow₀ n).symm -theorem fpow_bit1' (a : G₀) (n : ℤ) : a ^ bit1 n = (a * a) ^ n * a := -by rw [fpow_bit1, (commute.refl a).mul_fpow] +theorem zpow_bit1' (a : G₀) (n : ℤ) : a ^ bit1 n = (a * a) ^ n * a := +by rw [zpow_bit1₀, (commute.refl a).mul_zpow₀] -lemma fpow_eq_zero {x : G₀} {n : ℤ} (h : x ^ n = 0) : x = 0 := -classical.by_contradiction $ λ hx, fpow_ne_zero_of_ne_zero hx n h +lemma zpow_eq_zero {x : G₀} {n : ℤ} (h : x ^ n = 0) : x = 0 := +classical.by_contradiction $ λ hx, zpow_ne_zero_of_ne_zero hx n h -lemma fpow_ne_zero {x : G₀} (n : ℤ) : x ≠ 0 → x ^ n ≠ 0 := -mt fpow_eq_zero +lemma zpow_ne_zero {x : G₀} (n : ℤ) : x ≠ 0 → x ^ n ≠ 0 := +mt zpow_eq_zero -theorem fpow_neg_mul_fpow_self (n : ℤ) {x : G₀} (h : x ≠ 0) : +theorem zpow_neg_mul_zpow_self (n : ℤ) {x : G₀} (h : x ≠ 0) : x ^ (-n) * x ^ n = 1 := begin - rw [fpow_neg], - exact inv_mul_cancel (fpow_ne_zero n h) + rw [zpow_neg₀], + exact inv_mul_cancel (zpow_ne_zero n h) end theorem one_div_pow {a : G₀} (n : ℕ) : (1 / a) ^ n = 1 / a ^ n := by simp only [one_div, inv_pow₀] -theorem one_div_fpow {a : G₀} (n : ℤ) : +theorem one_div_zpow {a : G₀} (n : ℤ) : (1 / a) ^ n = 1 / a ^ n := -by simp only [one_div, inv_fpow] +by simp only [one_div, inv_zpow₀] -@[simp] lemma inv_fpow' {a : G₀} (n : ℤ) : +@[simp] lemma inv_zpow' {a : G₀} (n : ℤ) : (a ⁻¹) ^ n = a ^ (-n) := -by { rw [inv_fpow, ← fpow_neg_one, ← fpow_mul], simp } +by { rw [inv_zpow₀, ← zpow_neg_one₀, ← zpow_mul₀], simp } end int_pow @@ -231,9 +231,9 @@ variables {G₀ : Type*} [comm_group_with_zero G₀] (a / b) ^ n = a ^ n / b ^ n := by simp only [div_eq_mul_inv, mul_pow, inv_pow₀] -@[simp] theorem div_fpow (a : G₀) {b : G₀} (n : ℤ) : +@[simp] theorem div_zpow₀ (a : G₀) {b : G₀} (n : ℤ) : (a / b) ^ n = a ^ n / b ^ n := -by simp only [div_eq_mul_inv, mul_fpow, inv_fpow] +by simp only [div_eq_mul_inv, mul_zpow₀, inv_zpow₀] lemma div_sq_cancel {a : G₀} (ha : a ≠ 0) (b : G₀) : a ^ 2 * b / a = a * b := by rw [sq, mul_assoc, mul_div_cancel_left _ ha] @@ -242,26 +242,26 @@ end /-- If a monoid homomorphism `f` between two `group_with_zero`s maps `0` to `0`, then it maps `x^n`, `n : ℤ`, to `(f x)^n`. -/ -lemma monoid_with_zero_hom.map_fpow {G₀ G₀' : Type*} [group_with_zero G₀] [group_with_zero G₀'] +lemma monoid_with_zero_hom.map_zpow {G₀ G₀' : Type*} [group_with_zero G₀] [group_with_zero G₀'] (f : monoid_with_zero_hom G₀ G₀') (x : G₀) : ∀ n : ℤ, f (x ^ n) = f x ^ n -| (n : ℕ) := by { rw [gpow_coe_nat, gpow_coe_nat], exact f.to_monoid_hom.map_pow x n } +| (n : ℕ) := by { rw [zpow_coe_nat, zpow_coe_nat], exact f.to_monoid_hom.map_pow x n } | -[1+n] := begin - rw [gpow_neg_succ_of_nat, gpow_neg_succ_of_nat], + rw [zpow_neg_succ_of_nat, zpow_neg_succ_of_nat], exact ((f.map_inv _).trans $ congr_arg _ $ f.to_monoid_hom.map_pow x _) end -- I haven't been able to find a better home for this: -- it belongs with other lemmas on `linear_ordered_field`, but --- we need to wait until `fpow` has been defined in this file. +-- we need to wait until `zpow` has been defined in this file. section variables {R : Type*} [linear_ordered_field R] {a : R} lemma pow_minus_two_nonneg : 0 ≤ a^(-2 : ℤ) := begin - simp only [inv_nonneg, fpow_neg], + simp only [inv_nonneg, zpow_neg₀], change 0 ≤ a ^ ((2 : ℕ) : ℤ), - rw gpow_coe_nat, + rw zpow_coe_nat, apply sq_nonneg, end diff --git a/src/algebra/iterate_hom.lean b/src/algebra/iterate_hom.lean index 4e39f04e99897..6218c481c0c14 100644 --- a/src/algebra/iterate_hom.lean +++ b/src/algebra/iterate_hom.lean @@ -62,8 +62,8 @@ commute.iterate_left f.map_inv n x theorem iterate_map_pow (f : M →* M) (a) (n m : ℕ) : f^[n] (a^m) = (f^[n] a)^m := commute.iterate_left (λ x, f.map_pow x m) n a -theorem iterate_map_gpow (f : G →* G) (a) (n : ℕ) (m : ℤ) : f^[n] (a^m) = (f^[n] a)^m := -commute.iterate_left (λ x, f.map_gpow x m) n a +theorem iterate_map_zpow (f : G →* G) (a) (n : ℕ) (m : ℤ) : f^[n] (a^m) = (f^[n] a)^m := +commute.iterate_left (λ x, f.map_zpow x m) n a lemma coe_pow {M} [comm_monoid M] (f : monoid.End M) (n : ℕ) : ⇑(f^n) = (f^[n]) := hom_coe_pow _ rfl (λ f g, rfl) _ _ @@ -85,7 +85,7 @@ f.to_multiplicative.iterate_map_pow x n m theorem iterate_map_gsmul (f : G →+ G) (n : ℕ) (m : ℤ) (x : G) : f^[n] (m • x) = m • (f^[n] x) := -f.to_multiplicative.iterate_map_gpow x n m +f.to_multiplicative.iterate_map_zpow x n m end add_monoid_hom diff --git a/src/algebra/opposites.lean b/src/algebra/opposites.lean index 9948bcb52b086..f442786ee49e3 100644 --- a/src/algebra/opposites.lean +++ b/src/algebra/opposites.lean @@ -189,15 +189,15 @@ instance [cancel_comm_monoid α] : cancel_comm_monoid (opposite α) := { .. opposite.cancel_monoid α, ..opposite.comm_monoid α } instance [div_inv_monoid α] : div_inv_monoid (opposite α) := -{ gpow := λ n x, op $ div_inv_monoid.gpow n x.unop, - gpow_zero' := λ x, unop_injective $ div_inv_monoid.gpow_zero' x.unop, - gpow_succ' := λ n x, unop_injective $ (div_inv_monoid.gpow_succ' n x.unop).trans begin +{ zpow := λ n x, op $ div_inv_monoid.zpow n x.unop, + zpow_zero' := λ x, unop_injective $ div_inv_monoid.zpow_zero' x.unop, + zpow_succ' := λ n x, unop_injective $ (div_inv_monoid.zpow_succ' n x.unop).trans begin dsimp, induction n with n ih, - { rw [int.of_nat_zero, div_inv_monoid.gpow_zero', one_mul, mul_one] }, - { rw [div_inv_monoid.gpow_succ' n x.unop, mul_assoc, ih], }, + { rw [int.of_nat_zero, div_inv_monoid.zpow_zero', one_mul, mul_one] }, + { rw [div_inv_monoid.zpow_succ' n x.unop, mul_assoc, ih], }, end, - gpow_neg' := λ z x, unop_injective $ div_inv_monoid.gpow_neg' z x.unop, + zpow_neg' := λ z x, unop_injective $ div_inv_monoid.zpow_neg' z x.unop, .. opposite.monoid α, .. opposite.has_inv α } instance [group α] : group (opposite α) := diff --git a/src/algebra/order/archimedean.lean b/src/algebra/order/archimedean.lean index d408f5dc9ae6b..6f2ed8b756022 100644 --- a/src/algebra/order/archimedean.lean +++ b/src/algebra/order/archimedean.lean @@ -156,13 +156,13 @@ lemma exists_int_pow_near [archimedean α] by classical; exact let ⟨N, hN⟩ := pow_unbounded_of_one_lt x⁻¹ hy in have he: ∃ m : ℤ, y ^ m ≤ x, from - ⟨-N, le_of_lt (by { rw [fpow_neg y (↑N), gpow_coe_nat], + ⟨-N, le_of_lt (by { rw [zpow_neg₀ y (↑N), zpow_coe_nat], exact (inv_lt hx (lt_trans (inv_pos.2 hx) hN)).1 hN })⟩, let ⟨M, hM⟩ := pow_unbounded_of_one_lt x hy in have hb: ∃ b : ℤ, ∀ m, y ^ m ≤ x → m ≤ b, from ⟨M, λ m hm, le_of_not_lt (λ hlt, not_lt_of_ge - (fpow_le_of_le hy.le hlt.le) - (lt_of_le_of_lt hm (by rwa ← gpow_coe_nat at hM)))⟩, + (zpow_le_of_le hy.le hlt.le) + (lt_of_le_of_lt hm (by rwa ← zpow_coe_nat at hM)))⟩, let ⟨n, hn₁, hn₂⟩ := int.exists_greatest_of_bdd hb he in ⟨n, hn₁, lt_of_not_ge (λ hge, not_le_of_gt (int.lt_succ _) (hn₂ _ hge))⟩ @@ -175,9 +175,9 @@ lemma exists_int_pow_near' [archimedean α] let ⟨m, hle, hlt⟩ := exists_int_pow_near (inv_pos.2 hx) hy in have hyp : 0 < y, from lt_trans zero_lt_one hy, ⟨-(m+1), -by rwa [fpow_neg, inv_lt (fpow_pos_of_pos hyp _) hx], -by rwa [neg_add, neg_add_cancel_right, fpow_neg, - le_inv hx (fpow_pos_of_pos hyp _)]⟩ +by rwa [zpow_neg₀, inv_lt (zpow_pos_of_pos hyp _) hx], +by rwa [neg_add, neg_add_cancel_right, zpow_neg₀, + le_inv hx (zpow_pos_of_pos hyp _)]⟩ /-- For any `y < 1` and any positive `x`, there exists `n : ℕ` with `y ^ n < x`. -/ lemma exists_pow_lt_of_lt_one [archimedean α] {x y : α} (hx : 0 < x) (hy : y < 1) : diff --git a/src/algebra/punit_instances.lean b/src/algebra/punit_instances.lean index a293d5bc2eb52..1f736e393ea26 100644 --- a/src/algebra/punit_instances.lean +++ b/src/algebra/punit_instances.lean @@ -26,7 +26,7 @@ by refine_struct inv := λ _, star, div := λ _ _, star, npow := λ _ _, star, - gpow := λ _ _, star, + zpow := λ _ _, star, .. }; intros; exact subsingleton.elim _ _ diff --git a/src/algebra/star/basic.lean b/src/algebra/star/basic.lean index 6eb0d75ed515f..9a5fda012017b 100644 --- a/src/algebra/star/basic.lean +++ b/src/algebra/star/basic.lean @@ -106,9 +106,9 @@ op_injective $ op_injective $ ((star_mul_equiv : R ≃* Rᵒᵖ).to_monoid_hom.map_inv x).trans (op_inv (star x)).symm -@[simp] lemma star_gpow [group R] [star_monoid R] (x : R) (z : ℤ) : star (x ^ z) = star x ^ z := +@[simp] lemma star_zpow [group R] [star_monoid R] (x : R) (z : ℤ) : star (x ^ z) = star x ^ z := op_injective $ - ((star_mul_equiv : R ≃* Rᵒᵖ).to_monoid_hom.map_gpow x z).trans (op_gpow (star x) z).symm + ((star_mul_equiv : R ≃* Rᵒᵖ).to_monoid_hom.map_zpow x z).trans (op_zpow (star x) z).symm /-- When multiplication is commutative, `star` preserves division. -/ @[simp] lemma star_div [comm_group R] [star_monoid R] (x y : R) : star (x / y) = star x / star y := @@ -237,10 +237,10 @@ alias star_ring_aut_self_apply ← is_R_or_C.conj_conj op_injective $ ((star_ring_equiv : R ≃+* Rᵒᵖ).to_ring_hom.map_inv x).trans (op_inv (star x)).symm -@[simp] lemma star_fpow [division_ring R] [star_ring R] (x : R) (z : ℤ) : +@[simp] lemma star_zpow₀ [division_ring R] [star_ring R] (x : R) (z : ℤ) : star (x ^ z) = star x ^ z := op_injective $ - ((star_ring_equiv : R ≃+* Rᵒᵖ).to_ring_hom.map_fpow x z).trans (op_gpow (star x) z).symm + ((star_ring_equiv : R ≃+* Rᵒᵖ).to_ring_hom.map_zpow x z).trans (op_zpow (star x) z).symm /-- When multiplication is commutative, `star` preserves division. -/ @[simp] lemma star_div' [field R] [star_ring R] (x y : R) : star (x / y) = star x / star y := diff --git a/src/analysis/asymptotics/specific_asymptotics.lean b/src/analysis/asymptotics/specific_asymptotics.lean index 3d5227187c385..401c5baad85d6 100644 --- a/src/analysis/asymptotics/specific_asymptotics.lean +++ b/src/analysis/asymptotics/specific_asymptotics.lean @@ -24,21 +24,21 @@ lemma pow_div_pow_eventually_eq_at_top {p q : ℕ} : (λ x : 𝕜, x^p / x^q) =ᶠ[at_top] (λ x, x^((p : ℤ) -q)) := begin apply ((eventually_gt_at_top (0 : 𝕜)).mono (λ x hx, _)), - simp [fpow_sub hx.ne'], + simp [zpow_sub₀ hx.ne'], end lemma pow_div_pow_eventually_eq_at_bot {p q : ℕ} : (λ x : 𝕜, x^p / x^q) =ᶠ[at_bot] (λ x, x^((p : ℤ) -q)) := begin apply ((eventually_lt_at_bot (0 : 𝕜)).mono (λ x hx, _)), - simp [fpow_sub hx.ne'.symm], + simp [zpow_sub₀ hx.ne'.symm], end -lemma tendsto_fpow_at_top_at_top {n : ℤ} +lemma tendsto_zpow_at_top_at_top {n : ℤ} (hn : 0 < n) : tendsto (λ x : 𝕜, x^n) at_top at_top := begin lift n to ℕ using hn.le, - simp only [gpow_coe_nat], + simp only [zpow_coe_nat], exact tendsto_pow_at_top (nat.succ_le_iff.mpr $int.coe_nat_pos.mp hn) end @@ -46,7 +46,7 @@ lemma tendsto_pow_div_pow_at_top_at_top {p q : ℕ} (hpq : q < p) : tendsto (λ x : 𝕜, x^p / x^q) at_top at_top := begin rw tendsto_congr' pow_div_pow_eventually_eq_at_top, - apply tendsto_fpow_at_top_at_top, + apply tendsto_zpow_at_top_at_top, linarith end @@ -54,7 +54,7 @@ lemma tendsto_pow_div_pow_at_top_zero [topological_space 𝕜] [order_topology (hpq : p < q) : tendsto (λ x : 𝕜, x^p / x^q) at_top (𝓝 0) := begin rw tendsto_congr' pow_div_pow_eventually_eq_at_top, - apply tendsto_fpow_at_top_zero, + apply tendsto_zpow_at_top_zero, linarith end diff --git a/src/analysis/asymptotics/superpolynomial_decay.lean b/src/analysis/asymptotics/superpolynomial_decay.lean index 16cbc65af224a..ef96c64ac9283 100644 --- a/src/analysis/asymptotics/superpolynomial_decay.lean +++ b/src/analysis/asymptotics/superpolynomial_decay.lean @@ -72,7 +72,7 @@ begin split; intros h c; specialize h (-c), { simpa [div_eq_mul_inv] using div_is_bounded_under_of_is_O h }, { refine (is_O_iff_div_is_bounded_under _).2 _, - { exact hα.mono (λ x hx hx', absurd (fpow_eq_zero hx') hx) }, + { exact hα.mono (λ x hx hx', absurd (zpow_eq_zero hx') hx) }, { simpa [div_eq_mul_inv] using h } } end @@ -92,7 +92,7 @@ begin have := this.mul_is_O (h $ c - 1), simp only [one_mul] at this, refine this.trans_is_O (is_O.of_bound 1 (hα'.mono (λ x hx, le_of_eq _))), - rw [fpow_sub_one hx, mul_comm, mul_assoc, inv_mul_cancel hx, one_mul, mul_one] + rw [zpow_sub_one₀ hx, mul_comm, mul_assoc, inv_mul_cancel hx, one_mul, mul_one] end theorem superpolynomial_decay_iff_norm_tendsto_zero (f : α → 𝕜) @@ -184,12 +184,12 @@ begin haveI : nontrivial α := (algebra_map α 𝕜).domain_nontrivial, refine λ c, (is_O.mul (is_O_refl (algebra_map α 𝕜) at_top) (hf (c - 1))).trans _, refine is_O_of_div_tendsto_nhds (eventually_of_forall - (λ x hx, mul_eq_zero_of_left (fpow_eq_zero hx) _)) 1 (tendsto_nhds.2 _), + (λ x hx, mul_eq_zero_of_left (zpow_eq_zero hx) _)) 1 (tendsto_nhds.2 _), refine λ s hs hs', at_top.sets_of_superset (mem_at_top 1) (λ x hx, set.mem_preimage.2 _), have hx' : algebra_map α 𝕜 x ≠ 0 := λ hx', (ne_of_lt $ lt_of_lt_of_le zero_lt_one hx).symm (by simpa [algebra.algebra_map_eq_smul_one, smul_eq_zero] using hx'), convert hs', - rw [pi.div_apply, div_eq_one_iff_eq (fpow_ne_zero c hx'), fpow_sub_one hx' c, + rw [pi.div_apply, div_eq_one_iff_eq (zpow_ne_zero c hx'), zpow_sub_one₀ hx' c, mul_comm (algebra_map α 𝕜 x), mul_assoc, inv_mul_cancel hx', mul_one], end @@ -252,19 +252,19 @@ begin { exact hC c hc }, { refine (hC C le_rfl).trans (is_O.of_bound 1 (_)), refine at_top.sets_of_superset hα (λ x hx, _), - simp only [one_mul, normed_field.norm_fpow, set.mem_set_of_eq], - exact fpow_le_of_le hx (le_of_not_le hc) } + simp only [one_mul, normed_field.norm_zpow, set.mem_set_of_eq], + exact zpow_le_of_le hx (le_of_not_le hc) } end -lemma superpolynomial_decay_of_is_O_fpow_le (hα : ∀ᶠ (x : α) in at_top, 1 ≤ ∥algebra_map α 𝕜 x∥) +lemma superpolynomial_decay_of_is_O_zpow_le (hα : ∀ᶠ (x : α) in at_top, 1 ≤ ∥algebra_map α 𝕜 x∥) (C : ℤ) (h : ∀ c ≤ C, is_O f (λ n, (algebra_map α 𝕜 n) ^ c) at_top) : superpolynomial_decay f := superpolynomial_decay_of_eventually_is_O hα (eventually_at_bot.2 ⟨C, h⟩) -lemma superpolynomial_decay_of_is_O_fpow_lt (hα : ∀ᶠ (x : α) in at_top, 1 ≤ ∥algebra_map α 𝕜 x∥) +lemma superpolynomial_decay_of_is_O_zpow_lt (hα : ∀ᶠ (x : α) in at_top, 1 ≤ ∥algebra_map α 𝕜 x∥) (C : ℤ) (h : ∀ c < C, is_O f (λ n, (algebra_map α 𝕜 n) ^ c) at_top) : superpolynomial_decay f := -superpolynomial_decay_of_is_O_fpow_le hα C.pred +superpolynomial_decay_of_is_O_zpow_le hα C.pred (λ c hc, h c (lt_of_le_of_lt hc (int.pred_self_lt C))) section order_topology @@ -279,7 +279,7 @@ begin refine is_O.trans_tendsto (hf (-1)) _, have : (has_inv.inv : 𝕜 → 𝕜) ∘ (algebra_map α 𝕜 : α → 𝕜) = (λ (n : α), (algebra_map α 𝕜 n) ^ (-1 : ℤ)), - by simp only [gpow_one, fpow_neg], + by simp only [zpow_one, zpow_neg₀], exact this ▸ (tendsto_inv_at_top_zero).comp (hα) end diff --git a/src/analysis/calculus/deriv.lean b/src/analysis/calculus/deriv.lean index 0dd18d0c7a51d..7c2e5f4fc2c69 100644 --- a/src/analysis/calculus/deriv.lean +++ b/src/analysis/calculus/deriv.lean @@ -1876,88 +1876,88 @@ hc.has_deriv_at.pow.deriv end pow -section fpow +section zpow /-! ### Derivative of `x ↦ x^m` for `m : ℤ` -/ variables {x : 𝕜} {s : set 𝕜} {m : ℤ} -lemma has_strict_deriv_at_fpow (m : ℤ) (x : 𝕜) (h : x ≠ 0 ∨ 0 ≤ m) : +lemma has_strict_deriv_at_zpow (m : ℤ) (x : 𝕜) (h : x ≠ 0 ∨ 0 ≤ m) : has_strict_deriv_at (λx, x^m) ((m : 𝕜) * x^(m-1)) x := begin have : ∀ m : ℤ, 0 < m → has_strict_deriv_at (λx, x^m) ((m:𝕜) * x^(m-1)) x, { assume m hm, lift m to ℕ using (le_of_lt hm), - simp only [gpow_coe_nat, int.cast_coe_nat], + simp only [zpow_coe_nat, int.cast_coe_nat], convert has_strict_deriv_at_pow _ _ using 2, - rw [← int.coe_nat_one, ← int.coe_nat_sub, gpow_coe_nat], + rw [← int.coe_nat_one, ← int.coe_nat_sub, zpow_coe_nat], norm_cast at hm, exact nat.succ_le_of_lt hm }, rcases lt_trichotomy m 0 with hm|hm|hm, { have hx : x ≠ 0, from h.resolve_right hm.not_le, have := (has_strict_deriv_at_inv _).scomp _ (this (-m) (neg_pos.2 hm)); - [skip, exact fpow_ne_zero_of_ne_zero hx _], - simp only [(∘), fpow_neg, one_div, inv_inv₀, smul_eq_mul] at this, + [skip, exact zpow_ne_zero_of_ne_zero hx _], + simp only [(∘), zpow_neg₀, one_div, inv_inv₀, smul_eq_mul] at this, convert this using 1, rw [sq, mul_inv₀, inv_inv₀, int.cast_neg, ← neg_mul_eq_neg_mul, neg_mul_neg, - ← fpow_add hx, mul_assoc, ← fpow_add hx], congr, abel }, - { simp only [hm, gpow_zero, int.cast_zero, zero_mul, has_strict_deriv_at_const] }, + ← zpow_add₀ hx, mul_assoc, ← zpow_add₀ hx], congr, abel }, + { simp only [hm, zpow_zero, int.cast_zero, zero_mul, has_strict_deriv_at_const] }, { exact this m hm } end -lemma has_deriv_at_fpow (m : ℤ) (x : 𝕜) (h : x ≠ 0 ∨ 0 ≤ m) : +lemma has_deriv_at_zpow (m : ℤ) (x : 𝕜) (h : x ≠ 0 ∨ 0 ≤ m) : has_deriv_at (λx, x^m) ((m : 𝕜) * x^(m-1)) x := -(has_strict_deriv_at_fpow m x h).has_deriv_at +(has_strict_deriv_at_zpow m x h).has_deriv_at -theorem has_deriv_within_at_fpow (m : ℤ) (x : 𝕜) (h : x ≠ 0 ∨ 0 ≤ m) (s : set 𝕜) : +theorem has_deriv_within_at_zpow (m : ℤ) (x : 𝕜) (h : x ≠ 0 ∨ 0 ≤ m) (s : set 𝕜) : has_deriv_within_at (λx, x^m) ((m : 𝕜) * x^(m-1)) s x := -(has_deriv_at_fpow m x h).has_deriv_within_at +(has_deriv_at_zpow m x h).has_deriv_within_at -lemma differentiable_at_fpow : differentiable_at 𝕜 (λx, x^m) x ↔ x ≠ 0 ∨ 0 ≤ m := -⟨λ H, normed_field.continuous_at_fpow.1 H.continuous_at, - λ H, (has_deriv_at_fpow m x H).differentiable_at⟩ +lemma differentiable_at_zpow : differentiable_at 𝕜 (λx, x^m) x ↔ x ≠ 0 ∨ 0 ≤ m := +⟨λ H, normed_field.continuous_at_zpow.1 H.continuous_at, + λ H, (has_deriv_at_zpow m x H).differentiable_at⟩ -lemma differentiable_within_at_fpow (m : ℤ) (x : 𝕜) (h : x ≠ 0 ∨ 0 ≤ m) : +lemma differentiable_within_at_zpow (m : ℤ) (x : 𝕜) (h : x ≠ 0 ∨ 0 ≤ m) : differentiable_within_at 𝕜 (λx, x^m) s x := -(differentiable_at_fpow.mpr h).differentiable_within_at +(differentiable_at_zpow.mpr h).differentiable_within_at -lemma differentiable_on_fpow (m : ℤ) (s : set 𝕜) (h : (0 : 𝕜) ∉ s ∨ 0 ≤ m) : +lemma differentiable_on_zpow (m : ℤ) (s : set 𝕜) (h : (0 : 𝕜) ∉ s ∨ 0 ≤ m) : differentiable_on 𝕜 (λx, x^m) s := -λ x hxs, differentiable_within_at_fpow m x $ h.imp_left $ ne_of_mem_of_not_mem hxs +λ x hxs, differentiable_within_at_zpow m x $ h.imp_left $ ne_of_mem_of_not_mem hxs -lemma deriv_fpow (m : ℤ) (x : 𝕜) : deriv (λ x, x ^ m) x = m * x ^ (m - 1) := +lemma deriv_zpow (m : ℤ) (x : 𝕜) : deriv (λ x, x ^ m) x = m * x ^ (m - 1) := begin by_cases H : x ≠ 0 ∨ 0 ≤ m, - { exact (has_deriv_at_fpow m x H).deriv }, - { rw deriv_zero_of_not_differentiable_at (mt differentiable_at_fpow.1 H), + { exact (has_deriv_at_zpow m x H).deriv }, + { rw deriv_zero_of_not_differentiable_at (mt differentiable_at_zpow.1 H), push_neg at H, rcases H with ⟨rfl, hm⟩, - rw [zero_fpow _ ((sub_one_lt _).trans hm).ne, mul_zero] } + rw [zero_zpow _ ((sub_one_lt _).trans hm).ne, mul_zero] } end -@[simp] lemma deriv_fpow' (m : ℤ) : deriv (λ x : 𝕜, x ^ m) = λ x, m * x ^ (m - 1) := -funext $ deriv_fpow m +@[simp] lemma deriv_zpow' (m : ℤ) : deriv (λ x : 𝕜, x ^ m) = λ x, m * x ^ (m - 1) := +funext $ deriv_zpow m -lemma deriv_within_fpow (hxs : unique_diff_within_at 𝕜 s x) (h : x ≠ 0 ∨ 0 ≤ m) : +lemma deriv_within_zpow (hxs : unique_diff_within_at 𝕜 s x) (h : x ≠ 0 ∨ 0 ≤ m) : deriv_within (λx, x^m) s x = (m : 𝕜) * x^(m-1) := -(has_deriv_within_at_fpow m x h s).deriv_within hxs +(has_deriv_within_at_zpow m x h s).deriv_within hxs -@[simp] lemma iter_deriv_fpow' (m : ℤ) (k : ℕ) : +@[simp] lemma iter_deriv_zpow' (m : ℤ) (k : ℕ) : deriv^[k] (λ x : 𝕜, x ^ m) = λ x, (∏ i in finset.range k, (m - i)) * x ^ (m - k) := begin induction k with k ihk, { simp only [one_mul, int.coe_nat_zero, id, sub_zero, finset.prod_range_zero, function.iterate_zero] }, - { simp only [function.iterate_succ_apply', ihk, deriv_const_mul_field', deriv_fpow', + { simp only [function.iterate_succ_apply', ihk, deriv_const_mul_field', deriv_zpow', finset.prod_range_succ, int.coe_nat_succ, ← sub_sub, int.cast_sub, int.cast_coe_nat, mul_assoc], } end -lemma iter_deriv_fpow (m : ℤ) (x : 𝕜) (k : ℕ) : +lemma iter_deriv_zpow (m : ℤ) (x : 𝕜) (k : ℕ) : deriv^[k] (λ y, y ^ m) x = (∏ i in finset.range k, (m - i)) * x ^ (m - k) := -congr_fun (iter_deriv_fpow' m k) x +congr_fun (iter_deriv_zpow' m k) x lemma iter_deriv_pow (n : ℕ) (x : 𝕜) (k : ℕ) : deriv^[k] (λx:𝕜, x^n) x = (∏ i in finset.range k, (n - i)) * x^(n-k) := begin - simp only [← gpow_coe_nat, iter_deriv_fpow, int.cast_coe_nat], + simp only [← zpow_coe_nat, iter_deriv_zpow, int.cast_coe_nat], cases le_or_lt k n with hkn hnk, { rw int.coe_nat_sub hkn }, { have : ∏ i in finset.range k, (n - i : 𝕜) = 0, @@ -1971,13 +1971,13 @@ funext $ λ x, iter_deriv_pow n x k lemma iter_deriv_inv (k : ℕ) (x : 𝕜) : deriv^[k] has_inv.inv x = (∏ i in finset.range k, (-1 - i)) * x ^ (-1 - k : ℤ) := -by simpa only [fpow_neg_one, int.cast_neg, int.cast_one] using iter_deriv_fpow (-1) x k +by simpa only [zpow_neg_one₀, int.cast_neg, int.cast_one] using iter_deriv_zpow (-1) x k @[simp] lemma iter_deriv_inv' (k : ℕ) : deriv^[k] has_inv.inv = λ x : 𝕜, (∏ i in finset.range k, (-1 - i)) * x ^ (-1 - k : ℤ) := funext (iter_deriv_inv k) -end fpow +end zpow /-! ### Upper estimates on liminf and limsup -/ diff --git a/src/analysis/convex/specific_functions.lean b/src/analysis/convex/specific_functions.lean index 3bca1de1375c8..c281e0383dbc4 100644 --- a/src/analysis/convex/specific_functions.lean +++ b/src/analysis/convex/specific_functions.lean @@ -15,7 +15,7 @@ In this file we prove that the following functions are convex: * `convex_on_pow_of_even` : given an even natural number $n$, the function $f(x)=x^n$ is convex on $(-∞, +∞)$; * `convex_on_pow` : for a natural $n$, the function $f(x)=x^n$ is convex on $[0, +∞)$; -* `convex_on_fpow` : for an integer $m$, the function $f(x)=x^m$ is convex on $(0, +∞)$. +* `convex_on_zpow` : for an integer $m$, the function $f(x)=x^m$ is convex on $(0, +∞)$. * `convex_on_rpow : ∀ p : ℝ, 1 ≤ p → convex_on ℝ (Ici 0) (λ x, x ^ p)` * `concave_on_log_Ioi` and `concave_on_log_Iio`: log is concave on `Ioi 0` and `Iio 0` respectively. -/ @@ -76,18 +76,18 @@ begin end /-- `x^m`, `m : ℤ` is convex on `(0, +∞)` for all `m` -/ -lemma convex_on_fpow (m : ℤ) : convex_on ℝ (Ioi 0) (λ x : ℝ, x^m) := +lemma convex_on_zpow (m : ℤ) : convex_on ℝ (Ioi 0) (λ x : ℝ, x^m) := begin have : ∀ n : ℤ, differentiable_on ℝ (λ x, x ^ n) (Ioi (0 : ℝ)), - from λ n, differentiable_on_fpow _ _ (or.inl $ lt_irrefl _), + from λ n, differentiable_on_zpow _ _ (or.inl $ lt_irrefl _), apply convex_on_of_deriv2_nonneg (convex_Ioi 0); - try { simp only [interior_Ioi, deriv_fpow'] }, + try { simp only [interior_Ioi, deriv_zpow'] }, { exact (this _).continuous_on }, { exact this _ }, { exact (this _).const_mul _ }, { intros x hx, - simp only [iter_deriv_fpow, ← int.cast_coe_nat, ← int.cast_sub, ← int.cast_prod], - refine mul_nonneg (int.cast_nonneg.2 _) (fpow_nonneg (le_of_lt hx) _), + simp only [iter_deriv_zpow, ← int.cast_coe_nat, ← int.cast_sub, ← int.cast_prod], + refine mul_nonneg (int.cast_nonneg.2 _) (zpow_nonneg (le_of_lt hx) _), exact int_prod_range_nonneg _ _ (nat.even_bit0 1) } end diff --git a/src/analysis/fourier.lean b/src/analysis/fourier.lean index 3ef9081a630d5..2738621452e30 100644 --- a/src/analysis/fourier.lean +++ b/src/analysis/fourier.lean @@ -87,7 +87,7 @@ section fourier continuous maps from `circle` to `ℂ`. -/ @[simps] def fourier (n : ℤ) : C(circle, ℂ) := { to_fun := λ z, z ^ n, - continuous_to_fun := continuous_subtype_coe.fpow n $ λ z, or.inl (nonzero_of_mem_circle z) } + continuous_to_fun := continuous_subtype_coe.zpow n $ λ z, or.inl (nonzero_of_mem_circle z) } @[simp] lemma fourier_zero {z : circle} : fourier 0 z = 1 := rfl @@ -96,7 +96,7 @@ by simp [← coe_inv_circle_eq_conj z] @[simp] lemma fourier_add {m n : ℤ} {z : circle} : fourier (m + n) z = (fourier m z) * (fourier n z) := -by simp [fpow_add (nonzero_of_mem_circle z)] +by simp [zpow_add₀ (nonzero_of_mem_circle z)] /-- The subalgebra of `C(circle, ℂ)` generated by `z ^ n` for `n ∈ ℤ`; equivalently, polynomials in `z` and `conj z`. -/ @@ -188,7 +188,7 @@ begin { have : (n:ℂ) ≠ 0 := by exact_mod_cast hn, field_simp, ring }, - simp [mul_fpow, ← complex.exp_int_mul, complex.exp_pi_mul_I, this] + simp [mul_zpow₀, ← complex.exp_int_mul, complex.exp_pi_mul_I, this] end /-- The monomials `z ^ n` are an orthonormal set with respect to Haar measure on the circle. -/ diff --git a/src/analysis/mean_inequalities.lean b/src/analysis/mean_inequalities.lean index fc2623641a357..d93311f497ce6 100644 --- a/src/analysis/mean_inequalities.lean +++ b/src/analysis/mean_inequalities.lean @@ -44,7 +44,7 @@ $$ Currently we only prove this inequality for $p=1$. As in the rest of `mathlib`, we provide different theorems for natural exponents (`pow_arith_mean_le_arith_mean_pow`), integer exponents -(`fpow_arith_mean_le_arith_mean_fpow`), and real exponents (`rpow_arith_mean_le_arith_mean_rpow` and +(`zpow_arith_mean_le_arith_mean_zpow`), and real exponents (`rpow_arith_mean_le_arith_mean_rpow` and `arith_mean_le_rpow_mean`). In the first two cases we prove $$ \left(\sum_{i\in s} w_i z_i\right)^n ≤ \sum_{i\in s} w_i z_i^n @@ -148,10 +148,10 @@ theorem pow_arith_mean_le_arith_mean_pow_of_even (w z : ι → ℝ) (hw : ∀ i (∑ i in s, w i * z i) ^ n ≤ ∑ i in s, (w i * z i ^ n) := (convex_on_pow_of_even hn).map_sum_le hw hw' (λ _ _, trivial) -theorem fpow_arith_mean_le_arith_mean_fpow (w z : ι → ℝ) (hw : ∀ i ∈ s, 0 ≤ w i) +theorem zpow_arith_mean_le_arith_mean_zpow (w z : ι → ℝ) (hw : ∀ i ∈ s, 0 ≤ w i) (hw' : ∑ i in s, w i = 1) (hz : ∀ i ∈ s, 0 < z i) (m : ℤ) : (∑ i in s, w i * z i) ^ m ≤ ∑ i in s, (w i * z i ^ m) := -(convex_on_fpow m).map_sum_le hw hw' hz +(convex_on_zpow m).map_sum_le hw hw' hz theorem rpow_arith_mean_le_arith_mean_rpow (w z : ι → ℝ) (hw : ∀ i ∈ s, 0 ≤ w i) (hw' : ∑ i in s, w i = 1) (hz : ∀ i ∈ s, 0 ≤ z i) {p : ℝ} (hp : 1 ≤ p) : diff --git a/src/analysis/normed_space/basic.lean b/src/analysis/normed_space/basic.lean index c44c753ad0b24..c00eac81c89a1 100644 --- a/src/analysis/normed_space/basic.lean +++ b/src/analysis/normed_space/basic.lean @@ -305,11 +305,11 @@ nnreal.eq $ norm_mul a b @[simp] lemma nnnorm_inv (a : α) : ∥a⁻¹∥₊ = ∥a∥₊⁻¹ := nnreal.eq $ by simp -@[simp] lemma norm_fpow : ∀ (a : α) (n : ℤ), ∥a^n∥ = ∥a∥^n := -(norm_hom : monoid_with_zero_hom α ℝ).map_fpow +@[simp] lemma norm_zpow : ∀ (a : α) (n : ℤ), ∥a^n∥ = ∥a∥^n := +(norm_hom : monoid_with_zero_hom α ℝ).map_zpow -@[simp] lemma nnnorm_fpow : ∀ (a : α) (n : ℤ), ∥a ^ n∥₊ = ∥a∥₊ ^ n := -(nnnorm_hom : monoid_with_zero_hom α ℝ≥0).map_fpow +@[simp] lemma nnnorm_zpow : ∀ (a : α) (n : ℤ), ∥a ^ n∥₊ = ∥a∥₊ ^ n := +(nnnorm_hom : monoid_with_zero_hom α ℝ≥0).map_zpow @[priority 100] -- see Note [lower instance priority] instance : has_continuous_inv₀ α := @@ -354,8 +354,8 @@ let ⟨n, hn⟩ := pow_unbounded_of_one_lt r hw in lemma exists_norm_lt {r : ℝ} (hr : 0 < r) : ∃ x : α, 0 < ∥x∥ ∧ ∥x∥ < r := let ⟨w, hw⟩ := exists_one_lt_norm α in let ⟨n, hle, hlt⟩ := exists_int_pow_near' hr hw in -⟨w^n, by { rw norm_fpow; exact fpow_pos_of_pos (lt_trans zero_lt_one hw) _}, -by rwa norm_fpow⟩ +⟨w^n, by { rw norm_zpow; exact zpow_pos_of_pos (lt_trans zero_lt_one hw) _}, +by rwa norm_zpow⟩ variable {α} @@ -719,21 +719,21 @@ begin have xεpos : 0 < ∥x∥/ε := div_pos ((ne.symm hx).le_iff_lt.1 (norm_nonneg x)) εpos, rcases exists_int_pow_near xεpos hc with ⟨n, hn⟩, have cpos : 0 < ∥c∥ := lt_trans (zero_lt_one : (0 :ℝ) < 1) hc, - have cnpos : 0 < ∥c^(n+1)∥ := by { rw norm_fpow, exact lt_trans xεpos hn.2 }, + have cnpos : 0 < ∥c^(n+1)∥ := by { rw norm_zpow, exact lt_trans xεpos hn.2 }, refine ⟨(c^(n+1))⁻¹, _, _, _, _⟩, show (c ^ (n + 1))⁻¹ ≠ 0, by rwa [ne.def, inv_eq_zero, ← ne.def, ← norm_pos_iff], show ∥(c ^ (n + 1))⁻¹ • x∥ < ε, - { rw [norm_smul, norm_inv, ← div_eq_inv_mul, div_lt_iff cnpos, mul_comm, norm_fpow], + { rw [norm_smul, norm_inv, ← div_eq_inv_mul, div_lt_iff cnpos, mul_comm, norm_zpow], exact (div_lt_iff εpos).1 (hn.2) }, show ε / ∥c∥ ≤ ∥(c ^ (n + 1))⁻¹ • x∥, - { rw [div_le_iff cpos, norm_smul, norm_inv, norm_fpow, fpow_add (ne_of_gt cpos), - gpow_one, mul_inv_rev₀, mul_comm, ← mul_assoc, ← mul_assoc, mul_inv_cancel (ne_of_gt cpos), - one_mul, ← div_eq_inv_mul, le_div_iff (fpow_pos_of_pos cpos _), mul_comm], + { rw [div_le_iff cpos, norm_smul, norm_inv, norm_zpow, zpow_add₀ (ne_of_gt cpos), + zpow_one, mul_inv_rev₀, mul_comm, ← mul_assoc, ← mul_assoc, mul_inv_cancel (ne_of_gt cpos), + one_mul, ← div_eq_inv_mul, le_div_iff (zpow_pos_of_pos cpos _), mul_comm], exact (le_div_iff εpos).1 hn.1 }, show ∥(c ^ (n + 1))⁻¹∥⁻¹ ≤ ε⁻¹ * ∥c∥ * ∥x∥, { have : ε⁻¹ * ∥c∥ * ∥x∥ = ε⁻¹ * ∥x∥ * ∥c∥, by ring, - rw [norm_inv, inv_inv₀, norm_fpow, fpow_add (ne_of_gt cpos), gpow_one, this, ← div_eq_inv_mul], + rw [norm_inv, inv_inv₀, norm_zpow, zpow_add₀ (ne_of_gt cpos), zpow_one, this, ← div_eq_inv_mul], exact mul_le_mul_of_nonneg_right hn.1 (norm_nonneg _) } end diff --git a/src/analysis/special_functions/exp.lean b/src/analysis/special_functions/exp.lean index 817d76ba8c6af..c55a8ce2522d4 100644 --- a/src/analysis/special_functions/exp.lean +++ b/src/analysis/special_functions/exp.lean @@ -191,7 +191,7 @@ begin (((tendsto_exp_div_pow_at_top n).const_mul_at_top hb).at_top_add ((tendsto_pow_neg_at_top hn).mul (@tendsto_const_nhds _ _ _ c _))), intros x hx, - simp only [fpow_neg x n], + simp only [zpow_neg₀ x n], ring, end diff --git a/src/analysis/special_functions/polynomials.lean b/src/analysis/special_functions/polynomials.lean index 999114530fa0d..8e9ca78ab5afc 100644 --- a/src/analysis/special_functions/polynomials.lean +++ b/src/analysis/special_functions/polynomials.lean @@ -141,7 +141,7 @@ begin refine (P.is_equivalent_at_top_lead.symm.div Q.is_equivalent_at_top_lead.symm).symm.trans (eventually_eq.is_equivalent ((eventually_gt_at_top 0).mono $ λ x hx, _)), - simp [← div_mul_div, hP, hQ, fpow_sub hx.ne.symm] + simp [← div_mul_div, hP, hQ, zpow_sub₀ hx.ne.symm] end lemma div_tendsto_zero_of_degree_lt (hdeg : P.degree < Q.degree) : @@ -152,7 +152,7 @@ begin rw ← nat_degree_lt_nat_degree_iff hP at hdeg, refine (is_equivalent_at_top_div P Q).symm.tendsto_nhds _, rw ← mul_zero, - refine (tendsto_fpow_at_top_zero _).const_mul _, + refine (tendsto_zpow_at_top_zero _).const_mul _, linarith end @@ -167,7 +167,7 @@ begin exact bot_lt_iff_ne_bot.2 (λ hQ', hQ (degree_eq_bot.1 hQ')) }, { exact absurd (leading_coeff_eq_zero.1 hQ0) hQ } }, { have := (is_equivalent_at_top_div P Q).tendsto_nhds h, - rw tendsto_const_mul_fpow_at_top_zero_iff hPQ at this, + rw tendsto_const_mul_zpow_at_top_zero_iff hPQ at this, cases this with h h, { exact absurd h.2 hPQ }, { rw [sub_lt_iff_lt_add, zero_add, int.coe_nat_lt] at h, @@ -190,7 +190,7 @@ begin rw ← nat_degree_lt_nat_degree_iff hQ at hdeg, refine (is_equivalent_at_top_div P Q).symm.tendsto_at_top _, apply tendsto.const_mul_at_top hpos, - apply tendsto_fpow_at_top_at_top, + apply tendsto_zpow_at_top_at_top, linarith end @@ -211,7 +211,7 @@ begin rw ← nat_degree_lt_nat_degree_iff hQ at hdeg, refine (is_equivalent_at_top_div P Q).symm.tendsto_at_bot _, apply tendsto.neg_const_mul_at_top hneg, - apply tendsto_fpow_at_top_at_top, + apply tendsto_zpow_at_top_at_top, linarith end diff --git a/src/analysis/special_functions/pow.lean b/src/analysis/special_functions/pow.lean index 64d40f77544cf..22efd237b7240 100644 --- a/src/analysis/special_functions/pow.lean +++ b/src/analysis/special_functions/pow.lean @@ -109,7 +109,7 @@ by simpa using cpow_neg x 1 @[simp] lemma cpow_int_cast (x : ℂ) : ∀ (n : ℤ), x ^ (n : ℂ) = x ^ n | (n : ℕ) := by simp; refl -| -[1+ n] := by rw gpow_neg_succ_of_nat; +| -[1+ n] := by rw zpow_neg_succ_of_nat; simp only [int.neg_succ_of_nat_coe, int.cast_neg, complex.cpow_neg, inv_eq_one_div, int.cast_coe_nat, cpow_nat_cast] @@ -551,7 +551,7 @@ by { simp only [sub_eq_add_neg] at h ⊢, simp only [rpow_add' hx h, rpow_neg hx lemma rpow_add_int {x : ℝ} (hx : x ≠ 0) (y : ℝ) (n : ℤ) : x ^ (y + n) = x ^ y * x ^ n := by rw [rpow_def, complex.of_real_add, complex.cpow_add _ _ (complex.of_real_ne_zero.mpr hx), - complex.of_real_int_cast, complex.cpow_int_cast, ← complex.of_real_fpow, mul_comm, + complex.of_real_int_cast, complex.cpow_int_cast, ← complex.of_real_zpow, mul_comm, complex.of_real_mul_re, ← rpow_def, mul_comm] lemma rpow_add_nat {x : ℝ} (hx : x ≠ 0) (y : ℝ) (n : ℕ) : x ^ (y + n) = x ^ y * x ^ n := @@ -570,7 +570,7 @@ lemma rpow_sub_one {x : ℝ} (hx : x ≠ 0) (y : ℝ) : x ^ (y - 1) = x ^ y / x by simpa using rpow_sub_nat hx y 1 @[simp] lemma rpow_int_cast (x : ℝ) (n : ℤ) : x ^ (n : ℝ) = x ^ n := -by simp only [rpow_def, ← complex.of_real_fpow, complex.cpow_int_cast, +by simp only [rpow_def, ← complex.of_real_zpow, complex.cpow_int_cast, complex.of_real_int_cast, complex.of_real_re] @[simp] lemma rpow_nat_cast (x : ℝ) (n : ℕ) : x ^ (n : ℝ) = x ^ n := @@ -579,7 +579,7 @@ rpow_int_cast x n lemma rpow_neg_one (x : ℝ) : x ^ (-1 : ℝ) = x⁻¹ := begin suffices H : x ^ ((-1 : ℤ) : ℝ) = x⁻¹, by exact_mod_cast H, - simp only [rpow_int_cast, gpow_one, fpow_neg], + simp only [rpow_int_cast, zpow_one, zpow_neg₀], end lemma mul_rpow {x y z : ℝ} (h : 0 ≤ x) (h₁ : 0 ≤ y) : (x*y)^z = x^z * y^z := diff --git a/src/analysis/specific_limits.lean b/src/analysis/specific_limits.lean index beba708cba473..088d9031a7010 100644 --- a/src/analysis/specific_limits.lean +++ b/src/analysis/specific_limits.lean @@ -78,28 +78,28 @@ lemma tendsto_norm_inverse_nhds_within_0_at_top {𝕜 : Type*} [normed_field tendsto (λ x:𝕜, ∥x⁻¹∥) (𝓝[{0}ᶜ] 0) at_top := (tendsto_inv_zero_at_top.comp tendsto_norm_zero').congr $ λ x, (normed_field.norm_inv x).symm -lemma tendsto_norm_fpow_nhds_within_0_at_top {𝕜 : Type*} [normed_field 𝕜] {m : ℤ} +lemma tendsto_norm_zpow_nhds_within_0_at_top {𝕜 : Type*} [normed_field 𝕜] {m : ℤ} (hm : m < 0) : tendsto (λ x : 𝕜, ∥x ^ m∥) (𝓝[{0}ᶜ] 0) at_top := begin rcases neg_surjective m with ⟨m, rfl⟩, rw neg_lt_zero at hm, lift m to ℕ using hm.le, rw int.coe_nat_pos at hm, - simp only [normed_field.norm_pow, fpow_neg, gpow_coe_nat, ← inv_pow₀], + simp only [normed_field.norm_pow, zpow_neg₀, zpow_coe_nat, ← inv_pow₀], exact (tendsto_pow_at_top hm).comp normed_field.tendsto_norm_inverse_nhds_within_0_at_top end -@[simp] lemma continuous_at_fpow {𝕜 : Type*} [nondiscrete_normed_field 𝕜] {m : ℤ} {x : 𝕜} : +@[simp] lemma continuous_at_zpow {𝕜 : Type*} [nondiscrete_normed_field 𝕜] {m : ℤ} {x : 𝕜} : continuous_at (λ x, x ^ m) x ↔ x ≠ 0 ∨ 0 ≤ m := begin - refine ⟨_, continuous_at_fpow _ _⟩, + refine ⟨_, continuous_at_zpow _ _⟩, contrapose!, rintro ⟨rfl, hm⟩ hc, exact not_tendsto_at_top_of_tendsto_nhds (hc.tendsto.mono_left nhds_within_le_nhds).norm - (tendsto_norm_fpow_nhds_within_0_at_top hm) + (tendsto_norm_zpow_nhds_within_0_at_top hm) end @[simp] lemma continuous_at_inv {𝕜 : Type*} [nondiscrete_normed_field 𝕜] {x : 𝕜} : continuous_at has_inv.inv x ↔ x ≠ 0 := -by simpa [(@zero_lt_one ℤ _ _).not_le] using @continuous_at_fpow _ _ (-1) x +by simpa [(@zero_lt_one ℤ _ _).not_le] using @continuous_at_zpow _ _ (-1) x end normed_field diff --git a/src/category_theory/conj.lean b/src/category_theory/conj.lean index 6081164c66ed3..a3df2bd9d3853 100644 --- a/src/category_theory/conj.lean +++ b/src/category_theory/conj.lean @@ -114,8 +114,8 @@ conj_Aut_mul α g f @[simp] lemma conj_Aut_pow (f : Aut X) (n : ℕ) : α.conj_Aut (f^n) = (α.conj_Aut f)^n := α.conj_Aut.to_monoid_hom.map_pow f n -@[simp] lemma conj_Aut_gpow (f : Aut X) (n : ℤ) : α.conj_Aut (f^n) = (α.conj_Aut f)^n := -α.conj_Aut.to_monoid_hom.map_gpow f n +@[simp] lemma conj_Aut_zpow (f : Aut X) (n : ℤ) : α.conj_Aut (f^n) = (α.conj_Aut f)^n := +α.conj_Aut.to_monoid_hom.map_zpow f n end iso diff --git a/src/category_theory/endomorphism.lean b/src/category_theory/endomorphism.lean index cf9ad16917d5a..756d8b225e8fc 100644 --- a/src/category_theory/endomorphism.lean +++ b/src/category_theory/endomorphism.lean @@ -90,7 +90,7 @@ by refine_struct mul := flip iso.trans, div := _, npow := @npow_rec (Aut X) ⟨iso.refl X⟩ ⟨flip iso.trans⟩, - gpow := @gpow_rec (Aut X) ⟨iso.refl X⟩ ⟨flip iso.trans⟩ ⟨iso.symm⟩ }; + zpow := @zpow_rec (Aut X) ⟨iso.refl X⟩ ⟨flip iso.trans⟩ ⟨iso.symm⟩ }; intros; try { refl }; ext; simp [flip, (*), monoid.mul, mul_one_class.mul, mul_one_class.one, has_one.one, monoid.one, has_inv.inv] diff --git a/src/data/complex/basic.lean b/src/data/complex/basic.lean index 96f9722ffe706..bbb0b4d1b24ba 100644 --- a/src/data/complex/basic.lean +++ b/src/data/complex/basic.lean @@ -327,11 +327,11 @@ noncomputable instance : field ℂ := inv_zero := complex.inv_zero, ..complex.comm_ring } -@[simp] lemma I_fpow_bit0 (n : ℤ) : I ^ (bit0 n) = (-1) ^ n := -by rw [fpow_bit0', I_mul_I] +@[simp] lemma I_zpow_bit0 (n : ℤ) : I ^ (bit0 n) = (-1) ^ n := +by rw [zpow_bit0', I_mul_I] -@[simp] lemma I_fpow_bit1 (n : ℤ) : I ^ (bit1 n) = (-1) ^ n * I := -by rw [fpow_bit1', I_mul_I] +@[simp] lemma I_zpow_bit1 (n : ℤ) : I ^ (bit1 n) = (-1) ^ n * I := +by rw [zpow_bit1', I_mul_I] lemma div_re (z w : ℂ) : (z / w).re = z.re * w.re / norm_sq w + z.im * w.im / norm_sq w := by simp [div_eq_mul_inv, mul_assoc, sub_eq_add_neg] @@ -341,8 +341,8 @@ by simp [div_eq_mul_inv, mul_assoc, sub_eq_add_neg, add_comm] @[simp, norm_cast] lemma of_real_div (r s : ℝ) : ((r / s : ℝ) : ℂ) = r / s := of_real.map_div r s -@[simp, norm_cast] lemma of_real_fpow (r : ℝ) (n : ℤ) : ((r ^ n : ℝ) : ℂ) = (r : ℂ) ^ n := -of_real.map_fpow r n +@[simp, norm_cast] lemma of_real_zpow (r : ℝ) (n : ℤ) : ((r ^ n : ℝ) : ℂ) = (r : ℂ) ^ n := +of_real.map_zpow r n @[simp] lemma div_I (z : ℂ) : z / I = -(z * I) := (div_eq_iff_mul_eq I_ne_zero).2 $ by simp [mul_assoc] diff --git a/src/data/complex/is_R_or_C.lean b/src/data/complex/is_R_or_C.lean index 62ae83db7c497..0df2d3d605db7 100644 --- a/src/data/complex/is_R_or_C.lean +++ b/src/data/complex/is_R_or_C.lean @@ -330,8 +330,8 @@ begin simp [norm_sq, div_mul_eq_div_mul_one_div, div_self h] } end -@[simp, norm_cast, priority 900] lemma of_real_fpow (r : ℝ) (n : ℤ) : ((r ^ n : ℝ) : K) = r ^ n := -(@is_R_or_C.coe_hom K _).map_fpow r n +@[simp, norm_cast, priority 900] lemma of_real_zpow (r : ℝ) (n : ℤ) : ((r ^ n : ℝ) : K) = r ^ n := +(@is_R_or_C.coe_hom K _).map_zpow r n lemma I_mul_I_of_nonzero : (I : K) ≠ 0 → (I : K) * I = -1 := by { have := I_mul_I_ax, tauto } diff --git a/src/data/equiv/mul_add_aut.lean b/src/data/equiv/mul_add_aut.lean index 9eff2fccf31d1..5ab45565bf26f 100644 --- a/src/data/equiv/mul_add_aut.lean +++ b/src/data/equiv/mul_add_aut.lean @@ -46,7 +46,7 @@ by refine_struct inv := mul_equiv.symm, div := _, npow := @npow_rec _ ⟨mul_equiv.refl M⟩ ⟨λ g h, mul_equiv.trans h g⟩, - gpow := @gpow_rec _ ⟨mul_equiv.refl M⟩ ⟨λ g h, mul_equiv.trans h g⟩ ⟨mul_equiv.symm⟩ }; + zpow := @zpow_rec _ ⟨mul_equiv.refl M⟩ ⟨λ g h, mul_equiv.trans h g⟩ ⟨mul_equiv.symm⟩ }; intros; ext; try { refl }; apply equiv.left_inv instance : inhabited (mul_aut M) := ⟨1⟩ @@ -122,7 +122,7 @@ by refine_struct inv := add_equiv.symm, div := _, npow := @npow_rec _ ⟨add_equiv.refl A⟩ ⟨λ g h, add_equiv.trans h g⟩, - gpow := @gpow_rec _ ⟨add_equiv.refl A⟩ ⟨λ g h, add_equiv.trans h g⟩ ⟨add_equiv.symm⟩ }; + zpow := @zpow_rec _ ⟨add_equiv.refl A⟩ ⟨λ g h, add_equiv.trans h g⟩ ⟨add_equiv.symm⟩ }; intros; ext; try { refl }; apply equiv.left_inv instance : inhabited (add_aut A) := ⟨1⟩ diff --git a/src/data/equiv/ring_aut.lean b/src/data/equiv/ring_aut.lean index 78fa57982e7e0..42dbccc8a7e74 100644 --- a/src/data/equiv/ring_aut.lean +++ b/src/data/equiv/ring_aut.lean @@ -43,7 +43,7 @@ by refine_struct inv := ring_equiv.symm, div := _, npow := @npow_rec _ ⟨ring_equiv.refl R⟩ ⟨λ g h, ring_equiv.trans h g⟩, - gpow := @gpow_rec _ ⟨ring_equiv.refl R⟩ ⟨λ g h, ring_equiv.trans h g⟩ ⟨ring_equiv.symm⟩ }; + zpow := @zpow_rec _ ⟨ring_equiv.refl R⟩ ⟨λ g h, ring_equiv.trans h g⟩ ⟨ring_equiv.symm⟩ }; intros; ext; try { refl }; apply equiv.left_inv instance : inhabited (ring_aut R) := ⟨1⟩ diff --git a/src/data/int/gcd.lean b/src/data/int/gcd.lean index 16384bf0f6852..fe591b96c5297 100644 --- a/src/data/int/gcd.lean +++ b/src/data/int/gcd.lean @@ -352,8 +352,8 @@ begin obtain ⟨x, rfl⟩ : is_unit x, { apply is_unit_of_pow_eq_one _ _ hm m.succ_pos }, simp only [← units.coe_pow] at *, - rw [← units.coe_one, ← gpow_coe_nat, ← units.ext_iff] at *, - simp only [nat.gcd_eq_gcd_ab, gpow_add, gpow_mul, hm, hn, one_gpow, one_mul] + rw [← units.coe_one, ← zpow_coe_nat, ← units.ext_iff] at *, + simp only [nat.gcd_eq_gcd_ab, zpow_add, zpow_mul, hm, hn, one_zpow, one_mul] end lemma gcd_nsmul_eq_zero {M : Type*} [add_monoid M] (x : M) {m n : ℕ} (hm : m • x = 0) diff --git a/src/data/real/irrational.lean b/src/data/real/irrational.lean index 553061c675163..6ee0140247a32 100644 --- a/src/data/real/irrational.lean +++ b/src/data/real/irrational.lean @@ -353,9 +353,9 @@ theorem of_pow : ∀ n : ℕ, irrational (x^n) → irrational x | 0 := λ h, by { rw pow_zero at h, exact (h ⟨1, cast_one⟩).elim } | (n+1) := λ h, by { rw pow_succ at h, exact h.mul_cases.elim id (of_pow n) } -theorem of_fpow : ∀ m : ℤ, irrational (x^m) → irrational x +theorem of_zpow : ∀ m : ℤ, irrational (x^m) → irrational x | (n:ℕ) := of_pow n -| -[1+n] := λ h, by { rw gpow_neg_succ_of_nat at h, exact h.of_inv.of_pow _ } +| -[1+n] := λ h, by { rw zpow_neg_succ_of_nat at h, exact h.of_inv.of_pow _ } end irrational diff --git a/src/deprecated/subfield.lean b/src/deprecated/subfield.lean index 12239b30ba174..2ede8becb1b18 100644 --- a/src/deprecated/subfield.lean +++ b/src/deprecated/subfield.lean @@ -35,8 +35,8 @@ lemma is_subfield.pow_mem {a : F} {n : ℤ} {s : set F} (hs : is_subfield s) (h a ^ n ∈ s := begin cases n, - { rw gpow_of_nat, exact hs.to_is_subring.to_is_submonoid.pow_mem h }, - { rw gpow_neg_succ_of_nat, exact hs.inv_mem (hs.to_is_subring.to_is_submonoid.pow_mem h) }, + { rw zpow_of_nat, exact hs.to_is_subring.to_is_submonoid.pow_mem h }, + { rw zpow_neg_succ_of_nat, exact hs.inv_mem (hs.to_is_subring.to_is_submonoid.pow_mem h) }, end lemma univ.is_subfield : is_subfield (@set.univ F) := diff --git a/src/dynamics/circle/rotation_number/translation_number.lean b/src/dynamics/circle/rotation_number/translation_number.lean index 578ccef2af384..daf5706f9e3e0 100644 --- a/src/dynamics/circle/rotation_number/translation_number.lean +++ b/src/dynamics/circle/rotation_number/translation_number.lean @@ -247,13 +247,13 @@ by refine (units.map _).comp to_units.to_monoid_hom; exact @[simp] lemma translate_inv_apply (x y : ℝ) : (translate $ multiplicative.of_add x)⁻¹ y = -x + y := rfl -@[simp] lemma translate_gpow (x : ℝ) (n : ℤ) : +@[simp] lemma translate_zpow (x : ℝ) (n : ℤ) : (translate (multiplicative.of_add x))^n = translate (multiplicative.of_add $ ↑n * x) := -by simp only [← gsmul_eq_mul, of_add_gsmul, monoid_hom.map_gpow] +by simp only [← gsmul_eq_mul, of_add_gsmul, monoid_hom.map_zpow] @[simp] lemma translate_pow (x : ℝ) (n : ℕ) : (translate (multiplicative.of_add x))^n = translate (multiplicative.of_add $ ↑n * x) := -translate_gpow x n +translate_zpow x n @[simp] lemma translate_iterate (x : ℝ) (n : ℕ) : (translate (multiplicative.of_add x))^[n] = translate (multiplicative.of_add $ ↑n * x) := @@ -608,7 +608,7 @@ eq_neg_iff_add_eq_zero.2 $ | (n+1) := by rw [pow_succ', translation_number_mul_of_commute (commute.pow_self f n), translation_number_pow n, nat.cast_add_one, add_mul, one_mul] -@[simp] lemma translation_number_gpow (f : units circle_deg1_lift) : +@[simp] lemma translation_number_zpow (f : units circle_deg1_lift) : ∀ n : ℤ, τ (f ^ n : units _) = n * τ f | (n : ℕ) := by simp [translation_number_pow f n] | -[1+n] := by { simp, ring } @@ -856,8 +856,8 @@ lemma units_semiconj_of_translation_number_eq {f₁ f₂ : units circle_deg1_lif (h : τ f₁ = τ f₂) : ∃ F : circle_deg1_lift, semiconj F f₁ f₂ := begin - have : ∀ n : multiplicative ℤ, τ ((units.coe_hom _).comp (gpowers_hom _ f₁) n) = - τ ((units.coe_hom _).comp (gpowers_hom _ f₂) n), + have : ∀ n : multiplicative ℤ, τ ((units.coe_hom _).comp (zpowers_hom _ f₁) n) = + τ ((units.coe_hom _).comp (zpowers_hom _ f₂) n), { intro n, simp [h] }, exact (semiconj_of_group_action_of_forall_translation_number_eq _ _ this).imp (λ F hF, hF (multiplicative.of_add 1)) diff --git a/src/field_theory/finite/basic.lean b/src/field_theory/finite/basic.lean index 555fa1ef63a9a..023113b782a9a 100644 --- a/src/field_theory/finite/basic.lean +++ b/src/field_theory/finite/basic.lean @@ -169,11 +169,11 @@ lemma forall_pow_eq_one_iff (i : ℕ) : begin classical, obtain ⟨x, hx⟩ := is_cyclic.exists_generator (units K), - rw [←fintype.card_units, ←order_of_eq_card_of_forall_mem_gpowers hx, order_of_dvd_iff_pow_eq_one], + rw [←fintype.card_units, ←order_of_eq_card_of_forall_mem_zpowers hx, order_of_dvd_iff_pow_eq_one], split, { intro h, apply h }, { intros h y, - simp_rw ← mem_powers_iff_mem_gpowers at hx, + simp_rw ← mem_powers_iff_mem_zpowers at hx, rcases hx y with ⟨j, rfl⟩, rw [← pow_mul, mul_comm, pow_mul, h, one_pow], } end diff --git a/src/field_theory/intermediate_field.lean b/src/field_theory/intermediate_field.lean index 7d10b0776ea47..214d416b24038 100644 --- a/src/field_theory/intermediate_field.lean +++ b/src/field_theory/intermediate_field.lean @@ -146,8 +146,8 @@ lemma sum_mem {ι : Type*} {t : finset ι} {f : ι → L} (h : ∀ c ∈ t, f c S.to_subfield.sum_mem h lemma pow_mem {x : L} (hx : x ∈ S) : ∀ (n : ℤ), x^n ∈ S -| (n : ℕ) := by { rw gpow_coe_nat, exact S.to_subfield.pow_mem hx _, } -| -[1+ n] := by { rw [gpow_neg_succ_of_nat], +| (n : ℕ) := by { rw zpow_coe_nat, exact S.to_subfield.pow_mem hx _, } +| -[1+ n] := by { rw [zpow_neg_succ_of_nat], exact S.to_subfield.inv_mem (S.to_subfield.pow_mem hx _) } lemma gsmul_mem {x : L} (hx : x ∈ S) (n : ℤ) : diff --git a/src/group_theory/free_group.lean b/src/group_theory/free_group.lean index 49bb030f01cbd..5d64f7c3cd9bd 100644 --- a/src/group_theory/free_group.lean +++ b/src/group_theory/free_group.lean @@ -629,12 +629,12 @@ def free_group_unit_equiv_int : free_group unit ≃ ℤ := begin rintros ⟨L⟩, refine list.rec_on L rfl _, - exact (λ ⟨⟨⟩, b⟩ tl ih, by cases b; simp [gpow_add] at ih ⊢; rw ih; refl), + exact (λ ⟨⟨⟩, b⟩ tl ih, by cases b; simp [zpow_add] at ih ⊢; rw ih; refl), end, right_inv := λ x, int.induction_on x (by simp) - (λ i ih, by simp at ih; simp [gpow_add, ih]) - (λ i ih, by simp at ih; simp [gpow_add, ih, sub_eq_add_neg, -int.add_neg_one]) + (λ i ih, by simp at ih; simp [zpow_add, ih]) + (λ i ih, by simp at ih; simp [zpow_add, ih, sub_eq_add_neg, -int.add_neg_one]) } section category diff --git a/src/group_theory/group_action/group.lean b/src/group_theory/group_action/group.lean index a49f86283dbe1..2e669b76397d1 100644 --- a/src/group_theory/group_action/group.lean +++ b/src/group_theory/group_action/group.lean @@ -89,7 +89,7 @@ lemma smul_inv [group β] [smul_comm_class α β β] [is_scalar_tower α β β] (c • x)⁻¹ = c⁻¹ • x⁻¹ := by rw [inv_eq_iff_mul_eq_one, smul_mul_smul, mul_right_inv, mul_right_inv, one_smul] -lemma smul_gpow [group β] [smul_comm_class α β β] [is_scalar_tower α β β] +lemma smul_zpow [group β] [smul_comm_class α β β] [is_scalar_tower α β β] (c : α) (x : β) (p : ℤ) : (c • x) ^ p = c ^ p • x ^ p := by { cases p; simp [smul_pow, smul_inv] } diff --git a/src/group_theory/order_of_element.lean b/src/group_theory/order_of_element.lean index c020e5426f43e..b2a9ae6bd894e 100644 --- a/src/group_theory/order_of_element.lean +++ b/src/group_theory/order_of_element.lean @@ -365,11 +365,11 @@ section group variables [group G] [add_group A] {x a} {i : ℤ} @[to_additive add_order_of_dvd_iff_gsmul_eq_zero] -lemma order_of_dvd_iff_gpow_eq_one : (order_of x : ℤ) ∣ i ↔ x ^ i = 1 := +lemma order_of_dvd_iff_zpow_eq_one : (order_of x : ℤ) ∣ i ↔ x ^ i = 1 := begin rcases int.eq_coe_or_neg i with ⟨i, rfl|rfl⟩, - { rw [int.coe_nat_dvd, order_of_dvd_iff_pow_eq_one, gpow_coe_nat] }, - { rw [dvd_neg, int.coe_nat_dvd, gpow_neg, inv_eq_one, gpow_coe_nat, + { rw [int.coe_nat_dvd, order_of_dvd_iff_pow_eq_one, zpow_coe_nat] }, + { rw [dvd_neg, int.coe_nat_dvd, zpow_neg, inv_eq_one, zpow_coe_nat, order_of_dvd_iff_pow_eq_one] } end @@ -377,19 +377,19 @@ end (y: H) : order_of (y : G) = order_of y := order_of_injective H.subtype subtype.coe_injective y -lemma gpow_eq_mod_order_of : x ^ i = x ^ (i % order_of x) := +lemma zpow_eq_mod_order_of : x ^ i = x ^ (i % order_of x) := calc x ^ i = x ^ (i % order_of x + order_of x * (i / order_of x)) : by rw [int.mod_add_div] ... = x ^ (i % order_of x) : - by simp [gpow_add, gpow_mul, pow_order_of_eq_one] + by simp [zpow_add, zpow_mul, pow_order_of_eq_one] lemma gsmul_eq_mod_add_order_of : i • a = (i % add_order_of a) • a := begin apply multiplicative.of_add.injective, - simp [of_add_gsmul, gpow_eq_mod_order_of], + simp [of_add_gsmul, zpow_eq_mod_order_of], end -attribute [to_additive gsmul_eq_mod_add_order_of] gpow_eq_mod_order_of +attribute [to_additive gsmul_eq_mod_add_order_of] zpow_eq_mod_order_of @[to_additive nsmul_inj_iff_of_add_order_of_eq_zero] lemma pow_inj_iff_of_order_of_eq_zero (h : order_of x = 0) {n m : ℕ} : @@ -642,19 +642,19 @@ end finite_cancel_monoid section finite_group variables [group G] [add_group A] -lemma exists_gpow_eq_one (x : G) : ∃ (i : ℤ) (H : i ≠ 0), x ^ (i : ℤ) = 1 := ---lemma exists_gpow_eq_one (a : α) : ∃ (i : ℤ) (H : i ≠ 0), a ^ (i : ℤ) = 1 := +lemma exists_zpow_eq_one (x : G) : ∃ (i : ℤ) (H : i ≠ 0), x ^ (i : ℤ) = 1 := +--lemma exists_zpow_eq_one (a : α) : ∃ (i : ℤ) (H : i ≠ 0), a ^ (i : ℤ) = 1 := begin rcases exists_pow_eq_one x with ⟨w, hw1, hw2⟩, refine ⟨w, int.coe_nat_ne_zero.mpr (ne_of_gt hw1), _⟩, - rw gpow_coe_nat, + rw zpow_coe_nat, exact (is_periodic_pt_mul_iff_pow_eq_one _).mp hw2, end lemma exists_gsmul_eq_zero (a : A) : ∃ (i : ℤ) (H : i ≠ 0), i • a = 0 := -@exists_gpow_eq_one (multiplicative A) _ _ a +@exists_zpow_eq_one (multiplicative A) _ _ a -attribute [to_additive] exists_gpow_eq_one +attribute [to_additive] exists_zpow_eq_one lemma mem_multiples_iff_mem_gmultiples : b ∈ add_submonoid.multiples a ↔ b ∈ add_subgroup.gmultiples a := @@ -667,29 +667,29 @@ lemma mem_multiples_iff_mem_gmultiples : open subgroup @[to_additive mem_multiples_iff_mem_gmultiples] -lemma mem_powers_iff_mem_gpowers : y ∈ submonoid.powers x ↔ y ∈ gpowers x := +lemma mem_powers_iff_mem_zpowers : y ∈ submonoid.powers x ↔ y ∈ zpowers x := ⟨λ ⟨n, hn⟩, ⟨n, by simp * at *⟩, λ ⟨i, hi⟩, ⟨(i % order_of x).nat_abs, - by rwa [← gpow_coe_nat, int.nat_abs_of_nonneg (int.mod_nonneg _ + by rwa [← zpow_coe_nat, int.nat_abs_of_nonneg (int.mod_nonneg _ (int.coe_nat_ne_zero_iff_pos.2 (order_of_pos x))), - ← gpow_eq_mod_order_of]⟩⟩ + ← zpow_eq_mod_order_of]⟩⟩ lemma multiples_eq_gmultiples (a : A) : (add_submonoid.multiples a : set A) = add_subgroup.gmultiples a := set.ext $ λ y, mem_multiples_iff_mem_gmultiples @[to_additive multiples_eq_gmultiples] -lemma powers_eq_gpowers (x : G) : (submonoid.powers x : set G) = gpowers x := -set.ext $ λ x, mem_powers_iff_mem_gpowers +lemma powers_eq_zpowers (x : G) : (submonoid.powers x : set G) = zpowers x := +set.ext $ λ x, mem_powers_iff_mem_zpowers lemma mem_gmultiples_iff_mem_range_add_order_of [decidable_eq A] : b ∈ add_subgroup.gmultiples a ↔ b ∈ (finset.range (add_order_of a)).image (• a) := by rw [← mem_multiples_iff_mem_gmultiples, mem_multiples_iff_mem_range_add_order_of] @[to_additive mem_gmultiples_iff_mem_range_add_order_of] -lemma mem_gpowers_iff_mem_range_order_of [decidable_eq G] : - y ∈ subgroup.gpowers x ↔ y ∈ (finset.range (order_of x)).image ((^) x : ℕ → G) := -by rw [← mem_powers_iff_mem_gpowers, mem_powers_iff_mem_range_order_of] +lemma mem_zpowers_iff_mem_range_order_of [decidable_eq G] : + y ∈ subgroup.zpowers x ↔ y ∈ (finset.range (order_of x)).image ((^) x : ℕ → G) := +by rw [← mem_powers_iff_mem_zpowers, mem_powers_iff_mem_range_order_of] noncomputable instance decidable_gmultiples [decidable_eq A] : decidable_pred (∈ add_subgroup.gmultiples a) := @@ -700,55 +700,55 @@ begin end @[to_additive decidable_gmultiples] -noncomputable instance decidable_gpowers [decidable_eq G] : - decidable_pred (∈ subgroup.gpowers x) := +noncomputable instance decidable_zpowers [decidable_eq G] : + decidable_pred (∈ subgroup.zpowers x) := begin simp_rw ←set_like.mem_coe, - rw ← powers_eq_gpowers, + rw ← powers_eq_zpowers, exact decidable_powers, end -/-- The equivalence between `fin (order_of x)` and `subgroup.gpowers x`, sending `i` to `x ^ i`. -/ -noncomputable def fin_equiv_gpowers (x : G) : - fin (order_of x) ≃ (subgroup.gpowers x : set G) := -(fin_equiv_powers x).trans (equiv.set.of_eq (powers_eq_gpowers x)) +/-- The equivalence between `fin (order_of x)` and `subgroup.zpowers x`, sending `i` to `x ^ i`. -/ +noncomputable def fin_equiv_zpowers (x : G) : + fin (order_of x) ≃ (subgroup.zpowers x : set G) := +(fin_equiv_powers x).trans (equiv.set.of_eq (powers_eq_zpowers x)) /-- The equivalence between `fin (add_order_of a)` and `subgroup.gmultiples a`, sending `i` to `i • a`. -/ noncomputable def fin_equiv_gmultiples (a : A) : fin (add_order_of a) ≃ (add_subgroup.gmultiples a : set A) := -fin_equiv_gpowers (multiplicative.of_add a) +fin_equiv_zpowers (multiplicative.of_add a) -attribute [to_additive fin_equiv_gmultiples] fin_equiv_gpowers +attribute [to_additive fin_equiv_gmultiples] fin_equiv_zpowers -@[simp] lemma fin_equiv_gpowers_apply {n : fin (order_of x)} : - fin_equiv_gpowers x n = ⟨x ^ (n : ℕ), n, gpow_coe_nat x n⟩ := rfl +@[simp] lemma fin_equiv_zpowers_apply {n : fin (order_of x)} : + fin_equiv_zpowers x n = ⟨x ^ (n : ℕ), n, zpow_coe_nat x n⟩ := rfl @[simp] lemma fin_equiv_gmultiples_apply {n : fin (add_order_of a)} : fin_equiv_gmultiples a n = ⟨(n : ℕ) • a, n, gsmul_coe_nat a n⟩ := -fin_equiv_gpowers_apply +fin_equiv_zpowers_apply -attribute [to_additive fin_equiv_gmultiples_apply] fin_equiv_gpowers_apply +attribute [to_additive fin_equiv_gmultiples_apply] fin_equiv_zpowers_apply -@[simp] lemma fin_equiv_gpowers_symm_apply (x : G) (n : ℕ) +@[simp] lemma fin_equiv_zpowers_symm_apply (x : G) (n : ℕ) {hn : ∃ (m : ℤ), x ^ m = x ^ n} : - ((fin_equiv_gpowers x).symm ⟨x ^ n, hn⟩) = ⟨n % order_of x, nat.mod_lt _ (order_of_pos x)⟩ := -by { rw [fin_equiv_gpowers, equiv.symm_trans_apply, equiv.set.of_eq_symm_apply], + ((fin_equiv_zpowers x).symm ⟨x ^ n, hn⟩) = ⟨n % order_of x, nat.mod_lt _ (order_of_pos x)⟩ := +by { rw [fin_equiv_zpowers, equiv.symm_trans_apply, equiv.set.of_eq_symm_apply], exact fin_equiv_powers_symm_apply x n } @[simp] lemma fin_equiv_gmultiples_symm_apply (a : A) (n : ℕ) {hn : ∃ (m : ℤ), m • a = n • a} : ((fin_equiv_gmultiples a).symm ⟨n • a, hn⟩) = ⟨n % add_order_of a, nat.mod_lt _ (add_order_of_pos a)⟩ := -fin_equiv_gpowers_symm_apply (multiplicative.of_add a) n +fin_equiv_zpowers_symm_apply (multiplicative.of_add a) n -attribute [to_additive fin_equiv_gmultiples_symm_apply] fin_equiv_gpowers_symm_apply +attribute [to_additive fin_equiv_gmultiples_symm_apply] fin_equiv_zpowers_symm_apply -/-- The equivalence between `subgroup.gpowers` of two elements `x, y` of the same order, mapping +/-- The equivalence between `subgroup.zpowers` of two elements `x, y` of the same order, mapping `x ^ i` to `y ^ i`. -/ -noncomputable def gpowers_equiv_gpowers (h : order_of x = order_of y) : - (subgroup.gpowers x : set G) ≃ (subgroup.gpowers y : set G) := -(fin_equiv_gpowers x).symm.trans ((fin.cast h).to_equiv.trans (fin_equiv_gpowers y)) +noncomputable def zpowers_equiv_zpowers (h : order_of x = order_of y) : + (subgroup.zpowers x : set G) ≃ (subgroup.zpowers y : set G) := +(fin_equiv_zpowers x).symm.trans ((fin.cast h).to_equiv.trans (fin_equiv_zpowers y)) /-- The equivalence between `subgroup.gmultiples` of two elements `a, b` of the same additive order, mapping `i • a` to `i • b`. -/ @@ -756,33 +756,33 @@ noncomputable def gmultiples_equiv_gmultiples (h : add_order_of a = add_order_of (add_subgroup.gmultiples a : set A) ≃ (add_subgroup.gmultiples b : set A) := (fin_equiv_gmultiples a).symm.trans ((fin.cast h).to_equiv.trans (fin_equiv_gmultiples b)) -attribute [to_additive gmultiples_equiv_gmultiples] gpowers_equiv_gpowers +attribute [to_additive gmultiples_equiv_gmultiples] zpowers_equiv_zpowers @[simp] -lemma gpowers_equiv_gpowers_apply (h : order_of x = order_of y) - (n : ℕ) : gpowers_equiv_gpowers h ⟨x ^ n, n, gpow_coe_nat x n⟩ = ⟨y ^ n, n, gpow_coe_nat y n⟩ := +lemma zpowers_equiv_zpowers_apply (h : order_of x = order_of y) + (n : ℕ) : zpowers_equiv_zpowers h ⟨x ^ n, n, zpow_coe_nat x n⟩ = ⟨y ^ n, n, zpow_coe_nat y n⟩ := begin - rw [gpowers_equiv_gpowers, equiv.trans_apply, equiv.trans_apply, - fin_equiv_gpowers_symm_apply, ← equiv.eq_symm_apply, fin_equiv_gpowers_symm_apply], + rw [zpowers_equiv_zpowers, equiv.trans_apply, equiv.trans_apply, + fin_equiv_zpowers_symm_apply, ← equiv.eq_symm_apply, fin_equiv_zpowers_symm_apply], simp [h] end @[simp] lemma gmultiples_equiv_gmultiples_apply (h : add_order_of a = add_order_of b) (n : ℕ) : gmultiples_equiv_gmultiples h ⟨n • a, n, gsmul_coe_nat a n⟩ = ⟨n • b, n, gsmul_coe_nat b n⟩ := -gpowers_equiv_gpowers_apply h n +zpowers_equiv_zpowers_apply h n -attribute [to_additive gmultiples_equiv_gmultiples_apply] gpowers_equiv_gpowers_apply +attribute [to_additive gmultiples_equiv_gmultiples_apply] zpowers_equiv_zpowers_apply -lemma order_eq_card_gpowers [decidable_eq G] : - order_of x = fintype.card (subgroup.gpowers x : set G) := -(fintype.card_fin (order_of x)).symm.trans (fintype.card_eq.2 ⟨fin_equiv_gpowers x⟩) +lemma order_eq_card_zpowers [decidable_eq G] : + order_of x = fintype.card (subgroup.zpowers x : set G) := +(fintype.card_fin (order_of x)).symm.trans (fintype.card_eq.2 ⟨fin_equiv_zpowers x⟩) lemma add_order_eq_card_gmultiples [decidable_eq A] : add_order_of a = fintype.card (add_subgroup.gmultiples a : set A) := (fintype.card_fin (add_order_of a)).symm.trans (fintype.card_eq.2 ⟨fin_equiv_gmultiples a⟩) -attribute [to_additive add_order_eq_card_gmultiples] order_eq_card_gpowers +attribute [to_additive add_order_eq_card_gmultiples] order_eq_card_zpowers open quotient_group @@ -790,12 +790,12 @@ open quotient_group lemma order_of_dvd_card_univ : order_of x ∣ fintype.card G := begin classical, - have ft_prod : fintype (quotient (gpowers x) × (gpowers x)), + have ft_prod : fintype (quotient (zpowers x) × (zpowers x)), from fintype.of_equiv G group_equiv_quotient_times_subgroup, - have ft_s : fintype (gpowers x), + have ft_s : fintype (zpowers x), from @fintype.prod_right _ _ _ ft_prod _, - have ft_cosets : fintype (quotient (gpowers x)), - from @fintype.prod_left _ _ _ ft_prod ⟨⟨1, (gpowers x).one_mem⟩⟩, + have ft_cosets : fintype (quotient (zpowers x)), + from @fintype.prod_left _ _ _ ft_prod ⟨⟨1, (zpowers x).one_mem⟩⟩, have eq₁ : fintype.card G = @fintype.card _ ft_cosets * @fintype.card _ ft_s, from calc fintype.card G = @fintype.card _ ft_prod : @fintype.card_congr _ _ _ ft_prod group_equiv_quotient_times_subgroup @@ -804,9 +804,9 @@ begin ... = @fintype.card _ ft_cosets * @fintype.card _ ft_s : @fintype.card_prod _ _ ft_cosets ft_s, have eq₂ : order_of x = @fintype.card _ ft_s, - from calc order_of x = _ : order_eq_card_gpowers + from calc order_of x = _ : order_eq_card_zpowers ... = _ : congr_arg (@fintype.card _) $ subsingleton.elim _ _, - exact dvd.intro (@fintype.card (quotient (subgroup.gpowers x)) ft_cosets) + exact dvd.intro (@fintype.card (quotient (subgroup.zpowers x)) ft_cosets) (by rw [eq₁, eq₂, mul_comm]) end @@ -834,10 +834,10 @@ end by rw [pow_eq_mod_order_of, ←nat.mod_mod_of_dvd n order_of_dvd_card_univ, ← pow_eq_mod_order_of] -@[to_additive] lemma gpow_eq_mod_card (n : ℤ) : +@[to_additive] lemma zpow_eq_mod_card (n : ℤ) : x ^ n = x ^ (n % fintype.card G) := -by rw [gpow_eq_mod_order_of, ← int.mod_mod_of_dvd n (int.coe_nat_dvd.2 order_of_dvd_card_univ), - ← gpow_eq_mod_order_of] +by rw [zpow_eq_mod_order_of, ← int.mod_mod_of_dvd n (int.coe_nat_dvd.2 order_of_dvd_card_univ), + ← zpow_eq_mod_order_of] attribute [to_additive card_nsmul_eq_zero] pow_card_eq_one @@ -847,12 +847,12 @@ attribute [to_additive card_nsmul_eq_zero] pow_card_eq_one inv_fun := λ g, g ^ (nat.gcd_b (fintype.card G) n), left_inv := λ g, by { have key : g ^ _ = g ^ _ := congr_arg (λ n : ℤ, g ^ n) (nat.gcd_eq_gcd_ab (fintype.card G) n), - rwa [gpow_add, gpow_mul, gpow_mul, gpow_coe_nat, gpow_coe_nat, gpow_coe_nat, - h.gcd_eq_one, pow_one, pow_card_eq_one, one_gpow, one_mul, eq_comm] at key }, + rwa [zpow_add, zpow_mul, zpow_mul, zpow_coe_nat, zpow_coe_nat, zpow_coe_nat, + h.gcd_eq_one, pow_one, pow_card_eq_one, one_zpow, one_mul, eq_comm] at key }, right_inv := λ g, by { have key : g ^ _ = g ^ _ := congr_arg (λ n : ℤ, g ^ n) (nat.gcd_eq_gcd_ab (fintype.card G) n), - rwa [gpow_add, gpow_mul, gpow_mul', gpow_coe_nat, gpow_coe_nat, gpow_coe_nat, - h.gcd_eq_one, pow_one, pow_card_eq_one, one_gpow, one_mul, eq_comm] at key } } + rwa [zpow_add, zpow_mul, zpow_mul', zpow_coe_nat, zpow_coe_nat, zpow_coe_nat, + h.gcd_eq_one, pow_one, pow_card_eq_one, one_zpow, one_mul, eq_comm] at key } } @[simp] lemma pow_coprime_one (h : nat.coprime (fintype.card G) n) : pow_coprime h 1 = 1 := one_pow n @@ -880,8 +880,8 @@ by {ext x, rw [set.mem_to_finset, set_like.mem_coe, mem_gmultiples_iff_mem_range /-- TODO: Generalise to `submonoid.powers`.-/ @[to_additive image_range_add_order_of] lemma image_range_order_of [decidable_eq G] : - finset.image (λ i, x ^ i) (finset.range (order_of x)) = (gpowers x : set G).to_finset := -by { ext x, rw [set.mem_to_finset, set_like.mem_coe, mem_gpowers_iff_mem_range_order_of] } + finset.image (λ i, x ^ i) (finset.range (order_of x)) = (zpowers x : set G).to_finset := +by { ext x, rw [set.mem_to_finset, set_like.mem_coe, mem_zpowers_iff_mem_range_order_of] } lemma gcd_nsmul_card_eq_zero_iff : n • a = 0 ↔ (gcd n (fintype.card A)) • a = 0 := ⟨λ h, gcd_nsmul_eq_zero _ h $ card_nsmul_eq_zero, diff --git a/src/group_theory/perm/basic.lean b/src/group_theory/perm/basic.lean index 10a04b5d6ab29..d9bf15ce9bac2 100644 --- a/src/group_theory/perm/basic.lean +++ b/src/group_theory/perm/basic.lean @@ -51,9 +51,9 @@ lemma eq_inv_iff_eq {f : perm α} {x y : α} : x = f⁻¹ y ↔ f x = y := f.eq_ lemma inv_eq_iff_eq {f : perm α} {x y : α} : f⁻¹ x = y ↔ x = f y := f.symm_apply_eq -lemma gpow_apply_comm {α : Type*} (σ : equiv.perm α) (m n : ℤ) {x : α} : +lemma zpow_apply_comm {α : Type*} (σ : equiv.perm α) (m n : ℤ) {x : α} : (σ ^ m) ((σ ^ n) x) = (σ ^ n) ((σ ^ m) x) := -by rw [←equiv.perm.mul_apply, ←equiv.perm.mul_apply, gpow_mul_comm] +by rw [←equiv.perm.mul_apply, ←equiv.perm.mul_apply, zpow_mul_comm] /-! Lemmas about mixing `perm` with `equiv`. Because we have multiple ways to express `equiv.refl`, `equiv.symm`, and `equiv.trans`, we want simp lemmas for every combination. diff --git a/src/group_theory/perm/concrete_cycle.lean b/src/group_theory/perm/concrete_cycle.lean index 3aec95164d4c1..f52e50a7239b5 100644 --- a/src/group_theory/perm/concrete_cycle.lean +++ b/src/group_theory/perm/concrete_cycle.lean @@ -81,7 +81,7 @@ begin have : w ∈ (x :: y :: l) := mem_of_form_perm_ne_self _ _ hw, obtain ⟨k, hk, rfl⟩ := nth_le_of_mem this, use k, - simp only [gpow_coe_nat, form_perm_pow_apply_head _ _ hl k, nat.mod_eq_of_lt hk] } } + simp only [zpow_coe_nat, form_perm_pow_apply_head _ _ hl k, nat.mod_eq_of_lt hk] } } end lemma pairwise_same_cycle_form_perm (hl : nodup l) (hn : 2 ≤ l.length) : diff --git a/src/group_theory/perm/cycles.lean b/src/group_theory/perm/cycles.lean index d08dc2b260e3d..65bc66e0d63cd 100644 --- a/src/group_theory/perm/cycles.lean +++ b/src/group_theory/perm/cycles.lean @@ -68,7 +68,7 @@ lemma is_cycle_swap {α : Type*} [decidable_eq α] {x y : α} (hxy : x ≠ y) : ⟨y, by rwa swap_apply_right, λ a (ha : ite (a = x) y (ite (a = y) x a) ≠ a), if hya : y = a then ⟨0, hya⟩ - else ⟨1, by { rw [gpow_one, swap_apply_def], split_ifs at *; cc }⟩⟩ + else ⟨1, by { rw [zpow_one, swap_apply_def], split_ifs at *; cc }⟩⟩ lemma is_swap.is_cycle {α : Type*} [decidable_eq α] {f : perm α} (hf : is_swap f) : is_cycle f := begin @@ -81,7 +81,7 @@ let ⟨x, hx⟩ := hf in ⟨x, by { simp only [inv_eq_iff_eq, *, forall_prop_of_true, ne.def] at *, cc }, λ y hy, let ⟨i, hi⟩ := hx.2 y (by { simp only [inv_eq_iff_eq, *, forall_prop_of_true, ne.def] at *, cc }) in - ⟨-i, by rwa [gpow_neg, inv_gpow, inv_inv]⟩⟩ + ⟨-i, by rwa [zpow_neg, inv_zpow, inv_inv]⟩⟩ lemma is_cycle.is_cycle_conj {f g : perm β} (hf : is_cycle f) : is_cycle (g * f * g⁻¹) := begin @@ -89,25 +89,25 @@ begin refine ⟨g a, by simp [ha1], λ b hb, _⟩, obtain ⟨i, hi⟩ := ha2 (g⁻¹ b) _, { refine ⟨i, _⟩, - rw conj_gpow, + rw conj_zpow, simp [hi] }, { contrapose! hb, rw [perm.mul_apply, perm.mul_apply, hb, apply_inv_self] } end -lemma is_cycle.exists_gpow_eq {f : perm β} (hf : is_cycle f) {x y : β} +lemma is_cycle.exists_zpow_eq {f : perm β} (hf : is_cycle f) {x y : β} (hx : f x ≠ x) (hy : f y ≠ y) : ∃ i : ℤ, (f ^ i) x = y := let ⟨g, hg⟩ := hf in let ⟨a, ha⟩ := hg.2 x hx in let ⟨b, hb⟩ := hg.2 y hy in -⟨b - a, by rw [← ha, ← mul_apply, ← gpow_add, sub_add_cancel, hb]⟩ +⟨b - a, by rw [← ha, ← mul_apply, ← zpow_add, sub_add_cancel, hb]⟩ lemma is_cycle.exists_pow_eq [fintype β] {f : perm β} (hf : is_cycle f) {x y : β} (hx : f x ≠ x) (hy : f y ≠ y) : ∃ i : ℕ, (f ^ i) x = y := -let ⟨n, hn⟩ := hf.exists_gpow_eq hx hy in +let ⟨n, hn⟩ := hf.exists_zpow_eq hx hy in by classical; exact ⟨(n % order_of f).to_nat, by { have := n.mod_nonneg (int.coe_nat_ne_zero.mpr (ne_of_gt (order_of_pos f))), - rwa [← gpow_coe_nat, int.to_nat_of_nonneg this, ← gpow_eq_mod_order_of] }⟩ + rwa [← zpow_coe_nat, int.to_nat_of_nonneg this, ← zpow_eq_mod_order_of] }⟩ lemma is_cycle.exists_pow_eq_one [fintype β] {f : perm β} (hf : is_cycle f) : ∃ (k : ℕ) (hk : 1 < k), f ^ k = 1 := @@ -124,12 +124,12 @@ begin end /-- The subgroup generated by a cycle is in bijection with its support -/ -noncomputable def is_cycle.gpowers_equiv_support {σ : perm α} (hσ : is_cycle σ) : - (↑(subgroup.gpowers σ) : set (perm α)) ≃ (↑(σ.support) : set α) := +noncomputable def is_cycle.zpowers_equiv_support {σ : perm α} (hσ : is_cycle σ) : + (↑(subgroup.zpowers σ) : set (perm α)) ≃ (↑(σ.support) : set α) := equiv.of_bijective (λ τ, ⟨τ (classical.some hσ), begin obtain ⟨τ, n, rfl⟩ := τ, - rw [finset.mem_coe, coe_fn_coe_base', subtype.coe_mk, gpow_apply_mem_support, mem_support], + rw [finset.mem_coe, coe_fn_coe_base', subtype.coe_mk, zpow_apply_mem_support, mem_support], exact (classical.some_spec hσ).1, end⟩) begin @@ -137,9 +137,9 @@ begin { rintros ⟨a, m, rfl⟩ ⟨b, n, rfl⟩ h, ext y, by_cases hy : σ y = y, - { simp_rw [subtype.coe_mk, gpow_apply_eq_self_of_apply_eq_self hy] }, + { simp_rw [subtype.coe_mk, zpow_apply_eq_self_of_apply_eq_self hy] }, { obtain ⟨i, rfl⟩ := (classical.some_spec hσ).2 y hy, - rw [subtype.coe_mk, subtype.coe_mk, gpow_apply_comm σ m i, gpow_apply_comm σ n i], + rw [subtype.coe_mk, subtype.coe_mk, zpow_apply_comm σ m i, zpow_apply_comm σ n i], exact congr_arg _ (subtype.ext_iff.mp h) } }, by { rintros ⟨y, hy⟩, rw [finset.mem_coe, mem_support] at hy, @@ -147,21 +147,21 @@ begin exact ⟨⟨σ ^ n, n, rfl⟩, rfl⟩ }, end -@[simp] lemma is_cycle.gpowers_equiv_support_apply {σ : perm α} (hσ : is_cycle σ) {n : ℕ} : - hσ.gpowers_equiv_support ⟨σ ^ n, n, rfl⟩ = ⟨(σ ^ n) (classical.some hσ), +@[simp] lemma is_cycle.zpowers_equiv_support_apply {σ : perm α} (hσ : is_cycle σ) {n : ℕ} : + hσ.zpowers_equiv_support ⟨σ ^ n, n, rfl⟩ = ⟨(σ ^ n) (classical.some hσ), pow_apply_mem_support.2 (mem_support.2 (classical.some_spec hσ).1)⟩ := rfl -@[simp] lemma is_cycle.gpowers_equiv_support_symm_apply {σ : perm α} (hσ : is_cycle σ) (n : ℕ) : - hσ.gpowers_equiv_support.symm ⟨(σ ^ n) (classical.some hσ), +@[simp] lemma is_cycle.zpowers_equiv_support_symm_apply {σ : perm α} (hσ : is_cycle σ) (n : ℕ) : + hσ.zpowers_equiv_support.symm ⟨(σ ^ n) (classical.some hσ), pow_apply_mem_support.2 (mem_support.2 (classical.some_spec hσ).1)⟩ = ⟨σ ^ n, n, rfl⟩ := -(equiv.symm_apply_eq _).2 hσ.gpowers_equiv_support_apply +(equiv.symm_apply_eq _).2 hσ.zpowers_equiv_support_apply lemma order_of_is_cycle {σ : perm α} (hσ : is_cycle σ) : order_of σ = σ.support.card := begin - rw [order_eq_card_gpowers, ←fintype.card_coe], - convert fintype.card_congr (is_cycle.gpowers_equiv_support hσ), + rw [order_eq_card_zpowers, ←fintype.card_coe], + convert fintype.card_congr (is_cycle.zpowers_equiv_support hσ), end lemma is_cycle_swap_mul_aux₁ {α : Type*} [decidable_eq α] : ∀ (n : ℕ) {b x : α} {f : perm α} @@ -178,7 +178,7 @@ lemma is_cycle_swap_mul_aux₁ {α : Type*} [decidable_eq α] : ∀ (n : ℕ) {b exact this.1 }, let ⟨i, hi⟩ := is_cycle_swap_mul_aux₁ n hb' (f.injective $ by { rw [apply_inv_self], rwa [pow_succ, mul_apply] at h }) in - ⟨i + 1, by rw [add_comm, gpow_add, mul_apply, hi, gpow_one, mul_apply, apply_inv_self, + ⟨i + 1, by rw [add_comm, zpow_add, mul_apply, hi, zpow_one, mul_apply, apply_inv_self, swap_apply_of_ne_of_ne (ne_and_ne_of_swap_mul_apply_ne_self hb).2 (ne.symm hfbx)]⟩ lemma is_cycle_swap_mul_aux₂ {α : Type*} [decidable_eq α] : @@ -187,25 +187,25 @@ lemma is_cycle_swap_mul_aux₂ {α : Type*} [decidable_eq α] : | (n : ℕ) := λ b x f, is_cycle_swap_mul_aux₁ n | -[1+ n] := λ b x f hb h, if hfbx : f⁻¹ x = b then - ⟨-1, by rwa [gpow_neg, gpow_one, mul_inv_rev, mul_apply, swap_inv, swap_apply_right]⟩ + ⟨-1, by rwa [zpow_neg, zpow_one, mul_inv_rev, mul_apply, swap_inv, swap_apply_right]⟩ else if hfbx' : f x = b then ⟨0, hfbx'⟩ else have f b ≠ b ∧ b ≠ x := ne_and_ne_of_swap_mul_apply_ne_self hb, have hb : (swap x (f⁻¹ x) * f⁻¹) (f⁻¹ b) ≠ f⁻¹ b, by { rw [mul_apply, swap_apply_def], split_ifs; - simp only [inv_eq_iff_eq, perm.mul_apply, gpow_neg_succ_of_nat, ne.def, + simp only [inv_eq_iff_eq, perm.mul_apply, zpow_neg_succ_of_nat, ne.def, perm.apply_inv_self] at *; cc }, let ⟨i, hi⟩ := is_cycle_swap_mul_aux₁ n hb (show (f⁻¹ ^ n) (f⁻¹ x) = f⁻¹ b, by - rw [← gpow_coe_nat, ← h, ← mul_apply, ← mul_apply, ← mul_apply, gpow_neg_succ_of_nat, - ← inv_pow, pow_succ', mul_assoc, mul_assoc, inv_mul_self, mul_one, gpow_coe_nat, + rw [← zpow_coe_nat, ← h, ← mul_apply, ← mul_apply, ← mul_apply, zpow_neg_succ_of_nat, + ← inv_pow, pow_succ', mul_assoc, mul_assoc, inv_mul_self, mul_one, zpow_coe_nat, ← pow_succ', ← pow_succ]) in have h : (swap x (f⁻¹ x) * f⁻¹) (f x) = f⁻¹ x, by rw [mul_apply, inv_apply_self, swap_apply_left], - ⟨-i, by rw [← add_sub_cancel i 1, neg_sub, sub_eq_add_neg, gpow_add, gpow_one, gpow_neg, - ← inv_gpow, mul_inv_rev, swap_inv, mul_swap_eq_swap_mul, inv_apply_self, swap_comm _ x, - gpow_add, gpow_one, mul_apply, mul_apply (_ ^ i), h, hi, mul_apply, apply_inv_self, + ⟨-i, by rw [← add_sub_cancel i 1, neg_sub, sub_eq_add_neg, zpow_add, zpow_one, zpow_neg, + ← inv_zpow, mul_inv_rev, swap_inv, mul_swap_eq_swap_mul, inv_apply_self, swap_comm _ x, + zpow_add, zpow_one, mul_apply, mul_apply (_ ^ i), h, hi, mul_apply, apply_inv_self, swap_apply_of_ne_of_ne this.2 (ne.symm hfbx')]⟩ lemma is_cycle.eq_swap_of_apply_apply_eq_self {α : Type*} [decidable_eq α] @@ -220,8 +220,8 @@ else begin rw [swap_apply_of_ne_of_ne hyx hfyx], refine by_contradiction (λ hy, _), cases hz.2 y hy with j hj, - rw [← sub_add_cancel j i, gpow_add, mul_apply, hi] at hj, - cases gpow_apply_eq_of_apply_apply_eq_self hffx (j - i) with hji hji, + rw [← sub_add_cancel j i, zpow_add, mul_apply, hi] at hj, + cases zpow_apply_eq_of_apply_apply_eq_self hffx (j - i) with hji hji, { rw [← hj, hji] at hyx, cc }, { rw [← hj, hji] at hfyx, cc } end @@ -231,10 +231,10 @@ lemma is_cycle.swap_mul {α : Type*} [decidable_eq α] {f : perm α} (hf : is_cy ⟨f x, by { simp only [swap_apply_def, mul_apply], split_ifs; simp [f.injective.eq_iff] at *; cc }, λ y hy, - let ⟨i, hi⟩ := hf.exists_gpow_eq hx (ne_and_ne_of_swap_mul_apply_ne_self hy).1 in + let ⟨i, hi⟩ := hf.exists_zpow_eq hx (ne_and_ne_of_swap_mul_apply_ne_self hy).1 in have hi : (f ^ (i - 1)) (f x) = y, from - calc (f ^ (i - 1)) (f x) = (f ^ (i - 1) * f ^ (1 : ℤ)) x : by rw [gpow_one, mul_apply] - ... = y : by rwa [← gpow_add, sub_add_cancel], + calc (f ^ (i - 1)) (f x) = (f ^ (i - 1) * f ^ (1 : ℤ)) x : by rw [zpow_one, mul_apply] + ... = y : by rwa [← zpow_add, sub_add_cancel], is_cycle_swap_mul_aux₂ (i - 1) hy hi⟩ lemma is_cycle.sign : ∀ {f : perm α} (hf : is_cycle f), @@ -273,17 +273,17 @@ begin obtain ⟨x, hx1, hx2⟩ := h1, refine ⟨x, (key x).mp hx1, λ y hy, _⟩, cases (hx2 y ((key y).mpr hy)) with i _, - exact ⟨n * i, by rwa gpow_mul⟩ + exact ⟨n * i, by rwa zpow_mul⟩ end --- The lemma `support_gpow_le` is relevant. It means that `h2` is equivalent to +-- The lemma `support_zpow_le` is relevant. It means that `h2` is equivalent to -- `σ.support = (σ ^ n).support`, as well as to `σ.support.card ≤ (σ ^ n).support.card`. -lemma is_cycle_of_is_cycle_gpow {σ : perm α} {n : ℤ} +lemma is_cycle_of_is_cycle_zpow {σ : perm α} {n : ℤ} (h1 : is_cycle (σ ^ n)) (h2 : σ.support ≤ (σ ^ n).support) : is_cycle σ := begin cases n, { exact is_cycle_of_is_cycle_pow h1 h2 }, - { simp only [le_eq_subset, gpow_neg_succ_of_nat, perm.support_inv] at h1 h2, + { simp only [le_eq_subset, zpow_neg_succ_of_nat, perm.support_inv] at h1 h2, simpa using is_cycle_of_is_cycle_pow h1.inv h2 } end @@ -305,8 +305,8 @@ begin have hint : ∀ (k : ℤ) (a : α), (g.extend_domain f ^ k) ↑(f a) = f ((g ^ k) a), { intros k a, induction k with k k, - { rw [gpow_of_nat, gpow_of_nat, hnat] }, - rw [gpow_neg_succ_of_nat, gpow_neg_succ_of_nat, inv_eq_iff_eq, hnat, apply_inv_self] }, + { rw [zpow_of_nat, zpow_of_nat, hnat] }, + rw [zpow_neg_succ_of_nat, zpow_neg_succ_of_nat, inv_eq_iff_eq, hnat, apply_inv_self] }, rw [hint, hi, apply_symm_apply, subtype.coe_mk] }, { rw [extend_domain_apply_subtype _ _ pb, con, apply_symm_apply, subtype.coe_mk] } }, { exact (hb (extend_domain_apply_not_subtype _ _ pb)).elim } @@ -328,20 +328,20 @@ def same_cycle (f : perm β) (x y : β) : Prop := ∃ i : ℤ, (f ^ i) x = y @[refl] lemma same_cycle.refl (f : perm β) (x : β) : same_cycle f x x := ⟨0, rfl⟩ @[symm] lemma same_cycle.symm {f : perm β} {x y : β} : same_cycle f x y → same_cycle f y x := -λ ⟨i, hi⟩, ⟨-i, by rw [gpow_neg, ← hi, inv_apply_self]⟩ +λ ⟨i, hi⟩, ⟨-i, by rw [zpow_neg, ← hi, inv_apply_self]⟩ @[trans] lemma same_cycle.trans {f : perm β} {x y z : β} : same_cycle f x y → same_cycle f y z → same_cycle f x z := -λ ⟨i, hi⟩ ⟨j, hj⟩, ⟨j + i, by rw [gpow_add, mul_apply, hi, hj]⟩ +λ ⟨i, hi⟩ ⟨j, hj⟩, ⟨j + i, by rw [zpow_add, mul_apply, hi, hj]⟩ lemma same_cycle.apply_eq_self_iff {f : perm β} {x y : β} : same_cycle f x y → (f x = x ↔ f y = y) := -λ ⟨i, hi⟩, by rw [← hi, ← mul_apply, ← gpow_one_add, add_comm, gpow_add_one, mul_apply, +λ ⟨i, hi⟩, by rw [← hi, ← mul_apply, ← zpow_one_add, add_comm, zpow_add_one, mul_apply, (f ^ i).injective.eq_iff] lemma is_cycle.same_cycle {f : perm β} (hf : is_cycle f) {x y : β} (hx : f x ≠ x) (hy : f y ≠ y) : same_cycle f x y := -hf.exists_gpow_eq hx hy +hf.exists_zpow_eq hx hy lemma same_cycle.nat' [fintype β] {f : perm β} {x y : β} (h : same_cycle f x y) : ∃ (i : ℕ) (h : i < order_of f), (f ^ i) x = y := @@ -352,9 +352,9 @@ begin { use 0, rw ←int.dvd_iff_mod_eq_zero at hk, obtain ⟨m, rfl⟩ := hk, - simp [pow_order_of_eq_one, order_of_pos, gpow_mul] }, + simp [pow_order_of_eq_one, order_of_pos, zpow_mul] }, { use ((k % order_of f).nat_abs), - rw [←gpow_coe_nat, int.nat_abs_of_nonneg, ←gpow_eq_mod_order_of], + rw [←zpow_coe_nat, int.nat_abs_of_nonneg, ←zpow_eq_mod_order_of], { refine ⟨_, rfl⟩, rw [←int.coe_nat_lt, int.nat_abs_of_nonneg], { refine (int.mod_lt_of_pos _ _), @@ -385,25 +385,25 @@ instance [fintype α] (f : perm α) : decidable_rel (same_cycle f) := { simp [order_of_le_card_univ] }, exact fintype_perm }, exact fintype_perm, }), - by { rw [← gpow_coe_nat, int.nat_abs_of_nonneg (int.mod_nonneg _ - (int.coe_nat_ne_zero_iff_pos.2 (order_of_pos _))), ← gpow_eq_mod_order_of, hi], + by { rw [← zpow_coe_nat, int.nat_abs_of_nonneg (int.mod_nonneg _ + (int.coe_nat_ne_zero_iff_pos.2 (order_of_pos _))), ← zpow_eq_mod_order_of, hi], exact fintype_perm }⟩⟩ lemma same_cycle_apply {f : perm β} {x y : β} : same_cycle f x (f y) ↔ same_cycle f x y := -⟨λ ⟨i, hi⟩, ⟨-1 + i, by rw [gpow_add, mul_apply, hi, gpow_neg_one, inv_apply_self]⟩, - λ ⟨i, hi⟩, ⟨1 + i, by rw [gpow_add, mul_apply, hi, gpow_one]⟩⟩ +⟨λ ⟨i, hi⟩, ⟨-1 + i, by rw [zpow_add, mul_apply, hi, zpow_neg_one, inv_apply_self]⟩, + λ ⟨i, hi⟩, ⟨1 + i, by rw [zpow_add, mul_apply, hi, zpow_one]⟩⟩ lemma same_cycle_cycle {f : perm β} {x : β} (hx : f x ≠ x) : is_cycle f ↔ (∀ {y}, same_cycle f x y ↔ f y ≠ y) := ⟨λ hf y, ⟨λ ⟨i, hi⟩ hy, hx $ - by { rw [← gpow_apply_eq_self_of_apply_eq_self hy i, (f ^ i).injective.eq_iff] at hi, + by { rw [← zpow_apply_eq_self_of_apply_eq_self hy i, (f ^ i).injective.eq_iff] at hi, rw [hi, hy] }, - hf.exists_gpow_eq hx⟩, + hf.exists_zpow_eq hx⟩, λ h, ⟨x, hx, λ y hy, h.2 hy⟩⟩ lemma same_cycle_inv (f : perm β) {x y : β} : same_cycle f⁻¹ x y ↔ same_cycle f x y := -⟨λ ⟨i, hi⟩, ⟨-i, by rw [gpow_neg, ← inv_gpow, hi]⟩, - λ ⟨i, hi⟩, ⟨-i, by rw [gpow_neg, ← inv_gpow, inv_inv, hi]⟩ ⟩ +⟨λ ⟨i, hi⟩, ⟨-i, by rw [zpow_neg, ← inv_zpow, hi]⟩, + λ ⟨i, hi⟩, ⟨-i, by rw [zpow_neg, ← inv_zpow, inv_inv, hi]⟩ ⟩ lemma same_cycle_inv_apply {f : perm β} {x y : β} : same_cycle f x (f⁻¹ y) ↔ same_cycle f x y := by rw [← same_cycle_inv, same_cycle_apply, same_cycle_inv] @@ -414,18 +414,18 @@ begin split, { rintro ⟨k, rfl⟩, use (k + n), - simp [gpow_add] }, + simp [zpow_add] }, { rintro ⟨k, rfl⟩, use (k - n), - rw [←gpow_coe_nat, ←mul_apply, ←gpow_add, int.sub_add_cancel] } + rw [←zpow_coe_nat, ←mul_apply, ←zpow_add, int.sub_add_cancel] } end -@[simp] lemma same_cycle_gpow_left_iff {f : perm β} {x y : β} {n : ℤ} : +@[simp] lemma same_cycle_zpow_left_iff {f : perm β} {x y : β} {n : ℤ} : same_cycle f ((f ^ n) x) y ↔ same_cycle f x y := begin cases n, { exact same_cycle_pow_left_iff }, - { rw [gpow_neg_succ_of_nat, ←inv_pow, ←same_cycle_inv, same_cycle_pow_left_iff, same_cycle_inv] } + { rw [zpow_neg_succ_of_nat, ←inv_pow, ←same_cycle_inv, same_cycle_pow_left_iff, same_cycle_inv] } end /-- Unlike `support_congr`, which assumes that `∀ (x ∈ g.support), f x = g x)`, here @@ -579,11 +579,11 @@ end cycle_of_pow_apply_self, if_pos, pow_succ, mul_apply], exact ⟨n, rfl⟩ } -@[simp] lemma cycle_of_gpow_apply_self [fintype α] (f : perm α) (x : α) : +@[simp] lemma cycle_of_zpow_apply_self [fintype α] (f : perm α) (x : α) : ∀ n : ℤ, (cycle_of f x ^ n) x = (f ^ n) x | (n : ℕ) := cycle_of_pow_apply_self f x n -| -[1+ n] := by rw [gpow_neg_succ_of_nat, ← inv_pow, cycle_of_inv, - gpow_neg_succ_of_nat, ← inv_pow, cycle_of_pow_apply_self] +| -[1+ n] := by rw [zpow_neg_succ_of_nat, ← inv_pow, cycle_of_inv, + zpow_neg_succ_of_nat, ← inv_pow, cycle_of_pow_apply_self] lemma same_cycle.cycle_of_apply [fintype α] {f : perm α} {x y : α} (h : same_cycle f x y) : cycle_of f x y = f y := dif_pos h @@ -601,17 +601,17 @@ begin { exact (cycle_of_apply_of_not_same_cycle (mt h.trans hz)).symm } end -@[simp] lemma cycle_of_apply_apply_gpow_self [fintype α] (f : perm α) (x : α) (k : ℤ) : +@[simp] lemma cycle_of_apply_apply_zpow_self [fintype α] (f : perm α) (x : α) (k : ℤ) : cycle_of f x ((f ^ k) x) = (f ^ (k + 1)) x := begin rw same_cycle.cycle_of_apply, - { rw [add_comm, gpow_add, gpow_one, mul_apply] }, + { rw [add_comm, zpow_add, zpow_one, mul_apply] }, { exact ⟨k, rfl⟩ } end @[simp] lemma cycle_of_apply_apply_pow_self [fintype α] (f : perm α) (x : α) (k : ℕ) : cycle_of f x ((f ^ k) x) = (f ^ (k + 1)) x := -by convert cycle_of_apply_apply_gpow_self f x k using 1 +by convert cycle_of_apply_apply_zpow_self f x k using 1 @[simp] lemma cycle_of_apply_apply_self [fintype α] (f : perm α) (x : α) : cycle_of f x (f x) = f (f x) := @@ -643,9 +643,9 @@ end cycle_of f ((f ^ n) x) = cycle_of f x := (same_cycle_pow_left_iff.mpr (same_cycle.refl _ _)).cycle_of_eq -@[simp] lemma cycle_of_self_apply_gpow [fintype α] (f : perm α) (n : ℤ) (x : α) : +@[simp] lemma cycle_of_self_apply_zpow [fintype α] (f : perm α) (n : ℤ) (x : α) : cycle_of f ((f ^ n) x) = cycle_of f x := -(same_cycle_gpow_left_iff.mpr (same_cycle.refl _ _)).cycle_of_eq +(same_cycle_zpow_left_iff.mpr (same_cycle.refl _ _)).cycle_of_eq lemma is_cycle.cycle_of [fintype α] {f : perm α} (hf : is_cycle f) {x : α} : cycle_of f x = if f x = x then 1 else f := @@ -664,7 +664,7 @@ have cycle_of f x x ≠ x, by rwa [(same_cycle.refl _ _).cycle_of_apply], ⟨λ h, mt h.apply_eq_self_iff.2 this, λ h, if hxy : same_cycle f x y then let ⟨i, hi⟩ := hxy in - ⟨i, by rw [cycle_of_gpow_apply_self, hi]⟩ + ⟨i, by rw [cycle_of_zpow_apply_self, hi]⟩ else by { rw [cycle_of_apply_of_not_same_cycle hxy] at h, exact (h rfl).elim }⟩ @[simp] lemma two_le_card_support_cycle_of_iff [fintype α] {f : perm α} {x : α} : @@ -693,13 +693,13 @@ begin ext y, by_cases hxy : (f * g).same_cycle x y, { obtain ⟨z, rfl⟩ := hxy, - rw cycle_of_apply_apply_gpow_self, - simp [h.mul_gpow, gpow_apply_eq_self_of_apply_eq_self hx] }, + rw cycle_of_apply_apply_zpow_self, + simp [h.mul_zpow, zpow_apply_eq_self_of_apply_eq_self hx] }, { rw [cycle_of_apply_of_not_same_cycle hxy, cycle_of_apply_of_not_same_cycle], contrapose! hxy, obtain ⟨z, rfl⟩ := hxy, refine ⟨z, _⟩, - simp [h.mul_gpow, gpow_apply_eq_self_of_apply_eq_self hx] } + simp [h.mul_zpow, zpow_apply_eq_self_of_apply_eq_self hx] } end lemma disjoint.cycle_of_mul_distrib [fintype α] {f g : perm α} (h : f.disjoint g) (x : α) : @@ -1129,14 +1129,14 @@ begin { have hpow : ∀ (k : ℤ), ((g * h) ^ k) x = (h ^ k) x, { intro k, suffices : (g ^ k) x = x, - { simpa [hd.commute.eq, hd.commute.symm.mul_gpow] }, - rw gpow_apply_eq_self_of_apply_eq_self, + { simpa [hd.commute.eq, hd.commute.symm.mul_zpow] }, + rw zpow_apply_eq_self_of_apply_eq_self, simpa using hgx }, obtain ⟨k, hk, hk'⟩ := IH' _ _, { refine ⟨k, _, _⟩, { rw [←cycle_of_eq_one_iff] at hgx, rwa [hd.cycle_of_mul_distrib, hgx, one_mul] }, - { simpa [←gpow_coe_nat, hpow] using hk' } }, + { simpa [←zpow_coe_nat, hpow] using hk' } }, { use m, simp [hpow] }, { rw [mem_support, hd.commute.eq] at hx, @@ -1144,14 +1144,14 @@ begin { have hpow : ∀ (k : ℤ), ((g * h) ^ k) x = (g ^ k) x, { intro k, suffices : (h ^ k) x = x, - { simpa [hd.commute.mul_gpow] }, - rw gpow_apply_eq_self_of_apply_eq_self, + { simpa [hd.commute.mul_zpow] }, + rw zpow_apply_eq_self_of_apply_eq_self, simpa using hhx }, obtain ⟨k, hk, hk'⟩ := IH _ _, { refine ⟨k, _, _⟩, { rw [←cycle_of_eq_one_iff] at hhx, rwa [hd.cycle_of_mul_distrib, hhx, mul_one] }, - { simpa [←gpow_coe_nat, hpow] using hk' } }, + { simpa [←zpow_coe_nat, hpow] using hk' } }, { use m, simp [hpow] }, { simpa [hhx] using hx } } } @@ -1175,7 +1175,7 @@ begin obtain ⟨k, rfl⟩ := h, rw [not_mem_support] at hx, rw [pow_apply_eq_self_of_apply_eq_self hx, - gpow_apply_eq_self_of_apply_eq_self hx] } + zpow_apply_eq_self_of_apply_eq_self hx] } end section generation @@ -1292,18 +1292,18 @@ end theorem is_cycle.is_conj (hσ : is_cycle σ) (hτ : is_cycle τ) (h : σ.support.card = τ.support.card) : is_conj σ τ := begin - refine is_conj_of_support_equiv (hσ.gpowers_equiv_support.symm.trans - ((gpowers_equiv_gpowers begin + refine is_conj_of_support_equiv (hσ.zpowers_equiv_support.symm.trans + ((zpowers_equiv_zpowers begin rw [order_of_is_cycle hσ, h, order_of_is_cycle hτ], - end).trans hτ.gpowers_equiv_support)) _, + end).trans hτ.zpowers_equiv_support)) _, intros x hx, simp only [perm.mul_apply, equiv.trans_apply, equiv.sum_congr_apply], obtain ⟨n, rfl⟩ := hσ.exists_pow_eq (classical.some_spec hσ).1 (mem_support.1 hx), apply eq.trans _ (congr rfl (congr rfl (congr rfl - (congr rfl (hσ.gpowers_equiv_support_symm_apply n).symm)))), - apply (congr rfl (congr rfl (congr rfl (hσ.gpowers_equiv_support_symm_apply (n + 1))))).trans _, - simp only [ne.def, is_cycle.gpowers_equiv_support_apply, - subtype.coe_mk, gpowers_equiv_gpowers_apply], + (congr rfl (hσ.zpowers_equiv_support_symm_apply n).symm)))), + apply (congr rfl (congr rfl (congr rfl (hσ.zpowers_equiv_support_symm_apply (n + 1))))).trans _, + simp only [ne.def, is_cycle.zpowers_equiv_support_apply, + subtype.coe_mk, zpowers_equiv_zpowers_apply], rw [pow_succ, perm.mul_apply], end diff --git a/src/group_theory/perm/fin.lean b/src/group_theory/perm/fin.lean index ddaedcfaaaa89..e6f940f11e328 100644 --- a/src/group_theory/perm/fin.lean +++ b/src/group_theory/perm/fin.lean @@ -111,7 +111,7 @@ begin refine ⟨0, dec_trivial, λ x hx', ⟨x, _⟩⟩, clear hx', cases x with x hx, - rw [coe_coe, gpow_coe_nat, fin.ext_iff, fin.coe_mk], + rw [coe_coe, zpow_coe_nat, fin.ext_iff, fin.coe_mk], induction x with x ih, { refl }, rw [pow_succ, perm.mul_apply, coe_fin_rotate_of_ne_last, ih (lt_trans x.lt_succ_self hx)], rw [ne.def, fin.ext_iff, ih (lt_trans x.lt_succ_self hx), fin.coe_last], diff --git a/src/group_theory/perm/list.lean b/src/group_theory/perm/list.lean index 6f2ee0a5c8569..6afbf9fe5248f 100644 --- a/src/group_theory/perm/list.lean +++ b/src/group_theory/perm/list.lean @@ -420,14 +420,14 @@ begin { simp [-form_perm_cons_cons, form_perm_ext_iff hl hl'] } } end -lemma form_perm_gpow_apply_mem_imp_mem (l : list α) (x : α) (hx : x ∈ l) (n : ℤ) : +lemma form_perm_zpow_apply_mem_imp_mem (l : list α) (x : α) (hx : x ∈ l) (n : ℤ) : ((form_perm l) ^ n) x ∈ l := begin by_cases h : (l.form_perm ^ n) x = x, { simpa [h] using hx }, { have : x ∈ {x | (l.form_perm ^ n) x ≠ x} := h, rw ←set_support_apply_mem at this, - replace this := set_support_gpow_subset _ _ this, + replace this := set_support_zpow_subset _ _ this, simpa using support_form_perm_le' _ this } end @@ -441,7 +441,7 @@ begin { have : x ∉ {x | (l.form_perm ^ l.length) x ≠ x}, { intros H, refine hx _, - replace H := set_support_gpow_subset l.form_perm l.length H, + replace H := set_support_zpow_subset l.form_perm l.length H, simpa using support_form_perm_le' _ H }, simpa } end diff --git a/src/group_theory/perm/support.lean b/src/group_theory/perm/support.lean index eb92da5884402..173dd58756a30 100644 --- a/src/group_theory/perm/support.lean +++ b/src/group_theory/perm/support.lean @@ -122,10 +122,10 @@ lemma pow_apply_eq_self_of_apply_eq_self {x : α} (hfx : f x = x) : | 0 := rfl | (n+1) := by rw [pow_succ', mul_apply, hfx, pow_apply_eq_self_of_apply_eq_self] -lemma gpow_apply_eq_self_of_apply_eq_self {x : α} (hfx : f x = x) : +lemma zpow_apply_eq_self_of_apply_eq_self {x : α} (hfx : f x = x) : ∀ n : ℤ, (f ^ n) x = x | (n : ℕ) := pow_apply_eq_self_of_apply_eq_self hfx n -| -[1+ n] := by rw [gpow_neg_succ_of_nat, inv_eq_iff_eq, pow_apply_eq_self_of_apply_eq_self hfx] +| -[1+ n] := by rw [zpow_neg_succ_of_nat, inv_eq_iff_eq, pow_apply_eq_self_of_apply_eq_self hfx] lemma pow_apply_eq_of_apply_apply_eq_self {x : α} (hffx : f (f x) = x) : ∀ n : ℕ, (f ^ n) x = x ∨ (f ^ n) x = f x @@ -134,10 +134,10 @@ lemma pow_apply_eq_of_apply_apply_eq_self {x : α} (hffx : f (f x) = x) : (λ h, or.inr (by rw [pow_succ, mul_apply, h])) (λ h, or.inl (by rw [pow_succ, mul_apply, h, hffx])) -lemma gpow_apply_eq_of_apply_apply_eq_self {x : α} (hffx : f (f x) = x) : +lemma zpow_apply_eq_of_apply_apply_eq_self {x : α} (hffx : f (f x) = x) : ∀ i : ℤ, (f ^ i) x = x ∨ (f ^ i) x = f x | (n : ℕ) := pow_apply_eq_of_apply_apply_eq_self hffx n -| -[1+ n] := by { rw [gpow_neg_succ_of_nat, inv_eq_iff_eq, ← f.injective.eq_iff, ← mul_apply, +| -[1+ n] := by { rw [zpow_neg_succ_of_nat, inv_eq_iff_eq, ← f.injective.eq_iff, ← mul_apply, ← pow_succ, eq_comm, inv_eq_iff_eq, ← mul_apply, ← pow_succ', @eq_comm _ x, or.comm], exact pow_apply_eq_of_apply_apply_eq_self hffx _ } @@ -154,14 +154,14 @@ lemma disjoint.mul_eq_one_iff {σ τ : perm α} (hστ : disjoint σ τ) : σ * τ = 1 ↔ σ = 1 ∧ τ = 1 := by simp_rw [ext_iff, one_apply, hστ.mul_apply_eq_iff, forall_and_distrib] -lemma disjoint.gpow_disjoint_gpow {σ τ : perm α} (hστ : disjoint σ τ) (m n : ℤ) : +lemma disjoint.zpow_disjoint_zpow {σ τ : perm α} (hστ : disjoint σ τ) (m n : ℤ) : disjoint (σ ^ m) (τ ^ n) := -λ x, or.imp (λ h, gpow_apply_eq_self_of_apply_eq_self h m) - (λ h, gpow_apply_eq_self_of_apply_eq_self h n) (hστ x) +λ x, or.imp (λ h, zpow_apply_eq_self_of_apply_eq_self h m) + (λ h, zpow_apply_eq_self_of_apply_eq_self h n) (hστ x) lemma disjoint.pow_disjoint_pow {σ τ : perm α} (hστ : disjoint σ τ) (m n : ℕ) : disjoint (σ ^ m) (τ ^ n) := -hστ.gpow_disjoint_gpow m n +hστ.zpow_disjoint_zpow m n end disjoint @@ -212,13 +212,13 @@ lemma set_support_apply_mem {p : perm α} {a : α} : p a ∈ {x | p x ≠ x} ↔ a ∈ {x | p x ≠ x} := by simp -lemma set_support_gpow_subset (n : ℤ) : +lemma set_support_zpow_subset (n : ℤ) : {x | (p ^ n) x ≠ x} ⊆ {x | p x ≠ x} := begin intros x, simp only [set.mem_set_of_eq, ne.def], intros hx H, - simpa [gpow_apply_eq_self_of_apply_eq_self H] using hx + simpa [zpow_apply_eq_self_of_apply_eq_self H] using hx end lemma set_support_mul_subset : @@ -308,12 +308,12 @@ begin end @[simp] -lemma gpow_apply_mem_support {n : ℤ} {x : α} : +lemma zpow_apply_mem_support {n : ℤ} {x : α} : (f ^ n) x ∈ f.support ↔ x ∈ f.support := begin cases n, - { rw [int.of_nat_eq_coe, gpow_coe_nat, pow_apply_mem_support] }, - { rw [gpow_neg_succ_of_nat, ← support_inv, ← inv_pow, pow_apply_mem_support] } + { rw [int.of_nat_eq_coe, zpow_coe_nat, pow_apply_mem_support] }, + { rw [zpow_neg_succ_of_nat, ← support_inv, ← inv_pow, pow_apply_mem_support] } end lemma pow_eq_on_of_mem_support (h : ∀ (x ∈ f.support ∩ g.support), f x = g x) @@ -363,9 +363,9 @@ begin exact sup_le_sup (le_refl _) hl } end -lemma support_gpow_le (σ : perm α) (n : ℤ) : +lemma support_zpow_le (σ : perm α) (n : ℤ) : (σ ^ n).support ≤ σ.support := -λ x h1, mem_support.mpr (λ h2, mem_support.mp h1 (gpow_apply_eq_self_of_apply_eq_self h2 n)) +λ x h1, mem_support.mpr (λ h2, mem_support.mp h1 (zpow_apply_eq_self_of_apply_eq_self h2 n)) @[simp] lemma support_swap {x y : α} (h : x ≠ y) : support (swap x y) = {x, y} := begin diff --git a/src/group_theory/quotient_group.lean b/src/group_theory/quotient_group.lean index b36c97d522519..9946637bc5415 100644 --- a/src/group_theory/quotient_group.lean +++ b/src/group_theory/quotient_group.lean @@ -141,8 +141,8 @@ lemma coe_inv (a : G) : ((a⁻¹ : G) : Q) = a⁻¹ := rfl @[simp] lemma coe_pow (a : G) (n : ℕ) : ((a ^ n : G) : Q) = a ^ n := (mk' N).map_pow a n -@[simp] lemma coe_gpow (a : G) (n : ℤ) : ((a ^ n : G) : Q) = a ^ n := -(mk' N).map_gpow a n +@[simp] lemma coe_zpow (a : G) (n : ℤ) : ((a ^ n : G) : Q) = a ^ n := +(mk' N).map_zpow a n /-- A group homomorphism `φ : G →* H` with `N ⊆ ker(φ)` descends (i.e. `lift`s) to a group homomorphism `G/N →* H`. -/ diff --git a/src/group_theory/specific_groups/cyclic.lean b/src/group_theory/specific_groups/cyclic.lean index 34b0ae2eec603..f6508ca1c93dd 100644 --- a/src/group_theory/specific_groups/cyclic.lean +++ b/src/group_theory/specific_groups/cyclic.lean @@ -53,11 +53,11 @@ open subgroup /-- A group is called *cyclic* if it is generated by a single element. -/ class is_cyclic (α : Type u) [group α] : Prop := -(exists_generator [] : ∃ g : α, ∀ x, x ∈ gpowers g) +(exists_generator [] : ∃ g : α, ∀ x, x ∈ zpowers g) @[priority 100] instance is_cyclic_of_subsingleton [group α] [subsingleton α] : is_cyclic α := -⟨⟨1, λ x, by { rw subsingleton.elim x 1, exact mem_gpowers 1 }⟩⟩ +⟨⟨1, λ x, by { rw subsingleton.elim x 1, exact mem_zpowers 1 }⟩⟩ /-- A cyclic group is always commutative. This is not an `instance` because often we have a better proof of `comm_group`. -/ @@ -65,7 +65,7 @@ def is_cyclic.comm_group [hg : group α] [is_cyclic α] : comm_group α := { mul_comm := λ x y, show x * y = y * x, from let ⟨g, hg⟩ := is_cyclic.exists_generator α in let ⟨n, hn⟩ := hg x in let ⟨m, hm⟩ := hg y in - hm ▸ hn ▸ gpow_mul_comm _ _ _, + hm ▸ hn ▸ zpow_mul_comm _ _ _, ..hg } variables [group α] @@ -78,7 +78,7 @@ begin use m, intro g, obtain ⟨n, rfl⟩ := hG g, - rw [monoid_hom.map_gpow, ←hm, ←gpow_mul, ←gpow_mul'], + rw [monoid_hom.map_zpow, ←hm, ←zpow_mul, ←zpow_mul'], end lemma is_cyclic_of_order_of_eq_card [fintype α] (x : α) @@ -89,7 +89,7 @@ begin simp_rw ← set_like.mem_coe, rw ← set.eq_univ_iff_forall, apply set.eq_of_subset_of_card_le (set.subset_univ _), - rw [fintype.card_congr (equiv.set.univ α), ← hx, order_eq_card_gpowers], + rw [fintype.card_congr (equiv.set.univ α), ← hx, order_eq_card_zpowers], end /-- A finite group of prime order is cyclic. -/ @@ -98,8 +98,8 @@ lemma is_cyclic_of_prime_card {α : Type u} [group α] [fintype α] {p : ℕ} [h ⟨begin obtain ⟨g, hg⟩ : ∃ g : α, g ≠ 1, from fintype.exists_ne_of_one_lt_card (by { rw h, exact hp.1.one_lt }) 1, - classical, -- for fintype (subgroup.gpowers g) - have : fintype.card (subgroup.gpowers g) ∣ p, + classical, -- for fintype (subgroup.zpowers g) + have : fintype.card (subgroup.zpowers g) ∣ p, { rw ←h, apply card_subgroup_dvd_card }, rw nat.dvd_prime hp.1 at this, @@ -108,9 +108,9 @@ lemma is_cyclic_of_prime_card {α : Type u} [group α] [fintype α] {p : ℕ} [h cases this with t ht, suffices : g = 1, { contradiction }, - have hgt := ht ⟨g, by { change g ∈ subgroup.gpowers g, exact subgroup.mem_gpowers g }⟩, + have hgt := ht ⟨g, by { change g ∈ subgroup.zpowers g, exact subgroup.mem_zpowers g }⟩, rw [←ht 1] at hgt, - change (⟨_, _⟩ : subgroup.gpowers g) = ⟨_, _⟩ at hgt, + change (⟨_, _⟩ : subgroup.zpowers g) = ⟨_, _⟩ at hgt, simpa using hgt }, { use g, intro x, @@ -119,9 +119,9 @@ lemma is_cyclic_of_prime_card {α : Type u} [group α] [fintype α] {p : ℕ} [h exact subgroup.mem_top _ } end⟩ -lemma order_of_eq_card_of_forall_mem_gpowers [fintype α] - {g : α} (hx : ∀ x, x ∈ gpowers g) : order_of g = fintype.card α := -by { classical, rw [← fintype.card_congr (equiv.set.univ α), order_eq_card_gpowers], +lemma order_of_eq_card_of_forall_mem_zpowers [fintype α] + {g : α} (hx : ∀ x, x ∈ zpowers g) : order_of g = fintype.card α := +by { classical, rw [← fintype.card_congr (equiv.set.univ α), order_eq_card_zpowers], simp [hx], apply fintype.card_of_finset', simp, intro x, exact hx x} instance bot.is_cyclic {α : Type u} [group α] : is_cyclic (⊥ : subgroup α) := @@ -135,26 +135,26 @@ if hx : ∃ (x : α), x ∈ H ∧ x ≠ (1 : α) then let ⟨k, hk⟩ := hg x in have hex : ∃ n : ℕ, 0 < n ∧ g ^ n ∈ H, from ⟨k.nat_abs, nat.pos_of_ne_zero - (λ h, hx₂ $ by rw [← hk, int.eq_zero_of_nat_abs_eq_zero h, gpow_zero]), + (λ h, hx₂ $ by rw [← hk, int.eq_zero_of_nat_abs_eq_zero h, zpow_zero]), match k, hk with - | (k : ℕ), hk := by rw [int.nat_abs_of_nat, ← gpow_coe_nat, hk]; exact hx₁ + | (k : ℕ), hk := by rw [int.nat_abs_of_nat, ← zpow_coe_nat, hk]; exact hx₁ | -[1+ k], hk := by rw [int.nat_abs_of_neg_succ_of_nat, ← subgroup.inv_mem_iff H]; simp * at * end⟩, ⟨⟨⟨g ^ nat.find hex, (nat.find_spec hex).2⟩, λ ⟨x, hx⟩, let ⟨k, hk⟩ := hg x in - have hk₁ : g ^ ((nat.find hex : ℤ) * (k / nat.find hex)) ∈ gpowers (g ^ nat.find hex), - from ⟨k / nat.find hex, by rw [← gpow_coe_nat, gpow_mul]⟩, + have hk₁ : g ^ ((nat.find hex : ℤ) * (k / nat.find hex)) ∈ zpowers (g ^ nat.find hex), + from ⟨k / nat.find hex, by rw [← zpow_coe_nat, zpow_mul]⟩, have hk₂ : g ^ ((nat.find hex : ℤ) * (k / nat.find hex)) ∈ H, - by { rw gpow_mul, apply H.gpow_mem, exact_mod_cast (nat.find_spec hex).2 }, + by { rw zpow_mul, apply H.zpow_mem, exact_mod_cast (nat.find_spec hex).2 }, have hk₃ : g ^ (k % nat.find hex) ∈ H, from (subgroup.mul_mem_cancel_right H hk₂).1 $ - by rw [← gpow_add, int.mod_add_div, hk]; exact hx, + by rw [← zpow_add, int.mod_add_div, hk]; exact hx, have hk₄ : k % nat.find hex = (k % nat.find hex).nat_abs, by rw int.nat_abs_of_nonneg (int.mod_nonneg _ (int.coe_nat_ne_zero_iff_pos.2 (nat.find_spec hex).1)), have hk₅ : g ^ (k % nat.find hex ).nat_abs ∈ H, - by rwa [← gpow_coe_nat, ← hk₄], + by rwa [← zpow_coe_nat, ← hk₄], have hk₆ : (k % (nat.find hex : ℤ)).nat_abs = 0, from by_contradiction (λ h, nat.find_min hex (int.coe_nat_lt.1 $ by rw [← hk₄]; @@ -162,7 +162,7 @@ if hx : ∃ (x : α), x ∈ H ∧ x ≠ (1 : α) then ⟨nat.pos_of_ne_zero h, hk₅⟩), ⟨k / (nat.find hex : ℤ), subtype.ext_iff_val.2 begin suffices : g ^ ((nat.find hex : ℤ) * (k / nat.find hex)) = x, - { simpa [gpow_mul] }, + { simpa [zpow_mul] }, rw [int.mul_div_cancel' (int.dvd_of_mod_eq_zero (int.eq_zero_of_nat_abs_eq_zero hk₆)), hk] end⟩⟩⟩ else @@ -180,17 +180,17 @@ lemma is_cyclic.card_pow_eq_one_le [decidable_eq α] [fintype α] [is_cyclic α] (hn0 : 0 < n) : (univ.filter (λ a : α, a ^ n = 1)).card ≤ n := let ⟨g, hg⟩ := is_cyclic.exists_generator α in calc (univ.filter (λ a : α, a ^ n = 1)).card - ≤ ((gpowers (g ^ (fintype.card α / (gcd n (fintype.card α))))) : set α).to_finset.card : + ≤ ((zpowers (g ^ (fintype.card α / (gcd n (fintype.card α))))) : set α).to_finset.card : card_le_of_subset (λ x hx, let ⟨m, hm⟩ := show x ∈ submonoid.powers g, - from mem_powers_iff_mem_gpowers.2 $ hg x in + from mem_powers_iff_mem_zpowers.2 $ hg x in set.mem_to_finset.2 ⟨(m / (fintype.card α / (gcd n (fintype.card α))) : ℕ), have hgmn : g ^ (m * gcd n (fintype.card α)) = 1, by rw [pow_mul, hm, ← pow_gcd_card_eq_one_iff]; exact (mem_filter.1 hx).2, begin - rw [gpow_coe_nat, ← pow_mul, nat.mul_div_cancel_left', hm], + rw [zpow_coe_nat, ← pow_mul, nat.mul_div_cancel_left', hm], refine dvd_of_mul_dvd_mul_right (gcd_pos_of_pos_left (fintype.card α) hn0) _, conv {to_lhs, - rw [nat.div_mul_cancel (gcd_dvd_right _ _), ← order_of_eq_card_of_forall_mem_gpowers hg]}, + rw [nat.div_mul_cancel (gcd_dvd_right _ _), ← order_of_eq_card_of_forall_mem_zpowers hg]}, exact order_of_dvd_of_pow_eq_one hgmn end⟩) ... ≤ n : @@ -198,8 +198,8 @@ calc (univ.filter (λ a : α, a ^ n = 1)).card have hm0 : 0 < m, from nat.pos_of_ne_zero $ λ hm0, by { rw [hm0, mul_zero, fintype.card_eq_zero_iff] at hm, exact hm.elim' 1 }, begin - rw [← fintype.card_of_finset' _ (λ _, set.mem_to_finset), ← order_eq_card_gpowers, - order_of_pow g, order_of_eq_card_of_forall_mem_gpowers hg], + rw [← fintype.card_of_finset' _ (λ _, set.mem_to_finset), ← order_eq_card_zpowers, + order_of_pow g, order_of_eq_card_of_forall_mem_zpowers hg], rw [hm] {occs := occurrences.pos [2,3]}, rw [nat.mul_div_cancel_left _ (gcd_pos_of_pos_left _ hn0), gcd_mul_left_left, hm, nat.mul_div_cancel _ hm0], @@ -210,13 +210,13 @@ end classical lemma is_cyclic.exists_monoid_generator [fintype α] [is_cyclic α] : ∃ x : α, ∀ y : α, y ∈ submonoid.powers x := -by { simp_rw [mem_powers_iff_mem_gpowers], exact is_cyclic.exists_generator α } +by { simp_rw [mem_powers_iff_mem_zpowers], exact is_cyclic.exists_generator α } section variables [decidable_eq α] [fintype α] -lemma is_cyclic.image_range_order_of (ha : ∀ x : α, x ∈ gpowers a) : +lemma is_cyclic.image_range_order_of (ha : ∀ x : α, x ∈ zpowers a) : finset.image (λ i, a ^ i) (range (order_of a)) = univ := begin simp_rw [←set_like.mem_coe] at ha, @@ -224,9 +224,9 @@ begin convert set.to_finset_univ end -lemma is_cyclic.image_range_card (ha : ∀ x : α, x ∈ gpowers a) : +lemma is_cyclic.image_range_card (ha : ∀ x : α, x ∈ zpowers a) : finset.image (λ i, a ^ i) (range (fintype.card α)) = univ := -by rw [← order_of_eq_card_of_forall_mem_gpowers ha, is_cyclic.image_range_order_of ha] +by rw [← order_of_eq_card_of_forall_mem_zpowers ha, is_cyclic.image_range_order_of ha] end @@ -241,15 +241,15 @@ lemma card_pow_eq_one_eq_order_of_aux (a : α) : (finset.univ.filter (λ b : α, b ^ order_of a = 1)).card = order_of a := le_antisymm (hn _ (order_of_pos a)) - (calc order_of a = @fintype.card (gpowers a) (id _) : order_eq_card_gpowers + (calc order_of a = @fintype.card (zpowers a) (id _) : order_eq_card_zpowers ... ≤ @fintype.card (↑(univ.filter (λ b : α, b ^ order_of a = 1)) : set α) (fintype.of_finset _ (λ _, iff.rfl)) : - @fintype.card_le_of_injective (gpowers a) + @fintype.card_le_of_injective (zpowers a) (↑(univ.filter (λ b : α, b ^ order_of a = 1)) : set α) (id _) (id _) (λ b, ⟨b.1, mem_filter.2 ⟨mem_univ _, let ⟨i, hi⟩ := b.2 in - by rw [← hi, ← gpow_coe_nat, ← gpow_mul, mul_comm, gpow_mul, gpow_coe_nat, - pow_order_of_eq_one, one_gpow]⟩⟩) (λ _ _ h, subtype.eq (subtype.mk.inj h)) + by rw [← hi, ← zpow_coe_nat, ← zpow_mul, mul_comm, zpow_mul, zpow_coe_nat, + pow_order_of_eq_one, one_zpow]⟩⟩) (λ _ _ h, subtype.eq (subtype.mk.inj h)) ... = (univ.filter (λ b : α, b ^ order_of a = 1)).card : fintype.card_of_finset _ _) open_locale nat -- use φ for nat.totient @@ -375,16 +375,16 @@ variables {G : Type*} {H : Type*} [group G] [group H] Also see `comm_group_of_cycle_center_quotient` for the `comm_group` instance -/ lemma commutative_of_cyclic_center_quotient [is_cyclic H] (f : G →* H) (hf : f.ker ≤ center G) (a b : G) : a * b = b * a := -let ⟨⟨x, y, (hxy : f y = x)⟩, (hx : ∀ a : f.range, a ∈ gpowers _)⟩ := +let ⟨⟨x, y, (hxy : f y = x)⟩, (hx : ∀ a : f.range, a ∈ zpowers _)⟩ := is_cyclic.exists_generator f.range in let ⟨m, hm⟩ := hx ⟨f a, a, rfl⟩ in let ⟨n, hn⟩ := hx ⟨f b, b, rfl⟩ in have hm : x ^ m = f a, by simpa [subtype.ext_iff] using hm, have hn : x ^ n = f b, by simpa [subtype.ext_iff] using hn, have ha : y ^ (-m) * a ∈ center G, - from hf (by rw [f.mem_ker, f.map_mul, f.map_gpow, hxy, gpow_neg, hm, inv_mul_self]), + from hf (by rw [f.mem_ker, f.map_mul, f.map_zpow, hxy, zpow_neg, hm, inv_mul_self]), have hb : y ^ (-n) * b ∈ center G, - from hf (by rw [f.mem_ker, f.map_mul, f.map_gpow, hxy, gpow_neg, hn, inv_mul_self]), + from hf (by rw [f.mem_ker, f.map_mul, f.map_zpow, hxy, zpow_neg, hn, inv_mul_self]), calc a * b = y ^ m * ((y ^ (-m) * a) * y ^ n) * (y ^ (-n) * b) : by simp [mul_assoc] ... = y ^ m * (y ^ n * (y ^ (-m) * a)) * (y ^ (-n) * b) : by rw [mem_center_iff.1 ha] ... = y ^ m * y ^ n * y ^ (-m) * (a * (y ^ (-n) * b)) : by simp [mul_assoc] @@ -411,11 +411,11 @@ begin { apply is_cyclic_of_subsingleton }, { obtain ⟨g, hg⟩ := exists_ne (1 : α), refine ⟨⟨g, λ x, _⟩⟩, - cases is_simple_lattice.eq_bot_or_eq_top (subgroup.gpowers g) with hb ht, + cases is_simple_lattice.eq_bot_or_eq_top (subgroup.zpowers g) with hb ht, { exfalso, apply hg, rw [← subgroup.mem_bot, ← hb], - apply subgroup.mem_gpowers }, + apply subgroup.mem_zpowers }, { rw ht, apply subgroup.mem_top } } end @@ -425,11 +425,11 @@ begin have h0 : 0 < fintype.card α := fintype.card_pos_iff.2 (by apply_instance), obtain ⟨g, hg⟩ := is_cyclic.exists_generator α, refine ⟨fintype.one_lt_card_iff_nontrivial.2 infer_instance, λ n hn, _⟩, - refine (is_simple_lattice.eq_bot_or_eq_top (subgroup.gpowers (g ^ n))).symm.imp _ _, + refine (is_simple_lattice.eq_bot_or_eq_top (subgroup.zpowers (g ^ n))).symm.imp _ _, { intro h, have hgo := order_of_pow g, - rw [order_of_eq_card_of_forall_mem_gpowers hg, nat.gcd_eq_right_iff_dvd.1 hn, - order_of_eq_card_of_forall_mem_gpowers, eq_comm, + rw [order_of_eq_card_of_forall_mem_zpowers hg, nat.gcd_eq_right_iff_dvd.1 hn, + order_of_eq_card_of_forall_mem_zpowers, eq_comm, nat.div_eq_iff_eq_mul_left (nat.pos_of_dvd_of_pos hn h0) hn] at hgo, { exact (mul_left_cancel₀ (ne_of_gt h0) ((mul_one (fintype.card α)).trans hgo)).symm }, { intro x, @@ -437,10 +437,10 @@ begin exact subgroup.mem_top _ } }, { intro h, apply le_antisymm (nat.le_of_dvd h0 hn), - rw ← order_of_eq_card_of_forall_mem_gpowers hg, + rw ← order_of_eq_card_of_forall_mem_zpowers hg, apply order_of_le_of_pow_eq_one (nat.pos_of_dvd_of_pos hn h0), rw [← subgroup.mem_bot, ← h], - exact subgroup.mem_gpowers _ } + exact subgroup.mem_zpowers _ } end end comm_group diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 37e094f8af23c..92fea322c427b 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -303,9 +303,9 @@ K.to_submonoid.prod_mem h lemma pow_mem {x : G} (hx : x ∈ K) : ∀ n : ℕ, x ^ n ∈ K := K.to_submonoid.pow_mem hx @[to_additive] -lemma gpow_mem {x : G} (hx : x ∈ K) : ∀ n : ℤ, x ^ n ∈ K -| (n : ℕ) := by { rw [gpow_coe_nat], exact pow_mem _ hx n } -| -[1+ n] := by { rw [gpow_neg_succ_of_nat], exact K.inv_mem (K.pow_mem hx n.succ) } +lemma zpow_mem {x : G} (hx : x ∈ K) : ∀ n : ℤ, x ^ n ∈ K +| (n : ℕ) := by { rw [zpow_coe_nat], exact pow_mem _ hx n } +| -[1+ n] := by { rw [zpow_neg_succ_of_nat], exact K.inv_mem (K.pow_mem hx n.succ) } /-- Construct a subgroup from a nonempty set that is closed under division. -/ @[to_additive "Construct a subgroup from a nonempty set that is closed under subtraction"] @@ -370,8 +370,8 @@ def subtype : H →* G := ⟨coe, rfl, λ _ _, rfl⟩ @[simp, norm_cast] lemma coe_pow (x : H) (n : ℕ) : ((x ^ n : H) : G) = x ^ n := coe_subtype H ▸ monoid_hom.map_pow _ _ _ -@[simp, norm_cast] lemma coe_gpow (x : H) (n : ℤ) : ((x ^ n : H) : G) = x ^ n := -coe_subtype H ▸ monoid_hom.map_gpow _ _ _ +@[simp, norm_cast] lemma coe_zpow (x : H) (n : ℤ) : ((x ^ n : H) : G) = x ^ n := +coe_subtype H ▸ monoid_hom.map_zpow _ _ _ /-- The inclusion homomorphism from a subgroup `H` contained in `K` to `K`. -/ @[to_additive "The inclusion homomorphism from a additive subgroup `H` contained in `K` to `K`."] @@ -686,15 +686,15 @@ by { rw [← le_bot_iff], exact closure_le _} lemma mem_closure_singleton {x y : G} : y ∈ closure ({x} : set G) ↔ ∃ n : ℤ, x ^ n = y := begin refine ⟨λ hy, closure_induction hy _ _ _ _, - λ ⟨n, hn⟩, hn ▸ gpow_mem _ (subset_closure $ mem_singleton x) n⟩, + λ ⟨n, hn⟩, hn ▸ zpow_mem _ (subset_closure $ mem_singleton x) n⟩, { intros y hy, rw [eq_of_mem_singleton hy], - exact ⟨1, gpow_one x⟩ }, - { exact ⟨0, gpow_zero x⟩ }, + exact ⟨1, zpow_one x⟩ }, + { exact ⟨0, zpow_zero x⟩ }, { rintros _ _ ⟨n, rfl⟩ ⟨m, rfl⟩, - exact ⟨n + m, gpow_add x n m⟩ }, + exact ⟨n + m, zpow_add x n m⟩ }, rintros _ ⟨n, rfl⟩, - exact ⟨-n, gpow_neg x n⟩ + exact ⟨-n, zpow_neg x n⟩ end lemma closure_singleton_one : closure ({1} : set G) = ⊥ := @@ -1439,7 +1439,7 @@ coe_subtype H ▸ add_monoid_hom.map_nsmul _ _ _ coe_subtype H ▸ add_monoid_hom.map_gsmul _ _ _ attribute [to_additive add_subgroup.coe_smul] subgroup.coe_pow -attribute [to_additive add_subgroup.coe_gsmul] subgroup.coe_gpow +attribute [to_additive add_subgroup.coe_gsmul] subgroup.coe_zpow end add_subgroup @@ -2016,37 +2016,37 @@ end⟩ namespace subgroup /-- The subgroup generated by an element. -/ -def gpowers (g : G) : subgroup G := -subgroup.copy (gpowers_hom G g).range (set.range ((^) g : ℤ → G)) rfl +def zpowers (g : G) : subgroup G := +subgroup.copy (zpowers_hom G g).range (set.range ((^) g : ℤ → G)) rfl -@[simp] lemma mem_gpowers (g : G) : g ∈ gpowers g := ⟨1, gpow_one _⟩ +@[simp] lemma mem_zpowers (g : G) : g ∈ zpowers g := ⟨1, zpow_one _⟩ -lemma gpowers_eq_closure (g : G) : gpowers g = closure {g} := +lemma zpowers_eq_closure (g : G) : zpowers g = closure {g} := by { ext, exact mem_closure_singleton.symm } -@[simp] lemma range_gpowers_hom (g : G) : (gpowers_hom G g).range = gpowers g := rfl +@[simp] lemma range_zpowers_hom (g : G) : (zpowers_hom G g).range = zpowers g := rfl -lemma gpowers_subset {a : G} {K : subgroup G} (h : a ∈ K) : gpowers a ≤ K := -λ x hx, match x, hx with _, ⟨i, rfl⟩ := K.gpow_mem h i end +lemma zpowers_subset {a : G} {K : subgroup G} (h : a ∈ K) : zpowers a ≤ K := +λ x hx, match x, hx with _, ⟨i, rfl⟩ := K.zpow_mem h i end -lemma mem_gpowers_iff {g h : G} : - h ∈ gpowers g ↔ ∃ (k : ℤ), g ^ k = h := +lemma mem_zpowers_iff {g h : G} : + h ∈ zpowers g ↔ ∃ (k : ℤ), g ^ k = h := iff.rfl -@[simp] lemma forall_gpowers {x : G} {p : gpowers x → Prop} : +@[simp] lemma forall_zpowers {x : G} {p : zpowers x → Prop} : (∀ g, p g) ↔ ∀ m : ℤ, p ⟨x ^ m, m, rfl⟩ := set.forall_subtype_range_iff -@[simp] lemma exists_gpowers {x : G} {p : gpowers x → Prop} : +@[simp] lemma exists_zpowers {x : G} {p : zpowers x → Prop} : (∃ g, p g) ↔ ∃ m : ℤ, p ⟨x ^ m, m, rfl⟩ := set.exists_subtype_range_iff -lemma forall_mem_gpowers {x : G} {p : G → Prop} : - (∀ g ∈ gpowers x, p g) ↔ ∀ m : ℤ, p (x ^ m) := +lemma forall_mem_zpowers {x : G} {p : G → Prop} : + (∀ g ∈ zpowers x, p g) ↔ ∀ m : ℤ, p (x ^ m) := set.forall_range_iff -lemma exists_mem_gpowers {x : G} {p : G → Prop} : - (∃ g ∈ gpowers x, p g) ↔ ∃ m : ℤ, p (x ^ m) := +lemma exists_mem_zpowers {x : G} {p : G → Prop} : + (∃ g ∈ zpowers x, p g) ↔ ∃ m : ℤ, p (x ^ m) := set.exists_range_iff end subgroup @@ -2060,18 +2060,18 @@ add_subgroup.copy (gmultiples_hom A a).range (set.range ((• a) : ℤ → A)) r @[simp] lemma range_gmultiples_hom (a : A) : (gmultiples_hom A a).range = gmultiples a := rfl lemma gmultiples_subset {a : A} {B : add_subgroup A} (h : a ∈ B) : gmultiples a ≤ B := -@subgroup.gpowers_subset (multiplicative A) _ _ (B.to_subgroup) h - -attribute [to_additive add_subgroup.gmultiples] subgroup.gpowers -attribute [to_additive add_subgroup.mem_gmultiples] subgroup.mem_gpowers -attribute [to_additive add_subgroup.gmultiples_eq_closure] subgroup.gpowers_eq_closure -attribute [to_additive add_subgroup.range_gmultiples_hom] subgroup.range_gpowers_hom -attribute [to_additive add_subgroup.gmultiples_subset] subgroup.gpowers_subset -attribute [to_additive add_subgroup.mem_gmultiples_iff] subgroup.mem_gpowers_iff -attribute [to_additive add_subgroup.forall_gmultiples] subgroup.forall_gpowers -attribute [to_additive add_subgroup.forall_mem_gmultiples] subgroup.forall_mem_gpowers -attribute [to_additive add_subgroup.exists_gmultiples] subgroup.exists_gpowers -attribute [to_additive add_subgroup.exists_mem_gmultiples] subgroup.exists_mem_gpowers +@subgroup.zpowers_subset (multiplicative A) _ _ (B.to_subgroup) h + +attribute [to_additive add_subgroup.gmultiples] subgroup.zpowers +attribute [to_additive add_subgroup.mem_gmultiples] subgroup.mem_zpowers +attribute [to_additive add_subgroup.gmultiples_eq_closure] subgroup.zpowers_eq_closure +attribute [to_additive add_subgroup.range_gmultiples_hom] subgroup.range_zpowers_hom +attribute [to_additive add_subgroup.gmultiples_subset] subgroup.zpowers_subset +attribute [to_additive add_subgroup.mem_gmultiples_iff] subgroup.mem_zpowers_iff +attribute [to_additive add_subgroup.forall_gmultiples] subgroup.forall_zpowers +attribute [to_additive add_subgroup.forall_mem_gmultiples] subgroup.forall_mem_zpowers +attribute [to_additive add_subgroup.exists_gmultiples] subgroup.exists_zpowers +attribute [to_additive add_subgroup.exists_mem_gmultiples] subgroup.exists_mem_zpowers end add_subgroup @@ -2079,27 +2079,27 @@ lemma int.mem_gmultiples_iff {a b : ℤ} : b ∈ add_subgroup.gmultiples a ↔ a ∣ b := exists_congr (λ k, by rw [mul_comm, eq_comm, ← smul_eq_mul]) -lemma of_mul_image_gpowers_eq_gmultiples_of_mul { x : G } : - additive.of_mul '' ((subgroup.gpowers x) : set G) = add_subgroup.gmultiples (additive.of_mul x) := +lemma of_mul_image_zpowers_eq_gmultiples_of_mul { x : G } : + additive.of_mul '' ((subgroup.zpowers x) : set G) = add_subgroup.gmultiples (additive.of_mul x) := begin ext y, split, { rintro ⟨z, ⟨m, hm⟩, hz2⟩, use m, simp only, - rwa [← of_mul_gpow, hm] }, + rwa [← of_mul_zpow, hm] }, { rintros ⟨n, hn⟩, refine ⟨x ^ n, ⟨n, rfl⟩, _⟩, - rwa of_mul_gpow } + rwa of_mul_zpow } end -lemma of_add_image_gmultiples_eq_gpowers_of_add {x : A} : +lemma of_add_image_gmultiples_eq_zpowers_of_add {x : A} : multiplicative.of_add '' ((add_subgroup.gmultiples x) : set A) = - subgroup.gpowers (multiplicative.of_add x) := + subgroup.zpowers (multiplicative.of_add x) := begin symmetry, rw equiv.eq_image_iff_symm_image_eq, - exact of_mul_image_gpowers_eq_gmultiples_of_mul, + exact of_mul_image_zpowers_eq_gmultiples_of_mul, end namespace mul_equiv @@ -2474,19 +2474,19 @@ def saturated (H : subgroup G) : Prop := ∀ ⦃n g⦄, npow n g ∈ H → n = 0 @[to_additive] lemma saturated_iff_npow {H : subgroup G} : saturated H ↔ (∀ (n : ℕ) (g : G), g^n ∈ H → n = 0 ∨ g ∈ H) := iff.rfl -@[to_additive] lemma saturated_iff_gpow {H : subgroup G} : +@[to_additive] lemma saturated_iff_zpow {H : subgroup G} : saturated H ↔ (∀ (n : ℤ) (g : G), g^n ∈ H → n = 0 ∨ g ∈ H) := begin split, { rintros hH ⟨n⟩ g hgn, - { simp only [int.coe_nat_eq_zero, int.of_nat_eq_coe, gpow_coe_nat] at hgn ⊢, + { simp only [int.coe_nat_eq_zero, int.of_nat_eq_coe, zpow_coe_nat] at hgn ⊢, exact hH hgn }, { suffices : g ^ (n+1) ∈ H, { refine (hH this).imp _ id, simp only [forall_false_left, nat.succ_ne_zero], }, - simpa only [inv_mem_iff, gpow_neg_succ_of_nat] using hgn, } }, + simpa only [inv_mem_iff, zpow_neg_succ_of_nat] using hgn, } }, { intros h n g hgn, specialize h n g, - simp only [int.coe_nat_eq_zero, gpow_coe_nat] at h, + simp only [int.coe_nat_eq_zero, zpow_coe_nat] at h, apply h hgn } end diff --git a/src/group_theory/sylow.lean b/src/group_theory/sylow.lean index d39e2ef52a224..ce3a39ea6a19a 100644 --- a/src/group_theory/sylow.lean +++ b/src/group_theory/sylow.lean @@ -340,17 +340,17 @@ have hequiv : H ≃ (subgroup.comap ((normalizer H).subtype : normalizer H →* ⟨λ a, ⟨⟨a.1, le_normalizer a.2⟩, a.2⟩, λ a, ⟨a.1.1, a.2⟩, λ ⟨_, _⟩, rfl, λ ⟨⟨_, _⟩, _⟩, rfl⟩, ⟨subgroup.map ((normalizer H).subtype) (subgroup.comap - (quotient_group.mk' (comap H.normalizer.subtype H)) (gpowers x)), + (quotient_group.mk' (comap H.normalizer.subtype H)) (zpowers x)), begin show card ↥(map H.normalizer.subtype - (comap (mk' (comap H.normalizer.subtype H)) (subgroup.gpowers x))) = p ^ (n + 1), + (comap (mk' (comap H.normalizer.subtype H)) (subgroup.zpowers x))) = p ^ (n + 1), suffices : card ↥(subtype.val '' ((subgroup.comap (mk' (comap H.normalizer.subtype H)) - (gpowers x)) : set (↥(H.normalizer)))) = p^(n+1), + (zpowers x)) : set (↥(H.normalizer)))) = p^(n+1), { convert this using 2 }, rw [set.card_image_of_injective - (subgroup.comap (mk' (comap H.normalizer.subtype H)) (gpowers x) : set (H.normalizer)) + (subgroup.comap (mk' (comap H.normalizer.subtype H)) (zpowers x) : set (H.normalizer)) subtype.val_injective, - pow_succ', ← hH, fintype.card_congr hequiv, ← hx, order_eq_card_gpowers, + pow_succ', ← hH, fintype.card_congr hequiv, ← hx, order_eq_card_zpowers, ← fintype.card_prod], exact @fintype.card_congr _ _ (id _) (id _) (preimage_mk_equiv_subgroup_times_set _ _) end, @@ -358,7 +358,7 @@ begin assume y hy, simp only [exists_prop, subgroup.coe_subtype, mk'_apply, subgroup.mem_map, subgroup.mem_comap], refine ⟨⟨y, le_normalizer hy⟩, ⟨0, _⟩, rfl⟩, - rw [gpow_zero, eq_comm, quotient_group.eq_one_iff], + rw [zpow_zero, eq_comm, quotient_group.eq_one_iff], simpa using hy end⟩ diff --git a/src/linear_algebra/matrix/fpow.lean b/src/linear_algebra/matrix/zpow.lean similarity index 56% rename from src/linear_algebra/matrix/fpow.lean rename to src/linear_algebra/matrix/zpow.lean index 5ecce7134a849..aa93ac668d7d0 100644 --- a/src/linear_algebra/matrix/fpow.lean +++ b/src/linear_algebra/matrix/zpow.lean @@ -14,7 +14,7 @@ the nonsingular inverse definition for negative powers. ## Implementation details The main definition is a direct recursive call on the integer inductive type, -as provided by the `div_inv_monoid.gpow` default implementation. +as provided by the `div_inv_monoid.zpow` default implementation. The lemma names are taken from `algebra.group_with_zero.power`. ## Tags @@ -75,120 +75,120 @@ end nat_pow section int_pow open int -@[simp] theorem one_fpow : ∀ (n : ℤ), (1 : M) ^ n = 1 -| (n : ℕ) := by rw [gpow_coe_nat, one_pow] -| -[1+ n] := by rw [gpow_neg_succ_of_nat, one_pow, inv_one] +@[simp] theorem one_zpow : ∀ (n : ℤ), (1 : M) ^ n = 1 +| (n : ℕ) := by rw [zpow_coe_nat, one_pow] +| -[1+ n] := by rw [zpow_neg_succ_of_nat, one_pow, inv_one] -lemma zero_fpow : ∀ z : ℤ, z ≠ 0 → (0 : M) ^ z = 0 -| (n : ℕ) h := by { rw [gpow_coe_nat, zero_pow], refine lt_of_le_of_ne n.zero_le (ne.symm _), +lemma zero_zpow : ∀ z : ℤ, z ≠ 0 → (0 : M) ^ z = 0 +| (n : ℕ) h := by { rw [zpow_coe_nat, zero_pow], refine lt_of_le_of_ne n.zero_le (ne.symm _), simpa using h } | -[1+n] h := by simp [zero_pow n.zero_lt_succ] -lemma zero_fpow_eq (n : ℤ) : (0 : M) ^ n = if n = 0 then 1 else 0 := +lemma zero_zpow_eq (n : ℤ) : (0 : M) ^ n = if n = 0 then 1 else 0 := begin split_ifs with h, - { rw [h, gpow_zero] }, - { rw [zero_fpow _ h] } + { rw [h, zpow_zero] }, + { rw [zero_zpow _ h] } end -theorem inv_fpow (A : M) : ∀n:ℤ, A⁻¹ ^ n = (A ^ n)⁻¹ -| (n : ℕ) := by rw [gpow_coe_nat, gpow_coe_nat, inv_pow'] -| -[1+ n] := by rw [gpow_neg_succ_of_nat, gpow_neg_succ_of_nat, inv_pow'] +theorem inv_zpow (A : M) : ∀n:ℤ, A⁻¹ ^ n = (A ^ n)⁻¹ +| (n : ℕ) := by rw [zpow_coe_nat, zpow_coe_nat, inv_pow'] +| -[1+ n] := by rw [zpow_neg_succ_of_nat, zpow_neg_succ_of_nat, inv_pow'] -@[simp] lemma fpow_neg_one (A : M) : A ^ (-1 : ℤ) = A⁻¹ := +@[simp] lemma zpow_neg_one (A : M) : A ^ (-1 : ℤ) = A⁻¹ := begin - convert div_inv_monoid.gpow_neg' 0 A, - simp only [gpow_one, int.coe_nat_zero, int.coe_nat_succ, gpow_eq_pow, zero_add] + convert div_inv_monoid.zpow_neg' 0 A, + simp only [zpow_one, int.coe_nat_zero, int.coe_nat_succ, zpow_eq_pow, zero_add] end -theorem fpow_coe_nat (A : M) (n : ℕ) : A ^ (n : ℤ) = (A ^ n) := -gpow_coe_nat _ _ +theorem zpow_coe_nat (A : M) (n : ℕ) : A ^ (n : ℤ) = (A ^ n) := +zpow_coe_nat _ _ -@[simp] theorem fpow_neg_coe_nat (A : M) (n : ℕ) : A ^ (-n : ℤ) = (A ^ n)⁻¹ := +@[simp] theorem zpow_neg_coe_nat (A : M) (n : ℕ) : A ^ (-n : ℤ) = (A ^ n)⁻¹ := begin cases n, { simp }, - { exact div_inv_monoid.gpow_neg' _ _ } + { exact div_inv_monoid.zpow_neg' _ _ } end -lemma _root_.is_unit.det_fpow {A : M} (h : is_unit A.det) (n : ℤ) : is_unit (A ^ n).det := +lemma _root_.is_unit.det_zpow {A : M} (h : is_unit A.det) (n : ℤ) : is_unit (A ^ n).det := begin cases n, { simpa using h.pow n }, { simpa using h.pow n.succ } end -lemma is_unit_det_fpow_iff {A : M} {z : ℤ} : +lemma is_unit_det_zpow_iff {A : M} {z : ℤ} : is_unit (A ^ z).det ↔ is_unit A.det ∨ z = 0 := begin induction z using int.induction_on with z IH z IH, { simp }, - { rw [←int.coe_nat_succ, fpow_coe_nat, det_pow, is_unit_pos_pow_iff (z.zero_lt_succ), + { rw [←int.coe_nat_succ, zpow_coe_nat, det_pow, is_unit_pos_pow_iff (z.zero_lt_succ), ←int.coe_nat_zero, int.coe_nat_eq_coe_nat_iff], simp }, - { rw [←neg_add', ←int.coe_nat_succ, fpow_neg_coe_nat, is_unit_nonsing_inv_det_iff, + { rw [←neg_add', ←int.coe_nat_succ, zpow_neg_coe_nat, is_unit_nonsing_inv_det_iff, det_pow, is_unit_pos_pow_iff (z.zero_lt_succ), neg_eq_zero, ←int.coe_nat_zero, int.coe_nat_eq_coe_nat_iff], simp } end -theorem fpow_neg {A : M} (h : is_unit A.det) : ∀ (n : ℤ), A ^ -n = (A ^ n)⁻¹ -| (n : ℕ) := fpow_neg_coe_nat _ _ -| -[1+ n] := by { rw [gpow_neg_succ_of_nat, neg_neg_of_nat_succ, of_nat_eq_coe, fpow_coe_nat, +theorem zpow_neg {A : M} (h : is_unit A.det) : ∀ (n : ℤ), A ^ -n = (A ^ n)⁻¹ +| (n : ℕ) := zpow_neg_coe_nat _ _ +| -[1+ n] := by { rw [zpow_neg_succ_of_nat, neg_neg_of_nat_succ, of_nat_eq_coe, zpow_coe_nat, nonsing_inv_nonsing_inv], rw det_pow, exact h.pow _ } -lemma inv_fpow' {A : M} (h : is_unit A.det) (n : ℤ) : +lemma inv_zpow' {A : M} (h : is_unit A.det) (n : ℤ) : (A ⁻¹) ^ n = A ^ (-n) := -by rw [fpow_neg h, inv_fpow] +by rw [zpow_neg h, inv_zpow] -lemma fpow_add_one {A : M} (h : is_unit A.det) : ∀ n : ℤ, A ^ (n + 1) = A ^ n * A +lemma zpow_add_one {A : M} (h : is_unit A.det) : ∀ n : ℤ, A ^ (n + 1) = A ^ n * A | (n : ℕ) := by simp [← int.coe_nat_succ, pow_succ'] | -((n : ℕ) + 1) := calc A ^ (-(n + 1) + 1 : ℤ) - = (A ^ n)⁻¹ : by rw [neg_add, neg_add_cancel_right, fpow_neg h, fpow_coe_nat] + = (A ^ n)⁻¹ : by rw [neg_add, neg_add_cancel_right, zpow_neg h, zpow_coe_nat] ... = (A ⬝ A ^ n)⁻¹ ⬝ A : by rw [mul_inv_rev, matrix.mul_assoc, nonsing_inv_mul _ h, matrix.mul_one] ... = A ^ -(n + 1 : ℤ) * A : - by rw [fpow_neg h, ← int.coe_nat_succ, fpow_coe_nat, pow_succ, mul_eq_mul, mul_eq_mul] + by rw [zpow_neg h, ← int.coe_nat_succ, zpow_coe_nat, pow_succ, mul_eq_mul, mul_eq_mul] -lemma fpow_sub_one {A : M} (h : is_unit A.det) (n : ℤ) : A ^ (n - 1) = A ^ n * A⁻¹ := +lemma zpow_sub_one {A : M} (h : is_unit A.det) (n : ℤ) : A ^ (n - 1) = A ^ n * A⁻¹ := calc A ^ (n - 1) = A ^ (n - 1) * A * A⁻¹ : by rw [mul_assoc, mul_eq_mul A, mul_nonsing_inv _ h, mul_one] - ... = A^n * A⁻¹ : by rw [← fpow_add_one h, sub_add_cancel] + ... = A^n * A⁻¹ : by rw [← zpow_add_one h, sub_add_cancel] -lemma fpow_add {A : M} (ha : is_unit A.det) (m n : ℤ) : A ^ (m + n) = A ^ m * A ^ n := +lemma zpow_add {A : M} (ha : is_unit A.det) (m n : ℤ) : A ^ (m + n) = A ^ m * A ^ n := begin induction n using int.induction_on with n ihn n ihn, case hz : { simp }, - { simp only [← add_assoc, fpow_add_one ha, ihn, mul_assoc] }, - { rw [fpow_sub_one ha, ← mul_assoc, ← ihn, ← fpow_sub_one ha, add_sub_assoc] } + { simp only [← add_assoc, zpow_add_one ha, ihn, mul_assoc] }, + { rw [zpow_sub_one ha, ← mul_assoc, ← ihn, ← zpow_sub_one ha, add_sub_assoc] } end -lemma fpow_add_of_nonpos {A : M} {m n : ℤ} (hm : m ≤ 0) (hn : n ≤ 0) : +lemma zpow_add_of_nonpos {A : M} {m n : ℤ} (hm : m ≤ 0) (hn : n ≤ 0) : A ^ (m + n) = A ^ m * A ^ n := begin rcases nonsing_inv_cancel_or_zero A with ⟨h, h'⟩ | h, - { exact fpow_add (is_unit_det_of_left_inverse h) m n }, + { exact zpow_add (is_unit_det_of_left_inverse h) m n }, { obtain ⟨k, rfl⟩ := exists_eq_neg_of_nat hm, obtain ⟨l, rfl⟩ := exists_eq_neg_of_nat hn, - simp_rw [←neg_add, ←int.coe_nat_add, fpow_neg_coe_nat, ←inv_pow', h, pow_add] } + simp_rw [←neg_add, ←int.coe_nat_add, zpow_neg_coe_nat, ←inv_pow', h, pow_add] } end -lemma fpow_add_of_nonneg {A : M} {m n : ℤ} (hm : 0 ≤ m) (hn : 0 ≤ n) : +lemma zpow_add_of_nonneg {A : M} {m n : ℤ} (hm : 0 ≤ m) (hn : 0 ≤ n) : A ^ (m + n) = A ^ m * A ^ n := begin rcases nonsing_inv_cancel_or_zero A with ⟨h, h'⟩ | h, - { exact fpow_add (is_unit_det_of_left_inverse h) m n }, + { exact zpow_add (is_unit_det_of_left_inverse h) m n }, { obtain ⟨k, rfl⟩ := eq_coe_of_zero_le hm, obtain ⟨l, rfl⟩ := eq_coe_of_zero_le hn, - rw [←int.coe_nat_add, fpow_coe_nat, fpow_coe_nat, fpow_coe_nat, pow_add] } + rw [←int.coe_nat_add, zpow_coe_nat, zpow_coe_nat, zpow_coe_nat, pow_add] } end -theorem fpow_one_add {A : M} (h : is_unit A.det) (i : ℤ) : A ^ (1 + i) = A * A ^ i := -by rw [fpow_add h, gpow_one] +theorem zpow_one_add {A : M} (h : is_unit A.det) (i : ℤ) : A ^ (1 + i) = A * A ^ i := +by rw [zpow_add h, zpow_one] -theorem semiconj_by.fpow_right {A X Y : M} (hx : is_unit X.det) (hy : is_unit Y.det) +theorem semiconj_by.zpow_right {A X Y : M} (hx : is_unit X.det) (hy : is_unit Y.det) (h : semiconj_by A X Y) : ∀ m : ℤ, semiconj_by A (X^m) (Y^m) | (n : ℕ) := by simp [h.pow_right n] @@ -201,7 +201,7 @@ theorem semiconj_by.fpow_right {A X Y : M} (hx : is_unit X.det) (hy : is_unit Y. have hy' : is_unit (Y ^ n.succ).det, { rw det_pow, exact hy.pow n.succ }, - rw [gpow_neg_succ_of_nat, gpow_neg_succ_of_nat, nonsing_inv_apply _ hx', nonsing_inv_apply _ hy', + rw [zpow_neg_succ_of_nat, zpow_neg_succ_of_nat, nonsing_inv_apply _ hx', nonsing_inv_apply _ hy', semiconj_by], refine (is_regular_of_is_left_regular_det hy'.is_regular.left).left _, rw [←mul_assoc, ←(h.pow_right n.succ).eq, mul_assoc, mul_eq_mul (X ^ _), mul_smul, mul_adjugate, @@ -210,68 +210,68 @@ theorem semiconj_by.fpow_right {A X Y : M} (hx : is_unit X.det) (hy : is_unit Y. hy'.coe_inv_mul, one_smul, matrix.mul_one, matrix.one_mul], end -theorem commute.fpow_right {A B : M} (h : commute A B) (m : ℤ) : commute A (B^m) := +theorem commute.zpow_right {A B : M} (h : commute A B) (m : ℤ) : commute A (B^m) := begin rcases nonsing_inv_cancel_or_zero B with ⟨hB, hB'⟩ | hB, - { refine semiconj_by.fpow_right _ _ h _; + { refine semiconj_by.zpow_right _ _ h _; exact is_unit_det_of_left_inverse hB }, { cases m, { simpa using h.pow_right _ }, { simp [←inv_pow', hB] } } end -theorem commute.fpow_left {A B : M} (h : commute A B) (m : ℤ) : commute (A^m) B := -(commute.fpow_right h.symm m).symm +theorem commute.zpow_left {A B : M} (h : commute A B) (m : ℤ) : commute (A^m) B := +(commute.zpow_right h.symm m).symm -theorem commute.fpow_fpow {A B : M} (h : commute A B) (m n : ℤ) : commute (A^m) (B^n) := -commute.fpow_right (commute.fpow_left h _) _ +theorem commute.zpow_zpow {A B : M} (h : commute A B) (m n : ℤ) : commute (A^m) (B^n) := +commute.zpow_right (commute.zpow_left h _) _ -theorem commute.fpow_self (A : M) (n : ℤ) : commute (A^n) A := -commute.fpow_left (commute.refl A) _ +theorem commute.zpow_self (A : M) (n : ℤ) : commute (A^n) A := +commute.zpow_left (commute.refl A) _ -theorem commute.self_fpow (A : M) (n : ℤ) : commute A (A^n) := -commute.fpow_right (commute.refl A) _ +theorem commute.self_zpow (A : M) (n : ℤ) : commute A (A^n) := +commute.zpow_right (commute.refl A) _ -theorem commute.fpow_fpow_self (A : M) (m n : ℤ) : commute (A^m) (A^n) := -commute.fpow_fpow (commute.refl A) _ _ +theorem commute.zpow_zpow_self (A : M) (m n : ℤ) : commute (A^m) (A^n) := +commute.zpow_zpow (commute.refl A) _ _ -theorem fpow_bit0 (A : M) (n : ℤ) : A ^ bit0 n = A ^ n * A ^ n := +theorem zpow_bit0 (A : M) (n : ℤ) : A ^ bit0 n = A ^ n * A ^ n := begin cases le_total 0 n with nonneg nonpos, - { exact fpow_add_of_nonneg nonneg nonneg }, - { exact fpow_add_of_nonpos nonpos nonpos } + { exact zpow_add_of_nonneg nonneg nonneg }, + { exact zpow_add_of_nonpos nonpos nonpos } end -lemma fpow_add_one_of_ne_neg_one {A : M} : ∀ (n : ℤ), n ≠ -1 → A ^ (n + 1) = A ^ n * A +lemma zpow_add_one_of_ne_neg_one {A : M} : ∀ (n : ℤ), n ≠ -1 → A ^ (n + 1) = A ^ n * A | (n : ℕ) _ := by simp [← int.coe_nat_succ, pow_succ'] | (-1) h := absurd rfl h | (-((n : ℕ) + 2)) _ := begin rcases nonsing_inv_cancel_or_zero A with ⟨h, h'⟩ | h, - { apply fpow_add_one (is_unit_det_of_left_inverse h) }, + { apply zpow_add_one (is_unit_det_of_left_inverse h) }, { show A ^ (-((n + 1 : ℕ) : ℤ)) = A ^ -((n + 2 : ℕ) : ℤ) * A, - simp_rw [fpow_neg_coe_nat, ←inv_pow', h, zero_pow nat.succ_pos', zero_mul] } + simp_rw [zpow_neg_coe_nat, ←inv_pow', h, zero_pow nat.succ_pos', zero_mul] } end -theorem fpow_bit1 (A : M) (n : ℤ) : A ^ bit1 n = A ^ n * A ^ n * A := +theorem zpow_bit1 (A : M) (n : ℤ) : A ^ bit1 n = A ^ n * A ^ n * A := begin - rw [bit1, fpow_add_one_of_ne_neg_one, fpow_bit0], + rw [bit1, zpow_add_one_of_ne_neg_one, zpow_bit0], intro h, simpa using congr_arg bodd h end -theorem fpow_mul (A : M) (h : is_unit A.det) : ∀ m n : ℤ, A ^ (m * n) = (A ^ m) ^ n -| (m : ℕ) (n : ℕ) := by rw [gpow_coe_nat, gpow_coe_nat, ← pow_mul, ← gpow_coe_nat, int.coe_nat_mul] -| (m : ℕ) -[1+ n] := by rw [gpow_coe_nat, gpow_neg_succ_of_nat, ← pow_mul, coe_nat_mul_neg_succ, - ←int.coe_nat_mul, fpow_neg_coe_nat] -| -[1+ m] (n : ℕ) := by rw [gpow_coe_nat, gpow_neg_succ_of_nat, ← inv_pow', ← pow_mul, - neg_succ_mul_coe_nat, ←int.coe_nat_mul, fpow_neg_coe_nat, inv_pow'] -| -[1+ m] -[1+ n] := by { rw [gpow_neg_succ_of_nat, gpow_neg_succ_of_nat, neg_succ_mul_neg_succ, - ←int.coe_nat_mul, fpow_coe_nat, inv_pow', ←pow_mul, nonsing_inv_nonsing_inv], +theorem zpow_mul (A : M) (h : is_unit A.det) : ∀ m n : ℤ, A ^ (m * n) = (A ^ m) ^ n +| (m : ℕ) (n : ℕ) := by rw [zpow_coe_nat, zpow_coe_nat, ← pow_mul, ← zpow_coe_nat, int.coe_nat_mul] +| (m : ℕ) -[1+ n] := by rw [zpow_coe_nat, zpow_neg_succ_of_nat, ← pow_mul, coe_nat_mul_neg_succ, + ←int.coe_nat_mul, zpow_neg_coe_nat] +| -[1+ m] (n : ℕ) := by rw [zpow_coe_nat, zpow_neg_succ_of_nat, ← inv_pow', ← pow_mul, + neg_succ_mul_coe_nat, ←int.coe_nat_mul, zpow_neg_coe_nat, inv_pow'] +| -[1+ m] -[1+ n] := by { rw [zpow_neg_succ_of_nat, zpow_neg_succ_of_nat, neg_succ_mul_neg_succ, + ←int.coe_nat_mul, zpow_coe_nat, inv_pow', ←pow_mul, nonsing_inv_nonsing_inv], rw det_pow, exact h.pow _ } -theorem fpow_mul' (A : M) (h : is_unit A.det) (m n : ℤ) : A ^ (m * n) = (A ^ n) ^ m := -by rw [mul_comm, fpow_mul _ h] +theorem zpow_mul' (A : M) (h : is_unit A.det) (m n : ℤ) : A ^ (m * n) = (A ^ n) ^ m := +by rw [mul_comm, zpow_mul _ h] @[simp, norm_cast] lemma units.coe_inv'' (u : units M) : ((u⁻¹ : units M) : M) = u⁻¹ := @@ -280,48 +280,48 @@ begin rw [←mul_eq_mul, ←units.coe_mul, inv_mul_self, units.coe_one] end -@[simp, norm_cast] lemma units.coe_fpow (u : units M) : +@[simp, norm_cast] lemma units.coe_zpow (u : units M) : ∀ (n : ℤ), ((u ^ n : units M) : M) = u ^ n -| (n : ℕ) := by rw [gpow_coe_nat, gpow_coe_nat, units.coe_pow] -| -[1+k] := by rw [gpow_neg_succ_of_nat, gpow_neg_succ_of_nat, ←inv_pow, u⁻¹.coe_pow, ←inv_pow', +| (n : ℕ) := by rw [_root_.zpow_coe_nat, zpow_coe_nat, units.coe_pow] +| -[1+k] := by rw [zpow_neg_succ_of_nat, zpow_neg_succ_of_nat, ←inv_pow, u⁻¹.coe_pow, ←inv_pow', units.coe_inv''] -lemma fpow_ne_zero_of_is_unit_det [nonempty n'] [nontrivial R] {A : M} +lemma zpow_ne_zero_of_is_unit_det [nonempty n'] [nontrivial R] {A : M} (ha : is_unit A.det) (z : ℤ) : A ^ z ≠ 0 := begin - have := ha.det_fpow z, + have := ha.det_zpow z, contrapose! this, rw [this, det_zero ‹_›], exact not_is_unit_zero end -lemma fpow_sub {A : M} (ha : is_unit A.det) (z1 z2 : ℤ) : A ^ (z1 - z2) = A ^ z1 / A ^ z2 := -by rw [sub_eq_add_neg, fpow_add ha, fpow_neg ha, div_eq_mul_inv] +lemma zpow_sub {A : M} (ha : is_unit A.det) (z1 z2 : ℤ) : A ^ (z1 - z2) = A ^ z1 / A ^ z2 := +by rw [sub_eq_add_neg, zpow_add ha, zpow_neg ha, div_eq_mul_inv] -lemma commute.mul_fpow {A B : M} (h : commute A B) : +lemma commute.mul_zpow {A B : M} (h : commute A B) : ∀ (i : ℤ), (A * B) ^ i = (A ^ i) * (B ^ i) | (n : ℕ) := by simp [h.mul_pow n, -mul_eq_mul] -| -[1+n] := by rw [gpow_neg_succ_of_nat, gpow_neg_succ_of_nat, gpow_neg_succ_of_nat, +| -[1+n] := by rw [zpow_neg_succ_of_nat, zpow_neg_succ_of_nat, zpow_neg_succ_of_nat, mul_eq_mul (_⁻¹), ←mul_inv_rev, ←mul_eq_mul, h.mul_pow n.succ, (h.pow_pow _ _).eq] -theorem fpow_bit0' (A : M) (n : ℤ) : A ^ bit0 n = (A * A) ^ n := -(fpow_bit0 A n).trans (commute.mul_fpow (commute.refl A) n).symm +theorem zpow_bit0' (A : M) (n : ℤ) : A ^ bit0 n = (A * A) ^ n := +(zpow_bit0 A n).trans (commute.mul_zpow (commute.refl A) n).symm -theorem fpow_bit1' (A : M) (n : ℤ) : A ^ bit1 n = (A * A) ^ n * A := -by rw [fpow_bit1, commute.mul_fpow (commute.refl A)] +theorem zpow_bit1' (A : M) (n : ℤ) : A ^ bit1 n = (A * A) ^ n * A := +by rw [zpow_bit1, commute.mul_zpow (commute.refl A)] -theorem fpow_neg_mul_fpow_self (n : ℤ) {A : M} (h : is_unit A.det) : +theorem zpow_neg_mul_zpow_self (n : ℤ) {A : M} (h : is_unit A.det) : A ^ (-n) * A ^ n = 1 := -by rw [fpow_neg h, mul_eq_mul, nonsing_inv_mul _ (h.det_fpow _)] +by rw [zpow_neg h, mul_eq_mul, nonsing_inv_mul _ (h.det_zpow _)] theorem one_div_pow {A : M} (n : ℕ) : (1 / A) ^ n = 1 / A ^ n := by simp only [one_div, inv_pow'] -theorem one_div_fpow {A : M} (n : ℤ) : +theorem one_div_zpow {A : M} (n : ℤ) : (1 / A) ^ n = 1 / A ^ n := -by simp only [one_div, inv_fpow] +by simp only [one_div, inv_zpow] end int_pow diff --git a/src/measure_theory/group/arithmetic.lean b/src/measure_theory/group/arithmetic.lean index 58c0ebe94b521..77b0a851bcb61 100644 --- a/src/measure_theory/group/arithmetic.lean +++ b/src/measure_theory/group/arithmetic.lean @@ -360,17 +360,17 @@ attribute [measurability] measurable.neg ae_measurable.neg end inv /- There is something extremely strange here: copy-pasting the proof of this lemma in the proof -of `has_measurable_gpow` fails, while `pp.all` does not show any difference in the goal. +of `has_measurable_zpow` fails, while `pp.all` does not show any difference in the goal. Keep it as a separate lemmas as a workaround. -/ -private lemma has_measurable_gpow_aux (G : Type u) [div_inv_monoid G] [measurable_space G] +private lemma has_measurable_zpow_aux (G : Type u) [div_inv_monoid G] [measurable_space G] [has_measurable_mul₂ G] [has_measurable_inv G] (k : ℕ) : measurable (λ (x : G), x ^(-[1+ k])) := begin - simp_rw [gpow_neg_succ_of_nat], + simp_rw [zpow_neg_succ_of_nat], exact (measurable_id.pow_const (k + 1)).inv end -instance has_measurable_gpow (G : Type u) [div_inv_monoid G] [measurable_space G] +instance has_measurable_zpow (G : Type u) [div_inv_monoid G] [measurable_space G] [has_measurable_mul₂ G] [has_measurable_inv G] : has_measurable_pow G ℤ := begin @@ -380,7 +380,7 @@ begin dsimp, apply int.cases_on n, { simpa using measurable_id.pow_const }, - { exact has_measurable_gpow_aux G } + { exact has_measurable_zpow_aux G } end @[priority 100, to_additive] diff --git a/src/number_theory/arithmetic_function.lean b/src/number_theory/arithmetic_function.lean index 2326486d77b92..6fd49fba24697 100644 --- a/src/number_theory/arithmetic_function.lean +++ b/src/number_theory/arithmetic_function.lean @@ -872,7 +872,7 @@ begin prod_congr rfl _], intros x hx, rw [dif_pos (nat.pos_of_mem_divisors (nat.snd_mem_divisors_of_mem_antidiagonal hx)), - units.coe_hom_apply, units.coe_gpow₀, units.coe_mk0] } + units.coe_hom_apply, units.coe_zpow₀, units.coe_mk0] } end end special_functions diff --git a/src/number_theory/padics/padic_integers.lean b/src/number_theory/padics/padic_integers.lean index c6e081014b7ab..29074dc0b03e9 100644 --- a/src/number_theory/padics/padic_integers.lean +++ b/src/number_theory/padics/padic_integers.lean @@ -313,8 +313,8 @@ lemma exists_pow_neg_lt {ε : ℝ} (hε : 0 < ε) : begin obtain ⟨k, hk⟩ := exists_nat_gt ε⁻¹, use k, - rw ← inv_lt_inv hε (_root_.fpow_pos_of_pos _ _), - { rw [fpow_neg, inv_inv₀, gpow_coe_nat], + rw ← inv_lt_inv hε (_root_.zpow_pos_of_pos _ _), + { rw [zpow_neg₀, inv_inv₀, zpow_coe_nat], apply lt_of_lt_of_le hk, norm_cast, apply le_of_lt, @@ -369,7 +369,7 @@ begin by_cases hx : x = 0, { simp [hx] }, have h : (1 : ℝ) < p := by exact_mod_cast hp_prime.1.one_lt, - rw [← neg_nonpos, ← (fpow_strict_mono h).le_iff_le], + rw [← neg_nonpos, ← (zpow_strict_mono h).le_iff_le], show (p : ℝ) ^ -valuation x ≤ p ^ 0, rw [← norm_eq_pow_val hx], simpa using x.property, @@ -386,7 +386,7 @@ begin exact_mod_cast (pow_eq_zero hc) }, { exact hc } }, rwa [norm_eq_pow_val aux, norm_p_pow, norm_eq_pow_val hc, - ← fpow_add, ← neg_add, fpow_inj, neg_inj] at this, + ← zpow_add₀, ← neg_add, zpow_inj, neg_inj] at this, { exact_mod_cast hp_prime.1.pos }, { exact_mod_cast hp_prime.1.ne_one }, { exact_mod_cast hp_prime.1.ne_zero }, @@ -444,8 +444,8 @@ See `unit_coeff_spec`. -/ def unit_coeff {x : ℤ_[p]} (hx : x ≠ 0) : units ℤ_[p] := let u : ℚ_[p] := x*p^(-x.valuation) in have hu : ∥u∥ = 1, -by simp [hx, nat.fpow_ne_zero_of_pos (by exact_mod_cast hp_prime.1.pos) x.valuation, - norm_eq_pow_val, fpow_neg, inv_mul_cancel, -cast_eq_of_rat_of_nat], +by simp [hx, nat.zpow_ne_zero_of_pos (by exact_mod_cast hp_prime.1.pos) x.valuation, + norm_eq_pow_val, zpow_neg, inv_mul_cancel, -cast_eq_of_rat_of_nat], mk_units hu @[simp] lemma unit_coeff_coe {x : ℤ_[p]} (hx : x ≠ 0) : @@ -457,11 +457,11 @@ begin apply subtype.coe_injective, push_cast, have repr : (x : ℚ_[p]) = (unit_coeff hx) * p ^ x.valuation, - { rw [unit_coeff_coe, mul_assoc, ← fpow_add], + { rw [unit_coeff_coe, mul_assoc, ← zpow_add₀], { simp }, { exact_mod_cast hp_prime.1.ne_zero } }, convert repr using 2, - rw [← gpow_coe_nat, int.nat_abs_of_nonneg (valuation_nonneg x)], + rw [← zpow_coe_nat, int.nat_abs_of_nonneg (valuation_nonneg x)], end end units @@ -474,7 +474,7 @@ lemma norm_le_pow_iff_le_valuation (x : ℤ_[p]) (hx : x ≠ 0) (n : ℕ) : begin rw norm_eq_pow_val hx, lift x.valuation to ℕ using x.valuation_nonneg with k hk, - simp only [int.coe_nat_le, fpow_neg, gpow_coe_nat], + simp only [int.coe_nat_le, zpow_neg₀, zpow_coe_nat], have aux : ∀ n : ℕ, 0 < (p ^ n : ℝ), { apply pow_pos, exact_mod_cast hp_prime.1.pos }, rw [inv_le_inv (aux _) (aux _)], @@ -505,7 +505,7 @@ lemma norm_le_pow_iff_mem_span_pow (x : ℤ_[p]) (n : ℕ) : begin by_cases hx : x = 0, { subst hx, - simp only [norm_zero, fpow_neg, gpow_coe_nat, inv_nonneg, iff_true, submodule.zero_mem], + simp only [norm_zero, zpow_neg₀, zpow_coe_nat, inv_nonneg, iff_true, submodule.zero_mem], exact_mod_cast nat.zero_le _ }, rw [norm_le_pow_iff_le_valuation x hx, mem_span_pow_iff_le_valuation x hx], end @@ -525,7 +525,7 @@ begin have := norm_le_pow_iff_mem_span_pow x 1, rw [ideal.mem_span_singleton, pow_one] at this, rw [← this, norm_le_pow_iff_norm_lt_pow_add_one], - simp only [gpow_zero, int.coe_nat_zero, int.coe_nat_succ, add_left_neg, zero_add], + simp only [zpow_zero, int.coe_nat_zero, int.coe_nat_succ, add_left_neg, zero_add], end @[simp] lemma pow_p_dvd_int_iff (n : ℕ) (a : ℤ) : (p ^ n : ℤ_[p]) ∣ a ↔ ↑p ^ n ∣ a := @@ -583,7 +583,7 @@ instance : is_adic_complete (maximal_ideal ℤ_[p]) ℤ_[p] := { intros ε hε, obtain ⟨m, hm⟩ := exists_pow_neg_lt p hε, refine ⟨m, λ n hn, lt_of_le_of_lt _ hm⟩, rw [← neg_sub, norm_neg], exact hx hn }, { refine ⟨x'.lim, λ n, _⟩, - have : (0:ℝ) < p ^ (-n : ℤ), { apply fpow_pos_of_pos, exact_mod_cast hp_prime.1.pos }, + have : (0:ℝ) < p ^ (-n : ℤ), { apply zpow_pos_of_pos, exact_mod_cast hp_prime.1.pos }, obtain ⟨i, hi⟩ := equiv_def₃ (equiv_lim x') this, by_cases hin : i ≤ n, { exact (hi i le_rfl n hin).le, }, diff --git a/src/number_theory/padics/padic_norm.lean b/src/number_theory/padics/padic_norm.lean index 2b213f6745bbc..dca3e8b178dc5 100644 --- a/src/number_theory/padics/padic_norm.lean +++ b/src/number_theory/padics/padic_norm.lean @@ -532,7 +532,7 @@ variables (p : ℕ) /-- Unfolds the definition of the p-adic norm of `q` when `q ≠ 0`. -/ -@[simp] protected lemma eq_fpow_of_nonzero {q : ℚ} (hq : q ≠ 0) : +@[simp] protected lemma eq_zpow_of_nonzero {q : ℚ} (hq : q ≠ 0) : padic_norm p q = p ^ (-(padic_val_rat p q)) := by simp [hq, padic_norm] @@ -544,7 +544,7 @@ if hq : q = 0 then by simp [hq, padic_norm] else begin unfold padic_norm; split_ifs, - apply fpow_nonneg, + apply zpow_nonneg, exact_mod_cast nat.zero_le _ end @@ -624,8 +624,8 @@ If `q ≠ 0`, then `padic_norm p q ≠ 0`. -/ protected lemma nonzero {q : ℚ} (hq : q ≠ 0) : padic_norm p q ≠ 0 := begin - rw padic_norm.eq_fpow_of_nonzero p hq, - apply fpow_ne_zero_of_ne_zero, + rw padic_norm.eq_zpow_of_nonzero p hq, + apply zpow_ne_zero_of_ne_zero, exact_mod_cast ne_of_gt hp.1.pos end @@ -637,7 +637,7 @@ begin apply by_contradiction, intro hq, unfold padic_norm at h, rw if_neg hq at h, apply absurd h, - apply fpow_ne_zero_of_ne_zero, + apply zpow_ne_zero_of_ne_zero, exact_mod_cast hp.1.ne_zero end @@ -652,7 +652,7 @@ else if hr : r = 0 then else have q*r ≠ 0, from mul_ne_zero hq hr, have (↑p : ℚ) ≠ 0, by simp [hp.1.ne_zero], - by simp [padic_norm, *, padic_val_rat.mul, fpow_add this, mul_comm] + by simp [padic_norm, *, padic_val_rat.mul, zpow_add₀ this, mul_comm] /-- The p-adic norm respects division. @@ -669,7 +669,7 @@ if hz : z = 0 then by simp [hz, zero_le_one] else begin unfold padic_norm, rw [if_neg _], - { refine fpow_le_one_of_nonpos _ _, + { refine zpow_le_one_of_nonpos _ _, { exact_mod_cast le_of_lt hp.1.one_lt, }, { rw [padic_val_rat_of_int _ hp.1.ne_one hz, neg_nonpos], norm_cast, simp }}, @@ -691,7 +691,7 @@ else unfold padic_norm, split_ifs, apply le_max_iff.2, left, - apply fpow_le_of_le, + apply zpow_le_of_le, { exact_mod_cast le_of_lt hp.1.one_lt }, { apply neg_le_neg, have : padic_val_rat p q = @@ -778,7 +778,7 @@ begin { norm_cast at hz, have : 0 ≤ (p^n : ℚ), {apply pow_nonneg, exact_mod_cast le_of_lt hp.1.pos }, simp [hz, this] }, - { rw [fpow_le_iff_le, neg_le_neg_iff, padic_val_rat_of_int _ hp.1.ne_one _], + { rw [zpow_le_iff_le, neg_le_neg_iff, padic_val_rat_of_int _ hp.1.ne_one _], { norm_cast, rw [← enat.coe_le_coe, enat.coe_get, ← multiplicity.pow_dvd_iff_le_multiplicity], simp }, diff --git a/src/number_theory/padics/padic_numbers.lean b/src/number_theory/padics/padic_numbers.lean index b77c4f65eae3f..43369b148fd5a 100644 --- a/src/number_theory/padics/padic_numbers.lean +++ b/src/number_theory/padics/padic_numbers.lean @@ -220,7 +220,7 @@ end lemma val_eq_iff_norm_eq {f g : padic_seq p} (hf : ¬ f ≈ 0) (hg : ¬ g ≈ 0) : f.valuation = g.valuation ↔ f.norm = g.norm := begin - rw [norm_eq_pow_val hf, norm_eq_pow_val hg, ← neg_inj, fpow_inj], + rw [norm_eq_pow_val hf, norm_eq_pow_val hg, ← neg_inj, zpow_inj], { exact_mod_cast (fact.out p.prime).pos }, { exact_mod_cast (fact.out p.prime).ne_one }, end @@ -816,7 +816,7 @@ end begin have p₀ : p ≠ 0 := hp.1.ne_zero, have p₁ : p ≠ 1 := hp.1.ne_one, - simp [p₀, p₁, norm, padic_norm, padic_val_rat, fpow_neg, padic.cast_eq_of_rat_of_nat], + simp [p₀, p₁, norm, padic_norm, padic_val_rat, zpow_neg, padic.cast_eq_of_rat_of_nat], end lemma norm_p_lt_one : ∥(p : ℚ_[p])∥ < 1 := @@ -827,7 +827,7 @@ begin end @[simp] lemma norm_p_pow (n : ℤ) : ∥(p^n : ℚ_[p])∥ = p^-n := -by rw [normed_field.norm_fpow, norm_p]; field_simp +by rw [normed_field.norm_zpow, norm_p]; field_simp instance : nondiscrete_normed_field ℚ_[p] := { non_trivial := ⟨p⁻¹, begin @@ -864,11 +864,11 @@ theorem norm_rat_le_one : ∀ {q : ℚ} (hq : ¬ p ∣ q.denom), ∥(q : ℚ_[p] from mt rat.zero_iff_num_zero.1 hnz, rw [padic_norm_e.eq_padic_norm], norm_cast, - rw [padic_norm.eq_fpow_of_nonzero p hnz', padic_val_rat_def p hnz'], + rw [padic_norm.eq_zpow_of_nonzero p hnz', padic_val_rat_def p hnz'], have h : (multiplicity p d).get _ = 0, by simp [multiplicity_eq_zero_of_not_dvd, hq], simp only, norm_cast, rw_mod_cast [h, sub_zero], - apply fpow_le_one_of_nonpos, + apply zpow_le_one_of_nonpos, { exact_mod_cast le_of_lt hp.1.one_lt, }, { apply neg_nonpos_of_nonneg, norm_cast, simp, } end @@ -895,7 +895,7 @@ begin rw H, apply dvd_zero }, { norm_cast at H ⊢, - convert gpow_zero _, + convert zpow_zero _, simp only [neg_eq_zero], rw padic_val_rat.padic_val_rat_of_int _ hp.1.ne_one H, norm_cast, @@ -1003,7 +1003,7 @@ begin change (padic_seq.norm _ : ℝ) = (p : ℝ) ^ -padic_seq.valuation _, rw padic_seq.norm_eq_pow_val, change ↑((p : ℚ) ^ -padic_seq.valuation f) = (p : ℝ) ^ -padic_seq.valuation f, - { rw rat.cast_fpow, + { rw rat.cast_zpow, congr' 1, norm_cast }, { apply cau_seq.not_lim_zero_of_not_congr_zero, @@ -1016,7 +1016,7 @@ end begin have h : (1 : ℝ) < p := by exact_mod_cast (fact.out p.prime).one_lt, rw ← neg_inj, - apply (fpow_strict_mono h).injective, + apply (zpow_strict_mono h).injective, dsimp only, rw ← norm_eq_pow_val, { simp }, @@ -1029,11 +1029,11 @@ lemma norm_le_pow_iff_norm_lt_pow_add_one (x : ℚ_[p]) (n : ℤ) : ∥x∥ ≤ p ^ n ↔ ∥x∥ < p ^ (n + 1) := begin have aux : ∀ n : ℤ, 0 < (p ^ n : ℝ), - { apply nat.fpow_pos_of_pos, exact hp_prime.1.pos }, + { apply nat.zpow_pos_of_pos, exact hp_prime.1.pos }, by_cases hx0 : x = 0, { simp [hx0, norm_zero, aux, le_of_lt (aux _)], }, rw norm_eq_pow_val hx0, have h1p : 1 < (p : ℝ), { exact_mod_cast hp_prime.1.one_lt }, - have H := fpow_strict_mono h1p, + have H := zpow_strict_mono h1p, rw [H.le_iff_le, H.lt_iff_lt, int.lt_add_one_iff], end diff --git a/src/number_theory/padics/ring_homs.lean b/src/number_theory/padics/ring_homs.lean index 564f06fb34e54..c89e23076b962 100644 --- a/src/number_theory/padics/ring_homs.lean +++ b/src/number_theory/padics/ring_homs.lean @@ -373,7 +373,7 @@ begin apply to_zmod_spec }, obtain ⟨c, rfl⟩ : is_unit c, -- TODO: write a can_lift instance for units { rw int.nat_abs_eq_zero at hc0, - rw [is_unit_iff, norm_eq_pow_val hc', hc0, neg_zero, gpow_zero], }, + rw [is_unit_iff, norm_eq_pow_val hc', hc0, neg_zero, zpow_zero], }, rw discrete_valuation_ring.unit_mul_pow_congr_unit _ _ _ _ _ hc, exact irreducible_p }, { rw zero_pow hc0, @@ -627,7 +627,7 @@ lemma lift_sub_val_mem_span (r : R) (n : ℕ) : (lift f_compat r - (f n r).val) ∈ (ideal.span {↑p ^ n} : ideal ℤ_[p]) := begin obtain ⟨k, hk⟩ := lim_nth_hom_spec f_compat r _ - (show (0 : ℝ) < p ^ (-n : ℤ), from nat.fpow_pos_of_pos hp_prime.1.pos _), + (show (0 : ℝ) < p ^ (-n : ℤ), from nat.zpow_pos_of_pos hp_prime.1.pos _), have := le_of_lt (hk (max n k) (le_max_right _ _)), rw norm_le_pow_iff_mem_span_pow at this, dsimp [lift], diff --git a/src/number_theory/quadratic_reciprocity.lean b/src/number_theory/quadratic_reciprocity.lean index b3debc4dec428..3b7f9214cf0f4 100644 --- a/src/number_theory/quadratic_reciprocity.lean +++ b/src/number_theory/quadratic_reciprocity.lean @@ -41,13 +41,13 @@ begin cases nat.prime.eq_two_or_odd (fact.out p.prime) with hp2 hp_odd, { substI p, refine iff_of_true ⟨1, _⟩ _; apply subsingleton.elim }, obtain ⟨g, hg⟩ := is_cyclic.exists_generator (units (zmod p)), - obtain ⟨n, hn⟩ : x ∈ submonoid.powers g, { rw mem_powers_iff_mem_gpowers, apply hg }, + obtain ⟨n, hn⟩ : x ∈ submonoid.powers g, { rw mem_powers_iff_mem_zpowers, apply hg }, split, { rintro ⟨y, rfl⟩, rw [← pow_mul, two_mul_odd_div_two hp_odd, units_pow_card_sub_one_eq_one], }, { subst x, assume h, have key : 2 * (p / 2) ∣ n * (p / 2), { rw [← pow_mul] at h, - rw [two_mul_odd_div_two hp_odd, ← card_units, ← order_of_eq_card_of_forall_mem_gpowers hg], + rw [two_mul_odd_div_two hp_odd, ← card_units, ← order_of_eq_card_of_forall_mem_zpowers hg], apply order_of_dvd_of_pow_eq_one h }, have : 0 < p / 2 := nat.div_pos (fact.out (1 < p)) dec_trivial, obtain ⟨m, rfl⟩ := dvd_of_mul_dvd_mul_right this key, diff --git a/src/ring_theory/roots_of_unity.lean b/src/ring_theory/roots_of_unity.lean index 666bec8ac669a..9837b26626875 100644 --- a/src/ring_theory/roots_of_unity.lean +++ b/src/ring_theory/roots_of_unity.lean @@ -32,9 +32,9 @@ monoids, expressing that an element is a primitive root of unity. ## Main results * `roots_of_unity.is_cyclic`: the roots of unity in an integral domain form a cyclic group. -* `is_primitive_root.zmod_equiv_gpowers`: `zmod k` is equivalent to +* `is_primitive_root.zmod_equiv_zpowers`: `zmod k` is equivalent to the subgroup generated by a primitive `k`-th root of unity. -* `is_primitive_root.gpowers_eq`: in an integral domain, the subgroup generated by +* `is_primitive_root.zpowers_eq`: in an integral domain, the subgroup generated by a primitive `k`-th root of unity is equal to the `k`-th roots of unity. * `is_primitive_root.card_primitive_roots`: if an integral domain has a primitive `k`-th root of unity, then it has `φ k` of them. @@ -204,9 +204,9 @@ lemma ring_hom.map_root_of_unity_eq_pow_self (σ : R →+* R) (ζ : roots_of_uni ∃ m : ℕ, σ ζ = ζ ^ m := begin obtain ⟨m, hm⟩ := (σ.restrict_roots_of_unity k).map_cyclic, - rw [←σ.restrict_roots_of_unity_coe_apply, hm, gpow_eq_mod_order_of, ←int.to_nat_of_nonneg + rw [←σ.restrict_roots_of_unity_coe_apply, hm, zpow_eq_mod_order_of, ←int.to_nat_of_nonneg (m.mod_nonneg (int.coe_nat_ne_zero.mpr (pos_iff_ne_zero.mp (order_of_pos ζ)))), - gpow_coe_nat, roots_of_unity.coe_pow], + zpow_coe_nat, roots_of_unity.coe_pow], exact ⟨(m % (order_of ζ)).to_nat, rfl⟩, end @@ -315,9 +315,9 @@ begin dvd_of_pow_eq_one := _ }, intros l hl, apply h.dvd_of_pow_eq_one, - rw [← pow_one ζ, ← gpow_coe_nat ζ, ← hi.gcd_eq_one, nat.gcd_eq_gcd_ab, gpow_add, - mul_pow, ← gpow_coe_nat, ← gpow_mul, mul_right_comm], - simp only [gpow_mul, hl, h.pow_eq_one, one_gpow, one_pow, one_mul, gpow_coe_nat] + rw [← pow_one ζ, ← zpow_coe_nat ζ, ← hi.gcd_eq_one, nat.gcd_eq_gcd_ab, zpow_add, + mul_pow, ← zpow_coe_nat, ← zpow_mul, mul_right_comm], + simp only [zpow_mul, hl, h.pow_eq_one, one_zpow, one_pow, one_mul, zpow_coe_nat] end lemma pow_of_prime (h : is_primitive_root ζ k) {p : ℕ} (hprime : nat.prime p) (hdiv : ¬ p ∣ k) : @@ -345,19 +345,19 @@ section comm_group variables {ζ : G} -lemma gpow_eq_one (h : is_primitive_root ζ k) : ζ ^ (k : ℤ) = 1 := -by { rw gpow_coe_nat, exact h.pow_eq_one } +lemma zpow_eq_one (h : is_primitive_root ζ k) : ζ ^ (k : ℤ) = 1 := +by { rw zpow_coe_nat, exact h.pow_eq_one } -lemma gpow_eq_one_iff_dvd (h : is_primitive_root ζ k) (l : ℤ) : +lemma zpow_eq_one_iff_dvd (h : is_primitive_root ζ k) (l : ℤ) : ζ ^ l = 1 ↔ (k : ℤ) ∣ l := begin by_cases h0 : 0 ≤ l, - { lift l to ℕ using h0, rw [gpow_coe_nat], norm_cast, exact h.pow_eq_one_iff_dvd l }, + { lift l to ℕ using h0, rw [zpow_coe_nat], norm_cast, exact h.pow_eq_one_iff_dvd l }, { have : 0 ≤ -l, { simp only [not_le, neg_nonneg] at h0 ⊢, exact le_of_lt h0 }, lift -l to ℕ using this with l' hl', rw [← dvd_neg, ← hl'], norm_cast, - rw [← h.pow_eq_one_iff_dvd, ← inv_inj, ← gpow_neg, ← hl', gpow_coe_nat, one_inv] } + rw [← h.pow_eq_one_iff_dvd, ← inv_inj, ← zpow_neg, ← hl', zpow_coe_nat, one_inv] } end lemma inv (h : is_primitive_root ζ k) : is_primitive_root ζ⁻¹ k := @@ -372,16 +372,16 @@ lemma inv (h : is_primitive_root ζ k) : is_primitive_root ζ⁻¹ k := @[simp] lemma inv_iff : is_primitive_root ζ⁻¹ k ↔ is_primitive_root ζ k := by { refine ⟨_, λ h, inv h⟩, intro h, rw [← inv_inv ζ], exact inv h } -lemma gpow_of_gcd_eq_one (h : is_primitive_root ζ k) (i : ℤ) (hi : i.gcd k = 1) : +lemma zpow_of_gcd_eq_one (h : is_primitive_root ζ k) (i : ℤ) (hi : i.gcd k = 1) : is_primitive_root (ζ ^ i) k := begin by_cases h0 : 0 ≤ i, { lift i to ℕ using h0, - rw gpow_coe_nat, + rw zpow_coe_nat, exact h.pow_of_coprime i hi }, have : 0 ≤ -i, { simp only [not_le, neg_nonneg] at h0 ⊢, exact le_of_lt h0 }, lift -i to ℕ using this with i' hi', - rw [← inv_iff, ← gpow_neg, ← hi', gpow_coe_nat], + rw [← inv_iff, ← zpow_neg, ← hi', zpow_coe_nat], apply h.pow_of_coprime, rw [int.gcd, ← int.nat_abs_neg, ← hi'] at hi, exact hi @@ -397,19 +397,19 @@ section comm_group_with_zero variables {ζ : G₀} -lemma fpow_eq_one (h : is_primitive_root ζ k) : ζ ^ (k : ℤ) = 1 := -by { rw gpow_coe_nat, exact h.pow_eq_one } +lemma zpow_eq_one₀ (h : is_primitive_root ζ k) : ζ ^ (k : ℤ) = 1 := +by { rw zpow_coe_nat, exact h.pow_eq_one } -lemma fpow_eq_one_iff_dvd (h : is_primitive_root ζ k) (l : ℤ) : +lemma zpow_eq_one_iff_dvd₀ (h : is_primitive_root ζ k) (l : ℤ) : ζ ^ l = 1 ↔ (k : ℤ) ∣ l := begin by_cases h0 : 0 ≤ l, - { lift l to ℕ using h0, rw [gpow_coe_nat], norm_cast, exact h.pow_eq_one_iff_dvd l }, + { lift l to ℕ using h0, rw [zpow_coe_nat], norm_cast, exact h.pow_eq_one_iff_dvd l }, { have : 0 ≤ -l, { simp only [not_le, neg_nonneg] at h0 ⊢, exact le_of_lt h0 }, lift -l to ℕ using this with l' hl', rw [← dvd_neg, ← hl'], norm_cast, - rw [← h.pow_eq_one_iff_dvd, ← inv_inj₀, ← fpow_neg, ← hl', gpow_coe_nat, inv_one] } + rw [← h.pow_eq_one_iff_dvd, ← inv_inj₀, ← zpow_neg₀, ← hl', zpow_coe_nat, inv_one] } end lemma inv' (h : is_primitive_root ζ k) : is_primitive_root ζ⁻¹ k := @@ -424,16 +424,16 @@ lemma inv' (h : is_primitive_root ζ k) : is_primitive_root ζ⁻¹ k := @[simp] lemma inv_iff' : is_primitive_root ζ⁻¹ k ↔ is_primitive_root ζ k := by { refine ⟨_, λ h, inv' h⟩, intro h, rw [← inv_inv₀ ζ], exact inv' h } -lemma fpow_of_gcd_eq_one (h : is_primitive_root ζ k) (i : ℤ) (hi : i.gcd k = 1) : +lemma zpow_of_gcd_eq_one₀ (h : is_primitive_root ζ k) (i : ℤ) (hi : i.gcd k = 1) : is_primitive_root (ζ ^ i) k := begin by_cases h0 : 0 ≤ i, { lift i to ℕ using h0, - rw gpow_coe_nat, + rw zpow_coe_nat, exact h.pow_of_coprime i hi }, have : 0 ≤ -i, { simp only [not_le, neg_nonneg] at h0 ⊢, exact le_of_lt h0 }, lift -i to ℕ using this with i' hi', - rw [← inv_iff', ← fpow_neg, ← hi', gpow_coe_nat], + rw [← inv_iff', ← zpow_neg₀, ← hi', zpow_coe_nat], apply h.pow_of_coprime, rw [int.gcd, ← int.nat_abs_neg, ← hi'] at hi, exact hi @@ -501,18 +501,18 @@ h.pow_eq_one /-- The (additive) monoid equivalence between `zmod k` and the powers of a primitive root of unity `ζ`. -/ -def zmod_equiv_gpowers (h : is_primitive_root ζ k) : zmod k ≃+ additive (subgroup.gpowers ζ) := +def zmod_equiv_zpowers (h : is_primitive_root ζ k) : zmod k ≃+ additive (subgroup.zpowers ζ) := add_equiv.of_bijective (add_monoid_hom.lift_of_right_inverse (int.cast_add_hom $ zmod k) _ zmod.int_cast_right_inverse - ⟨{ to_fun := λ i, additive.of_mul (⟨_, i, rfl⟩ : subgroup.gpowers ζ), - map_zero' := by { simp only [gpow_zero], refl }, - map_add' := by { intros i j, simp only [gpow_add], refl } }, + ⟨{ to_fun := λ i, additive.of_mul (⟨_, i, rfl⟩ : subgroup.zpowers ζ), + map_zero' := by { simp only [zpow_zero], refl }, + map_add' := by { intros i j, simp only [zpow_add], refl } }, (λ i hi, begin simp only [add_monoid_hom.mem_ker, char_p.int_cast_eq_zero_iff (zmod k) k, add_monoid_hom.coe_mk, int.coe_cast_add_hom] at hi ⊢, obtain ⟨i, rfl⟩ := hi, - simp only [gpow_mul, h.pow_eq_one, one_gpow, gpow_coe_nat], + simp only [zpow_mul, h.pow_eq_one, one_zpow, zpow_coe_nat], refl end)⟩) begin @@ -520,7 +520,7 @@ add_equiv.of_bijective { rw add_monoid_hom.injective_iff, intros i hi, rw subtype.ext_iff at hi, - have := (h.gpow_eq_one_iff_dvd _).mp hi, + have := (h.zpow_eq_one_iff_dvd _).mp hi, rw [← (char_p.int_cast_eq_zero_iff (zmod k) k _).mpr this, eq_comm], exact zmod.int_cast_right_inverse i }, { rintro ⟨ξ, i, rfl⟩, @@ -529,33 +529,33 @@ add_equiv.of_bijective refl } end -@[simp] lemma zmod_equiv_gpowers_apply_coe_int (i : ℤ) : - h.zmod_equiv_gpowers i = additive.of_mul (⟨ζ ^ i, i, rfl⟩ : subgroup.gpowers ζ) := +@[simp] lemma zmod_equiv_zpowers_apply_coe_int (i : ℤ) : + h.zmod_equiv_zpowers i = additive.of_mul (⟨ζ ^ i, i, rfl⟩ : subgroup.zpowers ζ) := add_monoid_hom.lift_of_right_inverse_comp_apply _ _ zmod.int_cast_right_inverse _ _ -@[simp] lemma zmod_equiv_gpowers_apply_coe_nat (i : ℕ) : - h.zmod_equiv_gpowers i = additive.of_mul (⟨ζ ^ i, i, rfl⟩ : subgroup.gpowers ζ) := +@[simp] lemma zmod_equiv_zpowers_apply_coe_nat (i : ℕ) : + h.zmod_equiv_zpowers i = additive.of_mul (⟨ζ ^ i, i, rfl⟩ : subgroup.zpowers ζ) := begin have : (i : zmod k) = (i : ℤ), by norm_cast, - simp only [this, zmod_equiv_gpowers_apply_coe_int, gpow_coe_nat], + simp only [this, zmod_equiv_zpowers_apply_coe_int, zpow_coe_nat], refl end -@[simp] lemma zmod_equiv_gpowers_symm_apply_gpow (i : ℤ) : - h.zmod_equiv_gpowers.symm (additive.of_mul (⟨ζ ^ i, i, rfl⟩ : subgroup.gpowers ζ)) = i := -by rw [← h.zmod_equiv_gpowers.symm_apply_apply i, zmod_equiv_gpowers_apply_coe_int] +@[simp] lemma zmod_equiv_zpowers_symm_apply_zpow (i : ℤ) : + h.zmod_equiv_zpowers.symm (additive.of_mul (⟨ζ ^ i, i, rfl⟩ : subgroup.zpowers ζ)) = i := +by rw [← h.zmod_equiv_zpowers.symm_apply_apply i, zmod_equiv_zpowers_apply_coe_int] -@[simp] lemma zmod_equiv_gpowers_symm_apply_gpow' (i : ℤ) : - h.zmod_equiv_gpowers.symm ⟨ζ ^ i, i, rfl⟩ = i := -h.zmod_equiv_gpowers_symm_apply_gpow i +@[simp] lemma zmod_equiv_zpowers_symm_apply_zpow' (i : ℤ) : + h.zmod_equiv_zpowers.symm ⟨ζ ^ i, i, rfl⟩ = i := +h.zmod_equiv_zpowers_symm_apply_zpow i -@[simp] lemma zmod_equiv_gpowers_symm_apply_pow (i : ℕ) : - h.zmod_equiv_gpowers.symm (additive.of_mul (⟨ζ ^ i, i, rfl⟩ : subgroup.gpowers ζ)) = i := -by rw [← h.zmod_equiv_gpowers.symm_apply_apply i, zmod_equiv_gpowers_apply_coe_nat] +@[simp] lemma zmod_equiv_zpowers_symm_apply_pow (i : ℕ) : + h.zmod_equiv_zpowers.symm (additive.of_mul (⟨ζ ^ i, i, rfl⟩ : subgroup.zpowers ζ)) = i := +by rw [← h.zmod_equiv_zpowers.symm_apply_apply i, zmod_equiv_zpowers_apply_coe_nat] -@[simp] lemma zmod_equiv_gpowers_symm_apply_pow' (i : ℕ) : - h.zmod_equiv_gpowers.symm ⟨ζ ^ i, i, rfl⟩ = i := -h.zmod_equiv_gpowers_symm_apply_pow i +@[simp] lemma zmod_equiv_zpowers_symm_apply_pow' (i : ℕ) : + h.zmod_equiv_zpowers.symm ⟨ζ ^ i, i, rfl⟩ = i := +h.zmod_equiv_zpowers_symm_apply_pow i /-- If there is a `n`-th primitive root of unity in `R` and `b` divides `n`, then there is a `b`-th primitive root of unity in `R`. -/ @@ -573,35 +573,35 @@ end variables [is_domain R] -lemma gpowers_eq {k : ℕ+} {ζ : units R} (h : is_primitive_root ζ k) : - subgroup.gpowers ζ = roots_of_unity k R := +lemma zpowers_eq {k : ℕ+} {ζ : units R} (h : is_primitive_root ζ k) : + subgroup.zpowers ζ = roots_of_unity k R := begin apply set_like.coe_injective, haveI : fact (0 < (k : ℕ)) := ⟨k.pos⟩, - haveI F : fintype (subgroup.gpowers ζ) := fintype.of_equiv _ (h.zmod_equiv_gpowers).to_equiv, - refine @set.eq_of_subset_of_card_le (units R) (subgroup.gpowers ζ) (roots_of_unity k R) + haveI F : fintype (subgroup.zpowers ζ) := fintype.of_equiv _ (h.zmod_equiv_zpowers).to_equiv, + refine @set.eq_of_subset_of_card_le (units R) (subgroup.zpowers ζ) (roots_of_unity k R) F (roots_of_unity.fintype R k) - (subgroup.gpowers_subset $ show ζ ∈ roots_of_unity k R, from h.pow_eq_one) _, + (subgroup.zpowers_subset $ show ζ ∈ roots_of_unity k R, from h.pow_eq_one) _, calc fintype.card (roots_of_unity k R) ≤ k : card_roots_of_unity R k ... = fintype.card (zmod k) : (zmod.card k).symm - ... = fintype.card (subgroup.gpowers ζ) : fintype.card_congr (h.zmod_equiv_gpowers).to_equiv + ... = fintype.card (subgroup.zpowers ζ) : fintype.card_congr (h.zmod_equiv_zpowers).to_equiv end lemma eq_pow_of_mem_roots_of_unity {k : ℕ+} {ζ ξ : units R} (h : is_primitive_root ζ k) (hξ : ξ ∈ roots_of_unity k R) : ∃ (i : ℕ) (hi : i < k), ζ ^ i = ξ := begin - obtain ⟨n, rfl⟩ : ∃ n : ℤ, ζ ^ n = ξ, by rwa [← h.gpowers_eq] at hξ, + obtain ⟨n, rfl⟩ : ∃ n : ℤ, ζ ^ n = ξ, by rwa [← h.zpowers_eq] at hξ, have hk0 : (0 : ℤ) < k := by exact_mod_cast k.pos, let i := n % k, have hi0 : 0 ≤ i := int.mod_nonneg _ (ne_of_gt hk0), lift i to ℕ using hi0 with i₀ hi₀, refine ⟨i₀, _, _⟩, { zify, rw [hi₀], exact int.mod_lt_of_pos _ hk0 }, - { have aux := h.gpow_eq_one, rw [← coe_coe] at aux, - rw [← gpow_coe_nat, hi₀, ← int.mod_add_div n k, gpow_add, gpow_mul, - aux, one_gpow, mul_one] } + { have aux := h.zpow_eq_one, rw [← coe_coe] at aux, + rw [← zpow_coe_nat, hi₀, ← int.mod_add_div n k, zpow_add, zpow_mul, + aux, one_zpow, mul_one] } end lemma eq_pow_of_pow_eq_one {k : ℕ} {ζ ξ : R} @@ -643,10 +643,10 @@ lemma card_roots_of_unity' {n : ℕ+} (h : is_primitive_root ζ n) : fintype.card (roots_of_unity n R) = n := begin haveI : fact (0 < ↑n) := ⟨n.pos⟩, - let e := h.zmod_equiv_gpowers, - haveI F : fintype (subgroup.gpowers ζ) := fintype.of_equiv _ e.to_equiv, + let e := h.zmod_equiv_zpowers, + haveI F : fintype (subgroup.zpowers ζ) := fintype.of_equiv _ e.to_equiv, calc fintype.card (roots_of_unity n R) - = fintype.card (subgroup.gpowers ζ) : fintype.card_congr $ by rw h.gpowers_eq + = fintype.card (subgroup.zpowers ζ) : fintype.card_congr $ by rw h.zpowers_eq ... = fintype.card (zmod n) : fintype.card_congr e.to_equiv.symm ... = n : zmod.card n end diff --git a/src/tactic/group.lean b/src/tactic/group.lean index bc74780c84d5a..c5cff97553b13 100644 --- a/src/tactic/group.lean +++ b/src/tactic/group.lean @@ -10,7 +10,7 @@ import tactic.doc_commands # `group` Normalizes expressions in the language of groups. The basic idea is to use the simplifier -to put everything into a product of group powers (`gpow` which takes a group element and an +to put everything into a product of group powers (`zpow` which takes a group element and an integer), then simplify the exponents using the `ring` tactic. The process needs to be repeated since `ring` can normalize an exponent to zero, leading to a factor that can be removed before collecting exponents again. The simplifier step also uses some extra lemmas to avoid @@ -24,18 +24,18 @@ group_theory -- The next four lemmas are not general purpose lemmas, they are intended for use only by -- the `group` tactic. -lemma tactic.group.gpow_trick {G : Type*} [group G] (a b : G) (n m : ℤ) : a*b^n*b^m = a*b^(n+m) := -by rw [mul_assoc, ← gpow_add] +lemma tactic.group.zpow_trick {G : Type*} [group G] (a b : G) (n m : ℤ) : a*b^n*b^m = a*b^(n+m) := +by rw [mul_assoc, ← zpow_add] -lemma tactic.group.gpow_trick_one {G : Type*} [group G] (a b : G) (m : ℤ) : a*b*b^m = a*b^(m+1) := -by rw [mul_assoc, mul_self_gpow] +lemma tactic.group.zpow_trick_one {G : Type*} [group G] (a b : G) (m : ℤ) : a*b*b^m = a*b^(m+1) := +by rw [mul_assoc, mul_self_zpow] -lemma tactic.group.gpow_trick_one' {G : Type*} [group G] (a b : G) (n : ℤ) : a*b^n*b = a*b^(n+1) := -by rw [mul_assoc, mul_gpow_self] +lemma tactic.group.zpow_trick_one' {G : Type*} [group G] (a b : G) (n : ℤ) : a*b^n*b = a*b^(n+1) := +by rw [mul_assoc, mul_zpow_self] -lemma tactic.group.gpow_trick_sub {G : Type*} [group G] (a b : G) (n m : ℤ) : +lemma tactic.group.zpow_trick_sub {G : Type*} [group G] (a b : G) (n m : ℤ) : a*b^n*b^(-m) = a*b^(n-m) := -by rw [mul_assoc, ← gpow_add] ; refl +by rw [mul_assoc, ← zpow_add] ; refl namespace tactic @@ -49,7 +49,7 @@ meta def aux_group₁ (locat : loc) : tactic unit := expr ``(mul_one), expr ``(one_mul), expr ``(one_pow), - expr ``(one_gpow), + expr ``(one_zpow), expr ``(sub_self), expr ``(add_neg_self), expr ``(neg_add_self), @@ -63,20 +63,20 @@ meta def aux_group₁ (locat : loc) : tactic unit := expr ``(int.coe_nat_bit1), expr ``(int.mul_neg_eq_neg_mul_symm), expr ``(int.neg_mul_eq_neg_mul_symm), - symm_expr ``(gpow_coe_nat), - symm_expr ``(gpow_neg_one), - symm_expr ``(gpow_mul), - symm_expr ``(gpow_add_one), - symm_expr ``(gpow_one_add), - symm_expr ``(gpow_add), - expr ``(mul_gpow_neg_one), - expr ``(gpow_zero), - expr ``(mul_gpow), + symm_expr ``(zpow_coe_nat), + symm_expr ``(zpow_neg_one), + symm_expr ``(zpow_mul), + symm_expr ``(zpow_add_one), + symm_expr ``(zpow_one_add), + symm_expr ``(zpow_add), + expr ``(mul_zpow_neg_one), + expr ``(zpow_zero), + expr ``(mul_zpow), symm_expr ``(mul_assoc), - expr ``(gpow_trick), - expr ``(gpow_trick_one), - expr ``(gpow_trick_one'), - expr ``(gpow_trick_sub), + expr ``(zpow_trick), + expr ``(zpow_trick_one), + expr ``(zpow_trick_one'), + expr ``(zpow_trick_sub), expr ``(tactic.ring.horner)] [] locat >> skip diff --git a/src/topology/algebra/group_with_zero.lean b/src/topology/algebra/group_with_zero.lean index 7350a6627d944..a4129196efd1b 100644 --- a/src/topology/algebra/group_with_zero.lean +++ b/src/topology/algebra/group_with_zero.lean @@ -202,44 +202,44 @@ protected def mul_right₀ (c : α) (hc : c ≠ 0) : α ≃ₜ α := end homeomorph -section fpow +section zpow variables [group_with_zero G₀] [topological_space G₀] [has_continuous_inv₀ G₀] [has_continuous_mul G₀] -lemma continuous_at_fpow (x : G₀) (m : ℤ) (h : x ≠ 0 ∨ 0 ≤ m) : continuous_at (λ x, x ^ m) x := +lemma continuous_at_zpow (x : G₀) (m : ℤ) (h : x ≠ 0 ∨ 0 ≤ m) : continuous_at (λ x, x ^ m) x := begin cases m, - { simpa only [gpow_of_nat] using continuous_at_pow x m }, - { simp only [gpow_neg_succ_of_nat], + { simpa only [zpow_of_nat] using continuous_at_pow x m }, + { simp only [zpow_neg_succ_of_nat], have hx : x ≠ 0, from h.resolve_right (int.neg_succ_of_nat_lt_zero m).not_le, exact (continuous_at_pow x (m + 1)).inv₀ (pow_ne_zero _ hx) } end -lemma continuous_on_fpow (m : ℤ) : continuous_on (λ x : G₀, x ^ m) {0}ᶜ := -λ x hx, (continuous_at_fpow _ _ (or.inl hx)).continuous_within_at +lemma continuous_on_zpow (m : ℤ) : continuous_on (λ x : G₀, x ^ m) {0}ᶜ := +λ x hx, (continuous_at_zpow _ _ (or.inl hx)).continuous_within_at -lemma filter.tendsto.fpow {f : α → G₀} {l : filter α} {a : G₀} (hf : tendsto f l (𝓝 a)) (m : ℤ) +lemma filter.tendsto.zpow {f : α → G₀} {l : filter α} {a : G₀} (hf : tendsto f l (𝓝 a)) (m : ℤ) (h : a ≠ 0 ∨ 0 ≤ m) : tendsto (λ x, (f x) ^ m) l (𝓝 (a ^ m)) := -(continuous_at_fpow _ m h).tendsto.comp hf +(continuous_at_zpow _ m h).tendsto.comp hf variables {X : Type*} [topological_space X] {a : X} {s : set X} {f : X → G₀} -lemma continuous_at.fpow (hf : continuous_at f a) (m : ℤ) (h : f a ≠ 0 ∨ 0 ≤ m) : +lemma continuous_at.zpow (hf : continuous_at f a) (m : ℤ) (h : f a ≠ 0 ∨ 0 ≤ m) : continuous_at (λ x, (f x) ^ m) a := -hf.fpow m h +hf.zpow m h -lemma continuous_within_at.fpow (hf : continuous_within_at f s a) (m : ℤ) (h : f a ≠ 0 ∨ 0 ≤ m) : +lemma continuous_within_at.zpow (hf : continuous_within_at f s a) (m : ℤ) (h : f a ≠ 0 ∨ 0 ≤ m) : continuous_within_at (λ x, f x ^ m) s a := -hf.fpow m h +hf.zpow m h -lemma continuous_on.fpow (hf : continuous_on f s) (m : ℤ) (h : ∀ a ∈ s, f a ≠ 0 ∨ 0 ≤ m) : +lemma continuous_on.zpow (hf : continuous_on f s) (m : ℤ) (h : ∀ a ∈ s, f a ≠ 0 ∨ 0 ≤ m) : continuous_on (λ x, f x ^ m) s := -λ a ha, (hf a ha).fpow m (h a ha) +λ a ha, (hf a ha).zpow m (h a ha) -@[continuity] lemma continuous.fpow (hf : continuous f) (m : ℤ) (h0 : ∀ a, f a ≠ 0 ∨ 0 ≤ m) : +@[continuity] lemma continuous.zpow (hf : continuous f) (m : ℤ) (h0 : ∀ a, f a ≠ 0 ∨ 0 ≤ m) : continuous (λ x, (f x) ^ m) := -continuous_iff_continuous_at.2 $ λ x, (hf.tendsto x).fpow m (h0 x) +continuous_iff_continuous_at.2 $ λ x, (hf.tendsto x).zpow m (h0 x) -end fpow +end zpow diff --git a/src/topology/algebra/ordered/basic.lean b/src/topology/algebra/ordered/basic.lean index f59e0470de770..0e4a5bdd28ed2 100644 --- a/src/topology/algebra/ordered/basic.lean +++ b/src/topology/algebra/ordered/basic.lean @@ -1751,10 +1751,10 @@ tendsto_inv_zero_at_top.comp h /-- The function `x^(-n)` tends to `0` at `+∞` for any positive natural `n`. A version for positive real powers exists as `tendsto_rpow_neg_at_top`. -/ lemma tendsto_pow_neg_at_top {n : ℕ} (hn : 1 ≤ n) : tendsto (λ x : α, x ^ (-(n:ℤ))) at_top (𝓝 0) := -tendsto.congr (λ x, (fpow_neg x n).symm) - (filter.tendsto.inv_tendsto_at_top (by simpa [gpow_coe_nat] using tendsto_pow_at_top hn)) +tendsto.congr (λ x, (zpow_neg₀ x n).symm) + (filter.tendsto.inv_tendsto_at_top (by simpa [zpow_coe_nat] using tendsto_pow_at_top hn)) -lemma tendsto_fpow_at_top_zero {n : ℤ} (hn : n < 0) : +lemma tendsto_zpow_at_top_zero {n : ℤ} (hn : n < 0) : tendsto (λ x : α, x^n) at_top (𝓝 0) := begin have : 1 ≤ -n := le_neg.mp (int.le_of_lt_add_one (hn.trans_le (neg_add_self 1).symm.le)), @@ -1763,9 +1763,9 @@ begin exact tendsto_pow_neg_at_top (by exact_mod_cast this) end -lemma tendsto_const_mul_fpow_at_top_zero {n : ℤ} {c : α} (hn : n < 0) : +lemma tendsto_const_mul_zpow_at_top_zero {n : ℤ} {c : α} (hn : n < 0) : tendsto (λ x, c * x ^ n) at_top (𝓝 0) := -(mul_zero c) ▸ (filter.tendsto.const_mul c (tendsto_fpow_at_top_zero hn)) +(mul_zero c) ▸ (filter.tendsto.const_mul c (tendsto_zpow_at_top_zero hn)) lemma tendsto_const_mul_pow_nhds_iff {n : ℕ} {c d : α} (hc : c ≠ 0) : tendsto (λ x : α, c * x ^ n) at_top (𝓝 d) ↔ n = 0 ∧ c = d := @@ -1786,22 +1786,22 @@ begin simpa [hn, hcd] using tendsto_const_nhds } end -lemma tendsto_const_mul_fpow_at_top_zero_iff {n : ℤ} {c d : α} (hc : c ≠ 0) : +lemma tendsto_const_mul_zpow_at_top_zero_iff {n : ℤ} {c d : α} (hc : c ≠ 0) : tendsto (λ x : α, c * x ^ n) at_top (𝓝 d) ↔ (n = 0 ∧ c = d) ∨ (n < 0 ∧ d = 0) := begin refine ⟨λ h, _, λ h, _⟩, { by_cases hn : 0 ≤ n, { lift n to ℕ using hn, - simp only [gpow_coe_nat] at h, + simp only [zpow_coe_nat] at h, rw [tendsto_const_mul_pow_nhds_iff hc, ← int.coe_nat_eq_zero] at h, exact or.inl h }, { rw not_le at hn, - refine or.inr ⟨hn, tendsto_nhds_unique h (tendsto_const_mul_fpow_at_top_zero hn)⟩ } }, + refine or.inr ⟨hn, tendsto_nhds_unique h (tendsto_const_mul_zpow_at_top_zero hn)⟩ } }, { cases h, - { simp only [h.left, h.right, gpow_zero, mul_one], + { simp only [h.left, h.right, zpow_zero, mul_one], exact tendsto_const_nhds }, - { exact h.2.symm ▸ tendsto_const_mul_fpow_at_top_zero h.1} } + { exact h.2.symm ▸ tendsto_const_mul_zpow_at_top_zero h.1} } end end linear_ordered_field diff --git a/test/refine_struct.lean b/test/refine_struct.lean index cbfd6bd536c46..579dc3b122055 100644 --- a/test/refine_struct.lean +++ b/test/refine_struct.lean @@ -57,10 +57,10 @@ begin guard_tags _field inv group, admit, guard_tags _field div group, admit, guard_tags _field div_eq_mul_inv group, admit, - guard_tags _field gpow group, admit, - guard_tags _field gpow_zero' group, admit, - guard_tags _field gpow_succ' group, admit, - guard_tags _field gpow_neg' group, admit, + guard_tags _field zpow group, admit, + guard_tags _field zpow_zero' group, admit, + guard_tags _field zpow_succ' group, admit, + guard_tags _field zpow_neg' group, admit, guard_tags _field mul_left_inv group, admit, }, trivial end @@ -138,22 +138,22 @@ begin -- α : Type -- ⊢ α → α - guard_tags _field gpow group, admit, - -- case group, gpow + guard_tags _field zpow group, admit, + -- case group, zpow -- α : Type -- ⊢ ℤ → α → α - guard_tags _field gpow_zero' group, admit, - -- case group, gpow_zero' + guard_tags _field zpow_zero' group, admit, + -- case group, zpow_zero' -- α : Type -- ⊢ ∀ (a : α), sorry 0 a = 1 - guard_tags _field gpow_succ' group, admit, + guard_tags _field zpow_succ' group, admit, -- case group, inv -- α : Type -- ⊢ ∀ (n : ℕ) (a : α), sorry (int.of_nat n.succ) a = a * sorry (int.of_nat n) a - guard_tags _field gpow_neg' group, admit, + guard_tags _field zpow_neg' group, admit, -- case group, inv -- α : Type -- ⊢ ∀ (n : ℕ) (a : α), sorry -[1+ n] a = sorry (sorry ↑(n.succ) a) @@ -185,10 +185,10 @@ begin guard_tags _field inv group, admit, guard_tags _field div group, admit, guard_tags _field div_eq_mul_inv group, admit, - guard_tags _field gpow group, admit, - guard_tags _field gpow_zero' group, admit, - guard_tags _field gpow_succ' group, admit, - guard_tags _field gpow_neg' group, admit, + guard_tags _field zpow group, admit, + guard_tags _field zpow_zero' group, admit, + guard_tags _field zpow_succ' group, admit, + guard_tags _field zpow_neg' group, admit, guard_tags _field mul_left_inv group, admit, guard_tags _field mul_assoc monoid, admit, guard_tags _field one monoid, admit,