From 033927c6d816cb989e003d8d40c935c8d4fccd78 Mon Sep 17 00:00:00 2001 From: jtristan Date: Thu, 25 Apr 2024 08:48:30 -0400 Subject: [PATCH 01/20] Change in RNG --- Tests/Random.py | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/Tests/Random.py b/Tests/Random.py index adfaf5b8..cf190fc0 100644 --- a/Tests/Random.py +++ b/Tests/Random.py @@ -2,5 +2,4 @@ class Random: def UniformPowerOfTwoSample(n): - return secrets.randbelow(n) - #return secrets.randbits(n.bit_length()-1) + return secrets.randbits(n.bit_length()-1) From 66f459a102e3ed2cf6536bf9a8272a08d6a61677 Mon Sep 17 00:00:00 2001 From: jtristan Date: Thu, 25 Apr 2024 18:15:37 -0400 Subject: [PATCH 02/20] Exploring query verification --- SampCert/DiffPrivacy/Sensitivity.lean | 155 ++++++++++++++++++++++++++ 1 file changed, 155 insertions(+) create mode 100644 SampCert/DiffPrivacy/Sensitivity.lean diff --git a/SampCert/DiffPrivacy/Sensitivity.lean b/SampCert/DiffPrivacy/Sensitivity.lean new file mode 100644 index 00000000..06103902 --- /dev/null +++ b/SampCert/DiffPrivacy/Sensitivity.lean @@ -0,0 +1,155 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jean-Baptiste Tristan +-/ + +-- Data : sorted list of integers +-- integer can be added, removed, modified + +import Mathlib.Topology.Algebra.InfiniteSum.Ring +import Mathlib.Algebra.Group.Basic +import SampCert.DiffPrivacy.ConcentratedBound +import SampCert.SLang +import SampCert.Samplers.GaussianGen.Basic + +open Classical Nat Int Real + +noncomputable section + +namespace SLang + +variable {T : Type} + +inductive Neighbour (l₁ l₂ : List ℕ) : Prop where + | Addition : l₁ = a ++ b → l₂ = a ++ [n] ++ b → Neighbour l₁ l₂ + | Deletion : l₁ = a ++ [n] ++ b → l₂ = a ++ b → Neighbour l₁ l₂ + | Update : l₁ = a ++ [n] ++ b → l₂ = a ++ [m] ++ b -> Neighbour l₁ l₂ + +def sensitivity (q : List ℕ → ℤ) (Δ : ℕ+) : Prop := + ∀ l₁ l₂ : List ℕ, Neighbour l₁ l₂ → Int.natAbs (q l₁ - q l₂) ≤ Δ + +def DP (q : List ℕ → SLang ℤ) (ε₁ ε₂ : ℕ+) : Prop := + ∀ α : ℝ, 1 < α → ∀ l₁ l₂ : List ℕ, Neighbour l₁ l₂ → + RenyiDivergence + (fun x : ℤ => (q l₁ x).toReal) + (fun x : ℤ => (q l₂ x).toReal) α + ≤ (1/2) * (ε₁ / ε₂)^2 * α + +def NoisedQuery (query : List ℕ → ℤ) (Δ : ℕ+) (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℤ := do + DiscreteGaussianGenSample (Δ * ε₂) ε₁ (query l) + +theorem NoisedQueryDP (query : List ℕ → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_sensitivity : sensitivity query Δ) : + DP (NoisedQuery query Δ ε₁ ε₂) ε₁ ε₂ := by + simp [DP, NoisedQuery] + intros α h1 l₁ l₂ h2 + have A := @DiscreteGaussianGenSampleZeroConcentrated α h1 (Δ * ε₂) ε₁ (query l₁) (query l₂) + apply le_trans A + clear A + replace bounded_sensitivity := bounded_sensitivity l₁ l₂ h2 + ring_nf + simp + conv => + left + left + right + rw [mul_pow] + conv => + left + rw [mul_assoc] + right + rw [mul_comm] + rw [← mul_assoc] + conv => + left + rw [mul_assoc] + right + rw [← mul_assoc] + left + rw [mul_comm] + rw [← mul_assoc] + rw [← mul_assoc] + rw [← mul_assoc] + simp only [inv_pow] + rw [mul_inv_le_iff'] + . have A : (α * ↑↑ε₁ ^ 2 * (↑↑ε₂ ^ 2)⁻¹) ≤ (α * ↑↑ε₁ ^ 2 * (↑↑ε₂ ^ 2)⁻¹) := le_refl (α * ↑↑ε₁ ^ 2 * (↑↑ε₂ ^ 2)⁻¹) + have B : 0 ≤ (α * ↑↑ε₁ ^ 2 * (↑↑ε₂ ^ 2)⁻¹) := by + simp + apply @le_trans ℝ _ 0 1 α (instStrictOrderedCommRingReal.proof_3) (le_of_lt h1) + apply mul_le_mul A _ _ B + . apply sq_le_sq.mpr + simp only [abs_cast] + rw [← Int.cast_sub] + rw [← Int.cast_abs] + apply Int.cast_le.mpr + rw [← Int.natCast_natAbs] + apply Int.ofNat_le.mpr + trivial + . apply sq_nonneg + . rw [pow_two] + rw [mul_pos_iff] + left + simp + +def CountingQuery (l : List ℕ) : ℤ := List.length l + +theorem CountingQuery1Sensitive : + sensitivity CountingQuery 1 := by + simp [CountingQuery, sensitivity] + intros l₁ l₂ H + cases H + . rename_i a b n h1 h2 + subst h1 h2 + simp + . rename_i a b n h1 h2 + subst h1 h2 + simp + . rename_i a n b m h1 h2 + subst h1 h2 + simp + +example (ε₁ ε₂ : ℕ+) : DP (NoisedQuery CountingQuery 1 ε₁ ε₂) ε₁ ε₂ := by + apply NoisedQueryDP + apply CountingQuery1Sensitive + +example (clip : ℕ+) : sensitivity (fun l : List ℕ => List.sum (List.map (fun n : ℕ => if n > clip then clip else n) l)) clip := by + simp [sensitivity] + intros l₁ l₂ H + cases H + . rename_i a b n h1 h2 + subst h1 h2 + simp + split + . rename_i h3 + simp + . rename_i h3 + simp at h3 + simpa + . rename_i a b n h1 h2 + subst h1 h2 + simp + split + . rename_i h3 + simp + . rename_i h3 + simp at h3 + simpa + . rename_i a n b m h1 h2 + subst h1 h2 + simp + split + . split + . rename_i h3 h4 + simp + . rename_i h3 h4 + simp at h4 + sorry -- OK + . split + . rename_i h3 h4 + simp at h3 + sorry -- OK + . rename_i h3 h4 + simp at h3 h4 + sorry -- OK + +end SLang From 149729665e1dca4bece759cc772aa8672285b7c2 Mon Sep 17 00:00:00 2001 From: jtristan Date: Tue, 30 Apr 2024 10:02:24 -0400 Subject: [PATCH 03/20] Exploring avg query --- SampCert.lean | 1 + SampCert/DiffPrivacy/BoundedAvg.lean | 29 +++ SampCert/DiffPrivacy/BoundedSum.lean | 202 ++++++++++++++++++++ SampCert/DiffPrivacy/ConcentratedBound.lean | 4 +- SampCert/DiffPrivacy/Count.lean | 41 ++++ SampCert/DiffPrivacy/DP.lean | 95 +++++++++ SampCert/DiffPrivacy/Neighbours.lean | 12 ++ SampCert/DiffPrivacy/RenyiDivergence.lean | 13 ++ SampCert/DiffPrivacy/Sensitivity.lean | 141 +------------- 9 files changed, 397 insertions(+), 141 deletions(-) create mode 100644 SampCert/DiffPrivacy/BoundedAvg.lean create mode 100644 SampCert/DiffPrivacy/BoundedSum.lean create mode 100644 SampCert/DiffPrivacy/Count.lean create mode 100644 SampCert/DiffPrivacy/DP.lean create mode 100644 SampCert/DiffPrivacy/Neighbours.lean create mode 100644 SampCert/DiffPrivacy/RenyiDivergence.lean diff --git a/SampCert.lean b/SampCert.lean index e0074f4d..f2ca9379 100644 --- a/SampCert.lean +++ b/SampCert.lean @@ -7,3 +7,4 @@ Authors: Jean-Baptiste Tristan import SampCert.Samplers.Gaussian.Basic import SampCert.DiffPrivacy.ConcentratedBound import SampCert.Samplers.GaussianGen.Basic +import SampCert.DiffPrivacy.BoundedAvg diff --git a/SampCert/DiffPrivacy/BoundedAvg.lean b/SampCert/DiffPrivacy/BoundedAvg.lean new file mode 100644 index 00000000..36ac1d30 --- /dev/null +++ b/SampCert/DiffPrivacy/BoundedAvg.lean @@ -0,0 +1,29 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jean-Baptiste Tristan +-/ + +import SampCert.DiffPrivacy.DP +import SampCert.DiffPrivacy.Count +import SampCert.DiffPrivacy.BoundedSum + +open Classical Nat Int Real + +noncomputable section + +namespace SLang + +def NoisedBoundedAvgQuery (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) (l : List ℤ) : SLang ℤ := do + let S ← NoisedBoundedSumQuery L U h ε₁ (2 * ε₂) l + let C ← NoisedCountingQuery ε₁ (2 * ε₂) l + if C ≤ 1 then return (L + U) / 2 + else return S / C + +theorem BoundedSumQueryDP (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) : DP (NoisedBoundedAvgQuery L U h ε₁ ε₂) ε₁ ε₂ := by + have A := @NoisedCountingQueryDP ℤ ε₁ (2 * ε₂) + have B := @NoisedBoundedSumQueryDP L U h ε₁ (2 * ε₂) + sorry + + +end SLang diff --git a/SampCert/DiffPrivacy/BoundedSum.lean b/SampCert/DiffPrivacy/BoundedSum.lean new file mode 100644 index 00000000..d84e1ae3 --- /dev/null +++ b/SampCert/DiffPrivacy/BoundedSum.lean @@ -0,0 +1,202 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jean-Baptiste Tristan +-/ + +import SampCert.DiffPrivacy.DP +import Mathlib.Algebra.Group.Defs +import Mathlib.Init.Algebra.Classes +import Init.Data.Int.Order + +open Classical Nat Int Real + +noncomputable section + +namespace SLang + +-- N+ and not Z because of sensitivity measure +-- Can be generalized +-- variable {T : Type} [Add T] [LinearOrder T] [coe: Coe T ℤ] +-- variable {h : Monotone coe.coe} + +-- instance : OrderHom T ℤ where +-- toFun := coe.coe +-- monotone' := h + +-- def BoundedSumQuery (clip₁ clip₂ : ℤ) (l : List ℤ) : ℤ := +-- List.sum (List.map (fun n : ℤ => if n < clip₁ then clip₁ else if clip₂ < n then clip₂ else n) l) + +def BoundedSumQuery (L U : ℤ) (l : List ℤ) : ℤ := + List.sum (List.map (fun n : ℤ => max (min n U) L) l) + +-- theorem natAbs1 {n : ℕ} {x : ℤ} (h : n ≤ x) : +-- n ≤ natAbs x := by +-- cases x +-- . rename_i m +-- simp +-- exact Int.ofNat_le.mp h +-- . rename_i m +-- simp +-- have A : -[m+1] < 0 := negSucc_lt_zero m +-- have B : n < (0 : ℤ) := by +-- apply Int.lt_of_le_of_lt h A +-- contradiction + +-- theorem natAbs2 {n : ℕ} {x : ℤ} (h : x ≤ -[n+1]) : +-- Nat.succ n ≤ natAbs x := by +-- cases x +-- . rename_i m +-- simp +-- have A : -[m+1] < 0 := negSucc_lt_zero m +-- have B : n < (0 : ℤ) := by +-- apply Int.lt_of_le_of_lt h A +-- contradiction +-- . rename_i m +-- simp +-- exact Int.ofNat_le.mp h + +def maxBoundPos (L U : ℤ) (h : L < U) : + 0 < max (Int.natAbs L) (Int.natAbs U) := by + have A : ¬ (L = 0 ∧ U = 0) := by + by_contra h' + cases h' + rename_i h1 h2 + subst h1 h2 + contradiction + simp at A + have B : L = 0 ∨ L ≠ 0 := by exact eq_or_ne L 0 + cases B + . rename_i h' + subst h' + simp at A + simp + trivial + . rename_i h' + simp + left + trivial + +def maxBound (L U : ℤ) (h : L < U) : ℕ+ := + ⟨ max (Int.natAbs L) (Int.natAbs U) , maxBoundPos L U h⟩ + +theorem BoundedSumQuerySensitivity (L U : ℤ) (h : L < U) : sensitivity (BoundedSumQuery L U) (maxBound L U h) := by + sorry + -- simp [sensitivity, BoundedSumQuery] + -- intros l₁ l₂ H + -- cases H + -- . rename_i a b n h1 h2 + -- subst h1 h2 + -- simp + -- right + -- . sorry + -- . sorry + +def NoisedBoundedSumQuery (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) (l : List ℤ) : SLang ℤ := do + NoisedQuery (BoundedSumQuery L U) (maxBound L U h) ε₁ ε₂ l + +theorem NoisedBoundedSumQueryDP (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) : DP (NoisedBoundedSumQuery L U h ε₁ ε₂) ε₁ ε₂ := by + apply NoisedQueryDP + apply BoundedSumQuerySensitivity + +-- theorem BoundedSumSensitivity (clip₁ clip₂ : ℤ) (h : clip₁ < clip₂) : sensitivity (BoundedSumQuery clip₁ clip₂) (max (Int.natAbs clip₁) (Int.natAbs clip₂)) := by +-- simp [sensitivity, BoundedSumQuery] +-- intros l₁ l₂ H +-- cases H +-- . rename_i a b n h1 h2 +-- subst h1 h2 +-- simp +-- split +-- . left +-- apply le_refl +-- . split +-- . right +-- apply le_refl +-- . rename_i h3 h4 +-- simp at * +-- cases n +-- . rename_i n +-- simp +-- right +-- simp at h4 +-- apply natAbs1 h4 +-- . rename_i n +-- simp at * +-- left +-- apply natAbs2 h3 +-- . rename_i a n b h1 h2 +-- subst h1 h2 +-- simp +-- split +-- . left +-- apply le_refl +-- . split +-- . right +-- apply le_refl +-- . rename_i h3 h4 +-- simp at * +-- cases n +-- . rename_i n +-- simp +-- right +-- simp at h4 +-- apply natAbs1 h4 +-- . rename_i n +-- simp at * +-- left +-- apply natAbs2 h3 +-- . rename_i a n b m h1 h2 +-- subst h1 h2 +-- simp +-- split +-- . split +-- . rename_i h1 h2 +-- simp at * +-- . split +-- . rename_i h1 h2 h3 +-- simp at * +-- sorry +-- . rename_i h1 h2 h3 +-- simp at * +-- sorry +-- . split +-- . split +-- . rename_i h1 h2 h3 +-- simp at * + +-- cases clip₁ +-- . cases clip₂ +-- . rename_i n' m' +-- simp at * +-- sorry +-- . rename_i n' m' +-- simp at * +-- sorry +-- . cases clip₂ +-- . rename_i n' m' +-- simp at * +-- sorry +-- . rename_i n' m' +-- simp at * +-- sorry + + +-- . split +-- . rename_i h1 h2 h3 h4 +-- simp at * +-- . rename_i h1 h2 h3 h4 +-- simp at * +-- sorry +-- . split +-- . rename_i h1 h2 h3 +-- simp at * +-- sorry +-- . split +-- . rename_i h1 h2 h3 h4 +-- simp at * +-- sorry +-- . rename_i h1 h2 h3 h4 +-- simp at * +-- sorry + +end SLang diff --git a/SampCert/DiffPrivacy/ConcentratedBound.lean b/SampCert/DiffPrivacy/ConcentratedBound.lean index db08eb83..a6091200 100644 --- a/SampCert/DiffPrivacy/ConcentratedBound.lean +++ b/SampCert/DiffPrivacy/ConcentratedBound.lean @@ -9,12 +9,10 @@ import SampCert.DiffPrivacy.GaussBound import SampCert.DiffPrivacy.GaussConvergence import SampCert.DiffPrivacy.GaussPeriodicity import SampCert.Util.Shift +import SampCert.DiffPrivacy.RenyiDivergence open Real -noncomputable def RenyiDivergence (p q : ℤ → ℝ) (α : ℝ) : ℝ := - (1 / (α - 1)) * Real.log (∑' x : ℤ, (p x)^α * (q x)^(1 - α)) - theorem sg_sum_pos' {σ : ℝ} (h : σ ≠ 0) (μ : ℝ) (α : ℝ) : 0 < ((gauss_term_ℝ σ μ) x / ∑' (x : ℤ), (gauss_term_ℝ σ μ) x)^α := by apply rpow_pos_of_pos diff --git a/SampCert/DiffPrivacy/Count.lean b/SampCert/DiffPrivacy/Count.lean new file mode 100644 index 00000000..65312c2c --- /dev/null +++ b/SampCert/DiffPrivacy/Count.lean @@ -0,0 +1,41 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jean-Baptiste Tristan +-/ + +import SampCert.DiffPrivacy.DP + +open Classical Nat Int Real + +noncomputable section + +namespace SLang + +variable {T : Type} + +def CountingQuery (l : List T) : ℤ := List.length l + +theorem CountingQuery1Sensitive : + @sensitivity T CountingQuery 1 := by + simp [CountingQuery, sensitivity] + intros l₁ l₂ H + cases H + . rename_i a b n h1 h2 + subst h1 h2 + simp + . rename_i a b n h1 h2 + subst h1 h2 + simp + . rename_i a n b m h1 h2 + subst h1 h2 + simp + +def NoisedCountingQuery (ε₁ ε₂ : ℕ+) (l : List T) : SLang ℤ := do + NoisedQuery CountingQuery 1 ε₁ ε₂ l + +theorem NoisedCountingQueryDP (ε₁ ε₂ : ℕ+) : @DP T (NoisedCountingQuery ε₁ ε₂) ε₁ ε₂ := by + apply NoisedQueryDP + apply CountingQuery1Sensitive + +end SLang diff --git a/SampCert/DiffPrivacy/DP.lean b/SampCert/DiffPrivacy/DP.lean new file mode 100644 index 00000000..743d9515 --- /dev/null +++ b/SampCert/DiffPrivacy/DP.lean @@ -0,0 +1,95 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jean-Baptiste Tristan +-/ + +import Mathlib.Topology.Algebra.InfiniteSum.Ring +import Mathlib.Algebra.Group.Basic +import SampCert.DiffPrivacy.ConcentratedBound +import SampCert.SLang +import SampCert.Samplers.GaussianGen.Basic +import SampCert.DiffPrivacy.Neighbours +import SampCert.DiffPrivacy.Sensitivity + +noncomputable section + +open Classical Nat Int Real + +def DP (q : List T → SLang ℤ) (ε₁ ε₂ : ℕ+) : Prop := + ∀ α : ℝ, 1 < α → ∀ l₁ l₂ : List T, Neighbour l₁ l₂ → + RenyiDivergence + (fun x : ℤ => (q l₁ x).toReal) + (fun x : ℤ => (q l₂ x).toReal) α + ≤ (1/2) * (ε₁ / ε₂)^2 * α + +namespace SLang + +def NoisedQuery (query : List T → ℤ) (Δ : ℕ+) (ε₁ ε₂ : ℕ+) (l : List T) : SLang ℤ := do + DiscreteGaussianGenSample (Δ * ε₂) ε₁ (query l) + +theorem NoisedQueryDP (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_sensitivity : sensitivity query Δ) : + DP (NoisedQuery query Δ ε₁ ε₂) ε₁ ε₂ := by + simp [DP, NoisedQuery] + intros α h1 l₁ l₂ h2 + have A := @DiscreteGaussianGenSampleZeroConcentrated α h1 (Δ * ε₂) ε₁ (query l₁) (query l₂) + apply le_trans A + clear A + replace bounded_sensitivity := bounded_sensitivity l₁ l₂ h2 + ring_nf + simp + conv => + left + left + right + rw [mul_pow] + conv => + left + rw [mul_assoc] + right + rw [mul_comm] + rw [← mul_assoc] + conv => + left + rw [mul_assoc] + right + rw [← mul_assoc] + left + rw [mul_comm] + rw [← mul_assoc] + rw [← mul_assoc] + rw [← mul_assoc] + simp only [inv_pow] + rw [mul_inv_le_iff'] + . have A : (α * ↑↑ε₁ ^ 2 * (↑↑ε₂ ^ 2)⁻¹) ≤ (α * ↑↑ε₁ ^ 2 * (↑↑ε₂ ^ 2)⁻¹) := le_refl (α * ↑↑ε₁ ^ 2 * (↑↑ε₂ ^ 2)⁻¹) + have B : 0 ≤ (α * ↑↑ε₁ ^ 2 * (↑↑ε₂ ^ 2)⁻¹) := by + simp + apply @le_trans ℝ _ 0 1 α (instStrictOrderedCommRingReal.proof_3) (le_of_lt h1) + apply mul_le_mul A _ _ B + . apply sq_le_sq.mpr + simp only [abs_cast] + rw [← Int.cast_sub] + rw [← Int.cast_abs] + apply Int.cast_le.mpr + rw [← Int.natCast_natAbs] + apply Int.ofNat_le.mpr + trivial + . apply sq_nonneg + . rw [pow_two] + rw [mul_pos_iff] + left + simp + +def Compose (nq1 nq2 : List T → SLang ℤ) (l : List T) : SLang (ℤ × ℤ) := do + let A ← nq1 l + let B ← nq2 l + return (A,B) + +def RAdd (ε₁ ε₂ ε₃ ε₄ : ℕ+) : ℕ+ × ℕ+ := + (ε₁ * ε₃ + ε₂ * ε₄,ε₃ * ε₄) + +-- theorem DPCompose (nq1 nq2 : List T → SLang ℤ) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h1 : DP nq1 ε₁ ε₂) (h2 : DP nq2 ε₃ ε₄) : +-- DP (Compose nq1 nq2) (RAdd ε₁ ε₂ ε₃ ε₄).1 (RAdd ε₁ ε₂ ε₃ ε₄).2 := by +-- sorry + +end SLang diff --git a/SampCert/DiffPrivacy/Neighbours.lean b/SampCert/DiffPrivacy/Neighbours.lean new file mode 100644 index 00000000..2281d028 --- /dev/null +++ b/SampCert/DiffPrivacy/Neighbours.lean @@ -0,0 +1,12 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jean-Baptiste Tristan +-/ + +variable {T : Type} + +inductive Neighbour (l₁ l₂ : List T) : Prop where + | Addition : l₁ = a ++ b → l₂ = a ++ [n] ++ b → Neighbour l₁ l₂ + | Deletion : l₁ = a ++ [n] ++ b → l₂ = a ++ b → Neighbour l₁ l₂ + | Update : l₁ = a ++ [n] ++ b → l₂ = a ++ [m] ++ b -> Neighbour l₁ l₂ diff --git a/SampCert/DiffPrivacy/RenyiDivergence.lean b/SampCert/DiffPrivacy/RenyiDivergence.lean new file mode 100644 index 00000000..0788a20f --- /dev/null +++ b/SampCert/DiffPrivacy/RenyiDivergence.lean @@ -0,0 +1,13 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jean-Baptiste Tristan +-/ + +import Mathlib.Topology.Algebra.InfiniteSum.Ring +import Mathlib.NumberTheory.ModularForms.JacobiTheta.TwoVariable + +open Real + +noncomputable def RenyiDivergence (p q : ℤ → ℝ) (α : ℝ) : ℝ := + (1 / (α - 1)) * Real.log (∑' x : ℤ, (p x)^α * (q x)^(1 - α)) diff --git a/SampCert/DiffPrivacy/Sensitivity.lean b/SampCert/DiffPrivacy/Sensitivity.lean index 06103902..831c9722 100644 --- a/SampCert/DiffPrivacy/Sensitivity.lean +++ b/SampCert/DiffPrivacy/Sensitivity.lean @@ -12,144 +12,9 @@ import Mathlib.Algebra.Group.Basic import SampCert.DiffPrivacy.ConcentratedBound import SampCert.SLang import SampCert.Samplers.GaussianGen.Basic +import SampCert.DiffPrivacy.Neighbours open Classical Nat Int Real -noncomputable section - -namespace SLang - -variable {T : Type} - -inductive Neighbour (l₁ l₂ : List ℕ) : Prop where - | Addition : l₁ = a ++ b → l₂ = a ++ [n] ++ b → Neighbour l₁ l₂ - | Deletion : l₁ = a ++ [n] ++ b → l₂ = a ++ b → Neighbour l₁ l₂ - | Update : l₁ = a ++ [n] ++ b → l₂ = a ++ [m] ++ b -> Neighbour l₁ l₂ - -def sensitivity (q : List ℕ → ℤ) (Δ : ℕ+) : Prop := - ∀ l₁ l₂ : List ℕ, Neighbour l₁ l₂ → Int.natAbs (q l₁ - q l₂) ≤ Δ - -def DP (q : List ℕ → SLang ℤ) (ε₁ ε₂ : ℕ+) : Prop := - ∀ α : ℝ, 1 < α → ∀ l₁ l₂ : List ℕ, Neighbour l₁ l₂ → - RenyiDivergence - (fun x : ℤ => (q l₁ x).toReal) - (fun x : ℤ => (q l₂ x).toReal) α - ≤ (1/2) * (ε₁ / ε₂)^2 * α - -def NoisedQuery (query : List ℕ → ℤ) (Δ : ℕ+) (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℤ := do - DiscreteGaussianGenSample (Δ * ε₂) ε₁ (query l) - -theorem NoisedQueryDP (query : List ℕ → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_sensitivity : sensitivity query Δ) : - DP (NoisedQuery query Δ ε₁ ε₂) ε₁ ε₂ := by - simp [DP, NoisedQuery] - intros α h1 l₁ l₂ h2 - have A := @DiscreteGaussianGenSampleZeroConcentrated α h1 (Δ * ε₂) ε₁ (query l₁) (query l₂) - apply le_trans A - clear A - replace bounded_sensitivity := bounded_sensitivity l₁ l₂ h2 - ring_nf - simp - conv => - left - left - right - rw [mul_pow] - conv => - left - rw [mul_assoc] - right - rw [mul_comm] - rw [← mul_assoc] - conv => - left - rw [mul_assoc] - right - rw [← mul_assoc] - left - rw [mul_comm] - rw [← mul_assoc] - rw [← mul_assoc] - rw [← mul_assoc] - simp only [inv_pow] - rw [mul_inv_le_iff'] - . have A : (α * ↑↑ε₁ ^ 2 * (↑↑ε₂ ^ 2)⁻¹) ≤ (α * ↑↑ε₁ ^ 2 * (↑↑ε₂ ^ 2)⁻¹) := le_refl (α * ↑↑ε₁ ^ 2 * (↑↑ε₂ ^ 2)⁻¹) - have B : 0 ≤ (α * ↑↑ε₁ ^ 2 * (↑↑ε₂ ^ 2)⁻¹) := by - simp - apply @le_trans ℝ _ 0 1 α (instStrictOrderedCommRingReal.proof_3) (le_of_lt h1) - apply mul_le_mul A _ _ B - . apply sq_le_sq.mpr - simp only [abs_cast] - rw [← Int.cast_sub] - rw [← Int.cast_abs] - apply Int.cast_le.mpr - rw [← Int.natCast_natAbs] - apply Int.ofNat_le.mpr - trivial - . apply sq_nonneg - . rw [pow_two] - rw [mul_pos_iff] - left - simp - -def CountingQuery (l : List ℕ) : ℤ := List.length l - -theorem CountingQuery1Sensitive : - sensitivity CountingQuery 1 := by - simp [CountingQuery, sensitivity] - intros l₁ l₂ H - cases H - . rename_i a b n h1 h2 - subst h1 h2 - simp - . rename_i a b n h1 h2 - subst h1 h2 - simp - . rename_i a n b m h1 h2 - subst h1 h2 - simp - -example (ε₁ ε₂ : ℕ+) : DP (NoisedQuery CountingQuery 1 ε₁ ε₂) ε₁ ε₂ := by - apply NoisedQueryDP - apply CountingQuery1Sensitive - -example (clip : ℕ+) : sensitivity (fun l : List ℕ => List.sum (List.map (fun n : ℕ => if n > clip then clip else n) l)) clip := by - simp [sensitivity] - intros l₁ l₂ H - cases H - . rename_i a b n h1 h2 - subst h1 h2 - simp - split - . rename_i h3 - simp - . rename_i h3 - simp at h3 - simpa - . rename_i a b n h1 h2 - subst h1 h2 - simp - split - . rename_i h3 - simp - . rename_i h3 - simp at h3 - simpa - . rename_i a n b m h1 h2 - subst h1 h2 - simp - split - . split - . rename_i h3 h4 - simp - . rename_i h3 h4 - simp at h4 - sorry -- OK - . split - . rename_i h3 h4 - simp at h3 - sorry -- OK - . rename_i h3 h4 - simp at h3 h4 - sorry -- OK - -end SLang +noncomputable def sensitivity (q : List T → ℤ) (Δ : ℕ) : Prop := + ∀ l₁ l₂ : List T, Neighbour l₁ l₂ → Int.natAbs (q l₁ - q l₂) ≤ Δ From f8be9289ebb01d8a94c8a80f09cdcd97b7d8132b Mon Sep 17 00:00:00 2001 From: jtristan Date: Tue, 30 Apr 2024 10:48:19 -0400 Subject: [PATCH 04/20] Exploring switch to ENNReal --- SampCert/DiffPrivacy/Count.lean | 2 +- SampCert/DiffPrivacy/DP.lean | 18 ++++++++++++------ SampCert/DiffPrivacy/RenyiDivergence.lean | 6 ++++-- SampCert/DiffPrivacy/Sensitivity.lean | 2 ++ 4 files changed, 19 insertions(+), 9 deletions(-) diff --git a/SampCert/DiffPrivacy/Count.lean b/SampCert/DiffPrivacy/Count.lean index 65312c2c..45603cbb 100644 --- a/SampCert/DiffPrivacy/Count.lean +++ b/SampCert/DiffPrivacy/Count.lean @@ -34,7 +34,7 @@ theorem CountingQuery1Sensitive : def NoisedCountingQuery (ε₁ ε₂ : ℕ+) (l : List T) : SLang ℤ := do NoisedQuery CountingQuery 1 ε₁ ε₂ l -theorem NoisedCountingQueryDP (ε₁ ε₂ : ℕ+) : @DP T (NoisedCountingQuery ε₁ ε₂) ε₁ ε₂ := by +theorem NoisedCountingQueryDP (ε₁ ε₂ : ℕ+) : @DP T ℤ (NoisedCountingQuery ε₁ ε₂) ε₁ ε₂ := by apply NoisedQueryDP apply CountingQuery1Sensitive diff --git a/SampCert/DiffPrivacy/DP.lean b/SampCert/DiffPrivacy/DP.lean index 743d9515..73decd85 100644 --- a/SampCert/DiffPrivacy/DP.lean +++ b/SampCert/DiffPrivacy/DP.lean @@ -16,11 +16,11 @@ noncomputable section open Classical Nat Int Real -def DP (q : List T → SLang ℤ) (ε₁ ε₂ : ℕ+) : Prop := +def DP (q : List T → SLang U) (ε₁ ε₂ : ℕ+) : Prop := ∀ α : ℝ, 1 < α → ∀ l₁ l₂ : List T, Neighbour l₁ l₂ → RenyiDivergence - (fun x : ℤ => (q l₁ x).toReal) - (fun x : ℤ => (q l₂ x).toReal) α + (fun x : U => (q l₁ x).toReal) + (fun x : U => (q l₂ x).toReal) α ≤ (1/2) * (ε₁ / ε₂)^2 * α namespace SLang @@ -88,8 +88,14 @@ def Compose (nq1 nq2 : List T → SLang ℤ) (l : List T) : SLang (ℤ × ℤ) : def RAdd (ε₁ ε₂ ε₃ ε₄ : ℕ+) : ℕ+ × ℕ+ := (ε₁ * ε₃ + ε₂ * ε₄,ε₃ * ε₄) --- theorem DPCompose (nq1 nq2 : List T → SLang ℤ) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h1 : DP nq1 ε₁ ε₂) (h2 : DP nq2 ε₃ ε₄) : --- DP (Compose nq1 nq2) (RAdd ε₁ ε₂ ε₃ ε₄).1 (RAdd ε₁ ε₂ ε₃ ε₄).2 := by --- sorry +theorem DPCompose (nq1 nq2 : List T → SLang ℤ) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h1 : DP nq1 ε₁ ε₂) (h2 : DP nq2 ε₃ ε₄) : + DP (Compose nq1 nq2) (RAdd ε₁ ε₂ ε₃ ε₄).1 (RAdd ε₁ ε₂ ε₃ ε₄).2 := by + simp [Compose, RAdd, RenyiDivergence, DP] + intro α h1 l₁ l₂ h2 + rw [tsum_prod'] + . simp + sorry + . sorry + . sorry end SLang diff --git a/SampCert/DiffPrivacy/RenyiDivergence.lean b/SampCert/DiffPrivacy/RenyiDivergence.lean index 0788a20f..a8f37fae 100644 --- a/SampCert/DiffPrivacy/RenyiDivergence.lean +++ b/SampCert/DiffPrivacy/RenyiDivergence.lean @@ -9,5 +9,7 @@ import Mathlib.NumberTheory.ModularForms.JacobiTheta.TwoVariable open Real -noncomputable def RenyiDivergence (p q : ℤ → ℝ) (α : ℝ) : ℝ := - (1 / (α - 1)) * Real.log (∑' x : ℤ, (p x)^α * (q x)^(1 - α)) +variable {T : Type} + +noncomputable def RenyiDivergence (p q : T → ℝ) (α : ℝ) : ℝ := + (1 / (α - 1)) * Real.log (∑' x : T, (p x)^α * (q x)^(1 - α)) diff --git a/SampCert/DiffPrivacy/Sensitivity.lean b/SampCert/DiffPrivacy/Sensitivity.lean index 831c9722..1b07a645 100644 --- a/SampCert/DiffPrivacy/Sensitivity.lean +++ b/SampCert/DiffPrivacy/Sensitivity.lean @@ -16,5 +16,7 @@ import SampCert.DiffPrivacy.Neighbours open Classical Nat Int Real +variable {T : Type} + noncomputable def sensitivity (q : List T → ℤ) (Δ : ℕ) : Prop := ∀ l₁ l₂ : List T, Neighbour l₁ l₂ → Int.natAbs (q l₁ - q l₂) ≤ Δ From 4a45446c33cd55e5d1f3c6a4cc7d408af8ffce6a Mon Sep 17 00:00:00 2001 From: jtristan Date: Wed, 1 May 2024 09:29:16 -0400 Subject: [PATCH 05/20] Reorg and DP rewrite --- SampCert/DiffPrivacy/BoundedAvg.lean | 37 ++++- SampCert/DiffPrivacy/BoundedSum.lean | 168 +++------------------- SampCert/DiffPrivacy/DP.lean | 112 ++++++++++++++- SampCert/DiffPrivacy/RenyiDivergence.lean | 3 +- SampCert/DiffPrivacy/Sensitivity.lean | 3 - 5 files changed, 162 insertions(+), 161 deletions(-) diff --git a/SampCert/DiffPrivacy/BoundedAvg.lean b/SampCert/DiffPrivacy/BoundedAvg.lean index 36ac1d30..8e800a35 100644 --- a/SampCert/DiffPrivacy/BoundedAvg.lean +++ b/SampCert/DiffPrivacy/BoundedAvg.lean @@ -17,13 +17,42 @@ namespace SLang def NoisedBoundedAvgQuery (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) (l : List ℤ) : SLang ℤ := do let S ← NoisedBoundedSumQuery L U h ε₁ (2 * ε₂) l let C ← NoisedCountingQuery ε₁ (2 * ε₂) l - if C ≤ 1 then return (L + U) / 2 - else return S / C + return S / C + +def NoisedBoundedAvgQuery' (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) (l : List ℤ) : SLang ℤ := + let X := Compose (NoisedBoundedSumQuery L U h ε₁ (2 * ε₂)) (NoisedCountingQuery ε₁ (2 * ε₂)) + PostProcess X (fun z => z.1 / z.2) l + +theorem NoisedBoundedAvgQueryIdentical (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) : + NoisedBoundedAvgQuery' L U h ε₁ ε₂ = NoisedBoundedAvgQuery L U h ε₁ ε₂ := by + ext l x + simp [NoisedBoundedAvgQuery, NoisedBoundedAvgQuery', PostProcess, Compose] theorem BoundedSumQueryDP (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) : DP (NoisedBoundedAvgQuery L U h ε₁ ε₂) ε₁ ε₂ := by + rw [← NoisedBoundedAvgQueryIdentical] + unfold NoisedBoundedAvgQuery' + simp only + have A := @NoisedCountingQueryDP ℤ ε₁ (2 * ε₂) have B := @NoisedBoundedSumQueryDP L U h ε₁ (2 * ε₂) - sorry - + have C := DPCompose B A + have D := DPPostProcess C (fun z => z.1 / z.2) + simp [RAdd] at D + + apply DP_cancel_sigma (ε₁ * (2 * ε₂) + 2 * ε₂ * ε₁) (2 * ε₂ * (2 * ε₂)) ε₁ ε₂ D + ring_nf + simp + ring_nf + rw [pow_two] + rw [← mul_assoc] + have A : (ε₂ : ℝ) ≠ 0 := by + simp + conv => + left + left + rw [mul_assoc] + right + rw [mul_inv_cancel A] + simp end SLang diff --git a/SampCert/DiffPrivacy/BoundedSum.lean b/SampCert/DiffPrivacy/BoundedSum.lean index d84e1ae3..d192f999 100644 --- a/SampCert/DiffPrivacy/BoundedSum.lean +++ b/SampCert/DiffPrivacy/BoundedSum.lean @@ -15,46 +15,25 @@ noncomputable section namespace SLang --- N+ and not Z because of sensitivity measure --- Can be generalized --- variable {T : Type} [Add T] [LinearOrder T] [coe: Coe T ℤ] --- variable {h : Monotone coe.coe} - --- instance : OrderHom T ℤ where --- toFun := coe.coe --- monotone' := h - --- def BoundedSumQuery (clip₁ clip₂ : ℤ) (l : List ℤ) : ℤ := --- List.sum (List.map (fun n : ℤ => if n < clip₁ then clip₁ else if clip₂ < n then clip₂ else n) l) - def BoundedSumQuery (L U : ℤ) (l : List ℤ) : ℤ := - List.sum (List.map (fun n : ℤ => max (min n U) L) l) - --- theorem natAbs1 {n : ℕ} {x : ℤ} (h : n ≤ x) : --- n ≤ natAbs x := by --- cases x --- . rename_i m --- simp --- exact Int.ofNat_le.mp h --- . rename_i m --- simp --- have A : -[m+1] < 0 := negSucc_lt_zero m --- have B : n < (0 : ℤ) := by --- apply Int.lt_of_le_of_lt h A --- contradiction - --- theorem natAbs2 {n : ℕ} {x : ℤ} (h : x ≤ -[n+1]) : --- Nat.succ n ≤ natAbs x := by --- cases x --- . rename_i m --- simp --- have A : -[m+1] < 0 := negSucc_lt_zero m --- have B : n < (0 : ℤ) := by --- apply Int.lt_of_le_of_lt h A --- contradiction --- . rename_i m --- simp --- exact Int.ofNat_le.mp h + let center := |U - L| / 2 + let L := L - center + let U := U - center + let S := List.sum (List.map (fun n : ℤ => max (min (n - center) U) L) l) + S + center * List.length l + +theorem natAbs1 {n : ℕ} {x : ℤ} (h : n ≤ x) : + n ≤ natAbs x := by + cases x + . rename_i m + simp + exact Int.ofNat_le.mp h + . rename_i m + simp + have A : -[m+1] < 0 := negSucc_lt_zero m + have B : n < (0 : ℤ) := by + apply Int.lt_of_le_of_lt h A + contradiction def maxBoundPos (L U : ℤ) (h : L < U) : 0 < max (Int.natAbs L) (Int.natAbs U) := by @@ -78,19 +57,10 @@ def maxBoundPos (L U : ℤ) (h : L < U) : trivial def maxBound (L U : ℤ) (h : L < U) : ℕ+ := - ⟨ max (Int.natAbs L) (Int.natAbs U) , maxBoundPos L U h⟩ + ⟨ max (Int.natAbs L) (Int.natAbs U) , maxBoundPos L U h ⟩ theorem BoundedSumQuerySensitivity (L U : ℤ) (h : L < U) : sensitivity (BoundedSumQuery L U) (maxBound L U h) := by sorry - -- simp [sensitivity, BoundedSumQuery] - -- intros l₁ l₂ H - -- cases H - -- . rename_i a b n h1 h2 - -- subst h1 h2 - -- simp - -- right - -- . sorry - -- . sorry def NoisedBoundedSumQuery (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) (l : List ℤ) : SLang ℤ := do NoisedQuery (BoundedSumQuery L U) (maxBound L U h) ε₁ ε₂ l @@ -99,104 +69,4 @@ theorem NoisedBoundedSumQueryDP (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) : D apply NoisedQueryDP apply BoundedSumQuerySensitivity --- theorem BoundedSumSensitivity (clip₁ clip₂ : ℤ) (h : clip₁ < clip₂) : sensitivity (BoundedSumQuery clip₁ clip₂) (max (Int.natAbs clip₁) (Int.natAbs clip₂)) := by --- simp [sensitivity, BoundedSumQuery] --- intros l₁ l₂ H --- cases H --- . rename_i a b n h1 h2 --- subst h1 h2 --- simp --- split --- . left --- apply le_refl --- . split --- . right --- apply le_refl --- . rename_i h3 h4 --- simp at * --- cases n --- . rename_i n --- simp --- right --- simp at h4 --- apply natAbs1 h4 --- . rename_i n --- simp at * --- left --- apply natAbs2 h3 --- . rename_i a n b h1 h2 --- subst h1 h2 --- simp --- split --- . left --- apply le_refl --- . split --- . right --- apply le_refl --- . rename_i h3 h4 --- simp at * --- cases n --- . rename_i n --- simp --- right --- simp at h4 --- apply natAbs1 h4 --- . rename_i n --- simp at * --- left --- apply natAbs2 h3 --- . rename_i a n b m h1 h2 --- subst h1 h2 --- simp --- split --- . split --- . rename_i h1 h2 --- simp at * --- . split --- . rename_i h1 h2 h3 --- simp at * --- sorry --- . rename_i h1 h2 h3 --- simp at * --- sorry --- . split --- . split --- . rename_i h1 h2 h3 --- simp at * - --- cases clip₁ --- . cases clip₂ --- . rename_i n' m' --- simp at * --- sorry --- . rename_i n' m' --- simp at * --- sorry --- . cases clip₂ --- . rename_i n' m' --- simp at * --- sorry --- . rename_i n' m' --- simp at * --- sorry - - --- . split --- . rename_i h1 h2 h3 h4 --- simp at * --- . rename_i h1 h2 h3 h4 --- simp at * --- sorry --- . split --- . rename_i h1 h2 h3 --- simp at * --- sorry --- . split --- . rename_i h1 h2 h3 h4 --- simp at * --- sorry --- . rename_i h1 h2 h3 h4 --- simp at * --- sorry - end SLang diff --git a/SampCert/DiffPrivacy/DP.lean b/SampCert/DiffPrivacy/DP.lean index 73decd85..d93cef99 100644 --- a/SampCert/DiffPrivacy/DP.lean +++ b/SampCert/DiffPrivacy/DP.lean @@ -23,6 +23,38 @@ def DP (q : List T → SLang U) (ε₁ ε₂ : ℕ+) : Prop := (fun x : U => (q l₂ x).toReal) α ≤ (1/2) * (ε₁ / ε₂)^2 * α +theorem DP_cancel_sigma {q : List T → SLang U} (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h1 : DP q ε1 ε2) (h2 : @Eq.{1} Real + (@HDiv.hDiv.{0, 0, 0} Real Real Real + (@instHDiv.{0} Real + (@DivInvMonoid.toDiv.{0} Real + (@DivisionMonoid.toDivInvMonoid.{0} Real + (@DivisionCommMonoid.toDivisionMonoid.{0} Real + (@CommGroupWithZero.toDivisionCommMonoid.{0} Real + (@Semifield.toCommGroupWithZero.{0} Real + (@LinearOrderedSemifield.toSemifield.{0} Real + (@LinearOrderedField.toLinearOrderedSemifield.{0} Real Real.instLinearOrderedField)))))))) + (@Nat.cast.{0} Real Real.instNatCast (PNat.val ε1)) (@Nat.cast.{0} Real Real.instNatCast (PNat.val ε2))) + (@HDiv.hDiv.{0, 0, 0} Real Real Real + (@instHDiv.{0} Real + (@DivInvMonoid.toDiv.{0} Real + (@DivisionMonoid.toDivInvMonoid.{0} Real + (@DivisionCommMonoid.toDivisionMonoid.{0} Real + (@CommGroupWithZero.toDivisionCommMonoid.{0} Real + (@Semifield.toCommGroupWithZero.{0} Real + (@LinearOrderedSemifield.toSemifield.{0} Real + (@LinearOrderedField.toLinearOrderedSemifield.{0} Real Real.instLinearOrderedField)))))))) + (@Nat.cast.{0} Real Real.instNatCast (PNat.val ε3)) (@Nat.cast.{0} Real Real.instNatCast (PNat.val ε4)))) : + DP q ε3 ε4 := by + simp [DP] at * + intros α h3 l₁ l₂ h4 + replace h1 := h1 α h3 l₁ l₂ h4 + have A : (ε1 : ℝ) ^ 2 / ε2 ^ 2 = ε3 ^ 2 / ε4 ^ 2 := by + rw [← div_pow] + rw [← div_pow] + congr + rw [← A] + trivial + namespace SLang def NoisedQuery (query : List T → ℤ) (Δ : ℕ+) (ε₁ ε₂ : ℕ+) (l : List T) : SLang ℤ := do @@ -86,16 +118,88 @@ def Compose (nq1 nq2 : List T → SLang ℤ) (l : List T) : SLang (ℤ × ℤ) : return (A,B) def RAdd (ε₁ ε₂ ε₃ ε₄ : ℕ+) : ℕ+ × ℕ+ := - (ε₁ * ε₃ + ε₂ * ε₄,ε₃ * ε₄) + (ε₁ * ε₄ + ε₂ * ε₃,ε₂ * ε₄) -theorem DPCompose (nq1 nq2 : List T → SLang ℤ) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h1 : DP nq1 ε₁ ε₂) (h2 : DP nq2 ε₃ ε₄) : +theorem DPCompose {nq1 nq2 : List T → SLang ℤ} {ε₁ ε₂ ε₃ ε₄ : ℕ+} (h1 : DP nq1 ε₁ ε₂) (h2 : DP nq2 ε₃ ε₄) : DP (Compose nq1 nq2) (RAdd ε₁ ε₂ ε₃ ε₄).1 (RAdd ε₁ ε₂ ε₃ ε₄).2 := by simp [Compose, RAdd, RenyiDivergence, DP] - intro α h1 l₁ l₂ h2 + intro α h3 l₁ l₂ h4 + simp [DP] at h1 h2 + replace h1 := h1 α h3 l₁ l₂ h4 + replace h2 := h2 α h3 l₁ l₂ h4 + simp [RenyiDivergence] at h1 h2 rw [tsum_prod'] . simp - sorry + have A : ∀ b c : ℤ, ∀ l : List T, (∑' (a : ℤ), nq1 l a * ∑' (a_1 : ℤ), if b = a ∧ c = a_1 then nq2 l a_1 else 0) = nq1 l b * nq2 l c := by + sorry + conv => + left + right + right + right + intro b + right + intro c + rw [A] + rw [A] + rw [ENNReal.toReal_mul] + rw [ENNReal.toReal_mul] + rw [Real.mul_rpow ENNReal.toReal_nonneg ENNReal.toReal_nonneg] + rw [Real.mul_rpow ENNReal.toReal_nonneg ENNReal.toReal_nonneg] + rw [mul_assoc] + right + rw [mul_comm] + rw [mul_assoc] + right + rw [mul_comm] + conv => + left + right + right + right + intro b + right + intro c + rw [← mul_assoc] + have B : Summable fun c => (nq2 l₁ c).toReal ^ α * (nq2 l₂ c).toReal ^ (1 - α) := sorry + have C : Summable fun b => (nq1 l₁ b).toReal ^ α * (nq1 l₂ b).toReal ^ (1 - α) := sorry + conv => + left + right + right + right + intro b + rw [@Summable.tsum_mul_left ℤ ℝ _ _ _ (fun c => (nq2 l₁ c).toReal ^ α * (nq2 l₂ c).toReal ^ (1 - α)) _ ((nq1 l₁ b).toReal ^ α * (nq1 l₂ b).toReal ^ (1 - α)) B] + rw [Summable.tsum_mul_right ] + . rw [Real.log_mul] + . rw [mul_add] + have D := _root_.add_le_add h1 h2 + apply le_trans D + clear h1 h2 A B C D + rw [← add_mul] + rw [mul_le_mul_iff_of_pos_right] + . rw [← mul_add] + rw [mul_le_mul_iff_of_pos_left] + . ring_nf + simp + . sorry -- easy + . sorry -- easy + . sorry -- ≠ 0 , not obvious + . sorry -- ≠ 0 + . apply C . sorry . sorry +def PostProcess (nq : List T → SLang U) (pp : U → ℤ) (l : List T) : SLang ℤ := do + let A ← nq l + return pp A + +theorem DPPostProcess {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP nq ε₁ ε₂) (pp : U → ℤ) : + DP (PostProcess nq pp) ε₁ ε₂ := by + simp [PostProcess, DP, RenyiDivergence] + intro α h1 l₁ l₂ h2 + simp [DP, RenyiDivergence] at h + replace h := h α h1 l₁ l₂ h2 + sorry + end SLang diff --git a/SampCert/DiffPrivacy/RenyiDivergence.lean b/SampCert/DiffPrivacy/RenyiDivergence.lean index a8f37fae..9f380a36 100644 --- a/SampCert/DiffPrivacy/RenyiDivergence.lean +++ b/SampCert/DiffPrivacy/RenyiDivergence.lean @@ -6,8 +6,9 @@ Authors: Jean-Baptiste Tristan import Mathlib.Topology.Algebra.InfiniteSum.Ring import Mathlib.NumberTheory.ModularForms.JacobiTheta.TwoVariable +--import Mathlib.Analysis.SpecialFunctions.Pow.NNReal -open Real +open Real ENNReal variable {T : Type} diff --git a/SampCert/DiffPrivacy/Sensitivity.lean b/SampCert/DiffPrivacy/Sensitivity.lean index 1b07a645..ba903f7b 100644 --- a/SampCert/DiffPrivacy/Sensitivity.lean +++ b/SampCert/DiffPrivacy/Sensitivity.lean @@ -4,9 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jean-Baptiste Tristan -/ --- Data : sorted list of integers --- integer can be added, removed, modified - import Mathlib.Topology.Algebra.InfiniteSum.Ring import Mathlib.Algebra.Group.Basic import SampCert.DiffPrivacy.ConcentratedBound From 724920e1c5d69a244a371e697ea57601d73eb1bf Mon Sep 17 00:00:00 2001 From: jtristan Date: Wed, 1 May 2024 11:14:46 -0400 Subject: [PATCH 06/20] cleaner definition of DP --- SampCert/DiffPrivacy/BoundedAvg.lean | 23 +++---------- SampCert/DiffPrivacy/BoundedSum.lean | 2 +- SampCert/DiffPrivacy/Count.lean | 2 +- SampCert/DiffPrivacy/DP.lean | 51 +++++----------------------- 4 files changed, 15 insertions(+), 63 deletions(-) diff --git a/SampCert/DiffPrivacy/BoundedAvg.lean b/SampCert/DiffPrivacy/BoundedAvg.lean index 8e800a35..3138b0e3 100644 --- a/SampCert/DiffPrivacy/BoundedAvg.lean +++ b/SampCert/DiffPrivacy/BoundedAvg.lean @@ -28,7 +28,7 @@ theorem NoisedBoundedAvgQueryIdentical (L U : ℤ) (h : L < U) (ε₁ ε₂ : ext l x simp [NoisedBoundedAvgQuery, NoisedBoundedAvgQuery', PostProcess, Compose] -theorem BoundedSumQueryDP (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) : DP (NoisedBoundedAvgQuery L U h ε₁ ε₂) ε₁ ε₂ := by +theorem BoundedSumQueryDP (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) : DP (NoisedBoundedAvgQuery L U h ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by rw [← NoisedBoundedAvgQueryIdentical] unfold NoisedBoundedAvgQuery' simp only @@ -36,23 +36,10 @@ theorem BoundedSumQueryDP (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) : DP (Noi have A := @NoisedCountingQueryDP ℤ ε₁ (2 * ε₂) have B := @NoisedBoundedSumQueryDP L U h ε₁ (2 * ε₂) have C := DPCompose B A + simp at C + ring_nf at C + rw [← division_def] at C have D := DPPostProcess C (fun z => z.1 / z.2) - simp [RAdd] at D - - apply DP_cancel_sigma (ε₁ * (2 * ε₂) + 2 * ε₂ * ε₁) (2 * ε₂ * (2 * ε₂)) ε₁ ε₂ D - ring_nf - simp - ring_nf - rw [pow_two] - rw [← mul_assoc] - have A : (ε₂ : ℝ) ≠ 0 := by - simp - conv => - left - left - rw [mul_assoc] - right - rw [mul_inv_cancel A] - simp + exact D end SLang diff --git a/SampCert/DiffPrivacy/BoundedSum.lean b/SampCert/DiffPrivacy/BoundedSum.lean index d192f999..63694cf0 100644 --- a/SampCert/DiffPrivacy/BoundedSum.lean +++ b/SampCert/DiffPrivacy/BoundedSum.lean @@ -65,7 +65,7 @@ theorem BoundedSumQuerySensitivity (L U : ℤ) (h : L < U) : sensitivity (Bounde def NoisedBoundedSumQuery (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) (l : List ℤ) : SLang ℤ := do NoisedQuery (BoundedSumQuery L U) (maxBound L U h) ε₁ ε₂ l -theorem NoisedBoundedSumQueryDP (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) : DP (NoisedBoundedSumQuery L U h ε₁ ε₂) ε₁ ε₂ := by +theorem NoisedBoundedSumQueryDP (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) : DP (NoisedBoundedSumQuery L U h ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by apply NoisedQueryDP apply BoundedSumQuerySensitivity diff --git a/SampCert/DiffPrivacy/Count.lean b/SampCert/DiffPrivacy/Count.lean index 45603cbb..28108435 100644 --- a/SampCert/DiffPrivacy/Count.lean +++ b/SampCert/DiffPrivacy/Count.lean @@ -34,7 +34,7 @@ theorem CountingQuery1Sensitive : def NoisedCountingQuery (ε₁ ε₂ : ℕ+) (l : List T) : SLang ℤ := do NoisedQuery CountingQuery 1 ε₁ ε₂ l -theorem NoisedCountingQueryDP (ε₁ ε₂ : ℕ+) : @DP T ℤ (NoisedCountingQuery ε₁ ε₂) ε₁ ε₂ := by +theorem NoisedCountingQueryDP (ε₁ ε₂ : ℕ+) : @DP T ℤ (NoisedCountingQuery ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by apply NoisedQueryDP apply CountingQuery1Sensitive diff --git a/SampCert/DiffPrivacy/DP.lean b/SampCert/DiffPrivacy/DP.lean index d93cef99..b449eb95 100644 --- a/SampCert/DiffPrivacy/DP.lean +++ b/SampCert/DiffPrivacy/DP.lean @@ -16,44 +16,12 @@ noncomputable section open Classical Nat Int Real -def DP (q : List T → SLang U) (ε₁ ε₂ : ℕ+) : Prop := +def DP (q : List T → SLang U) (ε : ℝ) : Prop := ∀ α : ℝ, 1 < α → ∀ l₁ l₂ : List T, Neighbour l₁ l₂ → RenyiDivergence (fun x : U => (q l₁ x).toReal) (fun x : U => (q l₂ x).toReal) α - ≤ (1/2) * (ε₁ / ε₂)^2 * α - -theorem DP_cancel_sigma {q : List T → SLang U} (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h1 : DP q ε1 ε2) (h2 : @Eq.{1} Real - (@HDiv.hDiv.{0, 0, 0} Real Real Real - (@instHDiv.{0} Real - (@DivInvMonoid.toDiv.{0} Real - (@DivisionMonoid.toDivInvMonoid.{0} Real - (@DivisionCommMonoid.toDivisionMonoid.{0} Real - (@CommGroupWithZero.toDivisionCommMonoid.{0} Real - (@Semifield.toCommGroupWithZero.{0} Real - (@LinearOrderedSemifield.toSemifield.{0} Real - (@LinearOrderedField.toLinearOrderedSemifield.{0} Real Real.instLinearOrderedField)))))))) - (@Nat.cast.{0} Real Real.instNatCast (PNat.val ε1)) (@Nat.cast.{0} Real Real.instNatCast (PNat.val ε2))) - (@HDiv.hDiv.{0, 0, 0} Real Real Real - (@instHDiv.{0} Real - (@DivInvMonoid.toDiv.{0} Real - (@DivisionMonoid.toDivInvMonoid.{0} Real - (@DivisionCommMonoid.toDivisionMonoid.{0} Real - (@CommGroupWithZero.toDivisionCommMonoid.{0} Real - (@Semifield.toCommGroupWithZero.{0} Real - (@LinearOrderedSemifield.toSemifield.{0} Real - (@LinearOrderedField.toLinearOrderedSemifield.{0} Real Real.instLinearOrderedField)))))))) - (@Nat.cast.{0} Real Real.instNatCast (PNat.val ε3)) (@Nat.cast.{0} Real Real.instNatCast (PNat.val ε4)))) : - DP q ε3 ε4 := by - simp [DP] at * - intros α h3 l₁ l₂ h4 - replace h1 := h1 α h3 l₁ l₂ h4 - have A : (ε1 : ℝ) ^ 2 / ε2 ^ 2 = ε3 ^ 2 / ε4 ^ 2 := by - rw [← div_pow] - rw [← div_pow] - congr - rw [← A] - trivial + ≤ (1/2) * ε ^ 2 * α namespace SLang @@ -61,7 +29,7 @@ def NoisedQuery (query : List T → ℤ) (Δ : ℕ+) (ε₁ ε₂ : ℕ+) (l : L DiscreteGaussianGenSample (Δ * ε₂) ε₁ (query l) theorem NoisedQueryDP (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_sensitivity : sensitivity query Δ) : - DP (NoisedQuery query Δ ε₁ ε₂) ε₁ ε₂ := by + DP (NoisedQuery query Δ ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by simp [DP, NoisedQuery] intros α h1 l₁ l₂ h2 have A := @DiscreteGaussianGenSampleZeroConcentrated α h1 (Δ * ε₂) ε₁ (query l₁) (query l₂) @@ -117,12 +85,9 @@ def Compose (nq1 nq2 : List T → SLang ℤ) (l : List T) : SLang (ℤ × ℤ) : let B ← nq2 l return (A,B) -def RAdd (ε₁ ε₂ ε₃ ε₄ : ℕ+) : ℕ+ × ℕ+ := - (ε₁ * ε₄ + ε₂ * ε₃,ε₂ * ε₄) - -theorem DPCompose {nq1 nq2 : List T → SLang ℤ} {ε₁ ε₂ ε₃ ε₄ : ℕ+} (h1 : DP nq1 ε₁ ε₂) (h2 : DP nq2 ε₃ ε₄) : - DP (Compose nq1 nq2) (RAdd ε₁ ε₂ ε₃ ε₄).1 (RAdd ε₁ ε₂ ε₃ ε₄).2 := by - simp [Compose, RAdd, RenyiDivergence, DP] +theorem DPCompose {nq1 nq2 : List T → SLang ℤ} {ε₁ ε₂ ε₃ ε₄ : ℕ+} (h1 : DP nq1 ((ε₁ : ℝ) / ε₂)) (h2 : DP nq2 ((ε₃ : ℝ) / ε₄)) : + DP (Compose nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by + simp [Compose, RenyiDivergence, DP] intro α h3 l₁ l₂ h4 simp [DP] at h1 h2 replace h1 := h1 α h3 l₁ l₂ h4 @@ -194,8 +159,8 @@ def PostProcess (nq : List T → SLang U) (pp : U → ℤ) (l : List T) : SLang let A ← nq l return pp A -theorem DPPostProcess {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP nq ε₁ ε₂) (pp : U → ℤ) : - DP (PostProcess nq pp) ε₁ ε₂ := by +theorem DPPostProcess {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP nq ((ε₁ : ℝ) / ε₂)) (pp : U → ℤ) : + DP (PostProcess nq pp) ((ε₁ : ℝ) / ε₂) := by simp [PostProcess, DP, RenyiDivergence] intro α h1 l₁ l₂ h2 simp [DP, RenyiDivergence] at h From bec417803c2a0e77a0d9d7c505f42ace49dfc15c Mon Sep 17 00:00:00 2001 From: jtristan Date: Wed, 1 May 2024 14:14:09 -0400 Subject: [PATCH 07/20] Switched Renyi div to ENNReal --- SampCert/DiffPrivacy/ConcentratedBound.lean | 151 +++++++++++++++++- SampCert/DiffPrivacy/DP.lean | 57 +++---- SampCert/DiffPrivacy/RenyiDivergence.lean | 6 +- SampCert/Samplers/GaussianGen/Properties.lean | 4 +- 4 files changed, 172 insertions(+), 46 deletions(-) diff --git a/SampCert/DiffPrivacy/ConcentratedBound.lean b/SampCert/DiffPrivacy/ConcentratedBound.lean index a6091200..f7fcc1b7 100644 --- a/SampCert/DiffPrivacy/ConcentratedBound.lean +++ b/SampCert/DiffPrivacy/ConcentratedBound.lean @@ -75,12 +75,14 @@ theorem SG_Renyi_simplify {σ : ℝ} (h : σ ≠ 0) (μ ν : ℤ) (α : ℝ) : left rw [← mul_assoc] +noncomputable def RenyiDivergence' (p q : T → ℝ) (α : ℝ) : ℝ := + (1 / (α - 1)) * Real.log (∑' x : T, (p x)^α * (q x)^(1 - α)) -theorem RenyiDivergenceBound {σ : ℝ} (h : σ ≠ 0) (μ : ℤ) (α : ℝ) (h' : α > 1) : - RenyiDivergence (fun (x : ℤ) => (gauss_term_ℝ σ μ) x / ∑' x : ℤ, (gauss_term_ℝ σ μ) x) +theorem RenyiDivergenceBound {σ : ℝ} (h : σ ≠ 0) (μ : ℤ) (α : ℝ) (h' : 1 < α) : + RenyiDivergence' (fun (x : ℤ) => (gauss_term_ℝ σ μ) x / ∑' x : ℤ, (gauss_term_ℝ σ μ) x) (fun (x : ℤ) => (gauss_term_ℝ σ (0 : ℤ)) x / ∑' x : ℤ, (gauss_term_ℝ σ (0 : ℤ)) x) α ≤ α * (μ^2 / (2 * σ^2)) := by - unfold RenyiDivergence + unfold RenyiDivergence' have A : 0 < 1 / (α - 1) := by simp [h'] rw [← le_div_iff' A] @@ -300,9 +302,9 @@ theorem sg_mul_simplify (ss : ℝ) (x μ ν : ℤ) : rw [← neg_add] theorem SG_Renyi_shift {σ : ℝ} (h : σ ≠ 0) (α : ℝ) (μ ν τ : ℤ) : - RenyiDivergence (fun (x : ℤ) => (gauss_term_ℝ σ μ) x / ∑' x : ℤ, (gauss_term_ℝ σ μ) x) (fun (x : ℤ) => (gauss_term_ℝ σ ν) x / ∑' x : ℤ, (gauss_term_ℝ σ ν) x) α - = RenyiDivergence (fun (x : ℤ) => (gauss_term_ℝ σ ((μ + τ) : ℤ)) x / ∑' x : ℤ, (gauss_term_ℝ σ ((μ + τ) : ℤ)) x) (fun (x : ℤ) => (gauss_term_ℝ σ ((ν + τ) : ℤ)) x / ∑' x : ℤ, (gauss_term_ℝ σ ((ν + τ) : ℤ)) x) α := by - unfold RenyiDivergence + RenyiDivergence' (fun (x : ℤ) => (gauss_term_ℝ σ μ) x / ∑' x : ℤ, (gauss_term_ℝ σ μ) x) (fun (x : ℤ) => (gauss_term_ℝ σ ν) x / ∑' x : ℤ, (gauss_term_ℝ σ ν) x) α + = RenyiDivergence' (fun (x : ℤ) => (gauss_term_ℝ σ ((μ + τ) : ℤ)) x / ∑' x : ℤ, (gauss_term_ℝ σ ((μ + τ) : ℤ)) x) (fun (x : ℤ) => (gauss_term_ℝ σ ((ν + τ) : ℤ)) x / ∑' x : ℤ, (gauss_term_ℝ σ ((ν + τ) : ℤ)) x) α := by + unfold RenyiDivergence' congr 2 conv => left @@ -410,11 +412,144 @@ theorem SG_Renyi_shift {σ : ℝ} (h : σ ≠ 0) (α : ℝ) (μ ν τ : ℤ) : apply Summable.mul_right apply summable_gauss_term' h -theorem RenyiDivergenceBound' {σ α : ℝ} (h : σ ≠ 0) (h' : 1 < α) (μ ν : ℤ) : - RenyiDivergence (fun (x : ℤ) => discrete_gaussian σ μ x) +theorem RenyiDivergenceBound_pre {σ α : ℝ} (h : σ ≠ 0) (h' : 1 < α) (μ ν : ℤ) : + RenyiDivergence' (fun (x : ℤ) => discrete_gaussian σ μ x) (fun (x : ℤ) => discrete_gaussian σ ν x) α ≤ α * (((μ - ν) : ℤ)^2 / (2 * σ^2)) := by unfold discrete_gaussian rw [SG_Renyi_shift h α μ ν (-ν)] rw [add_right_neg] apply RenyiDivergenceBound h (μ + -ν) α h' + +theorem RenyiSumSG_nonneg {σ α : ℝ} (h : σ ≠ 0) (μ ν n : ℤ) : + 0 ≤ discrete_gaussian σ μ n ^ α * discrete_gaussian σ ν n ^ (1 - α) := by + have A := discrete_gaussian_nonneg h μ n + have B := discrete_gaussian_nonneg h ν n + rw [mul_nonneg_iff] + left + constructor + . apply Real.rpow_nonneg A + . apply Real.rpow_nonneg B + +theorem SummableRenyiGauss {σ : ℝ} (h : σ ≠ 0) (μ ν : ℤ) (α : ℝ) : + Summable fun (x : ℤ) => discrete_gaussian σ μ x ^ α * discrete_gaussian σ ν x ^ (1 - α) := by + simp [discrete_gaussian] + have B : ∀ μ : ℤ, ∀ x : ℝ, 0 ≤ (gauss_term_ℝ σ μ) x := by + intro μ x + unfold gauss_term_ℝ + apply exp_nonneg + have C : ∀ μ : ℤ, 0 ≤ (∑' (x : ℤ), (gauss_term_ℝ σ μ) x)⁻¹ := by + intro μ + rw [inv_nonneg] + apply le_of_lt + apply sum_gauss_term_pos h + conv => + right + intro x + rw [division_def] + rw [division_def] + rw [mul_rpow (B μ x) (C μ)] + rw [mul_rpow (B ν x) (C ν)] + conv => + right + intro x + rw [mul_assoc] + right + rw [← mul_assoc] + left + rw [mul_comm] + conv => + right + intro x + ring_nf + apply Summable.mul_right + apply Summable.mul_right + unfold gauss_term_ℝ + conv => + right + intro x + rw [← Real.exp_mul] + rw [← Real.exp_mul] + rw [← exp_add] + rw [← mul_div_right_comm] + rw [← mul_div_right_comm] + rw [div_add_div_same] + rw [mul_sub_left_distrib] + rw [mul_one] + right + left + ring_nf + + have X : ∀ x : ℤ, + ↑x * ↑μ * α * 2 - ↑x * α * ↑ν * 2 + ↑x * ↑ν * 2 + (-↑x ^ 2 - ↑μ ^ 2 * α) + (α * ↑ν ^ 2 - ↑ν ^ 2) + = - ((↑x ^ 2 - 2 * x * (↑μ * α - α * ↑ν + ↑ν)) + (↑μ ^ 2 * α - α * ↑ν ^ 2 + ↑ν ^ 2)) := by + intro x + ring_nf + conv => + right + intro x + rw [X] + clear X + + have X : (↑μ ^ 2 * α - α * ↑ν ^ 2 + ↑ν ^ 2) + = (↑μ * α - α * ↑ν + ↑ν)^2 + (- ↑μ * α * ↑ν * 2 + ↑μ * α ^ 2 * ↑ν * 2 - + ↑μ ^ 2 * α ^ 2 + α * ↑ν ^ 2 - α ^ 2 * ↑ν ^ 2 + α * ↑μ ^ 2) := by + ring_nf + conv => + right + intro x + rw [X] + rw [← add_assoc] + clear X + + have X : ∀ x : ℤ, (x - (↑μ * α - α * ↑ν + ↑ν))^2 = ↑x ^ 2 - 2 * ↑x * (↑μ * α - α * ↑ν + ↑ν) + (↑μ * α - α * ↑ν + ↑ν) ^ 2 := by + intro x + ring_nf + conv => + right + intro x + rw [← X] + rw [neg_add] + rw [← div_add_div_same] + rw [exp_add] + clear X + apply Summable.mul_right + apply summable_gauss_term' h + +theorem RenyiDivergenceBound' {σ α : ℝ} (h : σ ≠ 0) (h' : 1 < α) (μ ν : ℤ) : + RenyiDivergence (fun (x : ℤ) => ENNReal.ofReal (discrete_gaussian σ μ x)) + (fun (x : ℤ) => ENNReal.ofReal (discrete_gaussian σ ν x)) + α ≤ α * (((μ - ν) : ℤ)^2 / (2 * σ^2)) := by + have A : RenyiDivergence (fun (x : ℤ) => ENNReal.ofReal (discrete_gaussian σ μ x)) + (fun (x : ℤ) => ENNReal.ofReal (discrete_gaussian σ ν x)) + α = RenyiDivergence' (fun (x : ℤ) => discrete_gaussian σ μ x) + (fun (x : ℤ) => discrete_gaussian σ ν x) + α := by + unfold RenyiDivergence + unfold RenyiDivergence' + congr + simp + have A₁ : ∀ x : ℤ, 0 ≤ discrete_gaussian σ μ x ^ α := by + intro x + apply Real.rpow_nonneg + apply discrete_gaussian_nonneg h μ x + conv => + left + right + right + intro x + rw [ENNReal.ofReal_rpow_of_pos (discrete_gaussian_pos h μ x)] + rw [ENNReal.ofReal_rpow_of_pos (discrete_gaussian_pos h ν x)] + rw [← ENNReal.ofReal_mul (A₁ x)] + rw [← ENNReal.ofReal_tsum_of_nonneg] + . simp + apply tsum_nonneg + intro i + apply RenyiSumSG_nonneg h + . apply RenyiSumSG_nonneg h + . apply SummableRenyiGauss h + + + + rw [A] + apply RenyiDivergenceBound_pre h h' diff --git a/SampCert/DiffPrivacy/DP.lean b/SampCert/DiffPrivacy/DP.lean index b449eb95..929a745c 100644 --- a/SampCert/DiffPrivacy/DP.lean +++ b/SampCert/DiffPrivacy/DP.lean @@ -14,14 +14,11 @@ import SampCert.DiffPrivacy.Sensitivity noncomputable section -open Classical Nat Int Real +open Classical Nat Int Real ENNReal def DP (q : List T → SLang U) (ε : ℝ) : Prop := ∀ α : ℝ, 1 < α → ∀ l₁ l₂ : List T, Neighbour l₁ l₂ → - RenyiDivergence - (fun x : U => (q l₁ x).toReal) - (fun x : U => (q l₂ x).toReal) α - ≤ (1/2) * ε ^ 2 * α + RenyiDivergence (q l₁) (q l₂) α ≤ (1/2) * ε ^ 2 * α namespace SLang @@ -76,7 +73,7 @@ theorem NoisedQueryDP (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_ trivial . apply sq_nonneg . rw [pow_two] - rw [mul_pos_iff] + rw [_root_.mul_pos_iff] left simp @@ -93,7 +90,7 @@ theorem DPCompose {nq1 nq2 : List T → SLang ℤ} {ε₁ ε₂ ε₃ ε₄ : replace h1 := h1 α h3 l₁ l₂ h4 replace h2 := h2 α h3 l₁ l₂ h4 simp [RenyiDivergence] at h1 h2 - rw [tsum_prod'] + rw [tsum_prod' ENNReal.summable (fun b : ℤ => ENNReal.summable)] . simp have A : ∀ b c : ℤ, ∀ l : List T, (∑' (a : ℤ), nq1 l a * ∑' (a_1 : ℤ), if b = a ∧ c = a_1 then nq2 l a_1 else 0) = nq1 l b * nq2 l c := by sorry @@ -102,15 +99,14 @@ theorem DPCompose {nq1 nq2 : List T → SLang ℤ} {ε₁ ε₂ ε₃ ε₄ : right right right + right intro b right intro c rw [A] rw [A] - rw [ENNReal.toReal_mul] - rw [ENNReal.toReal_mul] - rw [Real.mul_rpow ENNReal.toReal_nonneg ENNReal.toReal_nonneg] - rw [Real.mul_rpow ENNReal.toReal_nonneg ENNReal.toReal_nonneg] + rw [ENNReal.mul_rpow_of_ne_zero sorry sorry] + rw [ENNReal.mul_rpow_of_ne_zero sorry sorry] rw [mul_assoc] right rw [mul_comm] @@ -122,38 +118,35 @@ theorem DPCompose {nq1 nq2 : List T → SLang ℤ} {ε₁ ε₂ ε₃ ε₄ : right right right + right intro b right intro c rw [← mul_assoc] - have B : Summable fun c => (nq2 l₁ c).toReal ^ α * (nq2 l₂ c).toReal ^ (1 - α) := sorry - have C : Summable fun b => (nq1 l₁ b).toReal ^ α * (nq1 l₂ b).toReal ^ (1 - α) := sorry conv => left right right right + right intro b - rw [@Summable.tsum_mul_left ℤ ℝ _ _ _ (fun c => (nq2 l₁ c).toReal ^ α * (nq2 l₂ c).toReal ^ (1 - α)) _ ((nq1 l₁ b).toReal ^ α * (nq1 l₂ b).toReal ^ (1 - α)) B] - rw [Summable.tsum_mul_right ] - . rw [Real.log_mul] - . rw [mul_add] - have D := _root_.add_le_add h1 h2 - apply le_trans D - clear h1 h2 A B C D - rw [← add_mul] - rw [mul_le_mul_iff_of_pos_right] - . rw [← mul_add] - rw [mul_le_mul_iff_of_pos_left] - . ring_nf - simp - . sorry -- easy + rw [ENNReal.tsum_mul_left] + rw [ENNReal.tsum_mul_right] + rw [ENNReal.toReal_mul] + rw [Real.log_mul] + . rw [mul_add] + have D := _root_.add_le_add h1 h2 + apply le_trans D + rw [← add_mul] + rw [mul_le_mul_iff_of_pos_right] + . rw [← mul_add] + rw [mul_le_mul_iff_of_pos_left] + . ring_nf + simp . sorry -- easy - . sorry -- ≠ 0 , not obvious - . sorry -- ≠ 0 - . apply C - . sorry - . sorry + . sorry -- easy + . sorry -- ≠ 0 , not obvious + . sorry -- ≠ 0 def PostProcess (nq : List T → SLang U) (pp : U → ℤ) (l : List T) : SLang ℤ := do let A ← nq l diff --git a/SampCert/DiffPrivacy/RenyiDivergence.lean b/SampCert/DiffPrivacy/RenyiDivergence.lean index 9f380a36..a360524f 100644 --- a/SampCert/DiffPrivacy/RenyiDivergence.lean +++ b/SampCert/DiffPrivacy/RenyiDivergence.lean @@ -6,11 +6,11 @@ Authors: Jean-Baptiste Tristan import Mathlib.Topology.Algebra.InfiniteSum.Ring import Mathlib.NumberTheory.ModularForms.JacobiTheta.TwoVariable ---import Mathlib.Analysis.SpecialFunctions.Pow.NNReal +import Mathlib.Analysis.SpecialFunctions.Pow.NNReal open Real ENNReal variable {T : Type} -noncomputable def RenyiDivergence (p q : T → ℝ) (α : ℝ) : ℝ := - (1 / (α - 1)) * Real.log (∑' x : T, (p x)^α * (q x)^(1 - α)) +noncomputable def RenyiDivergence (p q : T → ENNReal) (α : ℝ) : ℝ := + (1 / (α - 1)) * Real.log (∑' x : T, (p x)^α * (q x)^(1 - α)).toReal diff --git a/SampCert/Samplers/GaussianGen/Properties.lean b/SampCert/Samplers/GaussianGen/Properties.lean index 6145be69..e8027703 100644 --- a/SampCert/Samplers/GaussianGen/Properties.lean +++ b/SampCert/Samplers/GaussianGen/Properties.lean @@ -48,7 +48,7 @@ theorem DiscreteGaussianGenSample_apply (num : PNat) (den : PNat) (μ x : ℤ) : . rw [SG_periodic' A] theorem DiscreteGaussianGenSampleZeroConcentrated {α : ℝ} (h : 1 < α) (num : PNat) (den : PNat) (μ ν : ℤ) : - RenyiDivergence (fun x : ℤ => ((DiscreteGaussianGenSample num den μ) x).toReal) (fun x : ℤ => ((DiscreteGaussianGenSample num den ν) x).toReal) α ≤ + RenyiDivergence ((DiscreteGaussianGenSample num den μ)) (DiscreteGaussianGenSample num den ν) α ≤ α * (((μ - ν) : ℤ)^2 / (2 * ((num : ℝ) / (den : ℝ))^2)) := by have A : (num : ℝ) / (den : ℝ) ≠ 0 := by simp only [ne_eq, div_eq_zero_iff, cast_eq_zero, PNat.ne_zero, or_self, not_false_eq_true] @@ -57,10 +57,8 @@ theorem DiscreteGaussianGenSampleZeroConcentrated {α : ℝ} (h : 1 < α) (num : congr . intro x rw [DiscreteGaussianGenSample_apply] - rw [ENNReal.toReal_ofReal (discrete_gaussian_nonneg A μ x)] . intro x rw [DiscreteGaussianGenSample_apply] - rw [ENNReal.toReal_ofReal (discrete_gaussian_nonneg A ν x)] . skip apply RenyiDivergenceBound' A h From c5fcdc5d9b88df4fb8ee21ae12a3f88e74f747d0 Mon Sep 17 00:00:00 2001 From: jtristan Date: Thu, 2 May 2024 17:34:00 -0400 Subject: [PATCH 08/20] Proved composition and mechanism with additinoal properties --- SampCert/DiffPrivacy/BoundedAvg.lean | 20 +- SampCert/DiffPrivacy/ConcentratedBound.lean | 3 - SampCert/DiffPrivacy/Count.lean | 15 +- SampCert/DiffPrivacy/DP.lean | 371 +++++++++++++++++++- 4 files changed, 386 insertions(+), 23 deletions(-) diff --git a/SampCert/DiffPrivacy/BoundedAvg.lean b/SampCert/DiffPrivacy/BoundedAvg.lean index 3138b0e3..107e1e48 100644 --- a/SampCert/DiffPrivacy/BoundedAvg.lean +++ b/SampCert/DiffPrivacy/BoundedAvg.lean @@ -33,13 +33,17 @@ theorem BoundedSumQueryDP (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) : DP (Noi unfold NoisedBoundedAvgQuery' simp only - have A := @NoisedCountingQueryDP ℤ ε₁ (2 * ε₂) - have B := @NoisedBoundedSumQueryDP L U h ε₁ (2 * ε₂) - have C := DPCompose B A - simp at C - ring_nf at C - rw [← division_def] at C - have D := DPPostProcess C (fun z => z.1 / z.2) - exact D + have A₁ := @NoisedCountingQueryDP ℤ ε₁ (2 * ε₂) + have A₂ := @NoisedCountingQuery_NonZeroNQ ℤ ε₁ (2 * ε₂) + have A₃ := @NoisedCountingQuery_NonTopNQ ℤ ε₁ (2 * ε₂) + have A₄ := @NoisedCountingQuery_NonTopRDNQ ℤ ε₁ (2 * ε₂) + sorry + -- have B := @NoisedBoundedSumQueryDP L U h ε₁ (2 * ε₂) + -- have C := DPCompose B A + -- simp at C + -- ring_nf at C + -- rw [← division_def] at C + -- have D := DPPostProcess C (fun z => z.1 / z.2) + -- exact D end SLang diff --git a/SampCert/DiffPrivacy/ConcentratedBound.lean b/SampCert/DiffPrivacy/ConcentratedBound.lean index f7fcc1b7..e7ce1d6d 100644 --- a/SampCert/DiffPrivacy/ConcentratedBound.lean +++ b/SampCert/DiffPrivacy/ConcentratedBound.lean @@ -548,8 +548,5 @@ theorem RenyiDivergenceBound' {σ α : ℝ} (h : σ ≠ 0) (h' : 1 < α) (μ ν apply RenyiSumSG_nonneg h . apply RenyiSumSG_nonneg h . apply SummableRenyiGauss h - - - rw [A] apply RenyiDivergenceBound_pre h h' diff --git a/SampCert/DiffPrivacy/Count.lean b/SampCert/DiffPrivacy/Count.lean index 28108435..d5811df5 100644 --- a/SampCert/DiffPrivacy/Count.lean +++ b/SampCert/DiffPrivacy/Count.lean @@ -34,8 +34,21 @@ theorem CountingQuery1Sensitive : def NoisedCountingQuery (ε₁ ε₂ : ℕ+) (l : List T) : SLang ℤ := do NoisedQuery CountingQuery 1 ε₁ ε₂ l -theorem NoisedCountingQueryDP (ε₁ ε₂ : ℕ+) : @DP T ℤ (NoisedCountingQuery ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by +theorem NoisedCountingQueryDP (ε₁ ε₂ : ℕ+) : + @DP T ℤ (NoisedCountingQuery ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by apply NoisedQueryDP apply CountingQuery1Sensitive +theorem NoisedCountingQuery_NonZeroNQ (ε₁ ε₂ : ℕ+) : + @NonZeroNQ T ℤ (NoisedCountingQuery ε₁ ε₂) := by + apply NoisedQuery_NonZeroNQ + +theorem NoisedCountingQuery_NonTopNQ (ε₁ ε₂ : ℕ+) : + @NonTopNQ T ℤ (NoisedCountingQuery ε₁ ε₂) := by + apply NoisedQuery_NonTopNQ + +theorem NoisedCountingQuery_NonTopRDNQ (ε₁ ε₂ : ℕ+) : + @NonTopRDNQ T ℤ (NoisedCountingQuery ε₁ ε₂) := by + apply NoisedQuery_NonTopRDNQ + end SLang diff --git a/SampCert/DiffPrivacy/DP.lean b/SampCert/DiffPrivacy/DP.lean index 929a745c..5dd7c2e2 100644 --- a/SampCert/DiffPrivacy/DP.lean +++ b/SampCert/DiffPrivacy/DP.lean @@ -20,6 +20,16 @@ def DP (q : List T → SLang U) (ε : ℝ) : Prop := ∀ α : ℝ, 1 < α → ∀ l₁ l₂ : List T, Neighbour l₁ l₂ → RenyiDivergence (q l₁) (q l₂) α ≤ (1/2) * ε ^ 2 * α +def NonZeroNQ (nq : List T → SLang U) := + ∀ l : List T, ∀ n : U, nq l n ≠ 0 + +def NonTopNQ (nq : List T → SLang U) := + ∀ l : List T, ∀ n : U, nq l n ≠ ⊤ + +def NonTopRDNQ (nq : List T → SLang U) : Prop := + ∀ α : ℝ, 1 < α → ∀ l₁ l₂ : List T, Neighbour l₁ l₂ → + ∑' (x : U), nq l₁ x ^ α * nq l₂ x ^ (1 - α) ≠ ⊤ + namespace SLang def NoisedQuery (query : List T → ℤ) (Δ : ℕ+) (ε₁ ε₂ : ℕ+) (l : List T) : SLang ℤ := do @@ -77,23 +87,265 @@ theorem NoisedQueryDP (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_ left simp +theorem NoisedQuery_NonZeroNQ (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : + NonZeroNQ (NoisedQuery query Δ ε₁ ε₂) := by + simp [NonZeroNQ, NoisedQuery, DiscreteGaussianGenSample] + intros l n + exists (n - query l) + simp + have A : ((Δ : ℝ) * ε₂ / ε₁) ≠ 0 := by + simp + have X := @discrete_gaussian_pos (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) A 0 (n - query l) + simp at X + trivial + +theorem NoisedQuery_NonTopNQ (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : + NonTopNQ (NoisedQuery query Δ ε₁ ε₂) := by + simp [NonTopNQ, NoisedQuery, DiscreteGaussianGenSample] + intro l n + rw [ENNReal.tsum_eq_add_tsum_ite (n - query l)] + simp + have X : ∀ x : ℤ, (@ite ℝ≥0∞ (x = n - query l) (propDecidable (x = n - query l)) 0 + (@ite ℝ≥0∞ (n = x + query l) (instDecidableEqInt n (x + query l)) + (ENNReal.ofReal (discrete_gaussian (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) 0 ↑x)) 0)) = 0 := by + intro x + split + . simp + . split + . rename_i h1 h2 + subst h2 + simp at h1 + . simp + conv => + right + left + right + intro x + rw [X] + simp + +theorem discrete_gaussian_shift {σ : ℝ} (h : σ ≠ 0) (μ : ℝ) (τ x : ℤ) : + discrete_gaussian σ μ (x - τ) = discrete_gaussian σ (μ + τ) (x) := by + simp [discrete_gaussian] + congr 1 + . simp [gauss_term_ℝ] + congr 3 + ring_nf + . rw [SG_periodic h] + +theorem NoisedQuery_NonTopRDNQ (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : + NonTopRDNQ (NoisedQuery query Δ ε₁ ε₂) := by + simp [NonTopRDNQ, NoisedQuery, DiscreteGaussianGenSample] + intro α _ l₁ l₂ _ + have A : ∀ x_1 x : ℤ, (@ite ℝ≥0∞ (x_1 = x - query l₁) (propDecidable (x_1 = x - query l₁)) 0 + (@ite ℝ≥0∞ (x = x_1 + query l₁) (instDecidableEqInt x (x_1 + query l₁)) + (ENNReal.ofReal (discrete_gaussian (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) 0 ↑x_1)) 0 )) = 0 := by + intro x y + split + . simp + . split + . rename_i h1 h2 + subst h2 + simp at h1 + . simp + have B : ∀ x_1 x : ℤ, (@ite ℝ≥0∞ (x_1 = x - query l₂) (propDecidable (x_1 = x - query l₂)) 0 + (@ite ℝ≥0∞ (x = x_1 + query l₂) (instDecidableEqInt x (x_1 + query l₂)) + (ENNReal.ofReal (discrete_gaussian (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) 0 ↑x_1)) 0)) = 0 := by + intro x y + split + . simp + . split + . rename_i h1 h2 + subst h2 + simp at h1 + . simp + conv => + right + left + right + intro x + left + rw [ENNReal.tsum_eq_add_tsum_ite (x - query l₁)] + simp + left + right + right + intro y + rw [A] + simp + conv => + right + left + right + intro x + right + rw [ENNReal.tsum_eq_add_tsum_ite (x - query l₂)] + simp + left + right + right + intro y + rw [B] + simp + clear A B + have P : (Δ : ℝ) * ε₂ / ε₁ ≠ 0 := by + simp + have A : ∀ x : ℤ, ∀ l : List T, 0 < discrete_gaussian (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) 0 (↑x - ↑(query l)) := by + intro x l + have A' := @discrete_gaussian_pos _ P 0 (x - query l) + simp at A' + trivial + have B : ∀ x : ℤ, 0 ≤ discrete_gaussian (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) 0 (↑x - ↑(query l₁)) ^ α := by + intro x + have A' := @discrete_gaussian_nonneg _ P 0 (x - query l₁) + simp at A' + apply rpow_nonneg A' + conv => + right + left + right + intro x + rw [ENNReal.ofReal_rpow_of_pos (A x l₁)] + rw [ENNReal.ofReal_rpow_of_pos (A x l₂)] + rw [← ENNReal.ofReal_mul (B x)] + rw [← ENNReal.ofReal_tsum_of_nonneg] + . simp + . intro n + have X := @RenyiSumSG_nonneg _ α P (query l₁) (query l₂) n + rw [discrete_gaussian_shift P] + rw [discrete_gaussian_shift P] + simp [X] + . have X := @SummableRenyiGauss _ P (query l₁) (query l₂) + conv => + right + intro x + rw [discrete_gaussian_shift P] + rw [discrete_gaussian_shift P] + simp [X] + def Compose (nq1 nq2 : List T → SLang ℤ) (l : List T) : SLang (ℤ × ℤ) := do let A ← nq1 l let B ← nq2 l return (A,B) -theorem DPCompose {nq1 nq2 : List T → SLang ℤ} {ε₁ ε₂ ε₃ ε₄ : ℕ+} (h1 : DP nq1 ((ε₁ : ℝ) / ε₂)) (h2 : DP nq2 ((ε₃ : ℝ) / ε₄)) : +theorem ENNReal_toTeal_NZ (x : ENNReal) (h1 : x ≠ 0) (h2 : x ≠ ⊤) : + x.toReal ≠ 0 := by + unfold ENNReal.toReal + unfold ENNReal.toNNReal + simp + intro H + cases H + . contradiction + . contradiction + +theorem simp_α_1 {α : ℝ} (h : 1 < α) : 0 < α := by + apply @lt_trans _ _ _ 1 _ _ h + simp only [zero_lt_one] + +theorem RenyiNoisedQueryNonZero {nq : List T → SLang ℤ} {α ε : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (h2 : Neighbour l₁ l₂) (h3 : DP nq ε) (h4 : NonZeroNQ nq) (h5 : NonTopRDNQ nq) (nts : NonTopNQ nq) : + (∑' (i : ℤ), nq l₁ i ^ α * nq l₂ i ^ (1 - α)).toReal ≠ 0 := by + simp [DP] at h3 + replace h3 := h3 α h1 l₁ l₂ h2 + simp [RenyiDivergence] at h3 + simp [NonZeroNQ] at h4 + simp [NonTopRDNQ] at h5 + replace h5 := h5 α h1 l₁ l₂ h2 + have h6 := h4 l₁ + have h7 := h4 l₂ + apply ENNReal_toTeal_NZ + . by_contra CONTRA + rw [ENNReal.tsum_eq_zero] at CONTRA + replace CONTRA := CONTRA 42 + replace h6 := h6 42 + replace h7 := h7 42 + rw [_root_.mul_eq_zero] at CONTRA + cases CONTRA + . rename_i h8 + rw [rpow_eq_zero_iff_of_pos] at h8 + contradiction + apply simp_α_1 h1 + . rename_i h8 + rw [ENNReal.rpow_eq_zero_iff] at h8 + cases h8 + . rename_i h8 + cases h8 + contradiction + . rename_i h8 + cases h8 + rename_i h8 h9 + replace nts := nts l₂ 42 + contradiction + . exact h5 + +theorem compose_sum_rw (nq1 nq2 : List T → SLang ℤ) (b c : ℤ) (l : List T) : + (∑' (a : ℤ), nq1 l a * ∑' (a_1 : ℤ), if b = a ∧ c = a_1 then nq2 l a_1 else 0) = nq1 l b * nq2 l c := by + have A : ∀ a b : ℤ, (∑' (a_1 : ℤ), if b = a ∧ c = a_1 then nq2 l a_1 else 0) = if b = a then (∑' (a_1 : ℤ), if c = a_1 then nq2 l a_1 else 0) else 0 := by + intro x y + split + . rename_i h + subst h + simp + . rename_i h + simp + intro h + contradiction + conv => + left + right + intro a + right + rw [A] + rw [ENNReal.tsum_eq_add_tsum_ite b] + simp + have B : ∀ x : ℤ, (@ite ℝ≥0∞ (x = b) (instDecidableEqInt x b) 0 + (@ite ℝ≥0∞ (b = x) (instDecidableEqInt b x) (nq1 l x * ∑' (a_1 : ℤ), @ite ℝ≥0∞ (c = a_1) (instDecidableEqInt c a_1) (nq2 l a_1) 0) 0)) = 0 := by + intro x + split + . simp + . split + . rename_i h1 h2 + subst h2 + contradiction + . simp + conv => + left + right + right + intro x + rw [B] + simp + congr 1 + rw [ENNReal.tsum_eq_add_tsum_ite c] + simp + have C :∀ x : ℤ, (@ite ℝ≥0∞ (x = c) (propDecidable (x = c)) 0 (@ite ℝ≥0∞ (c = x) (instDecidableEqInt c x) (nq2 l x) 0)) = 0 := by + intro x + split + . simp + . split + . rename_i h1 h2 + subst h2 + contradiction + . simp + conv => + left + right + right + intro X + rw [C] + simp + +theorem DPCompose {nq1 nq2 : List T → SLang ℤ} {ε₁ ε₂ ε₃ ε₄ : ℕ+} (h1 : DP nq1 ((ε₁ : ℝ) / ε₂)) (h2 : DP nq2 ((ε₃ : ℝ) / ε₄)) (nn1 : NonZeroNQ nq1) (nn2 : NonZeroNQ nq2) (nt1 : NonTopRDNQ nq1) (nt2 : NonTopRDNQ nq2) (nts1 : NonTopNQ nq1) (nts2 : NonTopNQ nq2) : DP (Compose nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by simp [Compose, RenyiDivergence, DP] intro α h3 l₁ l₂ h4 + have X := h1 + have Y := h2 simp [DP] at h1 h2 replace h1 := h1 α h3 l₁ l₂ h4 replace h2 := h2 α h3 l₁ l₂ h4 simp [RenyiDivergence] at h1 h2 rw [tsum_prod' ENNReal.summable (fun b : ℤ => ENNReal.summable)] . simp - have A : ∀ b c : ℤ, ∀ l : List T, (∑' (a : ℤ), nq1 l a * ∑' (a_1 : ℤ), if b = a ∧ c = a_1 then nq2 l a_1 else 0) = nq1 l b * nq2 l c := by - sorry conv => left right @@ -103,10 +355,10 @@ theorem DPCompose {nq1 nq2 : List T → SLang ℤ} {ε₁ ε₂ ε₃ ε₄ : intro b right intro c - rw [A] - rw [A] - rw [ENNReal.mul_rpow_of_ne_zero sorry sorry] - rw [ENNReal.mul_rpow_of_ne_zero sorry sorry] + rw [compose_sum_rw] + rw [compose_sum_rw] + rw [ENNReal.mul_rpow_of_ne_zero (nn1 l₁ b) (nn2 l₁ c)] + rw [ENNReal.mul_rpow_of_ne_zero (nn1 l₂ b) (nn2 l₂ c)] rw [mul_assoc] right rw [mul_comm] @@ -143,10 +395,107 @@ theorem DPCompose {nq1 nq2 : List T → SLang ℤ} {ε₁ ε₂ ε₃ ε₄ : rw [mul_le_mul_iff_of_pos_left] . ring_nf simp - . sorry -- easy - . sorry -- easy - . sorry -- ≠ 0 , not obvious - . sorry -- ≠ 0 + . simp + . apply lt_trans zero_lt_one h3 + . apply RenyiNoisedQueryNonZero h3 h4 X nn1 nt1 nts1 + . apply RenyiNoisedQueryNonZero h3 h4 Y nn2 nt2 nts2 + +theorem DPCompose_NonZeroNQ {nq1 nq2 : List T → SLang ℤ} (nn1 : NonZeroNQ nq1) (nn2 : NonZeroNQ nq2) : + NonZeroNQ (Compose nq1 nq2) := by + simp [NonZeroNQ] at * + intro l a b + replace nn1 := nn1 l a + replace nn2 := nn2 l b + simp [Compose] + exists a + simp + intro H + cases H + . rename_i H + contradiction + . rename_i H + contradiction + +theorem DPCompose_NonTopNQ {nq1 nq2 : List T → SLang ℤ} (nt1 : NonTopNQ nq1) (nt2 : NonTopNQ nq2) : + NonTopNQ (Compose nq1 nq2) := by + simp [NonTopNQ] at * + intro l a b + replace nt1 := nt1 l a + replace nt2 := nt2 l b + simp [Compose] + rw [compose_sum_rw] + rw [mul_eq_top] + intro H + cases H + . rename_i H + cases H + contradiction + . rename_i H + cases H + contradiction + +theorem DPCompose_NonTopRDNQ {nq1 nq2 : List T → SLang ℤ} (nt1 : NonTopRDNQ nq1) (nt2 : NonTopRDNQ nq2) (nn1 : NonZeroNQ nq1) (nn2 : NonZeroNQ nq2) : + NonTopRDNQ (Compose nq1 nq2) := by + simp [NonTopRDNQ] at * + intro α h1 l₁ l₂ h2 + replace nt1 := nt1 α h1 l₁ l₂ h2 + replace nt2 := nt2 α h1 l₁ l₂ h2 + simp [Compose] + rw [ENNReal.tsum_prod'] + simp + conv => + right + left + right + intro x + right + intro y + congr + . left + rw [compose_sum_rw] + . left + rw [compose_sum_rw] + conv => + right + left + right + intro x + right + intro y + rw [ENNReal.mul_rpow_of_ne_zero (nn1 l₁ x) (nn2 l₁ y)] + rw [ENNReal.mul_rpow_of_ne_zero (nn1 l₂ x) (nn2 l₂ y)] + rw [mul_assoc] + right + rw [mul_comm] + rw [mul_assoc] + right + rw [mul_comm] + conv => + right + left + right + intro x + right + intro y + rw [← mul_assoc] + conv => + right + left + right + intro x + rw [ENNReal.tsum_mul_left] + rw [ENNReal.tsum_mul_right] + intro H + rw [mul_eq_top] at H + cases H + . rename_i h3 + cases h3 + rename_i h4 h5 + contradiction + . rename_i h3 + cases h3 + rename_i h4 h5 + contradiction def PostProcess (nq : List T → SLang U) (pp : U → ℤ) (l : List T) : SLang ℤ := do let A ← nq l From 6938ace89479d09b52ce42b04a60a6fde46fcb00 Mon Sep 17 00:00:00 2001 From: jtristan Date: Mon, 6 May 2024 21:14:50 -0400 Subject: [PATCH 09/20] Trying to prove post processing --- SampCert/DiffPrivacy/DP.lean | 257 +++++++++++++++++++++- SampCert/DiffPrivacy/RenyiDivergence.lean | 14 +- 2 files changed, 266 insertions(+), 5 deletions(-) diff --git a/SampCert/DiffPrivacy/DP.lean b/SampCert/DiffPrivacy/DP.lean index 5dd7c2e2..74f44273 100644 --- a/SampCert/DiffPrivacy/DP.lean +++ b/SampCert/DiffPrivacy/DP.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jean-Baptiste Tristan -/ +import Mathlib.Topology.Algebra.InfiniteSum.Basic import Mathlib.Topology.Algebra.InfiniteSum.Ring import Mathlib.Algebra.Group.Basic import SampCert.DiffPrivacy.ConcentratedBound @@ -11,10 +12,15 @@ import SampCert.SLang import SampCert.Samplers.GaussianGen.Basic import SampCert.DiffPrivacy.Neighbours import SampCert.DiffPrivacy.Sensitivity +import Mathlib.MeasureTheory.MeasurableSpace.Basic +import Mathlib.MeasureTheory.Measure.Count +import Mathlib.Probability.ProbabilityMassFunction.Integrals +import Mathlib.Analysis.Convex.SpecificFunctions.Basic +import Mathlib.Analysis.Convex.Integral noncomputable section -open Classical Nat Int Real ENNReal +open Classical Nat Int Real ENNReal MeasureTheory def DP (q : List T → SLang U) (ε : ℝ) : Prop := ∀ α : ℝ, 1 < α → ∀ l₁ l₂ : List T, Neighbour l₁ l₂ → @@ -501,12 +507,255 @@ def PostProcess (nq : List T → SLang U) (pp : U → ℤ) (l : List T) : SLang let A ← nq l return pp A -theorem DPPostProcess {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP nq ((ε₁ : ℝ) / ε₂)) (pp : U → ℤ) : - DP (PostProcess nq pp) ((ε₁ : ℝ) / ε₂) := by +theorem foo (f : U → ℤ) (g : U → ENNReal) (x : ℤ) : + (∑' a : U, if x = f a then g a else 0) = ∑' a : { a | x = f a }, g a := by + have A := @tsum_split_ite U (fun a : U => x = f a) g (fun _ => 0) + simp only [decide_eq_true_eq, tsum_zero, add_zero] at A + rw [A] + have B : ↑{i | decide (x = f i) = true} = ↑{a | x = f a} := by + simp + rw [B] + +variable {T : Type} +variable [m1 : MeasurableSpace T] +variable [m2 : MeasurableSingletonClass T] +variable [m3: MeasureSpace T] + +theorem Integrable_rpow (f : T → ℝ) (nn : ∀ x : T, 0 ≤ f x) (μ : Measure T) (α : ENNReal) (mem : Memℒp f α μ) (h1 : α ≠ 0) (h2 : α ≠ ⊤) : + MeasureTheory.Integrable (fun x : T => (f x) ^ α.toReal) μ := by + have X := @MeasureTheory.Memℒp.integrable_norm_rpow T ℝ m1 μ _ f α mem h1 h2 + revert X + conv => + left + left + intro x + rw [← norm_rpow_of_nonneg (nn x)] + intro X + simp [Integrable] at * + constructor + . cases X + rename_i left right + rw [@aestronglyMeasurable_iff_aemeasurable] + apply AEMeasurable.pow_const + simp [Memℒp] at mem + cases mem + rename_i left' right' + rw [aestronglyMeasurable_iff_aemeasurable] at left' + simp [left'] + . rw [← hasFiniteIntegral_norm_iff] + simp [X] + +theorem bar (f : T → ℝ) (q : PMF T) (α : ℝ) (h : 1 < α) (h2 : ∀ x : T, 0 ≤ f x) (mem : Memℒp f (ENNReal.ofReal α) (PMF.toMeasure q)) : + ((∑' x : T, (f x) * (q x).toReal)) ^ α ≤ (∑' x : T, (f x) ^ α * (q x).toReal) := by + + conv => + left + left + right + intro x + rw [mul_comm] + rw [← smul_eq_mul] + conv => + right + right + intro x + rw [mul_comm] + rw [← smul_eq_mul] + rw [← PMF.integral_eq_tsum] + rw [← PMF.integral_eq_tsum] + + have A := @convexOn_rpow α (le_of_lt h) + have B : ContinuousOn (fun (x : ℝ) => x ^ α) (Set.Ici 0) := by + apply ContinuousOn.rpow + . exact continuousOn_id' (Set.Ici 0) + . exact continuousOn_const + . intro x h' + simp at h' + have OR : x = 0 ∨ 0 < x := by exact LE.le.eq_or_gt h' + cases OR + . rename_i h'' + subst h'' + right + apply lt_trans zero_lt_one h + . rename_i h'' + left + by_contra + rename_i h3 + subst h3 + simp at h'' + have C : @IsClosed ℝ UniformSpace.toTopologicalSpace (Set.Ici 0) := by + exact isClosed_Ici + have D := @ConvexOn.map_integral_le T ℝ m1 _ _ _ (PMF.toMeasure q) (Set.Ici 0) f (fun (x : ℝ) => x ^ α) (PMF.toMeasure.isProbabilityMeasure q) A B C + simp at D + apply D + . exact MeasureTheory.ae_of_all (PMF.toMeasure q) h2 + . apply MeasureTheory.Memℒp.integrable _ mem + rw [one_le_ofReal] + apply le_of_lt h + . rw [Function.comp_def] + have X : ENNReal.ofReal α ≠ 0 := by + simp + apply lt_trans zero_lt_one h + have Y : ENNReal.ofReal α ≠ ⊤ := by + simp + have Z := @Integrable_rpow T m1 f h2 (PMF.toMeasure q) (ENNReal.ofReal α) mem X Y + rw [toReal_ofReal] at Z + . exact Z + . apply le_of_lt + apply lt_trans zero_lt_one h + . have X : ENNReal.ofReal α ≠ 0 := by + simp + apply lt_trans zero_lt_one h + have Y : ENNReal.ofReal α ≠ ⊤ := by + simp + have Z := @Integrable_rpow T m1 f h2 (PMF.toMeasure q) (ENNReal.ofReal α) mem X Y + rw [toReal_ofReal] at Z + . exact Z + . apply le_of_lt + apply lt_trans zero_lt_one h + . apply MeasureTheory.Memℒp.integrable _ mem + rw [one_le_ofReal] + apply le_of_lt h + + +theorem quux (f : U → ℤ) (g : U → ENNReal) : + (∑' (x : ℤ) (i : ↑{a | x = f a}), g i) + = ∑' i : U, g i := by + sorry + +variable {U : Type} +variable [m2 : MeasurableSpace U] [m2' : MeasurableSingletonClass U] + +def SLangToPMF (q : SLang U) : PMF U := sorry + +theorem DPPostProcess {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP nq ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (f : U → ℤ) : + DP (PostProcess nq f) ((ε₁ : ℝ) / ε₂) := by simp [PostProcess, DP, RenyiDivergence] intro α h1 l₁ l₂ h2 simp [DP, RenyiDivergence] at h replace h := h α h1 l₁ l₂ h2 - sorry + + have X₁ : ∀ (x : U), nq l₂ x ≠ 0 := by + intro x + apply nn l₂ x + have X₂ : ∀ (x : U), nq l₂ x ≠ ⊤ := by + intro x + apply nts l₂ x + have X := @RenyiDivergenceExpectation _ (nq l₁) (nq l₂) _ h1 X₁ X₂ + rw [X] at h + clear X + have Y : ∀ (a : U), (nq l₁ a / nq l₂ a) ^ α * nq l₂ a ≠ ⊤ := by + intro a + rw [ne_iff_lt_or_gt] + left + rw [mul_lt_top_iff] + left + constructor + . sorry + . rw [lt_top_iff_ne_top] + apply nts + rw [ENNReal.tsum_toReal_eq Y] at h + simp only [toReal_mul] at h + revert h + conv => + left + left + right + right + right + intro a + rw [← toReal_rpow] + intro h + + have R : (∀ (x : U), 0 ≤ (fun a => (nq l₁ a / nq l₂ a).toReal) x) := sorry + have Z := @bar U m2 m2' (fun a : U => (nq l₁ a / nq l₂ a).toReal) (SLangToPMF (nq l₂)) α h1 R + simp at Z + + + + rw [RenyiDivergenceExpectation _ _ h1 sorry sorry] + conv => + left + right + right + right + right + intro x + rw [foo f (nq l₁) x] + rw [foo f (nq l₂) x] + --rw [RenyiDivergenceExpectation _ _ h1 sorry sorry] + + -- ∑' (a : ↑{a | f a = x}), nq l₂ ↑a is pushforward + -- we start from the goal, so h is outside, and we want to put it inside + conv => + left + right + right + right + right + intro x + left + left + rw [division_def] + rw [← ENNReal.tsum_mul_right] + right + intro i + left + rw [← @mul_one _ _ (nq l₁ ↑i)] + right + rw [← @ENNReal.mul_inv_cancel (nq l₂ i) (nn l₂ i) (nts l₂ i)] + rw [mul_comm] + conv => + left + right + ring_nf + + -- not true, should be ≤, Jensen's inequality, exploring for now + have A : ∀ x, (∑' (i : ↑{a | x = f a}), nq l₁ ↑i * (nq l₂ ↑i)⁻¹ * nq l₂ ↑i * (∑' (a : ↑{a | x = f a}), nq l₂ ↑a)⁻¹) ^ α + = (∑' (i : ↑{a | x = f a}), (nq l₁ ↑i * (nq l₂ ↑i)⁻¹) ^ α * nq l₂ ↑i * (∑' (a : ↑{a | x = f a}), nq l₂ ↑a)⁻¹) := sorry + + conv => + left + right + right + right + right + intro x + rw [A] + + conv => + left + right + right + right + right + intro x + rw [← ENNReal.tsum_mul_right] + right + intro i + rw [mul_assoc] + right + rw [@ENNReal.inv_mul_cancel (∑' (a : ↑{a | x = f a}), nq l₂ ↑a) sorry sorry] + conv => + left + right + ring_nf + + have B : (∑' (x : ℤ) (i : ↑{a | x = f a}), (nq l₁ ↑i * (nq l₂ ↑i)⁻¹) ^ α * nq l₂ ↑i) + = ∑' i : U, (nq l₁ ↑i * (nq l₂ ↑i)⁻¹) ^ α * nq l₂ ↑i := by + sorry + + rw [B] + clear A B + conv => + left + right + right + right + right + intro i + rw [← division_def] + rw [← RenyiDivergenceExpectation (nq l₁) (nq l₂) h1] + exact h end SLang diff --git a/SampCert/DiffPrivacy/RenyiDivergence.lean b/SampCert/DiffPrivacy/RenyiDivergence.lean index a360524f..831455bc 100644 --- a/SampCert/DiffPrivacy/RenyiDivergence.lean +++ b/SampCert/DiffPrivacy/RenyiDivergence.lean @@ -13,4 +13,16 @@ open Real ENNReal variable {T : Type} noncomputable def RenyiDivergence (p q : T → ENNReal) (α : ℝ) : ℝ := - (1 / (α - 1)) * Real.log (∑' x : T, (p x)^α * (q x)^(1 - α)).toReal + (α - 1)⁻¹ * Real.log (∑' x : T, (p x)^α * (q x)^(1 - α)).toReal + +theorem RenyiDivergenceExpectation (p q : T → ENNReal) {α : ℝ} (h : 1 < α) (h1 : ∀ x : T, q x ≠ 0) (h2 : ∀ x : T, q x ≠ ⊤) : + (α - 1)⁻¹ * Real.log ((∑' x : T, (p x)^α * (q x)^(1 - α))).toReal = (α - 1)⁻¹ * Real.log (∑' x : T, (p x / q x)^α * (q x)).toReal := by + congr 4 + ext x + rw [ENNReal.rpow_sub] + . rw [← ENNReal.mul_comm_div] + rw [← ENNReal.div_rpow_of_nonneg] + . rw [ENNReal.rpow_one] + . apply le_of_lt (lt_trans Real.zero_lt_one h ) + . apply h1 x + . apply h2 x From 41c8d85d606b9866bf54767c5b3edd7183e5cc9f Mon Sep 17 00:00:00 2001 From: jtristan Date: Wed, 8 May 2024 09:02:34 -0400 Subject: [PATCH 10/20] complete proof sketch for postprocessing with issues for elements of the image of postprocessor that have probability 0 --- SampCert/DiffPrivacy/DP.lean | 259 +++++++++++++++------- SampCert/DiffPrivacy/RenyiDivergence.lean | 12 + 2 files changed, 190 insertions(+), 81 deletions(-) diff --git a/SampCert/DiffPrivacy/DP.lean b/SampCert/DiffPrivacy/DP.lean index 74f44273..d02e751c 100644 --- a/SampCert/DiffPrivacy/DP.lean +++ b/SampCert/DiffPrivacy/DP.lean @@ -617,7 +617,6 @@ theorem bar (f : T → ℝ) (q : PMF T) (α : ℝ) (h : 1 < α) (h2 : ∀ x : T, rw [one_le_ofReal] apply le_of_lt h - theorem quux (f : U → ℤ) (g : U → ENNReal) : (∑' (x : ℤ) (i : ↑{a | x = f a}), g i) = ∑' i : U, g i := by @@ -626,14 +625,37 @@ theorem quux (f : U → ℤ) (g : U → ENNReal) : variable {U : Type} variable [m2 : MeasurableSpace U] [m2' : MeasurableSingletonClass U] -def SLangToPMF (q : SLang U) : PMF U := sorry +def δ (nq : SLang U) (f : U → ℤ) (a : ℤ) : {n : U | a = f n} → ENNReal := fun x : {n : U | a = f n} => nq x * (∑' (x : {n | a = f n}), nq x)⁻¹ -theorem DPPostProcess {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP nq ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (f : U → ℤ) : - DP (PostProcess nq f) ((ε₁ : ℝ) / ε₂) := by - simp [PostProcess, DP, RenyiDivergence] - intro α h1 l₁ l₂ h2 - simp [DP, RenyiDivergence] at h - replace h := h α h1 l₁ l₂ h2 +theorem δ_normalizes (nq : SLang U) (f : U → ℤ) (a : ℤ) : + HasSum (δ nq f a) 1 := by + rw [Summable.hasSum_iff ENNReal.summable] + unfold δ + rw [ENNReal.tsum_mul_right] + rw [ENNReal.mul_inv_cancel] + . simp + sorry -- ∃ x, a = f x ∧ ¬nq x = 0 + . sorry -- ∑' (i : ↑{n | a = f n}), nq ↑i ≠ ⊤ + +def δpmf (nq : SLang U) (f : U → ℤ) (a : ℤ) : PMF {n : U | a = f n} := + ⟨ δ nq f a , δ_normalizes nq f a ⟩ + +theorem δpmf_conv (nq : List T → SLang U) (a : ℤ) (x : {n | a = f n}) : + nq l₂ x * (∑' (x : {n | a = f n}), nq l₂ x)⁻¹ = (δpmf (nq l₂) f a) x := by + simp [δpmf] + conv => + right + left + left + +theorem preparation {nq : List T → SLang U} (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (f : U → ℤ) {α : ℝ} (h1 : 1 < α) (mem : ∀ (a : ℤ), + Memℒp (fun a_1 : {n | a = f n} => (nq l₁ ↑a_1 / nq l₂ ↑a_1).toReal) (ENNReal.ofReal α) (PMF.toMeasure (δpmf (nq l₂) f a))): + (∑' (x : U), nq l₁ x ^ α * nq l₂ x ^ (1 - α)).toReal + ≥ + (∑' (x : ℤ), + (∑' (x_1 : ↑{n | x = f n}), nq l₂ ↑x_1).toReal * + (∑' (x_1 : ↑{n | x = f n}), (nq l₁ ↑x_1 / nq l₂ ↑x_1).toReal * ((δpmf (nq l₂) f x) x_1).toReal) ^ α) + := by have X₁ : ∀ (x : U), nq l₂ x ≠ 0 := by intro x @@ -641,121 +663,196 @@ theorem DPPostProcess {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP nq have X₂ : ∀ (x : U), nq l₂ x ≠ ⊤ := by intro x apply nts l₂ x - have X := @RenyiDivergenceExpectation _ (nq l₁) (nq l₂) _ h1 X₁ X₂ - rw [X] at h - clear X - have Y : ∀ (a : U), (nq l₁ a / nq l₂ a) ^ α * nq l₂ a ≠ ⊤ := by - intro a - rw [ne_iff_lt_or_gt] - left - rw [mul_lt_top_iff] - left - constructor - . sorry - . rw [lt_top_iff_ne_top] - apply nts - rw [ENNReal.tsum_toReal_eq Y] at h - simp only [toReal_mul] at h - revert h + + rw [@RenyiDivergenceExpectation _ (nq l₁) (nq l₂) _ h1 X₁ X₂] + clear X₁ X₂ + + rw [← quux f (fun x => (nq l₁ x / nq l₂ x) ^ α * nq l₂ x)] + + have P1 : ∀ (a : ℤ), ∑' (i : ↑{a_1 | a = f a_1}), (nq l₁ ↑i / nq l₂ ↑i) ^ α * nq l₂ ↑i ≠ ⊤ := sorry + have P2 : ∀ (a : ℤ), ∀ (a_1 : ↑{a_1 | a = f a_1}), (nq l₁ ↑a_1 / nq l₂ ↑a_1) ^ α * nq l₂ ↑a_1 ≠ ⊤ := sorry + + rw [ENNReal.tsum_toReal_eq P1] conv => - left left right - right - right intro a - rw [← toReal_rpow] - intro h - - have R : (∀ (x : U), 0 ≤ (fun a => (nq l₁ a / nq l₂ a).toReal) x) := sorry - have Z := @bar U m2 m2' (fun a : U => (nq l₁ a / nq l₂ a).toReal) (SLangToPMF (nq l₂)) α h1 R - simp at Z + rw [ENNReal.tsum_toReal_eq (P2 a)] + right + intro b + rw [toReal_mul] + -- have P3 : (∀ (x : U), 0 ≤ (fun a => (nq l₁ a / nq l₂ a).toReal) x) := by + -- intro x + -- simp only [toReal_nonneg] + let κ (a : ℤ) := ∑' x : {n : U | a = f n}, nq l₂ x + have P4 : ∀ n : ℤ, (κ n / κ n).toReal = 1 := by + intro n + rw [division_def] + rw [ENNReal.mul_inv_cancel] + . rw [one_toReal] + . simp [κ] + sorry -- ∃ x, n = f x ∧ ¬nq l₂ x = 0 + . simp only [κ] + sorry -- ∑' (x : ↑{n_1 | n = f n_1}), nq l₂ ↑x ≠ ⊤ - rw [RenyiDivergenceExpectation _ _ h1 sorry sorry] conv => left right + intro a + rw [← mul_one (∑' (b : ↑{a_1 | a = f a_1}), ((nq l₁ ↑b / nq l₂ ↑b) ^ α).toReal * (nq l₂ ↑b).toReal)] right - right - right - intro x - rw [foo f (nq l₁) x] - rw [foo f (nq l₂) x] - --rw [RenyiDivergenceExpectation _ _ h1 sorry sorry] + rw [← (P4 a)] + rw [toReal_div] - -- ∑' (a : ↑{a | f a = x}), nq l₂ ↑a is pushforward - -- we start from the goal, so h is outside, and we want to put it inside conv => left right - right - right - right - intro x - left - left + intro a rw [division_def] - rw [← ENNReal.tsum_mul_right] - right - intro i - left - rw [← @mul_one _ _ (nq l₁ ↑i)] right - rw [← @ENNReal.mul_inv_cancel (nq l₂ i) (nn l₂ i) (nts l₂ i)] rw [mul_comm] + conv => left right - ring_nf + intro a + rw [← mul_assoc] + rw [← tsum_mul_right] + left + right + intro x + rw [mul_assoc] + right + rw [← toReal_inv] + rw [← toReal_mul] + + simp only [κ] - -- not true, should be ≤, Jensen's inequality, exploring for now - have A : ∀ x, (∑' (i : ↑{a | x = f a}), nq l₁ ↑i * (nq l₂ ↑i)⁻¹ * nq l₂ ↑i * (∑' (a : ↑{a | x = f a}), nq l₂ ↑a)⁻¹) ^ α - = (∑' (i : ↑{a | x = f a}), (nq l₁ ↑i * (nq l₂ ↑i)⁻¹) ^ α * nq l₂ ↑i * (∑' (a : ↑{a | x = f a}), nq l₂ ↑a)⁻¹) := sorry + have P5 : ∀ a : ℤ, ∀ (x : ↑{n | a = f n}), 0 ≤ (nq l₁ ↑x / nq l₂ ↑x).toReal := by + intro a x + simp only [toReal_nonneg] + + have A := fun a : ℤ => @bar {n : U | a = f n} Subtype.instMeasurableSpace Subtype.instMeasurableSingletonClass (fun a => (nq l₁ a / nq l₂ a).toReal) (δpmf (nq l₂) f a) α h1 (P5 a) (mem a) + have P6 : ∀ a, 0 ≤ (∑' (x : ↑{n | a = f n}), nq l₂ ↑x).toReal := by + intro a + simp only [toReal_nonneg] + have A' := fun a : ℤ => mul_le_mul_of_nonneg_left (A a) (P6 a) + clear A P6 + + have Y₁ : Summable fun i => + (∑' (x : ↑{n | i = f n}), nq l₂ ↑x).toReal * + (∑' (x : ↑{n | i = f n}), (nq l₁ ↑x / nq l₂ ↑x).toReal * ((δpmf (nq l₂) f i) x).toReal) ^ α := sorry + + have Y₂ : Summable fun i => + (∑' (x : ↑{n | i = f n}), nq l₂ ↑x).toReal * + ∑' (x : ↑{n | i = f n}), (nq l₁ ↑x / nq l₂ ↑x).toReal ^ α * ((δpmf (nq l₂) f i) x).toReal := sorry + + have B := tsum_le_tsum A' Y₁ Y₂ + clear A' + clear Y₁ Y₂ conv => left right - right - right + intro a + left right intro x - rw [A] + rw [δpmf_conv] conv => left right - right + intro a + rw [mul_comm] right right intro x - rw [← ENNReal.tsum_mul_right] - right - intro i - rw [mul_assoc] - right - rw [@ENNReal.inv_mul_cancel (∑' (a : ↑{a | x = f a}), nq l₂ ↑a) sorry sorry] - conv => left - right - ring_nf + rw [← toReal_rpow] - have B : (∑' (x : ℤ) (i : ↑{a | x = f a}), (nq l₁ ↑i * (nq l₂ ↑i)⁻¹) ^ α * nq l₂ ↑i) - = ∑' i : U, (nq l₁ ↑i * (nq l₂ ↑i)⁻¹) ^ α * nq l₂ ↑i := by - sorry + simp only [ge_iff_le] + exact B - rw [B] - clear A B +theorem DPPostProcess {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP nq ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (f : U → ℤ) : + DP (PostProcess nq f) ((ε₁ : ℝ) / ε₂) := by + simp [PostProcess, DP, RenyiDivergence] + intro α h1 l₁ l₂ h2 + simp [DP, RenyiDivergence] at h + replace h := h α h1 l₁ l₂ h2 + + apply le_trans _ h + clear h + + have A : 0 ≤ (α - 1)⁻¹ := by + simp + apply le_of_lt h1 + apply mul_le_mul_of_nonneg_left _ A + clear A + have B : 0 < + (∑' (x : ℤ), + (∑' (a : U), if x = f a then nq l₁ a else 0) ^ α * (∑' (a : U), if x = f a then nq l₂ a else 0) ^ (1 - α)).toReal := sorry + apply log_le_log B + clear B + + have C := @preparation _ _ _ _ l₁ l₂ nq nn nt nts f α h1 sorry + + apply le_trans _ C + clear C + apply le_of_eq + + rw [ENNReal.tsum_toReal_eq sorry] + apply tsum_congr + intro b + rw [foo] + rw [foo] + rw [toReal_mul] conv => - left right right + left right + intro x + rw [division_def] + rw [toReal_mul] + rw [← δpmf_conv] + rw [toReal_mul] + rw [mul_assoc] right - intro i - rw [← division_def] - rw [← RenyiDivergenceExpectation (nq l₁) (nq l₂) h1] - exact h + rw [← mul_assoc] + left + rw [← toReal_mul] + rw [ENNReal.inv_mul_cancel (nn l₂ x) (nts l₂ x)] + rw [one_toReal] + simp only [one_mul] + + rw [tsum_mul_right] + rw [Real.mul_rpow sorry sorry] + + rw [ENNReal.rpow_sub _ _ sorry sorry] + rw [ENNReal.rpow_one] + rw [division_def] + rw [toReal_mul] + rw [← toReal_rpow] + rw [← mul_assoc] + rw [← mul_assoc] + congr 1 + . rw [mul_comm] + congr 2 + rw [ENNReal.tsum_toReal_eq sorry] + . rw [toReal_rpow] + congr 1 + rw [ENNReal.inv_rpow] + + + + + + + + + end SLang diff --git a/SampCert/DiffPrivacy/RenyiDivergence.lean b/SampCert/DiffPrivacy/RenyiDivergence.lean index 831455bc..3d8a06a9 100644 --- a/SampCert/DiffPrivacy/RenyiDivergence.lean +++ b/SampCert/DiffPrivacy/RenyiDivergence.lean @@ -16,6 +16,18 @@ noncomputable def RenyiDivergence (p q : T → ENNReal) (α : ℝ) : ℝ := (α - 1)⁻¹ * Real.log (∑' x : T, (p x)^α * (q x)^(1 - α)).toReal theorem RenyiDivergenceExpectation (p q : T → ENNReal) {α : ℝ} (h : 1 < α) (h1 : ∀ x : T, q x ≠ 0) (h2 : ∀ x : T, q x ≠ ⊤) : + (∑' x : T, (p x)^α * (q x)^(1 - α)) = ∑' x : T, (p x / q x)^α * (q x) := by + congr 4 + ext x + rw [ENNReal.rpow_sub] + . rw [← ENNReal.mul_comm_div] + rw [← ENNReal.div_rpow_of_nonneg] + . rw [ENNReal.rpow_one] + . apply le_of_lt (lt_trans Real.zero_lt_one h ) + . apply h1 x + . apply h2 x + +theorem RenyiDivergenceExpectation' (p q : T → ENNReal) {α : ℝ} (h : 1 < α) (h1 : ∀ x : T, q x ≠ 0) (h2 : ∀ x : T, q x ≠ ⊤) : (α - 1)⁻¹ * Real.log ((∑' x : T, (p x)^α * (q x)^(1 - α))).toReal = (α - 1)⁻¹ * Real.log (∑' x : T, (p x / q x)^α * (q x)).toReal := by congr 4 ext x From e6a761236d3f2f018707a38a3423719d524fd74a Mon Sep 17 00:00:00 2001 From: jtristan Date: Thu, 9 May 2024 15:00:11 -0400 Subject: [PATCH 11/20] proved the part postprocessing on Lp --- SampCert/DiffPrivacy/DP.lean | 464 ++++++++++++++++++++++------------- 1 file changed, 291 insertions(+), 173 deletions(-) diff --git a/SampCert/DiffPrivacy/DP.lean b/SampCert/DiffPrivacy/DP.lean index d02e751c..d2930097 100644 --- a/SampCert/DiffPrivacy/DP.lean +++ b/SampCert/DiffPrivacy/DP.lean @@ -20,7 +20,7 @@ import Mathlib.Analysis.Convex.Integral noncomputable section -open Classical Nat Int Real ENNReal MeasureTheory +open Classical Nat Int Real ENNReal MeasureTheory Measure def DP (q : List T → SLang U) (ε : ℝ) : Prop := ∀ α : ℝ, 1 < α → ∀ l₁ l₂ : List T, Neighbour l₁ l₂ → @@ -618,12 +618,14 @@ theorem bar (f : T → ℝ) (q : PMF T) (α : ℝ) (h : 1 < α) (h2 : ∀ x : T, apply le_of_lt h theorem quux (f : U → ℤ) (g : U → ENNReal) : - (∑' (x : ℤ) (i : ↑{a | x = f a}), g i) + (∑' (x : ℤ), if {a : U | x = f a} = {} then 0 else ∑'(i : {a : U | x = f a}), g i) = ∑' i : U, g i := by sorry variable {U : Type} variable [m2 : MeasurableSpace U] [m2' : MeasurableSingletonClass U] +variable [count : Countable U] +variable [disc : DiscreteMeasurableSpace U] def δ (nq : SLang U) (f : U → ℤ) (a : ℤ) : {n : U | a = f n} → ENNReal := fun x : {n : U | a = f n} => nq x * (∑' (x : {n | a = f n}), nq x)⁻¹ @@ -648,211 +650,327 @@ theorem δpmf_conv (nq : List T → SLang U) (a : ℤ) (x : {n | a = f n}) : left left -theorem preparation {nq : List T → SLang U} (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (f : U → ℤ) {α : ℝ} (h1 : 1 < α) (mem : ∀ (a : ℤ), - Memℒp (fun a_1 : {n | a = f n} => (nq l₁ ↑a_1 / nq l₂ ↑a_1).toReal) (ENNReal.ofReal α) (PMF.toMeasure (δpmf (nq l₂) f a))): - (∑' (x : U), nq l₁ x ^ α * nq l₂ x ^ (1 - α)).toReal - ≥ - (∑' (x : ℤ), - (∑' (x_1 : ↑{n | x = f n}), nq l₂ ↑x_1).toReal * - (∑' (x_1 : ↑{n | x = f n}), (nq l₁ ↑x_1 / nq l₂ ↑x_1).toReal * ((δpmf (nq l₂) f x) x_1).toReal) ^ α) - := by +theorem δpmf_conv' (nq : List T → SLang U) (l₂ : List T) (f : U → ℤ) (a : ℤ) : + (fun x : {n | a = f n} => nq l₂ x * (∑' (x : {n | a = f n}), nq l₂ x)⁻¹) = (δpmf (nq l₂) f a) := by + ext x + rw [δpmf_conv] - have X₁ : ∀ (x : U), nq l₂ x ≠ 0 := by - intro x - apply nn l₂ x - have X₂ : ∀ (x : U), nq l₂ x ≠ ⊤ := by - intro x - apply nts l₂ x - - rw [@RenyiDivergenceExpectation _ (nq l₁) (nq l₂) _ h1 X₁ X₂] - clear X₁ X₂ - - rw [← quux f (fun x => (nq l₁ x / nq l₂ x) ^ α * nq l₂ x)] - - have P1 : ∀ (a : ℤ), ∑' (i : ↑{a_1 | a = f a_1}), (nq l₁ ↑i / nq l₂ ↑i) ^ α * nq l₂ ↑i ≠ ⊤ := sorry - have P2 : ∀ (a : ℤ), ∀ (a_1 : ↑{a_1 | a = f a_1}), (nq l₁ ↑a_1 / nq l₂ ↑a_1) ^ α * nq l₂ ↑a_1 ≠ ⊤ := sorry - - rw [ENNReal.tsum_toReal_eq P1] - conv => - left - right - intro a - rw [ENNReal.tsum_toReal_eq (P2 a)] - right - intro b - rw [toReal_mul] - - -- have P3 : (∀ (x : U), 0 ≤ (fun a => (nq l₁ a / nq l₂ a).toReal) x) := by - -- intro x - -- simp only [toReal_nonneg] - - let κ (a : ℤ) := ∑' x : {n : U | a = f n}, nq l₂ x - have P4 : ∀ n : ℤ, (κ n / κ n).toReal = 1 := by - intro n - rw [division_def] - rw [ENNReal.mul_inv_cancel] - . rw [one_toReal] - . simp [κ] - sorry -- ∃ x, n = f x ∧ ¬nq l₂ x = 0 - . simp only [κ] - sorry -- ∑' (x : ↑{n_1 | n = f n_1}), nq l₂ ↑x ≠ ⊤ - - conv => - left - right - intro a - rw [← mul_one (∑' (b : ↑{a_1 | a = f a_1}), ((nq l₁ ↑b / nq l₂ ↑b) ^ α).toReal * (nq l₂ ↑b).toReal)] - right - rw [← (P4 a)] - rw [toReal_div] - - conv => - left - right - intro a - rw [division_def] - right - rw [mul_comm] - - conv => - left - right - intro a - rw [← mul_assoc] - rw [← tsum_mul_right] - left - right - intro x - rw [mul_assoc] - right - rw [← toReal_inv] - rw [← toReal_mul] - - simp only [κ] +theorem witness {f : U → ℤ} {i : ℤ} (h : ¬{b | i = f b} = ∅) : + ∃ x : U, i = f x := by + rw [← nonempty_subtype] + exact Set.nonempty_iff_ne_empty'.mpr h - have P5 : ∀ a : ℤ, ∀ (x : ↑{n | a = f n}), 0 ≤ (nq l₁ ↑x / nq l₂ ↑x).toReal := by - intro a x - simp only [toReal_nonneg] - - have A := fun a : ℤ => @bar {n : U | a = f n} Subtype.instMeasurableSpace Subtype.instMeasurableSingletonClass (fun a => (nq l₁ a / nq l₂ a).toReal) (δpmf (nq l₂) f a) α h1 (P5 a) (mem a) - have P6 : ∀ a, 0 ≤ (∑' (x : ↑{n | a = f n}), nq l₂ ↑x).toReal := by - intro a - simp only [toReal_nonneg] - have A' := fun a : ℤ => mul_le_mul_of_nonneg_left (A a) (P6 a) - clear A P6 - - have Y₁ : Summable fun i => - (∑' (x : ↑{n | i = f n}), nq l₂ ↑x).toReal * - (∑' (x : ↑{n | i = f n}), (nq l₁ ↑x / nq l₂ ↑x).toReal * ((δpmf (nq l₂) f i) x).toReal) ^ α := sorry - - have Y₂ : Summable fun i => - (∑' (x : ↑{n | i = f n}), nq l₂ ↑x).toReal * - ∑' (x : ↑{n | i = f n}), (nq l₁ ↑x / nq l₂ ↑x).toReal ^ α * ((δpmf (nq l₂) f i) x).toReal := sorry - - have B := tsum_le_tsum A' Y₁ Y₂ - clear A' - clear Y₁ Y₂ - - conv => - left - right - intro a - left - right - intro x - rw [δpmf_conv] - - conv => - left - right - intro a - rw [mul_comm] - right - right - intro x - left - rw [← toReal_rpow] +theorem norm_simplify (x : ENNReal) (h : x ≠ ⊤) : + @nnnorm ℝ SeminormedAddGroup.toNNNorm x.toReal = x := by + simp [nnnorm] + cases x + . contradiction + . rename_i v + simp + rfl - simp only [ge_iff_le] - exact B +theorem RD1 (p q : T → ENNReal) (α : ℝ) (h : 1 < α) (RD : ∑' (x : T), p x ^ α * q x ^ (1 - α) ≠ ⊤) (nz : ∀ x : T, q x ≠ 0) (nt : ∀ x : T, q x ≠ ⊤) : + ∑' (x : T), (p x / q x) ^ α * q x ≠ ⊤ := by + rw [← RenyiDivergenceExpectation p q h nz nt] + trivial -theorem DPPostProcess {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP nq ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (f : U → ℤ) : +theorem DPPostProcess_alt1 {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP nq ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (f : U → ℤ) : DP (PostProcess nq f) ((ε₁ : ℝ) / ε₂) := by simp [PostProcess, DP, RenyiDivergence] intro α h1 l₁ l₂ h2 simp [DP, RenyiDivergence] at h replace h := h α h1 l₁ l₂ h2 + -- Part 1, removing fluff + apply le_trans _ h clear h + -- remove the α scaling have A : 0 ≤ (α - 1)⁻¹ := by simp apply le_of_lt h1 apply mul_le_mul_of_nonneg_left _ A clear A + + -- remove the log have B : 0 < (∑' (x : ℤ), (∑' (a : U), if x = f a then nq l₁ a else 0) ^ α * (∑' (a : U), if x = f a then nq l₂ a else 0) ^ (1 - α)).toReal := sorry apply log_le_log B clear B - have C := @preparation _ _ _ _ l₁ l₂ nq nn nt nts f α h1 sorry + have RDConvegence : ∑' (x : U), nq l₁ x ^ α * nq l₂ x ^ (1 - α) ≠ ⊤ := by + simp [NonTopRDNQ] at nt + have nt := nt α h1 l₁ l₂ h2 + trivial - apply le_trans _ C - clear C - apply le_of_eq + apply toReal_mono RDConvegence - rw [ENNReal.tsum_toReal_eq sorry] - apply tsum_congr - intro b - rw [foo] - rw [foo] - rw [toReal_mul] - conv => - right - right - left - right - intro x - rw [division_def] - rw [toReal_mul] - rw [← δpmf_conv] - rw [toReal_mul] - rw [mul_assoc] - right - rw [← mul_assoc] - left - rw [← toReal_mul] - rw [ENNReal.inv_mul_cancel (nn l₂ x) (nts l₂ x)] - rw [one_toReal] - simp only [one_mul] - - rw [tsum_mul_right] - rw [Real.mul_rpow sorry sorry] - - rw [ENNReal.rpow_sub _ _ sorry sorry] - rw [ENNReal.rpow_one] - rw [division_def] - rw [toReal_mul] - rw [← toReal_rpow] - rw [← mul_assoc] - rw [← mul_assoc] - congr 1 - . rw [mul_comm] - congr 2 - rw [ENNReal.tsum_toReal_eq sorry] - . rw [toReal_rpow] - congr 1 - rw [ENNReal.inv_rpow] + -- Rewrite as cascading expectations + rw [@RenyiDivergenceExpectation _ (nq l₁) (nq l₂) _ h1 (nn l₂) (nts l₂)] + -- Shuffle the sum + rw [← quux f (fun x => (nq l₁ x / nq l₂ x) ^ α * nq l₂ x)] + apply ENNReal.tsum_le_tsum + intro i + have S2 : (∑' (a : ↑{n | i = f n}), nq l₁ ↑a / nq l₂ ↑a * (δpmf (nq l₂) f i) a) ^ α ≠ ⊤ := by + sorry + have S1 : ∀ (a : ↑{n | i = f n}), nq l₁ ↑a / nq l₂ ↑a * (δpmf (nq l₂) f i) a ≠ ⊤ := by + sorry + have S3 : ∑' (a : ↑{n | i = f n}), (nq l₁ ↑a / nq l₂ ↑a) ^ α * (δpmf (nq l₂) f i) a ≠ ⊤ := by + sorry + have S4 : ∀ (a : ↑{n | i = f n}), (nq l₁ ↑a / nq l₂ ↑a) ^ α * (δpmf (nq l₂) f i) a ≠ ⊤ := by + intro a + apply ENNReal.ne_top_of_tsum_ne_top S3 + have MasterRW : ∀ l : List T, ∑' (a : ↑{a | i = f a}), nq l ↑a ≠ ⊤ := sorry + -- Get rid of elements with probability 0 in the pushforward + split + . rename_i empty + rw [foo] + have ZE : (∑' (x_1 : ↑{n | i = f n}), nq l₁ ↑x_1) = 0 := by + simp + intro a H + have I₁ : a ∈ {b | i = f b} := by + simp [H] + have I2 : {b | i = f b} ≠ ∅ := by + apply ne_of_mem_of_not_mem' I₁ + simp + contradiction + rw [ZE] + simp only [toReal_mul, zero_toReal, ge_iff_le] + rw [ENNReal.zero_rpow_of_pos] + . simp + . apply lt_trans zero_lt_one h1 + + -- Part 2: apply Jensen's inequality + . rename_i NotEmpty + rw [foo] + rw [foo] + + -- Introduce Q(f⁻¹ i) + let κ := ∑' x : {n : U | i = f n}, nq l₂ x + have P4 : κ / κ = 1 := by + rw [division_def] + rw [ENNReal.mul_inv_cancel] + . simp [κ] -- Use here for δ normalization + have T := witness NotEmpty + cases T + rename_i z w + exists z + constructor + . trivial + . apply nn l₂ + . simp only [κ] + apply MasterRW l₂ + conv => + right + right + intro a + rw [← mul_one ((nq l₁ ↑a / nq l₂ ↑a) ^ α * nq l₂ ↑a)] + right + rw [← P4] + clear P4 + simp only [κ] + + conv => + right + right + intro a + right + rw [division_def] + rw [mul_comm] + + conv => + right + right + intro a + rw [← mul_assoc] + + rw [ENNReal.tsum_mul_right] + + -- Jensen's inequality + + have P5 : ∀ (x : ↑{n | i = f n}), 0 ≤ (fun a => (nq l₁ ↑a / nq l₂ ↑a).toReal) x := by + intro x + simp only [toReal_nonneg] + + -- stronglyMeasurable_iff_measurable + -- measurable_discrete + have XXX : @Memℒp ℝ Real.normedAddCommGroup (↑{n | i = f n}) Subtype.instMeasurableSpace (fun a => (nq l₁ ↑a / nq l₂ ↑a).toReal) + (ENNReal.ofReal α) (PMF.toMeasure (δpmf (nq l₂) f i)) := by + simp [Memℒp] + constructor + . apply MeasureTheory.StronglyMeasurable.aestronglyMeasurable + apply Measurable.stronglyMeasurable + apply Measurable.ennreal_toReal + conv => + right + intro x + rw [division_def] + apply Measurable.mul + . -- MeasurableSingletonClass.toDiscreteMeasurableSpace + apply measurable_discrete + . apply Measurable.inv + apply measurable_discrete + . simp [snorm] + split + . simp + . simp [snorm'] + rw [MeasureTheory.lintegral_countable'] -- Uses countable + rw [toReal_ofReal (le_of_lt (lt_trans zero_lt_one h1))] + have OTHER : ∀ a, nq l₁ a / nq l₂ a ≠ ⊤ := by + intro a + rw [division_def] + rw [ne_iff_lt_or_gt] + left + rw [mul_lt_top_iff] + left + constructor + . exact Ne.lt_top' (id (Ne.symm (nts l₁ a))) + . simp + exact pos_iff_ne_zero.mpr (nn l₂ a) + + conv => + left + left + right + intro a + rw [norm_simplify _ (OTHER a)] + have Z : 0 < α⁻¹ := by + simp + apply lt_trans zero_lt_one h1 + rw [rpow_lt_top_iff_of_pos Z] + conv => + left + right + intro a + rw [PMF.toMeasure_apply_singleton _ _ (measurableSet_singleton a)] + + apply Ne.lt_top' (id (Ne.symm _)) + apply S3 + + + have Jensen's := @bar {n : U | i = f n} Subtype.instMeasurableSpace Subtype.instMeasurableSingletonClass (fun a => (nq l₁ a / nq l₂ a).toReal) (δpmf (nq l₂) f i) α h1 P5 XXX + clear P5 + + have P6 : 0 ≤ (∑' (x : ↑{n | i = f n}), nq l₂ ↑x).toReal := by + simp only [toReal_nonneg] + have A' := mul_le_mul_of_nonneg_left Jensen's P6 + clear Jensen's P6 + + conv => + right + rw [mul_comm] + right + right + intro a + rw [mul_assoc] + rw [δpmf_conv] + + -- Here + + replace A' := ofReal_le_ofReal A' + rw [ofReal_mul toReal_nonneg] at A' + rw [ofReal_mul toReal_nonneg] at A' + rw [ofReal_toReal_eq_iff.2 (MasterRW l₂)] at A' + simp only at A' + + revert A' + conv => + left + right + right + right + right + intro x + rw [toReal_rpow] + rw [← toReal_mul] + conv => + left + right + right + right + rw [← ENNReal.tsum_toReal_eq S4] + intro A' + rw [ofReal_toReal_eq_iff.2 S3] at A' + + apply le_trans _ A' + clear A' + apply le_of_eq + + -- Part 3: + + conv => + right + right + right + left + right + intro x + rw [← toReal_mul] + rw [← ENNReal.tsum_toReal_eq S1] + rw [toReal_rpow] + rw [ofReal_toReal_eq_iff.2 S2] + + conv => + right + right + left + right + intro x + rw [division_def] + rw [← δpmf_conv] + rw [mul_assoc] + right + rw [← mul_assoc] + left + rw [ENNReal.inv_mul_cancel (nn l₂ x) (nts l₂ x)] + simp only [one_mul] + + rw [ENNReal.tsum_mul_right] + have H1 : 0 ≤ ∑' (x : ↑{n | i = f n}), (nq l₁ ↑x).toReal := by + apply tsum_nonneg + simp + have H2 : 0 ≤ (∑' (a : ↑{a | i = f a}), nq l₂ ↑a)⁻¹.toReal := by + apply toReal_nonneg + have H4 : (∑' (a : ↑{a | i = f a}), nq l₂ ↑a)⁻¹ ≠ ⊤ := by + apply inv_ne_top.mpr + simp + have T := witness NotEmpty + cases T + rename_i z w + exists z + constructor + . trivial + . apply nn l₂ + rw [ENNReal.mul_rpow_of_ne_top (MasterRW l₁) H4] + + have H3 : ∑' (a : ↑{a | i = f a}), nq l₂ ↑a ≠ 0 := by + simp + have T := witness NotEmpty + cases T + rename_i z w + exists z + constructor + . trivial + . apply nn l₂ + rw [ENNReal.rpow_sub _ _ H3 (MasterRW l₂)] + rw [ENNReal.rpow_one] + rw [division_def] + rw [← mul_assoc] + rw [← mul_assoc] + congr 1 + . rw [mul_comm] + . congr 1 + rw [ENNReal.inv_rpow] end SLang From 8e18c9046ea9757b0099a60958d8cc8561c21548 Mon Sep 17 00:00:00 2001 From: jtristan Date: Fri, 10 May 2024 17:15:54 -0400 Subject: [PATCH 12/20] Proving that the sum is greater than 0 --- SampCert/DiffPrivacy/DP.lean | 166 +++++++++++++++++++++++++++-------- 1 file changed, 129 insertions(+), 37 deletions(-) diff --git a/SampCert/DiffPrivacy/DP.lean b/SampCert/DiffPrivacy/DP.lean index d2930097..d9ad6d97 100644 --- a/SampCert/DiffPrivacy/DP.lean +++ b/SampCert/DiffPrivacy/DP.lean @@ -29,6 +29,9 @@ def DP (q : List T → SLang U) (ε : ℝ) : Prop := def NonZeroNQ (nq : List T → SLang U) := ∀ l : List T, ∀ n : U, nq l n ≠ 0 +def NonTopSum (nq : List T → SLang U) := + ∀ l : List T, ∑' n : U, nq l n ≠ ⊤ + def NonTopNQ (nq : List T → SLang U) := ∀ l : List T, ∀ n : U, nq l n ≠ ⊤ @@ -617,41 +620,33 @@ theorem bar (f : T → ℝ) (q : PMF T) (α : ℝ) (h : 1 < α) (h2 : ∀ x : T, rw [one_le_ofReal] apply le_of_lt h -theorem quux (f : U → ℤ) (g : U → ENNReal) : - (∑' (x : ℤ), if {a : U | x = f a} = {} then 0 else ∑'(i : {a : U | x = f a}), g i) - = ∑' i : U, g i := by - sorry - variable {U : Type} -variable [m2 : MeasurableSpace U] [m2' : MeasurableSingletonClass U] +variable [m2 : MeasurableSpace U] -- [m2' : MeasurableSingletonClass U] variable [count : Countable U] variable [disc : DiscreteMeasurableSpace U] def δ (nq : SLang U) (f : U → ℤ) (a : ℤ) : {n : U | a = f n} → ENNReal := fun x : {n : U | a = f n} => nq x * (∑' (x : {n | a = f n}), nq x)⁻¹ -theorem δ_normalizes (nq : SLang U) (f : U → ℤ) (a : ℤ) : +theorem δ_normalizes (nq : SLang U) (f : U → ℤ) (a : ℤ) (h1 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ 0) (h2 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ ⊤) : HasSum (δ nq f a) 1 := by rw [Summable.hasSum_iff ENNReal.summable] unfold δ rw [ENNReal.tsum_mul_right] - rw [ENNReal.mul_inv_cancel] - . simp - sorry -- ∃ x, a = f x ∧ ¬nq x = 0 - . sorry -- ∑' (i : ↑{n | a = f n}), nq ↑i ≠ ⊤ + rw [ENNReal.mul_inv_cancel h1 h2] -def δpmf (nq : SLang U) (f : U → ℤ) (a : ℤ) : PMF {n : U | a = f n} := - ⟨ δ nq f a , δ_normalizes nq f a ⟩ +def δpmf (nq : SLang U) (f : U → ℤ) (a : ℤ) (h1 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ 0) (h2 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ ⊤) : PMF {n : U | a = f n} := + ⟨ δ nq f a , δ_normalizes nq f a h1 h2 ⟩ -theorem δpmf_conv (nq : List T → SLang U) (a : ℤ) (x : {n | a = f n}) : - nq l₂ x * (∑' (x : {n | a = f n}), nq l₂ x)⁻¹ = (δpmf (nq l₂) f a) x := by +theorem δpmf_conv (nq : SLang U) (a : ℤ) (x : {n | a = f n}) (h1 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ 0) (h2 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ ⊤) : + nq x * (∑' (x : {n | a = f n}), nq x)⁻¹ = (δpmf nq f a h1 h2) x := by simp [δpmf] conv => right left left -theorem δpmf_conv' (nq : List T → SLang U) (l₂ : List T) (f : U → ℤ) (a : ℤ) : - (fun x : {n | a = f n} => nq l₂ x * (∑' (x : {n | a = f n}), nq l₂ x)⁻¹) = (δpmf (nq l₂) f a) := by +theorem δpmf_conv' (nq : SLang U) (f : U → ℤ) (a : ℤ) (h1 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ 0) (h2 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ ⊤) : + (fun x : {n | a = f n} => nq x * (∑' (x : {n | a = f n}), nq x)⁻¹) = (δpmf nq f a h1 h2) := by ext x rw [δpmf_conv] @@ -674,7 +669,86 @@ theorem RD1 (p q : T → ENNReal) (α : ℝ) (h : 1 < α) (RD : ∑' (x : T), p rw [← RenyiDivergenceExpectation p q h nz nt] trivial -theorem DPPostProcess_alt1 {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP nq ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (f : U → ℤ) : +theorem ENNReal.HasSum_fiberwise {f : T → ENNReal} {a : ENNReal} (hf : HasSum f a) (g : T → ℤ) : + HasSum (fun c : ℤ ↦ ∑' b : g ⁻¹' {c}, f b) a := by + let A := (Equiv.sigmaFiberEquiv g) + have B := @Equiv.hasSum_iff ENNReal T ((y : ℤ) × { x // g x = y }) _ _ f a A + replace B := B.2 hf + have C := @HasSum.sigma ENNReal ℤ _ _ _ _ (fun y : ℤ => { x // g x = y }) (f ∘ ⇑(Equiv.sigmaFiberEquiv g)) (fun c => ∑' (b : ↑(g ⁻¹' {c})), f ↑b) a B + apply C + intro b + have F := @Summable.hasSum_iff ENNReal _ _ _ (fun c => (f ∘ ⇑(Equiv.sigmaFiberEquiv g)) { fst := b, snd := c }) ((fun c => ∑' (b : ↑(g ⁻¹' {c})), f ↑b) b) _ + apply (F _).2 + . rfl + . apply ENNReal.summable + +theorem ENNReal.tsum_fiberwise (p : T → ENNReal) (f : T → ℤ) : + ∑' (x : ℤ), ∑' (b : (f ⁻¹' {x})), p b + = ∑' i : T, p i := by + apply HasSum.tsum_eq + apply ENNReal.HasSum_fiberwise + apply Summable.hasSum + exact ENNReal.summable + +theorem quux (p : T → ENNReal) (f : T → ℤ) : + (∑' i : T, p i) + = ∑' (x : ℤ), if {a : T | x = f a} = {} then 0 else ∑'(i : {a : T | x = f a}), p i := by + rw [← ENNReal.tsum_fiberwise p f] + have A : ∀ x, f ⁻¹' {x} = { a | x = f a } := by + intro x + simp [Set.preimage] + rw [Set.ext_iff] + simp + intro y + exact eq_comm + conv => + left + right + intro x + rw [A] + clear A + apply tsum_congr + intro b + split + . rename_i h' + rw [h'] + simp only [tsum_empty] + . simp + +theorem convergent_subset {p : T → ENNReal} (f : T → ℤ) (conv : ∑' (x : T), p x ≠ ⊤) : + ∑' (x : { y : T| x = f y }), p x ≠ ⊤ := by + rw [← foo] + have A : (∑' (y : T), if x = f y then p y else 0) ≤ ∑' (x : T), p x := by + apply tsum_le_tsum + . intro i + split + . trivial + . simp only [_root_.zero_le] + . exact ENNReal.summable + . exact ENNReal.summable + rw [← lt_top_iff_ne_top] + apply lt_of_le_of_lt A + rw [lt_top_iff_ne_top] + trivial + +theorem ENNReal.tsum_pos (f : T → ENNReal) (h1 : ∑' x : T, f x ≠ ⊤) (h2 : ∀ x : T, f x ≠ 0) (i : T) : + 0 < ∑' x : T, f x := by + apply (toNNReal_lt_toNNReal ENNReal.zero_ne_top h1).mp + simp only [zero_toNNReal] + rw [ENNReal.tsum_toNNReal_eq (ENNReal.ne_top_of_tsum_ne_top h1)] + have S : Summable fun a => (f a).toNNReal := by + rw [← tsum_coe_ne_top_iff_summable] + conv => + left + right + intro b + rw [ENNReal.coe_toNNReal (ENNReal.ne_top_of_tsum_ne_top h1 b)] + trivial + have B:= @NNReal.tsum_pos T (fun (a : T) => (f a).toNNReal) S i + apply B + apply ENNReal.toNNReal_pos (h2 i) (ENNReal.ne_top_of_tsum_ne_top h1 i) + +theorem DPPostProcess_alt1 {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP nq ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (conv : NonTopSum nq) (f : U → ℤ) : DP (PostProcess nq f) ((ε₁ : ℝ) / ε₂) := by simp [PostProcess, DP, RenyiDivergence] intro α h1 l₁ l₂ h2 @@ -696,7 +770,9 @@ theorem DPPostProcess_alt1 {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : D -- remove the log have B : 0 < (∑' (x : ℤ), - (∑' (a : U), if x = f a then nq l₁ a else 0) ^ α * (∑' (a : U), if x = f a then nq l₂ a else 0) ^ (1 - α)).toReal := sorry + (∑' (a : U), if x = f a then nq l₁ a else 0) ^ α * (∑' (a : U), if x = f a then nq l₂ a else 0) ^ (1 - α)).toReal := by + sorry + apply log_le_log B clear B @@ -711,25 +787,12 @@ theorem DPPostProcess_alt1 {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : D rw [@RenyiDivergenceExpectation _ (nq l₁) (nq l₂) _ h1 (nn l₂) (nts l₂)] -- Shuffle the sum - rw [← quux f (fun x => (nq l₁ x / nq l₂ x) ^ α * nq l₂ x)] + rw [quux (fun x => (nq l₁ x / nq l₂ x) ^ α * nq l₂ x) f] apply ENNReal.tsum_le_tsum intro i - have S2 : (∑' (a : ↑{n | i = f n}), nq l₁ ↑a / nq l₂ ↑a * (δpmf (nq l₂) f i) a) ^ α ≠ ⊤ := by - sorry - have S1 : ∀ (a : ↑{n | i = f n}), nq l₁ ↑a / nq l₂ ↑a * (δpmf (nq l₂) f i) a ≠ ⊤ := by - sorry - have S3 : ∑' (a : ↑{n | i = f n}), (nq l₁ ↑a / nq l₂ ↑a) ^ α * (δpmf (nq l₂) f i) a ≠ ⊤ := by - sorry - - have S4 : ∀ (a : ↑{n | i = f n}), (nq l₁ ↑a / nq l₂ ↑a) ^ α * (δpmf (nq l₂) f i) a ≠ ⊤ := by - intro a - apply ENNReal.ne_top_of_tsum_ne_top S3 - - have MasterRW : ∀ l : List T, ∑' (a : ↑{a | i = f a}), nq l ↑a ≠ ⊤ := sorry - -- Get rid of elements with probability 0 in the pushforward split . rename_i empty @@ -752,6 +815,37 @@ theorem DPPostProcess_alt1 {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : D -- Part 2: apply Jensen's inequality . rename_i NotEmpty + + have MasterRW : ∀ l : List T, ∑' (a : ↑{a | i = f a}), nq l ↑a ≠ ⊤ := by + intro l + apply convergent_subset + simp [NonTopSum] at conv + have conv := conv l + apply conv + + have MasterZero : ∀ l : List T, ∑' (a : ↑{a | i = f a}), nq l ↑a ≠ 0 := by + intro l + simp + have T := witness NotEmpty + cases T + rename_i z w + exists z + constructor + . trivial + . apply nn l + + have S2 : (∑' (a : ↑{n | i = f n}), nq l₁ ↑a / nq l₂ ↑a * (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) a) ^ α ≠ ⊤ := by + sorry + have S1 : ∀ (a : ↑{n | i = f n}), nq l₁ ↑a / nq l₂ ↑a * (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) a ≠ ⊤ := by + sorry + have S3 : ∑' (a : ↑{n | i = f n}), (nq l₁ ↑a / nq l₂ ↑a) ^ α * (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) a ≠ ⊤ := by + sorry + + have S4 : ∀ (a : ↑{n | i = f n}), (nq l₁ ↑a / nq l₂ ↑a) ^ α * (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) a ≠ ⊤ := by + intro a + apply ENNReal.ne_top_of_tsum_ne_top S3 + + rw [foo] rw [foo] @@ -803,10 +897,8 @@ theorem DPPostProcess_alt1 {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : D intro x simp only [toReal_nonneg] - -- stronglyMeasurable_iff_measurable - -- measurable_discrete have XXX : @Memℒp ℝ Real.normedAddCommGroup (↑{n | i = f n}) Subtype.instMeasurableSpace (fun a => (nq l₁ ↑a / nq l₂ ↑a).toReal) - (ENNReal.ofReal α) (PMF.toMeasure (δpmf (nq l₂) f i)) := by + (ENNReal.ofReal α) (PMF.toMeasure (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂))) := by simp [Memℒp] constructor . apply MeasureTheory.StronglyMeasurable.aestronglyMeasurable @@ -859,7 +951,7 @@ theorem DPPostProcess_alt1 {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : D apply S3 - have Jensen's := @bar {n : U | i = f n} Subtype.instMeasurableSpace Subtype.instMeasurableSingletonClass (fun a => (nq l₁ a / nq l₂ a).toReal) (δpmf (nq l₂) f i) α h1 P5 XXX + have Jensen's := @bar {n : U | i = f n} Subtype.instMeasurableSpace Subtype.instMeasurableSingletonClass (fun a => (nq l₁ a / nq l₂ a).toReal) (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) α h1 P5 XXX clear P5 have P6 : 0 ≤ (∑' (x : ↑{n | i = f n}), nq l₂ ↑x).toReal := by From ffd9d8683d83c6d0624eaf5b628cc2fb6b994feb Mon Sep 17 00:00:00 2001 From: jtristan Date: Mon, 13 May 2024 16:09:56 -0400 Subject: [PATCH 13/20] Proof post-processing complete --- SampCert/DiffPrivacy/DP.lean | 193 ++++++++++++++++++++++++++++------- 1 file changed, 158 insertions(+), 35 deletions(-) diff --git a/SampCert/DiffPrivacy/DP.lean b/SampCert/DiffPrivacy/DP.lean index d9ad6d97..a04cf1a9 100644 --- a/SampCert/DiffPrivacy/DP.lean +++ b/SampCert/DiffPrivacy/DP.lean @@ -731,7 +731,7 @@ theorem convergent_subset {p : T → ENNReal} (f : T → ℤ) (conv : ∑' (x : rw [lt_top_iff_ne_top] trivial -theorem ENNReal.tsum_pos (f : T → ENNReal) (h1 : ∑' x : T, f x ≠ ⊤) (h2 : ∀ x : T, f x ≠ 0) (i : T) : +theorem ENNReal.tsum_pos {f : T → ENNReal} (h1 : ∑' x : T, f x ≠ ⊤) (h2 : ∀ x : T, f x ≠ 0) (i : T) : 0 < ∑' x : T, f x := by apply (toNNReal_lt_toNNReal ENNReal.zero_ne_top h1).mp simp only [zero_toNNReal] @@ -748,40 +748,25 @@ theorem ENNReal.tsum_pos (f : T → ENNReal) (h1 : ∑' x : T, f x ≠ ⊤) (h2 apply B apply ENNReal.toNNReal_pos (h2 i) (ENNReal.ne_top_of_tsum_ne_top h1 i) -theorem DPPostProcess_alt1 {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP nq ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (conv : NonTopSum nq) (f : U → ℤ) : - DP (PostProcess nq f) ((ε₁ : ℝ) / ε₂) := by - simp [PostProcess, DP, RenyiDivergence] - intro α h1 l₁ l₂ h2 - simp [DP, RenyiDivergence] at h - replace h := h α h1 l₁ l₂ h2 +theorem ENNReal.tsum_pos_int {f : ℤ → ENNReal} (h1 : ∑' x : ℤ, f x ≠ ⊤) (h2 : ∀ x : ℤ, f x ≠ 0) : + 0 < ∑' x : ℤ, f x := by + apply ENNReal.tsum_pos h1 h2 42 - -- Part 1, removing fluff - - apply le_trans _ h - clear h +theorem tsum_pos_int {f : ℤ → ENNReal} (h1 : ∑' x : ℤ, f x ≠ ⊤) (h2 : ∀ x : ℤ, f x ≠ 0) : + 0 < (∑' x : ℤ, f x).toReal := by + have X : 0 = (0 : ENNReal).toReal := rfl + rw [X] + clear X + apply toReal_strict_mono h1 + apply ENNReal.tsum_pos_int h1 h2 - -- remove the α scaling - have A : 0 ≤ (α - 1)⁻¹ := by - simp - apply le_of_lt h1 - apply mul_le_mul_of_nonneg_left _ A - clear A - - -- remove the log - have B : 0 < +theorem DPostPocess_pre {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP nq ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (conv : NonTopSum nq) (f : U → ℤ) {α : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (h2 : Neighbour l₁ l₂) : (∑' (x : ℤ), - (∑' (a : U), if x = f a then nq l₁ a else 0) ^ α * (∑' (a : U), if x = f a then nq l₂ a else 0) ^ (1 - α)).toReal := by - sorry + (∑' (a : U), if x = f a then nq l₁ a else 0) ^ α * + (∑' (a : U), if x = f a then nq l₂ a else 0) ^ (1 - α)) ≤ + (∑' (x : U), nq l₁ x ^ α * nq l₂ x ^ (1 - α)) := by - apply log_le_log B - clear B - - have RDConvegence : ∑' (x : U), nq l₁ x ^ α * nq l₂ x ^ (1 - α) ≠ ⊤ := by - simp [NonTopRDNQ] at nt - have nt := nt α h1 l₁ l₂ h2 - trivial - - apply toReal_mono RDConvegence + simp [DP, RenyiDivergence] at h -- Rewrite as cascading expectations rw [@RenyiDivergenceExpectation _ (nq l₁) (nq l₂) _ h1 (nn l₂) (nts l₂)] @@ -835,17 +820,56 @@ theorem DPPostProcess_alt1 {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : D . apply nn l have S2 : (∑' (a : ↑{n | i = f n}), nq l₁ ↑a / nq l₂ ↑a * (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) a) ^ α ≠ ⊤ := by - sorry + conv => + left + left + right + intro a + rw [← δpmf_conv] + rw [division_def] + rw [mul_assoc] + right + rw [← mul_assoc] + rw [ENNReal.inv_mul_cancel (nn l₂ a) (nts l₂ a)] + rw [one_mul] + rw [ENNReal.tsum_mul_right] + apply ENNReal.rpow_ne_top_of_nonneg (le_of_lt (lt_trans zero_lt_one h1 )) + apply mul_ne_top + . apply convergent_subset _ (conv l₁) + . apply inv_ne_top.mpr (MasterZero l₂) + have S1 : ∀ (a : ↑{n | i = f n}), nq l₁ ↑a / nq l₂ ↑a * (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) a ≠ ⊤ := by - sorry + intro a + apply mul_ne_top + . rw [division_def] + apply mul_ne_top (nts l₁ a) + apply inv_ne_top.mpr (nn l₂ a) + . rw [← δpmf_conv] + apply mul_ne_top (nts l₂ a) + apply inv_ne_top.mpr (MasterZero l₂) + have S3 : ∑' (a : ↑{n | i = f n}), (nq l₁ ↑a / nq l₂ ↑a) ^ α * (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) a ≠ ⊤ := by - sorry + conv => + left + right + intro a + rw [← δpmf_conv] + rw [← mul_assoc] + rw [ENNReal.tsum_mul_right] + apply mul_ne_top + . rw [← RenyiDivergenceExpectation _ _ h1] + . replace nt := nt α h1 l₁ l₂ h2 + apply convergent_subset _ nt + . intro x + apply nn + . intro x + apply nts + . apply inv_ne_top.mpr (MasterZero l₂) have S4 : ∀ (a : ↑{n | i = f n}), (nq l₁ ↑a / nq l₂ ↑a) ^ α * (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) a ≠ ⊤ := by intro a apply ENNReal.ne_top_of_tsum_ne_top S3 - rw [foo] rw [foo] @@ -1065,4 +1089,103 @@ theorem DPPostProcess_alt1 {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : D . congr 1 rw [ENNReal.inv_rpow] +theorem tsum_ne_zero_of_ne_zero {T : Type} [Inhabited T] (f : T → ENNReal) (h : ∀ x : T, f x ≠ 0) : + ∑' x : T, f x ≠ 0 := by + by_contra CONTRA + rw [ENNReal.tsum_eq_zero] at CONTRA + have A := h default + have B := CONTRA default + contradiction + +variable [Inhabited U] + +theorem DPPostProcess_alt1 {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP nq ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (conv : NonTopSum nq) (f : U → ℤ) : + DP (PostProcess nq f) ((ε₁ : ℝ) / ε₂) := by + simp [PostProcess, DP, RenyiDivergence] + intro α h1 l₁ l₂ h2 + have h' := h + simp [DP, RenyiDivergence] at h' + replace h' := h' α h1 l₁ l₂ h2 + + -- Part 1, removing fluff + + apply le_trans _ h' + clear h' + + -- remove the α scaling + have A : 0 ≤ (α - 1)⁻¹ := by + simp + apply le_of_lt h1 + apply mul_le_mul_of_nonneg_left _ A + clear A + + have RDConvegence : ∑' (x : U), nq l₁ x ^ α * nq l₂ x ^ (1 - α) ≠ ⊤ := by + simp [NonTopRDNQ] at nt + have nt := nt α h1 l₁ l₂ h2 + trivial + + have B := DPostPocess_pre h nn nt nts conv f h1 h2 + have B' : ∑' (x : ℤ), (∑' (a : U), if x = f a then nq l₁ a else 0) ^ α * (∑' (a : U), if x = f a then nq l₂ a else 0) ^ (1 - α) ≠ ⊤ := by + by_contra CONTRA + rw [CONTRA] at B + simp at B + contradiction + + -- remove the log + apply log_le_log _ (toReal_mono RDConvegence B) + apply toReal_pos _ B' + apply (tsum_ne_zero_iff ENNReal.summable).mpr + exists (f default) + + rw [ENNReal.tsum_eq_add_tsum_ite default] + conv => + left + right + rw [ENNReal.tsum_eq_add_tsum_ite default] + simp only [reduceIte] + apply mul_ne_zero + . by_contra CONTRA + rw [ENNReal.rpow_eq_zero_iff_of_pos (lt_trans zero_lt_one h1)] at CONTRA + simp at CONTRA + cases CONTRA + rename_i left right + have Y := nn l₁ default + contradiction + . by_contra CONTRA + rw [ENNReal.rpow_eq_zero_iff] at CONTRA + cases CONTRA + . rename_i CONTRA + cases CONTRA + rename_i left right + simp at left + cases left + rename_i le1 le2 + have Y := nn l₂ default + contradiction + . rename_i CONTRA + cases CONTRA + rename_i left right + simp at left + cases left + . rename_i left + have Y := nts l₂ default + contradiction + . rename_i left + have Rem := conv l₂ + have X : (∑' (x : U), if x = default then 0 else if f default = f x then nq l₂ x else 0) ≤ ∑' (n : U), nq l₂ n := by + apply ENNReal.tsum_le_tsum + intro a + split + . simp + . split + . simp + . simp + replace Rem := Ne.symm Rem + have Y := Ne.lt_top' Rem + have Z : (∑' (x : U), if x = default then 0 else if f default = f x then nq l₂ x else 0) < ⊤ := by + apply lt_of_le_of_lt X Y + rw [lt_top_iff_ne_top] at Z + contradiction + + end SLang From bd8841525b959c4b1b4d7a48a3f835e35e0f2161 Mon Sep 17 00:00:00 2001 From: jtristan Date: Tue, 14 May 2024 10:46:46 -0400 Subject: [PATCH 14/20] complete set of DP primitives with properties --- SampCert/DiffPrivacy/BoundedAvg.lean | 48 +++++++------- SampCert/DiffPrivacy/BoundedSum.lean | 14 ++-- SampCert/DiffPrivacy/DP.lean | 76 ++++++++++++++++++++++ SampCert/DiffPrivacy/DiscreteGaussian.lean | 6 ++ 4 files changed, 113 insertions(+), 31 deletions(-) diff --git a/SampCert/DiffPrivacy/BoundedAvg.lean b/SampCert/DiffPrivacy/BoundedAvg.lean index 107e1e48..88c55d66 100644 --- a/SampCert/DiffPrivacy/BoundedAvg.lean +++ b/SampCert/DiffPrivacy/BoundedAvg.lean @@ -14,30 +14,30 @@ noncomputable section namespace SLang -def NoisedBoundedAvgQuery (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) (l : List ℤ) : SLang ℤ := do - let S ← NoisedBoundedSumQuery L U h ε₁ (2 * ε₂) l - let C ← NoisedCountingQuery ε₁ (2 * ε₂) l - return S / C - -def NoisedBoundedAvgQuery' (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) (l : List ℤ) : SLang ℤ := - let X := Compose (NoisedBoundedSumQuery L U h ε₁ (2 * ε₂)) (NoisedCountingQuery ε₁ (2 * ε₂)) - PostProcess X (fun z => z.1 / z.2) l - -theorem NoisedBoundedAvgQueryIdentical (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) : - NoisedBoundedAvgQuery' L U h ε₁ ε₂ = NoisedBoundedAvgQuery L U h ε₁ ε₂ := by - ext l x - simp [NoisedBoundedAvgQuery, NoisedBoundedAvgQuery', PostProcess, Compose] - -theorem BoundedSumQueryDP (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) : DP (NoisedBoundedAvgQuery L U h ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by - rw [← NoisedBoundedAvgQueryIdentical] - unfold NoisedBoundedAvgQuery' - simp only - - have A₁ := @NoisedCountingQueryDP ℤ ε₁ (2 * ε₂) - have A₂ := @NoisedCountingQuery_NonZeroNQ ℤ ε₁ (2 * ε₂) - have A₃ := @NoisedCountingQuery_NonTopNQ ℤ ε₁ (2 * ε₂) - have A₄ := @NoisedCountingQuery_NonTopRDNQ ℤ ε₁ (2 * ε₂) - sorry +-- def NoisedBoundedAvgQuery (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) (l : List ℤ) : SLang ℤ := do +-- let S ← NoisedBoundedSumQuery L U h ε₁ (2 * ε₂) l +-- let C ← NoisedCountingQuery ε₁ (2 * ε₂) l +-- return S / C + +-- def NoisedBoundedAvgQuery' (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) (l : List ℤ) : SLang ℤ := +-- let X := Compose (NoisedBoundedSumQuery L U h ε₁ (2 * ε₂)) (NoisedCountingQuery ε₁ (2 * ε₂)) +-- PostProcess X (fun z => z.1 / z.2) l + +-- theorem NoisedBoundedAvgQueryIdentical (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) : +-- NoisedBoundedAvgQuery' L U h ε₁ ε₂ = NoisedBoundedAvgQuery L U h ε₁ ε₂ := by +-- ext l x +-- simp [NoisedBoundedAvgQuery, NoisedBoundedAvgQuery', PostProcess, Compose] + +-- theorem BoundedSumQueryDP (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) : DP (NoisedBoundedAvgQuery L U h ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by +-- rw [← NoisedBoundedAvgQueryIdentical] +-- unfold NoisedBoundedAvgQuery' +-- simp only + +-- have A₁ := @NoisedCountingQueryDP ℤ ε₁ (2 * ε₂) +-- have A₂ := @NoisedCountingQuery_NonZeroNQ ℤ ε₁ (2 * ε₂) +-- have A₃ := @NoisedCountingQuery_NonTopNQ ℤ ε₁ (2 * ε₂) +-- have A₄ := @NoisedCountingQuery_NonTopRDNQ ℤ ε₁ (2 * ε₂) +-- sorry -- have B := @NoisedBoundedSumQueryDP L U h ε₁ (2 * ε₂) -- have C := DPCompose B A -- simp at C diff --git a/SampCert/DiffPrivacy/BoundedSum.lean b/SampCert/DiffPrivacy/BoundedSum.lean index 63694cf0..77d1116b 100644 --- a/SampCert/DiffPrivacy/BoundedSum.lean +++ b/SampCert/DiffPrivacy/BoundedSum.lean @@ -59,14 +59,14 @@ def maxBoundPos (L U : ℤ) (h : L < U) : def maxBound (L U : ℤ) (h : L < U) : ℕ+ := ⟨ max (Int.natAbs L) (Int.natAbs U) , maxBoundPos L U h ⟩ -theorem BoundedSumQuerySensitivity (L U : ℤ) (h : L < U) : sensitivity (BoundedSumQuery L U) (maxBound L U h) := by - sorry +-- theorem BoundedSumQuerySensitivity (L U : ℤ) (h : L < U) : sensitivity (BoundedSumQuery L U) (maxBound L U h) := by +-- sorry -def NoisedBoundedSumQuery (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) (l : List ℤ) : SLang ℤ := do - NoisedQuery (BoundedSumQuery L U) (maxBound L U h) ε₁ ε₂ l +-- def NoisedBoundedSumQuery (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) (l : List ℤ) : SLang ℤ := do +-- NoisedQuery (BoundedSumQuery L U) (maxBound L U h) ε₁ ε₂ l -theorem NoisedBoundedSumQueryDP (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) : DP (NoisedBoundedSumQuery L U h ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by - apply NoisedQueryDP - apply BoundedSumQuerySensitivity +-- theorem NoisedBoundedSumQueryDP (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) : DP (NoisedBoundedSumQuery L U h ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by +-- apply NoisedQueryDP +-- apply BoundedSumQuerySensitivity end SLang diff --git a/SampCert/DiffPrivacy/DP.lean b/SampCert/DiffPrivacy/DP.lean index a04cf1a9..16cf7a0c 100644 --- a/SampCert/DiffPrivacy/DP.lean +++ b/SampCert/DiffPrivacy/DP.lean @@ -142,6 +142,48 @@ theorem discrete_gaussian_shift {σ : ℝ} (h : σ ≠ 0) (μ : ℝ) (τ x : ℤ ring_nf . rw [SG_periodic h] +theorem NoisedQuery_NonTopSum (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : + NonTopSum (NoisedQuery query Δ ε₁ ε₂) := by + simp [NonTopSum, NoisedQuery, DiscreteGaussianGenSample] + intro l + have X : ∀ n: ℤ, ∀ x : ℤ, (@ite ℝ≥0∞ (x = n - query l) (propDecidable (x = n - query l)) 0 + (@ite ℝ≥0∞ (n = x + query l) (instDecidableEqInt n (x + query l)) + (ENNReal.ofReal (discrete_gaussian (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) 0 ↑x)) 0)) = 0 := by + intro n x + split + . simp + . split + . rename_i h1 h2 + subst h2 + simp at h1 + . simp + conv => + right + left + right + intro n + rw [ENNReal.tsum_eq_add_tsum_ite (n - query l)] + simp + right + right + intro x + rw [X] + simp + have A : (Δ : ℝ) * ε₂ / ε₁ ≠ 0 := by + simp + conv => + right + left + right + intro n + rw [discrete_gaussian_shift A] + simp + rw [← ENNReal.ofReal_tsum_of_nonneg] + . rw [discrete_gaussian_normalizes A] + simp + . apply discrete_gaussian_nonneg A + . apply discrete_gaussian_summable' A (query l) + theorem NoisedQuery_NonTopRDNQ (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : NonTopRDNQ (NoisedQuery query Δ ε₁ ε₂) := by simp [NonTopRDNQ, NoisedQuery, DiscreteGaussianGenSample] @@ -443,6 +485,40 @@ theorem DPCompose_NonTopNQ {nq1 nq2 : List T → SLang ℤ} (nt1 : NonTopNQ nq1) cases H contradiction +theorem DPCompose_NonTopSum {nq1 nq2 : List T → SLang ℤ} (nt1 : NonTopSum nq1) (nt2 : NonTopSum nq2) : + NonTopSum (Compose nq1 nq2) := by + simp [NonTopSum] at * + intro l + replace nt1 := nt1 l + replace nt2 := nt2 l + simp [Compose] + rw [ENNReal.tsum_prod'] + conv => + right + left + right + intro a + right + intro b + simp + rw [compose_sum_rw] + conv => + right + left + right + intro a + rw [ENNReal.tsum_mul_left] + rw [ENNReal.tsum_mul_right] + rw [mul_eq_top] + intro H + cases H + . rename_i H + cases H + contradiction + . rename_i H + cases H + contradiction + theorem DPCompose_NonTopRDNQ {nq1 nq2 : List T → SLang ℤ} (nt1 : NonTopRDNQ nq1) (nt2 : NonTopRDNQ nq2) (nn1 : NonZeroNQ nq1) (nn2 : NonZeroNQ nq2) : NonTopRDNQ (Compose nq1 nq2) := by simp [NonTopRDNQ] at * diff --git a/SampCert/DiffPrivacy/DiscreteGaussian.lean b/SampCert/DiffPrivacy/DiscreteGaussian.lean index b4ae8f6f..0118fd73 100644 --- a/SampCert/DiffPrivacy/DiscreteGaussian.lean +++ b/SampCert/DiffPrivacy/DiscreteGaussian.lean @@ -230,6 +230,12 @@ theorem discrete_gaussian_summable {σ : ℝ} (h : σ ≠ 0) (μ : ℝ) : apply Summable.div_const apply summable_gauss_term' h +theorem discrete_gaussian_summable' {σ : ℝ} (h : σ ≠ 0) (μ : ℝ) : + Summable fun (n : ℤ) => discrete_gaussian σ μ n := by + unfold discrete_gaussian + apply Summable.div_const + apply summable_gauss_term' h + theorem discrete_gaussian_normalizes {σ : ℝ} (h : σ ≠ 0) (μ : ℝ) : (∑' n : ℤ, discrete_gaussian σ μ n) = 1 := by unfold discrete_gaussian From 64f6dd5a1b0b19f7fa9d053e6786cc0fc748123f Mon Sep 17 00:00:00 2001 From: jtristan Date: Wed, 15 May 2024 09:56:55 -0400 Subject: [PATCH 15/20] Complete proof all the way to mean query --- SampCert/DiffPrivacy/BoundedAvg.lean | 72 +++++++------ SampCert/DiffPrivacy/BoundedSum.lean | 151 +++++++++++++++++++-------- SampCert/DiffPrivacy/Count.lean | 4 + SampCert/DiffPrivacy/DP.lean | 2 +- 4 files changed, 153 insertions(+), 76 deletions(-) diff --git a/SampCert/DiffPrivacy/BoundedAvg.lean b/SampCert/DiffPrivacy/BoundedAvg.lean index 88c55d66..3fcaa868 100644 --- a/SampCert/DiffPrivacy/BoundedAvg.lean +++ b/SampCert/DiffPrivacy/BoundedAvg.lean @@ -14,36 +14,46 @@ noncomputable section namespace SLang --- def NoisedBoundedAvgQuery (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) (l : List ℤ) : SLang ℤ := do --- let S ← NoisedBoundedSumQuery L U h ε₁ (2 * ε₂) l --- let C ← NoisedCountingQuery ε₁ (2 * ε₂) l --- return S / C - --- def NoisedBoundedAvgQuery' (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) (l : List ℤ) : SLang ℤ := --- let X := Compose (NoisedBoundedSumQuery L U h ε₁ (2 * ε₂)) (NoisedCountingQuery ε₁ (2 * ε₂)) --- PostProcess X (fun z => z.1 / z.2) l - --- theorem NoisedBoundedAvgQueryIdentical (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) : --- NoisedBoundedAvgQuery' L U h ε₁ ε₂ = NoisedBoundedAvgQuery L U h ε₁ ε₂ := by --- ext l x --- simp [NoisedBoundedAvgQuery, NoisedBoundedAvgQuery', PostProcess, Compose] - --- theorem BoundedSumQueryDP (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) : DP (NoisedBoundedAvgQuery L U h ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by --- rw [← NoisedBoundedAvgQueryIdentical] --- unfold NoisedBoundedAvgQuery' --- simp only - --- have A₁ := @NoisedCountingQueryDP ℤ ε₁ (2 * ε₂) --- have A₂ := @NoisedCountingQuery_NonZeroNQ ℤ ε₁ (2 * ε₂) --- have A₃ := @NoisedCountingQuery_NonTopNQ ℤ ε₁ (2 * ε₂) --- have A₄ := @NoisedCountingQuery_NonTopRDNQ ℤ ε₁ (2 * ε₂) --- sorry - -- have B := @NoisedBoundedSumQueryDP L U h ε₁ (2 * ε₂) - -- have C := DPCompose B A - -- simp at C - -- ring_nf at C - -- rw [← division_def] at C - -- have D := DPPostProcess C (fun z => z.1 / z.2) - -- exact D +def NoisedBoundedAvgQuery (U : ℕ+) (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℤ := do + let S ← NoisedBoundedSumQuery U ε₁ (2 * ε₂) l + let C ← NoisedCountingQuery ε₁ (2 * ε₂) l + return S / C + +def NoisedBoundedAvgQuery' (U : ℕ+) (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℤ := + let X := Compose (NoisedBoundedSumQuery U ε₁ (2 * ε₂)) (NoisedCountingQuery ε₁ (2 * ε₂)) + PostProcess X (fun z => z.1 / z.2) l + +theorem NoisedBoundedAvgQueryIdentical (U : ℕ+) (ε₁ ε₂ : ℕ+) : + NoisedBoundedAvgQuery' U ε₁ ε₂ = NoisedBoundedAvgQuery U ε₁ ε₂ := by + ext l x + simp [NoisedBoundedAvgQuery, NoisedBoundedAvgQuery', PostProcess, Compose] + +theorem BoundedSumQueryDP (U : ℕ+) (ε₁ ε₂ : ℕ+) : DP (NoisedBoundedAvgQuery U ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by + rw [← NoisedBoundedAvgQueryIdentical] + unfold NoisedBoundedAvgQuery' + simp only + + have A₁ := @NoisedCountingQueryDP ℕ ε₁ (2 * ε₂) + have A₂ := @NoisedCountingQuery_NonZeroNQ ℕ ε₁ (2 * ε₂) + have A₃ := @NoisedCountingQuery_NonTopNQ ℕ ε₁ (2 * ε₂) + have A₄ := @NoisedCountingQuery_NonTopRDNQ ℕ ε₁ (2 * ε₂) + have A₅ := @NoisedCountingQuery_NonTopSum ℕ ε₁ (2 * ε₂) + + have B₁ := NoisedBoundedSumQueryDP U ε₁ (2 * ε₂) + have B₂ := NoisedBoundedSumQuery_NonZeroNQ U ε₁ (2 * ε₂) + have B₃ := NoisedBoundedSumQuery_NonTopNQ U ε₁ (2 * ε₂) + have B₄ := NoisedBoundedSumQuery_NonTopRDNQ U ε₁ (2 * ε₂) + have B₅ := NoisedBoundedSumQuery_NonTopSum U ε₁ (2 * ε₂) + + have C₁ := DPCompose B₁ A₁ B₂ A₂ B₄ A₄ B₃ A₃ + have C₂ := DPCompose_NonZeroNQ B₂ A₂ + have C₃ := DPCompose_NonTopNQ B₃ A₃ + have C₄ := DPCompose_NonTopRDNQ B₄ A₄ B₂ A₂ + have C₅ := DPCompose_NonTopSum B₅ A₅ + simp at * + ring_nf at * + rw [← division_def] at * + have D := DPPostProcess C₁ C₂ C₄ C₃ C₅ (fun z => z.1 /z.2) + exact D end SLang diff --git a/SampCert/DiffPrivacy/BoundedSum.lean b/SampCert/DiffPrivacy/BoundedSum.lean index 77d1116b..7f2b17ee 100644 --- a/SampCert/DiffPrivacy/BoundedSum.lean +++ b/SampCert/DiffPrivacy/BoundedSum.lean @@ -15,58 +15,121 @@ noncomputable section namespace SLang -def BoundedSumQuery (L U : ℤ) (l : List ℤ) : ℤ := - let center := |U - L| / 2 - let L := L - center - let U := U - center - let S := List.sum (List.map (fun n : ℤ => max (min (n - center) U) L) l) - S + center * List.length l - -theorem natAbs1 {n : ℕ} {x : ℤ} (h : n ≤ x) : - n ≤ natAbs x := by - cases x - . rename_i m +def BoundedSumQuery (U : ℕ+) (l : List ℕ) : ℤ := + List.sum (List.map (fun n : ℕ => (Nat.min n U)) l) + +-- theorem natAbs1 {n : ℕ} {x : ℤ} (h : n ≤ x) : +-- n ≤ natAbs x := by +-- cases x +-- . rename_i m +-- simp +-- exact Int.ofNat_le.mp h +-- . rename_i m +-- simp +-- have A : -[m+1] < 0 := negSucc_lt_zero m +-- have B : n < (0 : ℤ) := by +-- apply Int.lt_of_le_of_lt h A +-- contradiction + +-- def maxBoundPos (L U : ℤ) (h : L < U) : +-- 0 < max (Int.natAbs L) (Int.natAbs U) := by +-- have A : ¬ (L = 0 ∧ U = 0) := by +-- by_contra h' +-- cases h' +-- rename_i h1 h2 +-- subst h1 h2 +-- contradiction +-- simp at A +-- have B : L = 0 ∨ L ≠ 0 := by exact eq_or_ne L 0 +-- cases B +-- . rename_i h' +-- subst h' +-- simp at A +-- simp +-- trivial +-- . rename_i h' +-- simp +-- left +-- trivial + +-- def maxBound (L U : ℤ) (h : L < U) : ℕ+ := +-- ⟨ max (Int.natAbs L) (Int.natAbs U) , maxBoundPos L U h ⟩ + +theorem BoundedSumQuerySensitivity (U : ℕ+) : sensitivity (BoundedSumQuery U) U := by + simp [sensitivity, BoundedSumQuery] + intros l₁ l₂ H + have A : ∀ n : ℕ, (@min ℤ instMinInt (n : ℤ) (U : ℤ) = n) ∨ (@min ℤ instMinInt n U = U) := by + intro n simp - exact Int.ofNat_le.mp h - . rename_i m + exact Nat.le_total n ↑U + cases H + . rename_i l l' n h1 h2 + subst h1 h2 simp - have A : -[m+1] < 0 := negSucc_lt_zero m - have B : n < (0 : ℤ) := by - apply Int.lt_of_le_of_lt h A - contradiction - -def maxBoundPos (L U : ℤ) (h : L < U) : - 0 < max (Int.natAbs L) (Int.natAbs U) := by - have A : ¬ (L = 0 ∧ U = 0) := by - by_contra h' - cases h' - rename_i h1 h2 + cases A n + . rename_i h + rw [h] + simp at * + trivial + . rename_i h + rw [h] + simp + . rename_i l n l' h1 h2 subst h1 h2 - contradiction - simp at A - have B : L = 0 ∨ L ≠ 0 := by exact eq_or_ne L 0 - cases B - . rename_i h' - subst h' - simp at A simp - trivial - . rename_i h' + cases A n + . rename_i h + rw [h] + simp at * + trivial + . rename_i h + rw [h] + simp + . rename_i l n l' m h1 h2 + subst h1 h2 simp - left - trivial + cases A n + . rename_i h + cases A m + . rename_i h' + rw [h, h'] + simp at * + apply Int.natAbs_coe_sub_coe_le_of_le h h' + . rename_i h' + rw [h, h'] + simp at * + apply Int.natAbs_coe_sub_coe_le_of_le h le_rfl + . rename_i h + cases A m + . rename_i h' + rw [h, h'] + simp at * + apply Int.natAbs_coe_sub_coe_le_of_le le_rfl h' + . rename_i h' + rw [h, h'] + simp at * + +def NoisedBoundedSumQuery (U : ℕ+) (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℤ := do + NoisedQuery (BoundedSumQuery U) U ε₁ ε₂ l + +theorem NoisedBoundedSumQueryDP (U : ℕ+) (ε₁ ε₂ : ℕ+) : DP (NoisedBoundedSumQuery U ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by + apply NoisedQueryDP + apply BoundedSumQuerySensitivity -def maxBound (L U : ℤ) (h : L < U) : ℕ+ := - ⟨ max (Int.natAbs L) (Int.natAbs U) , maxBoundPos L U h ⟩ +theorem NoisedBoundedSumQuery_NonZeroNQ (U : ℕ+) (ε₁ ε₂ : ℕ+) : + @NonZeroNQ ℕ ℤ (NoisedBoundedSumQuery U ε₁ ε₂) := by + apply NoisedQuery_NonZeroNQ --- theorem BoundedSumQuerySensitivity (L U : ℤ) (h : L < U) : sensitivity (BoundedSumQuery L U) (maxBound L U h) := by --- sorry +theorem NoisedBoundedSumQuery_NonTopNQ (U : ℕ+) (ε₁ ε₂ : ℕ+) : + @NonTopNQ ℕ ℤ (NoisedBoundedSumQuery U ε₁ ε₂) := by + apply NoisedQuery_NonTopNQ --- def NoisedBoundedSumQuery (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) (l : List ℤ) : SLang ℤ := do --- NoisedQuery (BoundedSumQuery L U) (maxBound L U h) ε₁ ε₂ l +theorem NoisedBoundedSumQuery_NonTopRDNQ (U : ℕ+) (ε₁ ε₂ : ℕ+) : + @NonTopRDNQ ℕ ℤ (NoisedBoundedSumQuery U ε₁ ε₂) := by + apply NoisedQuery_NonTopRDNQ --- theorem NoisedBoundedSumQueryDP (L U : ℤ) (h : L < U) (ε₁ ε₂ : ℕ+) : DP (NoisedBoundedSumQuery L U h ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by --- apply NoisedQueryDP --- apply BoundedSumQuerySensitivity +theorem NoisedBoundedSumQuery_NonTopSum (U : ℕ+) (ε₁ ε₂ : ℕ+) : + @NonTopSum ℕ ℤ (NoisedBoundedSumQuery U ε₁ ε₂) := by + apply NoisedQuery_NonTopSum end SLang diff --git a/SampCert/DiffPrivacy/Count.lean b/SampCert/DiffPrivacy/Count.lean index d5811df5..65902465 100644 --- a/SampCert/DiffPrivacy/Count.lean +++ b/SampCert/DiffPrivacy/Count.lean @@ -51,4 +51,8 @@ theorem NoisedCountingQuery_NonTopRDNQ (ε₁ ε₂ : ℕ+) : @NonTopRDNQ T ℤ (NoisedCountingQuery ε₁ ε₂) := by apply NoisedQuery_NonTopRDNQ +theorem NoisedCountingQuery_NonTopSum (ε₁ ε₂ : ℕ+) : + @NonTopSum T ℤ (NoisedCountingQuery ε₁ ε₂) := by + apply NoisedQuery_NonTopSum + end SLang diff --git a/SampCert/DiffPrivacy/DP.lean b/SampCert/DiffPrivacy/DP.lean index 16cf7a0c..1c47d405 100644 --- a/SampCert/DiffPrivacy/DP.lean +++ b/SampCert/DiffPrivacy/DP.lean @@ -1175,7 +1175,7 @@ theorem tsum_ne_zero_of_ne_zero {T : Type} [Inhabited T] (f : T → ENNReal) (h variable [Inhabited U] -theorem DPPostProcess_alt1 {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP nq ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (conv : NonTopSum nq) (f : U → ℤ) : +theorem DPPostProcess {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP nq ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (conv : NonTopSum nq) (f : U → ℤ) : DP (PostProcess nq f) ((ε₁ : ℝ) / ε₂) := by simp [PostProcess, DP, RenyiDivergence] intro α h1 l₁ l₂ h2 From a75e8ebf4584d725f4b92cec8e4364cb0c1e7f15 Mon Sep 17 00:00:00 2001 From: jtristan Date: Wed, 15 May 2024 15:10:50 -0400 Subject: [PATCH 16/20] Intermediate step in code reorg --- SampCert.lean | 5 +- SampCert/DiffPrivacy/DP.lean | 1267 ----------------- .../Neighbours.lean | 0 .../RenyiDivergence.lean | 0 .../Sensitivity.lean | 4 +- .../ZeroConcentrated}/ConcentratedBound.lean | 10 +- .../ZeroConcentrated/DP.lean | 40 + .../ZeroConcentrated}/DiscreteGaussian.lean | 0 .../ZeroConcentrated}/GaussBound.lean | 2 +- .../ZeroConcentrated}/GaussConvergence.lean | 2 +- .../ZeroConcentrated}/GaussPeriodicity.lean | 2 +- .../Queries/BoundedMean/BoundedMean.lean} | 7 +- .../Queries/BoundedSum}/BoundedSum.lean | 39 +- .../Queries/Count}/Count.lean | 2 +- SampCert/Samplers/Gaussian/Properties.lean | 2 +- SampCert/Samplers/GaussianGen/Properties.lean | 6 +- 16 files changed, 60 insertions(+), 1328 deletions(-) delete mode 100644 SampCert/DiffPrivacy/DP.lean rename SampCert/{DiffPrivacy => DifferentialPrivacy}/Neighbours.lean (100%) rename SampCert/{DiffPrivacy => DifferentialPrivacy}/RenyiDivergence.lean (100%) rename SampCert/{DiffPrivacy => DifferentialPrivacy}/Sensitivity.lean (79%) rename SampCert/{DiffPrivacy => DifferentialPrivacy/ZeroConcentrated}/ConcentratedBound.lean (97%) create mode 100644 SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean rename SampCert/{DiffPrivacy => DifferentialPrivacy/ZeroConcentrated}/DiscreteGaussian.lean (100%) rename SampCert/{DiffPrivacy => DifferentialPrivacy/ZeroConcentrated}/GaussBound.lean (98%) rename SampCert/{DiffPrivacy => DifferentialPrivacy/ZeroConcentrated}/GaussConvergence.lean (92%) rename SampCert/{DiffPrivacy => DifferentialPrivacy/ZeroConcentrated}/GaussPeriodicity.lean (98%) rename SampCert/{DiffPrivacy/BoundedAvg.lean => DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/BoundedMean.lean} (86%) rename SampCert/{DiffPrivacy => DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum}/BoundedSum.lean (74%) rename SampCert/{DiffPrivacy => DifferentialPrivacy/ZeroConcentrated/Queries/Count}/Count.lean (94%) diff --git a/SampCert.lean b/SampCert.lean index f2ca9379..5035d020 100644 --- a/SampCert.lean +++ b/SampCert.lean @@ -4,7 +4,4 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jean-Baptiste Tristan -/ -import SampCert.Samplers.Gaussian.Basic -import SampCert.DiffPrivacy.ConcentratedBound -import SampCert.Samplers.GaussianGen.Basic -import SampCert.DiffPrivacy.BoundedAvg +import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.BoundedMean.BoundedMean diff --git a/SampCert/DiffPrivacy/DP.lean b/SampCert/DiffPrivacy/DP.lean deleted file mode 100644 index 1c47d405..00000000 --- a/SampCert/DiffPrivacy/DP.lean +++ /dev/null @@ -1,1267 +0,0 @@ -/- -Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Jean-Baptiste Tristan --/ - -import Mathlib.Topology.Algebra.InfiniteSum.Basic -import Mathlib.Topology.Algebra.InfiniteSum.Ring -import Mathlib.Algebra.Group.Basic -import SampCert.DiffPrivacy.ConcentratedBound -import SampCert.SLang -import SampCert.Samplers.GaussianGen.Basic -import SampCert.DiffPrivacy.Neighbours -import SampCert.DiffPrivacy.Sensitivity -import Mathlib.MeasureTheory.MeasurableSpace.Basic -import Mathlib.MeasureTheory.Measure.Count -import Mathlib.Probability.ProbabilityMassFunction.Integrals -import Mathlib.Analysis.Convex.SpecificFunctions.Basic -import Mathlib.Analysis.Convex.Integral - -noncomputable section - -open Classical Nat Int Real ENNReal MeasureTheory Measure - -def DP (q : List T → SLang U) (ε : ℝ) : Prop := - ∀ α : ℝ, 1 < α → ∀ l₁ l₂ : List T, Neighbour l₁ l₂ → - RenyiDivergence (q l₁) (q l₂) α ≤ (1/2) * ε ^ 2 * α - -def NonZeroNQ (nq : List T → SLang U) := - ∀ l : List T, ∀ n : U, nq l n ≠ 0 - -def NonTopSum (nq : List T → SLang U) := - ∀ l : List T, ∑' n : U, nq l n ≠ ⊤ - -def NonTopNQ (nq : List T → SLang U) := - ∀ l : List T, ∀ n : U, nq l n ≠ ⊤ - -def NonTopRDNQ (nq : List T → SLang U) : Prop := - ∀ α : ℝ, 1 < α → ∀ l₁ l₂ : List T, Neighbour l₁ l₂ → - ∑' (x : U), nq l₁ x ^ α * nq l₂ x ^ (1 - α) ≠ ⊤ - -namespace SLang - -def NoisedQuery (query : List T → ℤ) (Δ : ℕ+) (ε₁ ε₂ : ℕ+) (l : List T) : SLang ℤ := do - DiscreteGaussianGenSample (Δ * ε₂) ε₁ (query l) - -theorem NoisedQueryDP (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_sensitivity : sensitivity query Δ) : - DP (NoisedQuery query Δ ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by - simp [DP, NoisedQuery] - intros α h1 l₁ l₂ h2 - have A := @DiscreteGaussianGenSampleZeroConcentrated α h1 (Δ * ε₂) ε₁ (query l₁) (query l₂) - apply le_trans A - clear A - replace bounded_sensitivity := bounded_sensitivity l₁ l₂ h2 - ring_nf - simp - conv => - left - left - right - rw [mul_pow] - conv => - left - rw [mul_assoc] - right - rw [mul_comm] - rw [← mul_assoc] - conv => - left - rw [mul_assoc] - right - rw [← mul_assoc] - left - rw [mul_comm] - rw [← mul_assoc] - rw [← mul_assoc] - rw [← mul_assoc] - simp only [inv_pow] - rw [mul_inv_le_iff'] - . have A : (α * ↑↑ε₁ ^ 2 * (↑↑ε₂ ^ 2)⁻¹) ≤ (α * ↑↑ε₁ ^ 2 * (↑↑ε₂ ^ 2)⁻¹) := le_refl (α * ↑↑ε₁ ^ 2 * (↑↑ε₂ ^ 2)⁻¹) - have B : 0 ≤ (α * ↑↑ε₁ ^ 2 * (↑↑ε₂ ^ 2)⁻¹) := by - simp - apply @le_trans ℝ _ 0 1 α (instStrictOrderedCommRingReal.proof_3) (le_of_lt h1) - apply mul_le_mul A _ _ B - . apply sq_le_sq.mpr - simp only [abs_cast] - rw [← Int.cast_sub] - rw [← Int.cast_abs] - apply Int.cast_le.mpr - rw [← Int.natCast_natAbs] - apply Int.ofNat_le.mpr - trivial - . apply sq_nonneg - . rw [pow_two] - rw [_root_.mul_pos_iff] - left - simp - -theorem NoisedQuery_NonZeroNQ (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : - NonZeroNQ (NoisedQuery query Δ ε₁ ε₂) := by - simp [NonZeroNQ, NoisedQuery, DiscreteGaussianGenSample] - intros l n - exists (n - query l) - simp - have A : ((Δ : ℝ) * ε₂ / ε₁) ≠ 0 := by - simp - have X := @discrete_gaussian_pos (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) A 0 (n - query l) - simp at X - trivial - -theorem NoisedQuery_NonTopNQ (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : - NonTopNQ (NoisedQuery query Δ ε₁ ε₂) := by - simp [NonTopNQ, NoisedQuery, DiscreteGaussianGenSample] - intro l n - rw [ENNReal.tsum_eq_add_tsum_ite (n - query l)] - simp - have X : ∀ x : ℤ, (@ite ℝ≥0∞ (x = n - query l) (propDecidable (x = n - query l)) 0 - (@ite ℝ≥0∞ (n = x + query l) (instDecidableEqInt n (x + query l)) - (ENNReal.ofReal (discrete_gaussian (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) 0 ↑x)) 0)) = 0 := by - intro x - split - . simp - . split - . rename_i h1 h2 - subst h2 - simp at h1 - . simp - conv => - right - left - right - intro x - rw [X] - simp - -theorem discrete_gaussian_shift {σ : ℝ} (h : σ ≠ 0) (μ : ℝ) (τ x : ℤ) : - discrete_gaussian σ μ (x - τ) = discrete_gaussian σ (μ + τ) (x) := by - simp [discrete_gaussian] - congr 1 - . simp [gauss_term_ℝ] - congr 3 - ring_nf - . rw [SG_periodic h] - -theorem NoisedQuery_NonTopSum (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : - NonTopSum (NoisedQuery query Δ ε₁ ε₂) := by - simp [NonTopSum, NoisedQuery, DiscreteGaussianGenSample] - intro l - have X : ∀ n: ℤ, ∀ x : ℤ, (@ite ℝ≥0∞ (x = n - query l) (propDecidable (x = n - query l)) 0 - (@ite ℝ≥0∞ (n = x + query l) (instDecidableEqInt n (x + query l)) - (ENNReal.ofReal (discrete_gaussian (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) 0 ↑x)) 0)) = 0 := by - intro n x - split - . simp - . split - . rename_i h1 h2 - subst h2 - simp at h1 - . simp - conv => - right - left - right - intro n - rw [ENNReal.tsum_eq_add_tsum_ite (n - query l)] - simp - right - right - intro x - rw [X] - simp - have A : (Δ : ℝ) * ε₂ / ε₁ ≠ 0 := by - simp - conv => - right - left - right - intro n - rw [discrete_gaussian_shift A] - simp - rw [← ENNReal.ofReal_tsum_of_nonneg] - . rw [discrete_gaussian_normalizes A] - simp - . apply discrete_gaussian_nonneg A - . apply discrete_gaussian_summable' A (query l) - -theorem NoisedQuery_NonTopRDNQ (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : - NonTopRDNQ (NoisedQuery query Δ ε₁ ε₂) := by - simp [NonTopRDNQ, NoisedQuery, DiscreteGaussianGenSample] - intro α _ l₁ l₂ _ - have A : ∀ x_1 x : ℤ, (@ite ℝ≥0∞ (x_1 = x - query l₁) (propDecidable (x_1 = x - query l₁)) 0 - (@ite ℝ≥0∞ (x = x_1 + query l₁) (instDecidableEqInt x (x_1 + query l₁)) - (ENNReal.ofReal (discrete_gaussian (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) 0 ↑x_1)) 0 )) = 0 := by - intro x y - split - . simp - . split - . rename_i h1 h2 - subst h2 - simp at h1 - . simp - have B : ∀ x_1 x : ℤ, (@ite ℝ≥0∞ (x_1 = x - query l₂) (propDecidable (x_1 = x - query l₂)) 0 - (@ite ℝ≥0∞ (x = x_1 + query l₂) (instDecidableEqInt x (x_1 + query l₂)) - (ENNReal.ofReal (discrete_gaussian (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) 0 ↑x_1)) 0)) = 0 := by - intro x y - split - . simp - . split - . rename_i h1 h2 - subst h2 - simp at h1 - . simp - conv => - right - left - right - intro x - left - rw [ENNReal.tsum_eq_add_tsum_ite (x - query l₁)] - simp - left - right - right - intro y - rw [A] - simp - conv => - right - left - right - intro x - right - rw [ENNReal.tsum_eq_add_tsum_ite (x - query l₂)] - simp - left - right - right - intro y - rw [B] - simp - clear A B - have P : (Δ : ℝ) * ε₂ / ε₁ ≠ 0 := by - simp - have A : ∀ x : ℤ, ∀ l : List T, 0 < discrete_gaussian (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) 0 (↑x - ↑(query l)) := by - intro x l - have A' := @discrete_gaussian_pos _ P 0 (x - query l) - simp at A' - trivial - have B : ∀ x : ℤ, 0 ≤ discrete_gaussian (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) 0 (↑x - ↑(query l₁)) ^ α := by - intro x - have A' := @discrete_gaussian_nonneg _ P 0 (x - query l₁) - simp at A' - apply rpow_nonneg A' - conv => - right - left - right - intro x - rw [ENNReal.ofReal_rpow_of_pos (A x l₁)] - rw [ENNReal.ofReal_rpow_of_pos (A x l₂)] - rw [← ENNReal.ofReal_mul (B x)] - rw [← ENNReal.ofReal_tsum_of_nonneg] - . simp - . intro n - have X := @RenyiSumSG_nonneg _ α P (query l₁) (query l₂) n - rw [discrete_gaussian_shift P] - rw [discrete_gaussian_shift P] - simp [X] - . have X := @SummableRenyiGauss _ P (query l₁) (query l₂) - conv => - right - intro x - rw [discrete_gaussian_shift P] - rw [discrete_gaussian_shift P] - simp [X] - -def Compose (nq1 nq2 : List T → SLang ℤ) (l : List T) : SLang (ℤ × ℤ) := do - let A ← nq1 l - let B ← nq2 l - return (A,B) - -theorem ENNReal_toTeal_NZ (x : ENNReal) (h1 : x ≠ 0) (h2 : x ≠ ⊤) : - x.toReal ≠ 0 := by - unfold ENNReal.toReal - unfold ENNReal.toNNReal - simp - intro H - cases H - . contradiction - . contradiction - -theorem simp_α_1 {α : ℝ} (h : 1 < α) : 0 < α := by - apply @lt_trans _ _ _ 1 _ _ h - simp only [zero_lt_one] - -theorem RenyiNoisedQueryNonZero {nq : List T → SLang ℤ} {α ε : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (h2 : Neighbour l₁ l₂) (h3 : DP nq ε) (h4 : NonZeroNQ nq) (h5 : NonTopRDNQ nq) (nts : NonTopNQ nq) : - (∑' (i : ℤ), nq l₁ i ^ α * nq l₂ i ^ (1 - α)).toReal ≠ 0 := by - simp [DP] at h3 - replace h3 := h3 α h1 l₁ l₂ h2 - simp [RenyiDivergence] at h3 - simp [NonZeroNQ] at h4 - simp [NonTopRDNQ] at h5 - replace h5 := h5 α h1 l₁ l₂ h2 - have h6 := h4 l₁ - have h7 := h4 l₂ - apply ENNReal_toTeal_NZ - . by_contra CONTRA - rw [ENNReal.tsum_eq_zero] at CONTRA - replace CONTRA := CONTRA 42 - replace h6 := h6 42 - replace h7 := h7 42 - rw [_root_.mul_eq_zero] at CONTRA - cases CONTRA - . rename_i h8 - rw [rpow_eq_zero_iff_of_pos] at h8 - contradiction - apply simp_α_1 h1 - . rename_i h8 - rw [ENNReal.rpow_eq_zero_iff] at h8 - cases h8 - . rename_i h8 - cases h8 - contradiction - . rename_i h8 - cases h8 - rename_i h8 h9 - replace nts := nts l₂ 42 - contradiction - . exact h5 - -theorem compose_sum_rw (nq1 nq2 : List T → SLang ℤ) (b c : ℤ) (l : List T) : - (∑' (a : ℤ), nq1 l a * ∑' (a_1 : ℤ), if b = a ∧ c = a_1 then nq2 l a_1 else 0) = nq1 l b * nq2 l c := by - have A : ∀ a b : ℤ, (∑' (a_1 : ℤ), if b = a ∧ c = a_1 then nq2 l a_1 else 0) = if b = a then (∑' (a_1 : ℤ), if c = a_1 then nq2 l a_1 else 0) else 0 := by - intro x y - split - . rename_i h - subst h - simp - . rename_i h - simp - intro h - contradiction - conv => - left - right - intro a - right - rw [A] - rw [ENNReal.tsum_eq_add_tsum_ite b] - simp - have B : ∀ x : ℤ, (@ite ℝ≥0∞ (x = b) (instDecidableEqInt x b) 0 - (@ite ℝ≥0∞ (b = x) (instDecidableEqInt b x) (nq1 l x * ∑' (a_1 : ℤ), @ite ℝ≥0∞ (c = a_1) (instDecidableEqInt c a_1) (nq2 l a_1) 0) 0)) = 0 := by - intro x - split - . simp - . split - . rename_i h1 h2 - subst h2 - contradiction - . simp - conv => - left - right - right - intro x - rw [B] - simp - congr 1 - rw [ENNReal.tsum_eq_add_tsum_ite c] - simp - have C :∀ x : ℤ, (@ite ℝ≥0∞ (x = c) (propDecidable (x = c)) 0 (@ite ℝ≥0∞ (c = x) (instDecidableEqInt c x) (nq2 l x) 0)) = 0 := by - intro x - split - . simp - . split - . rename_i h1 h2 - subst h2 - contradiction - . simp - conv => - left - right - right - intro X - rw [C] - simp - -theorem DPCompose {nq1 nq2 : List T → SLang ℤ} {ε₁ ε₂ ε₃ ε₄ : ℕ+} (h1 : DP nq1 ((ε₁ : ℝ) / ε₂)) (h2 : DP nq2 ((ε₃ : ℝ) / ε₄)) (nn1 : NonZeroNQ nq1) (nn2 : NonZeroNQ nq2) (nt1 : NonTopRDNQ nq1) (nt2 : NonTopRDNQ nq2) (nts1 : NonTopNQ nq1) (nts2 : NonTopNQ nq2) : - DP (Compose nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by - simp [Compose, RenyiDivergence, DP] - intro α h3 l₁ l₂ h4 - have X := h1 - have Y := h2 - simp [DP] at h1 h2 - replace h1 := h1 α h3 l₁ l₂ h4 - replace h2 := h2 α h3 l₁ l₂ h4 - simp [RenyiDivergence] at h1 h2 - rw [tsum_prod' ENNReal.summable (fun b : ℤ => ENNReal.summable)] - . simp - conv => - left - right - right - right - right - intro b - right - intro c - rw [compose_sum_rw] - rw [compose_sum_rw] - rw [ENNReal.mul_rpow_of_ne_zero (nn1 l₁ b) (nn2 l₁ c)] - rw [ENNReal.mul_rpow_of_ne_zero (nn1 l₂ b) (nn2 l₂ c)] - rw [mul_assoc] - right - rw [mul_comm] - rw [mul_assoc] - right - rw [mul_comm] - conv => - left - right - right - right - right - intro b - right - intro c - rw [← mul_assoc] - conv => - left - right - right - right - right - intro b - rw [ENNReal.tsum_mul_left] - rw [ENNReal.tsum_mul_right] - rw [ENNReal.toReal_mul] - rw [Real.log_mul] - . rw [mul_add] - have D := _root_.add_le_add h1 h2 - apply le_trans D - rw [← add_mul] - rw [mul_le_mul_iff_of_pos_right] - . rw [← mul_add] - rw [mul_le_mul_iff_of_pos_left] - . ring_nf - simp - . simp - . apply lt_trans zero_lt_one h3 - . apply RenyiNoisedQueryNonZero h3 h4 X nn1 nt1 nts1 - . apply RenyiNoisedQueryNonZero h3 h4 Y nn2 nt2 nts2 - -theorem DPCompose_NonZeroNQ {nq1 nq2 : List T → SLang ℤ} (nn1 : NonZeroNQ nq1) (nn2 : NonZeroNQ nq2) : - NonZeroNQ (Compose nq1 nq2) := by - simp [NonZeroNQ] at * - intro l a b - replace nn1 := nn1 l a - replace nn2 := nn2 l b - simp [Compose] - exists a - simp - intro H - cases H - . rename_i H - contradiction - . rename_i H - contradiction - -theorem DPCompose_NonTopNQ {nq1 nq2 : List T → SLang ℤ} (nt1 : NonTopNQ nq1) (nt2 : NonTopNQ nq2) : - NonTopNQ (Compose nq1 nq2) := by - simp [NonTopNQ] at * - intro l a b - replace nt1 := nt1 l a - replace nt2 := nt2 l b - simp [Compose] - rw [compose_sum_rw] - rw [mul_eq_top] - intro H - cases H - . rename_i H - cases H - contradiction - . rename_i H - cases H - contradiction - -theorem DPCompose_NonTopSum {nq1 nq2 : List T → SLang ℤ} (nt1 : NonTopSum nq1) (nt2 : NonTopSum nq2) : - NonTopSum (Compose nq1 nq2) := by - simp [NonTopSum] at * - intro l - replace nt1 := nt1 l - replace nt2 := nt2 l - simp [Compose] - rw [ENNReal.tsum_prod'] - conv => - right - left - right - intro a - right - intro b - simp - rw [compose_sum_rw] - conv => - right - left - right - intro a - rw [ENNReal.tsum_mul_left] - rw [ENNReal.tsum_mul_right] - rw [mul_eq_top] - intro H - cases H - . rename_i H - cases H - contradiction - . rename_i H - cases H - contradiction - -theorem DPCompose_NonTopRDNQ {nq1 nq2 : List T → SLang ℤ} (nt1 : NonTopRDNQ nq1) (nt2 : NonTopRDNQ nq2) (nn1 : NonZeroNQ nq1) (nn2 : NonZeroNQ nq2) : - NonTopRDNQ (Compose nq1 nq2) := by - simp [NonTopRDNQ] at * - intro α h1 l₁ l₂ h2 - replace nt1 := nt1 α h1 l₁ l₂ h2 - replace nt2 := nt2 α h1 l₁ l₂ h2 - simp [Compose] - rw [ENNReal.tsum_prod'] - simp - conv => - right - left - right - intro x - right - intro y - congr - . left - rw [compose_sum_rw] - . left - rw [compose_sum_rw] - conv => - right - left - right - intro x - right - intro y - rw [ENNReal.mul_rpow_of_ne_zero (nn1 l₁ x) (nn2 l₁ y)] - rw [ENNReal.mul_rpow_of_ne_zero (nn1 l₂ x) (nn2 l₂ y)] - rw [mul_assoc] - right - rw [mul_comm] - rw [mul_assoc] - right - rw [mul_comm] - conv => - right - left - right - intro x - right - intro y - rw [← mul_assoc] - conv => - right - left - right - intro x - rw [ENNReal.tsum_mul_left] - rw [ENNReal.tsum_mul_right] - intro H - rw [mul_eq_top] at H - cases H - . rename_i h3 - cases h3 - rename_i h4 h5 - contradiction - . rename_i h3 - cases h3 - rename_i h4 h5 - contradiction - -def PostProcess (nq : List T → SLang U) (pp : U → ℤ) (l : List T) : SLang ℤ := do - let A ← nq l - return pp A - -theorem foo (f : U → ℤ) (g : U → ENNReal) (x : ℤ) : - (∑' a : U, if x = f a then g a else 0) = ∑' a : { a | x = f a }, g a := by - have A := @tsum_split_ite U (fun a : U => x = f a) g (fun _ => 0) - simp only [decide_eq_true_eq, tsum_zero, add_zero] at A - rw [A] - have B : ↑{i | decide (x = f i) = true} = ↑{a | x = f a} := by - simp - rw [B] - -variable {T : Type} -variable [m1 : MeasurableSpace T] -variable [m2 : MeasurableSingletonClass T] -variable [m3: MeasureSpace T] - -theorem Integrable_rpow (f : T → ℝ) (nn : ∀ x : T, 0 ≤ f x) (μ : Measure T) (α : ENNReal) (mem : Memℒp f α μ) (h1 : α ≠ 0) (h2 : α ≠ ⊤) : - MeasureTheory.Integrable (fun x : T => (f x) ^ α.toReal) μ := by - have X := @MeasureTheory.Memℒp.integrable_norm_rpow T ℝ m1 μ _ f α mem h1 h2 - revert X - conv => - left - left - intro x - rw [← norm_rpow_of_nonneg (nn x)] - intro X - simp [Integrable] at * - constructor - . cases X - rename_i left right - rw [@aestronglyMeasurable_iff_aemeasurable] - apply AEMeasurable.pow_const - simp [Memℒp] at mem - cases mem - rename_i left' right' - rw [aestronglyMeasurable_iff_aemeasurable] at left' - simp [left'] - . rw [← hasFiniteIntegral_norm_iff] - simp [X] - -theorem bar (f : T → ℝ) (q : PMF T) (α : ℝ) (h : 1 < α) (h2 : ∀ x : T, 0 ≤ f x) (mem : Memℒp f (ENNReal.ofReal α) (PMF.toMeasure q)) : - ((∑' x : T, (f x) * (q x).toReal)) ^ α ≤ (∑' x : T, (f x) ^ α * (q x).toReal) := by - - conv => - left - left - right - intro x - rw [mul_comm] - rw [← smul_eq_mul] - conv => - right - right - intro x - rw [mul_comm] - rw [← smul_eq_mul] - rw [← PMF.integral_eq_tsum] - rw [← PMF.integral_eq_tsum] - - have A := @convexOn_rpow α (le_of_lt h) - have B : ContinuousOn (fun (x : ℝ) => x ^ α) (Set.Ici 0) := by - apply ContinuousOn.rpow - . exact continuousOn_id' (Set.Ici 0) - . exact continuousOn_const - . intro x h' - simp at h' - have OR : x = 0 ∨ 0 < x := by exact LE.le.eq_or_gt h' - cases OR - . rename_i h'' - subst h'' - right - apply lt_trans zero_lt_one h - . rename_i h'' - left - by_contra - rename_i h3 - subst h3 - simp at h'' - have C : @IsClosed ℝ UniformSpace.toTopologicalSpace (Set.Ici 0) := by - exact isClosed_Ici - have D := @ConvexOn.map_integral_le T ℝ m1 _ _ _ (PMF.toMeasure q) (Set.Ici 0) f (fun (x : ℝ) => x ^ α) (PMF.toMeasure.isProbabilityMeasure q) A B C - simp at D - apply D - . exact MeasureTheory.ae_of_all (PMF.toMeasure q) h2 - . apply MeasureTheory.Memℒp.integrable _ mem - rw [one_le_ofReal] - apply le_of_lt h - . rw [Function.comp_def] - have X : ENNReal.ofReal α ≠ 0 := by - simp - apply lt_trans zero_lt_one h - have Y : ENNReal.ofReal α ≠ ⊤ := by - simp - have Z := @Integrable_rpow T m1 f h2 (PMF.toMeasure q) (ENNReal.ofReal α) mem X Y - rw [toReal_ofReal] at Z - . exact Z - . apply le_of_lt - apply lt_trans zero_lt_one h - . have X : ENNReal.ofReal α ≠ 0 := by - simp - apply lt_trans zero_lt_one h - have Y : ENNReal.ofReal α ≠ ⊤ := by - simp - have Z := @Integrable_rpow T m1 f h2 (PMF.toMeasure q) (ENNReal.ofReal α) mem X Y - rw [toReal_ofReal] at Z - . exact Z - . apply le_of_lt - apply lt_trans zero_lt_one h - . apply MeasureTheory.Memℒp.integrable _ mem - rw [one_le_ofReal] - apply le_of_lt h - -variable {U : Type} -variable [m2 : MeasurableSpace U] -- [m2' : MeasurableSingletonClass U] -variable [count : Countable U] -variable [disc : DiscreteMeasurableSpace U] - -def δ (nq : SLang U) (f : U → ℤ) (a : ℤ) : {n : U | a = f n} → ENNReal := fun x : {n : U | a = f n} => nq x * (∑' (x : {n | a = f n}), nq x)⁻¹ - -theorem δ_normalizes (nq : SLang U) (f : U → ℤ) (a : ℤ) (h1 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ 0) (h2 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ ⊤) : - HasSum (δ nq f a) 1 := by - rw [Summable.hasSum_iff ENNReal.summable] - unfold δ - rw [ENNReal.tsum_mul_right] - rw [ENNReal.mul_inv_cancel h1 h2] - -def δpmf (nq : SLang U) (f : U → ℤ) (a : ℤ) (h1 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ 0) (h2 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ ⊤) : PMF {n : U | a = f n} := - ⟨ δ nq f a , δ_normalizes nq f a h1 h2 ⟩ - -theorem δpmf_conv (nq : SLang U) (a : ℤ) (x : {n | a = f n}) (h1 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ 0) (h2 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ ⊤) : - nq x * (∑' (x : {n | a = f n}), nq x)⁻¹ = (δpmf nq f a h1 h2) x := by - simp [δpmf] - conv => - right - left - left - -theorem δpmf_conv' (nq : SLang U) (f : U → ℤ) (a : ℤ) (h1 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ 0) (h2 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ ⊤) : - (fun x : {n | a = f n} => nq x * (∑' (x : {n | a = f n}), nq x)⁻¹) = (δpmf nq f a h1 h2) := by - ext x - rw [δpmf_conv] - -theorem witness {f : U → ℤ} {i : ℤ} (h : ¬{b | i = f b} = ∅) : - ∃ x : U, i = f x := by - rw [← nonempty_subtype] - exact Set.nonempty_iff_ne_empty'.mpr h - -theorem norm_simplify (x : ENNReal) (h : x ≠ ⊤) : - @nnnorm ℝ SeminormedAddGroup.toNNNorm x.toReal = x := by - simp [nnnorm] - cases x - . contradiction - . rename_i v - simp - rfl - -theorem RD1 (p q : T → ENNReal) (α : ℝ) (h : 1 < α) (RD : ∑' (x : T), p x ^ α * q x ^ (1 - α) ≠ ⊤) (nz : ∀ x : T, q x ≠ 0) (nt : ∀ x : T, q x ≠ ⊤) : - ∑' (x : T), (p x / q x) ^ α * q x ≠ ⊤ := by - rw [← RenyiDivergenceExpectation p q h nz nt] - trivial - -theorem ENNReal.HasSum_fiberwise {f : T → ENNReal} {a : ENNReal} (hf : HasSum f a) (g : T → ℤ) : - HasSum (fun c : ℤ ↦ ∑' b : g ⁻¹' {c}, f b) a := by - let A := (Equiv.sigmaFiberEquiv g) - have B := @Equiv.hasSum_iff ENNReal T ((y : ℤ) × { x // g x = y }) _ _ f a A - replace B := B.2 hf - have C := @HasSum.sigma ENNReal ℤ _ _ _ _ (fun y : ℤ => { x // g x = y }) (f ∘ ⇑(Equiv.sigmaFiberEquiv g)) (fun c => ∑' (b : ↑(g ⁻¹' {c})), f ↑b) a B - apply C - intro b - have F := @Summable.hasSum_iff ENNReal _ _ _ (fun c => (f ∘ ⇑(Equiv.sigmaFiberEquiv g)) { fst := b, snd := c }) ((fun c => ∑' (b : ↑(g ⁻¹' {c})), f ↑b) b) _ - apply (F _).2 - . rfl - . apply ENNReal.summable - -theorem ENNReal.tsum_fiberwise (p : T → ENNReal) (f : T → ℤ) : - ∑' (x : ℤ), ∑' (b : (f ⁻¹' {x})), p b - = ∑' i : T, p i := by - apply HasSum.tsum_eq - apply ENNReal.HasSum_fiberwise - apply Summable.hasSum - exact ENNReal.summable - -theorem quux (p : T → ENNReal) (f : T → ℤ) : - (∑' i : T, p i) - = ∑' (x : ℤ), if {a : T | x = f a} = {} then 0 else ∑'(i : {a : T | x = f a}), p i := by - rw [← ENNReal.tsum_fiberwise p f] - have A : ∀ x, f ⁻¹' {x} = { a | x = f a } := by - intro x - simp [Set.preimage] - rw [Set.ext_iff] - simp - intro y - exact eq_comm - conv => - left - right - intro x - rw [A] - clear A - apply tsum_congr - intro b - split - . rename_i h' - rw [h'] - simp only [tsum_empty] - . simp - -theorem convergent_subset {p : T → ENNReal} (f : T → ℤ) (conv : ∑' (x : T), p x ≠ ⊤) : - ∑' (x : { y : T| x = f y }), p x ≠ ⊤ := by - rw [← foo] - have A : (∑' (y : T), if x = f y then p y else 0) ≤ ∑' (x : T), p x := by - apply tsum_le_tsum - . intro i - split - . trivial - . simp only [_root_.zero_le] - . exact ENNReal.summable - . exact ENNReal.summable - rw [← lt_top_iff_ne_top] - apply lt_of_le_of_lt A - rw [lt_top_iff_ne_top] - trivial - -theorem ENNReal.tsum_pos {f : T → ENNReal} (h1 : ∑' x : T, f x ≠ ⊤) (h2 : ∀ x : T, f x ≠ 0) (i : T) : - 0 < ∑' x : T, f x := by - apply (toNNReal_lt_toNNReal ENNReal.zero_ne_top h1).mp - simp only [zero_toNNReal] - rw [ENNReal.tsum_toNNReal_eq (ENNReal.ne_top_of_tsum_ne_top h1)] - have S : Summable fun a => (f a).toNNReal := by - rw [← tsum_coe_ne_top_iff_summable] - conv => - left - right - intro b - rw [ENNReal.coe_toNNReal (ENNReal.ne_top_of_tsum_ne_top h1 b)] - trivial - have B:= @NNReal.tsum_pos T (fun (a : T) => (f a).toNNReal) S i - apply B - apply ENNReal.toNNReal_pos (h2 i) (ENNReal.ne_top_of_tsum_ne_top h1 i) - -theorem ENNReal.tsum_pos_int {f : ℤ → ENNReal} (h1 : ∑' x : ℤ, f x ≠ ⊤) (h2 : ∀ x : ℤ, f x ≠ 0) : - 0 < ∑' x : ℤ, f x := by - apply ENNReal.tsum_pos h1 h2 42 - -theorem tsum_pos_int {f : ℤ → ENNReal} (h1 : ∑' x : ℤ, f x ≠ ⊤) (h2 : ∀ x : ℤ, f x ≠ 0) : - 0 < (∑' x : ℤ, f x).toReal := by - have X : 0 = (0 : ENNReal).toReal := rfl - rw [X] - clear X - apply toReal_strict_mono h1 - apply ENNReal.tsum_pos_int h1 h2 - -theorem DPostPocess_pre {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP nq ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (conv : NonTopSum nq) (f : U → ℤ) {α : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (h2 : Neighbour l₁ l₂) : - (∑' (x : ℤ), - (∑' (a : U), if x = f a then nq l₁ a else 0) ^ α * - (∑' (a : U), if x = f a then nq l₂ a else 0) ^ (1 - α)) ≤ - (∑' (x : U), nq l₁ x ^ α * nq l₂ x ^ (1 - α)) := by - - simp [DP, RenyiDivergence] at h - - -- Rewrite as cascading expectations - rw [@RenyiDivergenceExpectation _ (nq l₁) (nq l₂) _ h1 (nn l₂) (nts l₂)] - - -- Shuffle the sum - rw [quux (fun x => (nq l₁ x / nq l₂ x) ^ α * nq l₂ x) f] - - apply ENNReal.tsum_le_tsum - - intro i - - -- Get rid of elements with probability 0 in the pushforward - split - . rename_i empty - rw [foo] - have ZE : (∑' (x_1 : ↑{n | i = f n}), nq l₁ ↑x_1) = 0 := by - simp - intro a H - have I₁ : a ∈ {b | i = f b} := by - simp [H] - have I2 : {b | i = f b} ≠ ∅ := by - apply ne_of_mem_of_not_mem' I₁ - simp - contradiction - rw [ZE] - simp only [toReal_mul, zero_toReal, ge_iff_le] - - rw [ENNReal.zero_rpow_of_pos] - . simp - . apply lt_trans zero_lt_one h1 - - -- Part 2: apply Jensen's inequality - . rename_i NotEmpty - - have MasterRW : ∀ l : List T, ∑' (a : ↑{a | i = f a}), nq l ↑a ≠ ⊤ := by - intro l - apply convergent_subset - simp [NonTopSum] at conv - have conv := conv l - apply conv - - have MasterZero : ∀ l : List T, ∑' (a : ↑{a | i = f a}), nq l ↑a ≠ 0 := by - intro l - simp - have T := witness NotEmpty - cases T - rename_i z w - exists z - constructor - . trivial - . apply nn l - - have S2 : (∑' (a : ↑{n | i = f n}), nq l₁ ↑a / nq l₂ ↑a * (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) a) ^ α ≠ ⊤ := by - conv => - left - left - right - intro a - rw [← δpmf_conv] - rw [division_def] - rw [mul_assoc] - right - rw [← mul_assoc] - rw [ENNReal.inv_mul_cancel (nn l₂ a) (nts l₂ a)] - rw [one_mul] - rw [ENNReal.tsum_mul_right] - apply ENNReal.rpow_ne_top_of_nonneg (le_of_lt (lt_trans zero_lt_one h1 )) - apply mul_ne_top - . apply convergent_subset _ (conv l₁) - . apply inv_ne_top.mpr (MasterZero l₂) - - have S1 : ∀ (a : ↑{n | i = f n}), nq l₁ ↑a / nq l₂ ↑a * (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) a ≠ ⊤ := by - intro a - apply mul_ne_top - . rw [division_def] - apply mul_ne_top (nts l₁ a) - apply inv_ne_top.mpr (nn l₂ a) - . rw [← δpmf_conv] - apply mul_ne_top (nts l₂ a) - apply inv_ne_top.mpr (MasterZero l₂) - - have S3 : ∑' (a : ↑{n | i = f n}), (nq l₁ ↑a / nq l₂ ↑a) ^ α * (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) a ≠ ⊤ := by - conv => - left - right - intro a - rw [← δpmf_conv] - rw [← mul_assoc] - rw [ENNReal.tsum_mul_right] - apply mul_ne_top - . rw [← RenyiDivergenceExpectation _ _ h1] - . replace nt := nt α h1 l₁ l₂ h2 - apply convergent_subset _ nt - . intro x - apply nn - . intro x - apply nts - . apply inv_ne_top.mpr (MasterZero l₂) - - have S4 : ∀ (a : ↑{n | i = f n}), (nq l₁ ↑a / nq l₂ ↑a) ^ α * (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) a ≠ ⊤ := by - intro a - apply ENNReal.ne_top_of_tsum_ne_top S3 - - rw [foo] - rw [foo] - - -- Introduce Q(f⁻¹ i) - let κ := ∑' x : {n : U | i = f n}, nq l₂ x - have P4 : κ / κ = 1 := by - rw [division_def] - rw [ENNReal.mul_inv_cancel] - . simp [κ] -- Use here for δ normalization - have T := witness NotEmpty - cases T - rename_i z w - exists z - constructor - . trivial - . apply nn l₂ - . simp only [κ] - apply MasterRW l₂ - - conv => - right - right - intro a - rw [← mul_one ((nq l₁ ↑a / nq l₂ ↑a) ^ α * nq l₂ ↑a)] - right - rw [← P4] - clear P4 - simp only [κ] - - conv => - right - right - intro a - right - rw [division_def] - rw [mul_comm] - - conv => - right - right - intro a - rw [← mul_assoc] - - rw [ENNReal.tsum_mul_right] - - -- Jensen's inequality - - have P5 : ∀ (x : ↑{n | i = f n}), 0 ≤ (fun a => (nq l₁ ↑a / nq l₂ ↑a).toReal) x := by - intro x - simp only [toReal_nonneg] - - have XXX : @Memℒp ℝ Real.normedAddCommGroup (↑{n | i = f n}) Subtype.instMeasurableSpace (fun a => (nq l₁ ↑a / nq l₂ ↑a).toReal) - (ENNReal.ofReal α) (PMF.toMeasure (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂))) := by - simp [Memℒp] - constructor - . apply MeasureTheory.StronglyMeasurable.aestronglyMeasurable - apply Measurable.stronglyMeasurable - apply Measurable.ennreal_toReal - conv => - right - intro x - rw [division_def] - apply Measurable.mul - . -- MeasurableSingletonClass.toDiscreteMeasurableSpace - apply measurable_discrete - . apply Measurable.inv - apply measurable_discrete - . simp [snorm] - split - . simp - . simp [snorm'] - rw [MeasureTheory.lintegral_countable'] -- Uses countable - rw [toReal_ofReal (le_of_lt (lt_trans zero_lt_one h1))] - have OTHER : ∀ a, nq l₁ a / nq l₂ a ≠ ⊤ := by - intro a - rw [division_def] - rw [ne_iff_lt_or_gt] - left - rw [mul_lt_top_iff] - left - constructor - . exact Ne.lt_top' (id (Ne.symm (nts l₁ a))) - . simp - exact pos_iff_ne_zero.mpr (nn l₂ a) - - conv => - left - left - right - intro a - rw [norm_simplify _ (OTHER a)] - have Z : 0 < α⁻¹ := by - simp - apply lt_trans zero_lt_one h1 - rw [rpow_lt_top_iff_of_pos Z] - conv => - left - right - intro a - rw [PMF.toMeasure_apply_singleton _ _ (measurableSet_singleton a)] - - apply Ne.lt_top' (id (Ne.symm _)) - apply S3 - - - have Jensen's := @bar {n : U | i = f n} Subtype.instMeasurableSpace Subtype.instMeasurableSingletonClass (fun a => (nq l₁ a / nq l₂ a).toReal) (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) α h1 P5 XXX - clear P5 - - have P6 : 0 ≤ (∑' (x : ↑{n | i = f n}), nq l₂ ↑x).toReal := by - simp only [toReal_nonneg] - have A' := mul_le_mul_of_nonneg_left Jensen's P6 - clear Jensen's P6 - - conv => - right - rw [mul_comm] - right - right - intro a - rw [mul_assoc] - rw [δpmf_conv] - - -- Here - - replace A' := ofReal_le_ofReal A' - rw [ofReal_mul toReal_nonneg] at A' - rw [ofReal_mul toReal_nonneg] at A' - rw [ofReal_toReal_eq_iff.2 (MasterRW l₂)] at A' - simp only at A' - - revert A' - conv => - left - right - right - right - right - intro x - rw [toReal_rpow] - rw [← toReal_mul] - conv => - left - right - right - right - rw [← ENNReal.tsum_toReal_eq S4] - intro A' - rw [ofReal_toReal_eq_iff.2 S3] at A' - - apply le_trans _ A' - clear A' - apply le_of_eq - - -- Part 3: - - conv => - right - right - right - left - right - intro x - rw [← toReal_mul] - rw [← ENNReal.tsum_toReal_eq S1] - rw [toReal_rpow] - rw [ofReal_toReal_eq_iff.2 S2] - - conv => - right - right - left - right - intro x - rw [division_def] - rw [← δpmf_conv] - rw [mul_assoc] - right - rw [← mul_assoc] - left - rw [ENNReal.inv_mul_cancel (nn l₂ x) (nts l₂ x)] - simp only [one_mul] - - rw [ENNReal.tsum_mul_right] - have H1 : 0 ≤ ∑' (x : ↑{n | i = f n}), (nq l₁ ↑x).toReal := by - apply tsum_nonneg - simp - have H2 : 0 ≤ (∑' (a : ↑{a | i = f a}), nq l₂ ↑a)⁻¹.toReal := by - apply toReal_nonneg - have H4 : (∑' (a : ↑{a | i = f a}), nq l₂ ↑a)⁻¹ ≠ ⊤ := by - apply inv_ne_top.mpr - simp - have T := witness NotEmpty - cases T - rename_i z w - exists z - constructor - . trivial - . apply nn l₂ - rw [ENNReal.mul_rpow_of_ne_top (MasterRW l₁) H4] - - have H3 : ∑' (a : ↑{a | i = f a}), nq l₂ ↑a ≠ 0 := by - simp - have T := witness NotEmpty - cases T - rename_i z w - exists z - constructor - . trivial - . apply nn l₂ - rw [ENNReal.rpow_sub _ _ H3 (MasterRW l₂)] - rw [ENNReal.rpow_one] - rw [division_def] - rw [← mul_assoc] - rw [← mul_assoc] - congr 1 - . rw [mul_comm] - . congr 1 - rw [ENNReal.inv_rpow] - -theorem tsum_ne_zero_of_ne_zero {T : Type} [Inhabited T] (f : T → ENNReal) (h : ∀ x : T, f x ≠ 0) : - ∑' x : T, f x ≠ 0 := by - by_contra CONTRA - rw [ENNReal.tsum_eq_zero] at CONTRA - have A := h default - have B := CONTRA default - contradiction - -variable [Inhabited U] - -theorem DPPostProcess {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP nq ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (conv : NonTopSum nq) (f : U → ℤ) : - DP (PostProcess nq f) ((ε₁ : ℝ) / ε₂) := by - simp [PostProcess, DP, RenyiDivergence] - intro α h1 l₁ l₂ h2 - have h' := h - simp [DP, RenyiDivergence] at h' - replace h' := h' α h1 l₁ l₂ h2 - - -- Part 1, removing fluff - - apply le_trans _ h' - clear h' - - -- remove the α scaling - have A : 0 ≤ (α - 1)⁻¹ := by - simp - apply le_of_lt h1 - apply mul_le_mul_of_nonneg_left _ A - clear A - - have RDConvegence : ∑' (x : U), nq l₁ x ^ α * nq l₂ x ^ (1 - α) ≠ ⊤ := by - simp [NonTopRDNQ] at nt - have nt := nt α h1 l₁ l₂ h2 - trivial - - have B := DPostPocess_pre h nn nt nts conv f h1 h2 - have B' : ∑' (x : ℤ), (∑' (a : U), if x = f a then nq l₁ a else 0) ^ α * (∑' (a : U), if x = f a then nq l₂ a else 0) ^ (1 - α) ≠ ⊤ := by - by_contra CONTRA - rw [CONTRA] at B - simp at B - contradiction - - -- remove the log - apply log_le_log _ (toReal_mono RDConvegence B) - apply toReal_pos _ B' - apply (tsum_ne_zero_iff ENNReal.summable).mpr - exists (f default) - - rw [ENNReal.tsum_eq_add_tsum_ite default] - conv => - left - right - rw [ENNReal.tsum_eq_add_tsum_ite default] - simp only [reduceIte] - apply mul_ne_zero - . by_contra CONTRA - rw [ENNReal.rpow_eq_zero_iff_of_pos (lt_trans zero_lt_one h1)] at CONTRA - simp at CONTRA - cases CONTRA - rename_i left right - have Y := nn l₁ default - contradiction - . by_contra CONTRA - rw [ENNReal.rpow_eq_zero_iff] at CONTRA - cases CONTRA - . rename_i CONTRA - cases CONTRA - rename_i left right - simp at left - cases left - rename_i le1 le2 - have Y := nn l₂ default - contradiction - . rename_i CONTRA - cases CONTRA - rename_i left right - simp at left - cases left - . rename_i left - have Y := nts l₂ default - contradiction - . rename_i left - have Rem := conv l₂ - have X : (∑' (x : U), if x = default then 0 else if f default = f x then nq l₂ x else 0) ≤ ∑' (n : U), nq l₂ n := by - apply ENNReal.tsum_le_tsum - intro a - split - . simp - . split - . simp - . simp - replace Rem := Ne.symm Rem - have Y := Ne.lt_top' Rem - have Z : (∑' (x : U), if x = default then 0 else if f default = f x then nq l₂ x else 0) < ⊤ := by - apply lt_of_le_of_lt X Y - rw [lt_top_iff_ne_top] at Z - contradiction - - -end SLang diff --git a/SampCert/DiffPrivacy/Neighbours.lean b/SampCert/DifferentialPrivacy/Neighbours.lean similarity index 100% rename from SampCert/DiffPrivacy/Neighbours.lean rename to SampCert/DifferentialPrivacy/Neighbours.lean diff --git a/SampCert/DiffPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean similarity index 100% rename from SampCert/DiffPrivacy/RenyiDivergence.lean rename to SampCert/DifferentialPrivacy/RenyiDivergence.lean diff --git a/SampCert/DiffPrivacy/Sensitivity.lean b/SampCert/DifferentialPrivacy/Sensitivity.lean similarity index 79% rename from SampCert/DiffPrivacy/Sensitivity.lean rename to SampCert/DifferentialPrivacy/Sensitivity.lean index ba903f7b..88b6f934 100644 --- a/SampCert/DiffPrivacy/Sensitivity.lean +++ b/SampCert/DifferentialPrivacy/Sensitivity.lean @@ -6,10 +6,8 @@ Authors: Jean-Baptiste Tristan import Mathlib.Topology.Algebra.InfiniteSum.Ring import Mathlib.Algebra.Group.Basic -import SampCert.DiffPrivacy.ConcentratedBound import SampCert.SLang -import SampCert.Samplers.GaussianGen.Basic -import SampCert.DiffPrivacy.Neighbours +import SampCert.DifferentialPrivacy.Neighbours open Classical Nat Int Real diff --git a/SampCert/DiffPrivacy/ConcentratedBound.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean similarity index 97% rename from SampCert/DiffPrivacy/ConcentratedBound.lean rename to SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean index e7ce1d6d..22636fb8 100644 --- a/SampCert/DiffPrivacy/ConcentratedBound.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean @@ -4,12 +4,12 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jean-Baptiste Tristan -/ -import SampCert.DiffPrivacy.DiscreteGaussian -import SampCert.DiffPrivacy.GaussBound -import SampCert.DiffPrivacy.GaussConvergence -import SampCert.DiffPrivacy.GaussPeriodicity +import SampCert.DifferentialPrivacy.ZeroConcentrated.DiscreteGaussian +import SampCert.DifferentialPrivacy.ZeroConcentrated.GaussBound +import SampCert.DifferentialPrivacy.ZeroConcentrated.GaussConvergence +import SampCert.DifferentialPrivacy.ZeroConcentrated.GaussPeriodicity import SampCert.Util.Shift -import SampCert.DiffPrivacy.RenyiDivergence +import SampCert.DifferentialPrivacy.RenyiDivergence open Real diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean new file mode 100644 index 00000000..b949e03e --- /dev/null +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean @@ -0,0 +1,40 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jean-Baptiste Tristan +-/ + +import Mathlib.Topology.Algebra.InfiniteSum.Basic +import Mathlib.Topology.Algebra.InfiniteSum.Ring +import Mathlib.Algebra.Group.Basic +import SampCert.DifferentialPrivacy.ZeroConcentrated.ConcentratedBound +import SampCert.SLang +import SampCert.Samplers.GaussianGen.Basic +import SampCert.DifferentialPrivacy.Neighbours +import SampCert.DifferentialPrivacy.Sensitivity +import Mathlib.MeasureTheory.MeasurableSpace.Basic +import Mathlib.MeasureTheory.Measure.Count +import Mathlib.Probability.ProbabilityMassFunction.Integrals +import Mathlib.Analysis.Convex.SpecificFunctions.Basic +import Mathlib.Analysis.Convex.Integral + +noncomputable section + +open Classical Nat Int Real ENNReal MeasureTheory Measure + +def DP (q : List T → SLang U) (ε : ℝ) : Prop := + ∀ α : ℝ, 1 < α → ∀ l₁ l₂ : List T, Neighbour l₁ l₂ → + RenyiDivergence (q l₁) (q l₂) α ≤ (1/2) * ε ^ 2 * α + +def NonZeroNQ (nq : List T → SLang U) := + ∀ l : List T, ∀ n : U, nq l n ≠ 0 + +def NonTopSum (nq : List T → SLang U) := + ∀ l : List T, ∑' n : U, nq l n ≠ ⊤ + +def NonTopNQ (nq : List T → SLang U) := + ∀ l : List T, ∀ n : U, nq l n ≠ ⊤ + +def NonTopRDNQ (nq : List T → SLang U) : Prop := + ∀ α : ℝ, 1 < α → ∀ l₁ l₂ : List T, Neighbour l₁ l₂ → + ∑' (x : U), nq l₁ x ^ α * nq l₂ x ^ (1 - α) ≠ ⊤ diff --git a/SampCert/DiffPrivacy/DiscreteGaussian.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/DiscreteGaussian.lean similarity index 100% rename from SampCert/DiffPrivacy/DiscreteGaussian.lean rename to SampCert/DifferentialPrivacy/ZeroConcentrated/DiscreteGaussian.lean diff --git a/SampCert/DiffPrivacy/GaussBound.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/GaussBound.lean similarity index 98% rename from SampCert/DiffPrivacy/GaussBound.lean rename to SampCert/DifferentialPrivacy/ZeroConcentrated/GaussBound.lean index 7377f806..40d51de2 100644 --- a/SampCert/DiffPrivacy/GaussBound.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/GaussBound.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jean-Baptiste Tristan -/ -import SampCert.DiffPrivacy.DiscreteGaussian +import SampCert.DifferentialPrivacy.ZeroConcentrated.DiscreteGaussian noncomputable section diff --git a/SampCert/DiffPrivacy/GaussConvergence.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/GaussConvergence.lean similarity index 92% rename from SampCert/DiffPrivacy/GaussConvergence.lean rename to SampCert/DifferentialPrivacy/ZeroConcentrated/GaussConvergence.lean index 288c47c7..69f23edf 100644 --- a/SampCert/DiffPrivacy/GaussConvergence.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/GaussConvergence.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jean-Baptiste Tristan -/ -import SampCert.DiffPrivacy.DiscreteGaussian +import SampCert.DifferentialPrivacy.ZeroConcentrated.DiscreteGaussian open Classical Nat Real diff --git a/SampCert/DiffPrivacy/GaussPeriodicity.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/GaussPeriodicity.lean similarity index 98% rename from SampCert/DiffPrivacy/GaussPeriodicity.lean rename to SampCert/DifferentialPrivacy/ZeroConcentrated/GaussPeriodicity.lean index f7a16c46..9f1256cb 100644 --- a/SampCert/DiffPrivacy/GaussPeriodicity.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/GaussPeriodicity.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jean-Baptiste Tristan -/ -import SampCert.DiffPrivacy.GaussConvergence +import SampCert.DifferentialPrivacy.ZeroConcentrated.GaussConvergence theorem SGShift (μ σ : ℝ) (n : ℤ) (k : ℤ) : (gauss_term_ℝ σ μ) (((n + k) : ℤ) : ℝ) = (gauss_term_ℝ σ (μ - k)) n := by diff --git a/SampCert/DiffPrivacy/BoundedAvg.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/BoundedMean.lean similarity index 86% rename from SampCert/DiffPrivacy/BoundedAvg.lean rename to SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/BoundedMean.lean index 3fcaa868..71af5711 100644 --- a/SampCert/DiffPrivacy/BoundedAvg.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/BoundedMean.lean @@ -4,9 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jean-Baptiste Tristan -/ -import SampCert.DiffPrivacy.DP -import SampCert.DiffPrivacy.Count -import SampCert.DiffPrivacy.BoundedSum +import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Composition.Composition +import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Postprocessing.Postprocessing +import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.Count.Count +import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.BoundedSum.BoundedSum open Classical Nat Int Real diff --git a/SampCert/DiffPrivacy/BoundedSum.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/BoundedSum.lean similarity index 74% rename from SampCert/DiffPrivacy/BoundedSum.lean rename to SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/BoundedSum.lean index 7f2b17ee..55b8413e 100644 --- a/SampCert/DiffPrivacy/BoundedSum.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/BoundedSum.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jean-Baptiste Tristan -/ -import SampCert.DiffPrivacy.DP +import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Mechanism.Mechanism import Mathlib.Algebra.Group.Defs import Mathlib.Init.Algebra.Classes import Init.Data.Int.Order @@ -18,43 +18,6 @@ namespace SLang def BoundedSumQuery (U : ℕ+) (l : List ℕ) : ℤ := List.sum (List.map (fun n : ℕ => (Nat.min n U)) l) --- theorem natAbs1 {n : ℕ} {x : ℤ} (h : n ≤ x) : --- n ≤ natAbs x := by --- cases x --- . rename_i m --- simp --- exact Int.ofNat_le.mp h --- . rename_i m --- simp --- have A : -[m+1] < 0 := negSucc_lt_zero m --- have B : n < (0 : ℤ) := by --- apply Int.lt_of_le_of_lt h A --- contradiction - --- def maxBoundPos (L U : ℤ) (h : L < U) : --- 0 < max (Int.natAbs L) (Int.natAbs U) := by --- have A : ¬ (L = 0 ∧ U = 0) := by --- by_contra h' --- cases h' --- rename_i h1 h2 --- subst h1 h2 --- contradiction --- simp at A --- have B : L = 0 ∨ L ≠ 0 := by exact eq_or_ne L 0 --- cases B --- . rename_i h' --- subst h' --- simp at A --- simp --- trivial --- . rename_i h' --- simp --- left --- trivial - --- def maxBound (L U : ℤ) (h : L < U) : ℕ+ := --- ⟨ max (Int.natAbs L) (Int.natAbs U) , maxBoundPos L U h ⟩ - theorem BoundedSumQuerySensitivity (U : ℕ+) : sensitivity (BoundedSumQuery U) U := by simp [sensitivity, BoundedSumQuery] intros l₁ l₂ H diff --git a/SampCert/DiffPrivacy/Count.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Count.lean similarity index 94% rename from SampCert/DiffPrivacy/Count.lean rename to SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Count.lean index 65902465..31846c1e 100644 --- a/SampCert/DiffPrivacy/Count.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Count.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jean-Baptiste Tristan -/ -import SampCert.DiffPrivacy.DP +import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Mechanism.Mechanism open Classical Nat Int Real diff --git a/SampCert/Samplers/Gaussian/Properties.lean b/SampCert/Samplers/Gaussian/Properties.lean index b10832b4..43def706 100644 --- a/SampCert/Samplers/Gaussian/Properties.lean +++ b/SampCert/Samplers/Gaussian/Properties.lean @@ -12,7 +12,7 @@ import SampCert.Samplers.Laplace.Basic import Mathlib.NumberTheory.ModularForms.JacobiTheta.OneVariable import SampCert.Util.UtilMathlib import SampCert.Samplers.Gaussian.Code -import SampCert.DiffPrivacy.DiscreteGaussian +import SampCert.DifferentialPrivacy.ZeroConcentrated.DiscreteGaussian noncomputable section diff --git a/SampCert/Samplers/GaussianGen/Properties.lean b/SampCert/Samplers/GaussianGen/Properties.lean index e8027703..75c8d6b3 100644 --- a/SampCert/Samplers/GaussianGen/Properties.lean +++ b/SampCert/Samplers/GaussianGen/Properties.lean @@ -6,9 +6,9 @@ Authors: Jean-Baptiste Tristan import SampCert.Samplers.GaussianGen.Code import SampCert.Samplers.Gaussian.Properties -import SampCert.DiffPrivacy.DiscreteGaussian -import SampCert.DiffPrivacy.GaussPeriodicity -import SampCert.DiffPrivacy.ConcentratedBound +import SampCert.DifferentialPrivacy.ZeroConcentrated.DiscreteGaussian +import SampCert.DifferentialPrivacy.ZeroConcentrated.GaussPeriodicity +import SampCert.DifferentialPrivacy.ZeroConcentrated.ConcentratedBound noncomputable section From 1b87266125876f560144d94d3306b21711fa05fc Mon Sep 17 00:00:00 2001 From: jtristan Date: Wed, 15 May 2024 15:11:39 -0400 Subject: [PATCH 17/20] Missing files --- .../Foundations/Composition/Basic.lean | 8 + .../Foundations/Composition/Code.lean | 20 + .../Foundations/Composition/Composition.lean | 317 ++++++++ .../Foundations/Composition/Properties.lean | 319 ++++++++ .../Foundations/Mechanism/Basic.lean | 8 + .../Foundations/Mechanism/Code.lean | 18 + .../Foundations/Mechanism/Mechanism.lean | 242 ++++++ .../Foundations/Mechanism/Properties.lean | 248 +++++++ .../Foundations/Postprocessing/Basic.lean | 8 + .../Foundations/Postprocessing/Code.lean | 19 + .../Postprocessing/Postprocessing.lean | 691 +++++++++++++++++ .../Postprocessing/Properties.lean | 697 ++++++++++++++++++ .../Queries/BoundedMean/Basic.lean | 8 + .../Queries/BoundedMean/Code.lean | 23 + .../Queries/BoundedMean/Properties.lean | 56 ++ .../Queries/BoundedSum/Basic.lean | 8 + .../Queries/BoundedSum/Code.lean | 21 + .../Queries/BoundedSum/Properties.lean | 96 +++ .../ZeroConcentrated/Queries/Count/Basic.lean | 8 + .../ZeroConcentrated/Queries/Count/Code.lean | 19 + .../Queries/Count/Properties.lean | 57 ++ 21 files changed, 2891 insertions(+) create mode 100644 SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Composition/Basic.lean create mode 100644 SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Composition/Code.lean create mode 100644 SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Composition/Composition.lean create mode 100644 SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Composition/Properties.lean create mode 100644 SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Mechanism/Basic.lean create mode 100644 SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Mechanism/Code.lean create mode 100644 SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Mechanism/Mechanism.lean create mode 100644 SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Mechanism/Properties.lean create mode 100644 SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Basic.lean create mode 100644 SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Code.lean create mode 100644 SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Postprocessing.lean create mode 100644 SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Properties.lean create mode 100644 SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/Basic.lean create mode 100644 SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/Code.lean create mode 100644 SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/Properties.lean create mode 100644 SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/Basic.lean create mode 100644 SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/Code.lean create mode 100644 SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/Properties.lean create mode 100644 SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Basic.lean create mode 100644 SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Code.lean create mode 100644 SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Properties.lean diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Composition/Basic.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Composition/Basic.lean new file mode 100644 index 00000000..abcaf138 --- /dev/null +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Composition/Basic.lean @@ -0,0 +1,8 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jean-Baptiste Tristan +-/ + +import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Composition.Code +import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Composition.Properties diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Composition/Code.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Composition/Code.lean new file mode 100644 index 00000000..d6d195e8 --- /dev/null +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Composition/Code.lean @@ -0,0 +1,20 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jean-Baptiste Tristan +-/ + +import SampCert.DifferentialPrivacy.ZeroConcentrated.DP + +noncomputable section + +open Classical Nat Int Real ENNReal MeasureTheory Measure + +namespace SLang + +def Compose (nq1 nq2 : List T → SLang ℤ) (l : List T) : SLang (ℤ × ℤ) := do + let A ← nq1 l + let B ← nq2 l + return (A,B) + +end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Composition/Composition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Composition/Composition.lean new file mode 100644 index 00000000..28e39a5c --- /dev/null +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Composition/Composition.lean @@ -0,0 +1,317 @@ +import SampCert.DifferentialPrivacy.ZeroConcentrated.DP + +noncomputable section + +open Classical Nat Int Real ENNReal MeasureTheory Measure + +namespace SLang + +def Compose (nq1 nq2 : List T → SLang ℤ) (l : List T) : SLang (ℤ × ℤ) := do + let A ← nq1 l + let B ← nq2 l + return (A,B) + +theorem ENNReal_toTeal_NZ (x : ENNReal) (h1 : x ≠ 0) (h2 : x ≠ ⊤) : + x.toReal ≠ 0 := by + unfold ENNReal.toReal + unfold ENNReal.toNNReal + simp + intro H + cases H + . contradiction + . contradiction + +theorem simp_α_1 {α : ℝ} (h : 1 < α) : 0 < α := by + apply @lt_trans _ _ _ 1 _ _ h + simp only [zero_lt_one] + +theorem RenyiNoisedQueryNonZero {nq : List T → SLang ℤ} {α ε : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (h2 : Neighbour l₁ l₂) (h3 : DP nq ε) (h4 : NonZeroNQ nq) (h5 : NonTopRDNQ nq) (nts : NonTopNQ nq) : + (∑' (i : ℤ), nq l₁ i ^ α * nq l₂ i ^ (1 - α)).toReal ≠ 0 := by + simp [DP] at h3 + replace h3 := h3 α h1 l₁ l₂ h2 + simp [RenyiDivergence] at h3 + simp [NonZeroNQ] at h4 + simp [NonTopRDNQ] at h5 + replace h5 := h5 α h1 l₁ l₂ h2 + have h6 := h4 l₁ + have h7 := h4 l₂ + apply ENNReal_toTeal_NZ + . by_contra CONTRA + rw [ENNReal.tsum_eq_zero] at CONTRA + replace CONTRA := CONTRA 42 + replace h6 := h6 42 + replace h7 := h7 42 + rw [_root_.mul_eq_zero] at CONTRA + cases CONTRA + . rename_i h8 + rw [rpow_eq_zero_iff_of_pos] at h8 + contradiction + apply simp_α_1 h1 + . rename_i h8 + rw [ENNReal.rpow_eq_zero_iff] at h8 + cases h8 + . rename_i h8 + cases h8 + contradiction + . rename_i h8 + cases h8 + rename_i h8 h9 + replace nts := nts l₂ 42 + contradiction + . exact h5 + +theorem compose_sum_rw (nq1 nq2 : List T → SLang ℤ) (b c : ℤ) (l : List T) : + (∑' (a : ℤ), nq1 l a * ∑' (a_1 : ℤ), if b = a ∧ c = a_1 then nq2 l a_1 else 0) = nq1 l b * nq2 l c := by + have A : ∀ a b : ℤ, (∑' (a_1 : ℤ), if b = a ∧ c = a_1 then nq2 l a_1 else 0) = if b = a then (∑' (a_1 : ℤ), if c = a_1 then nq2 l a_1 else 0) else 0 := by + intro x y + split + . rename_i h + subst h + simp + . rename_i h + simp + intro h + contradiction + conv => + left + right + intro a + right + rw [A] + rw [ENNReal.tsum_eq_add_tsum_ite b] + simp + have B : ∀ x : ℤ, (@ite ℝ≥0∞ (x = b) (instDecidableEqInt x b) 0 + (@ite ℝ≥0∞ (b = x) (instDecidableEqInt b x) (nq1 l x * ∑' (a_1 : ℤ), @ite ℝ≥0∞ (c = a_1) (instDecidableEqInt c a_1) (nq2 l a_1) 0) 0)) = 0 := by + intro x + split + . simp + . split + . rename_i h1 h2 + subst h2 + contradiction + . simp + conv => + left + right + right + intro x + rw [B] + simp + congr 1 + rw [ENNReal.tsum_eq_add_tsum_ite c] + simp + have C :∀ x : ℤ, (@ite ℝ≥0∞ (x = c) (propDecidable (x = c)) 0 (@ite ℝ≥0∞ (c = x) (instDecidableEqInt c x) (nq2 l x) 0)) = 0 := by + intro x + split + . simp + . split + . rename_i h1 h2 + subst h2 + contradiction + . simp + conv => + left + right + right + intro X + rw [C] + simp + +theorem DPCompose {nq1 nq2 : List T → SLang ℤ} {ε₁ ε₂ ε₃ ε₄ : ℕ+} (h1 : DP nq1 ((ε₁ : ℝ) / ε₂)) (h2 : DP nq2 ((ε₃ : ℝ) / ε₄)) (nn1 : NonZeroNQ nq1) (nn2 : NonZeroNQ nq2) (nt1 : NonTopRDNQ nq1) (nt2 : NonTopRDNQ nq2) (nts1 : NonTopNQ nq1) (nts2 : NonTopNQ nq2) : + DP (Compose nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by + simp [Compose, RenyiDivergence, DP] + intro α h3 l₁ l₂ h4 + have X := h1 + have Y := h2 + simp [DP] at h1 h2 + replace h1 := h1 α h3 l₁ l₂ h4 + replace h2 := h2 α h3 l₁ l₂ h4 + simp [RenyiDivergence] at h1 h2 + rw [tsum_prod' ENNReal.summable (fun b : ℤ => ENNReal.summable)] + . simp + conv => + left + right + right + right + right + intro b + right + intro c + rw [compose_sum_rw] + rw [compose_sum_rw] + rw [ENNReal.mul_rpow_of_ne_zero (nn1 l₁ b) (nn2 l₁ c)] + rw [ENNReal.mul_rpow_of_ne_zero (nn1 l₂ b) (nn2 l₂ c)] + rw [mul_assoc] + right + rw [mul_comm] + rw [mul_assoc] + right + rw [mul_comm] + conv => + left + right + right + right + right + intro b + right + intro c + rw [← mul_assoc] + conv => + left + right + right + right + right + intro b + rw [ENNReal.tsum_mul_left] + rw [ENNReal.tsum_mul_right] + rw [ENNReal.toReal_mul] + rw [Real.log_mul] + . rw [mul_add] + have D := _root_.add_le_add h1 h2 + apply le_trans D + rw [← add_mul] + rw [mul_le_mul_iff_of_pos_right] + . rw [← mul_add] + rw [mul_le_mul_iff_of_pos_left] + . ring_nf + simp + . simp + . apply lt_trans zero_lt_one h3 + . apply RenyiNoisedQueryNonZero h3 h4 X nn1 nt1 nts1 + . apply RenyiNoisedQueryNonZero h3 h4 Y nn2 nt2 nts2 + +theorem DPCompose_NonZeroNQ {nq1 nq2 : List T → SLang ℤ} (nn1 : NonZeroNQ nq1) (nn2 : NonZeroNQ nq2) : + NonZeroNQ (Compose nq1 nq2) := by + simp [NonZeroNQ] at * + intro l a b + replace nn1 := nn1 l a + replace nn2 := nn2 l b + simp [Compose] + exists a + simp + intro H + cases H + . rename_i H + contradiction + . rename_i H + contradiction + +theorem DPCompose_NonTopNQ {nq1 nq2 : List T → SLang ℤ} (nt1 : NonTopNQ nq1) (nt2 : NonTopNQ nq2) : + NonTopNQ (Compose nq1 nq2) := by + simp [NonTopNQ] at * + intro l a b + replace nt1 := nt1 l a + replace nt2 := nt2 l b + simp [Compose] + rw [compose_sum_rw] + rw [mul_eq_top] + intro H + cases H + . rename_i H + cases H + contradiction + . rename_i H + cases H + contradiction + +theorem DPCompose_NonTopSum {nq1 nq2 : List T → SLang ℤ} (nt1 : NonTopSum nq1) (nt2 : NonTopSum nq2) : + NonTopSum (Compose nq1 nq2) := by + simp [NonTopSum] at * + intro l + replace nt1 := nt1 l + replace nt2 := nt2 l + simp [Compose] + rw [ENNReal.tsum_prod'] + conv => + right + left + right + intro a + right + intro b + simp + rw [compose_sum_rw] + conv => + right + left + right + intro a + rw [ENNReal.tsum_mul_left] + rw [ENNReal.tsum_mul_right] + rw [mul_eq_top] + intro H + cases H + . rename_i H + cases H + contradiction + . rename_i H + cases H + contradiction + +theorem DPCompose_NonTopRDNQ {nq1 nq2 : List T → SLang ℤ} (nt1 : NonTopRDNQ nq1) (nt2 : NonTopRDNQ nq2) (nn1 : NonZeroNQ nq1) (nn2 : NonZeroNQ nq2) : + NonTopRDNQ (Compose nq1 nq2) := by + simp [NonTopRDNQ] at * + intro α h1 l₁ l₂ h2 + replace nt1 := nt1 α h1 l₁ l₂ h2 + replace nt2 := nt2 α h1 l₁ l₂ h2 + simp [Compose] + rw [ENNReal.tsum_prod'] + simp + conv => + right + left + right + intro x + right + intro y + congr + . left + rw [compose_sum_rw] + . left + rw [compose_sum_rw] + conv => + right + left + right + intro x + right + intro y + rw [ENNReal.mul_rpow_of_ne_zero (nn1 l₁ x) (nn2 l₁ y)] + rw [ENNReal.mul_rpow_of_ne_zero (nn1 l₂ x) (nn2 l₂ y)] + rw [mul_assoc] + right + rw [mul_comm] + rw [mul_assoc] + right + rw [mul_comm] + conv => + right + left + right + intro x + right + intro y + rw [← mul_assoc] + conv => + right + left + right + intro x + rw [ENNReal.tsum_mul_left] + rw [ENNReal.tsum_mul_right] + intro H + rw [mul_eq_top] at H + cases H + . rename_i h3 + cases h3 + rename_i h4 h5 + contradiction + . rename_i h3 + cases h3 + rename_i h4 h5 + contradiction + +end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Composition/Properties.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Composition/Properties.lean new file mode 100644 index 00000000..12f7efb2 --- /dev/null +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Composition/Properties.lean @@ -0,0 +1,319 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jean-Baptiste Tristan +-/ + +import SampCert.DifferentialPrivacy.ZeroConcentrated.DP +import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Composition.Code + +noncomputable section + +open Classical Nat Int Real ENNReal MeasureTheory Measure + +namespace SLang + +theorem ENNReal_toTeal_NZ (x : ENNReal) (h1 : x ≠ 0) (h2 : x ≠ ⊤) : + x.toReal ≠ 0 := by + unfold ENNReal.toReal + unfold ENNReal.toNNReal + simp + intro H + cases H + . contradiction + . contradiction + +theorem simp_α_1 {α : ℝ} (h : 1 < α) : 0 < α := by + apply @lt_trans _ _ _ 1 _ _ h + simp only [zero_lt_one] + +theorem RenyiNoisedQueryNonZero {nq : List T → SLang ℤ} {α ε : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (h2 : Neighbour l₁ l₂) (h3 : DP nq ε) (h4 : NonZeroNQ nq) (h5 : NonTopRDNQ nq) (nts : NonTopNQ nq) : + (∑' (i : ℤ), nq l₁ i ^ α * nq l₂ i ^ (1 - α)).toReal ≠ 0 := by + simp [DP] at h3 + replace h3 := h3 α h1 l₁ l₂ h2 + simp [RenyiDivergence] at h3 + simp [NonZeroNQ] at h4 + simp [NonTopRDNQ] at h5 + replace h5 := h5 α h1 l₁ l₂ h2 + have h6 := h4 l₁ + have h7 := h4 l₂ + apply ENNReal_toTeal_NZ + . by_contra CONTRA + rw [ENNReal.tsum_eq_zero] at CONTRA + replace CONTRA := CONTRA 42 + replace h6 := h6 42 + replace h7 := h7 42 + rw [_root_.mul_eq_zero] at CONTRA + cases CONTRA + . rename_i h8 + rw [rpow_eq_zero_iff_of_pos] at h8 + contradiction + apply simp_α_1 h1 + . rename_i h8 + rw [ENNReal.rpow_eq_zero_iff] at h8 + cases h8 + . rename_i h8 + cases h8 + contradiction + . rename_i h8 + cases h8 + rename_i h8 h9 + replace nts := nts l₂ 42 + contradiction + . exact h5 + +theorem compose_sum_rw (nq1 nq2 : List T → SLang ℤ) (b c : ℤ) (l : List T) : + (∑' (a : ℤ), nq1 l a * ∑' (a_1 : ℤ), if b = a ∧ c = a_1 then nq2 l a_1 else 0) = nq1 l b * nq2 l c := by + have A : ∀ a b : ℤ, (∑' (a_1 : ℤ), if b = a ∧ c = a_1 then nq2 l a_1 else 0) = if b = a then (∑' (a_1 : ℤ), if c = a_1 then nq2 l a_1 else 0) else 0 := by + intro x y + split + . rename_i h + subst h + simp + . rename_i h + simp + intro h + contradiction + conv => + left + right + intro a + right + rw [A] + rw [ENNReal.tsum_eq_add_tsum_ite b] + simp + have B : ∀ x : ℤ, (@ite ℝ≥0∞ (x = b) (instDecidableEqInt x b) 0 + (@ite ℝ≥0∞ (b = x) (instDecidableEqInt b x) (nq1 l x * ∑' (a_1 : ℤ), @ite ℝ≥0∞ (c = a_1) (instDecidableEqInt c a_1) (nq2 l a_1) 0) 0)) = 0 := by + intro x + split + . simp + . split + . rename_i h1 h2 + subst h2 + contradiction + . simp + conv => + left + right + right + intro x + rw [B] + simp + congr 1 + rw [ENNReal.tsum_eq_add_tsum_ite c] + simp + have C :∀ x : ℤ, (@ite ℝ≥0∞ (x = c) (propDecidable (x = c)) 0 (@ite ℝ≥0∞ (c = x) (instDecidableEqInt c x) (nq2 l x) 0)) = 0 := by + intro x + split + . simp + . split + . rename_i h1 h2 + subst h2 + contradiction + . simp + conv => + left + right + right + intro X + rw [C] + simp + +theorem DPCompose {nq1 nq2 : List T → SLang ℤ} {ε₁ ε₂ ε₃ ε₄ : ℕ+} (h1 : DP nq1 ((ε₁ : ℝ) / ε₂)) (h2 : DP nq2 ((ε₃ : ℝ) / ε₄)) (nn1 : NonZeroNQ nq1) (nn2 : NonZeroNQ nq2) (nt1 : NonTopRDNQ nq1) (nt2 : NonTopRDNQ nq2) (nts1 : NonTopNQ nq1) (nts2 : NonTopNQ nq2) : + DP (Compose nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by + simp [Compose, RenyiDivergence, DP] + intro α h3 l₁ l₂ h4 + have X := h1 + have Y := h2 + simp [DP] at h1 h2 + replace h1 := h1 α h3 l₁ l₂ h4 + replace h2 := h2 α h3 l₁ l₂ h4 + simp [RenyiDivergence] at h1 h2 + rw [tsum_prod' ENNReal.summable (fun b : ℤ => ENNReal.summable)] + . simp + conv => + left + right + right + right + right + intro b + right + intro c + rw [compose_sum_rw] + rw [compose_sum_rw] + rw [ENNReal.mul_rpow_of_ne_zero (nn1 l₁ b) (nn2 l₁ c)] + rw [ENNReal.mul_rpow_of_ne_zero (nn1 l₂ b) (nn2 l₂ c)] + rw [mul_assoc] + right + rw [mul_comm] + rw [mul_assoc] + right + rw [mul_comm] + conv => + left + right + right + right + right + intro b + right + intro c + rw [← mul_assoc] + conv => + left + right + right + right + right + intro b + rw [ENNReal.tsum_mul_left] + rw [ENNReal.tsum_mul_right] + rw [ENNReal.toReal_mul] + rw [Real.log_mul] + . rw [mul_add] + have D := _root_.add_le_add h1 h2 + apply le_trans D + rw [← add_mul] + rw [mul_le_mul_iff_of_pos_right] + . rw [← mul_add] + rw [mul_le_mul_iff_of_pos_left] + . ring_nf + simp + . simp + . apply lt_trans zero_lt_one h3 + . apply RenyiNoisedQueryNonZero h3 h4 X nn1 nt1 nts1 + . apply RenyiNoisedQueryNonZero h3 h4 Y nn2 nt2 nts2 + +theorem DPCompose_NonZeroNQ {nq1 nq2 : List T → SLang ℤ} (nn1 : NonZeroNQ nq1) (nn2 : NonZeroNQ nq2) : + NonZeroNQ (Compose nq1 nq2) := by + simp [NonZeroNQ] at * + intro l a b + replace nn1 := nn1 l a + replace nn2 := nn2 l b + simp [Compose] + exists a + simp + intro H + cases H + . rename_i H + contradiction + . rename_i H + contradiction + +theorem DPCompose_NonTopNQ {nq1 nq2 : List T → SLang ℤ} (nt1 : NonTopNQ nq1) (nt2 : NonTopNQ nq2) : + NonTopNQ (Compose nq1 nq2) := by + simp [NonTopNQ] at * + intro l a b + replace nt1 := nt1 l a + replace nt2 := nt2 l b + simp [Compose] + rw [compose_sum_rw] + rw [mul_eq_top] + intro H + cases H + . rename_i H + cases H + contradiction + . rename_i H + cases H + contradiction + +theorem DPCompose_NonTopSum {nq1 nq2 : List T → SLang ℤ} (nt1 : NonTopSum nq1) (nt2 : NonTopSum nq2) : + NonTopSum (Compose nq1 nq2) := by + simp [NonTopSum] at * + intro l + replace nt1 := nt1 l + replace nt2 := nt2 l + simp [Compose] + rw [ENNReal.tsum_prod'] + conv => + right + left + right + intro a + right + intro b + simp + rw [compose_sum_rw] + conv => + right + left + right + intro a + rw [ENNReal.tsum_mul_left] + rw [ENNReal.tsum_mul_right] + rw [mul_eq_top] + intro H + cases H + . rename_i H + cases H + contradiction + . rename_i H + cases H + contradiction + +theorem DPCompose_NonTopRDNQ {nq1 nq2 : List T → SLang ℤ} (nt1 : NonTopRDNQ nq1) (nt2 : NonTopRDNQ nq2) (nn1 : NonZeroNQ nq1) (nn2 : NonZeroNQ nq2) : + NonTopRDNQ (Compose nq1 nq2) := by + simp [NonTopRDNQ] at * + intro α h1 l₁ l₂ h2 + replace nt1 := nt1 α h1 l₁ l₂ h2 + replace nt2 := nt2 α h1 l₁ l₂ h2 + simp [Compose] + rw [ENNReal.tsum_prod'] + simp + conv => + right + left + right + intro x + right + intro y + congr + . left + rw [compose_sum_rw] + . left + rw [compose_sum_rw] + conv => + right + left + right + intro x + right + intro y + rw [ENNReal.mul_rpow_of_ne_zero (nn1 l₁ x) (nn2 l₁ y)] + rw [ENNReal.mul_rpow_of_ne_zero (nn1 l₂ x) (nn2 l₂ y)] + rw [mul_assoc] + right + rw [mul_comm] + rw [mul_assoc] + right + rw [mul_comm] + conv => + right + left + right + intro x + right + intro y + rw [← mul_assoc] + conv => + right + left + right + intro x + rw [ENNReal.tsum_mul_left] + rw [ENNReal.tsum_mul_right] + intro H + rw [mul_eq_top] at H + cases H + . rename_i h3 + cases h3 + rename_i h4 h5 + contradiction + . rename_i h3 + cases h3 + rename_i h4 h5 + contradiction + +end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Mechanism/Basic.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Mechanism/Basic.lean new file mode 100644 index 00000000..4f80389b --- /dev/null +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Mechanism/Basic.lean @@ -0,0 +1,8 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jean-Baptiste Tristan +-/ + +import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Mechanism.Code +import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Mechanism.Properties diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Mechanism/Code.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Mechanism/Code.lean new file mode 100644 index 00000000..71f57091 --- /dev/null +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Mechanism/Code.lean @@ -0,0 +1,18 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jean-Baptiste Tristan +-/ + +import SampCert.DifferentialPrivacy.ZeroConcentrated.DP + +noncomputable section + +open Classical Nat Int Real ENNReal MeasureTheory Measure + +namespace SLang + +def NoisedQuery (query : List T → ℤ) (Δ : ℕ+) (ε₁ ε₂ : ℕ+) (l : List T) : SLang ℤ := do + DiscreteGaussianGenSample (Δ * ε₂) ε₁ (query l) + +end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Mechanism/Mechanism.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Mechanism/Mechanism.lean new file mode 100644 index 00000000..29df6fb3 --- /dev/null +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Mechanism/Mechanism.lean @@ -0,0 +1,242 @@ +import SampCert.DifferentialPrivacy.ZeroConcentrated.DP + +noncomputable section + +open Classical Nat Int Real ENNReal MeasureTheory Measure + +namespace SLang + +def NoisedQuery (query : List T → ℤ) (Δ : ℕ+) (ε₁ ε₂ : ℕ+) (l : List T) : SLang ℤ := do + DiscreteGaussianGenSample (Δ * ε₂) ε₁ (query l) + +theorem NoisedQueryDP (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_sensitivity : sensitivity query Δ) : + DP (NoisedQuery query Δ ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by + simp [DP, NoisedQuery] + intros α h1 l₁ l₂ h2 + have A := @DiscreteGaussianGenSampleZeroConcentrated α h1 (Δ * ε₂) ε₁ (query l₁) (query l₂) + apply le_trans A + clear A + replace bounded_sensitivity := bounded_sensitivity l₁ l₂ h2 + ring_nf + simp + conv => + left + left + right + rw [mul_pow] + conv => + left + rw [mul_assoc] + right + rw [mul_comm] + rw [← mul_assoc] + conv => + left + rw [mul_assoc] + right + rw [← mul_assoc] + left + rw [mul_comm] + rw [← mul_assoc] + rw [← mul_assoc] + rw [← mul_assoc] + simp only [inv_pow] + rw [mul_inv_le_iff'] + . have A : (α * ↑↑ε₁ ^ 2 * (↑↑ε₂ ^ 2)⁻¹) ≤ (α * ↑↑ε₁ ^ 2 * (↑↑ε₂ ^ 2)⁻¹) := le_refl (α * ↑↑ε₁ ^ 2 * (↑↑ε₂ ^ 2)⁻¹) + have B : 0 ≤ (α * ↑↑ε₁ ^ 2 * (↑↑ε₂ ^ 2)⁻¹) := by + simp + apply @le_trans ℝ _ 0 1 α (instStrictOrderedCommRingReal.proof_3) (le_of_lt h1) + apply mul_le_mul A _ _ B + . apply sq_le_sq.mpr + simp only [abs_cast] + rw [← Int.cast_sub] + rw [← Int.cast_abs] + apply Int.cast_le.mpr + rw [← Int.natCast_natAbs] + apply Int.ofNat_le.mpr + trivial + . apply sq_nonneg + . rw [pow_two] + rw [_root_.mul_pos_iff] + left + simp + +theorem NoisedQuery_NonZeroNQ (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : + NonZeroNQ (NoisedQuery query Δ ε₁ ε₂) := by + simp [NonZeroNQ, NoisedQuery, DiscreteGaussianGenSample] + intros l n + exists (n - query l) + simp + have A : ((Δ : ℝ) * ε₂ / ε₁) ≠ 0 := by + simp + have X := @discrete_gaussian_pos (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) A 0 (n - query l) + simp at X + trivial + +theorem NoisedQuery_NonTopNQ (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : + NonTopNQ (NoisedQuery query Δ ε₁ ε₂) := by + simp [NonTopNQ, NoisedQuery, DiscreteGaussianGenSample] + intro l n + rw [ENNReal.tsum_eq_add_tsum_ite (n - query l)] + simp + have X : ∀ x : ℤ, (@ite ℝ≥0∞ (x = n - query l) (propDecidable (x = n - query l)) 0 + (@ite ℝ≥0∞ (n = x + query l) (instDecidableEqInt n (x + query l)) + (ENNReal.ofReal (discrete_gaussian (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) 0 ↑x)) 0)) = 0 := by + intro x + split + . simp + . split + . rename_i h1 h2 + subst h2 + simp at h1 + . simp + conv => + right + left + right + intro x + rw [X] + simp + +theorem discrete_gaussian_shift {σ : ℝ} (h : σ ≠ 0) (μ : ℝ) (τ x : ℤ) : + discrete_gaussian σ μ (x - τ) = discrete_gaussian σ (μ + τ) (x) := by + simp [discrete_gaussian] + congr 1 + . simp [gauss_term_ℝ] + congr 3 + ring_nf + . rw [SG_periodic h] + +theorem NoisedQuery_NonTopSum (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : + NonTopSum (NoisedQuery query Δ ε₁ ε₂) := by + simp [NonTopSum, NoisedQuery, DiscreteGaussianGenSample] + intro l + have X : ∀ n: ℤ, ∀ x : ℤ, (@ite ℝ≥0∞ (x = n - query l) (propDecidable (x = n - query l)) 0 + (@ite ℝ≥0∞ (n = x + query l) (instDecidableEqInt n (x + query l)) + (ENNReal.ofReal (discrete_gaussian (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) 0 ↑x)) 0)) = 0 := by + intro n x + split + . simp + . split + . rename_i h1 h2 + subst h2 + simp at h1 + . simp + conv => + right + left + right + intro n + rw [ENNReal.tsum_eq_add_tsum_ite (n - query l)] + simp + right + right + intro x + rw [X] + simp + have A : (Δ : ℝ) * ε₂ / ε₁ ≠ 0 := by + simp + conv => + right + left + right + intro n + rw [discrete_gaussian_shift A] + simp + rw [← ENNReal.ofReal_tsum_of_nonneg] + . rw [discrete_gaussian_normalizes A] + simp + . apply discrete_gaussian_nonneg A + . apply discrete_gaussian_summable' A (query l) + +theorem NoisedQuery_NonTopRDNQ (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : + NonTopRDNQ (NoisedQuery query Δ ε₁ ε₂) := by + simp [NonTopRDNQ, NoisedQuery, DiscreteGaussianGenSample] + intro α _ l₁ l₂ _ + have A : ∀ x_1 x : ℤ, (@ite ℝ≥0∞ (x_1 = x - query l₁) (propDecidable (x_1 = x - query l₁)) 0 + (@ite ℝ≥0∞ (x = x_1 + query l₁) (instDecidableEqInt x (x_1 + query l₁)) + (ENNReal.ofReal (discrete_gaussian (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) 0 ↑x_1)) 0 )) = 0 := by + intro x y + split + . simp + . split + . rename_i h1 h2 + subst h2 + simp at h1 + . simp + have B : ∀ x_1 x : ℤ, (@ite ℝ≥0∞ (x_1 = x - query l₂) (propDecidable (x_1 = x - query l₂)) 0 + (@ite ℝ≥0∞ (x = x_1 + query l₂) (instDecidableEqInt x (x_1 + query l₂)) + (ENNReal.ofReal (discrete_gaussian (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) 0 ↑x_1)) 0)) = 0 := by + intro x y + split + . simp + . split + . rename_i h1 h2 + subst h2 + simp at h1 + . simp + conv => + right + left + right + intro x + left + rw [ENNReal.tsum_eq_add_tsum_ite (x - query l₁)] + simp + left + right + right + intro y + rw [A] + simp + conv => + right + left + right + intro x + right + rw [ENNReal.tsum_eq_add_tsum_ite (x - query l₂)] + simp + left + right + right + intro y + rw [B] + simp + clear A B + have P : (Δ : ℝ) * ε₂ / ε₁ ≠ 0 := by + simp + have A : ∀ x : ℤ, ∀ l : List T, 0 < discrete_gaussian (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) 0 (↑x - ↑(query l)) := by + intro x l + have A' := @discrete_gaussian_pos _ P 0 (x - query l) + simp at A' + trivial + have B : ∀ x : ℤ, 0 ≤ discrete_gaussian (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) 0 (↑x - ↑(query l₁)) ^ α := by + intro x + have A' := @discrete_gaussian_nonneg _ P 0 (x - query l₁) + simp at A' + apply rpow_nonneg A' + conv => + right + left + right + intro x + rw [ENNReal.ofReal_rpow_of_pos (A x l₁)] + rw [ENNReal.ofReal_rpow_of_pos (A x l₂)] + rw [← ENNReal.ofReal_mul (B x)] + rw [← ENNReal.ofReal_tsum_of_nonneg] + . simp + . intro n + have X := @RenyiSumSG_nonneg _ α P (query l₁) (query l₂) n + rw [discrete_gaussian_shift P] + rw [discrete_gaussian_shift P] + simp [X] + . have X := @SummableRenyiGauss _ P (query l₁) (query l₂) + conv => + right + intro x + rw [discrete_gaussian_shift P] + rw [discrete_gaussian_shift P] + simp [X] + +end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Mechanism/Properties.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Mechanism/Properties.lean new file mode 100644 index 00000000..764842dc --- /dev/null +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Mechanism/Properties.lean @@ -0,0 +1,248 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jean-Baptiste Tristan +-/ + +import SampCert.DifferentialPrivacy.ZeroConcentrated.DP + +noncomputable section + +open Classical Nat Int Real ENNReal MeasureTheory Measure + +namespace SLang + +def NoisedQuery (query : List T → ℤ) (Δ : ℕ+) (ε₁ ε₂ : ℕ+) (l : List T) : SLang ℤ := do + DiscreteGaussianGenSample (Δ * ε₂) ε₁ (query l) + +theorem NoisedQueryDP (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_sensitivity : sensitivity query Δ) : + DP (NoisedQuery query Δ ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by + simp [DP, NoisedQuery] + intros α h1 l₁ l₂ h2 + have A := @DiscreteGaussianGenSampleZeroConcentrated α h1 (Δ * ε₂) ε₁ (query l₁) (query l₂) + apply le_trans A + clear A + replace bounded_sensitivity := bounded_sensitivity l₁ l₂ h2 + ring_nf + simp + conv => + left + left + right + rw [mul_pow] + conv => + left + rw [mul_assoc] + right + rw [mul_comm] + rw [← mul_assoc] + conv => + left + rw [mul_assoc] + right + rw [← mul_assoc] + left + rw [mul_comm] + rw [← mul_assoc] + rw [← mul_assoc] + rw [← mul_assoc] + simp only [inv_pow] + rw [mul_inv_le_iff'] + . have A : (α * ↑↑ε₁ ^ 2 * (↑↑ε₂ ^ 2)⁻¹) ≤ (α * ↑↑ε₁ ^ 2 * (↑↑ε₂ ^ 2)⁻¹) := le_refl (α * ↑↑ε₁ ^ 2 * (↑↑ε₂ ^ 2)⁻¹) + have B : 0 ≤ (α * ↑↑ε₁ ^ 2 * (↑↑ε₂ ^ 2)⁻¹) := by + simp + apply @le_trans ℝ _ 0 1 α (instStrictOrderedCommRingReal.proof_3) (le_of_lt h1) + apply mul_le_mul A _ _ B + . apply sq_le_sq.mpr + simp only [abs_cast] + rw [← Int.cast_sub] + rw [← Int.cast_abs] + apply Int.cast_le.mpr + rw [← Int.natCast_natAbs] + apply Int.ofNat_le.mpr + trivial + . apply sq_nonneg + . rw [pow_two] + rw [_root_.mul_pos_iff] + left + simp + +theorem NoisedQuery_NonZeroNQ (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : + NonZeroNQ (NoisedQuery query Δ ε₁ ε₂) := by + simp [NonZeroNQ, NoisedQuery, DiscreteGaussianGenSample] + intros l n + exists (n - query l) + simp + have A : ((Δ : ℝ) * ε₂ / ε₁) ≠ 0 := by + simp + have X := @discrete_gaussian_pos (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) A 0 (n - query l) + simp at X + trivial + +theorem NoisedQuery_NonTopNQ (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : + NonTopNQ (NoisedQuery query Δ ε₁ ε₂) := by + simp [NonTopNQ, NoisedQuery, DiscreteGaussianGenSample] + intro l n + rw [ENNReal.tsum_eq_add_tsum_ite (n - query l)] + simp + have X : ∀ x : ℤ, (@ite ℝ≥0∞ (x = n - query l) (propDecidable (x = n - query l)) 0 + (@ite ℝ≥0∞ (n = x + query l) (instDecidableEqInt n (x + query l)) + (ENNReal.ofReal (discrete_gaussian (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) 0 ↑x)) 0)) = 0 := by + intro x + split + . simp + . split + . rename_i h1 h2 + subst h2 + simp at h1 + . simp + conv => + right + left + right + intro x + rw [X] + simp + +theorem discrete_gaussian_shift {σ : ℝ} (h : σ ≠ 0) (μ : ℝ) (τ x : ℤ) : + discrete_gaussian σ μ (x - τ) = discrete_gaussian σ (μ + τ) (x) := by + simp [discrete_gaussian] + congr 1 + . simp [gauss_term_ℝ] + congr 3 + ring_nf + . rw [SG_periodic h] + +theorem NoisedQuery_NonTopSum (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : + NonTopSum (NoisedQuery query Δ ε₁ ε₂) := by + simp [NonTopSum, NoisedQuery, DiscreteGaussianGenSample] + intro l + have X : ∀ n: ℤ, ∀ x : ℤ, (@ite ℝ≥0∞ (x = n - query l) (propDecidable (x = n - query l)) 0 + (@ite ℝ≥0∞ (n = x + query l) (instDecidableEqInt n (x + query l)) + (ENNReal.ofReal (discrete_gaussian (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) 0 ↑x)) 0)) = 0 := by + intro n x + split + . simp + . split + . rename_i h1 h2 + subst h2 + simp at h1 + . simp + conv => + right + left + right + intro n + rw [ENNReal.tsum_eq_add_tsum_ite (n - query l)] + simp + right + right + intro x + rw [X] + simp + have A : (Δ : ℝ) * ε₂ / ε₁ ≠ 0 := by + simp + conv => + right + left + right + intro n + rw [discrete_gaussian_shift A] + simp + rw [← ENNReal.ofReal_tsum_of_nonneg] + . rw [discrete_gaussian_normalizes A] + simp + . apply discrete_gaussian_nonneg A + . apply discrete_gaussian_summable' A (query l) + +theorem NoisedQuery_NonTopRDNQ (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : + NonTopRDNQ (NoisedQuery query Δ ε₁ ε₂) := by + simp [NonTopRDNQ, NoisedQuery, DiscreteGaussianGenSample] + intro α _ l₁ l₂ _ + have A : ∀ x_1 x : ℤ, (@ite ℝ≥0∞ (x_1 = x - query l₁) (propDecidable (x_1 = x - query l₁)) 0 + (@ite ℝ≥0∞ (x = x_1 + query l₁) (instDecidableEqInt x (x_1 + query l₁)) + (ENNReal.ofReal (discrete_gaussian (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) 0 ↑x_1)) 0 )) = 0 := by + intro x y + split + . simp + . split + . rename_i h1 h2 + subst h2 + simp at h1 + . simp + have B : ∀ x_1 x : ℤ, (@ite ℝ≥0∞ (x_1 = x - query l₂) (propDecidable (x_1 = x - query l₂)) 0 + (@ite ℝ≥0∞ (x = x_1 + query l₂) (instDecidableEqInt x (x_1 + query l₂)) + (ENNReal.ofReal (discrete_gaussian (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) 0 ↑x_1)) 0)) = 0 := by + intro x y + split + . simp + . split + . rename_i h1 h2 + subst h2 + simp at h1 + . simp + conv => + right + left + right + intro x + left + rw [ENNReal.tsum_eq_add_tsum_ite (x - query l₁)] + simp + left + right + right + intro y + rw [A] + simp + conv => + right + left + right + intro x + right + rw [ENNReal.tsum_eq_add_tsum_ite (x - query l₂)] + simp + left + right + right + intro y + rw [B] + simp + clear A B + have P : (Δ : ℝ) * ε₂ / ε₁ ≠ 0 := by + simp + have A : ∀ x : ℤ, ∀ l : List T, 0 < discrete_gaussian (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) 0 (↑x - ↑(query l)) := by + intro x l + have A' := @discrete_gaussian_pos _ P 0 (x - query l) + simp at A' + trivial + have B : ∀ x : ℤ, 0 ≤ discrete_gaussian (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) 0 (↑x - ↑(query l₁)) ^ α := by + intro x + have A' := @discrete_gaussian_nonneg _ P 0 (x - query l₁) + simp at A' + apply rpow_nonneg A' + conv => + right + left + right + intro x + rw [ENNReal.ofReal_rpow_of_pos (A x l₁)] + rw [ENNReal.ofReal_rpow_of_pos (A x l₂)] + rw [← ENNReal.ofReal_mul (B x)] + rw [← ENNReal.ofReal_tsum_of_nonneg] + . simp + . intro n + have X := @RenyiSumSG_nonneg _ α P (query l₁) (query l₂) n + rw [discrete_gaussian_shift P] + rw [discrete_gaussian_shift P] + simp [X] + . have X := @SummableRenyiGauss _ P (query l₁) (query l₂) + conv => + right + intro x + rw [discrete_gaussian_shift P] + rw [discrete_gaussian_shift P] + simp [X] + +end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Basic.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Basic.lean new file mode 100644 index 00000000..467a3bd5 --- /dev/null +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Basic.lean @@ -0,0 +1,8 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jean-Baptiste Tristan +-/ + +import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Postprocessing.Code +import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Postprocessing.Properties diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Code.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Code.lean new file mode 100644 index 00000000..e0b6a899 --- /dev/null +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Code.lean @@ -0,0 +1,19 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jean-Baptiste Tristan +-/ + +import SampCert.DifferentialPrivacy.ZeroConcentrated.DP + +noncomputable section + +open Classical Nat Int Real ENNReal MeasureTheory Measure + +namespace SLang + +def PostProcess (nq : List T → SLang U) (pp : U → ℤ) (l : List T) : SLang ℤ := do + let A ← nq l + return pp A + +end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Postprocessing.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Postprocessing.lean new file mode 100644 index 00000000..94205f84 --- /dev/null +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Postprocessing.lean @@ -0,0 +1,691 @@ +import SampCert.DifferentialPrivacy.ZeroConcentrated.DP + +noncomputable section + +open Classical Nat Int Real ENNReal MeasureTheory Measure + +namespace SLang + +def PostProcess (nq : List T → SLang U) (pp : U → ℤ) (l : List T) : SLang ℤ := do + let A ← nq l + return pp A + +theorem foo (f : U → ℤ) (g : U → ENNReal) (x : ℤ) : + (∑' a : U, if x = f a then g a else 0) = ∑' a : { a | x = f a }, g a := by + have A := @tsum_split_ite U (fun a : U => x = f a) g (fun _ => 0) + simp only [decide_eq_true_eq, tsum_zero, add_zero] at A + rw [A] + have B : ↑{i | decide (x = f i) = true} = ↑{a | x = f a} := by + simp + rw [B] + +variable {T : Type} +variable [m1 : MeasurableSpace T] +variable [m2 : MeasurableSingletonClass T] +variable [m3: MeasureSpace T] + +theorem Integrable_rpow (f : T → ℝ) (nn : ∀ x : T, 0 ≤ f x) (μ : Measure T) (α : ENNReal) (mem : Memℒp f α μ) (h1 : α ≠ 0) (h2 : α ≠ ⊤) : + MeasureTheory.Integrable (fun x : T => (f x) ^ α.toReal) μ := by + have X := @MeasureTheory.Memℒp.integrable_norm_rpow T ℝ m1 μ _ f α mem h1 h2 + revert X + conv => + left + left + intro x + rw [← norm_rpow_of_nonneg (nn x)] + intro X + simp [Integrable] at * + constructor + . cases X + rename_i left right + rw [@aestronglyMeasurable_iff_aemeasurable] + apply AEMeasurable.pow_const + simp [Memℒp] at mem + cases mem + rename_i left' right' + rw [aestronglyMeasurable_iff_aemeasurable] at left' + simp [left'] + . rw [← hasFiniteIntegral_norm_iff] + simp [X] + +theorem bar (f : T → ℝ) (q : PMF T) (α : ℝ) (h : 1 < α) (h2 : ∀ x : T, 0 ≤ f x) (mem : Memℒp f (ENNReal.ofReal α) (PMF.toMeasure q)) : + ((∑' x : T, (f x) * (q x).toReal)) ^ α ≤ (∑' x : T, (f x) ^ α * (q x).toReal) := by + + conv => + left + left + right + intro x + rw [mul_comm] + rw [← smul_eq_mul] + conv => + right + right + intro x + rw [mul_comm] + rw [← smul_eq_mul] + rw [← PMF.integral_eq_tsum] + rw [← PMF.integral_eq_tsum] + + have A := @convexOn_rpow α (le_of_lt h) + have B : ContinuousOn (fun (x : ℝ) => x ^ α) (Set.Ici 0) := by + apply ContinuousOn.rpow + . exact continuousOn_id' (Set.Ici 0) + . exact continuousOn_const + . intro x h' + simp at h' + have OR : x = 0 ∨ 0 < x := by exact LE.le.eq_or_gt h' + cases OR + . rename_i h'' + subst h'' + right + apply lt_trans zero_lt_one h + . rename_i h'' + left + by_contra + rename_i h3 + subst h3 + simp at h'' + have C : @IsClosed ℝ UniformSpace.toTopologicalSpace (Set.Ici 0) := by + exact isClosed_Ici + have D := @ConvexOn.map_integral_le T ℝ m1 _ _ _ (PMF.toMeasure q) (Set.Ici 0) f (fun (x : ℝ) => x ^ α) (PMF.toMeasure.isProbabilityMeasure q) A B C + simp at D + apply D + . exact MeasureTheory.ae_of_all (PMF.toMeasure q) h2 + . apply MeasureTheory.Memℒp.integrable _ mem + rw [one_le_ofReal] + apply le_of_lt h + . rw [Function.comp_def] + have X : ENNReal.ofReal α ≠ 0 := by + simp + apply lt_trans zero_lt_one h + have Y : ENNReal.ofReal α ≠ ⊤ := by + simp + have Z := @Integrable_rpow T m1 f h2 (PMF.toMeasure q) (ENNReal.ofReal α) mem X Y + rw [toReal_ofReal] at Z + . exact Z + . apply le_of_lt + apply lt_trans zero_lt_one h + . have X : ENNReal.ofReal α ≠ 0 := by + simp + apply lt_trans zero_lt_one h + have Y : ENNReal.ofReal α ≠ ⊤ := by + simp + have Z := @Integrable_rpow T m1 f h2 (PMF.toMeasure q) (ENNReal.ofReal α) mem X Y + rw [toReal_ofReal] at Z + . exact Z + . apply le_of_lt + apply lt_trans zero_lt_one h + . apply MeasureTheory.Memℒp.integrable _ mem + rw [one_le_ofReal] + apply le_of_lt h + +variable {U : Type} +variable [m2 : MeasurableSpace U] -- [m2' : MeasurableSingletonClass U] +variable [count : Countable U] +variable [disc : DiscreteMeasurableSpace U] + +def δ (nq : SLang U) (f : U → ℤ) (a : ℤ) : {n : U | a = f n} → ENNReal := fun x : {n : U | a = f n} => nq x * (∑' (x : {n | a = f n}), nq x)⁻¹ + +theorem δ_normalizes (nq : SLang U) (f : U → ℤ) (a : ℤ) (h1 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ 0) (h2 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ ⊤) : + HasSum (δ nq f a) 1 := by + rw [Summable.hasSum_iff ENNReal.summable] + unfold δ + rw [ENNReal.tsum_mul_right] + rw [ENNReal.mul_inv_cancel h1 h2] + +def δpmf (nq : SLang U) (f : U → ℤ) (a : ℤ) (h1 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ 0) (h2 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ ⊤) : PMF {n : U | a = f n} := + ⟨ δ nq f a , δ_normalizes nq f a h1 h2 ⟩ + +theorem δpmf_conv (nq : SLang U) (a : ℤ) (x : {n | a = f n}) (h1 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ 0) (h2 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ ⊤) : + nq x * (∑' (x : {n | a = f n}), nq x)⁻¹ = (δpmf nq f a h1 h2) x := by + simp [δpmf] + conv => + right + left + left + +theorem δpmf_conv' (nq : SLang U) (f : U → ℤ) (a : ℤ) (h1 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ 0) (h2 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ ⊤) : + (fun x : {n | a = f n} => nq x * (∑' (x : {n | a = f n}), nq x)⁻¹) = (δpmf nq f a h1 h2) := by + ext x + rw [δpmf_conv] + +theorem witness {f : U → ℤ} {i : ℤ} (h : ¬{b | i = f b} = ∅) : + ∃ x : U, i = f x := by + rw [← nonempty_subtype] + exact Set.nonempty_iff_ne_empty'.mpr h + +theorem norm_simplify (x : ENNReal) (h : x ≠ ⊤) : + @nnnorm ℝ SeminormedAddGroup.toNNNorm x.toReal = x := by + simp [nnnorm] + cases x + . contradiction + . rename_i v + simp + rfl + +theorem RD1 (p q : T → ENNReal) (α : ℝ) (h : 1 < α) (RD : ∑' (x : T), p x ^ α * q x ^ (1 - α) ≠ ⊤) (nz : ∀ x : T, q x ≠ 0) (nt : ∀ x : T, q x ≠ ⊤) : + ∑' (x : T), (p x / q x) ^ α * q x ≠ ⊤ := by + rw [← RenyiDivergenceExpectation p q h nz nt] + trivial + +theorem ENNReal.HasSum_fiberwise {f : T → ENNReal} {a : ENNReal} (hf : HasSum f a) (g : T → ℤ) : + HasSum (fun c : ℤ ↦ ∑' b : g ⁻¹' {c}, f b) a := by + let A := (Equiv.sigmaFiberEquiv g) + have B := @Equiv.hasSum_iff ENNReal T ((y : ℤ) × { x // g x = y }) _ _ f a A + replace B := B.2 hf + have C := @HasSum.sigma ENNReal ℤ _ _ _ _ (fun y : ℤ => { x // g x = y }) (f ∘ ⇑(Equiv.sigmaFiberEquiv g)) (fun c => ∑' (b : ↑(g ⁻¹' {c})), f ↑b) a B + apply C + intro b + have F := @Summable.hasSum_iff ENNReal _ _ _ (fun c => (f ∘ ⇑(Equiv.sigmaFiberEquiv g)) { fst := b, snd := c }) ((fun c => ∑' (b : ↑(g ⁻¹' {c})), f ↑b) b) _ + apply (F _).2 + . rfl + . apply ENNReal.summable + +theorem ENNReal.tsum_fiberwise (p : T → ENNReal) (f : T → ℤ) : + ∑' (x : ℤ), ∑' (b : (f ⁻¹' {x})), p b + = ∑' i : T, p i := by + apply HasSum.tsum_eq + apply ENNReal.HasSum_fiberwise + apply Summable.hasSum + exact ENNReal.summable + +theorem quux (p : T → ENNReal) (f : T → ℤ) : + (∑' i : T, p i) + = ∑' (x : ℤ), if {a : T | x = f a} = {} then 0 else ∑'(i : {a : T | x = f a}), p i := by + rw [← ENNReal.tsum_fiberwise p f] + have A : ∀ x, f ⁻¹' {x} = { a | x = f a } := by + intro x + simp [Set.preimage] + rw [Set.ext_iff] + simp + intro y + exact eq_comm + conv => + left + right + intro x + rw [A] + clear A + apply tsum_congr + intro b + split + . rename_i h' + rw [h'] + simp only [tsum_empty] + . simp + +theorem convergent_subset {p : T → ENNReal} (f : T → ℤ) (conv : ∑' (x : T), p x ≠ ⊤) : + ∑' (x : { y : T| x = f y }), p x ≠ ⊤ := by + rw [← foo] + have A : (∑' (y : T), if x = f y then p y else 0) ≤ ∑' (x : T), p x := by + apply tsum_le_tsum + . intro i + split + . trivial + . simp only [_root_.zero_le] + . exact ENNReal.summable + . exact ENNReal.summable + rw [← lt_top_iff_ne_top] + apply lt_of_le_of_lt A + rw [lt_top_iff_ne_top] + trivial + +theorem ENNReal.tsum_pos {f : T → ENNReal} (h1 : ∑' x : T, f x ≠ ⊤) (h2 : ∀ x : T, f x ≠ 0) (i : T) : + 0 < ∑' x : T, f x := by + apply (toNNReal_lt_toNNReal ENNReal.zero_ne_top h1).mp + simp only [zero_toNNReal] + rw [ENNReal.tsum_toNNReal_eq (ENNReal.ne_top_of_tsum_ne_top h1)] + have S : Summable fun a => (f a).toNNReal := by + rw [← tsum_coe_ne_top_iff_summable] + conv => + left + right + intro b + rw [ENNReal.coe_toNNReal (ENNReal.ne_top_of_tsum_ne_top h1 b)] + trivial + have B:= @NNReal.tsum_pos T (fun (a : T) => (f a).toNNReal) S i + apply B + apply ENNReal.toNNReal_pos (h2 i) (ENNReal.ne_top_of_tsum_ne_top h1 i) + +theorem ENNReal.tsum_pos_int {f : ℤ → ENNReal} (h1 : ∑' x : ℤ, f x ≠ ⊤) (h2 : ∀ x : ℤ, f x ≠ 0) : + 0 < ∑' x : ℤ, f x := by + apply ENNReal.tsum_pos h1 h2 42 + +theorem tsum_pos_int {f : ℤ → ENNReal} (h1 : ∑' x : ℤ, f x ≠ ⊤) (h2 : ∀ x : ℤ, f x ≠ 0) : + 0 < (∑' x : ℤ, f x).toReal := by + have X : 0 = (0 : ENNReal).toReal := rfl + rw [X] + clear X + apply toReal_strict_mono h1 + apply ENNReal.tsum_pos_int h1 h2 + +theorem DPostPocess_pre {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP nq ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (conv : NonTopSum nq) (f : U → ℤ) {α : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (h2 : Neighbour l₁ l₂) : + (∑' (x : ℤ), + (∑' (a : U), if x = f a then nq l₁ a else 0) ^ α * + (∑' (a : U), if x = f a then nq l₂ a else 0) ^ (1 - α)) ≤ + (∑' (x : U), nq l₁ x ^ α * nq l₂ x ^ (1 - α)) := by + + simp [DP, RenyiDivergence] at h + + -- Rewrite as cascading expectations + rw [@RenyiDivergenceExpectation _ (nq l₁) (nq l₂) _ h1 (nn l₂) (nts l₂)] + + -- Shuffle the sum + rw [quux (fun x => (nq l₁ x / nq l₂ x) ^ α * nq l₂ x) f] + + apply ENNReal.tsum_le_tsum + + intro i + + -- Get rid of elements with probability 0 in the pushforward + split + . rename_i empty + rw [foo] + have ZE : (∑' (x_1 : ↑{n | i = f n}), nq l₁ ↑x_1) = 0 := by + simp + intro a H + have I₁ : a ∈ {b | i = f b} := by + simp [H] + have I2 : {b | i = f b} ≠ ∅ := by + apply ne_of_mem_of_not_mem' I₁ + simp + contradiction + rw [ZE] + simp only [toReal_mul, zero_toReal, ge_iff_le] + + rw [ENNReal.zero_rpow_of_pos] + . simp + . apply lt_trans zero_lt_one h1 + + -- Part 2: apply Jensen's inequality + . rename_i NotEmpty + + have MasterRW : ∀ l : List T, ∑' (a : ↑{a | i = f a}), nq l ↑a ≠ ⊤ := by + intro l + apply convergent_subset + simp [NonTopSum] at conv + have conv := conv l + apply conv + + have MasterZero : ∀ l : List T, ∑' (a : ↑{a | i = f a}), nq l ↑a ≠ 0 := by + intro l + simp + have T := witness NotEmpty + cases T + rename_i z w + exists z + constructor + . trivial + . apply nn l + + have S2 : (∑' (a : ↑{n | i = f n}), nq l₁ ↑a / nq l₂ ↑a * (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) a) ^ α ≠ ⊤ := by + conv => + left + left + right + intro a + rw [← δpmf_conv] + rw [division_def] + rw [mul_assoc] + right + rw [← mul_assoc] + rw [ENNReal.inv_mul_cancel (nn l₂ a) (nts l₂ a)] + rw [one_mul] + rw [ENNReal.tsum_mul_right] + apply ENNReal.rpow_ne_top_of_nonneg (le_of_lt (lt_trans zero_lt_one h1 )) + apply mul_ne_top + . apply convergent_subset _ (conv l₁) + . apply inv_ne_top.mpr (MasterZero l₂) + + have S1 : ∀ (a : ↑{n | i = f n}), nq l₁ ↑a / nq l₂ ↑a * (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) a ≠ ⊤ := by + intro a + apply mul_ne_top + . rw [division_def] + apply mul_ne_top (nts l₁ a) + apply inv_ne_top.mpr (nn l₂ a) + . rw [← δpmf_conv] + apply mul_ne_top (nts l₂ a) + apply inv_ne_top.mpr (MasterZero l₂) + + have S3 : ∑' (a : ↑{n | i = f n}), (nq l₁ ↑a / nq l₂ ↑a) ^ α * (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) a ≠ ⊤ := by + conv => + left + right + intro a + rw [← δpmf_conv] + rw [← mul_assoc] + rw [ENNReal.tsum_mul_right] + apply mul_ne_top + . rw [← RenyiDivergenceExpectation _ _ h1] + . replace nt := nt α h1 l₁ l₂ h2 + apply convergent_subset _ nt + . intro x + apply nn + . intro x + apply nts + . apply inv_ne_top.mpr (MasterZero l₂) + + have S4 : ∀ (a : ↑{n | i = f n}), (nq l₁ ↑a / nq l₂ ↑a) ^ α * (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) a ≠ ⊤ := by + intro a + apply ENNReal.ne_top_of_tsum_ne_top S3 + + rw [foo] + rw [foo] + + -- Introduce Q(f⁻¹ i) + let κ := ∑' x : {n : U | i = f n}, nq l₂ x + have P4 : κ / κ = 1 := by + rw [division_def] + rw [ENNReal.mul_inv_cancel] + . simp [κ] -- Use here for δ normalization + have T := witness NotEmpty + cases T + rename_i z w + exists z + constructor + . trivial + . apply nn l₂ + . simp only [κ] + apply MasterRW l₂ + + conv => + right + right + intro a + rw [← mul_one ((nq l₁ ↑a / nq l₂ ↑a) ^ α * nq l₂ ↑a)] + right + rw [← P4] + clear P4 + simp only [κ] + + conv => + right + right + intro a + right + rw [division_def] + rw [mul_comm] + + conv => + right + right + intro a + rw [← mul_assoc] + + rw [ENNReal.tsum_mul_right] + + -- Jensen's inequality + + have P5 : ∀ (x : ↑{n | i = f n}), 0 ≤ (fun a => (nq l₁ ↑a / nq l₂ ↑a).toReal) x := by + intro x + simp only [toReal_nonneg] + + have XXX : @Memℒp ℝ Real.normedAddCommGroup (↑{n | i = f n}) Subtype.instMeasurableSpace (fun a => (nq l₁ ↑a / nq l₂ ↑a).toReal) + (ENNReal.ofReal α) (PMF.toMeasure (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂))) := by + simp [Memℒp] + constructor + . apply MeasureTheory.StronglyMeasurable.aestronglyMeasurable + apply Measurable.stronglyMeasurable + apply Measurable.ennreal_toReal + conv => + right + intro x + rw [division_def] + apply Measurable.mul + . -- MeasurableSingletonClass.toDiscreteMeasurableSpace + apply measurable_discrete + . apply Measurable.inv + apply measurable_discrete + . simp [snorm] + split + . simp + . simp [snorm'] + rw [MeasureTheory.lintegral_countable'] -- Uses countable + rw [toReal_ofReal (le_of_lt (lt_trans zero_lt_one h1))] + have OTHER : ∀ a, nq l₁ a / nq l₂ a ≠ ⊤ := by + intro a + rw [division_def] + rw [ne_iff_lt_or_gt] + left + rw [mul_lt_top_iff] + left + constructor + . exact Ne.lt_top' (id (Ne.symm (nts l₁ a))) + . simp + exact pos_iff_ne_zero.mpr (nn l₂ a) + + conv => + left + left + right + intro a + rw [norm_simplify _ (OTHER a)] + have Z : 0 < α⁻¹ := by + simp + apply lt_trans zero_lt_one h1 + rw [rpow_lt_top_iff_of_pos Z] + conv => + left + right + intro a + rw [PMF.toMeasure_apply_singleton _ _ (measurableSet_singleton a)] + + apply Ne.lt_top' (id (Ne.symm _)) + apply S3 + + + have Jensen's := @bar {n : U | i = f n} Subtype.instMeasurableSpace Subtype.instMeasurableSingletonClass (fun a => (nq l₁ a / nq l₂ a).toReal) (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) α h1 P5 XXX + clear P5 + + have P6 : 0 ≤ (∑' (x : ↑{n | i = f n}), nq l₂ ↑x).toReal := by + simp only [toReal_nonneg] + have A' := mul_le_mul_of_nonneg_left Jensen's P6 + clear Jensen's P6 + + conv => + right + rw [mul_comm] + right + right + intro a + rw [mul_assoc] + rw [δpmf_conv] + + -- Here + + replace A' := ofReal_le_ofReal A' + rw [ofReal_mul toReal_nonneg] at A' + rw [ofReal_mul toReal_nonneg] at A' + rw [ofReal_toReal_eq_iff.2 (MasterRW l₂)] at A' + simp only at A' + + revert A' + conv => + left + right + right + right + right + intro x + rw [toReal_rpow] + rw [← toReal_mul] + conv => + left + right + right + right + rw [← ENNReal.tsum_toReal_eq S4] + intro A' + rw [ofReal_toReal_eq_iff.2 S3] at A' + + apply le_trans _ A' + clear A' + apply le_of_eq + + -- Part 3: + + conv => + right + right + right + left + right + intro x + rw [← toReal_mul] + rw [← ENNReal.tsum_toReal_eq S1] + rw [toReal_rpow] + rw [ofReal_toReal_eq_iff.2 S2] + + conv => + right + right + left + right + intro x + rw [division_def] + rw [← δpmf_conv] + rw [mul_assoc] + right + rw [← mul_assoc] + left + rw [ENNReal.inv_mul_cancel (nn l₂ x) (nts l₂ x)] + simp only [one_mul] + + rw [ENNReal.tsum_mul_right] + have H1 : 0 ≤ ∑' (x : ↑{n | i = f n}), (nq l₁ ↑x).toReal := by + apply tsum_nonneg + simp + have H2 : 0 ≤ (∑' (a : ↑{a | i = f a}), nq l₂ ↑a)⁻¹.toReal := by + apply toReal_nonneg + have H4 : (∑' (a : ↑{a | i = f a}), nq l₂ ↑a)⁻¹ ≠ ⊤ := by + apply inv_ne_top.mpr + simp + have T := witness NotEmpty + cases T + rename_i z w + exists z + constructor + . trivial + . apply nn l₂ + rw [ENNReal.mul_rpow_of_ne_top (MasterRW l₁) H4] + + have H3 : ∑' (a : ↑{a | i = f a}), nq l₂ ↑a ≠ 0 := by + simp + have T := witness NotEmpty + cases T + rename_i z w + exists z + constructor + . trivial + . apply nn l₂ + rw [ENNReal.rpow_sub _ _ H3 (MasterRW l₂)] + rw [ENNReal.rpow_one] + rw [division_def] + rw [← mul_assoc] + rw [← mul_assoc] + congr 1 + . rw [mul_comm] + . congr 1 + rw [ENNReal.inv_rpow] + +theorem tsum_ne_zero_of_ne_zero {T : Type} [Inhabited T] (f : T → ENNReal) (h : ∀ x : T, f x ≠ 0) : + ∑' x : T, f x ≠ 0 := by + by_contra CONTRA + rw [ENNReal.tsum_eq_zero] at CONTRA + have A := h default + have B := CONTRA default + contradiction + +variable [Inhabited U] + +theorem DPPostProcess {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP nq ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (conv : NonTopSum nq) (f : U → ℤ) : + DP (PostProcess nq f) ((ε₁ : ℝ) / ε₂) := by + simp [PostProcess, DP, RenyiDivergence] + intro α h1 l₁ l₂ h2 + have h' := h + simp [DP, RenyiDivergence] at h' + replace h' := h' α h1 l₁ l₂ h2 + + -- Part 1, removing fluff + + apply le_trans _ h' + clear h' + + -- remove the α scaling + have A : 0 ≤ (α - 1)⁻¹ := by + simp + apply le_of_lt h1 + apply mul_le_mul_of_nonneg_left _ A + clear A + + have RDConvegence : ∑' (x : U), nq l₁ x ^ α * nq l₂ x ^ (1 - α) ≠ ⊤ := by + simp [NonTopRDNQ] at nt + have nt := nt α h1 l₁ l₂ h2 + trivial + + have B := DPostPocess_pre h nn nt nts conv f h1 h2 + have B' : ∑' (x : ℤ), (∑' (a : U), if x = f a then nq l₁ a else 0) ^ α * (∑' (a : U), if x = f a then nq l₂ a else 0) ^ (1 - α) ≠ ⊤ := by + by_contra CONTRA + rw [CONTRA] at B + simp at B + contradiction + + -- remove the log + apply log_le_log _ (toReal_mono RDConvegence B) + apply toReal_pos _ B' + apply (tsum_ne_zero_iff ENNReal.summable).mpr + exists (f default) + + rw [ENNReal.tsum_eq_add_tsum_ite default] + conv => + left + right + rw [ENNReal.tsum_eq_add_tsum_ite default] + simp only [reduceIte] + apply mul_ne_zero + . by_contra CONTRA + rw [ENNReal.rpow_eq_zero_iff_of_pos (lt_trans zero_lt_one h1)] at CONTRA + simp at CONTRA + cases CONTRA + rename_i left right + have Y := nn l₁ default + contradiction + . by_contra CONTRA + rw [ENNReal.rpow_eq_zero_iff] at CONTRA + cases CONTRA + . rename_i CONTRA + cases CONTRA + rename_i left right + simp at left + cases left + rename_i le1 le2 + have Y := nn l₂ default + contradiction + . rename_i CONTRA + cases CONTRA + rename_i left right + simp at left + cases left + . rename_i left + have Y := nts l₂ default + contradiction + . rename_i left + have Rem := conv l₂ + have X : (∑' (x : U), if x = default then 0 else if f default = f x then nq l₂ x else 0) ≤ ∑' (n : U), nq l₂ n := by + apply ENNReal.tsum_le_tsum + intro a + split + . simp + . split + . simp + . simp + replace Rem := Ne.symm Rem + have Y := Ne.lt_top' Rem + have Z : (∑' (x : U), if x = default then 0 else if f default = f x then nq l₂ x else 0) < ⊤ := by + apply lt_of_le_of_lt X Y + rw [lt_top_iff_ne_top] at Z + contradiction + + +end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Properties.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Properties.lean new file mode 100644 index 00000000..faef6fec --- /dev/null +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Properties.lean @@ -0,0 +1,697 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jean-Baptiste Tristan +-/ + +import SampCert.DifferentialPrivacy.ZeroConcentrated.DP + +noncomputable section + +open Classical Nat Int Real ENNReal MeasureTheory Measure + +namespace SLang + +def PostProcess (nq : List T → SLang U) (pp : U → ℤ) (l : List T) : SLang ℤ := do + let A ← nq l + return pp A + +theorem foo (f : U → ℤ) (g : U → ENNReal) (x : ℤ) : + (∑' a : U, if x = f a then g a else 0) = ∑' a : { a | x = f a }, g a := by + have A := @tsum_split_ite U (fun a : U => x = f a) g (fun _ => 0) + simp only [decide_eq_true_eq, tsum_zero, add_zero] at A + rw [A] + have B : ↑{i | decide (x = f i) = true} = ↑{a | x = f a} := by + simp + rw [B] + +variable {T : Type} +variable [m1 : MeasurableSpace T] +variable [m2 : MeasurableSingletonClass T] +variable [m3: MeasureSpace T] + +theorem Integrable_rpow (f : T → ℝ) (nn : ∀ x : T, 0 ≤ f x) (μ : Measure T) (α : ENNReal) (mem : Memℒp f α μ) (h1 : α ≠ 0) (h2 : α ≠ ⊤) : + MeasureTheory.Integrable (fun x : T => (f x) ^ α.toReal) μ := by + have X := @MeasureTheory.Memℒp.integrable_norm_rpow T ℝ m1 μ _ f α mem h1 h2 + revert X + conv => + left + left + intro x + rw [← norm_rpow_of_nonneg (nn x)] + intro X + simp [Integrable] at * + constructor + . cases X + rename_i left right + rw [@aestronglyMeasurable_iff_aemeasurable] + apply AEMeasurable.pow_const + simp [Memℒp] at mem + cases mem + rename_i left' right' + rw [aestronglyMeasurable_iff_aemeasurable] at left' + simp [left'] + . rw [← hasFiniteIntegral_norm_iff] + simp [X] + +theorem bar (f : T → ℝ) (q : PMF T) (α : ℝ) (h : 1 < α) (h2 : ∀ x : T, 0 ≤ f x) (mem : Memℒp f (ENNReal.ofReal α) (PMF.toMeasure q)) : + ((∑' x : T, (f x) * (q x).toReal)) ^ α ≤ (∑' x : T, (f x) ^ α * (q x).toReal) := by + + conv => + left + left + right + intro x + rw [mul_comm] + rw [← smul_eq_mul] + conv => + right + right + intro x + rw [mul_comm] + rw [← smul_eq_mul] + rw [← PMF.integral_eq_tsum] + rw [← PMF.integral_eq_tsum] + + have A := @convexOn_rpow α (le_of_lt h) + have B : ContinuousOn (fun (x : ℝ) => x ^ α) (Set.Ici 0) := by + apply ContinuousOn.rpow + . exact continuousOn_id' (Set.Ici 0) + . exact continuousOn_const + . intro x h' + simp at h' + have OR : x = 0 ∨ 0 < x := by exact LE.le.eq_or_gt h' + cases OR + . rename_i h'' + subst h'' + right + apply lt_trans zero_lt_one h + . rename_i h'' + left + by_contra + rename_i h3 + subst h3 + simp at h'' + have C : @IsClosed ℝ UniformSpace.toTopologicalSpace (Set.Ici 0) := by + exact isClosed_Ici + have D := @ConvexOn.map_integral_le T ℝ m1 _ _ _ (PMF.toMeasure q) (Set.Ici 0) f (fun (x : ℝ) => x ^ α) (PMF.toMeasure.isProbabilityMeasure q) A B C + simp at D + apply D + . exact MeasureTheory.ae_of_all (PMF.toMeasure q) h2 + . apply MeasureTheory.Memℒp.integrable _ mem + rw [one_le_ofReal] + apply le_of_lt h + . rw [Function.comp_def] + have X : ENNReal.ofReal α ≠ 0 := by + simp + apply lt_trans zero_lt_one h + have Y : ENNReal.ofReal α ≠ ⊤ := by + simp + have Z := @Integrable_rpow T m1 f h2 (PMF.toMeasure q) (ENNReal.ofReal α) mem X Y + rw [toReal_ofReal] at Z + . exact Z + . apply le_of_lt + apply lt_trans zero_lt_one h + . have X : ENNReal.ofReal α ≠ 0 := by + simp + apply lt_trans zero_lt_one h + have Y : ENNReal.ofReal α ≠ ⊤ := by + simp + have Z := @Integrable_rpow T m1 f h2 (PMF.toMeasure q) (ENNReal.ofReal α) mem X Y + rw [toReal_ofReal] at Z + . exact Z + . apply le_of_lt + apply lt_trans zero_lt_one h + . apply MeasureTheory.Memℒp.integrable _ mem + rw [one_le_ofReal] + apply le_of_lt h + +variable {U : Type} +variable [m2 : MeasurableSpace U] -- [m2' : MeasurableSingletonClass U] +variable [count : Countable U] +variable [disc : DiscreteMeasurableSpace U] + +def δ (nq : SLang U) (f : U → ℤ) (a : ℤ) : {n : U | a = f n} → ENNReal := fun x : {n : U | a = f n} => nq x * (∑' (x : {n | a = f n}), nq x)⁻¹ + +theorem δ_normalizes (nq : SLang U) (f : U → ℤ) (a : ℤ) (h1 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ 0) (h2 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ ⊤) : + HasSum (δ nq f a) 1 := by + rw [Summable.hasSum_iff ENNReal.summable] + unfold δ + rw [ENNReal.tsum_mul_right] + rw [ENNReal.mul_inv_cancel h1 h2] + +def δpmf (nq : SLang U) (f : U → ℤ) (a : ℤ) (h1 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ 0) (h2 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ ⊤) : PMF {n : U | a = f n} := + ⟨ δ nq f a , δ_normalizes nq f a h1 h2 ⟩ + +theorem δpmf_conv (nq : SLang U) (a : ℤ) (x : {n | a = f n}) (h1 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ 0) (h2 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ ⊤) : + nq x * (∑' (x : {n | a = f n}), nq x)⁻¹ = (δpmf nq f a h1 h2) x := by + simp [δpmf] + conv => + right + left + left + +theorem δpmf_conv' (nq : SLang U) (f : U → ℤ) (a : ℤ) (h1 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ 0) (h2 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ ⊤) : + (fun x : {n | a = f n} => nq x * (∑' (x : {n | a = f n}), nq x)⁻¹) = (δpmf nq f a h1 h2) := by + ext x + rw [δpmf_conv] + +theorem witness {f : U → ℤ} {i : ℤ} (h : ¬{b | i = f b} = ∅) : + ∃ x : U, i = f x := by + rw [← nonempty_subtype] + exact Set.nonempty_iff_ne_empty'.mpr h + +theorem norm_simplify (x : ENNReal) (h : x ≠ ⊤) : + @nnnorm ℝ SeminormedAddGroup.toNNNorm x.toReal = x := by + simp [nnnorm] + cases x + . contradiction + . rename_i v + simp + rfl + +theorem RD1 (p q : T → ENNReal) (α : ℝ) (h : 1 < α) (RD : ∑' (x : T), p x ^ α * q x ^ (1 - α) ≠ ⊤) (nz : ∀ x : T, q x ≠ 0) (nt : ∀ x : T, q x ≠ ⊤) : + ∑' (x : T), (p x / q x) ^ α * q x ≠ ⊤ := by + rw [← RenyiDivergenceExpectation p q h nz nt] + trivial + +theorem ENNReal.HasSum_fiberwise {f : T → ENNReal} {a : ENNReal} (hf : HasSum f a) (g : T → ℤ) : + HasSum (fun c : ℤ ↦ ∑' b : g ⁻¹' {c}, f b) a := by + let A := (Equiv.sigmaFiberEquiv g) + have B := @Equiv.hasSum_iff ENNReal T ((y : ℤ) × { x // g x = y }) _ _ f a A + replace B := B.2 hf + have C := @HasSum.sigma ENNReal ℤ _ _ _ _ (fun y : ℤ => { x // g x = y }) (f ∘ ⇑(Equiv.sigmaFiberEquiv g)) (fun c => ∑' (b : ↑(g ⁻¹' {c})), f ↑b) a B + apply C + intro b + have F := @Summable.hasSum_iff ENNReal _ _ _ (fun c => (f ∘ ⇑(Equiv.sigmaFiberEquiv g)) { fst := b, snd := c }) ((fun c => ∑' (b : ↑(g ⁻¹' {c})), f ↑b) b) _ + apply (F _).2 + . rfl + . apply ENNReal.summable + +theorem ENNReal.tsum_fiberwise (p : T → ENNReal) (f : T → ℤ) : + ∑' (x : ℤ), ∑' (b : (f ⁻¹' {x})), p b + = ∑' i : T, p i := by + apply HasSum.tsum_eq + apply ENNReal.HasSum_fiberwise + apply Summable.hasSum + exact ENNReal.summable + +theorem quux (p : T → ENNReal) (f : T → ℤ) : + (∑' i : T, p i) + = ∑' (x : ℤ), if {a : T | x = f a} = {} then 0 else ∑'(i : {a : T | x = f a}), p i := by + rw [← ENNReal.tsum_fiberwise p f] + have A : ∀ x, f ⁻¹' {x} = { a | x = f a } := by + intro x + simp [Set.preimage] + rw [Set.ext_iff] + simp + intro y + exact eq_comm + conv => + left + right + intro x + rw [A] + clear A + apply tsum_congr + intro b + split + . rename_i h' + rw [h'] + simp only [tsum_empty] + . simp + +theorem convergent_subset {p : T → ENNReal} (f : T → ℤ) (conv : ∑' (x : T), p x ≠ ⊤) : + ∑' (x : { y : T| x = f y }), p x ≠ ⊤ := by + rw [← foo] + have A : (∑' (y : T), if x = f y then p y else 0) ≤ ∑' (x : T), p x := by + apply tsum_le_tsum + . intro i + split + . trivial + . simp only [_root_.zero_le] + . exact ENNReal.summable + . exact ENNReal.summable + rw [← lt_top_iff_ne_top] + apply lt_of_le_of_lt A + rw [lt_top_iff_ne_top] + trivial + +theorem ENNReal.tsum_pos {f : T → ENNReal} (h1 : ∑' x : T, f x ≠ ⊤) (h2 : ∀ x : T, f x ≠ 0) (i : T) : + 0 < ∑' x : T, f x := by + apply (toNNReal_lt_toNNReal ENNReal.zero_ne_top h1).mp + simp only [zero_toNNReal] + rw [ENNReal.tsum_toNNReal_eq (ENNReal.ne_top_of_tsum_ne_top h1)] + have S : Summable fun a => (f a).toNNReal := by + rw [← tsum_coe_ne_top_iff_summable] + conv => + left + right + intro b + rw [ENNReal.coe_toNNReal (ENNReal.ne_top_of_tsum_ne_top h1 b)] + trivial + have B:= @NNReal.tsum_pos T (fun (a : T) => (f a).toNNReal) S i + apply B + apply ENNReal.toNNReal_pos (h2 i) (ENNReal.ne_top_of_tsum_ne_top h1 i) + +theorem ENNReal.tsum_pos_int {f : ℤ → ENNReal} (h1 : ∑' x : ℤ, f x ≠ ⊤) (h2 : ∀ x : ℤ, f x ≠ 0) : + 0 < ∑' x : ℤ, f x := by + apply ENNReal.tsum_pos h1 h2 42 + +theorem tsum_pos_int {f : ℤ → ENNReal} (h1 : ∑' x : ℤ, f x ≠ ⊤) (h2 : ∀ x : ℤ, f x ≠ 0) : + 0 < (∑' x : ℤ, f x).toReal := by + have X : 0 = (0 : ENNReal).toReal := rfl + rw [X] + clear X + apply toReal_strict_mono h1 + apply ENNReal.tsum_pos_int h1 h2 + +theorem DPostPocess_pre {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP nq ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (conv : NonTopSum nq) (f : U → ℤ) {α : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (h2 : Neighbour l₁ l₂) : + (∑' (x : ℤ), + (∑' (a : U), if x = f a then nq l₁ a else 0) ^ α * + (∑' (a : U), if x = f a then nq l₂ a else 0) ^ (1 - α)) ≤ + (∑' (x : U), nq l₁ x ^ α * nq l₂ x ^ (1 - α)) := by + + simp [DP, RenyiDivergence] at h + + -- Rewrite as cascading expectations + rw [@RenyiDivergenceExpectation _ (nq l₁) (nq l₂) _ h1 (nn l₂) (nts l₂)] + + -- Shuffle the sum + rw [quux (fun x => (nq l₁ x / nq l₂ x) ^ α * nq l₂ x) f] + + apply ENNReal.tsum_le_tsum + + intro i + + -- Get rid of elements with probability 0 in the pushforward + split + . rename_i empty + rw [foo] + have ZE : (∑' (x_1 : ↑{n | i = f n}), nq l₁ ↑x_1) = 0 := by + simp + intro a H + have I₁ : a ∈ {b | i = f b} := by + simp [H] + have I2 : {b | i = f b} ≠ ∅ := by + apply ne_of_mem_of_not_mem' I₁ + simp + contradiction + rw [ZE] + simp only [toReal_mul, zero_toReal, ge_iff_le] + + rw [ENNReal.zero_rpow_of_pos] + . simp + . apply lt_trans zero_lt_one h1 + + -- Part 2: apply Jensen's inequality + . rename_i NotEmpty + + have MasterRW : ∀ l : List T, ∑' (a : ↑{a | i = f a}), nq l ↑a ≠ ⊤ := by + intro l + apply convergent_subset + simp [NonTopSum] at conv + have conv := conv l + apply conv + + have MasterZero : ∀ l : List T, ∑' (a : ↑{a | i = f a}), nq l ↑a ≠ 0 := by + intro l + simp + have T := witness NotEmpty + cases T + rename_i z w + exists z + constructor + . trivial + . apply nn l + + have S2 : (∑' (a : ↑{n | i = f n}), nq l₁ ↑a / nq l₂ ↑a * (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) a) ^ α ≠ ⊤ := by + conv => + left + left + right + intro a + rw [← δpmf_conv] + rw [division_def] + rw [mul_assoc] + right + rw [← mul_assoc] + rw [ENNReal.inv_mul_cancel (nn l₂ a) (nts l₂ a)] + rw [one_mul] + rw [ENNReal.tsum_mul_right] + apply ENNReal.rpow_ne_top_of_nonneg (le_of_lt (lt_trans zero_lt_one h1 )) + apply mul_ne_top + . apply convergent_subset _ (conv l₁) + . apply inv_ne_top.mpr (MasterZero l₂) + + have S1 : ∀ (a : ↑{n | i = f n}), nq l₁ ↑a / nq l₂ ↑a * (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) a ≠ ⊤ := by + intro a + apply mul_ne_top + . rw [division_def] + apply mul_ne_top (nts l₁ a) + apply inv_ne_top.mpr (nn l₂ a) + . rw [← δpmf_conv] + apply mul_ne_top (nts l₂ a) + apply inv_ne_top.mpr (MasterZero l₂) + + have S3 : ∑' (a : ↑{n | i = f n}), (nq l₁ ↑a / nq l₂ ↑a) ^ α * (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) a ≠ ⊤ := by + conv => + left + right + intro a + rw [← δpmf_conv] + rw [← mul_assoc] + rw [ENNReal.tsum_mul_right] + apply mul_ne_top + . rw [← RenyiDivergenceExpectation _ _ h1] + . replace nt := nt α h1 l₁ l₂ h2 + apply convergent_subset _ nt + . intro x + apply nn + . intro x + apply nts + . apply inv_ne_top.mpr (MasterZero l₂) + + have S4 : ∀ (a : ↑{n | i = f n}), (nq l₁ ↑a / nq l₂ ↑a) ^ α * (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) a ≠ ⊤ := by + intro a + apply ENNReal.ne_top_of_tsum_ne_top S3 + + rw [foo] + rw [foo] + + -- Introduce Q(f⁻¹ i) + let κ := ∑' x : {n : U | i = f n}, nq l₂ x + have P4 : κ / κ = 1 := by + rw [division_def] + rw [ENNReal.mul_inv_cancel] + . simp [κ] -- Use here for δ normalization + have T := witness NotEmpty + cases T + rename_i z w + exists z + constructor + . trivial + . apply nn l₂ + . simp only [κ] + apply MasterRW l₂ + + conv => + right + right + intro a + rw [← mul_one ((nq l₁ ↑a / nq l₂ ↑a) ^ α * nq l₂ ↑a)] + right + rw [← P4] + clear P4 + simp only [κ] + + conv => + right + right + intro a + right + rw [division_def] + rw [mul_comm] + + conv => + right + right + intro a + rw [← mul_assoc] + + rw [ENNReal.tsum_mul_right] + + -- Jensen's inequality + + have P5 : ∀ (x : ↑{n | i = f n}), 0 ≤ (fun a => (nq l₁ ↑a / nq l₂ ↑a).toReal) x := by + intro x + simp only [toReal_nonneg] + + have XXX : @Memℒp ℝ Real.normedAddCommGroup (↑{n | i = f n}) Subtype.instMeasurableSpace (fun a => (nq l₁ ↑a / nq l₂ ↑a).toReal) + (ENNReal.ofReal α) (PMF.toMeasure (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂))) := by + simp [Memℒp] + constructor + . apply MeasureTheory.StronglyMeasurable.aestronglyMeasurable + apply Measurable.stronglyMeasurable + apply Measurable.ennreal_toReal + conv => + right + intro x + rw [division_def] + apply Measurable.mul + . -- MeasurableSingletonClass.toDiscreteMeasurableSpace + apply measurable_discrete + . apply Measurable.inv + apply measurable_discrete + . simp [snorm] + split + . simp + . simp [snorm'] + rw [MeasureTheory.lintegral_countable'] -- Uses countable + rw [toReal_ofReal (le_of_lt (lt_trans zero_lt_one h1))] + have OTHER : ∀ a, nq l₁ a / nq l₂ a ≠ ⊤ := by + intro a + rw [division_def] + rw [ne_iff_lt_or_gt] + left + rw [mul_lt_top_iff] + left + constructor + . exact Ne.lt_top' (id (Ne.symm (nts l₁ a))) + . simp + exact pos_iff_ne_zero.mpr (nn l₂ a) + + conv => + left + left + right + intro a + rw [norm_simplify _ (OTHER a)] + have Z : 0 < α⁻¹ := by + simp + apply lt_trans zero_lt_one h1 + rw [rpow_lt_top_iff_of_pos Z] + conv => + left + right + intro a + rw [PMF.toMeasure_apply_singleton _ _ (measurableSet_singleton a)] + + apply Ne.lt_top' (id (Ne.symm _)) + apply S3 + + + have Jensen's := @bar {n : U | i = f n} Subtype.instMeasurableSpace Subtype.instMeasurableSingletonClass (fun a => (nq l₁ a / nq l₂ a).toReal) (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) α h1 P5 XXX + clear P5 + + have P6 : 0 ≤ (∑' (x : ↑{n | i = f n}), nq l₂ ↑x).toReal := by + simp only [toReal_nonneg] + have A' := mul_le_mul_of_nonneg_left Jensen's P6 + clear Jensen's P6 + + conv => + right + rw [mul_comm] + right + right + intro a + rw [mul_assoc] + rw [δpmf_conv] + + -- Here + + replace A' := ofReal_le_ofReal A' + rw [ofReal_mul toReal_nonneg] at A' + rw [ofReal_mul toReal_nonneg] at A' + rw [ofReal_toReal_eq_iff.2 (MasterRW l₂)] at A' + simp only at A' + + revert A' + conv => + left + right + right + right + right + intro x + rw [toReal_rpow] + rw [← toReal_mul] + conv => + left + right + right + right + rw [← ENNReal.tsum_toReal_eq S4] + intro A' + rw [ofReal_toReal_eq_iff.2 S3] at A' + + apply le_trans _ A' + clear A' + apply le_of_eq + + -- Part 3: + + conv => + right + right + right + left + right + intro x + rw [← toReal_mul] + rw [← ENNReal.tsum_toReal_eq S1] + rw [toReal_rpow] + rw [ofReal_toReal_eq_iff.2 S2] + + conv => + right + right + left + right + intro x + rw [division_def] + rw [← δpmf_conv] + rw [mul_assoc] + right + rw [← mul_assoc] + left + rw [ENNReal.inv_mul_cancel (nn l₂ x) (nts l₂ x)] + simp only [one_mul] + + rw [ENNReal.tsum_mul_right] + have H1 : 0 ≤ ∑' (x : ↑{n | i = f n}), (nq l₁ ↑x).toReal := by + apply tsum_nonneg + simp + have H2 : 0 ≤ (∑' (a : ↑{a | i = f a}), nq l₂ ↑a)⁻¹.toReal := by + apply toReal_nonneg + have H4 : (∑' (a : ↑{a | i = f a}), nq l₂ ↑a)⁻¹ ≠ ⊤ := by + apply inv_ne_top.mpr + simp + have T := witness NotEmpty + cases T + rename_i z w + exists z + constructor + . trivial + . apply nn l₂ + rw [ENNReal.mul_rpow_of_ne_top (MasterRW l₁) H4] + + have H3 : ∑' (a : ↑{a | i = f a}), nq l₂ ↑a ≠ 0 := by + simp + have T := witness NotEmpty + cases T + rename_i z w + exists z + constructor + . trivial + . apply nn l₂ + rw [ENNReal.rpow_sub _ _ H3 (MasterRW l₂)] + rw [ENNReal.rpow_one] + rw [division_def] + rw [← mul_assoc] + rw [← mul_assoc] + congr 1 + . rw [mul_comm] + . congr 1 + rw [ENNReal.inv_rpow] + +theorem tsum_ne_zero_of_ne_zero {T : Type} [Inhabited T] (f : T → ENNReal) (h : ∀ x : T, f x ≠ 0) : + ∑' x : T, f x ≠ 0 := by + by_contra CONTRA + rw [ENNReal.tsum_eq_zero] at CONTRA + have A := h default + have B := CONTRA default + contradiction + +variable [Inhabited U] + +theorem DPPostProcess {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP nq ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (conv : NonTopSum nq) (f : U → ℤ) : + DP (PostProcess nq f) ((ε₁ : ℝ) / ε₂) := by + simp [PostProcess, DP, RenyiDivergence] + intro α h1 l₁ l₂ h2 + have h' := h + simp [DP, RenyiDivergence] at h' + replace h' := h' α h1 l₁ l₂ h2 + + -- Part 1, removing fluff + + apply le_trans _ h' + clear h' + + -- remove the α scaling + have A : 0 ≤ (α - 1)⁻¹ := by + simp + apply le_of_lt h1 + apply mul_le_mul_of_nonneg_left _ A + clear A + + have RDConvegence : ∑' (x : U), nq l₁ x ^ α * nq l₂ x ^ (1 - α) ≠ ⊤ := by + simp [NonTopRDNQ] at nt + have nt := nt α h1 l₁ l₂ h2 + trivial + + have B := DPostPocess_pre h nn nt nts conv f h1 h2 + have B' : ∑' (x : ℤ), (∑' (a : U), if x = f a then nq l₁ a else 0) ^ α * (∑' (a : U), if x = f a then nq l₂ a else 0) ^ (1 - α) ≠ ⊤ := by + by_contra CONTRA + rw [CONTRA] at B + simp at B + contradiction + + -- remove the log + apply log_le_log _ (toReal_mono RDConvegence B) + apply toReal_pos _ B' + apply (tsum_ne_zero_iff ENNReal.summable).mpr + exists (f default) + + rw [ENNReal.tsum_eq_add_tsum_ite default] + conv => + left + right + rw [ENNReal.tsum_eq_add_tsum_ite default] + simp only [reduceIte] + apply mul_ne_zero + . by_contra CONTRA + rw [ENNReal.rpow_eq_zero_iff_of_pos (lt_trans zero_lt_one h1)] at CONTRA + simp at CONTRA + cases CONTRA + rename_i left right + have Y := nn l₁ default + contradiction + . by_contra CONTRA + rw [ENNReal.rpow_eq_zero_iff] at CONTRA + cases CONTRA + . rename_i CONTRA + cases CONTRA + rename_i left right + simp at left + cases left + rename_i le1 le2 + have Y := nn l₂ default + contradiction + . rename_i CONTRA + cases CONTRA + rename_i left right + simp at left + cases left + . rename_i left + have Y := nts l₂ default + contradiction + . rename_i left + have Rem := conv l₂ + have X : (∑' (x : U), if x = default then 0 else if f default = f x then nq l₂ x else 0) ≤ ∑' (n : U), nq l₂ n := by + apply ENNReal.tsum_le_tsum + intro a + split + . simp + . split + . simp + . simp + replace Rem := Ne.symm Rem + have Y := Ne.lt_top' Rem + have Z : (∑' (x : U), if x = default then 0 else if f default = f x then nq l₂ x else 0) < ⊤ := by + apply lt_of_le_of_lt X Y + rw [lt_top_iff_ne_top] at Z + contradiction + + +end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/Basic.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/Basic.lean new file mode 100644 index 00000000..6ad79a34 --- /dev/null +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/Basic.lean @@ -0,0 +1,8 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jean-Baptiste Tristan +-/ + +import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.BoundedMean.Code +import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.BoundedMean.Properties diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/Code.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/Code.lean new file mode 100644 index 00000000..70a25010 --- /dev/null +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/Code.lean @@ -0,0 +1,23 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jean-Baptiste Tristan +-/ + +import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Composition.Composition +import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Postprocessing.Postprocessing +import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.Count.Count +import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.BoundedSum.BoundedSum + +open Classical Nat Int Real + +noncomputable section + +namespace SLang + +def NoisedBoundedAvgQuery (U : ℕ+) (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℤ := do + let S ← NoisedBoundedSumQuery U ε₁ (2 * ε₂) l + let C ← NoisedCountingQuery ε₁ (2 * ε₂) l + return S / C + +end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/Properties.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/Properties.lean new file mode 100644 index 00000000..3ac6832e --- /dev/null +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/Properties.lean @@ -0,0 +1,56 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jean-Baptiste Tristan +-/ + +import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Composition.Composition +import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Postprocessing.Postprocessing +import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.Count.Count +import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.BoundedSum.BoundedSum +import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.BoundedMean.Code + +open Classical Nat Int Real + +noncomputable section + +namespace SLang + +def NoisedBoundedAvgQuery' (U : ℕ+) (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℤ := + let X := Compose (NoisedBoundedSumQuery U ε₁ (2 * ε₂)) (NoisedCountingQuery ε₁ (2 * ε₂)) + PostProcess X (fun z => z.1 / z.2) l + +theorem NoisedBoundedAvgQueryIdentical (U : ℕ+) (ε₁ ε₂ : ℕ+) : + NoisedBoundedAvgQuery' U ε₁ ε₂ = NoisedBoundedAvgQuery U ε₁ ε₂ := by + ext l x + simp [NoisedBoundedAvgQuery, NoisedBoundedAvgQuery', PostProcess, Compose] + +theorem BoundedSumQueryDP (U : ℕ+) (ε₁ ε₂ : ℕ+) : DP (NoisedBoundedAvgQuery U ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by + rw [← NoisedBoundedAvgQueryIdentical] + unfold NoisedBoundedAvgQuery' + simp only + + have A₁ := @NoisedCountingQueryDP ℕ ε₁ (2 * ε₂) + have A₂ := @NoisedCountingQuery_NonZeroNQ ℕ ε₁ (2 * ε₂) + have A₃ := @NoisedCountingQuery_NonTopNQ ℕ ε₁ (2 * ε₂) + have A₄ := @NoisedCountingQuery_NonTopRDNQ ℕ ε₁ (2 * ε₂) + have A₅ := @NoisedCountingQuery_NonTopSum ℕ ε₁ (2 * ε₂) + + have B₁ := NoisedBoundedSumQueryDP U ε₁ (2 * ε₂) + have B₂ := NoisedBoundedSumQuery_NonZeroNQ U ε₁ (2 * ε₂) + have B₃ := NoisedBoundedSumQuery_NonTopNQ U ε₁ (2 * ε₂) + have B₄ := NoisedBoundedSumQuery_NonTopRDNQ U ε₁ (2 * ε₂) + have B₅ := NoisedBoundedSumQuery_NonTopSum U ε₁ (2 * ε₂) + + have C₁ := DPCompose B₁ A₁ B₂ A₂ B₄ A₄ B₃ A₃ + have C₂ := DPCompose_NonZeroNQ B₂ A₂ + have C₃ := DPCompose_NonTopNQ B₃ A₃ + have C₄ := DPCompose_NonTopRDNQ B₄ A₄ B₂ A₂ + have C₅ := DPCompose_NonTopSum B₅ A₅ + simp at * + ring_nf at * + rw [← division_def] at * + have D := DPPostProcess C₁ C₂ C₄ C₃ C₅ (fun z => z.1 /z.2) + exact D + +end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/Basic.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/Basic.lean new file mode 100644 index 00000000..77d03c1b --- /dev/null +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/Basic.lean @@ -0,0 +1,8 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jean-Baptiste Tristan +-/ + +import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.BoundedSum.Code +import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.BoundedSum.Properties diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/Code.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/Code.lean new file mode 100644 index 00000000..6668165c --- /dev/null +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/Code.lean @@ -0,0 +1,21 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jean-Baptiste Tristan +-/ + +import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Mechanism.Mechanism +import Mathlib.Algebra.Group.Defs +import Mathlib.Init.Algebra.Classes +import Init.Data.Int.Order + +open Classical Nat Int Real + +noncomputable section + +namespace SLang + +def BoundedSumQuery (U : ℕ+) (l : List ℕ) : ℤ := + List.sum (List.map (fun n : ℕ => (Nat.min n U)) l) + +end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/Properties.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/Properties.lean new file mode 100644 index 00000000..cdf85f74 --- /dev/null +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/Properties.lean @@ -0,0 +1,96 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jean-Baptiste Tristan +-/ + +import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Mechanism.Mechanism +import Mathlib.Algebra.Group.Defs +import Mathlib.Init.Algebra.Classes +import Init.Data.Int.Order +import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.BoundedSum.Code + +open Classical Nat Int Real + +noncomputable section + +namespace SLang + +theorem BoundedSumQuerySensitivity (U : ℕ+) : sensitivity (BoundedSumQuery U) U := by + simp [sensitivity, BoundedSumQuery] + intros l₁ l₂ H + have A : ∀ n : ℕ, (@min ℤ instMinInt (n : ℤ) (U : ℤ) = n) ∨ (@min ℤ instMinInt n U = U) := by + intro n + simp + exact Nat.le_total n ↑U + cases H + . rename_i l l' n h1 h2 + subst h1 h2 + simp + cases A n + . rename_i h + rw [h] + simp at * + trivial + . rename_i h + rw [h] + simp + . rename_i l n l' h1 h2 + subst h1 h2 + simp + cases A n + . rename_i h + rw [h] + simp at * + trivial + . rename_i h + rw [h] + simp + . rename_i l n l' m h1 h2 + subst h1 h2 + simp + cases A n + . rename_i h + cases A m + . rename_i h' + rw [h, h'] + simp at * + apply Int.natAbs_coe_sub_coe_le_of_le h h' + . rename_i h' + rw [h, h'] + simp at * + apply Int.natAbs_coe_sub_coe_le_of_le h le_rfl + . rename_i h + cases A m + . rename_i h' + rw [h, h'] + simp at * + apply Int.natAbs_coe_sub_coe_le_of_le le_rfl h' + . rename_i h' + rw [h, h'] + simp at * + +def NoisedBoundedSumQuery (U : ℕ+) (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℤ := do + NoisedQuery (BoundedSumQuery U) U ε₁ ε₂ l + +theorem NoisedBoundedSumQueryDP (U : ℕ+) (ε₁ ε₂ : ℕ+) : DP (NoisedBoundedSumQuery U ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by + apply NoisedQueryDP + apply BoundedSumQuerySensitivity + +theorem NoisedBoundedSumQuery_NonZeroNQ (U : ℕ+) (ε₁ ε₂ : ℕ+) : + @NonZeroNQ ℕ ℤ (NoisedBoundedSumQuery U ε₁ ε₂) := by + apply NoisedQuery_NonZeroNQ + +theorem NoisedBoundedSumQuery_NonTopNQ (U : ℕ+) (ε₁ ε₂ : ℕ+) : + @NonTopNQ ℕ ℤ (NoisedBoundedSumQuery U ε₁ ε₂) := by + apply NoisedQuery_NonTopNQ + +theorem NoisedBoundedSumQuery_NonTopRDNQ (U : ℕ+) (ε₁ ε₂ : ℕ+) : + @NonTopRDNQ ℕ ℤ (NoisedBoundedSumQuery U ε₁ ε₂) := by + apply NoisedQuery_NonTopRDNQ + +theorem NoisedBoundedSumQuery_NonTopSum (U : ℕ+) (ε₁ ε₂ : ℕ+) : + @NonTopSum ℕ ℤ (NoisedBoundedSumQuery U ε₁ ε₂) := by + apply NoisedQuery_NonTopSum + +end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Basic.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Basic.lean new file mode 100644 index 00000000..42def315 --- /dev/null +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Basic.lean @@ -0,0 +1,8 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jean-Baptiste Tristan +-/ + +import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.Count.Code +import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.Count.Properties diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Code.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Code.lean new file mode 100644 index 00000000..1b3c83fa --- /dev/null +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Code.lean @@ -0,0 +1,19 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jean-Baptiste Tristan +-/ + +import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Mechanism.Mechanism + +open Classical Nat Int Real + +noncomputable section + +namespace SLang + +variable {T : Type} + +def CountingQuery (l : List T) : ℤ := List.length l + +end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Properties.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Properties.lean new file mode 100644 index 00000000..80f7875e --- /dev/null +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Properties.lean @@ -0,0 +1,57 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jean-Baptiste Tristan +-/ + +import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Mechanism.Mechanism +import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.Count.Code + +open Classical Nat Int Real + +noncomputable section + +namespace SLang + +variable {T : Type} + +theorem CountingQuery1Sensitive : + @sensitivity T CountingQuery 1 := by + simp [CountingQuery, sensitivity] + intros l₁ l₂ H + cases H + . rename_i a b n h1 h2 + subst h1 h2 + simp + . rename_i a b n h1 h2 + subst h1 h2 + simp + . rename_i a n b m h1 h2 + subst h1 h2 + simp + +def NoisedCountingQuery (ε₁ ε₂ : ℕ+) (l : List T) : SLang ℤ := do + NoisedQuery CountingQuery 1 ε₁ ε₂ l + +theorem NoisedCountingQueryDP (ε₁ ε₂ : ℕ+) : + @DP T ℤ (NoisedCountingQuery ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by + apply NoisedQueryDP + apply CountingQuery1Sensitive + +theorem NoisedCountingQuery_NonZeroNQ (ε₁ ε₂ : ℕ+) : + @NonZeroNQ T ℤ (NoisedCountingQuery ε₁ ε₂) := by + apply NoisedQuery_NonZeroNQ + +theorem NoisedCountingQuery_NonTopNQ (ε₁ ε₂ : ℕ+) : + @NonTopNQ T ℤ (NoisedCountingQuery ε₁ ε₂) := by + apply NoisedQuery_NonTopNQ + +theorem NoisedCountingQuery_NonTopRDNQ (ε₁ ε₂ : ℕ+) : + @NonTopRDNQ T ℤ (NoisedCountingQuery ε₁ ε₂) := by + apply NoisedQuery_NonTopRDNQ + +theorem NoisedCountingQuery_NonTopSum (ε₁ ε₂ : ℕ+) : + @NonTopSum T ℤ (NoisedCountingQuery ε₁ ε₂) := by + apply NoisedQuery_NonTopSum + +end SLang From 1b253b3cf36ba4988959126ac22a64a4732544c8 Mon Sep 17 00:00:00 2001 From: jtristan Date: Wed, 15 May 2024 15:27:12 -0400 Subject: [PATCH 18/20] Split between code and props --- SampCert.lean | 2 +- .../Foundations/Composition/Composition.lean | 317 -------- .../Foundations/Mechanism/Mechanism.lean | 242 ------ .../Foundations/Mechanism/Properties.lean | 4 +- .../Postprocessing/Postprocessing.lean | 691 ------------------ .../Postprocessing/Properties.lean | 5 +- .../Queries/BoundedMean/BoundedMean.lean | 60 -- .../Queries/BoundedMean/Code.lean | 8 +- .../Queries/BoundedMean/Properties.lean | 8 +- .../Queries/BoundedSum/BoundedSum.lean | 98 --- .../Queries/BoundedSum/Code.lean | 5 +- .../Queries/BoundedSum/Properties.lean | 5 +- .../ZeroConcentrated/Queries/Count/Code.lean | 5 +- .../ZeroConcentrated/Queries/Count/Count.lean | 58 -- .../Queries/Count/Properties.lean | 5 +- 15 files changed, 21 insertions(+), 1492 deletions(-) delete mode 100644 SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Composition/Composition.lean delete mode 100644 SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Mechanism/Mechanism.lean delete mode 100644 SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Postprocessing.lean delete mode 100644 SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/BoundedMean.lean delete mode 100644 SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/BoundedSum.lean delete mode 100644 SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Count.lean diff --git a/SampCert.lean b/SampCert.lean index 5035d020..f9cf72c1 100644 --- a/SampCert.lean +++ b/SampCert.lean @@ -4,4 +4,4 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jean-Baptiste Tristan -/ -import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.BoundedMean.BoundedMean +import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.BoundedMean.Code diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Composition/Composition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Composition/Composition.lean deleted file mode 100644 index 28e39a5c..00000000 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Composition/Composition.lean +++ /dev/null @@ -1,317 +0,0 @@ -import SampCert.DifferentialPrivacy.ZeroConcentrated.DP - -noncomputable section - -open Classical Nat Int Real ENNReal MeasureTheory Measure - -namespace SLang - -def Compose (nq1 nq2 : List T → SLang ℤ) (l : List T) : SLang (ℤ × ℤ) := do - let A ← nq1 l - let B ← nq2 l - return (A,B) - -theorem ENNReal_toTeal_NZ (x : ENNReal) (h1 : x ≠ 0) (h2 : x ≠ ⊤) : - x.toReal ≠ 0 := by - unfold ENNReal.toReal - unfold ENNReal.toNNReal - simp - intro H - cases H - . contradiction - . contradiction - -theorem simp_α_1 {α : ℝ} (h : 1 < α) : 0 < α := by - apply @lt_trans _ _ _ 1 _ _ h - simp only [zero_lt_one] - -theorem RenyiNoisedQueryNonZero {nq : List T → SLang ℤ} {α ε : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (h2 : Neighbour l₁ l₂) (h3 : DP nq ε) (h4 : NonZeroNQ nq) (h5 : NonTopRDNQ nq) (nts : NonTopNQ nq) : - (∑' (i : ℤ), nq l₁ i ^ α * nq l₂ i ^ (1 - α)).toReal ≠ 0 := by - simp [DP] at h3 - replace h3 := h3 α h1 l₁ l₂ h2 - simp [RenyiDivergence] at h3 - simp [NonZeroNQ] at h4 - simp [NonTopRDNQ] at h5 - replace h5 := h5 α h1 l₁ l₂ h2 - have h6 := h4 l₁ - have h7 := h4 l₂ - apply ENNReal_toTeal_NZ - . by_contra CONTRA - rw [ENNReal.tsum_eq_zero] at CONTRA - replace CONTRA := CONTRA 42 - replace h6 := h6 42 - replace h7 := h7 42 - rw [_root_.mul_eq_zero] at CONTRA - cases CONTRA - . rename_i h8 - rw [rpow_eq_zero_iff_of_pos] at h8 - contradiction - apply simp_α_1 h1 - . rename_i h8 - rw [ENNReal.rpow_eq_zero_iff] at h8 - cases h8 - . rename_i h8 - cases h8 - contradiction - . rename_i h8 - cases h8 - rename_i h8 h9 - replace nts := nts l₂ 42 - contradiction - . exact h5 - -theorem compose_sum_rw (nq1 nq2 : List T → SLang ℤ) (b c : ℤ) (l : List T) : - (∑' (a : ℤ), nq1 l a * ∑' (a_1 : ℤ), if b = a ∧ c = a_1 then nq2 l a_1 else 0) = nq1 l b * nq2 l c := by - have A : ∀ a b : ℤ, (∑' (a_1 : ℤ), if b = a ∧ c = a_1 then nq2 l a_1 else 0) = if b = a then (∑' (a_1 : ℤ), if c = a_1 then nq2 l a_1 else 0) else 0 := by - intro x y - split - . rename_i h - subst h - simp - . rename_i h - simp - intro h - contradiction - conv => - left - right - intro a - right - rw [A] - rw [ENNReal.tsum_eq_add_tsum_ite b] - simp - have B : ∀ x : ℤ, (@ite ℝ≥0∞ (x = b) (instDecidableEqInt x b) 0 - (@ite ℝ≥0∞ (b = x) (instDecidableEqInt b x) (nq1 l x * ∑' (a_1 : ℤ), @ite ℝ≥0∞ (c = a_1) (instDecidableEqInt c a_1) (nq2 l a_1) 0) 0)) = 0 := by - intro x - split - . simp - . split - . rename_i h1 h2 - subst h2 - contradiction - . simp - conv => - left - right - right - intro x - rw [B] - simp - congr 1 - rw [ENNReal.tsum_eq_add_tsum_ite c] - simp - have C :∀ x : ℤ, (@ite ℝ≥0∞ (x = c) (propDecidable (x = c)) 0 (@ite ℝ≥0∞ (c = x) (instDecidableEqInt c x) (nq2 l x) 0)) = 0 := by - intro x - split - . simp - . split - . rename_i h1 h2 - subst h2 - contradiction - . simp - conv => - left - right - right - intro X - rw [C] - simp - -theorem DPCompose {nq1 nq2 : List T → SLang ℤ} {ε₁ ε₂ ε₃ ε₄ : ℕ+} (h1 : DP nq1 ((ε₁ : ℝ) / ε₂)) (h2 : DP nq2 ((ε₃ : ℝ) / ε₄)) (nn1 : NonZeroNQ nq1) (nn2 : NonZeroNQ nq2) (nt1 : NonTopRDNQ nq1) (nt2 : NonTopRDNQ nq2) (nts1 : NonTopNQ nq1) (nts2 : NonTopNQ nq2) : - DP (Compose nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by - simp [Compose, RenyiDivergence, DP] - intro α h3 l₁ l₂ h4 - have X := h1 - have Y := h2 - simp [DP] at h1 h2 - replace h1 := h1 α h3 l₁ l₂ h4 - replace h2 := h2 α h3 l₁ l₂ h4 - simp [RenyiDivergence] at h1 h2 - rw [tsum_prod' ENNReal.summable (fun b : ℤ => ENNReal.summable)] - . simp - conv => - left - right - right - right - right - intro b - right - intro c - rw [compose_sum_rw] - rw [compose_sum_rw] - rw [ENNReal.mul_rpow_of_ne_zero (nn1 l₁ b) (nn2 l₁ c)] - rw [ENNReal.mul_rpow_of_ne_zero (nn1 l₂ b) (nn2 l₂ c)] - rw [mul_assoc] - right - rw [mul_comm] - rw [mul_assoc] - right - rw [mul_comm] - conv => - left - right - right - right - right - intro b - right - intro c - rw [← mul_assoc] - conv => - left - right - right - right - right - intro b - rw [ENNReal.tsum_mul_left] - rw [ENNReal.tsum_mul_right] - rw [ENNReal.toReal_mul] - rw [Real.log_mul] - . rw [mul_add] - have D := _root_.add_le_add h1 h2 - apply le_trans D - rw [← add_mul] - rw [mul_le_mul_iff_of_pos_right] - . rw [← mul_add] - rw [mul_le_mul_iff_of_pos_left] - . ring_nf - simp - . simp - . apply lt_trans zero_lt_one h3 - . apply RenyiNoisedQueryNonZero h3 h4 X nn1 nt1 nts1 - . apply RenyiNoisedQueryNonZero h3 h4 Y nn2 nt2 nts2 - -theorem DPCompose_NonZeroNQ {nq1 nq2 : List T → SLang ℤ} (nn1 : NonZeroNQ nq1) (nn2 : NonZeroNQ nq2) : - NonZeroNQ (Compose nq1 nq2) := by - simp [NonZeroNQ] at * - intro l a b - replace nn1 := nn1 l a - replace nn2 := nn2 l b - simp [Compose] - exists a - simp - intro H - cases H - . rename_i H - contradiction - . rename_i H - contradiction - -theorem DPCompose_NonTopNQ {nq1 nq2 : List T → SLang ℤ} (nt1 : NonTopNQ nq1) (nt2 : NonTopNQ nq2) : - NonTopNQ (Compose nq1 nq2) := by - simp [NonTopNQ] at * - intro l a b - replace nt1 := nt1 l a - replace nt2 := nt2 l b - simp [Compose] - rw [compose_sum_rw] - rw [mul_eq_top] - intro H - cases H - . rename_i H - cases H - contradiction - . rename_i H - cases H - contradiction - -theorem DPCompose_NonTopSum {nq1 nq2 : List T → SLang ℤ} (nt1 : NonTopSum nq1) (nt2 : NonTopSum nq2) : - NonTopSum (Compose nq1 nq2) := by - simp [NonTopSum] at * - intro l - replace nt1 := nt1 l - replace nt2 := nt2 l - simp [Compose] - rw [ENNReal.tsum_prod'] - conv => - right - left - right - intro a - right - intro b - simp - rw [compose_sum_rw] - conv => - right - left - right - intro a - rw [ENNReal.tsum_mul_left] - rw [ENNReal.tsum_mul_right] - rw [mul_eq_top] - intro H - cases H - . rename_i H - cases H - contradiction - . rename_i H - cases H - contradiction - -theorem DPCompose_NonTopRDNQ {nq1 nq2 : List T → SLang ℤ} (nt1 : NonTopRDNQ nq1) (nt2 : NonTopRDNQ nq2) (nn1 : NonZeroNQ nq1) (nn2 : NonZeroNQ nq2) : - NonTopRDNQ (Compose nq1 nq2) := by - simp [NonTopRDNQ] at * - intro α h1 l₁ l₂ h2 - replace nt1 := nt1 α h1 l₁ l₂ h2 - replace nt2 := nt2 α h1 l₁ l₂ h2 - simp [Compose] - rw [ENNReal.tsum_prod'] - simp - conv => - right - left - right - intro x - right - intro y - congr - . left - rw [compose_sum_rw] - . left - rw [compose_sum_rw] - conv => - right - left - right - intro x - right - intro y - rw [ENNReal.mul_rpow_of_ne_zero (nn1 l₁ x) (nn2 l₁ y)] - rw [ENNReal.mul_rpow_of_ne_zero (nn1 l₂ x) (nn2 l₂ y)] - rw [mul_assoc] - right - rw [mul_comm] - rw [mul_assoc] - right - rw [mul_comm] - conv => - right - left - right - intro x - right - intro y - rw [← mul_assoc] - conv => - right - left - right - intro x - rw [ENNReal.tsum_mul_left] - rw [ENNReal.tsum_mul_right] - intro H - rw [mul_eq_top] at H - cases H - . rename_i h3 - cases h3 - rename_i h4 h5 - contradiction - . rename_i h3 - cases h3 - rename_i h4 h5 - contradiction - -end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Mechanism/Mechanism.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Mechanism/Mechanism.lean deleted file mode 100644 index 29df6fb3..00000000 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Mechanism/Mechanism.lean +++ /dev/null @@ -1,242 +0,0 @@ -import SampCert.DifferentialPrivacy.ZeroConcentrated.DP - -noncomputable section - -open Classical Nat Int Real ENNReal MeasureTheory Measure - -namespace SLang - -def NoisedQuery (query : List T → ℤ) (Δ : ℕ+) (ε₁ ε₂ : ℕ+) (l : List T) : SLang ℤ := do - DiscreteGaussianGenSample (Δ * ε₂) ε₁ (query l) - -theorem NoisedQueryDP (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_sensitivity : sensitivity query Δ) : - DP (NoisedQuery query Δ ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by - simp [DP, NoisedQuery] - intros α h1 l₁ l₂ h2 - have A := @DiscreteGaussianGenSampleZeroConcentrated α h1 (Δ * ε₂) ε₁ (query l₁) (query l₂) - apply le_trans A - clear A - replace bounded_sensitivity := bounded_sensitivity l₁ l₂ h2 - ring_nf - simp - conv => - left - left - right - rw [mul_pow] - conv => - left - rw [mul_assoc] - right - rw [mul_comm] - rw [← mul_assoc] - conv => - left - rw [mul_assoc] - right - rw [← mul_assoc] - left - rw [mul_comm] - rw [← mul_assoc] - rw [← mul_assoc] - rw [← mul_assoc] - simp only [inv_pow] - rw [mul_inv_le_iff'] - . have A : (α * ↑↑ε₁ ^ 2 * (↑↑ε₂ ^ 2)⁻¹) ≤ (α * ↑↑ε₁ ^ 2 * (↑↑ε₂ ^ 2)⁻¹) := le_refl (α * ↑↑ε₁ ^ 2 * (↑↑ε₂ ^ 2)⁻¹) - have B : 0 ≤ (α * ↑↑ε₁ ^ 2 * (↑↑ε₂ ^ 2)⁻¹) := by - simp - apply @le_trans ℝ _ 0 1 α (instStrictOrderedCommRingReal.proof_3) (le_of_lt h1) - apply mul_le_mul A _ _ B - . apply sq_le_sq.mpr - simp only [abs_cast] - rw [← Int.cast_sub] - rw [← Int.cast_abs] - apply Int.cast_le.mpr - rw [← Int.natCast_natAbs] - apply Int.ofNat_le.mpr - trivial - . apply sq_nonneg - . rw [pow_two] - rw [_root_.mul_pos_iff] - left - simp - -theorem NoisedQuery_NonZeroNQ (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : - NonZeroNQ (NoisedQuery query Δ ε₁ ε₂) := by - simp [NonZeroNQ, NoisedQuery, DiscreteGaussianGenSample] - intros l n - exists (n - query l) - simp - have A : ((Δ : ℝ) * ε₂ / ε₁) ≠ 0 := by - simp - have X := @discrete_gaussian_pos (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) A 0 (n - query l) - simp at X - trivial - -theorem NoisedQuery_NonTopNQ (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : - NonTopNQ (NoisedQuery query Δ ε₁ ε₂) := by - simp [NonTopNQ, NoisedQuery, DiscreteGaussianGenSample] - intro l n - rw [ENNReal.tsum_eq_add_tsum_ite (n - query l)] - simp - have X : ∀ x : ℤ, (@ite ℝ≥0∞ (x = n - query l) (propDecidable (x = n - query l)) 0 - (@ite ℝ≥0∞ (n = x + query l) (instDecidableEqInt n (x + query l)) - (ENNReal.ofReal (discrete_gaussian (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) 0 ↑x)) 0)) = 0 := by - intro x - split - . simp - . split - . rename_i h1 h2 - subst h2 - simp at h1 - . simp - conv => - right - left - right - intro x - rw [X] - simp - -theorem discrete_gaussian_shift {σ : ℝ} (h : σ ≠ 0) (μ : ℝ) (τ x : ℤ) : - discrete_gaussian σ μ (x - τ) = discrete_gaussian σ (μ + τ) (x) := by - simp [discrete_gaussian] - congr 1 - . simp [gauss_term_ℝ] - congr 3 - ring_nf - . rw [SG_periodic h] - -theorem NoisedQuery_NonTopSum (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : - NonTopSum (NoisedQuery query Δ ε₁ ε₂) := by - simp [NonTopSum, NoisedQuery, DiscreteGaussianGenSample] - intro l - have X : ∀ n: ℤ, ∀ x : ℤ, (@ite ℝ≥0∞ (x = n - query l) (propDecidable (x = n - query l)) 0 - (@ite ℝ≥0∞ (n = x + query l) (instDecidableEqInt n (x + query l)) - (ENNReal.ofReal (discrete_gaussian (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) 0 ↑x)) 0)) = 0 := by - intro n x - split - . simp - . split - . rename_i h1 h2 - subst h2 - simp at h1 - . simp - conv => - right - left - right - intro n - rw [ENNReal.tsum_eq_add_tsum_ite (n - query l)] - simp - right - right - intro x - rw [X] - simp - have A : (Δ : ℝ) * ε₂ / ε₁ ≠ 0 := by - simp - conv => - right - left - right - intro n - rw [discrete_gaussian_shift A] - simp - rw [← ENNReal.ofReal_tsum_of_nonneg] - . rw [discrete_gaussian_normalizes A] - simp - . apply discrete_gaussian_nonneg A - . apply discrete_gaussian_summable' A (query l) - -theorem NoisedQuery_NonTopRDNQ (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : - NonTopRDNQ (NoisedQuery query Δ ε₁ ε₂) := by - simp [NonTopRDNQ, NoisedQuery, DiscreteGaussianGenSample] - intro α _ l₁ l₂ _ - have A : ∀ x_1 x : ℤ, (@ite ℝ≥0∞ (x_1 = x - query l₁) (propDecidable (x_1 = x - query l₁)) 0 - (@ite ℝ≥0∞ (x = x_1 + query l₁) (instDecidableEqInt x (x_1 + query l₁)) - (ENNReal.ofReal (discrete_gaussian (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) 0 ↑x_1)) 0 )) = 0 := by - intro x y - split - . simp - . split - . rename_i h1 h2 - subst h2 - simp at h1 - . simp - have B : ∀ x_1 x : ℤ, (@ite ℝ≥0∞ (x_1 = x - query l₂) (propDecidable (x_1 = x - query l₂)) 0 - (@ite ℝ≥0∞ (x = x_1 + query l₂) (instDecidableEqInt x (x_1 + query l₂)) - (ENNReal.ofReal (discrete_gaussian (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) 0 ↑x_1)) 0)) = 0 := by - intro x y - split - . simp - . split - . rename_i h1 h2 - subst h2 - simp at h1 - . simp - conv => - right - left - right - intro x - left - rw [ENNReal.tsum_eq_add_tsum_ite (x - query l₁)] - simp - left - right - right - intro y - rw [A] - simp - conv => - right - left - right - intro x - right - rw [ENNReal.tsum_eq_add_tsum_ite (x - query l₂)] - simp - left - right - right - intro y - rw [B] - simp - clear A B - have P : (Δ : ℝ) * ε₂ / ε₁ ≠ 0 := by - simp - have A : ∀ x : ℤ, ∀ l : List T, 0 < discrete_gaussian (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) 0 (↑x - ↑(query l)) := by - intro x l - have A' := @discrete_gaussian_pos _ P 0 (x - query l) - simp at A' - trivial - have B : ∀ x : ℤ, 0 ≤ discrete_gaussian (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) 0 (↑x - ↑(query l₁)) ^ α := by - intro x - have A' := @discrete_gaussian_nonneg _ P 0 (x - query l₁) - simp at A' - apply rpow_nonneg A' - conv => - right - left - right - intro x - rw [ENNReal.ofReal_rpow_of_pos (A x l₁)] - rw [ENNReal.ofReal_rpow_of_pos (A x l₂)] - rw [← ENNReal.ofReal_mul (B x)] - rw [← ENNReal.ofReal_tsum_of_nonneg] - . simp - . intro n - have X := @RenyiSumSG_nonneg _ α P (query l₁) (query l₂) n - rw [discrete_gaussian_shift P] - rw [discrete_gaussian_shift P] - simp [X] - . have X := @SummableRenyiGauss _ P (query l₁) (query l₂) - conv => - right - intro x - rw [discrete_gaussian_shift P] - rw [discrete_gaussian_shift P] - simp [X] - -end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Mechanism/Properties.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Mechanism/Properties.lean index 764842dc..9b5034b5 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Mechanism/Properties.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Mechanism/Properties.lean @@ -5,6 +5,7 @@ Authors: Jean-Baptiste Tristan -/ import SampCert.DifferentialPrivacy.ZeroConcentrated.DP +import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Mechanism.Code noncomputable section @@ -12,9 +13,6 @@ open Classical Nat Int Real ENNReal MeasureTheory Measure namespace SLang -def NoisedQuery (query : List T → ℤ) (Δ : ℕ+) (ε₁ ε₂ : ℕ+) (l : List T) : SLang ℤ := do - DiscreteGaussianGenSample (Δ * ε₂) ε₁ (query l) - theorem NoisedQueryDP (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_sensitivity : sensitivity query Δ) : DP (NoisedQuery query Δ ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by simp [DP, NoisedQuery] diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Postprocessing.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Postprocessing.lean deleted file mode 100644 index 94205f84..00000000 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Postprocessing.lean +++ /dev/null @@ -1,691 +0,0 @@ -import SampCert.DifferentialPrivacy.ZeroConcentrated.DP - -noncomputable section - -open Classical Nat Int Real ENNReal MeasureTheory Measure - -namespace SLang - -def PostProcess (nq : List T → SLang U) (pp : U → ℤ) (l : List T) : SLang ℤ := do - let A ← nq l - return pp A - -theorem foo (f : U → ℤ) (g : U → ENNReal) (x : ℤ) : - (∑' a : U, if x = f a then g a else 0) = ∑' a : { a | x = f a }, g a := by - have A := @tsum_split_ite U (fun a : U => x = f a) g (fun _ => 0) - simp only [decide_eq_true_eq, tsum_zero, add_zero] at A - rw [A] - have B : ↑{i | decide (x = f i) = true} = ↑{a | x = f a} := by - simp - rw [B] - -variable {T : Type} -variable [m1 : MeasurableSpace T] -variable [m2 : MeasurableSingletonClass T] -variable [m3: MeasureSpace T] - -theorem Integrable_rpow (f : T → ℝ) (nn : ∀ x : T, 0 ≤ f x) (μ : Measure T) (α : ENNReal) (mem : Memℒp f α μ) (h1 : α ≠ 0) (h2 : α ≠ ⊤) : - MeasureTheory.Integrable (fun x : T => (f x) ^ α.toReal) μ := by - have X := @MeasureTheory.Memℒp.integrable_norm_rpow T ℝ m1 μ _ f α mem h1 h2 - revert X - conv => - left - left - intro x - rw [← norm_rpow_of_nonneg (nn x)] - intro X - simp [Integrable] at * - constructor - . cases X - rename_i left right - rw [@aestronglyMeasurable_iff_aemeasurable] - apply AEMeasurable.pow_const - simp [Memℒp] at mem - cases mem - rename_i left' right' - rw [aestronglyMeasurable_iff_aemeasurable] at left' - simp [left'] - . rw [← hasFiniteIntegral_norm_iff] - simp [X] - -theorem bar (f : T → ℝ) (q : PMF T) (α : ℝ) (h : 1 < α) (h2 : ∀ x : T, 0 ≤ f x) (mem : Memℒp f (ENNReal.ofReal α) (PMF.toMeasure q)) : - ((∑' x : T, (f x) * (q x).toReal)) ^ α ≤ (∑' x : T, (f x) ^ α * (q x).toReal) := by - - conv => - left - left - right - intro x - rw [mul_comm] - rw [← smul_eq_mul] - conv => - right - right - intro x - rw [mul_comm] - rw [← smul_eq_mul] - rw [← PMF.integral_eq_tsum] - rw [← PMF.integral_eq_tsum] - - have A := @convexOn_rpow α (le_of_lt h) - have B : ContinuousOn (fun (x : ℝ) => x ^ α) (Set.Ici 0) := by - apply ContinuousOn.rpow - . exact continuousOn_id' (Set.Ici 0) - . exact continuousOn_const - . intro x h' - simp at h' - have OR : x = 0 ∨ 0 < x := by exact LE.le.eq_or_gt h' - cases OR - . rename_i h'' - subst h'' - right - apply lt_trans zero_lt_one h - . rename_i h'' - left - by_contra - rename_i h3 - subst h3 - simp at h'' - have C : @IsClosed ℝ UniformSpace.toTopologicalSpace (Set.Ici 0) := by - exact isClosed_Ici - have D := @ConvexOn.map_integral_le T ℝ m1 _ _ _ (PMF.toMeasure q) (Set.Ici 0) f (fun (x : ℝ) => x ^ α) (PMF.toMeasure.isProbabilityMeasure q) A B C - simp at D - apply D - . exact MeasureTheory.ae_of_all (PMF.toMeasure q) h2 - . apply MeasureTheory.Memℒp.integrable _ mem - rw [one_le_ofReal] - apply le_of_lt h - . rw [Function.comp_def] - have X : ENNReal.ofReal α ≠ 0 := by - simp - apply lt_trans zero_lt_one h - have Y : ENNReal.ofReal α ≠ ⊤ := by - simp - have Z := @Integrable_rpow T m1 f h2 (PMF.toMeasure q) (ENNReal.ofReal α) mem X Y - rw [toReal_ofReal] at Z - . exact Z - . apply le_of_lt - apply lt_trans zero_lt_one h - . have X : ENNReal.ofReal α ≠ 0 := by - simp - apply lt_trans zero_lt_one h - have Y : ENNReal.ofReal α ≠ ⊤ := by - simp - have Z := @Integrable_rpow T m1 f h2 (PMF.toMeasure q) (ENNReal.ofReal α) mem X Y - rw [toReal_ofReal] at Z - . exact Z - . apply le_of_lt - apply lt_trans zero_lt_one h - . apply MeasureTheory.Memℒp.integrable _ mem - rw [one_le_ofReal] - apply le_of_lt h - -variable {U : Type} -variable [m2 : MeasurableSpace U] -- [m2' : MeasurableSingletonClass U] -variable [count : Countable U] -variable [disc : DiscreteMeasurableSpace U] - -def δ (nq : SLang U) (f : U → ℤ) (a : ℤ) : {n : U | a = f n} → ENNReal := fun x : {n : U | a = f n} => nq x * (∑' (x : {n | a = f n}), nq x)⁻¹ - -theorem δ_normalizes (nq : SLang U) (f : U → ℤ) (a : ℤ) (h1 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ 0) (h2 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ ⊤) : - HasSum (δ nq f a) 1 := by - rw [Summable.hasSum_iff ENNReal.summable] - unfold δ - rw [ENNReal.tsum_mul_right] - rw [ENNReal.mul_inv_cancel h1 h2] - -def δpmf (nq : SLang U) (f : U → ℤ) (a : ℤ) (h1 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ 0) (h2 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ ⊤) : PMF {n : U | a = f n} := - ⟨ δ nq f a , δ_normalizes nq f a h1 h2 ⟩ - -theorem δpmf_conv (nq : SLang U) (a : ℤ) (x : {n | a = f n}) (h1 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ 0) (h2 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ ⊤) : - nq x * (∑' (x : {n | a = f n}), nq x)⁻¹ = (δpmf nq f a h1 h2) x := by - simp [δpmf] - conv => - right - left - left - -theorem δpmf_conv' (nq : SLang U) (f : U → ℤ) (a : ℤ) (h1 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ 0) (h2 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ ⊤) : - (fun x : {n | a = f n} => nq x * (∑' (x : {n | a = f n}), nq x)⁻¹) = (δpmf nq f a h1 h2) := by - ext x - rw [δpmf_conv] - -theorem witness {f : U → ℤ} {i : ℤ} (h : ¬{b | i = f b} = ∅) : - ∃ x : U, i = f x := by - rw [← nonempty_subtype] - exact Set.nonempty_iff_ne_empty'.mpr h - -theorem norm_simplify (x : ENNReal) (h : x ≠ ⊤) : - @nnnorm ℝ SeminormedAddGroup.toNNNorm x.toReal = x := by - simp [nnnorm] - cases x - . contradiction - . rename_i v - simp - rfl - -theorem RD1 (p q : T → ENNReal) (α : ℝ) (h : 1 < α) (RD : ∑' (x : T), p x ^ α * q x ^ (1 - α) ≠ ⊤) (nz : ∀ x : T, q x ≠ 0) (nt : ∀ x : T, q x ≠ ⊤) : - ∑' (x : T), (p x / q x) ^ α * q x ≠ ⊤ := by - rw [← RenyiDivergenceExpectation p q h nz nt] - trivial - -theorem ENNReal.HasSum_fiberwise {f : T → ENNReal} {a : ENNReal} (hf : HasSum f a) (g : T → ℤ) : - HasSum (fun c : ℤ ↦ ∑' b : g ⁻¹' {c}, f b) a := by - let A := (Equiv.sigmaFiberEquiv g) - have B := @Equiv.hasSum_iff ENNReal T ((y : ℤ) × { x // g x = y }) _ _ f a A - replace B := B.2 hf - have C := @HasSum.sigma ENNReal ℤ _ _ _ _ (fun y : ℤ => { x // g x = y }) (f ∘ ⇑(Equiv.sigmaFiberEquiv g)) (fun c => ∑' (b : ↑(g ⁻¹' {c})), f ↑b) a B - apply C - intro b - have F := @Summable.hasSum_iff ENNReal _ _ _ (fun c => (f ∘ ⇑(Equiv.sigmaFiberEquiv g)) { fst := b, snd := c }) ((fun c => ∑' (b : ↑(g ⁻¹' {c})), f ↑b) b) _ - apply (F _).2 - . rfl - . apply ENNReal.summable - -theorem ENNReal.tsum_fiberwise (p : T → ENNReal) (f : T → ℤ) : - ∑' (x : ℤ), ∑' (b : (f ⁻¹' {x})), p b - = ∑' i : T, p i := by - apply HasSum.tsum_eq - apply ENNReal.HasSum_fiberwise - apply Summable.hasSum - exact ENNReal.summable - -theorem quux (p : T → ENNReal) (f : T → ℤ) : - (∑' i : T, p i) - = ∑' (x : ℤ), if {a : T | x = f a} = {} then 0 else ∑'(i : {a : T | x = f a}), p i := by - rw [← ENNReal.tsum_fiberwise p f] - have A : ∀ x, f ⁻¹' {x} = { a | x = f a } := by - intro x - simp [Set.preimage] - rw [Set.ext_iff] - simp - intro y - exact eq_comm - conv => - left - right - intro x - rw [A] - clear A - apply tsum_congr - intro b - split - . rename_i h' - rw [h'] - simp only [tsum_empty] - . simp - -theorem convergent_subset {p : T → ENNReal} (f : T → ℤ) (conv : ∑' (x : T), p x ≠ ⊤) : - ∑' (x : { y : T| x = f y }), p x ≠ ⊤ := by - rw [← foo] - have A : (∑' (y : T), if x = f y then p y else 0) ≤ ∑' (x : T), p x := by - apply tsum_le_tsum - . intro i - split - . trivial - . simp only [_root_.zero_le] - . exact ENNReal.summable - . exact ENNReal.summable - rw [← lt_top_iff_ne_top] - apply lt_of_le_of_lt A - rw [lt_top_iff_ne_top] - trivial - -theorem ENNReal.tsum_pos {f : T → ENNReal} (h1 : ∑' x : T, f x ≠ ⊤) (h2 : ∀ x : T, f x ≠ 0) (i : T) : - 0 < ∑' x : T, f x := by - apply (toNNReal_lt_toNNReal ENNReal.zero_ne_top h1).mp - simp only [zero_toNNReal] - rw [ENNReal.tsum_toNNReal_eq (ENNReal.ne_top_of_tsum_ne_top h1)] - have S : Summable fun a => (f a).toNNReal := by - rw [← tsum_coe_ne_top_iff_summable] - conv => - left - right - intro b - rw [ENNReal.coe_toNNReal (ENNReal.ne_top_of_tsum_ne_top h1 b)] - trivial - have B:= @NNReal.tsum_pos T (fun (a : T) => (f a).toNNReal) S i - apply B - apply ENNReal.toNNReal_pos (h2 i) (ENNReal.ne_top_of_tsum_ne_top h1 i) - -theorem ENNReal.tsum_pos_int {f : ℤ → ENNReal} (h1 : ∑' x : ℤ, f x ≠ ⊤) (h2 : ∀ x : ℤ, f x ≠ 0) : - 0 < ∑' x : ℤ, f x := by - apply ENNReal.tsum_pos h1 h2 42 - -theorem tsum_pos_int {f : ℤ → ENNReal} (h1 : ∑' x : ℤ, f x ≠ ⊤) (h2 : ∀ x : ℤ, f x ≠ 0) : - 0 < (∑' x : ℤ, f x).toReal := by - have X : 0 = (0 : ENNReal).toReal := rfl - rw [X] - clear X - apply toReal_strict_mono h1 - apply ENNReal.tsum_pos_int h1 h2 - -theorem DPostPocess_pre {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP nq ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (conv : NonTopSum nq) (f : U → ℤ) {α : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (h2 : Neighbour l₁ l₂) : - (∑' (x : ℤ), - (∑' (a : U), if x = f a then nq l₁ a else 0) ^ α * - (∑' (a : U), if x = f a then nq l₂ a else 0) ^ (1 - α)) ≤ - (∑' (x : U), nq l₁ x ^ α * nq l₂ x ^ (1 - α)) := by - - simp [DP, RenyiDivergence] at h - - -- Rewrite as cascading expectations - rw [@RenyiDivergenceExpectation _ (nq l₁) (nq l₂) _ h1 (nn l₂) (nts l₂)] - - -- Shuffle the sum - rw [quux (fun x => (nq l₁ x / nq l₂ x) ^ α * nq l₂ x) f] - - apply ENNReal.tsum_le_tsum - - intro i - - -- Get rid of elements with probability 0 in the pushforward - split - . rename_i empty - rw [foo] - have ZE : (∑' (x_1 : ↑{n | i = f n}), nq l₁ ↑x_1) = 0 := by - simp - intro a H - have I₁ : a ∈ {b | i = f b} := by - simp [H] - have I2 : {b | i = f b} ≠ ∅ := by - apply ne_of_mem_of_not_mem' I₁ - simp - contradiction - rw [ZE] - simp only [toReal_mul, zero_toReal, ge_iff_le] - - rw [ENNReal.zero_rpow_of_pos] - . simp - . apply lt_trans zero_lt_one h1 - - -- Part 2: apply Jensen's inequality - . rename_i NotEmpty - - have MasterRW : ∀ l : List T, ∑' (a : ↑{a | i = f a}), nq l ↑a ≠ ⊤ := by - intro l - apply convergent_subset - simp [NonTopSum] at conv - have conv := conv l - apply conv - - have MasterZero : ∀ l : List T, ∑' (a : ↑{a | i = f a}), nq l ↑a ≠ 0 := by - intro l - simp - have T := witness NotEmpty - cases T - rename_i z w - exists z - constructor - . trivial - . apply nn l - - have S2 : (∑' (a : ↑{n | i = f n}), nq l₁ ↑a / nq l₂ ↑a * (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) a) ^ α ≠ ⊤ := by - conv => - left - left - right - intro a - rw [← δpmf_conv] - rw [division_def] - rw [mul_assoc] - right - rw [← mul_assoc] - rw [ENNReal.inv_mul_cancel (nn l₂ a) (nts l₂ a)] - rw [one_mul] - rw [ENNReal.tsum_mul_right] - apply ENNReal.rpow_ne_top_of_nonneg (le_of_lt (lt_trans zero_lt_one h1 )) - apply mul_ne_top - . apply convergent_subset _ (conv l₁) - . apply inv_ne_top.mpr (MasterZero l₂) - - have S1 : ∀ (a : ↑{n | i = f n}), nq l₁ ↑a / nq l₂ ↑a * (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) a ≠ ⊤ := by - intro a - apply mul_ne_top - . rw [division_def] - apply mul_ne_top (nts l₁ a) - apply inv_ne_top.mpr (nn l₂ a) - . rw [← δpmf_conv] - apply mul_ne_top (nts l₂ a) - apply inv_ne_top.mpr (MasterZero l₂) - - have S3 : ∑' (a : ↑{n | i = f n}), (nq l₁ ↑a / nq l₂ ↑a) ^ α * (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) a ≠ ⊤ := by - conv => - left - right - intro a - rw [← δpmf_conv] - rw [← mul_assoc] - rw [ENNReal.tsum_mul_right] - apply mul_ne_top - . rw [← RenyiDivergenceExpectation _ _ h1] - . replace nt := nt α h1 l₁ l₂ h2 - apply convergent_subset _ nt - . intro x - apply nn - . intro x - apply nts - . apply inv_ne_top.mpr (MasterZero l₂) - - have S4 : ∀ (a : ↑{n | i = f n}), (nq l₁ ↑a / nq l₂ ↑a) ^ α * (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) a ≠ ⊤ := by - intro a - apply ENNReal.ne_top_of_tsum_ne_top S3 - - rw [foo] - rw [foo] - - -- Introduce Q(f⁻¹ i) - let κ := ∑' x : {n : U | i = f n}, nq l₂ x - have P4 : κ / κ = 1 := by - rw [division_def] - rw [ENNReal.mul_inv_cancel] - . simp [κ] -- Use here for δ normalization - have T := witness NotEmpty - cases T - rename_i z w - exists z - constructor - . trivial - . apply nn l₂ - . simp only [κ] - apply MasterRW l₂ - - conv => - right - right - intro a - rw [← mul_one ((nq l₁ ↑a / nq l₂ ↑a) ^ α * nq l₂ ↑a)] - right - rw [← P4] - clear P4 - simp only [κ] - - conv => - right - right - intro a - right - rw [division_def] - rw [mul_comm] - - conv => - right - right - intro a - rw [← mul_assoc] - - rw [ENNReal.tsum_mul_right] - - -- Jensen's inequality - - have P5 : ∀ (x : ↑{n | i = f n}), 0 ≤ (fun a => (nq l₁ ↑a / nq l₂ ↑a).toReal) x := by - intro x - simp only [toReal_nonneg] - - have XXX : @Memℒp ℝ Real.normedAddCommGroup (↑{n | i = f n}) Subtype.instMeasurableSpace (fun a => (nq l₁ ↑a / nq l₂ ↑a).toReal) - (ENNReal.ofReal α) (PMF.toMeasure (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂))) := by - simp [Memℒp] - constructor - . apply MeasureTheory.StronglyMeasurable.aestronglyMeasurable - apply Measurable.stronglyMeasurable - apply Measurable.ennreal_toReal - conv => - right - intro x - rw [division_def] - apply Measurable.mul - . -- MeasurableSingletonClass.toDiscreteMeasurableSpace - apply measurable_discrete - . apply Measurable.inv - apply measurable_discrete - . simp [snorm] - split - . simp - . simp [snorm'] - rw [MeasureTheory.lintegral_countable'] -- Uses countable - rw [toReal_ofReal (le_of_lt (lt_trans zero_lt_one h1))] - have OTHER : ∀ a, nq l₁ a / nq l₂ a ≠ ⊤ := by - intro a - rw [division_def] - rw [ne_iff_lt_or_gt] - left - rw [mul_lt_top_iff] - left - constructor - . exact Ne.lt_top' (id (Ne.symm (nts l₁ a))) - . simp - exact pos_iff_ne_zero.mpr (nn l₂ a) - - conv => - left - left - right - intro a - rw [norm_simplify _ (OTHER a)] - have Z : 0 < α⁻¹ := by - simp - apply lt_trans zero_lt_one h1 - rw [rpow_lt_top_iff_of_pos Z] - conv => - left - right - intro a - rw [PMF.toMeasure_apply_singleton _ _ (measurableSet_singleton a)] - - apply Ne.lt_top' (id (Ne.symm _)) - apply S3 - - - have Jensen's := @bar {n : U | i = f n} Subtype.instMeasurableSpace Subtype.instMeasurableSingletonClass (fun a => (nq l₁ a / nq l₂ a).toReal) (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) α h1 P5 XXX - clear P5 - - have P6 : 0 ≤ (∑' (x : ↑{n | i = f n}), nq l₂ ↑x).toReal := by - simp only [toReal_nonneg] - have A' := mul_le_mul_of_nonneg_left Jensen's P6 - clear Jensen's P6 - - conv => - right - rw [mul_comm] - right - right - intro a - rw [mul_assoc] - rw [δpmf_conv] - - -- Here - - replace A' := ofReal_le_ofReal A' - rw [ofReal_mul toReal_nonneg] at A' - rw [ofReal_mul toReal_nonneg] at A' - rw [ofReal_toReal_eq_iff.2 (MasterRW l₂)] at A' - simp only at A' - - revert A' - conv => - left - right - right - right - right - intro x - rw [toReal_rpow] - rw [← toReal_mul] - conv => - left - right - right - right - rw [← ENNReal.tsum_toReal_eq S4] - intro A' - rw [ofReal_toReal_eq_iff.2 S3] at A' - - apply le_trans _ A' - clear A' - apply le_of_eq - - -- Part 3: - - conv => - right - right - right - left - right - intro x - rw [← toReal_mul] - rw [← ENNReal.tsum_toReal_eq S1] - rw [toReal_rpow] - rw [ofReal_toReal_eq_iff.2 S2] - - conv => - right - right - left - right - intro x - rw [division_def] - rw [← δpmf_conv] - rw [mul_assoc] - right - rw [← mul_assoc] - left - rw [ENNReal.inv_mul_cancel (nn l₂ x) (nts l₂ x)] - simp only [one_mul] - - rw [ENNReal.tsum_mul_right] - have H1 : 0 ≤ ∑' (x : ↑{n | i = f n}), (nq l₁ ↑x).toReal := by - apply tsum_nonneg - simp - have H2 : 0 ≤ (∑' (a : ↑{a | i = f a}), nq l₂ ↑a)⁻¹.toReal := by - apply toReal_nonneg - have H4 : (∑' (a : ↑{a | i = f a}), nq l₂ ↑a)⁻¹ ≠ ⊤ := by - apply inv_ne_top.mpr - simp - have T := witness NotEmpty - cases T - rename_i z w - exists z - constructor - . trivial - . apply nn l₂ - rw [ENNReal.mul_rpow_of_ne_top (MasterRW l₁) H4] - - have H3 : ∑' (a : ↑{a | i = f a}), nq l₂ ↑a ≠ 0 := by - simp - have T := witness NotEmpty - cases T - rename_i z w - exists z - constructor - . trivial - . apply nn l₂ - rw [ENNReal.rpow_sub _ _ H3 (MasterRW l₂)] - rw [ENNReal.rpow_one] - rw [division_def] - rw [← mul_assoc] - rw [← mul_assoc] - congr 1 - . rw [mul_comm] - . congr 1 - rw [ENNReal.inv_rpow] - -theorem tsum_ne_zero_of_ne_zero {T : Type} [Inhabited T] (f : T → ENNReal) (h : ∀ x : T, f x ≠ 0) : - ∑' x : T, f x ≠ 0 := by - by_contra CONTRA - rw [ENNReal.tsum_eq_zero] at CONTRA - have A := h default - have B := CONTRA default - contradiction - -variable [Inhabited U] - -theorem DPPostProcess {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP nq ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (conv : NonTopSum nq) (f : U → ℤ) : - DP (PostProcess nq f) ((ε₁ : ℝ) / ε₂) := by - simp [PostProcess, DP, RenyiDivergence] - intro α h1 l₁ l₂ h2 - have h' := h - simp [DP, RenyiDivergence] at h' - replace h' := h' α h1 l₁ l₂ h2 - - -- Part 1, removing fluff - - apply le_trans _ h' - clear h' - - -- remove the α scaling - have A : 0 ≤ (α - 1)⁻¹ := by - simp - apply le_of_lt h1 - apply mul_le_mul_of_nonneg_left _ A - clear A - - have RDConvegence : ∑' (x : U), nq l₁ x ^ α * nq l₂ x ^ (1 - α) ≠ ⊤ := by - simp [NonTopRDNQ] at nt - have nt := nt α h1 l₁ l₂ h2 - trivial - - have B := DPostPocess_pre h nn nt nts conv f h1 h2 - have B' : ∑' (x : ℤ), (∑' (a : U), if x = f a then nq l₁ a else 0) ^ α * (∑' (a : U), if x = f a then nq l₂ a else 0) ^ (1 - α) ≠ ⊤ := by - by_contra CONTRA - rw [CONTRA] at B - simp at B - contradiction - - -- remove the log - apply log_le_log _ (toReal_mono RDConvegence B) - apply toReal_pos _ B' - apply (tsum_ne_zero_iff ENNReal.summable).mpr - exists (f default) - - rw [ENNReal.tsum_eq_add_tsum_ite default] - conv => - left - right - rw [ENNReal.tsum_eq_add_tsum_ite default] - simp only [reduceIte] - apply mul_ne_zero - . by_contra CONTRA - rw [ENNReal.rpow_eq_zero_iff_of_pos (lt_trans zero_lt_one h1)] at CONTRA - simp at CONTRA - cases CONTRA - rename_i left right - have Y := nn l₁ default - contradiction - . by_contra CONTRA - rw [ENNReal.rpow_eq_zero_iff] at CONTRA - cases CONTRA - . rename_i CONTRA - cases CONTRA - rename_i left right - simp at left - cases left - rename_i le1 le2 - have Y := nn l₂ default - contradiction - . rename_i CONTRA - cases CONTRA - rename_i left right - simp at left - cases left - . rename_i left - have Y := nts l₂ default - contradiction - . rename_i left - have Rem := conv l₂ - have X : (∑' (x : U), if x = default then 0 else if f default = f x then nq l₂ x else 0) ≤ ∑' (n : U), nq l₂ n := by - apply ENNReal.tsum_le_tsum - intro a - split - . simp - . split - . simp - . simp - replace Rem := Ne.symm Rem - have Y := Ne.lt_top' Rem - have Z : (∑' (x : U), if x = default then 0 else if f default = f x then nq l₂ x else 0) < ⊤ := by - apply lt_of_le_of_lt X Y - rw [lt_top_iff_ne_top] at Z - contradiction - - -end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Properties.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Properties.lean index faef6fec..5b2f0b47 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Properties.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Properties.lean @@ -5,6 +5,7 @@ Authors: Jean-Baptiste Tristan -/ import SampCert.DifferentialPrivacy.ZeroConcentrated.DP +import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Postprocessing.Code noncomputable section @@ -12,10 +13,6 @@ open Classical Nat Int Real ENNReal MeasureTheory Measure namespace SLang -def PostProcess (nq : List T → SLang U) (pp : U → ℤ) (l : List T) : SLang ℤ := do - let A ← nq l - return pp A - theorem foo (f : U → ℤ) (g : U → ENNReal) (x : ℤ) : (∑' a : U, if x = f a then g a else 0) = ∑' a : { a | x = f a }, g a := by have A := @tsum_split_ite U (fun a : U => x = f a) g (fun _ => 0) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/BoundedMean.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/BoundedMean.lean deleted file mode 100644 index 71af5711..00000000 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/BoundedMean.lean +++ /dev/null @@ -1,60 +0,0 @@ -/- -Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Jean-Baptiste Tristan --/ - -import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Composition.Composition -import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Postprocessing.Postprocessing -import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.Count.Count -import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.BoundedSum.BoundedSum - -open Classical Nat Int Real - -noncomputable section - -namespace SLang - -def NoisedBoundedAvgQuery (U : ℕ+) (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℤ := do - let S ← NoisedBoundedSumQuery U ε₁ (2 * ε₂) l - let C ← NoisedCountingQuery ε₁ (2 * ε₂) l - return S / C - -def NoisedBoundedAvgQuery' (U : ℕ+) (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℤ := - let X := Compose (NoisedBoundedSumQuery U ε₁ (2 * ε₂)) (NoisedCountingQuery ε₁ (2 * ε₂)) - PostProcess X (fun z => z.1 / z.2) l - -theorem NoisedBoundedAvgQueryIdentical (U : ℕ+) (ε₁ ε₂ : ℕ+) : - NoisedBoundedAvgQuery' U ε₁ ε₂ = NoisedBoundedAvgQuery U ε₁ ε₂ := by - ext l x - simp [NoisedBoundedAvgQuery, NoisedBoundedAvgQuery', PostProcess, Compose] - -theorem BoundedSumQueryDP (U : ℕ+) (ε₁ ε₂ : ℕ+) : DP (NoisedBoundedAvgQuery U ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by - rw [← NoisedBoundedAvgQueryIdentical] - unfold NoisedBoundedAvgQuery' - simp only - - have A₁ := @NoisedCountingQueryDP ℕ ε₁ (2 * ε₂) - have A₂ := @NoisedCountingQuery_NonZeroNQ ℕ ε₁ (2 * ε₂) - have A₃ := @NoisedCountingQuery_NonTopNQ ℕ ε₁ (2 * ε₂) - have A₄ := @NoisedCountingQuery_NonTopRDNQ ℕ ε₁ (2 * ε₂) - have A₅ := @NoisedCountingQuery_NonTopSum ℕ ε₁ (2 * ε₂) - - have B₁ := NoisedBoundedSumQueryDP U ε₁ (2 * ε₂) - have B₂ := NoisedBoundedSumQuery_NonZeroNQ U ε₁ (2 * ε₂) - have B₃ := NoisedBoundedSumQuery_NonTopNQ U ε₁ (2 * ε₂) - have B₄ := NoisedBoundedSumQuery_NonTopRDNQ U ε₁ (2 * ε₂) - have B₅ := NoisedBoundedSumQuery_NonTopSum U ε₁ (2 * ε₂) - - have C₁ := DPCompose B₁ A₁ B₂ A₂ B₄ A₄ B₃ A₃ - have C₂ := DPCompose_NonZeroNQ B₂ A₂ - have C₃ := DPCompose_NonTopNQ B₃ A₃ - have C₄ := DPCompose_NonTopRDNQ B₄ A₄ B₂ A₂ - have C₅ := DPCompose_NonTopSum B₅ A₅ - simp at * - ring_nf at * - rw [← division_def] at * - have D := DPPostProcess C₁ C₂ C₄ C₃ C₅ (fun z => z.1 /z.2) - exact D - -end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/Code.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/Code.lean index 70a25010..76d93c2f 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/Code.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/Code.lean @@ -4,10 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jean-Baptiste Tristan -/ -import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Composition.Composition -import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Postprocessing.Postprocessing -import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.Count.Count -import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.BoundedSum.BoundedSum +import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Composition.Code +import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Postprocessing.Code +import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.Count.Code +import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.BoundedSum.Code open Classical Nat Int Real diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/Properties.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/Properties.lean index 3ac6832e..116db798 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/Properties.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/Properties.lean @@ -4,10 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jean-Baptiste Tristan -/ -import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Composition.Composition -import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Postprocessing.Postprocessing -import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.Count.Count -import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.BoundedSum.BoundedSum +import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Composition.Basic +import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Postprocessing.Basic +import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.Count.Basic +import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.BoundedSum.Basic import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.BoundedMean.Code open Classical Nat Int Real diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/BoundedSum.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/BoundedSum.lean deleted file mode 100644 index 55b8413e..00000000 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/BoundedSum.lean +++ /dev/null @@ -1,98 +0,0 @@ -/- -Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Jean-Baptiste Tristan --/ - -import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Mechanism.Mechanism -import Mathlib.Algebra.Group.Defs -import Mathlib.Init.Algebra.Classes -import Init.Data.Int.Order - -open Classical Nat Int Real - -noncomputable section - -namespace SLang - -def BoundedSumQuery (U : ℕ+) (l : List ℕ) : ℤ := - List.sum (List.map (fun n : ℕ => (Nat.min n U)) l) - -theorem BoundedSumQuerySensitivity (U : ℕ+) : sensitivity (BoundedSumQuery U) U := by - simp [sensitivity, BoundedSumQuery] - intros l₁ l₂ H - have A : ∀ n : ℕ, (@min ℤ instMinInt (n : ℤ) (U : ℤ) = n) ∨ (@min ℤ instMinInt n U = U) := by - intro n - simp - exact Nat.le_total n ↑U - cases H - . rename_i l l' n h1 h2 - subst h1 h2 - simp - cases A n - . rename_i h - rw [h] - simp at * - trivial - . rename_i h - rw [h] - simp - . rename_i l n l' h1 h2 - subst h1 h2 - simp - cases A n - . rename_i h - rw [h] - simp at * - trivial - . rename_i h - rw [h] - simp - . rename_i l n l' m h1 h2 - subst h1 h2 - simp - cases A n - . rename_i h - cases A m - . rename_i h' - rw [h, h'] - simp at * - apply Int.natAbs_coe_sub_coe_le_of_le h h' - . rename_i h' - rw [h, h'] - simp at * - apply Int.natAbs_coe_sub_coe_le_of_le h le_rfl - . rename_i h - cases A m - . rename_i h' - rw [h, h'] - simp at * - apply Int.natAbs_coe_sub_coe_le_of_le le_rfl h' - . rename_i h' - rw [h, h'] - simp at * - -def NoisedBoundedSumQuery (U : ℕ+) (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℤ := do - NoisedQuery (BoundedSumQuery U) U ε₁ ε₂ l - -theorem NoisedBoundedSumQueryDP (U : ℕ+) (ε₁ ε₂ : ℕ+) : DP (NoisedBoundedSumQuery U ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by - apply NoisedQueryDP - apply BoundedSumQuerySensitivity - -theorem NoisedBoundedSumQuery_NonZeroNQ (U : ℕ+) (ε₁ ε₂ : ℕ+) : - @NonZeroNQ ℕ ℤ (NoisedBoundedSumQuery U ε₁ ε₂) := by - apply NoisedQuery_NonZeroNQ - -theorem NoisedBoundedSumQuery_NonTopNQ (U : ℕ+) (ε₁ ε₂ : ℕ+) : - @NonTopNQ ℕ ℤ (NoisedBoundedSumQuery U ε₁ ε₂) := by - apply NoisedQuery_NonTopNQ - -theorem NoisedBoundedSumQuery_NonTopRDNQ (U : ℕ+) (ε₁ ε₂ : ℕ+) : - @NonTopRDNQ ℕ ℤ (NoisedBoundedSumQuery U ε₁ ε₂) := by - apply NoisedQuery_NonTopRDNQ - -theorem NoisedBoundedSumQuery_NonTopSum (U : ℕ+) (ε₁ ε₂ : ℕ+) : - @NonTopSum ℕ ℤ (NoisedBoundedSumQuery U ε₁ ε₂) := by - apply NoisedQuery_NonTopSum - -end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/Code.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/Code.lean index 6668165c..c159f7b6 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/Code.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/Code.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jean-Baptiste Tristan -/ -import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Mechanism.Mechanism +import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Mechanism.Code import Mathlib.Algebra.Group.Defs import Mathlib.Init.Algebra.Classes import Init.Data.Int.Order @@ -18,4 +18,7 @@ namespace SLang def BoundedSumQuery (U : ℕ+) (l : List ℕ) : ℤ := List.sum (List.map (fun n : ℕ => (Nat.min n U)) l) +def NoisedBoundedSumQuery (U : ℕ+) (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℤ := do + NoisedQuery (BoundedSumQuery U) U ε₁ ε₂ l + end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/Properties.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/Properties.lean index cdf85f74..84f22663 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/Properties.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/Properties.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jean-Baptiste Tristan -/ -import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Mechanism.Mechanism +import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Mechanism.Basic import Mathlib.Algebra.Group.Defs import Mathlib.Init.Algebra.Classes import Init.Data.Int.Order @@ -70,9 +70,6 @@ theorem BoundedSumQuerySensitivity (U : ℕ+) : sensitivity (BoundedSumQuery U) rw [h, h'] simp at * -def NoisedBoundedSumQuery (U : ℕ+) (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℤ := do - NoisedQuery (BoundedSumQuery U) U ε₁ ε₂ l - theorem NoisedBoundedSumQueryDP (U : ℕ+) (ε₁ ε₂ : ℕ+) : DP (NoisedBoundedSumQuery U ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by apply NoisedQueryDP apply BoundedSumQuerySensitivity diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Code.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Code.lean index 1b3c83fa..6841d98f 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Code.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Code.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jean-Baptiste Tristan -/ -import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Mechanism.Mechanism +import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Mechanism.Code open Classical Nat Int Real @@ -16,4 +16,7 @@ variable {T : Type} def CountingQuery (l : List T) : ℤ := List.length l +def NoisedCountingQuery (ε₁ ε₂ : ℕ+) (l : List T) : SLang ℤ := do + NoisedQuery CountingQuery 1 ε₁ ε₂ l + end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Count.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Count.lean deleted file mode 100644 index 31846c1e..00000000 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Count.lean +++ /dev/null @@ -1,58 +0,0 @@ -/- -Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Jean-Baptiste Tristan --/ - -import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Mechanism.Mechanism - -open Classical Nat Int Real - -noncomputable section - -namespace SLang - -variable {T : Type} - -def CountingQuery (l : List T) : ℤ := List.length l - -theorem CountingQuery1Sensitive : - @sensitivity T CountingQuery 1 := by - simp [CountingQuery, sensitivity] - intros l₁ l₂ H - cases H - . rename_i a b n h1 h2 - subst h1 h2 - simp - . rename_i a b n h1 h2 - subst h1 h2 - simp - . rename_i a n b m h1 h2 - subst h1 h2 - simp - -def NoisedCountingQuery (ε₁ ε₂ : ℕ+) (l : List T) : SLang ℤ := do - NoisedQuery CountingQuery 1 ε₁ ε₂ l - -theorem NoisedCountingQueryDP (ε₁ ε₂ : ℕ+) : - @DP T ℤ (NoisedCountingQuery ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by - apply NoisedQueryDP - apply CountingQuery1Sensitive - -theorem NoisedCountingQuery_NonZeroNQ (ε₁ ε₂ : ℕ+) : - @NonZeroNQ T ℤ (NoisedCountingQuery ε₁ ε₂) := by - apply NoisedQuery_NonZeroNQ - -theorem NoisedCountingQuery_NonTopNQ (ε₁ ε₂ : ℕ+) : - @NonTopNQ T ℤ (NoisedCountingQuery ε₁ ε₂) := by - apply NoisedQuery_NonTopNQ - -theorem NoisedCountingQuery_NonTopRDNQ (ε₁ ε₂ : ℕ+) : - @NonTopRDNQ T ℤ (NoisedCountingQuery ε₁ ε₂) := by - apply NoisedQuery_NonTopRDNQ - -theorem NoisedCountingQuery_NonTopSum (ε₁ ε₂ : ℕ+) : - @NonTopSum T ℤ (NoisedCountingQuery ε₁ ε₂) := by - apply NoisedQuery_NonTopSum - -end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Properties.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Properties.lean index 80f7875e..6b8b4f1c 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Properties.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Properties.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jean-Baptiste Tristan -/ -import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Mechanism.Mechanism +import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Mechanism.Basic import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.Count.Code open Classical Nat Int Real @@ -30,9 +30,6 @@ theorem CountingQuery1Sensitive : subst h1 h2 simp -def NoisedCountingQuery (ε₁ ε₂ : ℕ+) (l : List T) : SLang ℤ := do - NoisedQuery CountingQuery 1 ε₁ ε₂ l - theorem NoisedCountingQueryDP (ε₁ ε₂ : ℕ+) : @DP T ℤ (NoisedCountingQuery ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by apply NoisedQueryDP From b5f2c29600b03c0265a0559e2f2ed02aa982fa90 Mon Sep 17 00:00:00 2001 From: jtristan Date: Wed, 15 May 2024 16:22:26 -0400 Subject: [PATCH 19/20] Moved Gaussian to util --- .../ZeroConcentrated/ConcentratedBound.lean | 32 +++++++++++++++---- .../ZeroConcentrated/DP.lean | 4 --- .../Foundations/Composition/Code.lean | 4 +-- .../Foundations/Mechanism/Code.lean | 4 +-- .../Foundations/Postprocessing/Code.lean | 4 +-- .../Queries/BoundedMean/Code.lean | 2 -- .../Queries/BoundedSum/Code.lean | 2 -- .../ZeroConcentrated/Queries/Count/Code.lean | 2 -- SampCert/Samplers/Gaussian/Properties.lean | 2 +- SampCert/Samplers/GaussianGen/Properties.lean | 20 ++---------- .../Gaussian}/DiscreteGaussian.lean | 0 .../Gaussian}/GaussBound.lean | 2 +- .../Gaussian}/GaussConvergence.lean | 2 +- .../Gaussian}/GaussPeriodicity.lean | 2 +- 14 files changed, 35 insertions(+), 47 deletions(-) rename SampCert/{DifferentialPrivacy/ZeroConcentrated => Util/Gaussian}/DiscreteGaussian.lean (100%) rename SampCert/{DifferentialPrivacy/ZeroConcentrated => Util/Gaussian}/GaussBound.lean (98%) rename SampCert/{DifferentialPrivacy/ZeroConcentrated => Util/Gaussian}/GaussConvergence.lean (92%) rename SampCert/{DifferentialPrivacy/ZeroConcentrated => Util/Gaussian}/GaussPeriodicity.lean (98%) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean index 22636fb8..f810bf2f 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean @@ -4,14 +4,15 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jean-Baptiste Tristan -/ -import SampCert.DifferentialPrivacy.ZeroConcentrated.DiscreteGaussian -import SampCert.DifferentialPrivacy.ZeroConcentrated.GaussBound -import SampCert.DifferentialPrivacy.ZeroConcentrated.GaussConvergence -import SampCert.DifferentialPrivacy.ZeroConcentrated.GaussPeriodicity +import SampCert.Util.Gaussian.DiscreteGaussian +import SampCert.Util.Gaussian.GaussBound +import SampCert.Util.Gaussian.GaussConvergence +import SampCert.Util.Gaussian.GaussPeriodicity import SampCert.Util.Shift import SampCert.DifferentialPrivacy.RenyiDivergence +import SampCert.Samplers.GaussianGen.Basic -open Real +open Real Nat theorem sg_sum_pos' {σ : ℝ} (h : σ ≠ 0) (μ : ℝ) (α : ℝ) : 0 < ((gauss_term_ℝ σ μ) x / ∑' (x : ℤ), (gauss_term_ℝ σ μ) x)^α := by @@ -183,7 +184,7 @@ theorem RenyiDivergenceBound {σ : ℝ} (h : σ ≠ 0) (μ : ℤ) (α : ℝ) (h' right intro x rw [E] - rw [add_div] + rw [_root_.add_div] rw [exp_add] rw [tsum_mul_right] rw [mul_comm] @@ -550,3 +551,22 @@ theorem RenyiDivergenceBound' {σ α : ℝ} (h : σ ≠ 0) (h' : 1 < α) (μ ν . apply SummableRenyiGauss h rw [A] apply RenyiDivergenceBound_pre h h' + +namespace SLang + +theorem DiscreteGaussianGenSampleZeroConcentrated {α : ℝ} (h : 1 < α) (num : PNat) (den : PNat) (μ ν : ℤ) : + RenyiDivergence ((DiscreteGaussianGenSample num den μ)) (DiscreteGaussianGenSample num den ν) α ≤ + α * (((μ - ν) : ℤ)^2 / (2 * ((num : ℝ) / (den : ℝ))^2)) := by + have A : (num : ℝ) / (den : ℝ) ≠ 0 := by + simp only [ne_eq, div_eq_zero_iff, cast_eq_zero, PNat.ne_zero, or_self, not_false_eq_true] + conv => + left + congr + . intro x + rw [DiscreteGaussianGenSample_apply] + . intro x + rw [DiscreteGaussianGenSample_apply] + . skip + apply RenyiDivergenceBound' A h + +end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean index b949e03e..6d947ac6 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean @@ -18,10 +18,6 @@ import Mathlib.Probability.ProbabilityMassFunction.Integrals import Mathlib.Analysis.Convex.SpecificFunctions.Basic import Mathlib.Analysis.Convex.Integral -noncomputable section - -open Classical Nat Int Real ENNReal MeasureTheory Measure - def DP (q : List T → SLang U) (ε : ℝ) : Prop := ∀ α : ℝ, 1 < α → ∀ l₁ l₂ : List T, Neighbour l₁ l₂ → RenyiDivergence (q l₁) (q l₂) α ≤ (1/2) * ε ^ 2 * α diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Composition/Code.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Composition/Code.lean index d6d195e8..f8f95cb5 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Composition/Code.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Composition/Code.lean @@ -4,12 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jean-Baptiste Tristan -/ -import SampCert.DifferentialPrivacy.ZeroConcentrated.DP +import SampCert.SLang noncomputable section -open Classical Nat Int Real ENNReal MeasureTheory Measure - namespace SLang def Compose (nq1 nq2 : List T → SLang ℤ) (l : List T) : SLang (ℤ × ℤ) := do diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Mechanism/Code.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Mechanism/Code.lean index 71f57091..3825c572 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Mechanism/Code.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Mechanism/Code.lean @@ -4,12 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jean-Baptiste Tristan -/ -import SampCert.DifferentialPrivacy.ZeroConcentrated.DP +import SampCert.Samplers.GaussianGen.Code noncomputable section -open Classical Nat Int Real ENNReal MeasureTheory Measure - namespace SLang def NoisedQuery (query : List T → ℤ) (Δ : ℕ+) (ε₁ ε₂ : ℕ+) (l : List T) : SLang ℤ := do diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Code.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Code.lean index e0b6a899..61ba59dc 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Code.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Code.lean @@ -4,12 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jean-Baptiste Tristan -/ -import SampCert.DifferentialPrivacy.ZeroConcentrated.DP +import SampCert.SLang noncomputable section -open Classical Nat Int Real ENNReal MeasureTheory Measure - namespace SLang def PostProcess (nq : List T → SLang U) (pp : U → ℤ) (l : List T) : SLang ℤ := do diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/Code.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/Code.lean index 76d93c2f..98555e3e 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/Code.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedMean/Code.lean @@ -9,8 +9,6 @@ import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Postprocessing. import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.Count.Code import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.BoundedSum.Code -open Classical Nat Int Real - noncomputable section namespace SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/Code.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/Code.lean index c159f7b6..866b9631 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/Code.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/BoundedSum/Code.lean @@ -9,8 +9,6 @@ import Mathlib.Algebra.Group.Defs import Mathlib.Init.Algebra.Classes import Init.Data.Int.Order -open Classical Nat Int Real - noncomputable section namespace SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Code.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Code.lean index 6841d98f..74a91310 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Code.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Queries/Count/Code.lean @@ -6,8 +6,6 @@ Authors: Jean-Baptiste Tristan import SampCert.DifferentialPrivacy.ZeroConcentrated.Foundations.Mechanism.Code -open Classical Nat Int Real - noncomputable section namespace SLang diff --git a/SampCert/Samplers/Gaussian/Properties.lean b/SampCert/Samplers/Gaussian/Properties.lean index 43def706..9abd2ffc 100644 --- a/SampCert/Samplers/Gaussian/Properties.lean +++ b/SampCert/Samplers/Gaussian/Properties.lean @@ -12,7 +12,7 @@ import SampCert.Samplers.Laplace.Basic import Mathlib.NumberTheory.ModularForms.JacobiTheta.OneVariable import SampCert.Util.UtilMathlib import SampCert.Samplers.Gaussian.Code -import SampCert.DifferentialPrivacy.ZeroConcentrated.DiscreteGaussian +import SampCert.Util.Gaussian.DiscreteGaussian noncomputable section diff --git a/SampCert/Samplers/GaussianGen/Properties.lean b/SampCert/Samplers/GaussianGen/Properties.lean index 75c8d6b3..c14fb036 100644 --- a/SampCert/Samplers/GaussianGen/Properties.lean +++ b/SampCert/Samplers/GaussianGen/Properties.lean @@ -6,9 +6,8 @@ Authors: Jean-Baptiste Tristan import SampCert.Samplers.GaussianGen.Code import SampCert.Samplers.Gaussian.Properties -import SampCert.DifferentialPrivacy.ZeroConcentrated.DiscreteGaussian -import SampCert.DifferentialPrivacy.ZeroConcentrated.GaussPeriodicity -import SampCert.DifferentialPrivacy.ZeroConcentrated.ConcentratedBound +import SampCert.Util.Gaussian.DiscreteGaussian +import SampCert.Util.Gaussian.GaussPeriodicity noncomputable section @@ -47,19 +46,4 @@ theorem DiscreteGaussianGenSample_apply (num : PNat) (den : PNat) (μ x : ℤ) : . simp [gauss_term_ℝ] . rw [SG_periodic' A] -theorem DiscreteGaussianGenSampleZeroConcentrated {α : ℝ} (h : 1 < α) (num : PNat) (den : PNat) (μ ν : ℤ) : - RenyiDivergence ((DiscreteGaussianGenSample num den μ)) (DiscreteGaussianGenSample num den ν) α ≤ - α * (((μ - ν) : ℤ)^2 / (2 * ((num : ℝ) / (den : ℝ))^2)) := by - have A : (num : ℝ) / (den : ℝ) ≠ 0 := by - simp only [ne_eq, div_eq_zero_iff, cast_eq_zero, PNat.ne_zero, or_self, not_false_eq_true] - conv => - left - congr - . intro x - rw [DiscreteGaussianGenSample_apply] - . intro x - rw [DiscreteGaussianGenSample_apply] - . skip - apply RenyiDivergenceBound' A h - end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/DiscreteGaussian.lean b/SampCert/Util/Gaussian/DiscreteGaussian.lean similarity index 100% rename from SampCert/DifferentialPrivacy/ZeroConcentrated/DiscreteGaussian.lean rename to SampCert/Util/Gaussian/DiscreteGaussian.lean diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/GaussBound.lean b/SampCert/Util/Gaussian/GaussBound.lean similarity index 98% rename from SampCert/DifferentialPrivacy/ZeroConcentrated/GaussBound.lean rename to SampCert/Util/Gaussian/GaussBound.lean index 40d51de2..85e4f02e 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/GaussBound.lean +++ b/SampCert/Util/Gaussian/GaussBound.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jean-Baptiste Tristan -/ -import SampCert.DifferentialPrivacy.ZeroConcentrated.DiscreteGaussian +import SampCert.Util.Gaussian.DiscreteGaussian noncomputable section diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/GaussConvergence.lean b/SampCert/Util/Gaussian/GaussConvergence.lean similarity index 92% rename from SampCert/DifferentialPrivacy/ZeroConcentrated/GaussConvergence.lean rename to SampCert/Util/Gaussian/GaussConvergence.lean index 69f23edf..8dffa574 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/GaussConvergence.lean +++ b/SampCert/Util/Gaussian/GaussConvergence.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jean-Baptiste Tristan -/ -import SampCert.DifferentialPrivacy.ZeroConcentrated.DiscreteGaussian +import SampCert.Util.Gaussian.DiscreteGaussian open Classical Nat Real diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/GaussPeriodicity.lean b/SampCert/Util/Gaussian/GaussPeriodicity.lean similarity index 98% rename from SampCert/DifferentialPrivacy/ZeroConcentrated/GaussPeriodicity.lean rename to SampCert/Util/Gaussian/GaussPeriodicity.lean index 9f1256cb..379b2d5f 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/GaussPeriodicity.lean +++ b/SampCert/Util/Gaussian/GaussPeriodicity.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jean-Baptiste Tristan -/ -import SampCert.DifferentialPrivacy.ZeroConcentrated.GaussConvergence +import SampCert.Util.Gaussian.GaussConvergence theorem SGShift (μ σ : ℝ) (n : ℤ) (k : ℤ) : (gauss_term_ℝ σ μ) (((n + k) : ℤ) : ℝ) = (gauss_term_ℝ σ (μ - k)) n := by From 079b482dfc3d1df60ad68778312159c3e4230ca8 Mon Sep 17 00:00:00 2001 From: jtristan Date: Thu, 16 May 2024 09:01:05 -0400 Subject: [PATCH 20/20] refactoring postprocessing --- SampCert.lean | 2 +- .../Postprocessing/Properties.lean | 54 ++++++++----------- Tests/testing-kolmogorov-discretegaussian.py | 2 +- 3 files changed, 25 insertions(+), 33 deletions(-) diff --git a/SampCert.lean b/SampCert.lean index f9cf72c1..93901c30 100644 --- a/SampCert.lean +++ b/SampCert.lean @@ -4,4 +4,4 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jean-Baptiste Tristan -/ -import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.BoundedMean.Code +import SampCert.DifferentialPrivacy.ZeroConcentrated.Queries.BoundedMean.Basic diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Properties.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Properties.lean index 5b2f0b47..28ffd441 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Properties.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Foundations/Postprocessing/Properties.lean @@ -13,7 +13,17 @@ open Classical Nat Int Real ENNReal MeasureTheory Measure namespace SLang -theorem foo (f : U → ℤ) (g : U → ENNReal) (x : ℤ) : +variable {T : Type} +variable [t1 : MeasurableSpace T] +variable [t2 : MeasurableSingletonClass T] + +variable {U : Type} +variable [m2 : MeasurableSpace U] +variable [count : Countable U] +variable [disc : DiscreteMeasurableSpace U] +variable [Inhabited U] + +theorem condition_to_subset (f : U → ℤ) (g : U → ENNReal) (x : ℤ) : (∑' a : U, if x = f a then g a else 0) = ∑' a : { a | x = f a }, g a := by have A := @tsum_split_ite U (fun a : U => x = f a) g (fun _ => 0) simp only [decide_eq_true_eq, tsum_zero, add_zero] at A @@ -22,14 +32,9 @@ theorem foo (f : U → ℤ) (g : U → ENNReal) (x : ℤ) : simp rw [B] -variable {T : Type} -variable [m1 : MeasurableSpace T] -variable [m2 : MeasurableSingletonClass T] -variable [m3: MeasureSpace T] - theorem Integrable_rpow (f : T → ℝ) (nn : ∀ x : T, 0 ≤ f x) (μ : Measure T) (α : ENNReal) (mem : Memℒp f α μ) (h1 : α ≠ 0) (h2 : α ≠ ⊤) : MeasureTheory.Integrable (fun x : T => (f x) ^ α.toReal) μ := by - have X := @MeasureTheory.Memℒp.integrable_norm_rpow T ℝ m1 μ _ f α mem h1 h2 + have X := @MeasureTheory.Memℒp.integrable_norm_rpow T ℝ t1 μ _ f α mem h1 h2 revert X conv => left @@ -51,7 +56,7 @@ theorem Integrable_rpow (f : T → ℝ) (nn : ∀ x : T, 0 ≤ f x) (μ : Measur . rw [← hasFiniteIntegral_norm_iff] simp [X] -theorem bar (f : T → ℝ) (q : PMF T) (α : ℝ) (h : 1 < α) (h2 : ∀ x : T, 0 ≤ f x) (mem : Memℒp f (ENNReal.ofReal α) (PMF.toMeasure q)) : +theorem Renyi_Jensen (f : T → ℝ) (q : PMF T) (α : ℝ) (h : 1 < α) (h2 : ∀ x : T, 0 ≤ f x) (mem : Memℒp f (ENNReal.ofReal α) (PMF.toMeasure q)) : ((∑' x : T, (f x) * (q x).toReal)) ^ α ≤ (∑' x : T, (f x) ^ α * (q x).toReal) := by conv => @@ -91,7 +96,7 @@ theorem bar (f : T → ℝ) (q : PMF T) (α : ℝ) (h : 1 < α) (h2 : ∀ x : T, simp at h'' have C : @IsClosed ℝ UniformSpace.toTopologicalSpace (Set.Ici 0) := by exact isClosed_Ici - have D := @ConvexOn.map_integral_le T ℝ m1 _ _ _ (PMF.toMeasure q) (Set.Ici 0) f (fun (x : ℝ) => x ^ α) (PMF.toMeasure.isProbabilityMeasure q) A B C + have D := @ConvexOn.map_integral_le T ℝ t1 _ _ _ (PMF.toMeasure q) (Set.Ici 0) f (fun (x : ℝ) => x ^ α) (PMF.toMeasure.isProbabilityMeasure q) A B C simp at D apply D . exact MeasureTheory.ae_of_all (PMF.toMeasure q) h2 @@ -104,7 +109,7 @@ theorem bar (f : T → ℝ) (q : PMF T) (α : ℝ) (h : 1 < α) (h2 : ∀ x : T, apply lt_trans zero_lt_one h have Y : ENNReal.ofReal α ≠ ⊤ := by simp - have Z := @Integrable_rpow T m1 f h2 (PMF.toMeasure q) (ENNReal.ofReal α) mem X Y + have Z := @Integrable_rpow T t1 f h2 (PMF.toMeasure q) (ENNReal.ofReal α) mem X Y rw [toReal_ofReal] at Z . exact Z . apply le_of_lt @@ -114,7 +119,7 @@ theorem bar (f : T → ℝ) (q : PMF T) (α : ℝ) (h : 1 < α) (h2 : ∀ x : T, apply lt_trans zero_lt_one h have Y : ENNReal.ofReal α ≠ ⊤ := by simp - have Z := @Integrable_rpow T m1 f h2 (PMF.toMeasure q) (ENNReal.ofReal α) mem X Y + have Z := @Integrable_rpow T t1 f h2 (PMF.toMeasure q) (ENNReal.ofReal α) mem X Y rw [toReal_ofReal] at Z . exact Z . apply le_of_lt @@ -123,11 +128,6 @@ theorem bar (f : T → ℝ) (q : PMF T) (α : ℝ) (h : 1 < α) (h2 : ∀ x : T, rw [one_le_ofReal] apply le_of_lt h -variable {U : Type} -variable [m2 : MeasurableSpace U] -- [m2' : MeasurableSingletonClass U] -variable [count : Countable U] -variable [disc : DiscreteMeasurableSpace U] - def δ (nq : SLang U) (f : U → ℤ) (a : ℤ) : {n : U | a = f n} → ENNReal := fun x : {n : U | a = f n} => nq x * (∑' (x : {n | a = f n}), nq x)⁻¹ theorem δ_normalizes (nq : SLang U) (f : U → ℤ) (a : ℤ) (h1 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ 0) (h2 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ ⊤) : @@ -193,7 +193,7 @@ theorem ENNReal.tsum_fiberwise (p : T → ENNReal) (f : T → ℤ) : apply Summable.hasSum exact ENNReal.summable -theorem quux (p : T → ENNReal) (f : T → ℤ) : +theorem fiberwisation (p : T → ENNReal) (f : T → ℤ) : (∑' i : T, p i) = ∑' (x : ℤ), if {a : T | x = f a} = {} then 0 else ∑'(i : {a : T | x = f a}), p i := by rw [← ENNReal.tsum_fiberwise p f] @@ -220,7 +220,7 @@ theorem quux (p : T → ENNReal) (f : T → ℤ) : theorem convergent_subset {p : T → ENNReal} (f : T → ℤ) (conv : ∑' (x : T), p x ≠ ⊤) : ∑' (x : { y : T| x = f y }), p x ≠ ⊤ := by - rw [← foo] + rw [← condition_to_subset] have A : (∑' (y : T), if x = f y then p y else 0) ≤ ∑' (x : T), p x := by apply tsum_le_tsum . intro i @@ -275,7 +275,7 @@ theorem DPostPocess_pre {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP n rw [@RenyiDivergenceExpectation _ (nq l₁) (nq l₂) _ h1 (nn l₂) (nts l₂)] -- Shuffle the sum - rw [quux (fun x => (nq l₁ x / nq l₂ x) ^ α * nq l₂ x) f] + rw [fiberwisation (fun x => (nq l₁ x / nq l₂ x) ^ α * nq l₂ x) f] apply ENNReal.tsum_le_tsum @@ -284,7 +284,7 @@ theorem DPostPocess_pre {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP n -- Get rid of elements with probability 0 in the pushforward split . rename_i empty - rw [foo] + rw [condition_to_subset] have ZE : (∑' (x_1 : ↑{n | i = f n}), nq l₁ ↑x_1) = 0 := by simp intro a H @@ -373,8 +373,8 @@ theorem DPostPocess_pre {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP n intro a apply ENNReal.ne_top_of_tsum_ne_top S3 - rw [foo] - rw [foo] + rw [condition_to_subset] + rw [condition_to_subset] -- Introduce Q(f⁻¹ i) let κ := ∑' x : {n : U | i = f n}, nq l₂ x @@ -478,7 +478,7 @@ theorem DPostPocess_pre {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP n apply S3 - have Jensen's := @bar {n : U | i = f n} Subtype.instMeasurableSpace Subtype.instMeasurableSingletonClass (fun a => (nq l₁ a / nq l₂ a).toReal) (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) α h1 P5 XXX + have Jensen's := @Renyi_Jensen {n : U | i = f n} Subtype.instMeasurableSpace Subtype.instMeasurableSingletonClass (fun a => (nq l₁ a / nq l₂ a).toReal) (δpmf (nq l₂) f i (MasterZero l₂) (MasterRW l₂)) α h1 P5 XXX clear P5 have P6 : 0 ≤ (∑' (x : ↑{n | i = f n}), nq l₂ ↑x).toReal := by @@ -556,11 +556,6 @@ theorem DPostPocess_pre {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP n simp only [one_mul] rw [ENNReal.tsum_mul_right] - have H1 : 0 ≤ ∑' (x : ↑{n | i = f n}), (nq l₁ ↑x).toReal := by - apply tsum_nonneg - simp - have H2 : 0 ≤ (∑' (a : ↑{a | i = f a}), nq l₂ ↑a)⁻¹.toReal := by - apply toReal_nonneg have H4 : (∑' (a : ↑{a | i = f a}), nq l₂ ↑a)⁻¹ ≠ ⊤ := by apply inv_ne_top.mpr simp @@ -600,8 +595,6 @@ theorem tsum_ne_zero_of_ne_zero {T : Type} [Inhabited T] (f : T → ENNReal) (h have B := CONTRA default contradiction -variable [Inhabited U] - theorem DPPostProcess {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP nq ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (conv : NonTopSum nq) (f : U → ℤ) : DP (PostProcess nq f) ((ε₁ : ℝ) / ε₂) := by simp [PostProcess, DP, RenyiDivergence] @@ -690,5 +683,4 @@ theorem DPPostProcess {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : DP nq rw [lt_top_iff_ne_top] at Z contradiction - end SLang diff --git a/Tests/testing-kolmogorov-discretegaussian.py b/Tests/testing-kolmogorov-discretegaussian.py index 34af11e6..20ddab1d 100644 --- a/Tests/testing-kolmogorov-discretegaussian.py +++ b/Tests/testing-kolmogorov-discretegaussian.py @@ -157,7 +157,7 @@ def test_kolmogorov_dist(N, sigma2, with_plots=False): print("* Calling the 'test_kolmogorov_dist' function with N=1000 and location parameter sigma^2=10 (without plots):") # How to use the "test_kolmogorov_dist" function: on N=10000 samples, with sigma^2 = 10 (no plots) diff = test_kolmogorov_dist(N,sig2) - if diff < 0.01: + if diff < 0.02: print("Test passed!") exit(0) else: