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

Commit c400677

Browse files
urkudjcommelin
andcommitted
feat(algebra/module/basic): module rat E is a subsingleton (#9570)
* allow different (semi)rings in `add_monoid_hom.map_int_cast_smul` and `add_monoid_hom.map_nat_cast_smul`; * add `add_monoid_hom.map_inv_int_cast_smul` and `add_monoid_hom.map_inv_nat_cast_smul`; * allow different division rings in `add_monoid_hom.map_rat_cast_smul`, drop `char_zero` assumption; * prove `subsingleton (module rat M)`; * add a few convenience lemmas. Co-authored-by: Johan Commelin <johan@commelin.net>
1 parent eb3595e commit c400677

File tree

4 files changed

+86
-39
lines changed

4 files changed

+86
-39
lines changed

src/algebra/module/basic.lean

Lines changed: 60 additions & 22 deletions
Original file line numberDiff line numberDiff line change
@@ -121,7 +121,7 @@ def module.to_add_monoid_End : R →+* add_monoid.End M :=
121121
map_add' := λ x y, add_monoid_hom.ext $ λ r, by simp [add_smul],
122122
..distrib_mul_action.to_add_monoid_End R M }
123123

124-
/-- A convenience alias for `module.to_add_monoid_End` as a `monoid_hom`, usually to allow the
124+
/-- A convenience alias for `module.to_add_monoid_End` as an `add_monoid_hom`, usually to allow the
125125
use of `add_monoid_hom.flip`. -/
126126
def smul_add_hom : R →+ M →+ M :=
127127
(module.to_add_monoid_End R M).to_add_monoid_hom
@@ -380,40 +380,78 @@ lemma map_int_module_smul [add_comm_group M] [add_comm_group M₂]
380380
(f : M →+ M₂) (x : ℤ) (a : M) : f (x • a) = x • f a :=
381381
by simp only [f.map_gsmul]
382382

383-
lemma map_int_cast_smul
384-
[ring R] [add_comm_group M] [add_comm_group M₂] [module R M] [module R M₂]
385-
(f : M →+ M₂) (x : ℤ) (a : M) : f ((x : R) • a) = (x : R) • f a :=
383+
lemma map_int_cast_smul [add_comm_group M] [add_comm_group M₂]
384+
(f : M →+ M₂) (R S : Type*) [ring R] [ring S] [module R M] [module S M₂]
385+
(x : ℤ) (a : M) : f ((x : R) • a) = (x : S) • f a :=
386386
by simp only [←gsmul_eq_smul_cast, f.map_gsmul]
387387

388-
lemma map_nat_cast_smul
389-
[semiring R] [add_comm_monoid M] [add_comm_monoid M₂]
390-
[module R M] [module R M₂] (f : M →+ M₂) (x : ℕ) (a : M) :
391-
f ((x : R) • a) = (x : R) • f a :=
388+
lemma map_nat_cast_smul [add_comm_monoid M] [add_comm_monoid M₂] (f : M →+ M₂)
389+
(R S : Type*) [semiring R] [semiring S] [module R M] [module S M₂] (x : ℕ) (a : M) :
390+
f ((x : R) • a) = (x : S) • f a :=
392391
by simp only [←nsmul_eq_smul_cast, f.map_nsmul]
393392

394-
lemma map_rat_cast_smul {R : Type*} [division_ring R] [char_zero R]
395-
{E : Type*} [add_comm_group E] [module R E] {F : Type*} [add_comm_group F] [module R F]
396-
(f : E →+ F) (c : ℚ) (x : E) :
397-
f ((c : R) • x) = (c : R) • f x :=
393+
lemma map_inv_int_cast_smul {E F : Type*} [add_comm_group E] [add_comm_group F] (f : E →+ F)
394+
(R S : Type*) [division_ring R] [division_ring S] [module R E] [module S F]
395+
(n : ) (x : E) :
396+
f ((n⁻¹ : R) • x) = (n⁻¹ : S) • f x :=
398397
begin
399-
have : ∀ (x : E) (n : ℕ), 0 < n → f (((n⁻¹ : ℚ) : R) • x) = ((n⁻¹ : ℚ) : R) • f x,
400-
{ intros x n hn,
401-
replace hn : (n : R) ≠ 0 := nat.cast_ne_zero.2 (ne_of_gt hn),
402-
conv_rhs { congr, skip, rw [← one_smul R x, ← mul_inv_cancel hn, mul_smul] },
403-
rw [f.map_nat_cast_smul, smul_smul, rat.cast_inv, rat.cast_coe_nat,
404-
inv_mul_cancel hn, one_smul] },
405-
refine c.num_denom_cases_on (λ m n hn hmn, _),
406-
rw [rat.mk_eq_div, div_eq_mul_inv, rat.cast_mul, int.cast_coe_nat, mul_smul, mul_smul,
407-
rat.cast_coe_int, f.map_int_cast_smul, this _ n hn]
398+
by_cases hR : (n : R) = 0; by_cases hS : (n : S) = 0,
399+
{ simp [hR, hS] },
400+
{ suffices : ∀ y, f y = 0, by simp [this], clear x, intro x,
401+
rw [← inv_smul_smul₀ hS (f x), ← map_int_cast_smul f R S], simp [hR] },
402+
{ suffices : ∀ y, f y = 0, by simp [this], clear x, intro x,
403+
rw [← smul_inv_smul₀ hR x, map_int_cast_smul f R S, hS, zero_smul] },
404+
{ rw [← inv_smul_smul₀ hS (f _), ← map_int_cast_smul f R S, smul_inv_smul₀ hR] }
408405
end
409406

