@@ -121,7 +121,7 @@ def module.to_add_monoid_End : R →+* add_monoid.End M :=
121
121
map_add' := λ x y, add_monoid_hom.ext $ λ r, by simp [add_smul],
122
122
..distrib_mul_action.to_add_monoid_End R M }
123
123
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
125
125
use of `add_monoid_hom.flip`. -/
126
126
def smul_add_hom : R →+ M →+ M :=
127
127
(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₂]
380
380
(f : M →+ M₂) (x : ℤ) (a : M) : f (x • a) = x • f a :=
381
381
by simp only [f.map_gsmul]
382
382
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 :=
386
386
by simp only [←gsmul_eq_smul_cast, f.map_gsmul]
387
387
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 :=
392
391
by simp only [←nsmul_eq_smul_cast, f.map_nsmul]
393
392
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 :=
398
397
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] }
408
405
end
409
406
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
+
410
420
lemma map_rat_module_smul {E : Type *} [add_comm_group E] [module ℚ E]
411
421
{F : Type *} [add_comm_group F] [module ℚ F] (f : E →+ F) (c : ℚ) (x : E) :
412
422
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
414
424
415
425
end add_monoid_hom
416
426
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
+
417
455
section no_zero_smul_divisors
418
456
/-! ### `no_zero_smul_divisors`
419
457
0 commit comments