Documentation

Mathlib.Algebra.Group.Action.Pointwise.Finset

Pointwise actions of finsets #

Instances #

instance Finset.smulCommClass_finset {α : Type u_2} {β : Type u_3} {γ : Type u_4} [DecidableEq γ] [SMul α γ] [SMul β γ] [SMulCommClass α β γ] :
instance Finset.vaddCommClass_finset {α : Type u_2} {β : Type u_3} {γ : Type u_4} [DecidableEq γ] [VAdd α γ] [VAdd β γ] [VAddCommClass α β γ] :
instance Finset.smulCommClass_finset' {α : Type u_2} {β : Type u_3} {γ : Type u_4} [DecidableEq γ] [SMul α γ] [SMul β γ] [SMulCommClass α β γ] :
instance Finset.vaddCommClass_finset' {α : Type u_2} {β : Type u_3} {γ : Type u_4} [DecidableEq γ] [VAdd α γ] [VAdd β γ] [VAddCommClass α β γ] :
instance Finset.smulCommClass_finset'' {α : Type u_2} {β : Type u_3} {γ : Type u_4} [DecidableEq γ] [SMul α γ] [SMul β γ] [SMulCommClass α β γ] :
instance Finset.vaddCommClass_finset'' {α : Type u_2} {β : Type u_3} {γ : Type u_4} [DecidableEq γ] [VAdd α γ] [VAdd β γ] [VAddCommClass α β γ] :
instance Finset.smulCommClass {α : Type u_2} {β : Type u_3} {γ : Type u_4} [DecidableEq γ] [SMul α γ] [SMul β γ] [SMulCommClass α β γ] :
instance Finset.vaddCommClass {α : Type u_2} {β : Type u_3} {γ : Type u_4} [DecidableEq γ] [VAdd α γ] [VAdd β γ] [VAddCommClass α β γ] :
instance Finset.isScalarTower {α : Type u_2} {β : Type u_3} {γ : Type u_4} [DecidableEq γ] [SMul α β] [SMul α γ] [SMul β γ] [IsScalarTower α β γ] :
instance Finset.vaddAssocClass {α : Type u_2} {β : Type u_3} {γ : Type u_4} [DecidableEq γ] [VAdd α β] [VAdd α γ] [VAdd β γ] [VAddAssocClass α β γ] :
instance Finset.isScalarTower' {α : Type u_2} {β : Type u_3} {γ : Type u_4} [DecidableEq γ] [DecidableEq β] [SMul α β] [SMul α γ] [SMul β γ] [IsScalarTower α β γ] :
instance Finset.vaddAssocClass' {α : Type u_2} {β : Type u_3} {γ : Type u_4} [DecidableEq γ] [DecidableEq β] [VAdd α β] [VAdd α γ] [VAdd β γ] [VAddAssocClass α β γ] :
instance Finset.isScalarTower'' {α : Type u_2} {β : Type u_3} {γ : Type u_4} [DecidableEq γ] [DecidableEq β] [SMul α β] [SMul α γ] [SMul β γ] [IsScalarTower α β γ] :
instance Finset.vaddAssocClass'' {α : Type u_2} {β : Type u_3} {γ : Type u_4} [DecidableEq γ] [DecidableEq β] [VAdd α β] [VAdd α γ] [VAdd β γ] [VAddAssocClass α β γ] :
instance Finset.isCentralScalar {α : Type u_2} {β : Type u_3} [DecidableEq β] [SMul α β] [SMul αᵐᵒᵖ β] [IsCentralScalar α β] :
instance Finset.isCentralVAdd {α : Type u_2} {β : Type u_3} [DecidableEq β] [VAdd α β] [VAdd αᵃᵒᵖ β] [IsCentralVAdd α β] :
def Finset.mulAction {α : Type u_2} {β : Type u_3} [DecidableEq β] [DecidableEq α] [Monoid α] [MulAction α β] :

A multiplicative action of a monoid α on a type β gives a multiplicative action of Finset α on Finset β.

