|
| 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 |
0 commit comments