Skip to content

Commit

Permalink
Another formalization of lemma
Browse files Browse the repository at this point in the history
  • Loading branch information
teorth committed Dec 10, 2023
1 parent e87fcee commit 6234e18
Show file tree
Hide file tree
Showing 2 changed files with 17 additions and 2 deletions.
17 changes: 16 additions & 1 deletion PFR/ImprovedPFR.lean
Original file line number Diff line number Diff line change
Expand Up @@ -97,12 +97,27 @@ $$ \leq I(U : V \, | \, S) + I(V : W \, | \,S) + I(W : U \, | \, S) + \frac{\eta
proof_wanted averaged_construct_good : 0 = 1


section GeneralInequality

variable {Ω₀ : Type*} [MeasureSpace Ω₀] [IsProbabilityMeasure (ℙ : Measure Ω₀₁)]

variable (Y : Ω₀ → G) (hY : Measurable Y)

variable (Z₁ Z₂ Z₃ Z₄ : Ω → G)
(hZ₁ : Measurable Z₁) (hZ₂ : Measurable Z₂) (hZ₃ : Measurable Z₃) (hZ₄ : Measurable Z₄)

variable (h_indep : iIndepFun (fun _i => hG) ![Z₁, Z₂, Z₃, Z₄])

local notation3 "Sum" => Z₁ + Z₂ + Z₃ + Z₄

/-- Let $X_1, X_2, X_3, X_4$ be independent $G$-valued random variables, and let $Y$ be another $G$-valued random variable. Set $S := X_1+X_2+X_3+X_4$. Then $d[Y; X_1+X_2|X_1 + X_3, S] - d[Y; X_1]$ is at most
$$ \tfrac{1}{4} (2d[X_1;X_2] + d[X_1;X_3] + d[X_3;X_4])$$
$$+ \tfrac{1}{8} (\bbH[X_1+X_2] - \bbH[X_1+X_3] + \bbH[X_2] - \bbH[X_3]$$
$$ + \bbH[X_2|X_2+X_4] - \bbH[X_1|X_1+X_3]).$$
-/
proof_wanted gen_ineq: 0 = 1
lemma gen_ineq: d[Y # Z₁ + Z₂ | ⟨ Z₁ + Z₃, Sum ⟩ ] - d[Y # Z₁ ] ≤ (2 * d[Z₁ # Z₂] + d[Z₁ # Z₃] + d[Z₃ # Z₄])/4 + (H[Z₁+Z₂] - H[Z₁+Z₃] + H[Z₂] - H[Z₃]) + H[Z₂|Z₂+Z₄] - H[Z₁|Z₁+Z₃] / 8 := sorry

end GeneralInequality

/-- The quantity
$$ \sum_{i=1}^2 \sum_{A,B \in \{U,V,W\}: A \neq B} d[X_i^0;A|B, S] - d[X_i^0;X_i]$$
Expand Down
2 changes: 1 addition & 1 deletion blueprint/src/chapter/improved_exponent.tex
Original file line number Diff line number Diff line change
Expand Up @@ -82,7 +82,7 @@ \chapter{Improving the exponents}

To control the expressions in the right-hand side of this lemma we need a general entropy inequality.

\begin{lemma}[General inequality]\label{gen-ineq}\lean{gen_ineq} Let $X_1, X_2, X_3, X_4$ be independent $G$-valued random variables, and let $Y$ be another $G$-valued random variable. Set $S := X_1+X_2+X_3+X_4$. Then
\begin{lemma}[General inequality]\label{gen-ineq}\lean{gen_ineq} \leanok Let $X_1, X_2, X_3, X_4$ be independent $G$-valued random variables, and let $Y$ be another $G$-valued random variable. Set $S := X_1+X_2+X_3+X_4$. Then
\begin{align*}
& d[Y; X_1+X_2|X_1 + X_3, S] - d[Y; X_1] \\
&\quad \leq \tfrac{1}{4} (2d[X_1;X_2] + d[X_1;X_3] + d[X_3;X_4])\\
Expand Down

0 comments on commit 6234e18

Please sign in to comment.