Skip to content

Commit 801ae80

Browse files
adomaniurkud
andcommitted
feat(Tactic/ComputeDegree): add helper lemmas for compute_degree_le (#5978)
This PR is a prequel to #5882: it simply adds the helper lemmas about polynomials that the tactic uses. [Zulip discussion ](https://leanprover.zulipchat.com/#narrow/stream/144837-PR-reviews/topic/.235882.20.60compute_degree_le.60) Co-authored-by: Yury G. Kudryashov <urkud@urkud.name>
1 parent 2a2e08c commit 801ae80

File tree

4 files changed

+90
-0
lines changed

4 files changed

+90
-0
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -2925,6 +2925,7 @@ import Mathlib.Tactic.ClearExcept
29252925
import Mathlib.Tactic.Clear_
29262926
import Mathlib.Tactic.Coe
29272927
import Mathlib.Tactic.Common
2928+
import Mathlib.Tactic.ComputeDegree
29282929
import Mathlib.Tactic.Congr!
29292930
import Mathlib.Tactic.Constructor
29302931
import Mathlib.Tactic.Continuity

Mathlib/Data/Polynomial/Degree/Definitions.lean

Lines changed: 48 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -285,6 +285,8 @@ theorem natDegree_nat_cast (n : ℕ) : natDegree (n : R[X]) = 0 := by
285285
simp only [← C_eq_nat_cast, natDegree_C]
286286
#align polynomial.nat_degree_nat_cast Polynomial.natDegree_nat_cast
287287

288+
theorem degree_nat_cast_le (n : ℕ) : degree (n : R[X]) ≤ 0 := degree_le_of_natDegree_le (by simp)
289+
288290
@[simp]
289291
theorem degree_monomial (n : ℕ) (ha : a ≠ 0) : degree (monomial n a) = n := by
290292
rw [degree, support_monomial n ha]; rfl
@@ -537,15 +539,23 @@ theorem coeff_mul_X_sub_C {p : R[X]} {r : R} {a : ℕ} :
537539
theorem degree_neg (p : R[X]) : degree (-p) = degree p := by unfold degree; rw [support_neg]
538540
#align polynomial.degree_neg Polynomial.degree_neg
539541

542+
theorem degree_neg_le_of_le {a : WithBot ℕ} {p : R[X]} (hp : degree p ≤ a) : degree (- p) ≤ a :=
543+
p.degree_neg.le.trans ‹_›
544+
540545
@[simp]
541546
theorem natDegree_neg (p : R[X]) : natDegree (-p) = natDegree p := by simp [natDegree]
542547
#align polynomial.nat_degree_neg Polynomial.natDegree_neg
543548

549+
theorem natDegree_neg_le_of_le {p : R[X]} (hp : natDegree p ≤ m) : natDegree (- p) ≤ m :=
550+
(natDegree_neg p).le.trans ‹_›
551+
544552
@[simp]
545553
theorem natDegree_int_cast (n : ℤ) : natDegree (n : R[X]) = 0 := by
546554
rw [← C_eq_int_cast, natDegree_C]
547555
#align polynomial.nat_degree_int_cast Polynomial.natDegree_int_cast
548556

557+
theorem degree_int_cast_le (n : ℤ) : degree (n : R[X]) ≤ 0 := degree_le_of_natDegree_le (by simp)
558+
549559
@[simp]
550560
theorem leadingCoeff_neg (p : R[X]) : (-p).leadingCoeff = -p.leadingCoeff := by
551561
rw [leadingCoeff, leadingCoeff, natDegree_neg, coeff_neg]
@@ -643,6 +653,10 @@ theorem degree_add_le_of_degree_le {p q : R[X]} {n : ℕ} (hp : degree p ≤ n)
643653
(degree_add_le p q).trans <| max_le hp hq
644654
#align polynomial.degree_add_le_of_degree_le Polynomial.degree_add_le_of_degree_le
645655

656+
theorem degree_add_le_of_le {a b : WithBot ℕ} (hp : degree p ≤ a) (hq : degree q ≤ b) :
657+
degree (p + q) ≤ max a b :=
658+
(p.degree_add_le q).trans <| max_le_max ‹_› ‹_›
659+
646660
theorem natDegree_add_le (p q : R[X]) : natDegree (p + q) ≤ max (natDegree p) (natDegree q) := by
647661
cases' le_max_iff.1 (degree_add_le p q) with h h <;> simp [natDegree_le_natDegree h]
648662
#align polynomial.nat_degree_add_le Polynomial.natDegree_add_le
@@ -652,6 +666,10 @@ theorem natDegree_add_le_of_degree_le {p q : R[X]} {n : ℕ} (hp : natDegree p
652666
(natDegree_add_le p q).trans <| max_le hp hq
653667
#align polynomial.nat_degree_add_le_of_degree_le Polynomial.natDegree_add_le_of_degree_le
654668

669+
theorem natDegree_add_le_of_le (hp : natDegree p ≤ m) (hq : natDegree q ≤ n) :
670+
natDegree (p + q) ≤ max m n :=
671+
(p.natDegree_add_le q).trans <| max_le_max ‹_› ‹_›
672+
655673
@[simp]
656674
theorem leadingCoeff_zero : leadingCoeff (0 : R[X]) = 0 :=
657675
rfl
@@ -782,6 +800,10 @@ theorem degree_mul_le (p q : R[X]) : degree (p * q) ≤ degree p + degree q :=
782800
exact add_le_add (le_degree_of_ne_zero ha) (le_degree_of_ne_zero hb)
783801
#align polynomial.degree_mul_le Polynomial.degree_mul_le
784802

803+
theorem degree_mul_le_of_le {a b : WithBot ℕ} (hp : degree p ≤ a) (hq : degree q ≤ b) :
804+
degree (p * q) ≤ a + b :=
805+
(p.degree_mul_le _).trans <| add_le_add ‹_› ‹_›
806+
785807
theorem degree_pow_le (p : R[X]) : ∀ n : ℕ, degree (p ^ n) ≤ n • degree p
786808
| 0 => by rw [pow_zero, zero_nsmul]; exact degree_one_le
787809
| n + 1 =>
@@ -791,6 +813,14 @@ theorem degree_pow_le (p : R[X]) : ∀ n : ℕ, degree (p ^ n) ≤ n • degree
791813
_ ≤ _ := by rw [succ_nsmul]; exact add_le_add le_rfl (degree_pow_le _ _)
792814
#align polynomial.degree_pow_le Polynomial.degree_pow_le
793815

816+
theorem degree_pow_le_of_le {a : WithBot ℕ} (b : ℕ) (hp : degree p ≤ a) :
817+
degree (p ^ b) ≤ b * a := by
818+
induction b with
819+
| zero => simp [degree_one_le]
820+
| succ n hn =>
821+
rw [Nat.cast_succ, add_mul, one_mul, pow_succ']
822+
exact degree_mul_le_of_le hn hp
823+
794824
@[simp]
795825
theorem leadingCoeff_monomial (a : R) (n : ℕ) : leadingCoeff (monomial n a) = a := by
796826
by_cases ha : a = 0
@@ -1034,6 +1064,10 @@ theorem natDegree_mul_le {p q : R[X]} : natDegree (p * q) ≤ natDegree p + natD
10341064
refine' add_le_add _ _ <;> apply degree_le_natDegree
10351065
#align polynomial.nat_degree_mul_le Polynomial.natDegree_mul_le
10361066

1067+
theorem natDegree_mul_le_of_le (hp : natDegree p ≤ m) (hg : natDegree q ≤ n) :
1068+
natDegree (p * q) ≤ m + n :=
1069+
natDegree_mul_le.trans <| add_le_add ‹_› ‹_›
1070+
10371071
theorem natDegree_pow_le {p : R[X]} {n : ℕ} : (p ^ n).natDegree ≤ n * p.natDegree := by
10381072
induction' n with i hi
10391073
· simp
@@ -1042,6 +1076,10 @@ theorem natDegree_pow_le {p : R[X]} {n : ℕ} : (p ^ n).natDegree ≤ n * p.natD
10421076
exact add_le_add_left hi _
10431077
#align polynomial.nat_degree_pow_le Polynomial.natDegree_pow_le
10441078

1079+
theorem natDegree_pow_le_of_le (n : ℕ) (hp : natDegree p ≤ m) :
1080+
natDegree (p ^ n) ≤ n * m :=
1081+
natDegree_pow_le.trans (Nat.mul_le_mul rfl.le ‹_›)
1082+
10451083
@[simp]
10461084
theorem coeff_pow_mul_natDegree (p : R[X]) (n : ℕ) :
10471085
(p ^ n).coeff (n * p.natDegree) = p.leadingCoeff ^ n := by
@@ -1075,6 +1113,8 @@ theorem natDegree_eq_zero_iff_degree_le_zero : p.natDegree = 0 ↔ p.degree ≤
10751113
Nat.cast_withBot, WithBot.coe_zero]
10761114
#align polynomial.nat_degree_eq_zero_iff_degree_le_zero Polynomial.natDegree_eq_zero_iff_degree_le_zero
10771115

1116+
theorem degree_zero_le : degree (0 : R[X]) ≤ 0 := natDegree_eq_zero_iff_degree_le_zero.mp rfl
1117+
10781118
theorem degree_le_iff_coeff_zero (f : R[X]) (n : WithBot ℕ) :
10791119
degree f ≤ n ↔ ∀ m : ℕ, n < m → coeff f m = 0 := by
10801120
-- Porting note: `Nat.cast_withBot` is required.
@@ -1294,10 +1334,18 @@ theorem degree_sub_le (p q : R[X]) : degree (p - q) ≤ max (degree p) (degree q
12941334
simpa only [degree_neg q] using degree_add_le p (-q)
12951335
#align polynomial.degree_sub_le Polynomial.degree_sub_le
12961336

1337+
theorem degree_sub_le_of_le {a b : WithBot ℕ} (hp : degree p ≤ a) (hq : degree q ≤ b) :
1338+
degree (p - q) ≤ max a b :=
1339+
(p.degree_sub_le q).trans <| max_le_max ‹_› ‹_›
1340+
12971341
theorem natDegree_sub_le (p q : R[X]) : natDegree (p - q) ≤ max (natDegree p) (natDegree q) := by
12981342
simpa only [← natDegree_neg q] using natDegree_add_le p (-q)
12991343
#align polynomial.nat_degree_sub_le Polynomial.natDegree_sub_le
13001344

1345+
theorem natDegree_sub_le_of_le (hp : natDegree p ≤ m) (hq : natDegree q ≤ n) :
1346+
natDegree (p - q) ≤ max m n :=
1347+
(p.natDegree_sub_le q).trans <| max_le_max ‹_› ‹_›
1348+
13011349
theorem degree_sub_lt (hd : degree p = degree q) (hp0 : p ≠ 0)
13021350
(hlc : leadingCoeff p = leadingCoeff q) : degree (p - q) < degree p :=
13031351
have hp : monomial (natDegree p) (leadingCoeff p) + p.erase (natDegree p) = p :=

Mathlib/Tactic.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -24,6 +24,7 @@ import Mathlib.Tactic.ClearExcept
2424
import Mathlib.Tactic.Clear_
2525
import Mathlib.Tactic.Coe
2626
import Mathlib.Tactic.Common
27+
import Mathlib.Tactic.ComputeDegree
2728
import Mathlib.Tactic.Congr!
2829
import Mathlib.Tactic.Constructor
2930
import Mathlib.Tactic.Continuity

Mathlib/Tactic/ComputeDegree.lean

Lines changed: 40 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,40 @@
1+
/-
2+
Copyright (c) 2023 Damiano Testa. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Damiano Testa
5+
-/
6+
7+
import Mathlib.Data.Polynomial.Degree.Definitions
8+
9+
/-!
10+
### Simple lemmas about `natDegree`
11+
12+
The lemmas in this section all have the form `natDegree <some form of cast> ≤ 0`.
13+
Their proofs are weakenings of the stronger lemmas `natDegree <same> = 0`.
14+
These are the lemmas called by `compute_degree_le` on (almost) all the leaves of its recursion.
15+
-/
16+
17+
open Polynomial
18+
19+
namespace Mathlib.Tactic.ComputeDegree
20+
21+
section leaf_lemmas
22+
23+
variable {R : Type _}
24+
25+
section semiring
26+
variable [Semiring R]
27+
28+
theorem natDegree_C_le (a : R) : natDegree (C a) ≤ 0 := (natDegree_C a).le
29+
theorem natDegree_nat_cast_le (n : ℕ) : natDegree (n : R[X]) ≤ 0 := (natDegree_nat_cast _).le
30+
theorem natDegree_zero_le : natDegree (0 : R[X]) ≤ 0 := natDegree_zero.le
31+
theorem natDegree_one_le : natDegree (1 : R[X]) ≤ 0 := natDegree_one.le
32+
33+
end semiring
34+
35+
variable [Ring R] in
36+
theorem natDegree_int_cast_le (n : ℤ) : natDegree (n : R[X]) ≤ 0 := (natDegree_int_cast _).le
37+
38+
end leaf_lemmas
39+
40+
end Mathlib.Tactic.ComputeDegree

0 commit comments

Comments
 (0)