diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean index d3cc6dc2264fe..fc127a34958cc 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean @@ -164,7 +164,6 @@ lemma irreducible_polynomial [IsDomain R] : Irreducible W.polynomial := by iterate 2 rw [degree_add_eq_right_of_degree_lt] <;> simp only [h] <;> decide iterate 2 rw [degree_add_eq_left_of_degree_lt] <;> simp only [h] <;> decide --- Porting note (#10619): removed `@[simp]` to avoid a `simpNF` linter error lemma evalEval_polynomial (x y : R) : W.polynomial.evalEval x y = y ^ 2 + W.a₁ * x * y + W.a₃ * y - (x ^ 3 + W.a₂ * x ^ 2 + W.a₄ * x + W.a₆) := by simp only [polynomial] @@ -183,7 +182,6 @@ lemma equation_iff' (x y : R) : W.Equation x y ↔ y ^ 2 + W.a₁ * x * y + W.a₃ * y - (x ^ 3 + W.a₂ * x ^ 2 + W.a₄ * x + W.a₆) = 0 := by rw [Equation, evalEval_polynomial] --- Porting note (#10619): removed `@[simp]` to avoid a `simpNF` linter error lemma equation_iff (x y : R) : W.Equation x y ↔ y ^ 2 + W.a₁ * x * y + W.a₃ * y = x ^ 3 + W.a₂ * x ^ 2 + W.a₄ * x + W.a₆ := by rw [equation_iff', sub_eq_zero] @@ -210,7 +208,6 @@ TODO: define this in terms of `Polynomial.derivative`. -/ noncomputable def polynomialX : R[X][Y] := C (C W.a₁) * Y - C (C 3 * X ^ 2 + C (2 * W.a₂) * X + C W.a₄) --- Porting note (#10619): removed `@[simp]` to avoid a `simpNF` linter error lemma evalEval_polynomialX (x y : R) : W.polynomialX.evalEval x y = W.a₁ * y - (3 * x ^ 2 + 2 * W.a₂ * x + W.a₄) := by simp only [polynomialX] @@ -226,7 +223,6 @@ TODO: define this in terms of `Polynomial.derivative`. -/ noncomputable def polynomialY : R[X][Y] := C (C 2) * Y + C (C W.a₁ * X + C W.a₃) --- Porting note (#10619): removed `@[simp]` to avoid a `simpNF` linter error lemma evalEval_polynomialY (x y : R) : W.polynomialY.evalEval x y = 2 * y + W.a₁ * x + W.a₃ := by simp only [polynomialY] @@ -256,7 +252,6 @@ lemma nonsingular_iff' (x y : R) : W.Nonsingular x y ↔ W.Equation x y ∧ (W.a₁ * y - (3 * x ^ 2 + 2 * W.a₂ * x + W.a₄) ≠ 0 ∨ 2 * y + W.a₁ * x + W.a₃ ≠ 0) := by rw [Nonsingular, equation_iff', evalEval_polynomialX, evalEval_polynomialY] --- Porting note (#10619): removed `@[simp]` to avoid a `simpNF` linter error lemma nonsingular_iff (x y : R) : W.Nonsingular x y ↔ W.Equation x y ∧ (W.a₁ * y ≠ 3 * x ^ 2 + 2 * W.a₂ * x + W.a₄ ∨ y ≠ -y - W.a₁ * x - W.a₃) := by rw [nonsingular_iff', sub_ne_zero, ← sub_ne_zero (a := y)] @@ -314,7 +309,6 @@ lemma negY_negY (x y : R) : W.negY x (W.negY x y) = y := by simp only [negY] ring1 --- Porting note (#10619): removed `@[simp]` to avoid a `simpNF` linter error lemma eval_negPolynomial (x y : R) : W.negPolynomial.evalEval x y = W.negY x y := by rw [negY, sub_sub, negPolynomial] eval_simp @@ -609,7 +603,6 @@ instance : Inhabited W.Point := instance : Zero W.Point := ⟨zero⟩ --- Porting note (#10619): removed `@[simp]` to avoid a `simpNF` linter error lemma zero_def : (zero : W.Point) = 0 := rfl @@ -625,7 +618,6 @@ def neg : W.Point → W.Point instance : Neg W.Point := ⟨neg⟩ --- Porting note (#10619): removed `@[simp]` to avoid a `simpNF` linter error lemma neg_def (P : W.Point) : P.neg = -P := rfl @@ -656,7 +648,6 @@ noncomputable def add : W.Point → W.Point → W.Point noncomputable instance instAddPoint : Add W.Point := ⟨add⟩ --- Porting note (#10619): removed `@[simp]` to avoid a `simpNF` linter error lemma add_def (P Q : W.Point) : P.add Q = P + Q := rfl diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean index 28d32c6a3fea5..d81a89407576b 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean @@ -213,7 +213,6 @@ section Ring /-! ### Ideals in the coordinate ring over a ring -/ --- Porting note (#10619): removed `@[simp]` to avoid a `simpNF` linter error /-- The class of the element $X - x$ in $R[W]$ for some $x \in R$. -/ noncomputable def XClass (x : R) : W.CoordinateRing := mk W <| C <| X - C x @@ -222,7 +221,6 @@ lemma XClass_ne_zero [Nontrivial R] (x : R) : XClass W x ≠ 0 := AdjoinRoot.mk_ne_zero_of_natDegree_lt W.monic_polynomial (C_ne_zero.mpr <| X_sub_C_ne_zero x) <| by rw [natDegree_polynomial, natDegree_C]; norm_num1 --- Porting note (#10619): removed `@[simp]` to avoid a `simpNF` linter error /-- The class of the element $Y - y(X)$ in $R[W]$ for some $y(X) \in R[X]$. -/ noncomputable def YClass (y : R[X]) : W.CoordinateRing := mk W <| Y - C y @@ -235,17 +233,14 @@ lemma C_addPolynomial (x y L : R) : mk W (C <| W.addPolynomial x y L) = mk W ((Y - C (linePolynomial x y L)) * (W.negPolynomial - C (linePolynomial x y L))) := AdjoinRoot.mk_eq_mk.mpr ⟨1, by rw [W.C_addPolynomial, add_sub_cancel_left, mul_one]⟩ --- Porting note (#10619): removed `@[simp]` to avoid a `simpNF` linter error /-- The ideal $\langle X - x \rangle$ of $R[W]$ for some $x \in R$. -/ noncomputable def XIdeal (x : R) : Ideal W.CoordinateRing := span {XClass W x} --- Porting note (#10619): removed `@[simp]` to avoid a `simpNF` linter error /-- The ideal $\langle Y - y(X) \rangle$ of $R[W]$ for some $y(X) \in R[X]$. -/ noncomputable def YIdeal (y : R[X]) : Ideal W.CoordinateRing := span {YClass W y} --- Porting note (#10619): removed `@[simp]` to avoid a `simpNF` linter error /-- The ideal $\langle X - x, Y - y(X) \rangle$ of $R[W]$ for some $x \in R$ and $y(X) \in R[X]$. -/ noncomputable def XYIdeal (x : R) (y : R[X]) : Ideal W.CoordinateRing := span {XClass W x, YClass W y} @@ -387,7 +382,6 @@ lemma XYIdeal_mul_XYIdeal {x₁ x₂ y₁ y₂ : F} (h₁ : W.Equation x₁ y₁ C_simp ring1 --- Porting note (#10619): removed `@[simp]` to avoid a `simpNF` linter error /-- The non-zero fractional ideal $\langle X - x, Y - y \rangle$ of $F(W)$ for some $x, y \in F$. -/ noncomputable def XYIdeal' {x y : F} (h : W.Nonsingular x y) : (FractionalIdeal W.CoordinateRing⁰ W.FunctionField)ˣ := @@ -518,7 +512,6 @@ noncomputable def toClass : W.Point →+ Additive (ClassGroup W.CoordinateRing) rw [add_of_imp h] exact (CoordinateRing.mk_XYIdeal'_mul_mk_XYIdeal' h₁ h₂ h).symm --- Porting note (#10619): removed `@[simp]` to avoid a `simpNF` linter error lemma toClass_zero : toClass (0 : W.Point) = 0 := rfl diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Weierstrass.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Weierstrass.lean index 87e3695c1b4d9..9d730cf3c9239 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Weierstrass.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Weierstrass.lean @@ -103,22 +103,18 @@ section Quantity /-! ### Standard quantities -/ --- Porting note (#10619): removed `@[simp]` to avoid a `simpNF` linter error /-- The `b₂` coefficient of a Weierstrass curve. -/ def b₂ : R := W.a₁ ^ 2 + 4 * W.a₂ --- Porting note (#10619): removed `@[simp]` to avoid a `simpNF` linter error /-- The `b₄` coefficient of a Weierstrass curve. -/ def b₄ : R := 2 * W.a₄ + W.a₁ * W.a₃ --- Porting note (#10619): removed `@[simp]` to avoid a `simpNF` linter error /-- The `b₆` coefficient of a Weierstrass curve. -/ def b₆ : R := W.a₃ ^ 2 + 4 * W.a₆ --- Porting note (#10619): removed `@[simp]` to avoid a `simpNF` linter error /-- The `b₈` coefficient of a Weierstrass curve. -/ def b₈ : R := W.a₁ ^ 2 * W.a₆ + 4 * W.a₂ * W.a₆ - W.a₁ * W.a₃ * W.a₄ + W.a₂ * W.a₃ ^ 2 - W.a₄ ^ 2 @@ -127,17 +123,14 @@ lemma b_relation : 4 * W.b₈ = W.b₂ * W.b₆ - W.b₄ ^ 2 := by simp only [b₂, b₄, b₆, b₈] ring1 --- Porting note (#10619): removed `@[simp]` to avoid a `simpNF` linter error /-- The `c₄` coefficient of a Weierstrass curve. -/ def c₄ : R := W.b₂ ^ 2 - 24 * W.b₄ --- Porting note (#10619): removed `@[simp]` to avoid a `simpNF` linter error /-- The `c₆` coefficient of a Weierstrass curve. -/ def c₆ : R := -W.b₂ ^ 3 + 36 * W.b₂ * W.b₄ - 216 * W.b₆ --- Porting note (#10619): removed `@[simp]` to avoid a `simpNF` linter error /-- The discriminant `Δ` of a Weierstrass curve. If `R` is a field, then this polynomial vanishes if and only if the cubic curve cut out by this equation is singular. Sometimes only defined up to sign in the literature; we choose the sign used by the LMFDB. For more discussion, see @@ -438,7 +431,6 @@ theorem ext {x y : EllipticCurve R} (h₁ : x.a₁ = y.a₁) (h₂ : x.a₂ = y. variable (E : EllipticCurve R) --- Porting note (#10619): removed `@[simp]` to avoid a `simpNF` linter error /-- The j-invariant `j` of an elliptic curve, which is invariant under isomorphisms over `R`. -/ def j : R := E.Δ'⁻¹ * E.c₄ ^ 3