Skip to content

Commit db5f318

Browse files
Parcly-Taxelkim-emjcommelin
committed
feat: port Data.MvPolynomial.Funext (#3225)
Co-authored-by: Scott Morrison <scott.morrison@gmail.com> Co-authored-by: Johan Commelin <johan@commelin.net>
1 parent c80a95d commit db5f318

File tree

5 files changed

+128
-0
lines changed

5 files changed

+128
-0
lines changed

Mathlib.lean

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -921,8 +921,10 @@ import Mathlib.Data.MvPolynomial.Counit
921921
import Mathlib.Data.MvPolynomial.Division
922922
import Mathlib.Data.MvPolynomial.Equiv
923923
import Mathlib.Data.MvPolynomial.Expand
924+
import Mathlib.Data.MvPolynomial.Funext
924925
import Mathlib.Data.MvPolynomial.Invertible
925926
import Mathlib.Data.MvPolynomial.Monad
927+
import Mathlib.Data.MvPolynomial.Polynomial
926928
import Mathlib.Data.MvPolynomial.Rename
927929
import Mathlib.Data.MvPolynomial.Supported
928930
import Mathlib.Data.MvPolynomial.Variables

Mathlib/Data/MvPolynomial/Basic.lean

Lines changed: 17 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1148,6 +1148,17 @@ theorem eval_assoc {τ} (f : σ → MvPolynomial τ R) (g : τ → R) (p : MvPol
11481148
congr with a; simp
11491149
#align mv_polynomial.eval_assoc MvPolynomial.eval_assoc
11501150

1151+
-- Porting note: new theorem
1152+
theorem eval_eval₂ [CommSemiring R] [CommSemiring S]
1153+
(f : R →+* MvPolynomial τ S) (g : σ → MvPolynomial τ S) (p : MvPolynomial σ R) :
1154+
eval x (eval₂ f g p) = eval₂ ((eval x).comp f) (fun s => eval x (g s)) p := by
1155+
apply induction_on p
1156+
· simp
1157+
· intro p q hp hq
1158+
simp [hp, hq]
1159+
· intro p n hp
1160+
simp [hp]
1161+
11511162
end Eval
11521163

11531164
section Map
@@ -1204,6 +1215,12 @@ theorem eval₂_eq_eval_map (g : σ → S₁) (p : MvPolynomial σ R) : p.eval
12041215
simp only [comp_apply, eval₂_X]
12051216
#align mv_polynomial.eval₂_eq_eval_map MvPolynomial.eval₂_eq_eval_map
12061217

1218+
-- Porting note: new theorem
1219+
-- This probably belongs earlier, but it breaks the fragile proof of `eval₂_eq_eval_map`
1220+
@[simp]
1221+
theorem eval₂_id (p : MvPolynomial σ R) : eval₂ (RingHom.id _) g p = eval g p :=
1222+
rfl
1223+
12071224
theorem eval₂_comp_right {S₂} [CommSemiring S₂] (k : S₁ →+* S₂) (f : R →+* S₁) (g : σ → S₁) (p) :
12081225
k (eval₂ f g p) = eval₂ k (k ∘ g) (map f p) := by
12091226
apply MvPolynomial.induction_on p
Lines changed: 70 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,70 @@
1+
/-
2+
Copyright (c) 2020 Johan Commelin. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Johan Commelin
5+
6+
! This file was ported from Lean 3 source module data.mv_polynomial.funext
7+
! leanprover-community/mathlib commit da01792ca4894d4f3a98d06b6c50455e5ed25da3
8+
! Please do not edit these lines, except to modify the commit id
9+
! if you have ported upstream changes.
10+
-/
11+
import Mathlib.Data.Polynomial.RingDivision
12+
import Mathlib.Data.MvPolynomial.Rename
13+
import Mathlib.Data.MvPolynomial.Polynomial
14+
import Mathlib.RingTheory.Polynomial.Basic
15+
16+
/-!
17+
## Function extensionality for multivariate polynomials
18+
19+
In this file we show that two multivariate polynomials over an infinite integral domain are equal
20+
if they are equal upon evaluating them on an arbitrary assignment of the variables.
21+
22+
# Main declaration
23+
24+
* `MvPolynomial.funext`: two polynomials `φ ψ : MvPolynomial σ R`
25+
over an infinite integral domain `R` are equal if `eval x φ = eval x ψ` for all `x : σ → R`.
26+
27+
-/
28+
29+
namespace MvPolynomial
30+
31+
variable {R : Type _} [CommRing R] [IsDomain R] [Infinite R]
32+
33+
private theorem funext_fin {n : ℕ} {p : MvPolynomial (Fin n) R}
34+
(h : ∀ x : Fin n → R, eval x p = 0) : p = 0 := by
35+
induction' n with n ih
36+
· apply (MvPolynomial.isEmptyRingEquiv R (Fin 0)).injective
37+
rw [RingEquiv.map_zero]
38+
convert h finZeroElim
39+
· apply (finSuccEquiv R n).injective
40+
simp only [AlgEquiv.map_zero]
41+
refine Polynomial.funext fun q => ?_
42+
rw [Polynomial.eval_zero]
43+
apply ih fun x => ?_
44+
calc _ = _ := eval_polynomial_eval_finSuccEquiv p _
45+
_ = 0 := h _
46+
47+
/-- Two multivariate polynomials over an infinite integral domain are equal
48+
if they are equal upon evaluating them on an arbitrary assignment of the variables. -/
49+
theorem funext {σ : Type _} {p q : MvPolynomial σ R} (h : ∀ x : σ → R, eval x p = eval x q) :
50+
p = q := by
51+
suffices ∀ p, (∀ x : σ → R, eval x p = 0) → p = 0 by
52+
rw [← sub_eq_zero, this (p - q)]
53+
simp only [h, RingHom.map_sub, forall_const, sub_self]
54+
clear h p q
55+
intro p h
56+
obtain ⟨n, f, hf, p, rfl⟩ := exists_fin_rename p
57+
suffices p = 0 by rw [this, AlgHom.map_zero]
58+
apply funext_fin
59+
intro x
60+
classical
61+
convert h (Function.extend f x 0)
62+
simp only [eval, eval₂Hom_rename, Function.extend_comp hf]
63+
#align mv_polynomial.funext MvPolynomial.funext
64+
65+
theorem funext_iff {σ : Type _} {p q : MvPolynomial σ R} :
66+
p = q ↔ ∀ x : σ → R, eval x p = eval x q :=
67+
by rintro rfl; simp only [forall_const, eq_self_iff_true], funext⟩
68+
#align mv_polynomial.funext_iff MvPolynomial.funext_iff
69+
70+
end MvPolynomial
Lines changed: 36 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,36 @@
1+
/-
2+
Copyright (c) 2023 Scott Morrison. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Scott Morrison
5+
-/
6+
import Mathlib.Data.MvPolynomial.Equiv
7+
import Mathlib.Data.Polynomial.Eval
8+
9+
/-!
10+
# Some lemmas relating polynomials and multivariable polynomials.
11+
-/
12+
13+
namespace MvPolynomial
14+
15+
theorem polynomial_eval_eval₂ [CommSemiring R] [CommSemiring S]
16+
(f : R →+* Polynomial S) (g : σ → Polynomial S) (p : MvPolynomial σ R) :
17+
Polynomial.eval x (eval₂ f g p) =
18+
eval₂ ((Polynomial.evalRingHom x).comp f) (fun s => Polynomial.eval x (g s)) p := by
19+
apply induction_on p
20+
· simp
21+
· intro p q hp hq
22+
simp [hp, hq]
23+
· intro p n hp
24+
simp [hp]
25+
26+
theorem eval_polynomial_eval_finSuccEquiv
27+
[CommSemiring R] (f : MvPolynomial (Fin (n + 1)) R) (q : MvPolynomial (Fin n) R) :
28+
(eval x) (Polynomial.eval q (finSuccEquiv R n f)) = eval (Fin.cases (eval x q) x) f := by
29+
simp only [finSuccEquiv_apply, coe_eval₂Hom, polynomial_eval_eval₂, eval_eval₂]
30+
conv in RingHom.comp _ _ =>
31+
{ refine @RingHom.ext _ _ _ _ _ (RingHom.id _) fun r => ?_
32+
simp }
33+
simp only [eval₂_id]
34+
congr
35+
funext i
36+
refine Fin.cases (by simp) (by simp) i

Mathlib/Tactic/Conv.lean

Lines changed: 3 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -50,6 +50,9 @@ syntax (name := dischargeConv) "discharge" (" => " tacticSeq)? : conv
5050
setGoals (m.mvarId! :: gs)
5151
| _ => Elab.throwUnsupportedSyntax
5252

53+
/-- Use `refine` in `conv` mode. -/
54+
macro "refine " e:term : conv => `(conv| tactic => refine $e)
55+
5356
open Elab Tactic
5457
/--
5558
The command `#conv tac => e` will run a conv tactic `tac` on `e`, and display the resulting

0 commit comments

Comments
 (0)