diff --git a/src/algebra/algebra/tower.lean b/src/algebra/algebra/tower.lean index 6921d5323dd2e..f48478914c174 100644 --- a/src/algebra/algebra/tower.lean +++ b/src/algebra/algebra/tower.lean @@ -311,7 +311,7 @@ variables [add_comm_group M] [module A M] [module R M] [is_scalar_tower R A M] lemma lsmul_injective [no_zero_smul_divisors A M] {x : A} (hx : x ≠ 0) : function.injective (lsmul R M x) := -smul_left_injective _ hx +smul_right_injective _ hx end algebra diff --git a/src/algebra/module/basic.lean b/src/algebra/module/basic.lean index 99f763ddcf874..9d87a38096089 100644 --- a/src/algebra/module/basic.lean +++ b/src/algebra/module/basic.lean @@ -457,7 +457,7 @@ section smul_injective variables (M) -lemma smul_left_injective [no_zero_smul_divisors R M] {c : R} (hc : c ≠ 0) : +lemma smul_right_injective [no_zero_smul_divisors R M] {c : R} (hc : c ≠ 0) : function.injective (λ (x : M), c • x) := λ x y h, sub_eq_zero.mp ((smul_eq_zero.mp (calc c • (x - y) = c • x - c • y : smul_sub c x y @@ -490,7 +490,7 @@ section smul_injective variables (R) -lemma smul_right_injective {x : M} (hx : x ≠ 0) : +lemma smul_left_injective {x : M} (hx : x ≠ 0) : function.injective (λ (c : R), c • x) := λ c d h, sub_eq_zero.mp ((smul_eq_zero.mp (calc (c - d) • x = c • x - d • x : sub_smul c d x diff --git a/src/analysis/calculus/fderiv_symmetric.lean b/src/analysis/calculus/fderiv_symmetric.lean index 95104c89eecbf..b5cd69507ceb8 100644 --- a/src/analysis/calculus/fderiv_symmetric.lean +++ b/src/analysis/calculus/fderiv_symmetric.lean @@ -321,7 +321,7 @@ begin simp only [continuous_linear_map.map_add, continuous_linear_map.map_smul, add_right_inj, continuous_linear_map.add_apply, pi.smul_apply, continuous_linear_map.coe_smul', add_zero, continuous_linear_map.zero_apply, smul_zero, continuous_linear_map.map_zero] at this, - exact smul_left_injective F (tpos m).ne' this }, + exact smul_right_injective F (tpos m).ne' this }, -- applying `second_derivative_within_at_symmetric_of_mem_interior` to the vectors `z + (t v) v` -- and `z + (t w) w`, we deduce that `f'' v w = f'' w v`. Cross terms involving `z` can be -- eliminated thanks to the fact proved above that `f'' m z = f'' z m`. @@ -332,7 +332,7 @@ begin rw ← sub_eq_zero at this, abel at this, simp only [one_gsmul, neg_smul, sub_eq_zero, mul_comm, ← sub_eq_add_neg] at this, - apply smul_left_injective F _ this, + apply smul_right_injective F _ this, simp [(tpos v).ne', (tpos w).ne'] end diff --git a/src/analysis/convex/basic.lean b/src/analysis/convex/basic.lean index 8ab2187ab9780..e295667f1ebca 100644 --- a/src/analysis/convex/basic.lean +++ b/src/analysis/convex/basic.lean @@ -164,7 +164,7 @@ set.ext $ λ z, ⟨λ ⟨a, b, ha, hb, hab, hz⟩, begin split, { rintro ⟨a, b, ha, hb, hab, hx⟩, - refine smul_left_injective _ hb.ne' ((add_right_inj (a • x)).1 _), + refine smul_right_injective _ hb.ne' ((add_right_inj (a • x)).1 _), rw [hx, ←add_smul, hab, one_smul] }, rintro rfl, simp only [open_segment_same, mem_singleton], diff --git a/src/analysis/normed_space/conformal_linear_map.lean b/src/analysis/normed_space/conformal_linear_map.lean index 00d0ea3779d50..2ca1d2a0a6c12 100644 --- a/src/analysis/normed_space/conformal_linear_map.lean +++ b/src/analysis/normed_space/conformal_linear_map.lean @@ -113,7 +113,7 @@ end lemma injective {f' : M →L[R] N} (h : is_conformal_map f') : function.injective f' := let ⟨c, hc, li, hf'⟩ := h in by simp only [hf', pi.smul_def]; - exact (smul_left_injective _ hc).comp li.injective + exact (smul_right_injective _ hc).comp li.injective lemma ne_zero [nontrivial M] {f' : M →L[R] N} (hf' : is_conformal_map f') : f' ≠ 0 := diff --git a/src/group_theory/complement.lean b/src/group_theory/complement.lean index b8529e315df56..e20a2320602b2 100644 --- a/src/group_theory/complement.lean +++ b/src/group_theory/complement.lean @@ -271,11 +271,10 @@ quotient.induction_on α (quotient.induction_on β change pow_coprime hH ((pow_coprime hH).symm (diff α β)⁻¹) * (diff α β) = 1, rw [equiv.apply_symm_apply, inv_mul_self] }⟩)) -lemma smul_injective [H.normal] (α : H.quotient_diff) - (hH : nat.coprime (fintype.card H) (fintype.card (quotient_group.quotient H))) - {h₁ h₂ : H} (hh : h₁ • α = h₂ • α) : h₁ = h₂ := -begin - revert hh, +lemma smul_left_injective [H.normal] (α : H.quotient_diff) + (hH : nat.coprime (fintype.card H) (fintype.card (quotient_group.quotient H))) : + function.injective (λ h : H, h • α) := +λ h₁ h₂, begin refine quotient.induction_on α (λ α hα, _), replace hα : diff (h₁ • α) (h₂ • α) = 1 := quotient.exact hα, rw [smul_diff, ←diff_inv, smul_diff, diff_self, mul_one, mul_inv_eq_one] at hα, @@ -288,7 +287,7 @@ lemma is_complement_stabilizer_of_coprime [fintype G] [H.normal] {α : H.quotien begin classical, let ϕ : H ≃ mul_action.orbit G α := equiv.of_bijective (λ h, ⟨h • α, h, rfl⟩) - ⟨λ h₁ h₂ hh, smul_injective α hH (subtype.ext_iff.mp hh), + ⟨λ h₁ h₂ hh, smul_left_injective α hH (subtype.ext_iff.mp hh), λ β, exists_imp_exists (λ h hh, subtype.ext hh) (exists_smul_eq α β hH)⟩, have key := card_eq_card_quotient_mul_card_subgroup (mul_action.stabilizer G α), rw ← fintype.card_congr (ϕ.trans (mul_action.orbit_equiv_quotient_stabilizer G α)) at key, diff --git a/src/linear_algebra/basis.lean b/src/linear_algebra/basis.lean index 768ddad829021..9d2ea02ca638b 100644 --- a/src/linear_algebra/basis.lean +++ b/src/linear_algebra/basis.lean @@ -591,7 +591,7 @@ begin map_smul' := λ c y, _ }⟩, { rw [finsupp.add_apply, add_smul] }, { rw [finsupp.smul_apply, smul_assoc] }, - { refine smul_right_injective _ nz _, + { refine smul_left_injective _ nz _, simp only [finsupp.single_eq_same], exact (w (f (default ι) • x)).some_spec }, { simp only [finsupp.single_eq_same], diff --git a/src/linear_algebra/tensor_product.lean b/src/linear_algebra/tensor_product.lean index bd6ba00172b8b..1ee3a2c206cbb 100644 --- a/src/linear_algebra/tensor_product.lean +++ b/src/linear_algebra/tensor_product.lean @@ -202,7 +202,7 @@ variables {R M : Type*} [comm_ring R] [add_comm_group M] [module R M] lemma lsmul_injective [no_zero_smul_divisors R M] {x : R} (hx : x ≠ 0) : function.injective (lsmul R M x) := -smul_left_injective _ hx +smul_right_injective _ hx lemma ker_lsmul [no_zero_smul_divisors R M] {a : R} (ha : a ≠ 0) : (linear_map.lsmul R M a).ker = ⊥ :=