Built with Alectryon, running Lean4 v4.19.0. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
/- - Created in 2024 by Gaëtan Serré-//-- https://github.com/gaetanserre/SBS-Proofs-/import Mathlib.Data.EReal.Basic
import Mathlib.Analysis.InnerProductSpace.Basic
import Mathlib.MeasureTheory.Measure.Lebesgue.Basic
import Mathlib.MeasureTheory.Integral.Bochner.Basic
import Mathlib.MeasureTheory.Integral.Bochner.L1
import Mathlib.MeasureTheory.Integral.Bochner.VitaliCaratheodory
import SBSProofs.Utils
import SBSProofs.Measure
import SBSProofs.RKHS.Basic
local macro_rules |`($x ^$y) =>`(HPow.hPow $x $y)
open scoped RealInnerProductSpace
open BigOperators Finset ENNReal NNReal MeasureTheory RKHS
set_option maxHeartbeats 400000/- We defined measures μ and π (ν is considered as the standard Lebesgue measure) along with their densities (finite and non-zero on the entire space)-/variable {d : ℕ} (hd : d ≠0) {Ω : Set (Vector ℝ d)}
variable [MeasureSpace Ω]
variable (μ π : DensityMeasure Ω) [IsProbabilityMeasure μ.toMeasure] [IsProbabilityMeasure π.toMeasure]
variable (hdμ : ∀x, μ.d x ≠0) (hdπ : ∀x, π.d x ≠0)
/- We define a RKHS of Ω → ℝ functions.-/variable (H₀ : Set (Ω→ ℝ)) [NormedAddCommGroup H₀] [InnerProductSpace ℝ H₀] [s : RKHS H₀]
/--We consider that the left-hand side of the equivalence holds for all x. In the future, we want to take into account that it only holds for almost all x w.r.t. μ.-/def positive_definite_kernel := ∀ (f : Fin d →Ω→ ℝ), (0≤∫ x in Set.univ, (∫ x' in Set.univ, (∑ i ∈ Set.univ, f i x * s.k x x' * f i x') ∂μ.toMeasure) ∂μ.toMeasure) ∧ (∫ x in Set.univ, (∫ x' in Set.univ, (∑ i ∈ Set.univ, f i x * s.k x x' * f i x') ∂μ.toMeasure) ∂μ.toMeasure =0↔∀i, ∀ᵐ x ∂ μ.toMeasure, f i x =0)
variable (h_kernel_positive : positive_definite_kernel μ H₀)
/-===============================KERNEL STEIN DISCREPANCY===============================-//-Here, we prove that KSD(μ | π) is a valid discrepancy measure and that π is the only fixed point of Φₜ(μ).-//- From here, as the derivative of multivariate function are hard to define and to manipulate (defining the gradient, the divergence...), we define the gradient of *f* as follows: f : Ω → ℝ df : Fin d → Ω → ℝ i ↦ x ↦ ∂xⁱ f(x) For vector-valued function, we defined them as follows: f : Fin d → Ω → ℝ i ↦ x ↦ f(x)ⁱ df : Fin d → Ω → ℝ i ↦ x ↦ ∂xⁱ f(x)ⁱ Also, we assume some simple lemmas using the above formalism. Sometimes, these lemmas are not rigorously defined but, in our framework, it is more than enough.-//- dk : x ↦ i ↦ y ↦ ∂xⁱ k(x, y) -/variable (dk : Ω→ Fin d →Ω→ ℝ)
/- d_ln_π : i ↦ x ↦ ∂xⁱ ln (μ(x) / π(x)) -/variable (d_ln_π : Fin d →Ω→ ℝ)
/- Definition of the steepest direction ϕ and its derivative.-/variable (ϕ : product_RKHS H₀ hd) (dϕ : Fin d →Ω→ ℝ)
/- As Ω is supposed to be compact, we will use this assumption only when the function is trivially integrable (e.g. continuous on compact). -/axiom is_integrable_H₀ : ∀ (f : Ω→ ℝ), Integrable f μ.toMeasure
axiom is_integrable_H₀_volume : ∀ (f : Ω→ ℝ), Integrable f
axiom is_measurable_H₀ : ∀ (f : Ω→ ℝ), Measurable f
axiom is_measurable_H₀_enn : ∀ (f : Ω→ ℝ≥0∞), Measurable f
/-d_ln_π_μ : i ↦ x ↦ ∂xⁱ ln (π(x) / μ(x))-/variable (d_ln_π_μ : Fin d →Ω→ ℝ)
/-Simple derivative rule: if the derivative is 0 ∀x, then the function is constant.-/variable (hd_ln_π_μ : ∀ (ν : Measure Ω), (∀i, ∀ᵐ x ∂ν, d_ln_π_μ i x =0) → (∃ c, ∀ᵐ x ∂ν, log (μ.d x / π.d x) = c))
/-dπ' : i ↦ x ↦ ∂xⁱ π(x)-/variable (dπ' : Fin d →Ω→ ℝ)
/-Log-trick: ∂xⁱ ln (π(x)) * π(x) = ∂xⁱ π(x).-/variable (hπ' : ∀x, ∀i, ENNReal.toReal (π.d x) * d_ln_π i x = dπ' i x)
variable [Norm Ω]
/-- Stein class of measure. f is in the Stein class of μ if, ∀i ∈ Set.univ, lim_(‖x‖ → ∞) μ(x) * ϕ(x)ⁱ = 0.-/def SteinClass (f : Fin d →Ω→ ℝ) := ∀ x, tends_to_infty (λ (x : Ω) ↦‖x‖) →∀i, ENNReal.toReal (μ.d x) * f i x =0/- Kernel Stein Discrepancy-/variable (KSD : Measure Ω→ Measure Ω→ ℝ)
/--KSD(μ | π) = ⟪∇ln π/μ, Pμ ∇ln π/μ⟫_L²(μ). We assume here that KSD is also equal to ∫ x, ∑ i ∈ Set.univ, (d_ln_π i x * ϕ i x + dϕ i x) ∂μ.-/def is_ksd := KSD μ.toMeasure π.toMeasure = (∫ x in Set.univ, (∫ x' in Set.univ, (∑ i ∈ Set.univ, d_ln_π_μ i x * s.k x x' * d_ln_π_μ i x') ∂μ.toMeasure) ∂μ.toMeasure) ∧ (KSD μ.toMeasure π.toMeasure =∫ x, ∑ i ∈ Set.univ, (d_ln_π i x * (ϕ i).1 x + dϕ i x) ∂μ.toMeasure)
/- KSD(μ | π) is originally defined as ‖ϕ^⋆‖²_H, it is therefore non-negative.-/def is_ksd_norm := KSD μ.toMeasure π.toMeasure =‖ϕ‖^2/- ϕ is in the Stein class of π-/variable (hstein : SteinClass π (λ i ↦ (ϕ i).1))
include hπ' h_kernel_positive hstein hd_ln_π_μ hdπ hdμ in/-- We show that, if ϕ is in the Stein class of π, KSD is a valid discrepancy measure i.e. μ = π ↔ KSD(μ | π) = 0.-/
Goals accomplished!
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x
mp
∫ (x : ↑Ω), ∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) ∂μ.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x
mp
∫ (x : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x ∂μ.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x
mp
∫ (x : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x ∂μ.toMeasure =0
-- Split the integral of sum into sum of integral.
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure
mp
∫ (x : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x ∂μ.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure
mp
∫ (x : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x ∂μ.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure
mp
∫ (x : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x ∂μ.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure
mp
∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure +∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂μ.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure
mp
∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure +∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂μ.toMeasure =0
-- Make the `Set.univ` appears for using the density later.
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure
mp
∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure +∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂μ.toMeasure =0
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure
mp
∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure +∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂μ.toMeasure =0
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure
mp
∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure +∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂μ.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure
mp
∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure +∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂μ.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure
mp
∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure +∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂μ.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure
mp
∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure +∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂μ.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure
mp
∫ (x : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x ∂μ.toMeasure +∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂μ.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure
mp
∫ (x : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x ∂μ.toMeasure +∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂μ.toMeasure =0
-- Replace μ by π in the integration.
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure
mp
∫ (x : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x ∂μ.toMeasure +∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂μ.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure
mp
∫ (x : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x ∂π.toMeasure +∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure
mp
∫ (x : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x ∂π.toMeasure +∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =0
-- Replace by its density.
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure
mp
∫ (x : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x ∂π.toMeasure +∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure
mp
∫ (x : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x ∂π.toMeasure +∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure
mp
(∫ (x : ↑Ω), (π.d x).toReal *∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) +∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure
mp
(∫ (x : ↑Ω), (π.d x).toReal *∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) +∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =0
-- Get ENNReal.toReal (π.d x) in the sum (a * ∑ b = ∑ b * a).
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure
mp
(∫ (x : ↑Ω), (π.d x).toReal *∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) +∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =0
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure
∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure
∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure
∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure mul_dist_sum: ∀ (a : ℝ) (f : Fin d → ℝ), (∑ i ∈ Set.univ.toFinset, f i) * a =∑ i ∈ Set.univ.toFinset, f i * a
∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure
∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure mul_dist_sum: ∀ (a : ℝ) (f : Fin d → ℝ), (∑ i ∈ Set.univ.toFinset, f i) * a =∑ i ∈ Set.univ.toFinset, f i * a x: ↑Ω
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure
∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure mul_dist_sum: ∀ (a : ℝ) (f : Fin d → ℝ), (∑ i ∈ Set.univ.toFinset, f i) * a =∑ i ∈ Set.univ.toFinset, f i * a x: ↑Ω
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure mul_dist_sum: ∀ (a : ℝ) (f : Fin d → ℝ), (∑ i ∈ Set.univ.toFinset, f i) * a =∑ i ∈ Set.univ.toFinset, f i * a x: ↑Ω
(∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i) * (π.d x).toReal =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure mul_dist_sum: ∀ (a : ℝ) (f : Fin d → ℝ), (∑ i ∈ Set.univ.toFinset, f i) * a =∑ i ∈ Set.univ.toFinset, f i * a x: ↑Ω
(∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i) * (π.d x).toReal =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure
∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure mul_dist: ∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal
mp
(∫ (x : ↑Ω), (π.d x).toReal *∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) +∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure mul_dist: ∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal
mp
(∫ (x : ↑Ω), ∑ x_1 ∈ Set.univ.toFinset, d_ln_π x_1 x *↑(ϕ x_1) x * (π.d x).toReal) +∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure mul_dist: ∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal
mp
(∫ (x : ↑Ω), ∑ x_1 ∈ Set.univ.toFinset, d_ln_π x_1 x *↑(ϕ x_1) x * (π.d x).toReal) +∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =0
-- Make the product ENNReal.toReal (π.d x) * d_ln_π i x appears to use the ln derivative rule.
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure mul_dist: ∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal mul_comm: ∀ (x : ↑Ω) (i : Fin d), d_ln_π i x *↑(ϕ i) x * (π.d x).toReal = (π.d x).toReal * d_ln_π i x *↑(ϕ i) x
mp
(∫ (x : ↑Ω), ∑ x_1 ∈ Set.univ.toFinset, d_ln_π x_1 x *↑(ϕ x_1) x * (π.d x).toReal) +∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure mul_dist: ∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal mul_comm: ∀ (x : ↑Ω) (i : Fin d), d_ln_π i x *↑(ϕ i) x * (π.d x).toReal = (π.d x).toReal * d_ln_π i x *↑(ϕ i) x
mp
(∫ (x : ↑Ω), ∑ x_1 ∈ Set.univ.toFinset, d_ln_π x_1 x *↑(ϕ x_1) x * (π.d x).toReal) +∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure mul_dist: ∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal mul_comm: ∀ (x : ↑Ω) (i : Fin d), d_ln_π i x *↑(ϕ i) x * (π.d x).toReal = (π.d x).toReal * d_ln_π i x *↑(ϕ i) x
mp
(∫ (x : ↑Ω), ∑ x_1 ∈ Set.univ.toFinset, (π.d x).toReal * d_ln_π x_1 x *↑(ϕ x_1) x) +∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure mul_dist: ∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal mul_comm: ∀ (x : ↑Ω) (i : Fin d), d_ln_π i x *↑(ϕ i) x * (π.d x).toReal = (π.d x).toReal * d_ln_π i x *↑(ϕ i) x
mp
(∫ (x : ↑Ω), ∑ x_1 ∈ Set.univ.toFinset, d_ln_π x_1 x *↑(ϕ x_1) x * (π.d x).toReal) +∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure mul_dist: ∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal mul_comm: ∀ (x : ↑Ω) (i : Fin d), d_ln_π i x *↑(ϕ i) x * (π.d x).toReal = (π.d x).toReal * d_ln_π i x *↑(ϕ i) x
mp
(∫ (x : ↑Ω), ∑ x_1 ∈ Set.univ.toFinset, dπ' x_1 x *↑(ϕ x_1) x) +∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure mul_dist: ∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal mul_comm: ∀ (x : ↑Ω) (i : Fin d), d_ln_π i x *↑(ϕ i) x * (π.d x).toReal = (π.d x).toReal * d_ln_π i x *↑(ϕ i) x
mp
(∫ (x : ↑Ω), ∑ x_1 ∈ Set.univ.toFinset, dπ' x_1 x *↑(ϕ x_1) x) +∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =0
-- Make the `Set.univ` appears to use the density.
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure mul_dist: ∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal mul_comm: ∀ (x : ↑Ω) (i : Fin d), d_ln_π i x *↑(ϕ i) x * (π.d x).toReal = (π.d x).toReal * d_ln_π i x *↑(ϕ i) x
mp
(∫ (x : ↑Ω), ∑ x_1 ∈ Set.univ.toFinset, dπ' x_1 x *↑(ϕ x_1) x) +∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =0
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure mul_dist: ∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal mul_comm: ∀ (x : ↑Ω) (i : Fin d), d_ln_π i x *↑(ϕ i) x * (π.d x).toReal = (π.d x).toReal * d_ln_π i x *↑(ϕ i) x
mp
(∫ (x : ↑Ω), ∑ x_1 ∈ Set.univ.toFinset, dπ' x_1 x *↑(ϕ x_1) x) +∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =0
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ✝: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure mul_dist: ∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal mul_comm: ∀ (x : ↑Ω) (i : Fin d), d_ln_π i x *↑(ϕ i) x * (π.d x).toReal = (π.d x).toReal * d_ln_π i x *↑(ϕ i) x int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure
mp
(∫ (x : ↑Ω), ∑ x_1 ∈ Set.univ.toFinset, dπ' x_1 x *↑(ϕ x_1) x) +∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ✝: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure mul_dist: ∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal mul_comm: ∀ (x : ↑Ω) (i : Fin d), d_ln_π i x *↑(ϕ i) x * (π.d x).toReal = (π.d x).toReal * d_ln_π i x *↑(ϕ i) x int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure
mp
(∫ (x : ↑Ω), ∑ x_1 ∈ Set.univ.toFinset, dπ' x_1 x *↑(ϕ x_1) x) +∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ✝: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure mul_dist: ∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal mul_comm: ∀ (x : ↑Ω) (i : Fin d), d_ln_π i x *↑(ϕ i) x * (π.d x).toReal = (π.d x).toReal * d_ln_π i x *↑(ϕ i) x int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure
mp
(∫ (x : ↑Ω), ∑ x_1 ∈ Set.univ.toFinset, dπ' x_1 x *↑(ϕ x_1) x) +∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ✝: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure mul_dist: ∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal mul_comm: ∀ (x : ↑Ω) (i : Fin d), d_ln_π i x *↑(ϕ i) x * (π.d x).toReal = (π.d x).toReal * d_ln_π i x *↑(ϕ i) x int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure
mp
(∫ (x : ↑Ω), ∑ x_1 ∈ Set.univ.toFinset, dπ' x_1 x *↑(ϕ x_1) x) +∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ✝: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure mul_dist: ∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal mul_comm: ∀ (x : ↑Ω) (i : Fin d), d_ln_π i x *↑(ϕ i) x * (π.d x).toReal = (π.d x).toReal * d_ln_π i x *↑(ϕ i) x int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure
mp
(∫ (x : ↑Ω), ∑ x_1 ∈ Set.univ.toFinset, dπ' x_1 x *↑(ϕ x_1) x) +∫ (x : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i x ∂π.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ✝: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure mul_dist: ∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal mul_comm: ∀ (x : ↑Ω) (i : Fin d), d_ln_π i x *↑(ϕ i) x * (π.d x).toReal = (π.d x).toReal * d_ln_π i x *↑(ϕ i) x int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure
mp
(∫ (x : ↑Ω), ∑ x_1 ∈ Set.univ.toFinset, dπ' x_1 x *↑(ϕ x_1) x) +∫ (x : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i x ∂π.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ✝: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi✝: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure mul_dist: ∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal mul_comm: ∀ (x : ↑Ω) (i : Fin d), d_ln_π i x *↑(ϕ i) x * (π.d x).toReal = (π.d x).toReal * d_ln_π i x *↑(ϕ i) x int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) π.toMeasure
mp
(∫ (x : ↑Ω), ∑ x_1 ∈ Set.univ.toFinset, dπ' x_1 x *↑(ϕ x_1) x) +∫ (x : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i x ∂π.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ✝: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi✝: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure mul_dist: ∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal mul_comm: ∀ (x : ↑Ω) (i : Fin d), d_ln_π i x *↑(ϕ i) x * (π.d x).toReal = (π.d x).toReal * d_ln_π i x *↑(ϕ i) x int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) π.toMeasure
mp
(∫ (x : ↑Ω), ∑ x_1 ∈ Set.univ.toFinset, dπ' x_1 x *↑(ϕ x_1) x) +∫ (x : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i x ∂π.toMeasure =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ✝: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi✝: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure mul_dist: ∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal mul_comm: ∀ (x : ↑Ω) (i : Fin d), d_ln_π i x *↑(ϕ i) x * (π.d x).toReal = (π.d x).toReal * d_ln_π i x *↑(ϕ i) x int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) π.toMeasure
mp
(∫ (x : ↑Ω), ∑ x_1 ∈ Set.univ.toFinset, dπ' x_1 x *↑(ϕ x_1) x) +∫ (x : ↑Ω), (π.d x).toReal *∑ i ∈ Set.univ.toFinset, dϕ i x =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ✝: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi✝: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure mul_dist: ∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal mul_comm: ∀ (x : ↑Ω) (i : Fin d), d_ln_π i x *↑(ϕ i) x * (π.d x).toReal = (π.d x).toReal * d_ln_π i x *↑(ϕ i) x int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) π.toMeasure
mp
(∫ (x : ↑Ω), ∑ x_1 ∈ Set.univ.toFinset, dπ' x_1 x *↑(ϕ x_1) x) +∫ (x : ↑Ω), (π.d x).toReal *∑ i ∈ Set.univ.toFinset, dϕ i x =0
-- Use the integration by parts on the right-hand side integral.
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ✝: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi✝: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure mul_dist: ∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal mul_comm: ∀ (x : ↑Ω) (i : Fin d), d_ln_π i x *↑(ϕ i) x * (π.d x).toReal = (π.d x).toReal * d_ln_π i x *↑(ϕ i) x int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) π.toMeasure
mp
(∫ (x : ↑Ω), ∑ x_1 ∈ Set.univ.toFinset, dπ' x_1 x *↑(ϕ x_1) x) +∫ (x : ↑Ω), (π.d x).toReal *∑ i ∈ Set.univ.toFinset, dϕ i x =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ✝: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi✝: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure mul_dist: ∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal mul_comm: ∀ (x : ↑Ω) (i : Fin d), d_ln_π i x *↑(ϕ i) x * (π.d x).toReal = (π.d x).toReal * d_ln_π i x *↑(ϕ i) x int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) π.toMeasure
mp
(∫ (x : ↑Ω), ∑ x_1 ∈ Set.univ.toFinset, dπ' x_1 x *↑(ϕ x_1) x) +-∫ (x : ↑Ω), ∑ i ∈ Set.univ.toFinset, dπ' i x *↑(ϕ i) x =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h: μ.toMeasure = π.toMeasure split_sum: ∀ (x : ↑Ω),
∑ i ∈ Set.univ.toFinset, (d_ln_π i x *↑(ϕ i) x + dϕ i x) =∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x +∑ i ∈ Set.univ.toFinset, dϕ i x h1: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) μ.toMeasure h2: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) μ.toMeasure int_univ✝: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, d_ln_π i a *↑(ϕ i) a ∂μ.toMeasure hi✝: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, d_ln_π i x *↑(ϕ i) x) π.toMeasure mul_dist: ∀ (x : ↑Ω),
(π.d x).toReal *∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i =∑ i ∈ Set.univ.toFinset, (fun i => d_ln_π i x *↑(ϕ i) x) i * (π.d x).toReal mul_comm: ∀ (x : ↑Ω) (i : Fin d), d_ln_π i x *↑(ϕ i) x * (π.d x).toReal = (π.d x).toReal * d_ln_π i x *↑(ϕ i) x int_univ: ∫ (a : ↑Ω), ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure =∫ (a : ↑Ω) in Set.univ, ∑ i ∈ Set.univ.toFinset, dϕ i a ∂π.toMeasure hi: Integrable (fun x =>∑ i ∈ Set.univ.toFinset, dϕ i x) π.toMeasure
mp
(∫ (x : ↑Ω), ∑ x_1 ∈ Set.univ.toFinset, dπ' x_1 x *↑(ϕ x_1) x) +-∫ (x : ↑Ω), ∑ i ∈ Set.univ.toFinset, dπ' i x *↑(ϕ i) x =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c
mpr.intro
μ.toMeasure = π.toMeasure
-- We show that, since, for almost all x, log (μ.d x / π.d x) = c, then μ.d x = exp(c) * π.d x.
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c
mpr.intro
μ.toMeasure = π.toMeasure
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c
∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c
∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x
-- We show that the two sets are equal to show that the complement of right-hand side one is of null measure.
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c
∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c
{x | log (μ.d x / π.d x) = c} = {x | μ.d x = ENNReal.ofReal (Real.exp c) * π.d x}
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c
{x | log (μ.d x / π.d x) = c} = {x | μ.d x = ENNReal.ofReal (Real.exp c) * π.d x}
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c
{x | log (μ.d x / π.d x) = c} = {x | μ.d x = ENNReal.ofReal (Real.exp c) * π.d x}
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω
h
x ∈ {x | log (μ.d x / π.d x) = c} ↔ x ∈ {x | μ.d x = ENNReal.ofReal (Real.exp c) * π.d x}
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c
{x | log (μ.d x / π.d x) = c} = {x | μ.d x = ENNReal.ofReal (Real.exp c) * π.d x}
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω
h.mp
x ∈ {x | log (μ.d x / π.d x) = c} → x ∈ {x | μ.d x = ENNReal.ofReal (Real.exp c) * π.d x}
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω
x ∈ {x | μ.d x = ENNReal.ofReal (Real.exp c) * π.d x} → x ∈ {x | log (μ.d x / π.d x) = c}
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c
{x | log (μ.d x / π.d x) = c} = {x | μ.d x = ENNReal.ofReal (Real.exp c) * π.d x}
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω
h.mp
x ∈ {x | log (μ.d x / π.d x) = c} → x ∈ {x | μ.d x = ENNReal.ofReal (Real.exp c) * π.d x}
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω
h.mp
x ∈ {x | log (μ.d x / π.d x) = c} → x ∈ {x | μ.d x = ENNReal.ofReal (Real.exp c) * π.d x}
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω
x ∈ {x | μ.d x = ENNReal.ofReal (Real.exp c) * π.d x} → x ∈ {x | log (μ.d x / π.d x) = c}
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c
h.mp
x ∈ {x | μ.d x = ENNReal.ofReal (Real.exp c) * π.d x}
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω
h.mp
x ∈ {x | log (μ.d x / π.d x) = c} → x ∈ {x | μ.d x = ENNReal.ofReal (Real.exp c) * π.d x}
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c
h.mp
x ∈ {x | μ.d x = ENNReal.ofReal (Real.exp c) * π.d x}
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c
μ.d x / π.d x ≠0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c
μ.d x / π.d x ≠0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c
μ.d x / π.d x ≠0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c frac_pos: 0< μ.d x / π.d x
μ.d x / π.d x ≠0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c
μ.d x / π.d x ≠0
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω
h.mp
x ∈ {x | log (μ.d x / π.d x) = c} → x ∈ {x | μ.d x = ENNReal.ofReal (Real.exp c) * π.d x}
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c frac_neq_zero: μ.d x / π.d x ≠0
h.mp
x ∈ {x | μ.d x = ENNReal.ofReal (Real.exp c) * π.d x}
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c frac_neq_zero: μ.d x / π.d x ≠0
μ.d x / π.d x ≠⊤
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c frac_neq_zero: μ.d x / π.d x ≠0
μ.d x / π.d x ≠⊤
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c frac_neq_zero: μ.d x / π.d x ≠0
μ.d x / π.d x ≠⊤
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝¹: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h✝: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c frac_neq_zero: μ.d x / π.d x ≠0 h: μ.d x / π.d x =⊤
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c frac_neq_zero: μ.d x / π.d x ≠0
μ.d x / π.d x ≠⊤
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝¹: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h✝: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c frac_neq_zero: μ.d x / π.d x ≠0 h: μ.d x / π.d x =⊤
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝¹: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h✝: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c frac_neq_zero: μ.d x / π.d x ≠0 h: μ.d x ≠0∧ π.d x =0∨ μ.d x =⊤∧ π.d x ≠⊤
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝¹: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h✝: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c frac_neq_zero: μ.d x / π.d x ≠0 h: μ.d x ≠0∧ π.d x =0∨ μ.d x =⊤∧ π.d x ≠⊤
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝¹: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h✝: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c frac_neq_zero: μ.d x / π.d x ≠0 h: μ.d x ≠0∧ π.d x =0∨ μ.d x =⊤∧ π.d x ≠⊤
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c frac_neq_zero: μ.d x / π.d x ≠0
μ.d x / π.d x ≠⊤
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝¹: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h✝: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c frac_neq_zero: μ.d x / π.d x ≠0 h: μ.d x ≠0∧ π.d x =0∨ μ.d x =⊤∧ π.d x ≠⊤
False
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝¹: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h✝: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c frac_neq_zero: μ.d x / π.d x ≠0 h: μ.d x ≠0∧ π.d x =0∨ μ.d x =⊤∧ π.d x ≠⊤
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c frac_neq_zero: μ.d x / π.d x ≠0 hp: μ.d x ≠0∧ π.d x =0
inl
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c frac_neq_zero: μ.d x / π.d x ≠0 hp: μ.d x ≠0∧ π.d x =0
inl
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c frac_neq_zero: μ.d x / π.d x ≠0 left✝: μ.d x ≠0 hpr: π.d x =0
inl.intro
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c frac_neq_zero: μ.d x / π.d x ≠0 hp: μ.d x ≠0∧ π.d x =0
inl
False
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝¹: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h✝: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c frac_neq_zero: μ.d x / π.d x ≠0 h: μ.d x ≠0∧ π.d x =0∨ μ.d x =⊤∧ π.d x ≠⊤
False
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝¹: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h✝: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c frac_neq_zero: μ.d x / π.d x ≠0 h: μ.d x ≠0∧ π.d x =0∨ μ.d x =⊤∧ π.d x ≠⊤
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c frac_neq_zero: μ.d x / π.d x ≠0 hq: μ.d x =⊤∧ π.d x ≠⊤
inr
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c frac_neq_zero: μ.d x / π.d x ≠0 hq: μ.d x =⊤∧ π.d x ≠⊤
inr
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c frac_neq_zero: μ.d x / π.d x ≠0 hql: μ.d x =⊤ right✝: π.d x ≠⊤
inr.intro
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c frac_neq_zero: μ.d x / π.d x ≠0 hq: μ.d x =⊤∧ π.d x ≠⊤
inr
False
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω
h.mp
x ∈ {x | log (μ.d x / π.d x) = c} → x ∈ {x | μ.d x = ENNReal.ofReal (Real.exp c) * π.d x}
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: log (μ.d x / π.d x) = c frac_neq_zero: μ.d x / π.d x ≠0 frac_finite: μ.d x / π.d x ≠⊤ cancel_log_exp: μ.d x / π.d x = ENNReal.ofReal (Real.exp c)
h.mp
x ∈ {x | μ.d x = ENNReal.ofReal (Real.exp c) * π.d x}
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω
h.mp
x ∈ {x | log (μ.d x / π.d x) = c} → x ∈ {x | μ.d x = ENNReal.ofReal (Real.exp c) * π.d x}
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c
{x | log (μ.d x / π.d x) = c} = {x | μ.d x = ENNReal.ofReal (Real.exp c) * π.d x}
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: μ.d x = ENNReal.ofReal (Real.exp c) * π.d x
h.mpr
x ∈ {x | log (μ.d x / π.d x) = c}
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c
{x | log (μ.d x / π.d x) = c} = {x | μ.d x = ENNReal.ofReal (Real.exp c) * π.d x}
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: μ.d x = ENNReal.ofReal (Real.exp c) * π.d x log_exp: log (ENNReal.ofReal (Real.exp c)) = c
h.mpr
x ∈ {x | log (μ.d x / π.d x) = c}
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c
{x | log (μ.d x / π.d x) = c} = {x | μ.d x = ENNReal.ofReal (Real.exp c) * π.d x}
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: μ.d x = ENNReal.ofReal (Real.exp c) * π.d x log_exp: log (ENNReal.ofReal (Real.exp c)) = c
h.mpr
x ∈ {x | log (μ.d x / π.d x) = c}
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: μ.d x = ENNReal.ofReal (Real.exp c) * π.d x log_exp: log (ENNReal.ofReal (Real.exp c)) = c
μ.d x / π.d x = ENNReal.ofReal (Real.exp c)
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: μ.d x = ENNReal.ofReal (Real.exp c) * π.d x log_exp: log (ENNReal.ofReal (Real.exp c)) = c
μ.d x / π.d x = ENNReal.ofReal (Real.exp c)
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: μ.d x = ENNReal.ofReal (Real.exp c) * π.d x log_exp: log (ENNReal.ofReal (Real.exp c)) = c
μ.d x / π.d x = ENNReal.ofReal (Real.exp c)
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: μ.d x = ENNReal.ofReal (Real.exp c) * π.d x log_exp: log (ENNReal.ofReal (Real.exp c)) = c
μ.d x / π.d x = ENNReal.ofReal (Real.exp c)
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: μ.d x = ENNReal.ofReal (Real.exp c) * π.d x log_exp: log (ENNReal.ofReal (Real.exp c)) = c
ENNReal.ofReal (Real.exp c) * π.d x / π.d x = ENNReal.ofReal (Real.exp c)
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: μ.d x = ENNReal.ofReal (Real.exp c) * π.d x log_exp: log (ENNReal.ofReal (Real.exp c)) = c
μ.d x / π.d x = ENNReal.ofReal (Real.exp c)
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: μ.d x = ENNReal.ofReal (Real.exp c) * π.d x log_exp: log (ENNReal.ofReal (Real.exp c)) = c
ENNReal.ofReal (Real.exp c) * (π.d x / π.d x) = ENNReal.ofReal (Real.exp c)
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: μ.d x = ENNReal.ofReal (Real.exp c) * π.d x log_exp: log (ENNReal.ofReal (Real.exp c)) = c
ENNReal.ofReal (Real.exp c) * (π.d x / π.d x) = ENNReal.ofReal (Real.exp c)
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: μ.d x = ENNReal.ofReal (Real.exp c) * π.d x log_exp: log (ENNReal.ofReal (Real.exp c)) = c
μ.d x / π.d x = ENNReal.ofReal (Real.exp c)
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: μ.d x = ENNReal.ofReal (Real.exp c) * π.d x log_exp: log (ENNReal.ofReal (Real.exp c)) = c
ENNReal.ofReal (Real.exp c) * (π.d x / π.d x) = ENNReal.ofReal (Real.exp c)
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: μ.d x = ENNReal.ofReal (Real.exp c) * π.d x log_exp: log (ENNReal.ofReal (Real.exp c)) = c
ENNReal.ofReal (Real.exp c) * (π.d x / π.d x) = ENNReal.ofReal (Real.exp c)
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: μ.d x = ENNReal.ofReal (Real.exp c) * π.d x log_exp: log (ENNReal.ofReal (Real.exp c)) = c
ENNReal.ofReal (Real.exp c) * (π.d x * (π.d x)⁻¹) = ENNReal.ofReal (Real.exp c)
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: μ.d x = ENNReal.ofReal (Real.exp c) * π.d x log_exp: log (ENNReal.ofReal (Real.exp c)) = c
μ.d x / π.d x = ENNReal.ofReal (Real.exp c)
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: μ.d x = ENNReal.ofReal (Real.exp c) * π.d x log_exp: log (ENNReal.ofReal (Real.exp c)) = c
ENNReal.ofReal (Real.exp c) * (π.d x * (π.d x)⁻¹) = ENNReal.ofReal (Real.exp c)
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: μ.d x = ENNReal.ofReal (Real.exp c) * π.d x log_exp: log (ENNReal.ofReal (Real.exp c)) = c
ENNReal.ofReal (Real.exp c) *1= ENNReal.ofReal (Real.exp c)
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: μ.d x = ENNReal.ofReal (Real.exp c) * π.d x log_exp: log (ENNReal.ofReal (Real.exp c)) = c
ENNReal.ofReal (Real.exp c) *1= ENNReal.ofReal (Real.exp c)
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: μ.d x = ENNReal.ofReal (Real.exp c) * π.d x log_exp: log (ENNReal.ofReal (Real.exp c)) = c
μ.d x / π.d x = ENNReal.ofReal (Real.exp c)
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c
{x | log (μ.d x / π.d x) = c} = {x | μ.d x = ENNReal.ofReal (Real.exp c) * π.d x}
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: μ.d x = ENNReal.ofReal (Real.exp c) * π.d x log_exp: log (ENNReal.ofReal (Real.exp c)) = c frac: μ.d x / π.d x = ENNReal.ofReal (Real.exp c)
h.mpr
x ∈ {x | log (μ.d x / π.d x) = c}
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: μ.d x = ENNReal.ofReal (Real.exp c) * π.d x log_exp: log (μ.d x / π.d x) = c frac: μ.d x / π.d x = ENNReal.ofReal (Real.exp c)
h.mpr
x ∈ {x | log (μ.d x / π.d x) = c}
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c x: ↑Ω hx: μ.d x = ENNReal.ofReal (Real.exp c) * π.d x log_exp: log (μ.d x / π.d x) = c frac: μ.d x / π.d x = ENNReal.ofReal (Real.exp c)
h.mpr
x ∈ {x | log (μ.d x / π.d x) = c}
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c
∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c eq_sets: {x | log (μ.d x / π.d x) = c} = {x | μ.d x = ENNReal.ofReal (Real.exp c) * π.d x} null_measure: μ.toMeasure {x | log (μ.d x / π.d x) = c}ᶜ=0
∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c
∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c eq_sets: {x | log (μ.d x / π.d x) = c} = {x | μ.d x = ENNReal.ofReal (Real.exp c) * π.d x} null_measure: μ.toMeasure {x | log (μ.d x / π.d x) = c}ᶜ=0
∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c eq_sets: {x | log (μ.d x / π.d x) = c} = {x | μ.d x = ENNReal.ofReal (Real.exp c) * π.d x} null_measure: μ.toMeasure {x | μ.d x = ENNReal.ofReal (Real.exp c) * π.d x}ᶜ=0
∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c eq_sets: {x | log (μ.d x / π.d x) = c} = {x | μ.d x = ENNReal.ofReal (Real.exp c) * π.d x} null_measure: μ.toMeasure {x | μ.d x = ENNReal.ofReal (Real.exp c) * π.d x}ᶜ=0
∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
Goals accomplished!🐙
-- We show by contradiction that ENNReal.ofReal (Real.exp c) = 1. If it is ≠ 1, this implies a contradiction as μ.d x = ENNReal.ofReal (Real.exp c) * π.d x and ∫⁻ x, μ.d x ∂ν = 1.
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x
mpr.intro
μ.toMeasure = π.toMeasure
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x
ENNReal.ofReal (Real.exp c) =1
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x
ENNReal.ofReal (Real.exp c) =1
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x
ENNReal.ofReal (Real.exp c) =1
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ¬ENNReal.ofReal (Real.exp c) =1
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ¬ENNReal.ofReal (Real.exp c) =1
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x
ENNReal.ofReal (Real.exp c) =1
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x
ENNReal.ofReal (Real.exp c) =1
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1
False
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1
False
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x
ENNReal.ofReal (Real.exp c) =1
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1 univ_eq_one_μ: ∫⁻ (x : ↑Ω), 1∂μ.toMeasure =1
False
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1 univ_eq_one_μ: ∫⁻ (x : ↑Ω), 1∂μ.toMeasure =1
False
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x
ENNReal.ofReal (Real.exp c) =1
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1 univ_eq_one_μ: ∫⁻ (x : ↑Ω), 1∂μ.toMeasure =1 univ_eq_one_π: ∫⁻ (x : ↑Ω), 1∂π.toMeasure =1
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1 univ_eq_one_μ: ∫⁻ (x : ↑Ω), μ.d x *1=1 univ_eq_one_π: ∫⁻ (x : ↑Ω), 1∂π.toMeasure =1
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1 univ_eq_one_μ: ∫⁻ (x : ↑Ω), μ.d x *1=1 univ_eq_one_π: ∫⁻ (x : ↑Ω), 1∂π.toMeasure =1
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1 univ_eq_one_μ: ∫⁻ (x : ↑Ω), μ.d x *1=1 univ_eq_one_π: ∫⁻ (x : ↑Ω), 1∂π.toMeasure =1
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x
ENNReal.ofReal (Real.exp c) =1
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1 univ_eq_one_μ: ∫⁻ (x : ↑Ω), μ.d x *1=1 univ_eq_one_π: ∫⁻ (x : ↑Ω), 1∂π.toMeasure =1 ae_μ_to_ae_volume: μ.d =ᶠ[ae volume] fun x => ENNReal.ofReal (Real.exp c) * π.d x
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x
ENNReal.ofReal (Real.exp c) =1
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1 univ_eq_one_μ: ∫⁻ (x : ↑Ω), μ.d x *1=1 univ_eq_one_π: ∫⁻ (x : ↑Ω), 1∂π.toMeasure =1 ae_μ_to_ae_volume: μ.d =ᶠ[ae volume] fun x => ENNReal.ofReal (Real.exp c) * π.d x
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1 univ_eq_one_π: ∫⁻ (x : ↑Ω), 1∂π.toMeasure =1 ae_μ_to_ae_volume: μ.d =ᶠ[ae volume] fun x => ENNReal.ofReal (Real.exp c) * π.d x univ_eq_one_μ: ∫⁻ (x : ↑Ω), μ.d x =1
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1 univ_eq_one_π: ∫⁻ (x : ↑Ω), 1∂π.toMeasure =1 ae_μ_to_ae_volume: μ.d =ᶠ[ae volume] fun x => ENNReal.ofReal (Real.exp c) * π.d x univ_eq_one_μ: ∫⁻ (x : ↑Ω), μ.d x =1
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x
ENNReal.ofReal (Real.exp c) =1
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1 univ_eq_one_π: ∫⁻ (x : ↑Ω), 1∂π.toMeasure =1 ae_μ_to_ae_volume: μ.d =ᶠ[ae volume] fun x => ENNReal.ofReal (Real.exp c) * π.d x univ_eq_one_μ: ∫⁻ (x : ↑Ω), μ.d x =1
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1 univ_eq_one_π: ∫⁻ (x : ↑Ω), 1∂π.toMeasure =1 ae_μ_to_ae_volume: μ.d =ᶠ[ae volume] fun x => ENNReal.ofReal (Real.exp c) * π.d x univ_eq_one_μ: ∫⁻ (a : ↑Ω), ENNReal.ofReal (Real.exp c) * π.d a =1
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1 univ_eq_one_π: ∫⁻ (x : ↑Ω), 1∂π.toMeasure =1 ae_μ_to_ae_volume: μ.d =ᶠ[ae volume] fun x => ENNReal.ofReal (Real.exp c) * π.d x univ_eq_one_μ: ∫⁻ (a : ↑Ω), ENNReal.ofReal (Real.exp c) * π.d a =1
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1 univ_eq_one_π: ∫⁻ (x : ↑Ω), 1∂π.toMeasure =1 ae_μ_to_ae_volume: μ.d =ᶠ[ae volume] fun x => ENNReal.ofReal (Real.exp c) * π.d x univ_eq_one_μ: ∫⁻ (a : ↑Ω), ENNReal.ofReal (Real.exp c) * π.d a =1
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x
ENNReal.ofReal (Real.exp c) =1
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1 univ_eq_one_π: ∫⁻ (x : ↑Ω), 1∂π.toMeasure =1 ae_μ_to_ae_volume: μ.d =ᶠ[ae volume] fun x => ENNReal.ofReal (Real.exp c) * π.d x univ_eq_one_μ: ∫⁻ (a : ↑Ω), ENNReal.ofReal (Real.exp c) * π.d a =1
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1 univ_eq_one_π: ∫⁻ (x : ↑Ω), π.d x *1=1 ae_μ_to_ae_volume: μ.d =ᶠ[ae volume] fun x => ENNReal.ofReal (Real.exp c) * π.d x univ_eq_one_μ: ∫⁻ (a : ↑Ω), ENNReal.ofReal (Real.exp c) * π.d a =1
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1 univ_eq_one_π: ∫⁻ (x : ↑Ω), π.d x *1=1 ae_μ_to_ae_volume: μ.d =ᶠ[ae volume] fun x => ENNReal.ofReal (Real.exp c) * π.d x univ_eq_one_μ: ∫⁻ (a : ↑Ω), ENNReal.ofReal (Real.exp c) * π.d a =1
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1 univ_eq_one_π: ∫⁻ (x : ↑Ω), π.d x *1=1 ae_μ_to_ae_volume: μ.d =ᶠ[ae volume] fun x => ENNReal.ofReal (Real.exp c) * π.d x univ_eq_one_μ: ∫⁻ (a : ↑Ω), ENNReal.ofReal (Real.exp c) * π.d a =1
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x
ENNReal.ofReal (Real.exp c) =1
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1 univ_eq_one_π: ∫⁻ (x : ↑Ω), π.d x *1=1 ae_μ_to_ae_volume: μ.d =ᶠ[ae volume] fun x => ENNReal.ofReal (Real.exp c) * π.d x univ_eq_one_μ: ∫⁻ (a : ↑Ω), ENNReal.ofReal (Real.exp c) * π.d a =1
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1 ae_μ_to_ae_volume: μ.d =ᶠ[ae volume] fun x => ENNReal.ofReal (Real.exp c) * π.d x univ_eq_one_μ: ∫⁻ (a : ↑Ω), ENNReal.ofReal (Real.exp c) * π.d a =1 univ_eq_one_π: ∫⁻ (x : ↑Ω), π.d x =1
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1 ae_μ_to_ae_volume: μ.d =ᶠ[ae volume] fun x => ENNReal.ofReal (Real.exp c) * π.d x univ_eq_one_μ: ∫⁻ (a : ↑Ω), ENNReal.ofReal (Real.exp c) * π.d a =1 univ_eq_one_π: ∫⁻ (x : ↑Ω), π.d x =1
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x
ENNReal.ofReal (Real.exp c) =1
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1 ae_μ_to_ae_volume: μ.d =ᶠ[ae volume] fun x => ENNReal.ofReal (Real.exp c) * π.d x univ_eq_one_μ: ∫⁻ (a : ↑Ω), ENNReal.ofReal (Real.exp c) * π.d a =1 univ_eq_one_π: ∫⁻ (x : ↑Ω), π.d x =1
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1 ae_μ_to_ae_volume: μ.d =ᶠ[ae volume] fun x => ENNReal.ofReal (Real.exp c) * π.d x univ_eq_one_μ: ENNReal.ofReal (Real.exp c) *∫⁻ (a : ↑Ω), π.d a =1 univ_eq_one_π: ∫⁻ (x : ↑Ω), π.d x =1
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1 ae_μ_to_ae_volume: μ.d =ᶠ[ae volume] fun x => ENNReal.ofReal (Real.exp c) * π.d x univ_eq_one_μ: ∫⁻ (a : ↑Ω), ENNReal.ofReal (Real.exp c) * π.d a =1 univ_eq_one_π: ∫⁻ (x : ↑Ω), π.d x =1
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1 ae_μ_to_ae_volume: μ.d =ᶠ[ae volume] fun x => ENNReal.ofReal (Real.exp c) * π.d x univ_eq_one_μ: ENNReal.ofReal (Real.exp c) *1=1 univ_eq_one_π: ∫⁻ (x : ↑Ω), π.d x =1
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1 ae_μ_to_ae_volume: μ.d =ᶠ[ae volume] fun x => ENNReal.ofReal (Real.exp c) * π.d x univ_eq_one_μ: ∫⁻ (a : ↑Ω), ENNReal.ofReal (Real.exp c) * π.d a =1 univ_eq_one_π: ∫⁻ (x : ↑Ω), π.d x =1
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1 ae_μ_to_ae_volume: μ.d =ᶠ[ae volume] fun x => ENNReal.ofReal (Real.exp c) * π.d x univ_eq_one_μ: ENNReal.ofReal (Real.exp c) =1 univ_eq_one_π: ∫⁻ (x : ↑Ω), π.d x =1
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1 ae_μ_to_ae_volume: μ.d =ᶠ[ae volume] fun x => ENNReal.ofReal (Real.exp c) * π.d x univ_eq_one_μ: ENNReal.ofReal (Real.exp c) =1 univ_eq_one_π: ∫⁻ (x : ↑Ω), π.d x =1
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x hc: ENNReal.ofReal (Real.exp c) ≠1 ae_μ_to_ae_volume: μ.d =ᶠ[ae volume] fun x => ENNReal.ofReal (Real.exp c) * π.d x univ_eq_one_μ: ENNReal.ofReal (Real.exp c) =1 univ_eq_one_π: ∫⁻ (x : ↑Ω), π.d x =1
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x
ENNReal.ofReal (Real.exp c) =1
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
Goals accomplished!🐙
-- We rewrite μ = π as ∀s, ∫⁻ x in s, μ.d ∂ν = ∀s, ∫⁻ x in s, π.d ∂ν and use μ.d = 1 * π.d.
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x exp_c_eq_one: ENNReal.ofReal (Real.exp c) =1
mpr.intro
μ.toMeasure = π.toMeasure
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c exp_c_eq_one: ENNReal.ofReal (Real.exp c) =1 dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x =1* π.d x
mpr.intro
μ.toMeasure = π.toMeasure
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = ENNReal.ofReal (Real.exp c) * π.d x exp_c_eq_one: ENNReal.ofReal (Real.exp c) =1
mpr.intro
μ.toMeasure = π.toMeasure
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c exp_c_eq_one: ENNReal.ofReal (Real.exp c) =1 dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = π.d x
mpr.intro
μ.toMeasure = π.toMeasure
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ, dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD h✝: ∫ (x : ↑Ω) in Set.univ,
∫ (x' : ↑Ω) in Set.univ,
∑ i ∈ Set.univ.toFinset, d_ln_π_μ i x * k H₀ x x' * d_ln_π_μ i x' ∂μ.toMeasure ∂μ.toMeasure =0 d_ln_π_μ_eq_0: ∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, d_ln_π_μ i x =0 c: ℝ h: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, log (μ.d x / π.d x) = c exp_c_eq_one: ENNReal.ofReal (Real.exp c) =1 dμ_propor: ∀ᵐ (x : ↑Ω) ∂μ.toMeasure, μ.d x = π.d x
mpr.intro
μ.toMeasure = π.toMeasure
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD
include hd_ln_π_μ dπ' hπ' hstein hdμ hdπ h_kernel_positive
/-- π is the only fixed point of Φₜ(μ). We proved that by showing that, if μ = π, ϕ^* = 0 and ϕ^* ≠ 0 otherwise.-/
Goals accomplished!
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD
μ.toMeasure = π.toMeasure ↔∀ (i : Fin d), ϕ i =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0
μ.toMeasure = π.toMeasure ↔∀ (i : Fin d), ϕ i =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD
μ.toMeasure = π.toMeasure ↔∀ (i : Fin d), ϕ i =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0
mp
μ.toMeasure = π.toMeasure →∀ (i : Fin d), ϕ i =0
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0
(∀ (i : Fin d), ϕ i =0) → μ.toMeasure = π.toMeasure
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD
μ.toMeasure = π.toMeasure ↔∀ (i : Fin d), ϕ i =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0
mp
μ.toMeasure = π.toMeasure →∀ (i : Fin d), ϕ i =0
-- μ = π → ϕ^* = 0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0 μ_eq_π: μ.toMeasure = π.toMeasure
mp
∀ (i : Fin d), ϕ i =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0
mp
μ.toMeasure = π.toMeasure →∀ (i : Fin d), ϕ i =0
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0
(∀ (i : Fin d), ϕ i =0) → μ.toMeasure = π.toMeasure
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0 μ_eq_π: μ.toMeasure = π.toMeasure
mp
∀ (i : Fin d), ϕ i =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔‖ϕ‖^2=0 μ_eq_π: μ.toMeasure = π.toMeasure
mp
∀ (i : Fin d), ϕ i =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0 μ_eq_π: μ.toMeasure = π.toMeasure
mp
∀ (i : Fin d), ϕ i =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔‖ϕ‖=0 μ_eq_π: μ.toMeasure = π.toMeasure
mp
∀ (i : Fin d), ϕ i =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔‖ϕ‖=0 μ_eq_π: μ.toMeasure = π.toMeasure
mp
∀ (i : Fin d), ϕ i =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔‖ϕ‖=0 μ_eq_π: μ.toMeasure = π.toMeasure
mp
∀ (i : Fin d), ϕ i =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0
mp
μ.toMeasure = π.toMeasure →∀ (i : Fin d), ϕ i =0
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0
(∀ (i : Fin d), ϕ i =0) → μ.toMeasure = π.toMeasure
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔‖ϕ‖=0 μ_eq_π: μ.toMeasure = π.toMeasure
mp
∀ (i : Fin d), ϕ i =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔‖ϕ‖=0 μ_eq_π: μ.toMeasure = π.toMeasure
mp
‖ϕ‖=0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔‖ϕ‖=0 μ_eq_π: μ.toMeasure = π.toMeasure
mp
‖ϕ‖=0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0
mp
μ.toMeasure = π.toMeasure →∀ (i : Fin d), ϕ i =0
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0
(∀ (i : Fin d), ϕ i =0) → μ.toMeasure = π.toMeasure
Goals accomplished!
Goals accomplished!🐙
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0
mpr
(∀ (i : Fin d), ϕ i =0) → μ.toMeasure = π.toMeasure
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD
μ.toMeasure = π.toMeasure ↔∀ (i : Fin d), ϕ i =0
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0
mpr
(∀ (i : Fin d), ϕ i =0) → μ.toMeasure = π.toMeasure
-- ϕ^* ≠ 0 → μ = π
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0 phi_norm: ∀ (i : Fin d), ϕ i =0
mpr
μ.toMeasure = π.toMeasure
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0
mpr
(∀ (i : Fin d), ϕ i =0) → μ.toMeasure = π.toMeasure
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0 phi_norm: ∀ (i : Fin d), ϕ i =0
mpr
μ.toMeasure = π.toMeasure
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0 phi_norm: ‖ϕ‖=0
mpr
μ.toMeasure = π.toMeasure
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0 phi_norm: ‖ϕ‖=0
mpr
μ.toMeasure = π.toMeasure
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0 phi_norm: ‖ϕ‖=0
mpr
μ.toMeasure = π.toMeasure
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0
mpr
(∀ (i : Fin d), ϕ i =0) → μ.toMeasure = π.toMeasure
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0 phi_norm: ‖ϕ‖=0 h: ¬μ.toMeasure = π.toMeasure
mpr
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0 phi_norm: ‖ϕ‖=0 h: ¬μ.toMeasure = π.toMeasure
mpr
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0
mpr
(∀ (i : Fin d), ϕ i =0) → μ.toMeasure = π.toMeasure
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0 phi_norm: ‖ϕ‖=0 h: ¬μ.toMeasure = π.toMeasure
mpr
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0 phi_norm: ‖ϕ‖^2=0 h: ¬μ.toMeasure = π.toMeasure
mpr
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0 phi_norm: ‖ϕ‖=0 h: ¬μ.toMeasure = π.toMeasure
mpr
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0 phi_norm: KSD μ.toMeasure π.toMeasure =0 h: ¬μ.toMeasure = π.toMeasure
mpr
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0 phi_norm: KSD μ.toMeasure π.toMeasure =0 h: ¬μ.toMeasure = π.toMeasure
mpr
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0 phi_norm: KSD μ.toMeasure π.toMeasure =0 h: ¬μ.toMeasure = π.toMeasure
mpr
False
Goals accomplished!
d: ℕ hd: d ≠0 Ω: Set (Vector ℝ d) inst✝⁵: MeasureSpace ↑Ω μ, π: DensityMeasure ↑Ω inst✝⁴: IsProbabilityMeasure μ.toMeasure inst✝³: IsProbabilityMeasure π.toMeasure hdμ: ∀ (x : ↑Ω), μ.d x ≠0 hdπ: ∀ (x : ↑Ω), π.d x ≠0 H₀: Set (↑Ω→ ℝ) inst✝²: NormedAddCommGroup ↑H₀ inst✝¹: InnerProductSpace ℝ ↑H₀ s: RKHS H₀ h_kernel_positive: positive_definite_kernel μ H₀ d_ln_π: Fin d →↑Ω→ ℝ ϕ: product_RKHS H₀ hd dϕ, d_ln_π_μ: Fin d →↑Ω→ ℝ hd_ln_π_μ: ∀ (ν : Measure ↑Ω), (∀ (i : Fin d), ∀ᵐ (x : ↑Ω) ∂ν, d_ln_π_μ i x =0) →∃ c, ∀ᵐ (x : ↑Ω) ∂ν, log (μ.d x / π.d x) = c dπ': Fin d →↑Ω→ ℝ hπ': ∀ (x : ↑Ω) (i : Fin d), (π.d x).toReal * d_ln_π i x = dπ' i x inst✝: Norm ↑Ω KSD: Measure ↑Ω→ Measure ↑Ω→ ℝ hstein: SteinClass π fun i =>↑(ϕ i) hksd: is_ksd hd μ π H₀ d_ln_π ϕ dϕ d_ln_π_μ KSD ksd_norm: is_ksd_norm hd μ π H₀ ϕ KSD KSD_discrepancy: μ.toMeasure = π.toMeasure ↔ KSD μ.toMeasure π.toMeasure =0
mpr
(∀ (i : Fin d), ϕ i =0) → μ.toMeasure = π.toMeasure