407+
lemma map_inv_nat_cast_smul {E F : Type*} [add_comm_group E] [add_comm_group F] (f : E →+ F)
408+
(R S : Type*) [division_ring R] [division_ring S] [module R E] [module S F]
409+
(n : ℕ) (x : E) :
410+
f ((n⁻¹ : R) • x) = (n⁻¹ : S) • f x :=
411+
f.map_inv_int_cast_smul R S n x
412+
413+
lemma map_rat_cast_smul {E F : Type*} [add_comm_group E] [add_comm_group F] (f : E →+ F)
414+
(R S : Type*) [division_ring R] [division_ring S] [module R E] [module S F]
415+
(c : ℚ) (x : E) :
416+
f ((c : R) • x) = (c : S) • f x :=
417+
by rw [rat.cast_def, rat.cast_def, div_eq_mul_inv, div_eq_mul_inv, mul_smul, mul_smul,
418+
map_int_cast_smul f R S, map_inv_nat_cast_smul f R S]
419+
410420
lemma map_rat_module_smul {E : Type*} [add_comm_group E] [module ℚ E]
411421
{F : Type*} [add_comm_group F] [module ℚ F] (f : E →+ F) (c : ℚ) (x : E) :
412422
f (c • x) = c • f x :=
413-
rat.cast_id c ▸ f.map_rat_cast_smul c x
423+
rat.cast_id c ▸ f.map_rat_cast_smul ℚ ℚ c x
414424

415425
end add_monoid_hom
416426

427+
/-- There can be at most one `module ℚ E` structure on an additive commutative group. This is not
428+
an instance because `simp` becomes very slow if we have many `subsingleton` instances,
429+
see [gh-6025]. -/
430+
lemma subsingleton_rat_module (E : Type*) [add_comm_group E] : subsingleton (module ℚ E) :=
431+
⟨λ P Q, module_ext P Q $ λ r x,
432+
@add_monoid_hom.map_rat_module_smul E ‹_› P E ‹_› Q (add_monoid_hom.id _) r x⟩
433+
434+
/-- If `E` is a vector space over two division rings `R` and `S`, then scalar multiplications
435+
agree on inverses of integer numbers in `R` and `S`. -/
436+
lemma inv_int_cast_smul_eq {E : Type*} (R S : Type*) [add_comm_group E] [division_ring R]
437+
[division_ring S] [module R E] [module S E] (n : ℤ) (x : E) :
438+
(n⁻¹ : R) • x = (n⁻¹ : S) • x :=
439+
(add_monoid_hom.id E).map_inv_int_cast_smul R S n x
440+
441+
/-- If `E` is a vector space over two division rings `R` and `S`, then scalar multiplications
442+
agree on inverses of natural numbers in `R` and `S`. -/
443+
lemma inv_nat_cast_smul_eq {E : Type*} (R S : Type*) [add_comm_group E] [division_ring R]
444+
[division_ring S] [module R E] [module S E] (n : ℕ) (x : E) :
445+
(n⁻¹ : R) • x = (n⁻¹ : S) • x :=
446+
(add_monoid_hom.id E).map_inv_nat_cast_smul R S n x
447+
448+
/-- If `E` is a vector space over two division rings `R` and `S`, then scalar multiplications
449+
agree on rational numbers in `R` and `S`. -/
450+
lemma rat_cast_smul_eq {E : Type*} (R S : Type*) [add_comm_group E] [division_ring R]
451+
[division_ring S] [module R E] [module S E] (r : ℚ) (x : E) :
452+
(r : R) • x = (r : S) • x :=
453+
(add_monoid_hom.id E).map_rat_cast_smul R S r x
454+
417455
section no_zero_smul_divisors
418456
/-! ### `no_zero_smul_divisors`
419457

src/data/rat/cast.lean

Lines changed: 16 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -37,6 +37,8 @@ variable [division_ring α]
3737
-- see Note [coercion into rings]
3838
@[priority 900] instance cast_coe : has_coe_t ℚ α := ⟨λ r, r.1 / r.2
3939

40+
theorem cast_def (r : ℚ) : (r : α) = r.num / r.denom := rfl
41+
4042
@[simp] theorem cast_of_int (n : ℤ) : (of_int n : α) = n :=
4143
show (n / (1:ℕ) : α) = n, by rw [nat.cast_one, div_one]
4244