Equations
Instances For
    def Finset.addAction {α : Type u_2} {β : Type u_3} [DecidableEq β] [DecidableEq α] [AddMonoid α] [AddAction α β] :

    An additive action of an additive monoid α on a type β gives an additive action of Finset α on Finset β

    Equations
    Instances For
      def Finset.mulActionFinset {α : Type u_2} {β : Type u_3} [DecidableEq β] [Monoid α] [MulAction α β] :

      A multiplicative action of a monoid on a type β gives a multiplicative action on Finset β.

      Equations
      Instances For
        def Finset.addActionFinset {α : Type u_2} {β : Type u_3} [DecidableEq β] [AddMonoid α] [AddAction α β] :

        An additive action of an additive monoid on a type β gives an additive action on Finset β.

        Equations
        Instances For
          theorem Finset.mul_singleton {α : Type u_2} [Mul α] [DecidableEq α] {s : Finset α} (a : α) :
          theorem Finset.add_singleton {α : Type u_2} [Add α] [DecidableEq α] {s : Finset α} (a : α) :
          theorem Finset.singleton_mul {α : Type u_2} [Mul α] [DecidableEq α] {s : Finset α} (a : α) :
          {a} * s = a s
          theorem Finset.singleton_add {α : Type u_2} [Add α] [DecidableEq α] {s : Finset α} (a : α) :
          {a} + s = a +ᵥ s
          theorem Finset.smul_finset_subset_mul {α : Type u_2} [Mul α] [DecidableEq α] {s t : Finset α} {a : α} :
          a sa t s * t
          theorem Finset.vadd_finset_subset_add {α : Type u_2} [Add α] [DecidableEq α] {s t : Finset α} {a : α} :
          a sa +ᵥ t s + t
          theorem Finset.op_smul_finset_subset_mul {α : Type u_2} [Mul α] [DecidableEq α] {s t : Finset α} {a : α} :
          a tMulOpposite.op a s s * t
          theorem Finset.op_vadd_finset_subset_add {α : Type u_2} [Add α] [DecidableEq α] {s t : Finset α} {a : α} :
          a tAddOpposite.op a +ᵥ s s + t
          @[simp]
          theorem Finset.biUnion_op_smul_finset {α : Type u_2} [Mul α] [DecidableEq α] (s t : Finset α) :
          (t.biUnion fun (a : α) => MulOpposite.op a s) = s * t
          @[simp]
          theorem Finset.biUnion_op_vadd_finset {α : Type u_2} [Add α] [DecidableEq α] (s t : Finset α) :
          (t.biUnion fun (a : α) => AddOpposite.op a +ᵥ s) = s + t
          theorem Finset.mul_subset_iff_left {α : Type u_2} [Mul α] [DecidableEq α] {s t u : Finset α} :
          s * t u as, a t u
          theorem Finset.add_subset_iff_left {α : Type u_2} [Add α] [DecidableEq α] {s t u : Finset α} :
          s + t u as, a +ᵥ t u
          theorem Finset.mul_subset_iff_right {α : Type u_2} [Mul α] [DecidableEq α] {s t u : Finset α} :
          s * t u bt, MulOpposite.op b s u
          theorem Finset.add_subset_iff_right {α : Type u_2} [Add α] [DecidableEq α] {s t u : Finset α} :
          s + t u bt, AddOpposite.op b +ᵥ s u
          theorem Finset.op_smul_finset_mul_eq_mul_smul_finset {α : Type u_2} [Semigroup α] [DecidableEq α] (a : α) (s t : Finset α) :
          MulOpposite.op a s * t = s * a t
          theorem Finset.op_vadd_finset_add_eq_add_vadd_finset {α : Type u_2} [AddSemigroup α] [DecidableEq α] (a : α) (s t : Finset α) :
          (AddOpposite.op a +ᵥ s) + t = s + (a +ᵥ t)
          theorem Finset.pairwiseDisjoint_smul_iff {α : Type u_2} [Mul α] [IsLeftCancelMul α] [DecidableEq α] {s : Set α} {t : Finset α} :
          (s.PairwiseDisjoint fun (x : α) => x t) Set.InjOn (fun (p : α × α) => p.1 * p.2) (s ×ˢ t)
          theorem Finset.pairwiseDisjoint_vadd_iff {α : Type u_2} [Add α] [IsLeftCancelAdd α] [DecidableEq α] {s : Set α} {t : Finset α} :
          (s.PairwiseDisjoint fun (x : α) => x +ᵥ t) Set.InjOn (fun (p : α × α) => p.1 + p.2) (s ×ˢ t)
          theorem Finset.image_smul_distrib {F : Type u_1} {α : Type u_2} {β : Type u_3} [DecidableEq α] [DecidableEq β] [Monoid α] [Monoid β] [FunLike F α β] [MonoidHomClass F α β] (f : F) (a : α) (s : Finset α) :
          image (⇑f) (a s) = f a image (⇑f) s
          theorem Finset.image_vadd_distrib {F : Type u_1} {α : Type u_2} {β : Type u_3} [DecidableEq α] [DecidableEq β] [AddMonoid α] [AddMonoid β] [FunLike F α β] [AddMonoidHomClass F α β] (f : F) (a : α) (s : Finset α) :
          image (⇑f) (a +ᵥ s) = f a +ᵥ image (⇑f) s
          @[simp]
          theorem Finset.smul_mem_smul_finset_iff {α : Type u_2} {β : Type u_3} [DecidableEq β] [Group α] [MulAction α β] {s : Finset β} {b : β} (a : α) :
          a b a s b s
          @[simp]
          theorem Finset.vadd_mem_vadd_finset_iff {α : Type u_2} {β : Type u_3} [DecidableEq β] [AddGroup α] [AddAction α β] {s : Finset β} {b : β} (a : α) :
          a +ᵥ b a +ᵥ s b s
          theorem Finset.inv_smul_mem_iff {α : Type u_2} {β : Type u_3} [DecidableEq β] [Group α] [MulAction α β] {s : Finset β} {a : α} {b : β} :
          a⁻¹ b s b a s
          theorem Finset.neg_vadd_mem_iff {α : Type u_2} {β : Type u_3} [DecidableEq β] [AddGroup α] [AddAction α β] {s : Finset β} {a : α} {b : β} :
          -a +ᵥ b s b a +ᵥ s
          theorem Finset.mem_inv_smul_finset_iff {α : Type u_2} {β : Type u_3} [DecidableEq β] [Group α] [MulAction α β] {s : Finset β} {a : α} {b : β} :
          b a⁻¹ s a b s
          theorem Finset.mem_neg_vadd_finset_iff {α : Type u_2} {β : Type u_3} [DecidableEq β] [AddGroup α] [AddAction α β] {s : Finset β} {a : α} {b : β} :
          b -a +ᵥ s a +ᵥ b s
          @[simp]
          theorem Finset.smul_finset_subset_smul_finset_iff {α : Type u_2} {β : Type u_3} [DecidableEq β] [Group α] [MulAction α β] {s t : Finset β} {a : α} :
          a s a t s t
          @[simp]
          theorem Finset.vadd_finset_subset_vadd_finset_iff {α : Type u_2} {β : Type u_3} [DecidableEq β] [AddGroup α] [AddAction α β] {s t : Finset β} {a : α} :
          a +ᵥ s a +ᵥ t s t
          theorem Finset.smul_finset_subset_iff {α : Type u_2} {β : Type u_3} [DecidableEq β] [Group α] [MulAction α β] {s t : Finset β} {a : α} :
          a s t s a⁻¹ t
          theorem Finset.vadd_finset_subset_iff {α : Type u_2} {β : Type u_3} [DecidableEq β] [AddGroup α] [AddAction α β] {s t : Finset β} {a : α} :
          a +ᵥ s t s -a +ᵥ t
          theorem Finset.subset_smul_finset_iff {α : Type u_2} {β : Type u_3} [DecidableEq β] [Group α] [MulAction α β] {s t : Finset β} {a : α} :
          s a t a⁻¹ s t
          theorem Finset.subset_vadd_finset_iff {α : Type u_2} {β : Type u_3} [DecidableEq β] [AddGroup α] [AddAction α β] {s t : Finset β} {a : α} :
          s a +ᵥ t -a +ᵥ s t
          theorem Finset.smul_finset_inter {α : Type u_2} {β : Type u_3} [DecidableEq β] [Group α] [MulAction α β] {s t : Finset β} {a : α} :
          a (s t) = a s a t
          theorem Finset.vadd_finset_inter {α : Type u_2} {β : Type u_3} [DecidableEq β] [AddGroup α] [AddAction α β] {s t : Finset β} {a : α} :
          a +ᵥ s t = (a +ᵥ s) (a +ᵥ t)
          theorem Finset.smul_finset_sdiff {α : Type u_2} {β : Type u_3} [DecidableEq β] [Group α] [MulAction α β] {s t : Finset β} {a : α} :
          a (s \ t) = a s \ a t
          theorem Finset.vadd_finset_sdiff {α : Type u_2} {β : Type u_3} [DecidableEq β] [AddGroup α] [AddAction α β] {s t : Finset β} {a : α} :
          a +ᵥ s \ t = (a +ᵥ s) \ (a +ᵥ t)
          theorem Finset.smul_finset_symmDiff {α : Type u_2} {β : Type u_3} [DecidableEq β] [Group α] [MulAction α β] {s t : Finset β} {a : α} :
          a symmDiff s t = symmDiff (a s) (a t)
          theorem Finset.vadd_finset_symmDiff {α : Type u_2} {β : Type u_3} [DecidableEq β] [AddGroup α] [AddAction α β] {s t : Finset β} {a : α} :
          a +ᵥ symmDiff s t = symmDiff (a +ᵥ s) (a +ᵥ t)
          @[simp]
          theorem Finset.smul_finset_univ {α : Type u_2} {β : Type u_3} [DecidableEq β] [Group α] [MulAction α β] {a : α} [Fintype β] :
          @[simp]
          theorem Finset.vadd_finset_univ {α : Type u_2} {β : Type u_3} [DecidableEq β] [AddGroup α] [AddAction α β] {a : α} [Fintype β] :
          @[simp]
          theorem Finset.smul_univ {α : Type u_2} {β : Type u_3} [DecidableEq β] [Group α] [MulAction α β] [Fintype β] {s : Finset α} (hs : s.Nonempty) :
          @[simp]
          theorem Finset.vadd_univ {α : Type u_2} {β : Type u_3} [DecidableEq β] [AddGroup α] [AddAction α β] [Fintype β] {s : Finset α} (hs : s.Nonempty) :
          @[simp]
          theorem Finset.card_smul_finset {α : Type u_2} {β : Type u_3} [DecidableEq β] [Group α] [MulAction α β] (a : α) (s : Finset β) :
          (a s).card = s.card
          @[simp]
          theorem Finset.card_vadd_finset {α : Type u_2} {β : Type u_3} [DecidableEq β] [AddGroup α] [AddAction α β] (a : α) (s : Finset β) :
          (a +ᵥ s).card = s.card
          @[simp]
          theorem Finset.dens_smul_finset {α : Type u_2} {β : Type u_3} [DecidableEq β] [Group α] [MulAction α β] [Fintype β] (a : α) (s : Finset β) :
          (a s).dens = s.dens
          @[simp]
          theorem Finset.dens_vadd_finset {α : Type u_2} {β : Type u_3} [DecidableEq β] [AddGroup α] [AddAction α β] [Fintype β] (a : α) (s : Finset β) :
          (a +ᵥ s).dens = s.dens
          theorem Finset.card_dvd_card_smul_right {α : Type u_2} {β : Type u_3} [DecidableEq β] [Group α] [MulAction α β] {t : Finset β} {s : Finset α} :
          ((fun (x : α) => x t) '' s).PairwiseDisjoint idt.card (s t).card

          If the left cosets of t by elements of s are disjoint (but not necessarily distinct!), then the size of t divides the size of s • t.

          theorem Finset.card_dvd_card_vadd_right {α : Type u_2} {β : Type u_3} [DecidableEq β] [AddGroup α] [AddAction α β] {t : Finset β} {s : Finset α} :
          ((fun (x : α) => x +ᵥ t) '' s).PairwiseDisjoint idt.card (s +ᵥ t).card

          If the left cosets of t by elements of s are disjoint (but not necessarily distinct!), then the size of t divides the size of s +ᵥ t.

          theorem Finset.card_dvd_card_mul_left {α : Type u_2} [Group α] [DecidableEq α] {s t : Finset α} :
          ((fun (b : α) => image (fun (a : α) => a * b) s) '' t).PairwiseDisjoint ids.card (s * t).card

          If the right cosets of s by elements of t are disjoint (but not necessarily distinct!), then the size of s divides the size of s * t.

          theorem Finset.card_dvd_card_add_left {α : Type u_2} [AddGroup α] [DecidableEq α] {s t : Finset α} :
          ((fun (b : α) => image (fun (a : α) => a + b) s) '' t).PairwiseDisjoint ids.card (s + t).card

          If the right cosets of s by elements of t are disjoint (but not necessarily distinct!), then the size of s divides the size of s + t.

          theorem Finset.card_dvd_card_mul_right {α : Type u_2} [Group α] [DecidableEq α] {s t : Finset α} :
          ((fun (x : α) => x t) '' s).PairwiseDisjoint idt.card (s * t).card

          If the left cosets of t by elements of s are disjoint (but not necessarily distinct!), then the size of t divides the size of s * t.

          theorem Finset.card_dvd_card_add_right {α : Type u_2} [AddGroup α] [DecidableEq α] {s t : Finset α} :
          ((fun (x : α) => x +ᵥ t) '' s).PairwiseDisjoint idt.card (s + t).card

          If the left cosets of t by elements of s are disjoint (but not necessarily distinct!), then the size of t divides the size of s + t.

          @[simp]
          theorem Finset.inv_smul_finset_distrib {α : Type u_2} [Group α] [DecidableEq α] (a : α) (s : Finset α) :
          @[simp]
          theorem Finset.neg_vadd_finset_distrib {α : Type u_2} [AddGroup α] [DecidableEq α] (a : α) (s : Finset α) :
          @[simp]
          theorem Finset.inv_op_smul_finset_distrib {α : Type u_2} [Group α] [DecidableEq α] (a : α) (s : Finset α) :
          @[simp]
          theorem Finset.neg_op_vadd_finset_distrib {α : Type u_2} [AddGroup α] [DecidableEq α] (a : α) (s : Finset α) :
          theorem Fintype.piFinset_smul {ι : Type u_5} {α : ιType u_6} {β : ιType u_7} [Fintype ι] [DecidableEq ι] [(i : ι) → DecidableEq (β i)] [(i : ι) → SMul (α i) (β i)] (s : (i : ι) → Finset (α i)) (t : (i : ι) → Finset (β i)) :
          (piFinset fun (i : ι) => s i t i) = piFinset s piFinset t
          theorem Fintype.piFinset_vadd {ι : Type u_5} {α : ιType u_6} {β : ιType u_7} [Fintype ι] [DecidableEq ι] [(i : ι) → DecidableEq (β i)] [(i : ι) → VAdd (α i) (β i)] (s : (i : ι) → Finset (α i)) (t : (i : ι) → Finset (β i)) :
          (piFinset fun (i : ι) => s i +ᵥ t i) = piFinset s +ᵥ piFinset t
          theorem Fintype.piFinset_smul_finset {ι : Type u_5} {α : ιType u_6} {β : ιType u_7} [Fintype ι] [DecidableEq ι] [(i : ι) → DecidableEq (β i)] [(i : ι) → SMul (α i) (β i)] (a : (i : ι) → α i) (s : (i : ι) → Finset (β i)) :
          (piFinset fun (i : ι) => a i s i) = a piFinset s
          theorem Fintype.piFinset_vadd_finset {ι : Type u_5} {α : ιType u_6} {β : ιType u_7} [Fintype ι] [DecidableEq ι] [(i : ι) → DecidableEq (β i)] [(i : ι) → VAdd (α i) (β i)] (a : (i : ι) → α i) (s : (i : ι) → Finset (β i)) :
          (piFinset fun (i : ι) => a i +ᵥ s i) = a +ᵥ piFinset s
          instance Nat.decidablePred_mem_vadd_set {s : Set } [DecidablePred fun (x : ) => x s] (a : ) :
          DecidablePred fun (x : ) => x a +ᵥ s
          Equations