Documentation

Mathlib.Data.MvPolynomial.Monad

Monad operations on MvPolynomial #

This file defines two monadic operations on MvPolynomial. Given p : MvPolynomial σ R,

These operations themselves have algebraic structure: MvPolynomial.bind₁ and MvPolynomial.join₁ are algebra homs and MvPolynomial.bind₂ and MvPolynomial.join₂ are ring homs.

They interact in convenient ways with MvPolynomial.rename, MvPolynomial.map, MvPolynomial.vars, and other polynomial operations. Indeed, MvPolynomial.rename is the "map" operation for the (bind₁, join₁) pair, whereas MvPolynomial.map is the "map" operation for the other pair.

Implementation notes #

We add a LawfulMonad instance for the (bind₁, join₁) pair. The second pair cannot be instantiated as a Monad, since it is not a monad in Type but in CommRing (or rather CommSemiRing).

def MvPolynomial.bind₁ {σ : Type u_1} {τ : Type u_2} {R : Type u_3} [CommSemiring R] (f : σMvPolynomial τ R) :

bind₁ is the "left hand side" bind operation on MvPolynomial, operating on the variable type. Given a polynomial p : MvPolynomial σ R and a map f : σ → MvPolynomial τ R taking variables in p to polynomials in the variable type τ, bind₁ f p replaces each variable in p with its value under f, producing a new polynomial in τ. The coefficient type remains the same. This operation is an algebra hom.

