@@ -104,7 +104,7 @@ def GradedMonoid (A : ι → Type _) :=
104
104
105
105
namespace GradedMonoid
106
106
107
- instance {A : ι → Type _} [Inhabited ι] [Inhabited (A default)] : Inhabited (GradedMonoid A) where
107
+ instance {A : ι → Type _} [Inhabited ι] [Inhabited (A default)] : Inhabited (GradedMonoid A) where
108
108
default := Sigma.instInhabitedSigma.default
109
109
110
110
/-- Construct an element of a graded monoid. -/
@@ -194,7 +194,7 @@ class GMonoid [AddMonoid ι] extends GMul A, GOne A where
194
194
/-- Successor powers behave as expected -/
195
195
gnpow_succ' :
196
196
∀ (n : ℕ) (a : GradedMonoid A),
197
- (GradedMonoid.mk _ <| gnpow n.succ a.snd) = a * ⟨_, gnpow n a.snd⟩ := by
197
+ (GradedMonoid.mk _ <| gnpow n.succ a.snd) = a * ⟨_, gnpow n a.snd⟩ := by
198
198
apply_gmonoid_gnpowRec_succ_tac
199
199
#align graded_monoid.gmonoid GradedMonoid.GMonoid
200
200
@@ -214,11 +214,11 @@ instance GMonoid.toMonoid [AddMonoid ι] [GMonoid A] : Monoid (GradedMonoid A)
214
214
theorem mk_pow [AddMonoid ι] [GMonoid A] {i} (a : A i) (n : ℕ) :
215
215
mk i a ^ n = mk (n • i) (GMonoid.gnpow _ a) :=
216
216
by
217
- match n with
218
- | 0 =>
217
+ match n with
218
+ | 0 =>
219
219
rw [pow_zero]
220
220
exact (GMonoid.gnpow_zero' ⟨_, a⟩).symm
221
- | n+1 =>
221
+ | n+1 =>
222
222
rw [pow_succ, mk_pow a n, mk_mul_mk]
223
223
exact (GMonoid.gnpow_succ' n ⟨_, a⟩).symm
224
224
#align graded_monoid.mk_pow GradedMonoid.mk_pow
@@ -268,7 +268,7 @@ variable [AddZeroClass ι] [GMul A]
268
268
an `Eq.rec` to turn `A (0 + i)` into `A i`.
269
269
-/
270
270
instance GradeZero.smul (i : ι) : SMul (A 0 ) (A i)
271
- where smul x y := @Eq.rec ι (0 +i) (fun a _ => A a) (GMul.mul x y) i (zero_add i)
271
+ where smul x y := @Eq.rec ι (0 +i) (fun a _ => A a) (GMul.mul x y) i (zero_add i)
272
272
#align graded_monoid.grade_zero.has_smul GradedMonoid.GradeZero.smul
273
273
274
274
/-- `(*) : A 0 → A 0 → A 0` is the value provided in `GradedMonoid.GMul.mul`, composed with
@@ -281,7 +281,7 @@ variable {A}
281
281
282
282
@[simp]
283
283
theorem mk_zero_smul {i} (a : A 0 ) (b : A i) : mk _ (a • b) = mk _ a * mk _ b :=
284
- Sigma.ext (zero_add _).symm <| eq_rec_heq _ _
284
+ Sigma.ext (zero_add _).symm <| eq_rec_heq _ _
285
285
#align graded_monoid.mk_zero_smul GradedMonoid.mk_zero_smul
286
286
287
287
@[simp]
@@ -295,7 +295,7 @@ section Monoid
295
295
296
296
variable [AddMonoid ι] [GMonoid A]
297
297
298
- instance : Pow (A 0 ) ℕ where
298
+ instance : Pow (A 0 ) ℕ where
299
299
pow x n := @Eq.rec ι (n • (0 :ι)) (fun a _ => A a) (GMonoid.gnpow n x) 0 (nsmul_zero n)
300
300
301
301
variable {A}
@@ -376,9 +376,9 @@ theorem List.dProdIndex_cons (a : α) (l : List α) (fι : α → ι) :
376
376
377
377
theorem List.dProdIndex_eq_map_sum (l : List α) (fι : α → ι) : l.dProdIndex fι = (l.map fι).sum :=
378
378
by
379
- match l with
380
- | [] => simp
381
- | head::tail => simp [List.dProdIndex_eq_map_sum tail fι]
379
+ match l with
380
+ | [] => simp
381
+ | head::tail => simp [List.dProdIndex_eq_map_sum tail fι]
382
382
#align list.dprod_index_eq_map_sum List.dProdIndex_eq_map_sum
383
383
384
384
/-- A dependent product for graded monoids represented by the indexed family of types `A i`.
@@ -407,9 +407,9 @@ theorem List.dProd_cons (fι : α → ι) (fA : ∀ a, A (fι a)) (a : α) (l :
407
407
theorem GradedMonoid.mk_list_dProd (l : List α) (fι : α → ι) (fA : ∀ a, A (fι a)) :
408
408
GradedMonoid.mk _ (l.dProd fι fA) = (l.map fun a => GradedMonoid.mk (fι a) (fA a)).prod :=
409
409
by
410
- match l with
411
- | [] => simp; rfl
412
- | head::tail =>
410
+ match l with
411
+ | [] => simp; rfl
412
+ | head::tail =>
413
413
simp[← GradedMonoid.mk_list_dProd tail _ _, GradedMonoid.mk_mul_mk, List.prod_cons]
414
414
#align graded_monoid.mk_list_dprod GradedMonoid.mk_list_dProd
415
415
@@ -450,7 +450,7 @@ structure. -/
450
450
instance Monoid.gMonoid [AddMonoid ι] [Monoid R] : GradedMonoid.GMonoid fun _ : ι => R :=
451
451
-- { Mul.gMul ι, One.gOne ι with
452
452
{ One.gOne ι with
453
- mul := fun x y => x * y
453
+ mul := fun x y => x * y
454
454
one_mul := fun _ => Sigma.ext (zero_add _) (heq_of_eq (one_mul _))
455
455
mul_one := fun _ => Sigma.ext (add_zero _) (heq_of_eq (mul_one _))
456
456
mul_assoc := fun _ _ _ => Sigma.ext (add_assoc _ _ _) (heq_of_eq (mul_assoc _ _ _))
@@ -473,11 +473,11 @@ instance CommMonoid.gCommMonoid [AddCommMonoid ι] [CommMonoid R] :
473
473
theorem List.dProd_monoid {α} [AddMonoid ι] [Monoid R] (l : List α) (fι : α → ι) (fA : α → R) :
474
474
@List.dProd _ _ (fun _:ι => R) _ _ l fι fA = (l.map fA).prod :=
475
475
by
476
- match l with
477
- | [] =>
476
+ match l with
477
+ | [] =>
478
478
rw [List.dProd_nil, List.map_nil, List.prod_nil]
479
479
rfl
480
- | head::tail =>
480
+ | head::tail =>
481
481
rw [List.dProd_cons, List.map_cons, List.prod_cons, List.dProd_monoid tail _ _]
482
482
rfl
483
483
#align list.dprod_monoid List.dProd_monoid
@@ -530,18 +530,18 @@ instance SetLike.gMul {S : Type _} [SetLike S R] [Mul R] [Add ι] (A : ι → S)
530
530
#align set_like.ghas_mul SetLike.gMul
531
531
532
532
/-
533
- Porting note: simpNF linter returns
533
+ Porting note: simpNF linter returns
534
534
535
535
"Left-hand side does not simplify, when using the simp lemma on itself."
536
536
537
537
However, simp does indeed solve the following. Possibly related std#71,std#78
538
538
539
539
example {S : Type _} [SetLike S R] [Mul R] [Add ι] (A : ι → S)
540
540
[SetLike.GradedMul A] {i j : ι} (x : A i) (y : A j) :
541
- ↑(@GradedMonoid.GMul.mul _ (fun i => A i) _ _ _ _ x y) = (x * y : R) := by simp
541
+ ↑(@GradedMonoid.GMul.mul _ (fun i => A i) _ _ _ _ x y) = (x * y : R) := by simp
542
542
543
543
-/
544
- @[simp,nolint simpNF]
544
+ @[simp,nolint simpNF]
545
545
theorem SetLike.coe_gMul {S : Type _} [SetLike S R] [Mul R] [Add ι] (A : ι → S)
546
546
[SetLike.GradedMul A] {i j : ι} (x : A i) (y : A j) :
547
547
↑(@GradedMonoid.GMul.mul _ (fun i => A i) _ _ _ _ x y) = (x * y : R) :=
@@ -562,23 +562,23 @@ variable {A : ι → S} [SetLike.GradedMonoid A]
562
562
563
563
theorem pow_mem_graded (n : ℕ) {r : R} {i : ι} (h : r ∈ A i) : r ^ n ∈ A (n • i) :=
564
564
by
565
- match n with
566
- | 0 =>
565
+ match n with
566
+ | 0 =>
567
567
rw [pow_zero, zero_nsmul]
568
568
exact one_mem_graded _
569
- | n+1 =>
569
+ | n+1 =>
570
570
rw [pow_succ', succ_nsmul']
571
571
exact mul_mem_graded (pow_mem_graded n h) h
572
572
#align set_like.pow_mem_graded SetLike.pow_mem_graded
573
573
574
574
theorem list_prod_map_mem_graded {ι'} (l : List ι') (i : ι' → ι) (r : ι' → R)
575
575
(h : ∀ j ∈ l, r j ∈ A (i j)) : (l.map r).prod ∈ A (l.map i).sum :=
576
576
by
577
- match l with
578
- | [] =>
577
+ match l with
578
+ | [] =>
579
579
rw [List.map_nil, List.map_nil, List.prod_nil, List.sum_nil]
580
580
exact one_mem_graded _
581
- | head::tail =>
581
+ | head::tail =>
582
582
rw [List.map_cons, List.map_cons, List.prod_cons, List.sum_cons]
583
583
exact
584
584
mul_mem_graded (h _ <| List.mem_cons_self _ _)
@@ -598,8 +598,7 @@ end SetLike
598
598
instance SetLike.gMonoid {S : Type _} [SetLike S R] [Monoid R] [AddMonoid ι] (A : ι → S)
599
599
[SetLike.GradedMonoid A] : GradedMonoid.GMonoid fun i => A i :=
600
600
{ SetLike.gOne A,
601
- SetLike.gMul
602
- A with
601
+ SetLike.gMul A with
603
602
one_mul := fun ⟨_, _, _⟩ => Sigma.subtype_ext (zero_add _) (one_mul _)
604
603
mul_one := fun ⟨_, _, _⟩ => Sigma.subtype_ext (add_zero _) (mul_one _)
605
604
mul_assoc := fun ⟨_, _, _⟩ ⟨_, _, _⟩ ⟨_, _, _⟩ =>
@@ -610,7 +609,7 @@ instance SetLike.gMonoid {S : Type _} [SetLike S R] [Monoid R] [AddMonoid ι] (A
610
609
#align set_like.gmonoid SetLike.gMonoid
611
610
612
611
/-
613
- Porting note: simpNF linter returns
612
+ Porting note: simpNF linter returns
614
613
615
614
"Left-hand side does not simplify, when using the simp lemma on itself."
616
615
@@ -621,7 +620,7 @@ example {S : Type _} [SetLike S R] [Monoid R] [AddMonoid ι] (A : ι → S)
621
620
↑(@GradedMonoid.GMonoid.gnpow _ (fun i => A i) _ _ n _ x) = (x:R)^n := by simp
622
621
623
622
-/
624
- @[simp,nolint simpNF]
623
+ @[simp,nolint simpNF]
625
624
theorem SetLike.coe_gnpow {S : Type _} [SetLike S R] [Monoid R] [AddMonoid ι] (A : ι → S)
626
625
[SetLike.GradedMonoid A] {i : ι} (x : A i) (n : ℕ) :
627
626
↑(@GradedMonoid.GMonoid.gnpow _ (fun i => A i) _ _ n _ x) = (x:R)^n :=
@@ -642,30 +641,30 @@ open SetLike SetLike.GradedMonoid
642
641
variable {α S : Type _} [SetLike S R] [Monoid R] [AddMonoid ι]
643
642
644
643
/-
645
- Porting note: simpNF linter returns
644
+ Porting note: simpNF linter returns
646
645
647
646
"Left-hand side does not simplify, when using the simp lemma on itself."
648
647
649
648
However, simp does indeed solve the following. Possibly related std#71,std#78
650
649
651
650
example (A : ι → S) [SetLike.GradedMonoid A] (fι : α → ι)
652
- (fA : ∀ a, A (fι a)) (l : List α) : ↑(@List.dProd _ _ (fun i => ↥(A i)) _ _ l fι fA)
651
+ (fA : ∀ a, A (fι a)) (l : List α) : ↑(@List.dProd _ _ (fun i => ↥(A i)) _ _ l fι fA)
653
652
= (List.prod (l.map fun a => fA a) : R) := by simp
654
653
-/
655
654
/-- Coercing a dependent product of subtypes is the same as taking the regular product of the
656
655
coercions. -/
657
- @[simp,nolint simpNF]
656
+ @[simp,nolint simpNF]
658
657
theorem SetLike.coe_list_dProd (A : ι → S) [SetLike.GradedMonoid A] (fι : α → ι)
659
- (fA : ∀ a, A (fι a)) (l : List α) : ↑(@List.dProd _ _ (fun i => ↥(A i)) _ _ l fι fA)
658
+ (fA : ∀ a, A (fι a)) (l : List α) : ↑(@List.dProd _ _ (fun i => ↥(A i)) _ _ l fι fA)
660
659
= (List.prod (l.map fun a => fA a) : R) := by
661
- match l with
662
- | [] =>
660
+ match l with
661
+ | [] =>
663
662
rw [List.dProd_nil, coe_gOne, List.map_nil, List.prod_nil]
664
- | head::tail =>
665
- rw [List.dProd_cons, coe_gMul, List.map_cons, List.prod_cons,
663
+ | head::tail =>
664
+ rw [List.dProd_cons, coe_gMul, List.map_cons, List.prod_cons,
666
665
SetLike.coe_list_dProd _ _ _ tail]
667
666
#align set_like.coe_list_dprod SetLike.coe_list_dProd
668
-
667
+
669
668
/-- A version of `List.coe_dProd_set_like` with `Subtype.mk`. -/
670
669
theorem SetLike.list_dProd_eq (A : ι → S) [SetLike.GradedMonoid A] (fι : α → ι) (fA : ∀ a, A (fι a))
671
670
(l : List α) :
@@ -709,8 +708,7 @@ def SetLike.homogeneousSubmonoid [AddMonoid ι] [Monoid R] (A : ι → S) [SetLi
709
708
Submonoid R where
710
709
carrier := { a | SetLike.Homogeneous A a }
711
710
one_mem' := SetLike.homogeneous_one A
712
- mul_mem' a b := SetLike.homogeneous_mul a b
711
+ mul_mem' a b := SetLike.homogeneous_mul a b
713
712
#align set_like.homogeneous_submonoid SetLike.homogeneousSubmonoid
714
713
715
714
end HomogeneousElements
716
-
0 commit comments