@@ -519,7 +519,8 @@ end add_comm_group
519
519
section
520
520
local attribute [instance] add_comm_monoid.nat_semimodule
521
521
522
- lemma module.smul_eq_smul {R : Type *} [ring R] {β : Type *} [add_comm_group β] [module R β]
522
+ lemma semimodule.smul_eq_smul (R : Type *) [semiring R]
523
+ {β : Type *} [add_comm_group β] [semimodule R β]
523
524
(n : ℕ) (b : β) : n • b = (n : R) • b :=
524
525
begin
525
526
induction n with n ih,
@@ -528,9 +529,9 @@ begin
528
529
rw [add_smul, add_smul, one_smul, ih, one_smul] }
529
530
end
530
531
531
- lemma module .add_monoid_smul_eq_smul { R : Type *} [ring R] {β : Type *} [add_comm_group β] [module R β]
532
- (n : ℕ) (b : β) : add_monoid.smul n b = (n : R) • b :=
533
- module .smul_eq_smul n b
532
+ lemma semimodule .add_monoid_smul_eq_smul ( R : Type *) [semiring R] {β : Type *} [add_comm_group β]
533
+ [semimodule R β] (n : ℕ) (b : β) : add_monoid.smul n b = (n : R) • b :=
534
+ semimodule .smul_eq_smul R n b
534
535
535
536
lemma nat.smul_def {M : Type *} [add_comm_monoid M] (n : ℕ) (x : M) :
536
537
n • x = add_monoid.smul n x :=
@@ -543,32 +544,20 @@ local attribute [instance] add_comm_group.int_module
543
544
544
545
lemma gsmul_eq_smul {M : Type *} [add_comm_group M] (n : ℤ) (x : M) : gsmul n x = n • x := rfl
545
546
546
- def is_add_group_hom.to_linear_map [add_comm_group α] [add_comm_group β]
547
- (f : α → β) [is_add_group_hom f] : α →ₗ[ℤ] β :=
548
- { to_fun := f,
549
- add := is_add_hom.map_add f,
550
- smul := λ i x, int.induction_on i (by rw [zero_smul, zero_smul, is_add_group_hom.map_zero f])
551
- (λ i ih, by rw [add_smul, add_smul, is_add_hom.map_add f, ih, one_smul, one_smul])
552
- (λ i ih, by rw [sub_smul, sub_smul, is_add_group_hom.map_sub f, ih, one_smul, one_smul]) }
553
-
554
- lemma module.gsmul_eq_smul_cast {R : Type *} [ring R] {β : Type *} [add_comm_group β] [module R β]
547
+ lemma module.gsmul_eq_smul_cast (R : Type *) [ring R] {β : Type *} [add_comm_group β] [module R β]
555
548
(n : ℤ) (b : β) : gsmul n b = (n : R) • b :=
556
549
begin
557
550
cases n,
551
+ { apply semimodule.add_monoid_smul_eq_smul, },
558
552
{ dsimp,
559
- apply module.add_monoid_smul_eq_smul, },
560
- { dsimp,
561
- rw module.add_monoid_smul_eq_smul (n.succ) b,
553
+ rw semimodule.add_monoid_smul_eq_smul R,
562
554
push_cast,
563
555
rw neg_smul, }
564
556
end
565
557
566
558
lemma module.gsmul_eq_smul {β : Type *} [add_comm_group β] [module ℤ β]
567
559
(n : ℤ) (b : β) : gsmul n b = n • b :=
568
- begin
569
- convert module.gsmul_eq_smul_cast n b,
570
- simp,
571
- end
560
+ by rw [module.gsmul_eq_smul_cast ℤ, int.cast_id]
572
561
573
562
end
574
563
@@ -577,30 +566,60 @@ end
577
566
lemma add_monoid_hom.map_int_module_smul
578
567
{α : Type *} {β : Type *} [add_comm_group α] [add_comm_group β]
579
568
[module ℤ α] [module ℤ β] (f : α →+ β) (x : ℤ) (a : α) : f (x • a) = x • f a :=
580
- begin
581
- rw ←module.gsmul_eq_smul,
582
- rw ←module.gsmul_eq_smul,
583
- rw add_monoid_hom.map_gsmul,
584
- end
569
+ by simp only [← module.gsmul_eq_smul, f.map_gsmul]
585
570
586
- lemma add_monoid_hom.map_smul_cast
571
+ lemma add_monoid_hom.map_int_cast_smul
587
572
{R : Type *} [ring R] {α : Type *} {β : Type *} [add_comm_group α] [add_comm_group β]
588
573
[module R α] [module R β] (f : α →+ β) (x : ℤ) (a : α) : f ((x : R) • a) = (x : R) • f a :=
574
+ by simp only [← module.gsmul_eq_smul_cast, f.map_gsmul]
575
+
576
+ lemma add_monoid_hom.map_nat_cast_smul
577
+ {R : Type *} [semiring R] {α : Type *} {β : Type *} [add_comm_group α] [add_comm_group β]
578
+ [semimodule R α] [semimodule R β] (f : α →+ β) (x : ℕ) (a : α) :
579
+ f ((x : R) • a) = (x : R) • f a :=
580
+ by simp only [← semimodule.add_monoid_smul_eq_smul, f.map_smul]
581
+
582
+ lemma add_monoid_hom.map_rat_cast_smul {R : Type *} [division_ring R] [char_zero R]
583
+ {E : Type *} [add_comm_group E] [module R E] {F : Type *} [add_comm_group F] [module R F]
584
+ (f : E →+ F) (c : ℚ) (x : E) :
585
+ f ((c : R) • x) = (c : R) • f x :=
589
586
begin
590
- rw ←module.gsmul_eq_smul_cast,
591
- rw ←module.gsmul_eq_smul_cast,
592
- rw add_monoid_hom.map_gsmul,
587
+ have : ∀ (x : E) (n : ℕ), 0 < n → f (((n⁻¹ : ℚ) : R) • x) = ((n⁻¹ : ℚ) : R) • f x,
588
+ { intros x n hn,
589
+ replace hn : (n : R) ≠ 0 := nat.cast_ne_zero.2 (ne_of_gt hn),
590
+ conv_rhs { congr, skip, rw [← one_smul R x, ← mul_inv_cancel hn, mul_smul] },
591
+ rw [f.map_nat_cast_smul, smul_smul, rat.cast_inv, rat.cast_coe_nat,
592
+ inv_mul_cancel hn, one_smul] },
593
+ refine c.num_denom_cases_on (λ m n hn hmn, _),
594
+ rw [rat.mk_eq_div, div_eq_mul_inv, rat.cast_mul, int.cast_coe_nat, mul_smul, mul_smul,
595
+ rat.cast_coe_int, f.map_int_cast_smul, this _ n hn]
593
596
end
594
597
598
+ lemma add_monoid_hom.map_rat_module_smul {E : Type *} [add_comm_group E] [vector_space ℚ E]
599
+ {F : Type *} [add_comm_group F] [module ℚ F] (f : E →+ F) (c : ℚ) (x : E) :
600
+ f (c • x) = c • f x :=
601
+ rat.cast_id c ▸ f.map_rat_cast_smul c x
602
+
595
603
-- We finally turn on these instances globally:
596
604
attribute [instance] add_comm_monoid.nat_semimodule add_comm_group.int_module
597
605
606
+ /-- Reinterpret an additive homomorphism as a `ℤ`-linear map. -/
607
+ def add_monoid_hom.to_int_linear_map [add_comm_group α] [add_comm_group β] (f : α →+ β) :
608
+ α →ₗ[ℤ] β :=
609
+ ⟨f, f.map_add, f.map_int_module_smul⟩
610
+
611
+ /-- Reinterpret an additive homomorphism as a `ℚ`-linear map. -/
612
+ def add_monoid_hom.to_rat_linear_map [add_comm_group α] [vector_space ℚ α]
613
+ [add_comm_group β] [vector_space ℚ β] (f : α →+ β) :
614
+ α →ₗ[ℚ] β :=
615
+ ⟨f, f.map_add, f.map_rat_module_smul⟩
616
+
598
617
namespace finset
599
618
600
619
lemma sum_const' {α : Type *} (R : Type *) [ring R] {β : Type *}
601
620
[add_comm_group β] [module R β] {s : finset α} (b : β) :
602
621
finset.sum s (λ (a : α), b) = (finset.card s : R) • b :=
603
- by rw [finset.sum_const, ← module .smul_eq_smul]; refl
622
+ by rw [finset.sum_const, ← semimodule .smul_eq_smul]; refl
604
623
605
624
variables {M : Type *} [decidable_linear_ordered_cancel_comm_monoid M]
606
625
{s : finset α} (f : α → M)
0 commit comments