Equations
Instances For
    def MvPolynomial.bind₂ {σ : Type u_1} {R : Type u_3} {S : Type u_4} [CommSemiring R] [CommSemiring S] (f : R →+* MvPolynomial σ S) :

    bind₂ is the "right hand side" bind operation on MvPolynomial, operating on the coefficient type. Given a polynomial p : MvPolynomial σ R and a map f : R → MvPolynomial σ S taking coefficients in p to polynomials over a new ring S, bind₂ f p replaces each coefficient in p with its value under f, producing a new polynomial over S. The variable type remains the same. This operation is a ring hom.

    Equations
    Instances For

      join₁ is the monadic join operation corresponding to MvPolynomial.bind₁. Given a polynomial p with coefficients in R whose variables are polynomials in σ with coefficients in R, join₁ p collapses p to a polynomial with variables in σ and coefficients in R. This operation is an algebra hom.

      Equations
      Instances For

        join₂ is the monadic join operation corresponding to MvPolynomial.bind₂. Given a polynomial p with variables in σ whose coefficients are polynomials in σ with coefficients in R, join₂ p collapses p to a polynomial with variables in σ and coefficients in R. This operation is a ring hom.

        Equations
        Instances For
          @[simp]
          theorem MvPolynomial.aeval_eq_bind₁ {σ : Type u_1} {τ : Type u_2} {R : Type u_3} [CommSemiring R] (f : σMvPolynomial τ R) :
          @[simp]
          theorem MvPolynomial.eval₂Hom_C_eq_bind₁ {σ : Type u_1} {τ : Type u_2} {R : Type u_3} [CommSemiring R] (f : σMvPolynomial τ R) :
          @[simp]
          theorem MvPolynomial.eval₂Hom_eq_bind₂ {σ : Type u_1} {R : Type u_3} {S : Type u_4} [CommSemiring R] [CommSemiring S] (f : R →+* MvPolynomial σ S) :
          @[simp]
          theorem MvPolynomial.aeval_id_eq_join₁ (σ : Type u_1) (R : Type u_3) [CommSemiring R] :
          MvPolynomial.aeval id = MvPolynomial.join₁
          theorem MvPolynomial.eval₂Hom_C_id_eq_join₁ (σ : Type u_1) (R : Type u_3) [CommSemiring R] (φ : MvPolynomial (MvPolynomial σ R) R) :
          (MvPolynomial.eval₂Hom MvPolynomial.C id) φ = MvPolynomial.join₁ φ
          @[simp]
          theorem MvPolynomial.eval₂Hom_id_X_eq_join₂ (σ : Type u_1) (R : Type u_3) [CommSemiring R] :
          MvPolynomial.eval₂Hom (RingHom.id (MvPolynomial σ R)) MvPolynomial.X = MvPolynomial.join₂
          @[simp]
          theorem MvPolynomial.bind₁_X_right {σ : Type u_1} {τ : Type u_2} {R : Type u_3} [CommSemiring R] (f : σMvPolynomial τ R) (i : σ) :
          @[simp]
          theorem MvPolynomial.bind₂_X_right {σ : Type u_1} {R : Type u_3} {S : Type u_4} [CommSemiring R] [CommSemiring S] (f : R →+* MvPolynomial σ S) (i : σ) :
          @[simp]
          theorem MvPolynomial.bind₁_X_left {σ : Type u_1} {R : Type u_3} [CommSemiring R] :
          theorem MvPolynomial.bind₁_C_right {σ : Type u_1} {τ : Type u_2} {R : Type u_3} [CommSemiring R] (f : σMvPolynomial τ R) (x : R) :
          (MvPolynomial.bind₁ f) (MvPolynomial.C x) = MvPolynomial.C x
          @[simp]
          theorem MvPolynomial.bind₂_C_right {σ : Type u_1} {R : Type u_3} {S : Type u_4} [CommSemiring R] [CommSemiring S] (f : R →+* MvPolynomial σ S) (r : R) :
          (MvPolynomial.bind₂ f) (MvPolynomial.C r) = f r
          @[simp]
          theorem MvPolynomial.bind₂_C_left {σ : Type u_1} {R : Type u_3} [CommSemiring R] :
          @[simp]
          theorem MvPolynomial.bind₂_comp_C {σ : Type u_1} {R : Type u_3} {S : Type u_4} [CommSemiring R] [CommSemiring S] (f : R →+* MvPolynomial σ S) :
          RingHom.comp (MvPolynomial.bind₂ f) MvPolynomial.C = f
          @[simp]
          theorem MvPolynomial.join₂_map {σ : Type u_1} {R : Type u_3} {S : Type u_4} [CommSemiring R] [CommSemiring S] (f : R →+* MvPolynomial σ S) (φ : MvPolynomial σ R) :
          MvPolynomial.join₂ ((MvPolynomial.map f) φ) = (MvPolynomial.bind₂ f) φ
          @[simp]
          theorem MvPolynomial.join₂_comp_map {σ : Type u_1} {R : Type u_3} {S : Type u_4} [CommSemiring R] [CommSemiring S] (f : R →+* MvPolynomial σ S) :
          theorem MvPolynomial.aeval_id_rename {σ : Type u_1} {τ : Type u_2} {R : Type u_3} [CommSemiring R] (f : σMvPolynomial τ R) (p : MvPolynomial σ R) :
          @[simp]
          theorem MvPolynomial.join₁_rename {σ : Type u_1} {τ : Type u_2} {R : Type u_3} [CommSemiring R] (f : σMvPolynomial τ R) (φ : MvPolynomial σ R) :
          MvPolynomial.join₁ ((MvPolynomial.rename f) φ) = (MvPolynomial.bind₁ f) φ
          @[simp]
          theorem MvPolynomial.bind₁_id {σ : Type u_1} {R : Type u_3} [CommSemiring R] :
          MvPolynomial.bind₁ id = MvPolynomial.join₁
          @[simp]
          theorem MvPolynomial.bind₂_id {σ : Type u_1} {R : Type u_3} [CommSemiring R] :
          MvPolynomial.bind₂ (RingHom.id (MvPolynomial σ R)) = MvPolynomial.join₂
          theorem MvPolynomial.bind₁_bind₁ {σ : Type u_1} {τ : Type u_2} {R : Type u_3} [CommSemiring R] {υ : Type u_6} (f : σMvPolynomial τ R) (g : τMvPolynomial υ R) (φ : MvPolynomial σ R) :
          theorem MvPolynomial.bind₁_comp_bind₁ {σ : Type u_1} {τ : Type u_2} {R : Type u_3} [CommSemiring R] {υ : Type u_6} (f : σMvPolynomial τ R) (g : τMvPolynomial υ R) :
          theorem MvPolynomial.rename_comp_bind₁ {σ : Type u_1} {τ : Type u_2} {R : Type u_3} [CommSemiring R] {υ : Type u_6} (f : σMvPolynomial τ R) (g : τυ) :
          theorem MvPolynomial.rename_bind₁ {σ : Type u_1} {τ : Type u_2} {R : Type u_3} [CommSemiring R] {υ : Type u_6} (f : σMvPolynomial τ R) (g : τυ) (φ : MvPolynomial σ R) :
          theorem MvPolynomial.map_bind₂ {σ : Type u_1} {R : Type u_3} {S : Type u_4} {T : Type u_5} [CommSemiring R] [CommSemiring S] [CommSemiring T] (f : R →+* MvPolynomial σ S) (g : S →+* T) (φ : MvPolynomial σ R) :
          theorem MvPolynomial.bind₁_comp_rename {σ : Type u_1} {τ : Type u_2} {R : Type u_3} [CommSemiring R] {υ : Type u_6} (f : τMvPolynomial υ R) (g : στ) :
          theorem MvPolynomial.bind₁_rename {σ : Type u_1} {τ : Type u_2} {R : Type u_3} [CommSemiring R] {υ : Type u_6} (f : τMvPolynomial υ R) (g : στ) (φ : MvPolynomial σ R) :
          theorem MvPolynomial.bind₂_map {σ : Type u_1} {R : Type u_3} {S : Type u_4} {T : Type u_5} [CommSemiring R] [CommSemiring S] [CommSemiring T] (f : S →+* MvPolynomial σ T) (g : R →+* S) (φ : MvPolynomial σ R) :
          @[simp]
          theorem MvPolynomial.map_comp_C {σ : Type u_1} {R : Type u_3} {S : Type u_4} [CommSemiring R] [CommSemiring S] (f : R →+* S) :
          RingHom.comp (MvPolynomial.map f) MvPolynomial.C = RingHom.comp MvPolynomial.C f
          theorem MvPolynomial.hom_bind₁ {σ : Type u_1} {τ : Type u_2} {R : Type u_3} {S : Type u_4} [CommSemiring R] [CommSemiring S] (f : MvPolynomial τ R →+* S) (g : σMvPolynomial τ R) (φ : MvPolynomial σ R) :
          f ((MvPolynomial.bind₁ g) φ) = (MvPolynomial.eval₂Hom (RingHom.comp f MvPolynomial.C) fun (i : σ) => f (g i)) φ
          theorem MvPolynomial.map_bind₁ {σ : Type u_1} {τ : Type u_2} {R : Type u_3} {S : Type u_4} [CommSemiring R] [CommSemiring S] (f : R →+* S) (g : σMvPolynomial τ R) (φ : MvPolynomial σ R) :
          @[simp]
          theorem MvPolynomial.eval₂Hom_comp_C {σ : Type u_1} {R : Type u_3} {S : Type u_4} [CommSemiring R] [CommSemiring S] (f : R →+* S) (g : σS) :
          RingHom.comp (MvPolynomial.eval₂Hom f g) MvPolynomial.C = f
          theorem MvPolynomial.eval₂Hom_bind₁ {σ : Type u_1} {τ : Type u_2} {R : Type u_3} {S : Type u_4} [CommSemiring R] [CommSemiring S] (f : R →+* S) (g : τS) (h : σMvPolynomial τ R) (φ : MvPolynomial σ R) :
          theorem MvPolynomial.aeval_bind₁ {σ : Type u_1} {τ : Type u_2} {R : Type u_3} {S : Type u_4} [CommSemiring R] [CommSemiring S] [Algebra R S] (f : τS) (g : σMvPolynomial τ R) (φ : MvPolynomial σ R) :
          theorem MvPolynomial.aeval_comp_bind₁ {σ : Type u_1} {τ : Type u_2} {R : Type u_3} {S : Type u_4} [CommSemiring R] [CommSemiring S] [Algebra R S] (f : τS) (g : σMvPolynomial τ R) :
          theorem MvPolynomial.eval₂Hom_bind₂ {σ : Type u_1} {R : Type u_3} {S : Type u_4} {T : Type u_5} [CommSemiring R] [CommSemiring S] [CommSemiring T] (f : S →+* T) (g : σT) (h : R →+* MvPolynomial σ S) (φ : MvPolynomial σ R) :
          theorem MvPolynomial.aeval_bind₂ {σ : Type u_1} {R : Type u_3} {S : Type u_4} {T : Type u_5} [CommSemiring R] [CommSemiring S] [CommSemiring T] [Algebra S T] (f : σT) (g : R →+* MvPolynomial σ S) (φ : MvPolynomial σ R) :
          theorem MvPolynomial.eval₂Hom_C_left {σ : Type u_1} {τ : Type u_2} {R : Type u_3} [CommSemiring R] (f : σMvPolynomial τ R) :
          theorem MvPolynomial.bind₁_monomial {σ : Type u_1} {τ : Type u_2} {R : Type u_3} [CommSemiring R] (f : σMvPolynomial τ R) (d : σ →₀ ) (r : R) :
          (MvPolynomial.bind₁ f) ((MvPolynomial.monomial d) r) = MvPolynomial.C r * Finset.prod d.support fun (i : σ) => f i ^ d i
          theorem MvPolynomial.bind₂_monomial {σ : Type u_1} {R : Type u_3} {S : Type u_4} [CommSemiring R] [CommSemiring S] (f : R →+* MvPolynomial σ S) (d : σ →₀ ) (r : R) :
          theorem MvPolynomial.vars_bind₁ {σ : Type u_1} {τ : Type u_2} {R : Type u_3} [CommSemiring R] [DecidableEq τ] (f : σMvPolynomial τ R) (φ : MvPolynomial σ R) :
          theorem MvPolynomial.mem_vars_bind₁ {σ : Type u_1} {τ : Type u_2} {R : Type u_3} [CommSemiring R] (f : σMvPolynomial τ R) (φ : MvPolynomial σ R) {j : τ} (h : j MvPolynomial.vars ((MvPolynomial.bind₁ f) φ)) :
          instance MvPolynomial.monad {R : Type u_3} [CommSemiring R] :
          Monad fun (σ : Type u_6) => MvPolynomial σ R
          Equations
          • MvPolynomial.monad = Monad.mk
          instance MvPolynomial.lawfulFunctor {R : Type u_3} [CommSemiring R] :
          LawfulFunctor fun (σ : Type u_6) => MvPolynomial σ R
          Equations
          • =
          instance MvPolynomial.lawfulMonad {R : Type u_3} [CommSemiring R] :
          LawfulMonad fun (σ : Type u_6) => MvPolynomial σ R
          Equations
          • =