@@ -116,6 +118,20 @@ by simp [sub_eq_add_neg, (cast_add_of_ne_zero m0 this)]
116118
rw [(d₁.commute_cast (_:α)).inv_right₀.eq]
117119
end
118120

121+
@[simp] theorem cast_inv_nat (n : ℕ) : ((n⁻¹ : ℚ) : α) = n⁻¹ :=
122+
begin
123+
cases n, { simp },
124+
simp_rw [coe_nat_eq_mk, inv_def, mk, mk_nat, dif_neg n.succ_ne_zero, mk_pnat],
125+
simp [cast_def]
126+
end
127+
128+
@[simp] theorem cast_inv_int (n : ℤ) : ((n⁻¹ : ℚ) : α) = n⁻¹ :=
129+
begin
130+
cases n,
131+
{ exact cast_inv_nat _ },
132+
{ simp only [int.cast_neg_succ_of_nat, ← nat.cast_succ, cast_neg, inv_neg, cast_inv_nat] }
133+
end
134+
119135
@[norm_cast] theorem cast_inv_of_ne_zero : ∀ {n : ℚ},
120136
(n.num : α) ≠ 0 → (n.denom : α) ≠ 0 → ((n⁻¹ : ℚ) : α) = n⁻¹
121137
| ⟨n, d, h, c⟩ := λ (n0 : (n:α) ≠ 0) (d0 : (d:α) ≠ 0), begin

src/number_theory/padics/padic_numbers.lean

Lines changed: 9 additions & 16 deletions
Original file line numberDiff line numberDiff line change
@@ -812,19 +812,6 @@ begin
812812
rw [← padic_norm_e.eq_padic_norm', ← padic.cast_eq_of_rat]
813813
end
814814

815-
instance : nondiscrete_normed_field ℚ_[p] :=
816-
{ non_trivial := ⟨padic.of_rat p (p⁻¹), begin
817-
have h0 : p ≠ 0 := ne_of_gt (hp.1.pos),
818-
have h1 : 1 < p := hp.1.one_lt,
819-
rw [← padic.cast_eq_of_rat, eq_padic_norm],
820-
simp only [padic_norm, inv_eq_zero],
821-
simp only [if_neg] {discharger := `[exact_mod_cast h0]},
822-
norm_cast,
823-
simp only [padic_val_rat.inv] {discharger := `[exact_mod_cast h0]},
824-
rw [neg_neg, padic_val_rat.padic_val_rat_self h1, gpow_one],
825-
exact_mod_cast h1,
826-
end⟩ }
827-
828815
@[simp] lemma norm_p : ∥(p : ℚ_[p])∥ = p⁻¹ :=
829816
begin
830817
have p₀ : p ≠ 0 := hp.1.ne_zero,
@@ -834,14 +821,20 @@ end
834821

835822
lemma norm_p_lt_one : ∥(p : ℚ_[p])∥ < 1 :=
836823
begin
837-
rw [norm_p, inv_eq_one_div, div_lt_iff, one_mul],
838-
{ exact_mod_cast hp.1.one_lt },
839-
{ exact_mod_cast hp.1.pos }
824+
rw norm_p,
825+
apply inv_lt_one,
826+
exact_mod_cast hp.1.one_lt
840827
end
841828

842829
@[simp] lemma norm_p_pow (n : ℤ) : ∥(p^n : ℚ_[p])∥ = p^-n :=
843830
by rw [normed_field.norm_fpow, norm_p]; field_simp
844831

832+
instance : nondiscrete_normed_field ℚ_[p] :=
833+
{ non_trivial := ⟨p⁻¹, begin
834+
rw [normed_field.norm_inv, norm_p, inv_inv₀],
835+
exact_mod_cast hp.1.one_lt
836+
end⟩ }
837+
845838
protected theorem image {q : ℚ_[p]} : q ≠ 0 → ∃ n : ℤ, ∥q∥ = ↑((↑p : ℚ) ^ (-n)) :=
846839
quotient.induction_on q $ λ f hf,
847840
have ¬ f ≈ 0, from (padic_seq.ne_zero_iff_nequiv_zero f).1 hf,

src/topology/instances/real_vector_space.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -26,7 +26,7 @@ suffices (λ c : ℝ, f (c • x)) = λ c : ℝ, c • f x, from _root_.congr_fu
2626
dense_embedding_of_rat.dense.equalizer
2727
(hf.comp $ continuous_id.smul continuous_const)
2828
(continuous_id.smul continuous_const)
29-
(funext $ λ r, f.map_rat_cast_smul r x)
29+
(funext $ λ r, f.map_rat_cast_smul ℝ ℝ r x)
3030

3131
/-- Reinterpret a continuous additive homomorphism between two real vector spaces
3232
as a continuous real-linear map. -/

0 commit comments

Comments
 (0)