From 1f7f8cb03e8825370e8bcb46cc66880fbf22a854 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Mon, 3 Jun 2024 15:12:17 -0400 Subject: [PATCH 001/122] proof skeleton --- SampCert/DifferentialPrivacy/Abstract.lean | 12 ++- .../DifferentialPrivacy/Pure/Composition.lean | 87 +++++++++++++++++++ 2 files changed, 96 insertions(+), 3 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Abstract.lean b/SampCert/DifferentialPrivacy/Abstract.lean index cdf2720d..b5f43107 100644 --- a/SampCert/DifferentialPrivacy/Abstract.lean +++ b/SampCert/DifferentialPrivacy/Abstract.lean @@ -28,11 +28,17 @@ namespace SLang abbrev Query (T U : Type) := List T → U abbrev Mechanism (T U : Type) := List T → SLang U + /-- -Product of mechanisms. +General (value-dependent) composition of mechanisms +-/ +def privComposeAdaptive (nq1 : Mechanism T U) (nq2 : U -> Mechanism T V) (l : List T) : SLang V := do + let A <- nq1 l + let B <- nq2 A l + return B -Note that the second mechanism does not depend on the output of the first; this is in currently -in contrast to the notions of composition found in the DP literature. +/-- +Composition of independent mechanisms -/ def privCompose (nq1 : Mechanism T U) (nq2 : Mechanism T V) (l : List T) : SLang (U × V) := do let A ← nq1 l diff --git a/SampCert/DifferentialPrivacy/Pure/Composition.lean b/SampCert/DifferentialPrivacy/Pure/Composition.lean index 18504502..ae9847ce 100644 --- a/SampCert/DifferentialPrivacy/Pure/Composition.lean +++ b/SampCert/DifferentialPrivacy/Pure/Composition.lean @@ -12,6 +12,21 @@ noncomputable section open Classical Set + +-- FIXME: move +/-- +Mediant inequality +-/ +lemma tsum_mediant (f g : U -> ENNReal) : (∑' (u : U), f u) / (∑' (u : U), g u) ≤ sSup (range (fun (u : U) => f u / g u)) := + sorry + +lemma bounded_quotient (f g : U -> ENNReal) (b : ENNReal) (h_bound : ∀ (u : U), f u / g u ≤ b) : + (∑' (u : U), f u) / (∑' (u : U), g u) ≤ b := by + apply le_trans + · apply tsum_mediant + · simp -- wow, what a simplification + assumption + namespace SLang theorem PureDP_Compose' {nq1 : Mechanism T U} {nq2 : List T → SLang V} {ε₁ ε₂ ε₃ ε₄ : ℕ+} (h1 : PureDP nq1 ((ε₁ : ℝ) / ε₂)) (h2 : PureDP nq2 ((ε₃ : ℝ) / ε₄)) : @@ -19,12 +34,16 @@ theorem PureDP_Compose' {nq1 : Mechanism T U} {nq2 : List T → SLang V} {ε₁ simp [PureDP] at * rcases h1 with ⟨h1a, _⟩ rcases h2 with ⟨h2a, _⟩ + -- Suffices to bound at each point in the resulting pmf rw [event_eq_singleton] at * simp [DP_singleton] at * + -- For all neighbouring DB's and output points intros l₁ l₂ neighbours x y replace h1a := h1a l₁ l₂ neighbours x replace h2a := h2a l₁ l₂ neighbours y + -- Open the implementation simp [privCompose] + -- Rearrange nested sum to double sum conv => left congr @@ -68,4 +87,72 @@ theorem PureDP_Compose (nq1 : List T → SLang U) (nq2 : List T → SLang V) (ε . apply PureDP_Compose' hc h'c . apply privCompose_NonZeroNQ h2 h'2 + +-- set_option pp.notation false + +theorem PureDP_ComposeAdaptive' (nq1 : List T → SLang U) (nq2 : U -> List T → SLang V) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h1 : PureDP nq1 ((ε₁ : ℝ) / ε₂)) (h2 : ∀ u : U, PureDP (nq2 u) ((ε₃ : ℝ) / ε₄)) : + PureDP (privComposeAdaptive nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by + simp [PureDP] at * + rcases h1 with ⟨h1a, _⟩ + rw [event_eq_singleton] at * + simp [DP_singleton] at * + apply And.intro + · -- Composition satisfies DP bound + + -- For all neighbours... + intros l₁ l₂ neighbours x + replace h1a := h1a l₁ l₂ neighbours + + -- Simplify under the ∀ in h2 + have h2' : ∀ (u : U), (nq2 u l₁ x) / (nq2 u l₂ x) <= ENNReal.ofReal (((↑↑ε₃ : ℝ ) / ↑↑ε₄).exp) := by + intro u + replace h2 := h2 u + rw [event_eq_singleton] at h2 + simp [DP_singleton] at h2 + rcases h2 with ⟨h2a, _⟩ + apply h2a + apply neighbours + + simp [privComposeAdaptive] + + have h3 : ∀ (a : U), nq1 l₁ a * nq2 a l₁ x / (nq1 l₂ a * nq2 a l₂ x) ≤ ENNReal.ofReal (↑↑ε₁ / ↑↑ε₂ + ↑↑ε₃ / ↑↑ε₄ : ℝ).exp := by + intro a + have h1a' := h1a a + have h2a' := h2' a + + rw [Real.exp_add] + -- How to focus individual goals in lean? + rw [ENNReal.ofReal_mul]; all_goals (try apply Real.exp_nonneg) + + -- Some inequality pushing I beg lean to let me Search + sorry + + -- Put a name to the summands (why is this so hard) + let f := (fun (a : U) => nq1 l₁ a * nq2 a l₁ x) + let g := (fun (a : U) => nq1 l₂ a * nq2 a l₂ x) + have hf : (∑' (a : U), nq1 l₁ a * nq2 a l₁ x) = (∑' (a : U), f a) := sorry + have hg : (∑' (a : U), nq1 l₂ a * nq2 a l₂ x) = (∑' (a : U), g a) := sorry + rw [hf, hg] + + -- Mediant lemma + apply bounded_quotient + + -- Split the exponential by the sum + rw [Real.exp_add] + -- How to focus individual goals in lean? + rw [ENNReal.ofReal_mul]; all_goals (try apply Real.exp_nonneg) + + -- Apply the inequalities from above + intro u + have h1a' := h1a u + have h2a' := h2' u + simp [f, g] + + -- Push around inequalities + sorry + + · -- Composition is nonzero at all elements + sorry + + end SLang From f7020472d152ddc3a7104a624e615c0dbfd0b3fc Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Mon, 3 Jun 2024 17:04:13 -0400 Subject: [PATCH 002/122] eod checkpoint --- .../DifferentialPrivacy/Pure/Composition.lean | 62 ++++++++++--------- 1 file changed, 34 insertions(+), 28 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Pure/Composition.lean b/SampCert/DifferentialPrivacy/Pure/Composition.lean index ae9847ce..ab198c19 100644 --- a/SampCert/DifferentialPrivacy/Pure/Composition.lean +++ b/SampCert/DifferentialPrivacy/Pure/Composition.lean @@ -13,20 +13,36 @@ noncomputable section open Classical Set +-- set_option pp.notation false -- FIXME: move /-- Mediant inequality -/ -lemma tsum_mediant (f g : U -> ENNReal) : (∑' (u : U), f u) / (∑' (u : U), g u) ≤ sSup (range (fun (u : U) => f u / g u)) := + + +lemma tsum_mediant (f g : U -> ENNReal) (hg0 : ∀ u, g u ≠ 0) (hgT : ∀ u, g u ≠ ⊤): + (∑' (u : U), f u) / (∑' (u : U), g u) ≤ ⨆ u, f u / g u := by + let b := ⨆ u, f u / g u + have Hb : ∀ (u : U), f u ≤ b * g u := by + simp [b] + intro u + refine (ENNReal.div_le_iff_le_mul ?hb0 ?hbt).mp ?a + · left; apply hg0 + · left; apply hgT -- I want to delete this one if possible + · refine (le_iSup (fun u => HDiv.hDiv (f u) (g u)) u) + have Hs : (∑' (u : U), g u) ≤ b * (∑' (u : U), f u) := sorry -- pointwise bound, using Hb + -- by rearrange Hs sorry -lemma bounded_quotient (f g : U -> ENNReal) (b : ENNReal) (h_bound : ∀ (u : U), f u / g u ≤ b) : + +lemma bounded_quotient (f g : U -> ENNReal) (b : ENNReal) (h_bound : ∀ (u : U), f u / g u ≤ b) (hg0 : ∀ u, g u ≠ 0) : (∑' (u : U), f u) / (∑' (u : U), g u) ≤ b := by apply le_trans · apply tsum_mediant - · simp -- wow, what a simplification + · apply hg0 + · sorry -- Can get this as an extra hypothesis if necessary... g should not be ∞ anywhere (though it's a little annoying that this isn't intrinsic... maybe try to eliminate this downstream?) + · simp assumption - namespace SLang theorem PureDP_Compose' {nq1 : Mechanism T U} {nq2 : List T → SLang V} {ε₁ ε₂ ε₃ ε₄ : ℕ+} (h1 : PureDP nq1 ((ε₁ : ℝ) / ε₂)) (h2 : PureDP nq2 ((ε₃ : ℝ) / ε₄)) : @@ -87,13 +103,10 @@ theorem PureDP_Compose (nq1 : List T → SLang U) (nq2 : List T → SLang V) (ε . apply PureDP_Compose' hc h'c . apply privCompose_NonZeroNQ h2 h'2 - --- set_option pp.notation false - theorem PureDP_ComposeAdaptive' (nq1 : List T → SLang U) (nq2 : U -> List T → SLang V) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h1 : PureDP nq1 ((ε₁ : ℝ) / ε₂)) (h2 : ∀ u : U, PureDP (nq2 u) ((ε₃ : ℝ) / ε₄)) : PureDP (privComposeAdaptive nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by simp [PureDP] at * - rcases h1 with ⟨h1a, _⟩ + rcases h1 with ⟨h1a, h1nz⟩ rw [event_eq_singleton] at * simp [DP_singleton] at * apply And.intro @@ -121,37 +134,30 @@ theorem PureDP_ComposeAdaptive' (nq1 : List T → SLang U) (nq2 : U -> List T have h2a' := h2' a rw [Real.exp_add] - -- How to focus individual goals in lean? + -- How to focus individual goals in lean? This is a mess rw [ENNReal.ofReal_mul]; all_goals (try apply Real.exp_nonneg) - -- Some inequality pushing I beg lean to let me Search - sorry + -- What a mess + apply (@le_trans' _ _ _ ( nq1 l₁ a / nq1 l₂ a * ENNReal.ofReal (↑↑ε₃ / ↑↑ε₄ : ℝ).exp)) + · sorry + · apply (@le_trans' _ _ _ ( nq1 l₁ a / nq1 l₂ a * nq2 a l₁ x / nq2 a l₂ x)) + · sorry + · sorry -- Put a name to the summands (why is this so hard) let f := (fun (a : U) => nq1 l₁ a * nq2 a l₁ x) let g := (fun (a : U) => nq1 l₂ a * nq2 a l₂ x) - have hf : (∑' (a : U), nq1 l₁ a * nq2 a l₁ x) = (∑' (a : U), f a) := sorry - have hg : (∑' (a : U), nq1 l₂ a * nq2 a l₂ x) = (∑' (a : U), g a) := sorry + have hf : (∑' (a : U), nq1 l₁ a * nq2 a l₁ x) = (∑' (a : U), f a) := by congr + have hg : (∑' (a : U), nq1 l₂ a * nq2 a l₂ x) = (∑' (a : U), g a) := by congr rw [hf, hg] - -- Mediant lemma + -- Conclude by Mediant lemma apply bounded_quotient - - -- Split the exponential by the sum - rw [Real.exp_add] - -- How to focus individual goals in lean? - rw [ENNReal.ofReal_mul]; all_goals (try apply Real.exp_nonneg) - - -- Apply the inequalities from above - intro u - have h1a' := h1a u - have h2a' := h2' u - simp [f, g] - - -- Push around inequalities - sorry + apply h3 · -- Composition is nonzero at all elements + simp [NonZeroNQ] + intros l n sorry From 45ad46e9db399bb715149ed9c799e358e8fd5774 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Mon, 3 Jun 2024 17:07:09 -0400 Subject: [PATCH 003/122] nit --- SampCert/DifferentialPrivacy/Pure/Composition.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/SampCert/DifferentialPrivacy/Pure/Composition.lean b/SampCert/DifferentialPrivacy/Pure/Composition.lean index ab198c19..489561e4 100644 --- a/SampCert/DifferentialPrivacy/Pure/Composition.lean +++ b/SampCert/DifferentialPrivacy/Pure/Composition.lean @@ -155,6 +155,11 @@ theorem PureDP_ComposeAdaptive' (nq1 : List T → SLang U) (nq2 : U -> List T apply bounded_quotient apply h3 + -- Nonzero side condition + rw [NonZeroNQ] at * + sorry + + · -- Composition is nonzero at all elements simp [NonZeroNQ] intros l n From 6c536b700949b5d2c7c82b49d97df08e5be70c7f Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Tue, 4 Jun 2024 11:11:15 -0400 Subject: [PATCH 004/122] checkpoint compose proof before changing side conditions --- .../DifferentialPrivacy/Pure/Composition.lean | 91 +++++++++++-------- 1 file changed, 52 insertions(+), 39 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Pure/Composition.lean b/SampCert/DifferentialPrivacy/Pure/Composition.lean index 489561e4..162ac87d 100644 --- a/SampCert/DifferentialPrivacy/Pure/Composition.lean +++ b/SampCert/DifferentialPrivacy/Pure/Composition.lean @@ -13,36 +13,32 @@ noncomputable section open Classical Set --- set_option pp.notation false --- FIXME: move + +-- Solving the side conditions needs to be done separately depending on if u is inhabited or not /-- Mediant inequality -/ - - -lemma tsum_mediant (f g : U -> ENNReal) (hg0 : ∀ u, g u ≠ 0) (hgT : ∀ u, g u ≠ ⊤): +lemma tsum_mediant (f g : U -> ENNReal) (hg0 : ∀ u, g u ≠ 0) (hf0 : ∀ u, f u ≠ 0): (∑' (u : U), f u) / (∑' (u : U), g u) ≤ ⨆ u, f u / g u := by - let b := ⨆ u, f u / g u - have Hb : ∀ (u : U), f u ≤ b * g u := by - simp [b] + apply (ENNReal.div_le_iff_le_mul _ _).mpr + · rw [← ENNReal.tsum_mul_left] + apply ENNReal.tsum_le_tsum intro u - refine (ENNReal.div_le_iff_le_mul ?hb0 ?hbt).mp ?a - · left; apply hg0 - · left; apply hgT -- I want to delete this one if possible + apply (ENNReal.div_le_iff_le_mul _ _).mp · refine (le_iSup (fun u => HDiv.hDiv (f u) (g u)) u) - have Hs : (∑' (u : U), g u) ≤ b * (∑' (u : U), f u) := sorry -- pointwise bound, using Hb - -- by rearrange Hs - sorry - + · left; apply hg0 + · sorry + · sorry + · sorry -lemma bounded_quotient (f g : U -> ENNReal) (b : ENNReal) (h_bound : ∀ (u : U), f u / g u ≤ b) (hg0 : ∀ u, g u ≠ 0) : +lemma bounded_quotient (f g : U -> ENNReal) (b : ENNReal) (h_bound : ∀ (u : U), f u / g u ≤ b) (hg0 : ∀ u, g u ≠ 0) (hf0 : ∀ u, f u ≠ 0) : (∑' (u : U), f u) / (∑' (u : U), g u) ≤ b := by apply le_trans - · apply tsum_mediant - · apply hg0 - · sorry -- Can get this as an extra hypothesis if necessary... g should not be ∞ anywhere (though it's a little annoying that this isn't intrinsic... maybe try to eliminate this downstream?) + · refine (tsum_mediant _ _ hg0 hf0) · simp assumption + + namespace SLang theorem PureDP_Compose' {nq1 : Mechanism T U} {nq2 : List T → SLang V} {ε₁ ε₂ ε₃ ε₄ : ℕ+} (h1 : PureDP nq1 ((ε₁ : ℝ) / ε₂)) (h2 : PureDP nq2 ((ε₃ : ℝ) / ε₄)) : @@ -103,6 +99,7 @@ theorem PureDP_Compose (nq1 : List T → SLang U) (nq2 : List T → SLang V) (ε . apply PureDP_Compose' hc h'c . apply privCompose_NonZeroNQ h2 h'2 +-- set_option pp.notation false theorem PureDP_ComposeAdaptive' (nq1 : List T → SLang U) (nq2 : U -> List T → SLang V) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h1 : PureDP nq1 ((ε₁ : ℝ) / ε₂)) (h2 : ∀ u : U, PureDP (nq2 u) ((ε₃ : ℝ) / ε₄)) : PureDP (privComposeAdaptive nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by simp [PureDP] at * @@ -122,27 +119,39 @@ theorem PureDP_ComposeAdaptive' (nq1 : List T → SLang U) (nq2 : U -> List T replace h2 := h2 u rw [event_eq_singleton] at h2 simp [DP_singleton] at h2 - rcases h2 with ⟨h2a, _⟩ + rcases h2 with ⟨h2a, h2nz⟩ apply h2a apply neighbours simp [privComposeAdaptive] - have h3 : ∀ (a : U), nq1 l₁ a * nq2 a l₁ x / (nq1 l₂ a * nq2 a l₂ x) ≤ ENNReal.ofReal (↑↑ε₁ / ↑↑ε₂ + ↑↑ε₃ / ↑↑ε₄ : ℝ).exp := by + have h3 : ∀ (a : U), nq1 l₁ a * nq2 a l₁ x / (nq1 l₂ a * nq2 a l₂ x) ≤ ENNReal.ofReal (↑↑ε₁ / ↑↑ε₂ + ↑↑ε₃ / ↑↑ε₄ : ℝ).exp := by intro a - have h1a' := h1a a - have h2a' := h2' a - + -- Split the expontntial rw [Real.exp_add] - -- How to focus individual goals in lean? This is a mess - rw [ENNReal.ofReal_mul]; all_goals (try apply Real.exp_nonneg) - - -- What a mess - apply (@le_trans' _ _ _ ( nq1 l₁ a / nq1 l₂ a * ENNReal.ofReal (↑↑ε₃ / ↑↑ε₄ : ℝ).exp)) - · sorry - · apply (@le_trans' _ _ _ ( nq1 l₁ a / nq1 l₂ a * nq2 a l₁ x / nq2 a l₂ x)) - · sorry - · sorry + rw [ENNReal.ofReal_mul] <;> try apply Real.exp_nonneg -- How to focus individual goals in lean? This is a mess + + -- Push around inequalities + rw [ENNReal.div_eq_inv_mul] + rw [ENNReal.mul_inv] + · rw [<- mul_assoc] + rw [mul_right_comm] + conv => + lhs + arg 1 + rw [mul_assoc] + rw [mul_right_comm] + rw [← ENNReal.div_eq_inv_mul] + rw [← ENNReal.div_eq_inv_mul] + exact mul_le_mul' (h1a a) (h2' a) + · left + apply h1nz + · right + rcases (h2 a) with ⟨ _ , h2nz ⟩ + apply h2nz + + -- apply? -- gives me this but it times out. odd. + -- refine (bounded_quotient (fun u => nq1 l₁ u * nq2 u l₁ x) (fun u => nq1 l₂ u * nq2 u l₂ x) (ENNReal.ofReal (↑↑ε₁ / ↑↑ε₂ + ↑↑ε₃ / ↑↑ε₄ : ℝ).exp) h3 ?intro.left.hg0 ?intro.left.hf0) -- Put a name to the summands (why is this so hard) let f := (fun (a : U) => nq1 l₁ a * nq2 a l₁ x) @@ -154,15 +163,19 @@ theorem PureDP_ComposeAdaptive' (nq1 : List T → SLang U) (nq2 : U -> List T -- Conclude by Mediant lemma apply bounded_quotient apply h3 - - -- Nonzero side condition - rw [NonZeroNQ] at * - sorry - + all_goals (intro u; rcases (h2 u) with ⟨ _ , h2nz ⟩) + all_goals (simp only [f, g]) + · exact mul_ne_zero (h1nz l₂ u) (h2nz l₂ x) + · exact mul_ne_zero (h1nz l₁ u) (h2nz l₁ x) · -- Composition is nonzero at all elements - simp [NonZeroNQ] + simp only [NonZeroNQ] intros l n + + simp? only [privComposeAdaptive, bind, pure, bind_pure, bind_apply] + -- Interesting: This relies on U being nonempty again + + -- ne_eq, ENNReal.tsum_eq_zero] sorry From f9a515dafb4fac6ed52c3a792ccdc33d235ec6f0 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Tue, 4 Jun 2024 11:12:38 -0400 Subject: [PATCH 005/122] nit --- SampCert/DifferentialPrivacy/Pure/Composition.lean | 10 +--------- 1 file changed, 1 insertion(+), 9 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Pure/Composition.lean b/SampCert/DifferentialPrivacy/Pure/Composition.lean index 162ac87d..3fa40298 100644 --- a/SampCert/DifferentialPrivacy/Pure/Composition.lean +++ b/SampCert/DifferentialPrivacy/Pure/Composition.lean @@ -46,16 +46,12 @@ theorem PureDP_Compose' {nq1 : Mechanism T U} {nq2 : List T → SLang V} {ε₁ simp [PureDP] at * rcases h1 with ⟨h1a, _⟩ rcases h2 with ⟨h2a, _⟩ - -- Suffices to bound at each point in the resulting pmf rw [event_eq_singleton] at * simp [DP_singleton] at * - -- For all neighbouring DB's and output points intros l₁ l₂ neighbours x y replace h1a := h1a l₁ l₂ neighbours x replace h2a := h2a l₁ l₂ neighbours y - -- Open the implementation simp [privCompose] - -- Rearrange nested sum to double sum conv => left congr @@ -107,13 +103,9 @@ theorem PureDP_ComposeAdaptive' (nq1 : List T → SLang U) (nq2 : U -> List T rw [event_eq_singleton] at * simp [DP_singleton] at * apply And.intro - · -- Composition satisfies DP bound - - -- For all neighbours... - intros l₁ l₂ neighbours x + · intros l₁ l₂ neighbours x replace h1a := h1a l₁ l₂ neighbours - -- Simplify under the ∀ in h2 have h2' : ∀ (u : U), (nq2 u l₁ x) / (nq2 u l₂ x) <= ENNReal.ofReal (((↑↑ε₃ : ℝ ) / ↑↑ε₄).exp) := by intro u replace h2 := h2 u From 6262ba813adc0bb812dd79e9857431bc7cbca7db Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Tue, 4 Jun 2024 13:06:39 -0400 Subject: [PATCH 006/122] done up to strengthening --- .../DifferentialPrivacy/Pure/Composition.lean | 93 +++++++++++++++---- 1 file changed, 75 insertions(+), 18 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Pure/Composition.lean b/SampCert/DifferentialPrivacy/Pure/Composition.lean index 3fa40298..c5cb8411 100644 --- a/SampCert/DifferentialPrivacy/Pure/Composition.lean +++ b/SampCert/DifferentialPrivacy/Pure/Composition.lean @@ -7,6 +7,7 @@ import SampCert.DifferentialPrivacy.Abstract import SampCert.DifferentialPrivacy.Pure.DP import Mathlib.Data.Set.Defs import Mathlib.Data.Set.Prod +import Mathlib.Logic.IsEmpty noncomputable section @@ -20,16 +21,51 @@ Mediant inequality -/ lemma tsum_mediant (f g : U -> ENNReal) (hg0 : ∀ u, g u ≠ 0) (hf0 : ∀ u, f u ≠ 0): (∑' (u : U), f u) / (∑' (u : U), g u) ≤ ⨆ u, f u / g u := by - apply (ENNReal.div_le_iff_le_mul _ _).mpr - · rw [← ENNReal.tsum_mul_left] - apply ENNReal.tsum_le_tsum - intro u - apply (ENNReal.div_le_iff_le_mul _ _).mp - · refine (le_iSup (fun u => HDiv.hDiv (f u) (g u)) u) - · left; apply hg0 - · sorry - · sorry - · sorry + cases (Classical.em (U -> False)) + · rename_i Hu_empty + let (HU : IsEmpty U) := { false := Hu_empty } + rw [iSup_of_empty] + rw [tsum_empty] + rw [tsum_empty] + simp + · rename_i Hu + simp at Hu + rcases Hu with ⟨ u0 ⟩ + apply (ENNReal.div_le_iff_le_mul _ _).mpr + · rw [← ENNReal.tsum_mul_left] + apply ENNReal.tsum_le_tsum + intro u + apply (ENNReal.div_le_iff_le_mul _ _).mp + · refine (le_iSup (fun u => HDiv.hDiv (f u) (g u)) u) + · left; apply hg0 + · -- Can this side condition can follow from the nonempty property? + right + apply ne_of_gt + apply (LT.lt.trans_le ?g1 ?g2) + case g2 => + apply le_iSup + apply u + refine (ENNReal.div_pos (hf0 u) ?g1.hb) + -- It would seem that we need g u ≠ ⊤ in either left or right case + -- since ∞^-1 = 0 + sorry + · left + apply ne_of_gt + apply (LT.lt.trans_le ?z1 ?z2) + case z2 => + apply ENNReal.le_tsum + apply u0 + exact pos_iff_ne_zero.mpr (hg0 u0) + · right + apply ne_of_gt + apply (LT.lt.trans_le ?z3 ?z4) + case z4 => + apply le_iSup + apply u0 + refine (ENNReal.div_pos (hf0 u0) ?z6) + -- Either case for this side condition needs g to not be top too + sorry + lemma bounded_quotient (f g : U -> ENNReal) (b : ENNReal) (h_bound : ∀ (u : U), f u / g u ≤ b) (hg0 : ∀ u, g u ≠ 0) (hf0 : ∀ u, f u ≠ 0) : (∑' (u : U), f u) / (∑' (u : U), g u) ≤ b := by @@ -121,7 +157,8 @@ theorem PureDP_ComposeAdaptive' (nq1 : List T → SLang U) (nq2 : U -> List T intro a -- Split the expontntial rw [Real.exp_add] - rw [ENNReal.ofReal_mul] <;> try apply Real.exp_nonneg -- How to focus individual goals in lean? This is a mess + rw [ENNReal.ofReal_mul ?g0] + case g0 => apply Real.exp_nonneg -- Push around inequalities rw [ENNReal.div_eq_inv_mul] @@ -164,11 +201,31 @@ theorem PureDP_ComposeAdaptive' (nq1 : List T → SLang U) (nq2 : U -> List T simp only [NonZeroNQ] intros l n - simp? only [privComposeAdaptive, bind, pure, bind_pure, bind_apply] - -- Interesting: This relies on U being nonempty again - - -- ne_eq, ENNReal.tsum_eq_zero] - sorry - - + simp only [privComposeAdaptive, bind, pure, bind_pure, bind_apply] + + cases (Classical.em (U -> False)) + · rename_i Hu_empty + -- let (HU : IsEmpty U) := { false := Hu_empty } + -- rw [tsum_empty] + -- simp + -- Conclusion is false. + -- Can I get a contradiction out of either NonZeroNQ or do I have to strengthen + -- U to be nonempty? + exfalso + simp [NonZeroNQ] at h1nz + sorry + + · -- Nonempty: series is bounded below by value at u0 (which is positive) + rename_i Hu + simp at Hu + rcases Hu with ⟨ u0 ⟩ + rcases h2 u0 with ⟨ _ , h2nz ⟩ + apply ne_of_gt + apply (LT.lt.trans_le ?g1 ?g2) + case g2 => + apply ENNReal.le_tsum + apply u0 + apply ENNReal.mul_pos + · apply h1nz + · apply h2nz end SLang From 1b7cd138aa0061cb0bce6f52cd378c945bb060a9 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Tue, 4 Jun 2024 13:47:10 -0400 Subject: [PATCH 007/122] eliminate g top issue with classical --- .../DifferentialPrivacy/Pure/Composition.lean | 117 +++++++++--------- 1 file changed, 60 insertions(+), 57 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Pure/Composition.lean b/SampCert/DifferentialPrivacy/Pure/Composition.lean index c5cb8411..722464de 100644 --- a/SampCert/DifferentialPrivacy/Pure/Composition.lean +++ b/SampCert/DifferentialPrivacy/Pure/Composition.lean @@ -21,50 +21,58 @@ Mediant inequality -/ lemma tsum_mediant (f g : U -> ENNReal) (hg0 : ∀ u, g u ≠ 0) (hf0 : ∀ u, f u ≠ 0): (∑' (u : U), f u) / (∑' (u : U), g u) ≤ ⨆ u, f u / g u := by - cases (Classical.em (U -> False)) - · rename_i Hu_empty - let (HU : IsEmpty U) := { false := Hu_empty } - rw [iSup_of_empty] - rw [tsum_empty] - rw [tsum_empty] + + -- We need this to discharge side conditions in this proof, but we can get away with + -- some classical reasoning instead + cases (Classical.em (∀ u , g u ≠ ⊤)) + case inr => + rename_i Hk + -- rcases Hk with ⟨ ucont, Hcont ⟩ + have hgtop : ∃ u, g u = ⊤ := by apply (Decidable.not_forall_not.mp Hk) + have sumtop : ∑' (u : U) , g u = ⊤ := by exact ENNReal.tsum_eq_top_of_eq_top hgtop + rw [sumtop] simp - · rename_i Hu - simp at Hu - rcases Hu with ⟨ u0 ⟩ - apply (ENNReal.div_le_iff_le_mul _ _).mpr - · rw [← ENNReal.tsum_mul_left] - apply ENNReal.tsum_le_tsum - intro u - apply (ENNReal.div_le_iff_le_mul _ _).mp - · refine (le_iSup (fun u => HDiv.hDiv (f u) (g u)) u) - · left; apply hg0 - · -- Can this side condition can follow from the nonempty property? - right + case inl => + rename_i assumption_g + cases (isEmpty_or_nonempty U) + · rename_i HU + rw [iSup_of_empty] + rw [tsum_empty] + rw [tsum_empty] + simp + · rename_i Hu + simp at Hu + rcases Hu with ⟨ u0 ⟩ + apply (ENNReal.div_le_iff_le_mul _ _).mpr + · rw [← ENNReal.tsum_mul_left] + apply ENNReal.tsum_le_tsum + intro u + apply (ENNReal.div_le_iff_le_mul _ _).mp + · refine (le_iSup (fun u => HDiv.hDiv (f u) (g u)) u) + · left; apply hg0 + · right + apply ne_of_gt + apply (LT.lt.trans_le ?g1 ?g2) + case g2 => + apply le_iSup + apply u + refine (ENNReal.div_pos (hf0 u) ?g1.hb) + apply assumption_g + · left + apply ne_of_gt + apply (LT.lt.trans_le ?z1 ?z2) + case z2 => + apply ENNReal.le_tsum + apply u0 + exact pos_iff_ne_zero.mpr (hg0 u0) + · right apply ne_of_gt - apply (LT.lt.trans_le ?g1 ?g2) - case g2 => + apply (LT.lt.trans_le ?z3 ?z4) + case z4 => apply le_iSup - apply u - refine (ENNReal.div_pos (hf0 u) ?g1.hb) - -- It would seem that we need g u ≠ ⊤ in either left or right case - -- since ∞^-1 = 0 - sorry - · left - apply ne_of_gt - apply (LT.lt.trans_le ?z1 ?z2) - case z2 => - apply ENNReal.le_tsum - apply u0 - exact pos_iff_ne_zero.mpr (hg0 u0) - · right - apply ne_of_gt - apply (LT.lt.trans_le ?z3 ?z4) - case z4 => - apply le_iSup - apply u0 - refine (ENNReal.div_pos (hf0 u0) ?z6) - -- Either case for this side condition needs g to not be top too - sorry + apply u0 + refine (ENNReal.div_pos (hf0 u0) ?z6) + apply assumption_g lemma bounded_quotient (f g : U -> ENNReal) (b : ENNReal) (h_bound : ∀ (u : U), f u / g u ≤ b) (hg0 : ∀ u, g u ≠ 0) (hf0 : ∀ u, f u ≠ 0) : @@ -131,8 +139,7 @@ theorem PureDP_Compose (nq1 : List T → SLang U) (nq2 : List T → SLang V) (ε . apply PureDP_Compose' hc h'c . apply privCompose_NonZeroNQ h2 h'2 --- set_option pp.notation false -theorem PureDP_ComposeAdaptive' (nq1 : List T → SLang U) (nq2 : U -> List T → SLang V) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h1 : PureDP nq1 ((ε₁ : ℝ) / ε₂)) (h2 : ∀ u : U, PureDP (nq2 u) ((ε₃ : ℝ) / ε₄)) : +theorem PureDP_ComposeAdaptive (nq1 : List T → SLang U) (nq2 : U -> List T → SLang V) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h1 : PureDP nq1 ((ε₁ : ℝ) / ε₂)) (h2 : ∀ u : U, PureDP (nq2 u) ((ε₃ : ℝ) / ε₄)) : PureDP (privComposeAdaptive nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by simp [PureDP] at * rcases h1 with ⟨h1a, h1nz⟩ @@ -147,9 +154,8 @@ theorem PureDP_ComposeAdaptive' (nq1 : List T → SLang U) (nq2 : U -> List T replace h2 := h2 u rw [event_eq_singleton] at h2 simp [DP_singleton] at h2 - rcases h2 with ⟨h2a, h2nz⟩ - apply h2a - apply neighbours + rcases h2 with ⟨h2a, _⟩ + exact h2a l₁ l₂ neighbours x simp [privComposeAdaptive] @@ -189,7 +195,6 @@ theorem PureDP_ComposeAdaptive' (nq1 : List T → SLang U) (nq2 : U -> List T have hg : (∑' (a : U), nq1 l₂ a * nq2 a l₂ x) = (∑' (a : U), g a) := by congr rw [hf, hg] - -- Conclude by Mediant lemma apply bounded_quotient apply h3 all_goals (intro u; rcases (h2 u) with ⟨ _ , h2nz ⟩) @@ -203,20 +208,16 @@ theorem PureDP_ComposeAdaptive' (nq1 : List T → SLang U) (nq2 : U -> List T simp only [privComposeAdaptive, bind, pure, bind_pure, bind_apply] - cases (Classical.em (U -> False)) + cases (isEmpty_or_nonempty U) · rename_i Hu_empty - -- let (HU : IsEmpty U) := { false := Hu_empty } - -- rw [tsum_empty] - -- simp - -- Conclusion is false. - -- Can I get a contradiction out of either NonZeroNQ or do I have to strengthen - -- U to be nonempty? exfalso - simp [NonZeroNQ] at h1nz + -- Because ``SLang U`` values are functions out of U, we can't get a contradiction out of + -- "run nq1 to get a value of type U" like we would in an operational approach. + -- + -- We should restrict U to be nonempty. sorry - · -- Nonempty: series is bounded below by value at u0 (which is positive) - rename_i Hu + · rename_i Hu simp at Hu rcases Hu with ⟨ u0 ⟩ rcases h2 u0 with ⟨ _ , h2nz ⟩ @@ -228,4 +229,6 @@ theorem PureDP_ComposeAdaptive' (nq1 : List T → SLang U) (nq2 : U -> List T apply ENNReal.mul_pos · apply h1nz · apply h2nz + + end SLang From 81749c4f36f949d3c663e3ed1ab0b9263d6b052e Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Tue, 4 Jun 2024 14:08:02 -0400 Subject: [PATCH 008/122] restrict composition to only nonempty middle types --- .../DifferentialPrivacy/Pure/Composition.lean | 111 +++++++++--------- 1 file changed, 54 insertions(+), 57 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Pure/Composition.lean b/SampCert/DifferentialPrivacy/Pure/Composition.lean index 722464de..ee84583a 100644 --- a/SampCert/DifferentialPrivacy/Pure/Composition.lean +++ b/SampCert/DifferentialPrivacy/Pure/Composition.lean @@ -13,7 +13,7 @@ noncomputable section open Classical Set - +variable [Hu : Nonempty U] -- Solving the side conditions needs to be done separately depending on if u is inhabited or not /-- @@ -34,45 +34,44 @@ lemma tsum_mediant (f g : U -> ENNReal) (hg0 : ∀ u, g u ≠ 0) (hf0 : ∀ u, f simp case inl => rename_i assumption_g - cases (isEmpty_or_nonempty U) - · rename_i HU - rw [iSup_of_empty] - rw [tsum_empty] - rw [tsum_empty] - simp - · rename_i Hu - simp at Hu - rcases Hu with ⟨ u0 ⟩ - apply (ENNReal.div_le_iff_le_mul _ _).mpr - · rw [← ENNReal.tsum_mul_left] - apply ENNReal.tsum_le_tsum - intro u - apply (ENNReal.div_le_iff_le_mul _ _).mp - · refine (le_iSup (fun u => HDiv.hDiv (f u) (g u)) u) - · left; apply hg0 - · right - apply ne_of_gt - apply (LT.lt.trans_le ?g1 ?g2) - case g2 => - apply le_iSup - apply u - refine (ENNReal.div_pos (hf0 u) ?g1.hb) - apply assumption_g - · left - apply ne_of_gt - apply (LT.lt.trans_le ?z1 ?z2) - case z2 => - apply ENNReal.le_tsum - apply u0 - exact pos_iff_ne_zero.mpr (hg0 u0) + -- cases (isEmpty_or_nonempty U) + -- · rename_i HU + -- rw [iSup_of_empty] + -- rw [tsum_empty] + -- rw [tsum_empty] + -- simp + -- · rename_i Hu + rcases Hu with ⟨ u0 ⟩ + apply (ENNReal.div_le_iff_le_mul _ _).mpr + · rw [← ENNReal.tsum_mul_left] + apply ENNReal.tsum_le_tsum + intro u + apply (ENNReal.div_le_iff_le_mul _ _).mp + · refine (le_iSup (fun u => HDiv.hDiv (f u) (g u)) u) + · left; apply hg0 · right apply ne_of_gt - apply (LT.lt.trans_le ?z3 ?z4) - case z4 => + apply (LT.lt.trans_le ?g1 ?g2) + case g2 => apply le_iSup - apply u0 - refine (ENNReal.div_pos (hf0 u0) ?z6) + apply u + refine (ENNReal.div_pos (hf0 u) ?g1.hb) apply assumption_g + · left + apply ne_of_gt + apply (LT.lt.trans_le ?z1 ?z2) + case z2 => + apply ENNReal.le_tsum + apply u0 + exact pos_iff_ne_zero.mpr (hg0 u0) + · right + apply ne_of_gt + apply (LT.lt.trans_le ?z3 ?z4) + case z4 => + apply le_iSup + apply u0 + refine (ENNReal.div_pos (hf0 u0) ?z6) + apply assumption_g lemma bounded_quotient (f g : U -> ENNReal) (b : ENNReal) (h_bound : ∀ (u : U), f u / g u ≤ b) (hg0 : ∀ u, g u ≠ 0) (hf0 : ∀ u, f u ≠ 0) : @@ -208,27 +207,25 @@ theorem PureDP_ComposeAdaptive (nq1 : List T → SLang U) (nq2 : U -> List T → simp only [privComposeAdaptive, bind, pure, bind_pure, bind_apply] - cases (isEmpty_or_nonempty U) - · rename_i Hu_empty - exfalso - -- Because ``SLang U`` values are functions out of U, we can't get a contradiction out of - -- "run nq1 to get a value of type U" like we would in an operational approach. - -- - -- We should restrict U to be nonempty. - sorry - - · rename_i Hu - simp at Hu - rcases Hu with ⟨ u0 ⟩ - rcases h2 u0 with ⟨ _ , h2nz ⟩ - apply ne_of_gt - apply (LT.lt.trans_le ?g1 ?g2) - case g2 => - apply ENNReal.le_tsum - apply u0 - apply ENNReal.mul_pos - · apply h1nz - · apply h2nz + -- cases (isEmpty_or_nonempty U) + -- · rename_i Hu_empty + -- exfalso + -- -- Because ``SLang U`` values are functions out of U, we can't get a contradiction out of + -- -- "run nq1 to get a value of type U" like we would in an operational approach. + -- -- + -- -- We should restrict U to be nonempty. + -- sorry + -- · rename_i Hu + rcases Hu with ⟨ u0 ⟩ + rcases h2 u0 with ⟨ _ , h2nz ⟩ + apply ne_of_gt + apply (LT.lt.trans_le ?g1 ?g2) + case g2 => + apply ENNReal.le_tsum + apply u0 + apply ENNReal.mul_pos + · apply h1nz + · apply h2nz end SLang From 257633553318f72a6547a054861f4f0f77bdd9b6 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Tue, 4 Jun 2024 14:41:58 -0400 Subject: [PATCH 009/122] template adaptive composition --- .../ZeroConcentrated/AdaptiveComposition.lean | 135 ++++++++++++++++++ .../ZeroConcentrated/System.lean | 1 + 2 files changed, 136 insertions(+) create mode 100644 SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean new file mode 100644 index 00000000..23eecb9d --- /dev/null +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean @@ -0,0 +1,135 @@ +/- +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.Abstract +import SampCert.DifferentialPrivacy.ZeroConcentrated.DP + +/-! +# zCDP Adaptive Composition + +This file builds up to a zCDP bound on adaptively composed zCDP queries. +-/ + +noncomputable section + +open Classical Nat Int Real ENNReal MeasureTheory Measure + +namespace SLang + +variable { T U V : Type } +variable [Inhabited U] +variable [Inhabited V] + +/-- +Adaptively Composed queries satisfy zCDP Renyi divergence bound. +-/ +theorem privComposeAdaptive_zCDPBound {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} {ε₁ ε₂ ε₃ ε₄ : ℕ+} (h1 : zCDPBound nq1 ((ε₁ : ℝ) / ε₂)) (h2 : ∀ u, zCDPBound (nq2 u) ((ε₃ : ℝ) / ε₄)) (nn1 : NonZeroNQ nq1) (nn2 : ∀ u, NonZeroNQ (nq2 u)) (nt1 : NonTopRDNQ nq1) (nt2 : ∀ u, NonTopRDNQ (nq2 u)) (nts1 : NonTopNQ nq1) (nts2 : ∀ u, NonTopNQ (nq2 u)) : + zCDPBound (privComposeAdaptive nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by + sorry + +/-- +Adaptive composed query distribution is nowhere zero +-/ +theorem privComposeAdaptive_NonZeroNQ {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (nt1 : NonZeroNQ nq1) (nt2 : ∀ u, NonZeroNQ (nq2 u)) : + NonZeroNQ (privComposeAdaptive nq1 nq2) := by + admit + +/-- +All outputs of a adaptive composed query have finite probability. +-/ +theorem privComposeAdaptive_NonTopNQ {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (nt1 : NonTopNQ nq1) (nt2 : ∀ u, NonTopNQ (nq2 u)) : + NonTopNQ (privComposeAdaptive nq1 nq2) := by + admit + +/-- +Adaptive composed query is a proper distribution +-/ +theorem privComposeAdaptive_NonTopSum {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (nt1 : NonTopNQ nq1) (nt2 : ∀ u, NonTopNQ (nq2 u)) : + NonTopSum (privComposeAdaptive nq1 nq2) := by + admit + + + +/-- +Renyi divergence beteeen adaptive composed queries on neighbours are finite. +-/ +theorem privCompose_NonTopRDNQ {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (nt1 : NonTopRDNQ nq1) (nt2 : ∀ u, NonTopRDNQ (nq2 u)) (nn1 : NonTopNQ nq1) (nn2 : ∀ u, NonTopNQ (nq2 u)) : + NonTopRDNQ (privComposeAdaptive nq1 nq2) := by admit + -- simp [NonTopRDNQ] at * + -- intro α h1 l₁ l₂ h2 + -- replace nt1 := nt1 α h1 l₁ l₂ h2 + -- replace nt2 := nt2 α h1 l₁ l₂ h2 + -- simp [privCompose] + -- 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_nonneg _ _ (le_of_lt (lt_trans zero_lt_one h1))] + -- rw [ENNReal.mul_rpow_of_ne_top (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 + +/-- +``privComposeAdaptive`` satisfies zCDP +-/ +theorem privComposeAdaptive_zCDP (nq1 : List T → SLang U) (nq2 : U -> List T → SLang V) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h : zCDP nq1 ((ε₁ : ℝ) / ε₂)) (h' : ∀ u, zCDP (nq2 u) ((ε₃ : ℝ) / ε₄)) : + zCDP (privComposeAdaptive nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by + simp [zCDP] at * + rcases h with ⟨ _ , _ , _ , _ , _ ⟩ + repeat any_goals constructor + · admit + · admit + · admit + · admit + · admit + +end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean index a7bbbe8b..35977cbb 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean @@ -7,6 +7,7 @@ import SampCert.DifferentialPrivacy.Abstract import SampCert.DifferentialPrivacy.ZeroConcentrated.DP import SampCert.DifferentialPrivacy.ZeroConcentrated.Mechanism.Basic import SampCert.DifferentialPrivacy.ZeroConcentrated.Composition +import SampCert.DifferentialPrivacy.ZeroConcentrated.AdaptiveComposition import SampCert.DifferentialPrivacy.ZeroConcentrated.Postprocessing /-! From 737d95535ea0874b28a1fe88d1ceeabe37785775 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Tue, 4 Jun 2024 14:50:50 -0400 Subject: [PATCH 010/122] NonZeroNQ --- .../ZeroConcentrated/AdaptiveComposition.lean | 13 ++++++++++--- 1 file changed, 10 insertions(+), 3 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean index 23eecb9d..fa62083a 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean @@ -19,8 +19,8 @@ open Classical Nat Int Real ENNReal MeasureTheory Measure namespace SLang variable { T U V : Type } -variable [Inhabited U] -variable [Inhabited V] +variable [HU : Inhabited U] +variable [HV : Inhabited V] /-- Adaptively Composed queries satisfy zCDP Renyi divergence bound. @@ -34,7 +34,14 @@ Adaptive composed query distribution is nowhere zero -/ theorem privComposeAdaptive_NonZeroNQ {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (nt1 : NonZeroNQ nq1) (nt2 : ∀ u, NonZeroNQ (nq2 u)) : NonZeroNQ (privComposeAdaptive nq1 nq2) := by - admit + simp [NonZeroNQ] at * + simp [privComposeAdaptive] + intros l n + rcases HU with ⟨ u0 ⟩ + exists u0 + apply And.intro + · apply nt1 + · apply nt2 /-- All outputs of a adaptive composed query have finite probability. From a6d03f9a9d96ca467d7cd45d6732e147ab34189c Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Tue, 4 Jun 2024 16:50:47 -0400 Subject: [PATCH 011/122] skech marginal-returning adaptive and conditioning chain rule --- SampCert/DifferentialPrivacy/Abstract.lean | 25 ++++++++++++++++++- .../DifferentialPrivacy/Pure/Composition.lean | 3 ++- .../DifferentialPrivacy/RenyiDivergence.lean | 4 +-- .../ZeroConcentrated/AdaptiveComposition.lean | 20 +++++++++++++-- 4 files changed, 46 insertions(+), 6 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Abstract.lean b/SampCert/DifferentialPrivacy/Abstract.lean index b5f43107..95a9eb1f 100644 --- a/SampCert/DifferentialPrivacy/Abstract.lean +++ b/SampCert/DifferentialPrivacy/Abstract.lean @@ -28,15 +28,38 @@ namespace SLang abbrev Query (T U : Type) := List T → U abbrev Mechanism (T U : Type) := List T → SLang U - /-- General (value-dependent) composition of mechanisms -/ +def privComposeAdaptive' (nq1 : Mechanism T U) (nq2 : U -> Mechanism T V) (l : List T) : SLang (U × V) := do + let A <- nq1 l + let B <- nq2 A l + return (A, B) + +/-- +Chain rule relating the adaptive composition definitions + +The joint distribution decomposes into the conditional and marginal (ie, nq1 l) distributions +-/ +lemma privComposeChainRule (nq1 : Mechanism T U) (nq2 : U -> Mechanism T V) (l : List T) : + ∀ (u : U), ∀ (v : V), privComposeAdaptive' nq1 nq2 l (u, v) = nq1 l u * nq2 u l v := by + intros u v + simp [privComposeAdaptive'] + -- Add an ite to the outer sum and simplify + admit + +/-- +Conditional composition of mechanisms +-/ def privComposeAdaptive (nq1 : Mechanism T U) (nq2 : U -> Mechanism T V) (l : List T) : SLang V := do let A <- nq1 l let B <- nq2 A l return B + + + + /-- Composition of independent mechanisms -/ diff --git a/SampCert/DifferentialPrivacy/Pure/Composition.lean b/SampCert/DifferentialPrivacy/Pure/Composition.lean index ee84583a..253f6838 100644 --- a/SampCert/DifferentialPrivacy/Pure/Composition.lean +++ b/SampCert/DifferentialPrivacy/Pure/Composition.lean @@ -215,7 +215,8 @@ theorem PureDP_ComposeAdaptive (nq1 : List T → SLang U) (nq2 : U -> List T → -- -- -- -- We should restrict U to be nonempty. -- sorry - -- · rename_i Hu + -- · + rcases Hu with ⟨ u0 ⟩ rcases h2 u0 with ⟨ _ , h2nz ⟩ apply ne_of_gt diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index be404036..a962fbee 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -24,7 +24,7 @@ noncomputable def RenyiDivergence (p q : T → ENNReal) (α : ℝ) : ℝ := (α - 1)⁻¹ * Real.log (∑' x : T, (p x)^α * (q x)^(1 - α)).toReal /-- -Closed form of the series in the definition of hte Renyi divergence. +Closed form of the series in the definition of the Renyi divergence. -/ 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 @@ -39,7 +39,7 @@ theorem RenyiDivergenceExpectation (p q : T → ENNReal) {α : ℝ} (h : 1 < α) . apply h2 x /-- -Closed form for the Renti Divergence. +Closed form for the Renyi Divergence. -/ 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 diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean index fa62083a..7071e92c 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean @@ -22,11 +22,20 @@ variable { T U V : Type } variable [HU : Inhabited U] variable [HV : Inhabited V] +/-- +The Renyi divergence is monotonic in its sum +-/ +lemma renyi_mono_sum (x y : ENNReal) {α : ℝ} (h : 1 < α) : Real.log (((ENNReal.ofReal α - 1) * x).toReal) ≤ Real.log (((ENNReal.ofReal α - 1) * y).toReal) -> (x ≤ y) := + sorry + + /-- Adaptively Composed queries satisfy zCDP Renyi divergence bound. -/ theorem privComposeAdaptive_zCDPBound {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} {ε₁ ε₂ ε₃ ε₄ : ℕ+} (h1 : zCDPBound nq1 ((ε₁ : ℝ) / ε₂)) (h2 : ∀ u, zCDPBound (nq2 u) ((ε₃ : ℝ) / ε₄)) (nn1 : NonZeroNQ nq1) (nn2 : ∀ u, NonZeroNQ (nq2 u)) (nt1 : NonTopRDNQ nq1) (nt2 : ∀ u, NonTopRDNQ (nq2 u)) (nts1 : NonTopNQ nq1) (nts2 : ∀ u, NonTopNQ (nq2 u)) : zCDPBound (privComposeAdaptive nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by + rw [zCDPBound] + sorry /-- @@ -48,13 +57,20 @@ All outputs of a adaptive composed query have finite probability. -/ theorem privComposeAdaptive_NonTopNQ {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (nt1 : NonTopNQ nq1) (nt2 : ∀ u, NonTopNQ (nq2 u)) : NonTopNQ (privComposeAdaptive nq1 nq2) := by + simp [NonTopNQ] at * + simp [privComposeAdaptive] admit /-- Adaptive composed query is a proper distribution -/ -theorem privComposeAdaptive_NonTopSum {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (nt1 : NonTopNQ nq1) (nt2 : ∀ u, NonTopNQ (nq2 u)) : +theorem privComposeAdaptive_NonTopSum {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (nt1 : NonTopSum nq1) (nt2 : ∀ u, NonTopSum (nq2 u)) : NonTopSum (privComposeAdaptive nq1 nq2) := by + rw [NonTopSum] at * + simp only [privComposeAdaptive, Bind.bind, pure, bind_pure, bind_apply] + intro l + -- forall n, a, nq1 l a * nq2 a l n ≤ (nq1 l a + nq2 a l n )^2 + -- Bound the series above by ∑' (n : V) (a : U), (nq1 l a + nq2 a l n )^2 admit @@ -62,7 +78,7 @@ theorem privComposeAdaptive_NonTopSum {nq1 : List T → SLang U} {nq2 : U -> Lis /-- Renyi divergence beteeen adaptive composed queries on neighbours are finite. -/ -theorem privCompose_NonTopRDNQ {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (nt1 : NonTopRDNQ nq1) (nt2 : ∀ u, NonTopRDNQ (nq2 u)) (nn1 : NonTopNQ nq1) (nn2 : ∀ u, NonTopNQ (nq2 u)) : +theorem privComposeAdaptive_NonTopRDNQ {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (nt1 : NonTopRDNQ nq1) (nt2 : ∀ u, NonTopRDNQ (nq2 u)) (nn1 : NonTopNQ nq1) (nn2 : ∀ u, NonTopNQ (nq2 u)) : NonTopRDNQ (privComposeAdaptive nq1 nq2) := by admit -- simp [NonTopRDNQ] at * -- intro α h1 l₁ l₂ h2 From 7d659b4d9ea584d83b853493fff917960611f488 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Wed, 5 Jun 2024 10:24:36 -0400 Subject: [PATCH 012/122] chain rule --- SampCert/DifferentialPrivacy/Abstract.lean | 31 ++++++++++++++++++++-- 1 file changed, 29 insertions(+), 2 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Abstract.lean b/SampCert/DifferentialPrivacy/Abstract.lean index 95a9eb1f..3008a7c4 100644 --- a/SampCert/DifferentialPrivacy/Abstract.lean +++ b/SampCert/DifferentialPrivacy/Abstract.lean @@ -36,6 +36,9 @@ def privComposeAdaptive' (nq1 : Mechanism T U) (nq2 : U -> Mechanism T V) (l : L let B <- nq2 A l return (A, B) + +-- set_option pp.notation false + /-- Chain rule relating the adaptive composition definitions @@ -45,8 +48,32 @@ lemma privComposeChainRule (nq1 : Mechanism T U) (nq2 : U -> Mechanism T V) (l : ∀ (u : U), ∀ (v : V), privComposeAdaptive' nq1 nq2 l (u, v) = nq1 l u * nq2 u l v := by intros u v simp [privComposeAdaptive'] - -- Add an ite to the outer sum and simplify - admit + -- How to simplify cases without explicit ite? this is silly + have hrw1 : ∀ (a : U), nq1 l a * (∑' (a_1 : V), if u = a ∧ v = a_1 then nq2 a l a_1 else 0) = if (u = a) then (nq1 l a * ∑' (a_1 : V), if u = a ∧ v = a_1 then nq2 a l a_1 else 0) else 0 := by + intro a + split + · simp + · aesop + have hrw2 : ∀ (a : U), (if (u = a) then (nq1 l a * ∑' (a_1 : V), if u = a ∧ v = a_1 then nq2 a l a_1 else 0) else 0) = (nq1 l u * (∑' (a_1 : V), if u = a ∧ v = a_1 then nq2 u l a_1 else 0)) := by + intro a + split + · aesop + · aesop + rw [tsum_congr hrw1] + rw [tsum_congr hrw2] + clear hrw1 hrw2 + rw [ENNReal.tsum_mul_left] + congr 1 + rw [<- ENNReal.tsum_prod] + have hrw3 : ∀ (p : U × V), (if u = p.1 ∧ v = p.2 then nq2 u l p.2 else 0) = nq2 u l v * (if p = (u, v) then 1 else 0) := by + intro p + split + · aesop + · aesop + rw [tsum_congr hrw3] + rw [ENNReal.tsum_mul_left] + rw [tsum_ite_eq] + exact MulOneClass.mul_one (nq2 u l v) /-- Conditional composition of mechanisms From e21a51051186ea3fff1aacb046d70017664250c6 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Wed, 5 Jun 2024 10:49:08 -0400 Subject: [PATCH 013/122] better pure proof w/ chain rule --- SampCert/DifferentialPrivacy/Abstract.lean | 1 + .../DifferentialPrivacy/Pure/Composition.lean | 49 +++++++++++++++++++ 2 files changed, 50 insertions(+) diff --git a/SampCert/DifferentialPrivacy/Abstract.lean b/SampCert/DifferentialPrivacy/Abstract.lean index 3008a7c4..c274c550 100644 --- a/SampCert/DifferentialPrivacy/Abstract.lean +++ b/SampCert/DifferentialPrivacy/Abstract.lean @@ -75,6 +75,7 @@ lemma privComposeChainRule (nq1 : Mechanism T U) (nq2 : U -> Mechanism T V) (l : rw [tsum_ite_eq] exact MulOneClass.mul_one (nq2 u l v) + /-- Conditional composition of mechanisms -/ diff --git a/SampCert/DifferentialPrivacy/Pure/Composition.lean b/SampCert/DifferentialPrivacy/Pure/Composition.lean index 253f6838..a97f3855 100644 --- a/SampCert/DifferentialPrivacy/Pure/Composition.lean +++ b/SampCert/DifferentialPrivacy/Pure/Composition.lean @@ -228,5 +228,54 @@ theorem PureDP_ComposeAdaptive (nq1 : List T → SLang U) (nq2 : U -> List T → · apply h1nz · apply h2nz +-- Better proof for Pure DP adaptive composition +theorem PureDP_ComposeAdaptive' (nq1 : List T → SLang U) (nq2 : U -> List T → SLang V) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h1 : PureDP nq1 ((ε₁ : ℝ) / ε₂)) (h2 : ∀ u : U, PureDP (nq2 u) ((ε₃ : ℝ) / ε₄)) : + PureDP (privComposeAdaptive' nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by + simp [PureDP] at * + rcases h1 with ⟨h1a, h1nz⟩ + rw [event_eq_singleton] at * + simp [DP_singleton] at * + apply And.intro + · intros l₁ l₂ Hl₁l₂ u v + rw [privComposeChainRule] + rw [privComposeChainRule] + rw [Real.exp_add] + rw [ENNReal.ofReal_mul ?s1] + case s1 => apply Real.exp_nonneg + rw [ENNReal.div_eq_inv_mul] + rw [ENNReal.mul_inv] + · rw [<- mul_assoc] + rw [mul_right_comm] + conv => + lhs + arg 1 + rw [mul_assoc] + rw [mul_right_comm] + rw [← ENNReal.div_eq_inv_mul] + rw [← ENNReal.div_eq_inv_mul] + rcases h2 u with ⟨ h2a'_pre, _ ⟩ + rw [event_eq_singleton] at h2a'_pre + simp [DP_singleton] at h2a'_pre + exact (mul_le_mul' (h1a l₁ l₂ Hl₁l₂ u) (h2a'_pre l₁ l₂ Hl₁l₂ v)) + · left + apply h1nz + · right + rcases (h2 u) with ⟨ _ , h2nz ⟩ + apply h2nz + · simp only [NonZeroNQ] + intros l n + simp only [privComposeAdaptive', bind, pure, bind_pure, bind_apply] + apply ne_of_gt + apply (LT.lt.trans_le ?g1 ?g2) + case g2 => + apply ENNReal.le_tsum + apply n.1 + apply ENNReal.mul_pos + · apply h1nz + · simp + exists n.2 + simp + rcases (h2 n.1) with ⟨ _ , H2nz ⟩ + apply H2nz end SLang From 1d4ad5edcf07fca7a68dba10ceca4ba3a4d111fa Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Wed, 5 Jun 2024 12:00:41 -0400 Subject: [PATCH 014/122] adaptive composition proof skeleton --- .../DifferentialPrivacy/RenyiDivergence.lean | 11 ++++ .../ZeroConcentrated/AdaptiveComposition.lean | 56 +++++++++++-------- 2 files changed, 45 insertions(+), 22 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index a962fbee..93291bc8 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -52,3 +52,14 @@ theorem RenyiDivergenceExpectation' (p q : T → ENNReal) {α : ℝ} (h : 1 < α . apply le_of_lt (lt_trans Real.zero_lt_one h ) . apply h1 x . apply h2 x + + +/-- +The Renyi divergence is monotonic in the value of its sum. +-/ +lemma RenyiDivergence_mono_sum (x y : ℝ) {α : ℝ} (h : 1 < α) : Real.log ((α - 1) * x) ≤ Real.log ((α - 1) * y) -> (x ≤ y) := + sorry + +lemma RenyiDivergence_exp {p q : T → ENNReal} (α : ℝ) (h : 1 < α) : + Real.log ((α - 1) * RenyiDivergence f g α) = (∑' x : T, (p x)^α * (q x)^(1 - α)).toReal := by + sorry diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean index 7071e92c..b3a71ecc 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean @@ -22,55 +22,67 @@ variable { T U V : Type } variable [HU : Inhabited U] variable [HV : Inhabited V] + + /-- -The Renyi divergence is monotonic in its sum +Bound on Renyi divergence on adaptively composed queries -/ -lemma renyi_mono_sum (x y : ENNReal) {α : ℝ} (h : 1 < α) : Real.log (((ENNReal.ofReal α - 1) * x).toReal) ≤ Real.log (((ENNReal.ofReal α - 1) * y).toReal) -> (x ≤ y) := - sorry +lemma primComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (α : ℝ) (Hα : 1 < α) : + RenyiDivergence (privComposeAdaptive' nq1 nq2 l₁) (privComposeAdaptive' nq1 nq2 l₂) α ≤ RenyiDivergence (nq1 l₁) (nq1 l₂) α + ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α := by + apply RenyiDivergence_mono_sum + -- rw [<- RenyiDivergence_exp] + sorry + sorry + sorry /-- Adaptively Composed queries satisfy zCDP Renyi divergence bound. -/ theorem privComposeAdaptive_zCDPBound {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} {ε₁ ε₂ ε₃ ε₄ : ℕ+} (h1 : zCDPBound nq1 ((ε₁ : ℝ) / ε₂)) (h2 : ∀ u, zCDPBound (nq2 u) ((ε₃ : ℝ) / ε₄)) (nn1 : NonZeroNQ nq1) (nn2 : ∀ u, NonZeroNQ (nq2 u)) (nt1 : NonTopRDNQ nq1) (nt2 : ∀ u, NonTopRDNQ (nq2 u)) (nts1 : NonTopNQ nq1) (nts2 : ∀ u, NonTopNQ (nq2 u)) : - zCDPBound (privComposeAdaptive nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by + zCDPBound (privComposeAdaptive' nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by rw [zCDPBound] - - sorry + intro α Hα l₁ l₂ Hneighbours + -- Loose + apply (@LE.le.trans _ _ _ (1/2 * (↑↑ε₁ / ↑↑ε₂)^2 * α + 1/2 * (↑↑ε₃ / ↑↑ε₄)^2 * α) _ _ ?case_sq) + case case_sq => + -- Binomial bound + sorry + + -- Rewrite the upper bounds in terms of Renyi divergences of nq1/nq2 + rw [zCDPBound] at h1 + have marginal_ub := h1 α Hα l₁ l₂ Hneighbours + have conditional_ub : (⨆ (u : U), RenyiDivergence (nq2 u l₁) (nq2 u l₂) α ≤ 1 / 2 * (↑↑ε₃ / ↑↑ε₄) ^ 2 * α) := + ciSup_le fun x => h2 x α Hα l₁ l₂ Hneighbours + apply (@LE.le.trans _ _ _ (RenyiDivergence (nq1 l₁) (nq1 l₂) α + ⨆ (u : U), RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) _ _ ?case_alg) + case case_alg => linarith + apply (primComposeAdaptive_renyi_bound _ Hα) /-- Adaptive composed query distribution is nowhere zero -/ theorem privComposeAdaptive_NonZeroNQ {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (nt1 : NonZeroNQ nq1) (nt2 : ∀ u, NonZeroNQ (nq2 u)) : - NonZeroNQ (privComposeAdaptive nq1 nq2) := by + NonZeroNQ (privComposeAdaptive' nq1 nq2) := by simp [NonZeroNQ] at * - simp [privComposeAdaptive] - intros l n - rcases HU with ⟨ u0 ⟩ - exists u0 - apply And.intro - · apply nt1 - · apply nt2 + simp [privComposeAdaptive'] + sorry /-- All outputs of a adaptive composed query have finite probability. -/ theorem privComposeAdaptive_NonTopNQ {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (nt1 : NonTopNQ nq1) (nt2 : ∀ u, NonTopNQ (nq2 u)) : - NonTopNQ (privComposeAdaptive nq1 nq2) := by + NonTopNQ (privComposeAdaptive' nq1 nq2) := by simp [NonTopNQ] at * - simp [privComposeAdaptive] admit /-- Adaptive composed query is a proper distribution -/ theorem privComposeAdaptive_NonTopSum {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (nt1 : NonTopSum nq1) (nt2 : ∀ u, NonTopSum (nq2 u)) : - NonTopSum (privComposeAdaptive nq1 nq2) := by + NonTopSum (privComposeAdaptive' nq1 nq2) := by rw [NonTopSum] at * - simp only [privComposeAdaptive, Bind.bind, pure, bind_pure, bind_apply] + simp only [privComposeAdaptive', Bind.bind, pure, bind_pure, bind_apply] intro l - -- forall n, a, nq1 l a * nq2 a l n ≤ (nq1 l a + nq2 a l n )^2 - -- Bound the series above by ∑' (n : V) (a : U), (nq1 l a + nq2 a l n )^2 admit @@ -79,7 +91,7 @@ theorem privComposeAdaptive_NonTopSum {nq1 : List T → SLang U} {nq2 : U -> Lis Renyi divergence beteeen adaptive composed queries on neighbours are finite. -/ theorem privComposeAdaptive_NonTopRDNQ {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (nt1 : NonTopRDNQ nq1) (nt2 : ∀ u, NonTopRDNQ (nq2 u)) (nn1 : NonTopNQ nq1) (nn2 : ∀ u, NonTopNQ (nq2 u)) : - NonTopRDNQ (privComposeAdaptive nq1 nq2) := by admit + NonTopRDNQ (privComposeAdaptive' nq1 nq2) := by admit -- simp [NonTopRDNQ] at * -- intro α h1 l₁ l₂ h2 -- replace nt1 := nt1 α h1 l₁ l₂ h2 @@ -145,7 +157,7 @@ theorem privComposeAdaptive_NonTopRDNQ {nq1 : List T → SLang U} {nq2 : U -> Li ``privComposeAdaptive`` satisfies zCDP -/ theorem privComposeAdaptive_zCDP (nq1 : List T → SLang U) (nq2 : U -> List T → SLang V) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h : zCDP nq1 ((ε₁ : ℝ) / ε₂)) (h' : ∀ u, zCDP (nq2 u) ((ε₃ : ℝ) / ε₄)) : - zCDP (privComposeAdaptive nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by + zCDP (privComposeAdaptive' nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by simp [zCDP] at * rcases h with ⟨ _ , _ , _ , _ , _ ⟩ repeat any_goals constructor From 96ef09efb1031b7640c01302161198c052078902 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Wed, 5 Jun 2024 14:26:11 -0400 Subject: [PATCH 015/122] skeleton for renyi bound proof --- .../DifferentialPrivacy/RenyiDivergence.lean | 6 +- .../ZeroConcentrated/AdaptiveComposition.lean | 72 +++++++++++++++++-- 2 files changed, 70 insertions(+), 8 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index 93291bc8..5654fbb7 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -57,9 +57,9 @@ theorem RenyiDivergenceExpectation' (p q : T → ENNReal) {α : ℝ} (h : 1 < α /-- The Renyi divergence is monotonic in the value of its sum. -/ -lemma RenyiDivergence_mono_sum (x y : ℝ) {α : ℝ} (h : 1 < α) : Real.log ((α - 1) * x) ≤ Real.log ((α - 1) * y) -> (x ≤ y) := +lemma RenyiDivergence_mono_sum (x y : ℝ) (α : ℝ) (h : 1 < α) : Real.exp ((α - 1) * x) ≤ Real.exp ((α - 1) * y) -> (x ≤ y) := sorry -lemma RenyiDivergence_exp {p q : T → ENNReal} (α : ℝ) (h : 1 < α) : - Real.log ((α - 1) * RenyiDivergence f g α) = (∑' x : T, (p x)^α * (q x)^(1 - α)).toReal := by +lemma RenyiDivergence_exp (p q : T → ENNReal) {α : ℝ} (h : 1 < α) : + Real.exp ((α - 1) * RenyiDivergence p q α) = (∑' x : T, (p x)^α * (q x)^(1 - α)).toReal := by sorry diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean index b3a71ecc..c807661b 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean @@ -24,17 +24,79 @@ variable [HV : Inhabited V] + /-- Bound on Renyi divergence on adaptively composed queries -/ lemma primComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (α : ℝ) (Hα : 1 < α) : RenyiDivergence (privComposeAdaptive' nq1 nq2 l₁) (privComposeAdaptive' nq1 nq2 l₂) α ≤ RenyiDivergence (nq1 l₁) (nq1 l₂) α + ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α := by - apply RenyiDivergence_mono_sum - -- rw [<- RenyiDivergence_exp] - sorry - sorry - sorry + apply (RenyiDivergence_mono_sum _ _ α Hα) + rw [RenyiDivergence_exp (privComposeAdaptive' nq1 nq2 l₁) (privComposeAdaptive' nq1 nq2 l₂) Hα] + rw [left_distrib] + rw [Real.exp_add] + have hmono_1 : rexp ((α - 1) * ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) = ⨆ u, rexp ((α - 1) * RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) := by + sorry + rw [hmono_1] + clear hmono_1 + rw [RenyiDivergence_exp (nq1 l₁) (nq1 l₂) Hα] + rw [mul_comm] + rw [<- (ENNReal.toReal_ofReal_mul _ _ ?h)] + case h => sorry + rw [mul_comm] + rw [← ENNReal.tsum_mul_right] + + apply (@LE.le.trans _ _ _ ((∑' (i : U), nq1 l₁ i ^ α * nq1 l₂ i ^ (1 - α) * ENNReal.ofReal (rexp ((α - 1) * RenyiDivergence (nq2 i l₁) (nq2 i l₂) α))).toReal) _ _ ?goal2) + case goal2 => + -- Can I do this without summability? + apply toReal_mono' + · sorry + · sorry + + -- After this point the argument is tight + apply Eq.le + conv => + rhs + congr + congr + intro i + rw [RenyiDivergence_exp (nq2 i l₁) (nq2 i l₂) Hα] + + conv => + lhs + congr + congr + intro + rw [privComposeChainRule] + rw [privComposeChainRule] + + have test : ∀ i, ENNReal.ofReal (∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α)).toReal = ∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α) := by + intro i + apply ofReal_toReal + -- Need a summability thing here... try to avoid this lemma if possible + sorry + conv => + rhs + congr + congr + intro + rw [test] + rw [← ENNReal.tsum_mul_left] + clear test + + rw [<- ENNReal.tsum_prod] + congr + apply funext + intro p + rcases p with ⟨ u , v ⟩ + simp + rw [ENNReal.mul_rpow_of_nonneg _ _ ?sc1] + case sc1 => + sorry + rw [ENNReal.mul_rpow_of_nonneg _ _ ?sc2] + case sc2 => + sorry + rw [mul_mul_mul_comm] /-- Adaptively Composed queries satisfy zCDP Renyi divergence bound. From 9c8cb6c72f7ea7ac99544e84fc4efae0aef63919 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Wed, 5 Jun 2024 14:38:07 -0400 Subject: [PATCH 016/122] nits --- .../ZeroConcentrated/AdaptiveComposition.lean | 14 ++++++-------- 1 file changed, 6 insertions(+), 8 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean index c807661b..c67cc42a 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean @@ -105,7 +105,7 @@ theorem privComposeAdaptive_zCDPBound {nq1 : List T → SLang U} {nq2 : U -> Lis zCDPBound (privComposeAdaptive' nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by rw [zCDPBound] intro α Hα l₁ l₂ Hneighbours - -- Loose + -- Loose step apply (@LE.le.trans _ _ _ (1/2 * (↑↑ε₁ / ↑↑ε₂)^2 * α + 1/2 * (↑↑ε₃ / ↑↑ε₄)^2 * α) _ _ ?case_sq) case case_sq => -- Binomial bound @@ -148,7 +148,6 @@ theorem privComposeAdaptive_NonTopSum {nq1 : List T → SLang U} {nq2 : U -> Lis admit - /-- Renyi divergence beteeen adaptive composed queries on neighbours are finite. -/ @@ -221,12 +220,11 @@ theorem privComposeAdaptive_NonTopRDNQ {nq1 : List T → SLang U} {nq2 : U -> Li theorem privComposeAdaptive_zCDP (nq1 : List T → SLang U) (nq2 : U -> List T → SLang V) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h : zCDP nq1 ((ε₁ : ℝ) / ε₂)) (h' : ∀ u, zCDP (nq2 u) ((ε₃ : ℝ) / ε₄)) : zCDP (privComposeAdaptive' nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by simp [zCDP] at * - rcases h with ⟨ _ , _ , _ , _ , _ ⟩ repeat any_goals constructor - · admit - · admit - · admit - · admit - · admit + · apply privComposeAdaptive_zCDPBound <;> aesop + · apply privComposeAdaptive_NonZeroNQ <;> aesop + · apply privComposeAdaptive_NonTopSum <;> aesop + · apply privComposeAdaptive_NonTopNQ <;> aesop + · apply privComposeAdaptive_NonTopRDNQ <;> aesop end SLang From d8c6cea422e8742f86931e83a633fa1c64a6a0bc Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Wed, 5 Jun 2024 15:47:24 -0400 Subject: [PATCH 017/122] some progress on main Renyi lemmas --- .../DifferentialPrivacy/RenyiDivergence.lean | 27 ++++++++++++++++--- 1 file changed, 24 insertions(+), 3 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index 5654fbb7..77fa417f 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -53,13 +53,34 @@ theorem RenyiDivergenceExpectation' (p q : T → ENNReal) {α : ℝ} (h : 1 < α . apply h1 x . apply h2 x - /-- The Renyi divergence is monotonic in the value of its sum. -/ -lemma RenyiDivergence_mono_sum (x y : ℝ) (α : ℝ) (h : 1 < α) : Real.exp ((α - 1) * x) ≤ Real.exp ((α - 1) * y) -> (x ≤ y) := +lemma RenyiDivergence_mono_sum (x y : ℝ) (α : ℝ) (h : 1 < α) : Real.exp ((α - 1) * x) ≤ Real.exp ((α - 1) * y) -> (x ≤ y) := by + intro H + have H1 : ((α - 1) * x) ≤ ((α - 1) * y) := exp_le_exp.mp H + -- refine (mul_le_mul_iff_right ?a).mp ?a + -- apply? + -- refine ((@ENNReal.mul_le_mul_left x y _ _ ).mp _) + -- apply mul_lt_mul_left_of_neg + -- refine (mul_le_mul_left ?a0).mp ?a -- ????? + + -- How annoying + sorry lemma RenyiDivergence_exp (p q : T → ENNReal) {α : ℝ} (h : 1 < α) : Real.exp ((α - 1) * RenyiDivergence p q α) = (∑' x : T, (p x)^α * (q x)^(1 - α)).toReal := by - sorry + simp only [RenyiDivergence] + rw [<- mul_assoc] + have test : (α - 1) * (α - 1)⁻¹ = 1 := by + refine mul_inv_cancel ?h + linarith + rw [test] + clear test + simp + rw [Real.exp_log] + apply ENNReal.toReal_pos_iff.mpr + apply And.intro + · sorry + · sorry From 2d0590abaa073c70e1587c2186e54f20035710b7 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Wed, 5 Jun 2024 16:39:45 -0400 Subject: [PATCH 018/122] close some side conditions --- .../ZeroConcentrated/AdaptiveComposition.lean | 33 +++++++++++++------ 1 file changed, 23 insertions(+), 10 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean index c67cc42a..53ac9de2 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean @@ -23,8 +23,6 @@ variable [HU : Inhabited U] variable [HV : Inhabited V] - - /-- Bound on Renyi divergence on adaptively composed queries -/ @@ -42,7 +40,10 @@ lemma primComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> Lis rw [RenyiDivergence_exp (nq1 l₁) (nq1 l₂) Hα] rw [mul_comm] rw [<- (ENNReal.toReal_ofReal_mul _ _ ?h)] - case h => sorry + case h => + refine Real.iSup_nonneg ?hf + intro i + exact exp_nonneg ((α - 1) * RenyiDivergence (nq2 i l₁) (nq2 i l₂) α) rw [mul_comm] rw [← ENNReal.tsum_mul_right] @@ -50,7 +51,15 @@ lemma primComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> Lis case goal2 => -- Can I do this without summability? apply toReal_mono' - · sorry + · apply tsum_le_tsum + · intro i + refine (ENNReal.mul_le_mul_left ?h.h.h0 ?h.h.hinf).mpr ?h.h.a + · sorry + · sorry + · apply ENNReal.ofReal_le_ofReal + sorry + · sorry + · sorry · sorry -- After this point the argument is tight @@ -110,7 +119,6 @@ theorem privComposeAdaptive_zCDPBound {nq1 : List T → SLang U} {nq2 : U -> Lis case case_sq => -- Binomial bound sorry - -- Rewrite the upper bounds in terms of Renyi divergences of nq1/nq2 rw [zCDPBound] at h1 have marginal_ub := h1 α Hα l₁ l₂ Hneighbours @@ -127,7 +135,7 @@ theorem privComposeAdaptive_NonZeroNQ {nq1 : List T → SLang U} {nq2 : U -> Lis NonZeroNQ (privComposeAdaptive' nq1 nq2) := by simp [NonZeroNQ] at * simp [privComposeAdaptive'] - sorry + aesop /-- All outputs of a adaptive composed query have finite probability. @@ -135,7 +143,11 @@ All outputs of a adaptive composed query have finite probability. theorem privComposeAdaptive_NonTopNQ {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (nt1 : NonTopNQ nq1) (nt2 : ∀ u, NonTopNQ (nq2 u)) : NonTopNQ (privComposeAdaptive' nq1 nq2) := by simp [NonTopNQ] at * - admit + intros l u v + rw [privComposeChainRule] + apply ENNReal.mul_ne_top + · aesop + · aesop /-- Adaptive composed query is a proper distribution @@ -143,8 +155,7 @@ Adaptive composed query is a proper distribution theorem privComposeAdaptive_NonTopSum {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (nt1 : NonTopSum nq1) (nt2 : ∀ u, NonTopSum (nq2 u)) : NonTopSum (privComposeAdaptive' nq1 nq2) := by rw [NonTopSum] at * - simp only [privComposeAdaptive', Bind.bind, pure, bind_pure, bind_apply] - intro l + -- Chain rule won't help here admit @@ -152,7 +163,9 @@ theorem privComposeAdaptive_NonTopSum {nq1 : List T → SLang U} {nq2 : U -> Lis Renyi divergence beteeen adaptive composed queries on neighbours are finite. -/ theorem privComposeAdaptive_NonTopRDNQ {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (nt1 : NonTopRDNQ nq1) (nt2 : ∀ u, NonTopRDNQ (nq2 u)) (nn1 : NonTopNQ nq1) (nn2 : ∀ u, NonTopNQ (nq2 u)) : - NonTopRDNQ (privComposeAdaptive' nq1 nq2) := by admit + NonTopRDNQ (privComposeAdaptive' nq1 nq2) := by + rw [NonTopRDNQ] at * + admit -- simp [NonTopRDNQ] at * -- intro α h1 l₁ l₂ h2 -- replace nt1 := nt1 α h1 l₁ l₂ h2 From ca93118301b176cc43c07b921e8b524a31dea69e Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Thu, 6 Jun 2024 10:00:40 -0400 Subject: [PATCH 019/122] refactor --- SampCert/DifferentialPrivacy/Abstract.lean | 19 +- .../Pure/AdaptiveComposition.lean | 134 ++++++++++++ .../DifferentialPrivacy/Pure/Composition.lean | 207 ------------------ SampCert/DifferentialPrivacy/Pure/System.lean | 1 + .../ZeroConcentrated/AdaptiveComposition.lean | 18 +- 5 files changed, 147 insertions(+), 232 deletions(-) create mode 100644 SampCert/DifferentialPrivacy/Pure/AdaptiveComposition.lean diff --git a/SampCert/DifferentialPrivacy/Abstract.lean b/SampCert/DifferentialPrivacy/Abstract.lean index c274c550..8808a184 100644 --- a/SampCert/DifferentialPrivacy/Abstract.lean +++ b/SampCert/DifferentialPrivacy/Abstract.lean @@ -31,7 +31,7 @@ abbrev Mechanism (T U : Type) := List T → SLang U /-- General (value-dependent) composition of mechanisms -/ -def privComposeAdaptive' (nq1 : Mechanism T U) (nq2 : U -> Mechanism T V) (l : List T) : SLang (U × V) := do +def privComposeAdaptive (nq1 : Mechanism T U) (nq2 : U -> Mechanism T V) (l : List T) : SLang (U × V) := do let A <- nq1 l let B <- nq2 A l return (A, B) @@ -45,9 +45,9 @@ Chain rule relating the adaptive composition definitions The joint distribution decomposes into the conditional and marginal (ie, nq1 l) distributions -/ lemma privComposeChainRule (nq1 : Mechanism T U) (nq2 : U -> Mechanism T V) (l : List T) : - ∀ (u : U), ∀ (v : V), privComposeAdaptive' nq1 nq2 l (u, v) = nq1 l u * nq2 u l v := by + ∀ (u : U), ∀ (v : V), privComposeAdaptive nq1 nq2 l (u, v) = nq1 l u * nq2 u l v := by intros u v - simp [privComposeAdaptive'] + simp [privComposeAdaptive] -- How to simplify cases without explicit ite? this is silly have hrw1 : ∀ (a : U), nq1 l a * (∑' (a_1 : V), if u = a ∧ v = a_1 then nq2 a l a_1 else 0) = if (u = a) then (nq1 l a * ∑' (a_1 : V), if u = a ∧ v = a_1 then nq2 a l a_1 else 0) else 0 := by intro a @@ -75,19 +75,6 @@ lemma privComposeChainRule (nq1 : Mechanism T U) (nq2 : U -> Mechanism T V) (l : rw [tsum_ite_eq] exact MulOneClass.mul_one (nq2 u l v) - -/-- -Conditional composition of mechanisms --/ -def privComposeAdaptive (nq1 : Mechanism T U) (nq2 : U -> Mechanism T V) (l : List T) : SLang V := do - let A <- nq1 l - let B <- nq2 A l - return B - - - - - /-- Composition of independent mechanisms -/ diff --git a/SampCert/DifferentialPrivacy/Pure/AdaptiveComposition.lean b/SampCert/DifferentialPrivacy/Pure/AdaptiveComposition.lean new file mode 100644 index 00000000..0d9db6a7 --- /dev/null +++ b/SampCert/DifferentialPrivacy/Pure/AdaptiveComposition.lean @@ -0,0 +1,134 @@ +/- +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, Markus de Medeiros +-/ +import SampCert.DifferentialPrivacy.Abstract +import SampCert.DifferentialPrivacy.Pure.DP +import Mathlib.Data.Set.Defs +import Mathlib.Data.Set.Prod +import Mathlib.Logic.IsEmpty + +noncomputable section + +open Classical Set + +variable [Hu : Nonempty U] + +/-- +Mediant inequality +-/ +lemma tsum_mediant (f g : U -> ENNReal) (hg0 : ∀ u, g u ≠ 0) (hf0 : ∀ u, f u ≠ 0): + (∑' (u : U), f u) / (∑' (u : U), g u) ≤ ⨆ u, f u / g u := by + + -- We need this to discharge side conditions in this proof, but we can get away with + -- some classical reasoning instead + cases (Classical.em (∀ u , g u ≠ ⊤)) + case inr => + rename_i Hk + -- rcases Hk with ⟨ ucont, Hcont ⟩ + have hgtop : ∃ u, g u = ⊤ := by apply (Decidable.not_forall_not.mp Hk) + have sumtop : ∑' (u : U) , g u = ⊤ := by exact ENNReal.tsum_eq_top_of_eq_top hgtop + rw [sumtop] + simp + case inl => + rename_i assumption_g + -- cases (isEmpty_or_nonempty U) + -- · rename_i HU + -- rw [iSup_of_empty] + -- rw [tsum_empty] + -- rw [tsum_empty] + -- simp + -- · rename_i Hu + rcases Hu with ⟨ u0 ⟩ + apply (ENNReal.div_le_iff_le_mul _ _).mpr + · rw [← ENNReal.tsum_mul_left] + apply ENNReal.tsum_le_tsum + intro u + apply (ENNReal.div_le_iff_le_mul _ _).mp + · refine (le_iSup (fun u => HDiv.hDiv (f u) (g u)) u) + · left; apply hg0 + · right + apply ne_of_gt + apply (LT.lt.trans_le ?g1 ?g2) + case g2 => + apply le_iSup + apply u + refine (ENNReal.div_pos (hf0 u) ?g1.hb) + apply assumption_g + · left + apply ne_of_gt + apply (LT.lt.trans_le ?z1 ?z2) + case z2 => + apply ENNReal.le_tsum + apply u0 + exact pos_iff_ne_zero.mpr (hg0 u0) + · right + apply ne_of_gt + apply (LT.lt.trans_le ?z3 ?z4) + case z4 => + apply le_iSup + apply u0 + refine (ENNReal.div_pos (hf0 u0) ?z6) + apply assumption_g + +lemma bounded_quotient (f g : U -> ENNReal) (b : ENNReal) (h_bound : ∀ (u : U), f u / g u ≤ b) (hg0 : ∀ u, g u ≠ 0) (hf0 : ∀ u, f u ≠ 0) : + (∑' (u : U), f u) / (∑' (u : U), g u) ≤ b := by + apply le_trans + · refine (tsum_mediant _ _ hg0 hf0) + · simp + assumption + +namespace SLang + +-- Better proof for Pure DP adaptive composition +theorem PureDP_ComposeAdaptive' (nq1 : List T → SLang U) (nq2 : U -> List T → SLang V) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h1 : PureDP nq1 ((ε₁ : ℝ) / ε₂)) (h2 : ∀ u : U, PureDP (nq2 u) ((ε₃ : ℝ) / ε₄)) : + PureDP (privComposeAdaptive nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by + simp [PureDP] at * + rcases h1 with ⟨h1a, h1nz⟩ + rw [event_eq_singleton] at * + simp [DP_singleton] at * + apply And.intro + · intros l₁ l₂ Hl₁l₂ u v + rw [privComposeChainRule] + rw [privComposeChainRule] + rw [Real.exp_add] + rw [ENNReal.ofReal_mul ?s1] + case s1 => apply Real.exp_nonneg + rw [ENNReal.div_eq_inv_mul] + rw [ENNReal.mul_inv] + · rw [<- mul_assoc] + rw [mul_right_comm] + conv => + lhs + arg 1 + rw [mul_assoc] + rw [mul_right_comm] + rw [← ENNReal.div_eq_inv_mul] + rw [← ENNReal.div_eq_inv_mul] + rcases h2 u with ⟨ h2a'_pre, _ ⟩ + rw [event_eq_singleton] at h2a'_pre + simp [DP_singleton] at h2a'_pre + exact (mul_le_mul' (h1a l₁ l₂ Hl₁l₂ u) (h2a'_pre l₁ l₂ Hl₁l₂ v)) + · left + apply h1nz + · right + rcases (h2 u) with ⟨ _ , h2nz ⟩ + apply h2nz + · simp only [NonZeroNQ] + intros l n + simp only [privComposeAdaptive , bind, pure, bind_pure, bind_apply] + apply ne_of_gt + apply (LT.lt.trans_le ?g1 ?g2) + case g2 => + apply ENNReal.le_tsum + apply n.1 + apply ENNReal.mul_pos + · apply h1nz + · simp + exists n.2 + simp + rcases (h2 n.1) with ⟨ _ , H2nz ⟩ + apply H2nz + +end SLang diff --git a/SampCert/DifferentialPrivacy/Pure/Composition.lean b/SampCert/DifferentialPrivacy/Pure/Composition.lean index a97f3855..f57566e2 100644 --- a/SampCert/DifferentialPrivacy/Pure/Composition.lean +++ b/SampCert/DifferentialPrivacy/Pure/Composition.lean @@ -15,73 +15,6 @@ open Classical Set variable [Hu : Nonempty U] --- Solving the side conditions needs to be done separately depending on if u is inhabited or not -/-- -Mediant inequality --/ -lemma tsum_mediant (f g : U -> ENNReal) (hg0 : ∀ u, g u ≠ 0) (hf0 : ∀ u, f u ≠ 0): - (∑' (u : U), f u) / (∑' (u : U), g u) ≤ ⨆ u, f u / g u := by - - -- We need this to discharge side conditions in this proof, but we can get away with - -- some classical reasoning instead - cases (Classical.em (∀ u , g u ≠ ⊤)) - case inr => - rename_i Hk - -- rcases Hk with ⟨ ucont, Hcont ⟩ - have hgtop : ∃ u, g u = ⊤ := by apply (Decidable.not_forall_not.mp Hk) - have sumtop : ∑' (u : U) , g u = ⊤ := by exact ENNReal.tsum_eq_top_of_eq_top hgtop - rw [sumtop] - simp - case inl => - rename_i assumption_g - -- cases (isEmpty_or_nonempty U) - -- · rename_i HU - -- rw [iSup_of_empty] - -- rw [tsum_empty] - -- rw [tsum_empty] - -- simp - -- · rename_i Hu - rcases Hu with ⟨ u0 ⟩ - apply (ENNReal.div_le_iff_le_mul _ _).mpr - · rw [← ENNReal.tsum_mul_left] - apply ENNReal.tsum_le_tsum - intro u - apply (ENNReal.div_le_iff_le_mul _ _).mp - · refine (le_iSup (fun u => HDiv.hDiv (f u) (g u)) u) - · left; apply hg0 - · right - apply ne_of_gt - apply (LT.lt.trans_le ?g1 ?g2) - case g2 => - apply le_iSup - apply u - refine (ENNReal.div_pos (hf0 u) ?g1.hb) - apply assumption_g - · left - apply ne_of_gt - apply (LT.lt.trans_le ?z1 ?z2) - case z2 => - apply ENNReal.le_tsum - apply u0 - exact pos_iff_ne_zero.mpr (hg0 u0) - · right - apply ne_of_gt - apply (LT.lt.trans_le ?z3 ?z4) - case z4 => - apply le_iSup - apply u0 - refine (ENNReal.div_pos (hf0 u0) ?z6) - apply assumption_g - - -lemma bounded_quotient (f g : U -> ENNReal) (b : ENNReal) (h_bound : ∀ (u : U), f u / g u ≤ b) (hg0 : ∀ u, g u ≠ 0) (hf0 : ∀ u, f u ≠ 0) : - (∑' (u : U), f u) / (∑' (u : U), g u) ≤ b := by - apply le_trans - · refine (tsum_mediant _ _ hg0 hf0) - · simp - assumption - - namespace SLang theorem PureDP_Compose' {nq1 : Mechanism T U} {nq2 : List T → SLang V} {ε₁ ε₂ ε₃ ε₄ : ℕ+} (h1 : PureDP nq1 ((ε₁ : ℝ) / ε₂)) (h2 : PureDP nq2 ((ε₃ : ℝ) / ε₄)) : @@ -138,144 +71,4 @@ theorem PureDP_Compose (nq1 : List T → SLang U) (nq2 : List T → SLang V) (ε . apply PureDP_Compose' hc h'c . apply privCompose_NonZeroNQ h2 h'2 -theorem PureDP_ComposeAdaptive (nq1 : List T → SLang U) (nq2 : U -> List T → SLang V) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h1 : PureDP nq1 ((ε₁ : ℝ) / ε₂)) (h2 : ∀ u : U, PureDP (nq2 u) ((ε₃ : ℝ) / ε₄)) : - PureDP (privComposeAdaptive nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by - simp [PureDP] at * - rcases h1 with ⟨h1a, h1nz⟩ - rw [event_eq_singleton] at * - simp [DP_singleton] at * - apply And.intro - · intros l₁ l₂ neighbours x - replace h1a := h1a l₁ l₂ neighbours - - have h2' : ∀ (u : U), (nq2 u l₁ x) / (nq2 u l₂ x) <= ENNReal.ofReal (((↑↑ε₃ : ℝ ) / ↑↑ε₄).exp) := by - intro u - replace h2 := h2 u - rw [event_eq_singleton] at h2 - simp [DP_singleton] at h2 - rcases h2 with ⟨h2a, _⟩ - exact h2a l₁ l₂ neighbours x - - simp [privComposeAdaptive] - - have h3 : ∀ (a : U), nq1 l₁ a * nq2 a l₁ x / (nq1 l₂ a * nq2 a l₂ x) ≤ ENNReal.ofReal (↑↑ε₁ / ↑↑ε₂ + ↑↑ε₃ / ↑↑ε₄ : ℝ).exp := by - intro a - -- Split the expontntial - rw [Real.exp_add] - rw [ENNReal.ofReal_mul ?g0] - case g0 => apply Real.exp_nonneg - - -- Push around inequalities - rw [ENNReal.div_eq_inv_mul] - rw [ENNReal.mul_inv] - · rw [<- mul_assoc] - rw [mul_right_comm] - conv => - lhs - arg 1 - rw [mul_assoc] - rw [mul_right_comm] - rw [← ENNReal.div_eq_inv_mul] - rw [← ENNReal.div_eq_inv_mul] - exact mul_le_mul' (h1a a) (h2' a) - · left - apply h1nz - · right - rcases (h2 a) with ⟨ _ , h2nz ⟩ - apply h2nz - - -- apply? -- gives me this but it times out. odd. - -- refine (bounded_quotient (fun u => nq1 l₁ u * nq2 u l₁ x) (fun u => nq1 l₂ u * nq2 u l₂ x) (ENNReal.ofReal (↑↑ε₁ / ↑↑ε₂ + ↑↑ε₃ / ↑↑ε₄ : ℝ).exp) h3 ?intro.left.hg0 ?intro.left.hf0) - - -- Put a name to the summands (why is this so hard) - let f := (fun (a : U) => nq1 l₁ a * nq2 a l₁ x) - let g := (fun (a : U) => nq1 l₂ a * nq2 a l₂ x) - have hf : (∑' (a : U), nq1 l₁ a * nq2 a l₁ x) = (∑' (a : U), f a) := by congr - have hg : (∑' (a : U), nq1 l₂ a * nq2 a l₂ x) = (∑' (a : U), g a) := by congr - rw [hf, hg] - - apply bounded_quotient - apply h3 - all_goals (intro u; rcases (h2 u) with ⟨ _ , h2nz ⟩) - all_goals (simp only [f, g]) - · exact mul_ne_zero (h1nz l₂ u) (h2nz l₂ x) - · exact mul_ne_zero (h1nz l₁ u) (h2nz l₁ x) - - · -- Composition is nonzero at all elements - simp only [NonZeroNQ] - intros l n - - simp only [privComposeAdaptive, bind, pure, bind_pure, bind_apply] - - -- cases (isEmpty_or_nonempty U) - -- · rename_i Hu_empty - -- exfalso - -- -- Because ``SLang U`` values are functions out of U, we can't get a contradiction out of - -- -- "run nq1 to get a value of type U" like we would in an operational approach. - -- -- - -- -- We should restrict U to be nonempty. - -- sorry - -- · - - rcases Hu with ⟨ u0 ⟩ - rcases h2 u0 with ⟨ _ , h2nz ⟩ - apply ne_of_gt - apply (LT.lt.trans_le ?g1 ?g2) - case g2 => - apply ENNReal.le_tsum - apply u0 - apply ENNReal.mul_pos - · apply h1nz - · apply h2nz - --- Better proof for Pure DP adaptive composition -theorem PureDP_ComposeAdaptive' (nq1 : List T → SLang U) (nq2 : U -> List T → SLang V) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h1 : PureDP nq1 ((ε₁ : ℝ) / ε₂)) (h2 : ∀ u : U, PureDP (nq2 u) ((ε₃ : ℝ) / ε₄)) : - PureDP (privComposeAdaptive' nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by - simp [PureDP] at * - rcases h1 with ⟨h1a, h1nz⟩ - rw [event_eq_singleton] at * - simp [DP_singleton] at * - apply And.intro - · intros l₁ l₂ Hl₁l₂ u v - rw [privComposeChainRule] - rw [privComposeChainRule] - rw [Real.exp_add] - rw [ENNReal.ofReal_mul ?s1] - case s1 => apply Real.exp_nonneg - rw [ENNReal.div_eq_inv_mul] - rw [ENNReal.mul_inv] - · rw [<- mul_assoc] - rw [mul_right_comm] - conv => - lhs - arg 1 - rw [mul_assoc] - rw [mul_right_comm] - rw [← ENNReal.div_eq_inv_mul] - rw [← ENNReal.div_eq_inv_mul] - rcases h2 u with ⟨ h2a'_pre, _ ⟩ - rw [event_eq_singleton] at h2a'_pre - simp [DP_singleton] at h2a'_pre - exact (mul_le_mul' (h1a l₁ l₂ Hl₁l₂ u) (h2a'_pre l₁ l₂ Hl₁l₂ v)) - · left - apply h1nz - · right - rcases (h2 u) with ⟨ _ , h2nz ⟩ - apply h2nz - · simp only [NonZeroNQ] - intros l n - simp only [privComposeAdaptive', bind, pure, bind_pure, bind_apply] - apply ne_of_gt - apply (LT.lt.trans_le ?g1 ?g2) - case g2 => - apply ENNReal.le_tsum - apply n.1 - apply ENNReal.mul_pos - · apply h1nz - · simp - exists n.2 - simp - rcases (h2 n.1) with ⟨ _ , H2nz ⟩ - apply H2nz - end SLang diff --git a/SampCert/DifferentialPrivacy/Pure/System.lean b/SampCert/DifferentialPrivacy/Pure/System.lean index 0106fc7e..fc3a415b 100644 --- a/SampCert/DifferentialPrivacy/Pure/System.lean +++ b/SampCert/DifferentialPrivacy/Pure/System.lean @@ -7,6 +7,7 @@ import SampCert.DifferentialPrivacy.Abstract import SampCert.DifferentialPrivacy.Pure.DP import SampCert.DifferentialPrivacy.Pure.Mechanism.Basic import SampCert.DifferentialPrivacy.Pure.Composition +import SampCert.DifferentialPrivacy.Pure.AdaptiveComposition import SampCert.DifferentialPrivacy.Pure.Postprocessing namespace SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean index 53ac9de2..3bfacf64 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean @@ -27,9 +27,9 @@ variable [HV : Inhabited V] Bound on Renyi divergence on adaptively composed queries -/ lemma primComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (α : ℝ) (Hα : 1 < α) : - RenyiDivergence (privComposeAdaptive' nq1 nq2 l₁) (privComposeAdaptive' nq1 nq2 l₂) α ≤ RenyiDivergence (nq1 l₁) (nq1 l₂) α + ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α := by + RenyiDivergence (privComposeAdaptive nq1 nq2 l₁) (privComposeAdaptive nq1 nq2 l₂) α ≤ RenyiDivergence (nq1 l₁) (nq1 l₂) α + ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α := by apply (RenyiDivergence_mono_sum _ _ α Hα) - rw [RenyiDivergence_exp (privComposeAdaptive' nq1 nq2 l₁) (privComposeAdaptive' nq1 nq2 l₂) Hα] + rw [RenyiDivergence_exp (privComposeAdaptive nq1 nq2 l₁) (privComposeAdaptive nq1 nq2 l₂) Hα] rw [left_distrib] rw [Real.exp_add] @@ -111,7 +111,7 @@ lemma primComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> Lis Adaptively Composed queries satisfy zCDP Renyi divergence bound. -/ theorem privComposeAdaptive_zCDPBound {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} {ε₁ ε₂ ε₃ ε₄ : ℕ+} (h1 : zCDPBound nq1 ((ε₁ : ℝ) / ε₂)) (h2 : ∀ u, zCDPBound (nq2 u) ((ε₃ : ℝ) / ε₄)) (nn1 : NonZeroNQ nq1) (nn2 : ∀ u, NonZeroNQ (nq2 u)) (nt1 : NonTopRDNQ nq1) (nt2 : ∀ u, NonTopRDNQ (nq2 u)) (nts1 : NonTopNQ nq1) (nts2 : ∀ u, NonTopNQ (nq2 u)) : - zCDPBound (privComposeAdaptive' nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by + zCDPBound (privComposeAdaptive nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by rw [zCDPBound] intro α Hα l₁ l₂ Hneighbours -- Loose step @@ -132,16 +132,16 @@ theorem privComposeAdaptive_zCDPBound {nq1 : List T → SLang U} {nq2 : U -> Lis Adaptive composed query distribution is nowhere zero -/ theorem privComposeAdaptive_NonZeroNQ {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (nt1 : NonZeroNQ nq1) (nt2 : ∀ u, NonZeroNQ (nq2 u)) : - NonZeroNQ (privComposeAdaptive' nq1 nq2) := by + NonZeroNQ (privComposeAdaptive nq1 nq2) := by simp [NonZeroNQ] at * - simp [privComposeAdaptive'] + simp [privComposeAdaptive] aesop /-- All outputs of a adaptive composed query have finite probability. -/ theorem privComposeAdaptive_NonTopNQ {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (nt1 : NonTopNQ nq1) (nt2 : ∀ u, NonTopNQ (nq2 u)) : - NonTopNQ (privComposeAdaptive' nq1 nq2) := by + NonTopNQ (privComposeAdaptive nq1 nq2) := by simp [NonTopNQ] at * intros l u v rw [privComposeChainRule] @@ -153,7 +153,7 @@ theorem privComposeAdaptive_NonTopNQ {nq1 : List T → SLang U} {nq2 : U -> List Adaptive composed query is a proper distribution -/ theorem privComposeAdaptive_NonTopSum {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (nt1 : NonTopSum nq1) (nt2 : ∀ u, NonTopSum (nq2 u)) : - NonTopSum (privComposeAdaptive' nq1 nq2) := by + NonTopSum (privComposeAdaptive nq1 nq2) := by rw [NonTopSum] at * -- Chain rule won't help here admit @@ -163,7 +163,7 @@ theorem privComposeAdaptive_NonTopSum {nq1 : List T → SLang U} {nq2 : U -> Lis Renyi divergence beteeen adaptive composed queries on neighbours are finite. -/ theorem privComposeAdaptive_NonTopRDNQ {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (nt1 : NonTopRDNQ nq1) (nt2 : ∀ u, NonTopRDNQ (nq2 u)) (nn1 : NonTopNQ nq1) (nn2 : ∀ u, NonTopNQ (nq2 u)) : - NonTopRDNQ (privComposeAdaptive' nq1 nq2) := by + NonTopRDNQ (privComposeAdaptive nq1 nq2) := by rw [NonTopRDNQ] at * admit -- simp [NonTopRDNQ] at * @@ -231,7 +231,7 @@ theorem privComposeAdaptive_NonTopRDNQ {nq1 : List T → SLang U} {nq2 : U -> Li ``privComposeAdaptive`` satisfies zCDP -/ theorem privComposeAdaptive_zCDP (nq1 : List T → SLang U) (nq2 : U -> List T → SLang V) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h : zCDP nq1 ((ε₁ : ℝ) / ε₂)) (h' : ∀ u, zCDP (nq2 u) ((ε₃ : ℝ) / ε₄)) : - zCDP (privComposeAdaptive' nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by + zCDP (privComposeAdaptive nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by simp [zCDP] at * repeat any_goals constructor · apply privComposeAdaptive_zCDPBound <;> aesop From db375c3769f8b93cca7336390f684aba733a033d Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Thu, 6 Jun 2024 11:59:28 -0400 Subject: [PATCH 020/122] close some side conditions --- .../DifferentialPrivacy/RenyiDivergence.lean | 18 ++++------ .../ZeroConcentrated/AdaptiveComposition.lean | 34 +++++++++++++++++-- 2 files changed, 38 insertions(+), 14 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index 77fa417f..8ba72971 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -53,21 +53,17 @@ theorem RenyiDivergenceExpectation' (p q : T → ENNReal) {α : ℝ} (h : 1 < α . apply h1 x . apply h2 x + + + /-- The Renyi divergence is monotonic in the value of its sum. -/ -lemma RenyiDivergence_mono_sum (x y : ℝ) (α : ℝ) (h : 1 < α) : Real.exp ((α - 1) * x) ≤ Real.exp ((α - 1) * y) -> (x ≤ y) := by +lemma RenyiDivergence_mono_sum (x y : ℝ) (α : ℝ) (h : 1 < α) : (Real.exp ((α - 1) * x)) ≤ (Real.exp ((α - 1) * y)) -> (x ≤ y) := by intro H - have H1 : ((α - 1) * x) ≤ ((α - 1) * y) := exp_le_exp.mp H - -- refine (mul_le_mul_iff_right ?a).mp ?a - -- apply? - -- refine ((@ENNReal.mul_le_mul_left x y _ _ ).mp _) - -- apply mul_lt_mul_left_of_neg - -- refine (mul_le_mul_left ?a0).mp ?a -- ????? - - -- How annoying - - sorry + apply le_of_mul_le_mul_left + · exact exp_le_exp.mp H + · linarith lemma RenyiDivergence_exp (p q : T → ENNReal) {α : ℝ} (h : 1 < α) : Real.exp ((α - 1) * RenyiDivergence p q α) = (∑' x : T, (p x)^α * (q x)^(1 - α)).toReal := by diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean index 3bfacf64..2325f8ba 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean @@ -23,18 +23,30 @@ variable [HU : Inhabited U] variable [HV : Inhabited V] +-- Maybe would be better as ENNReal +lemma iSup_smul_l (a : ℝ) (f : U -> ℝ) : a * ⨆ u, f u = ⨆ u, a * f u := by + rw [iSup, iSup] + sorry + +lemma iSup_exp (f : U -> ℝ) : ⨆ u, rexp (f u) = rexp (⨆ u, f u) := by + sorry + +-- rexp ((α - 1) * ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) = +-- ⨆ u, rexp ((α - 1) * RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) + + /-- Bound on Renyi divergence on adaptively composed queries -/ lemma primComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (α : ℝ) (Hα : 1 < α) : - RenyiDivergence (privComposeAdaptive nq1 nq2 l₁) (privComposeAdaptive nq1 nq2 l₂) α ≤ RenyiDivergence (nq1 l₁) (nq1 l₂) α + ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α := by + RenyiDivergence (privComposeAdaptive nq1 nq2 l₁) (privComposeAdaptive nq1 nq2 l₂) α ≤ RenyiDivergence (nq1 l₁) (nq1 l₂) α + ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α := by apply (RenyiDivergence_mono_sum _ _ α Hα) rw [RenyiDivergence_exp (privComposeAdaptive nq1 nq2 l₁) (privComposeAdaptive nq1 nq2 l₂) Hα] rw [left_distrib] rw [Real.exp_add] have hmono_1 : rexp ((α - 1) * ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) = ⨆ u, rexp ((α - 1) * RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) := by - sorry + rw [iSup_smul_l, iSup_exp] rw [hmono_1] clear hmono_1 rw [RenyiDivergence_exp (nq1 l₁) (nq1 l₂) Hα] @@ -118,7 +130,23 @@ theorem privComposeAdaptive_zCDPBound {nq1 : List T → SLang U} {nq2 : U -> Lis apply (@LE.le.trans _ _ _ (1/2 * (↑↑ε₁ / ↑↑ε₂)^2 * α + 1/2 * (↑↑ε₃ / ↑↑ε₄)^2 * α) _ _ ?case_sq) case case_sq => -- Binomial bound - sorry + rw [add_sq] + rw [<- right_distrib] + apply (mul_le_mul_of_nonneg_right _ ?goal1) + case goal1 => linarith + rw [<- left_distrib] + apply (mul_le_mul_of_nonneg_left _ ?goal1) + case goal1 => linarith + apply add_le_add_right + have hrw : (↑↑ε₁ / ↑↑ε₂ : ℝ) ^ 2 = (↑↑ε₁ / ↑↑ε₂) ^ 2 + 0 := by linarith + conv => + lhs + rw [hrw] + clear hrw + apply add_le_add_left + have h : 0 <= (↑↑ε₁ / ↑↑ε₂) * (↑↑ε₃ / ↑↑ε₄ : ℝ) := by + apply mul_nonneg <;> apply div_nonneg <;> linarith + linarith -- Rewrite the upper bounds in terms of Renyi divergences of nq1/nq2 rw [zCDPBound] at h1 have marginal_ub := h1 α Hα l₁ l₂ Hneighbours From a452caf0c5e772b827f24844ab63d7204b14451b Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Thu, 6 Jun 2024 13:55:34 -0400 Subject: [PATCH 021/122] close some side conditions --- .../ZeroConcentrated/AdaptiveComposition.lean | 86 ++++++++++++------- 1 file changed, 56 insertions(+), 30 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean index 2325f8ba..55b3f4a2 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean @@ -23,7 +23,7 @@ variable [HU : Inhabited U] variable [HV : Inhabited V] --- Maybe would be better as ENNReal +-- Maybe would be better as ENNReal? lemma iSup_smul_l (a : ℝ) (f : U -> ℝ) : a * ⨆ u, f u = ⨆ u, a * f u := by rw [iSup, iSup] sorry @@ -35,10 +35,13 @@ lemma iSup_exp (f : U -> ℝ) : ⨆ u, rexp (f u) = rexp (⨆ u, f u) := by -- ⨆ u, rexp ((α - 1) * RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) +lemma exp_non_top : ∀ (z : ENNReal) (β : ℝ), z ≠ ⊤ -> z ^ β ≠ ⊤ := by + sorry + /-- Bound on Renyi divergence on adaptively composed queries -/ -lemma primComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (α : ℝ) (Hα : 1 < α) : +lemma primComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (α : ℝ) (Hα : 1 < α) (HNT1 : NonTopNQ nq1) (HNTRDNQ2 : ∀ u, NonTopRDNQ (nq2 u)) (HN : Neighbour l₁ l₂): RenyiDivergence (privComposeAdaptive nq1 nq2 l₁) (privComposeAdaptive nq1 nq2 l₂) α ≤ RenyiDivergence (nq1 l₁) (nq1 l₂) α + ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α := by apply (RenyiDivergence_mono_sum _ _ α Hα) rw [RenyiDivergence_exp (privComposeAdaptive nq1 nq2 l₁) (privComposeAdaptive nq1 nq2 l₂) Hα] @@ -59,51 +62,75 @@ lemma primComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> Lis rw [mul_comm] rw [← ENNReal.tsum_mul_right] - apply (@LE.le.trans _ _ _ ((∑' (i : U), nq1 l₁ i ^ α * nq1 l₂ i ^ (1 - α) * ENNReal.ofReal (rexp ((α - 1) * RenyiDivergence (nq2 i l₁) (nq2 i l₂) α))).toReal) _ _ ?goal2) + apply (toReal_mono' _ ?goal2) case goal2 => - -- Can I do this without summability? - apply toReal_mono' - · apply tsum_le_tsum - · intro i - refine (ENNReal.mul_le_mul_left ?h.h.h0 ?h.h.hinf).mpr ?h.h.a - · sorry - · sorry - · apply ENNReal.ofReal_le_ofReal - sorry - · sorry - · sorry - · sorry + intro H + -- ?? + sorry + + + apply (@LE.le.trans _ _ _ ((∑' (i : U), nq1 l₁ i ^ α * nq1 l₂ i ^ (1 - α) * ENNReal.ofReal (rexp ((α - 1) * RenyiDivergence (nq2 i l₁) (nq2 i l₂) α)))) _ _ ?goal2) + case goal2 => + apply (tsum_le_tsum _ ENNReal.summable ENNReal.summable) + intro i + refine (ENNReal.mul_le_mul_left ?h.h.h0 ?h.h.hinf).mpr ?h.h.a + · -- Easy + sorry + · apply ENNReal.mul_ne_top + · apply exp_non_top + apply HNT1 + · apply exp_non_top + apply HNT1 + · apply ENNReal.ofReal_le_ofReal -- I don't see any way to make use of the ENNReal here + -- Can I cancel out the usage of hmono_1? (this step undoes it) + sorry + -- After this point the argument is tight apply Eq.le conv => rhs congr - congr intro i rw [RenyiDivergence_exp (nq2 i l₁) (nq2 i l₂) Hα] conv => lhs congr - congr intro rw [privComposeChainRule] rw [privComposeChainRule] - have test : ∀ i, ENNReal.ofReal (∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α)).toReal = ∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α) := by - intro i - apply ofReal_toReal - -- Need a summability thing here... try to avoid this lemma if possible - sorry conv => rhs congr - congr - intro - rw [test] - rw [← ENNReal.tsum_mul_left] - clear test + intro x + rw [<- (@ENNReal.ofReal_toReal (nq1 l₁ x ^ α * nq1 l₂ x ^ (1 - α)) ?goal2)] + · rw [<- ENNReal.ofReal_mul] + · rw [<- ENNReal.toReal_mul] + rw [(@ENNReal.ofReal_toReal (nq1 l₁ x ^ α * nq1 l₂ x ^ (1 - α) * ∑' (x_1 : V), nq2 x l₁ x_1 ^ α * nq2 x l₂ x_1 ^ (1 - α)) ?goal4)] + rfl + apply ENNReal.mul_ne_top + · apply ENNReal.mul_ne_top + · apply exp_non_top + apply HNT1 + · apply exp_non_top + apply HNT1 + · apply HNTRDNQ2 + apply Hα + apply HN + · apply ENNReal.toReal_nonneg + · apply ENNReal.mul_ne_top + · apply exp_non_top + apply HNT1 + · apply exp_non_top + apply HNT1 + + conv => + rhs + arg 1 + intro x + rw [<- ENNReal.tsum_mul_left] rw [<- ENNReal.tsum_prod] congr @@ -112,8 +139,7 @@ lemma primComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> Lis rcases p with ⟨ u , v ⟩ simp rw [ENNReal.mul_rpow_of_nonneg _ _ ?sc1] - case sc1 => - sorry + case sc1 => linarith rw [ENNReal.mul_rpow_of_nonneg _ _ ?sc2] case sc2 => sorry @@ -154,7 +180,7 @@ theorem privComposeAdaptive_zCDPBound {nq1 : List T → SLang U} {nq2 : U -> Lis ciSup_le fun x => h2 x α Hα l₁ l₂ Hneighbours apply (@LE.le.trans _ _ _ (RenyiDivergence (nq1 l₁) (nq1 l₂) α + ⨆ (u : U), RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) _ _ ?case_alg) case case_alg => linarith - apply (primComposeAdaptive_renyi_bound _ Hα) + apply (primComposeAdaptive_renyi_bound _ Hα _ _) <;> aesop /-- Adaptive composed query distribution is nowhere zero From 44813a18f3b794e5d1987f143bf2cc2524a7a5e8 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Thu, 6 Jun 2024 15:53:06 -0400 Subject: [PATCH 022/122] close some side conditions --- .../ZeroConcentrated/AdaptiveComposition.lean | 85 +++++++++++++------ 1 file changed, 61 insertions(+), 24 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean index 55b3f4a2..b7a9c9cf 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean @@ -22,37 +22,62 @@ variable { T U V : Type } variable [HU : Inhabited U] variable [HV : Inhabited V] - +-- set_option pp.all true -- Maybe would be better as ENNReal? -lemma iSup_smul_l (a : ℝ) (f : U -> ℝ) : a * ⨆ u, f u = ⨆ u, a * f u := by - rw [iSup, iSup] - sorry +lemma iSup_smul_l (a : ℝ) (Ha : 0 <= a) (f : U -> ℝ) : a * ⨆ u, f u = ⨆ u, a * f u := by + refine (mul_iSup_of_nonneg ?ha fun i => f i) + apply Ha + +-- #check @sSupHomClass (ℝ -> ℝ) ℝ ℝ _ _ funlike_inst +-- #check map_sSup + -lemma iSup_exp (f : U -> ℝ) : ⨆ u, rexp (f u) = rexp (⨆ u, f u) := by - sorry +def funlike_inst : FunLike (ℝ → ℝ) ℝ ℝ := by + constructor + case coe => + intro f + apply f + case coe_injective' => + exact fun ⦃a₁ a₂⦄ a => a --- rexp ((α - 1) * ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) = --- ⨆ u, rexp ((α - 1) * RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) +-- set_option pp.all true +lemma iSup_exp (f : U -> ℝ) : ⨆ u, Real.exp (f u) = Real.exp (⨆ u, f u) := by + symm + apply (@map_iSup (ℝ -> ℝ) ℝ ℝ U funlike_inst _ _ ?SHC rexp f) + · -- sSupHomClass + -- refine { map_sSup := ?SHC.map_sSup : @sSupHomClass (ℝ -> ℝ) ℝ ℝ _ _ _ funlike_inst} + -- May be circular + sorry +lemma exp_non_top : ∀ (z : ENNReal) (β : ℝ), z ≠ 0 -> z ≠ ⊤ -> z ^ β ≠ ⊤ := by + intro z β Hz0 HzT + intro W + have h : z = 0 ∧ β < 0 ∨ z = ⊤ ∧ 0 < β := by + apply rpow_eq_top_iff.mp + apply W + cases h + · aesop + · aesop -lemma exp_non_top : ∀ (z : ENNReal) (β : ℝ), z ≠ ⊤ -> z ^ β ≠ ⊤ := by - sorry /-- Bound on Renyi divergence on adaptively composed queries -/ -lemma primComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (α : ℝ) (Hα : 1 < α) (HNT1 : NonTopNQ nq1) (HNTRDNQ2 : ∀ u, NonTopRDNQ (nq2 u)) (HN : Neighbour l₁ l₂): +lemma primComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (α : ℝ) (Hα : 1 < α) (HNT1 : NonTopNQ nq1) (HNTRDNQ2 : ∀ u, NonTopRDNQ (nq2 u)) (HN : Neighbour l₁ l₂) (HNZ1 : NonZeroNQ nq1) (HNZ2 : ∀ u, NonZeroNQ (nq2 u)): RenyiDivergence (privComposeAdaptive nq1 nq2 l₁) (privComposeAdaptive nq1 nq2 l₂) α ≤ RenyiDivergence (nq1 l₁) (nq1 l₂) α + ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α := by apply (RenyiDivergence_mono_sum _ _ α Hα) rw [RenyiDivergence_exp (privComposeAdaptive nq1 nq2 l₁) (privComposeAdaptive nq1 nq2 l₂) Hα] rw [left_distrib] rw [Real.exp_add] + rw [RenyiDivergence_exp (nq1 l₁) (nq1 l₂) Hα] + have hmono_1 : rexp ((α - 1) * ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) = ⨆ u, rexp ((α - 1) * RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) := by rw [iSup_smul_l, iSup_exp] + linarith rw [hmono_1] clear hmono_1 - rw [RenyiDivergence_exp (nq1 l₁) (nq1 l₂) Hα] + rw [mul_comm] rw [<- (ENNReal.toReal_ofReal_mul _ _ ?h)] case h => @@ -74,17 +99,25 @@ lemma primComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> Lis apply (tsum_le_tsum _ ENNReal.summable ENNReal.summable) intro i refine (ENNReal.mul_le_mul_left ?h.h.h0 ?h.h.hinf).mpr ?h.h.a - · -- Easy - sorry + · apply mul_ne_zero_iff.mpr + apply And.intro + · sorry + · sorry · apply ENNReal.mul_ne_top · apply exp_non_top - apply HNT1 + · apply HNZ1 + · apply HNT1 + linarith · apply exp_non_top - apply HNT1 - · apply ENNReal.ofReal_le_ofReal -- I don't see any way to make use of the ENNReal here - -- Can I cancel out the usage of hmono_1? (this step undoes it) - sorry - + · apply HNZ1 + · apply HNT1 + linarith + · apply ENNReal.ofReal_le_ofReal + rw [iSup_exp] + rw [<- iSup_smul_l] + · -- Should be easy + sorry + · linarith -- After this point the argument is tight apply Eq.le @@ -113,18 +146,22 @@ lemma primComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> Lis apply ENNReal.mul_ne_top · apply ENNReal.mul_ne_top · apply exp_non_top - apply HNT1 + · apply HNZ1 + · apply HNT1 · apply exp_non_top - apply HNT1 + · apply HNZ1 + · apply HNT1 · apply HNTRDNQ2 apply Hα apply HN · apply ENNReal.toReal_nonneg · apply ENNReal.mul_ne_top · apply exp_non_top - apply HNT1 + · apply HNZ1 + · apply HNT1 · apply exp_non_top - apply HNT1 + · apply HNZ1 + · apply HNT1 conv => rhs From 4b40cc1f53b702a39a2e860802bee47bd98dda3b Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Thu, 6 Jun 2024 17:04:00 -0400 Subject: [PATCH 023/122] inline RentyDivergence_exp and fix build --- .../DifferentialPrivacy/RenyiDivergence.lean | 16 ------ .../ZeroConcentrated/AdaptiveComposition.lean | 57 +++++++++++++------ 2 files changed, 41 insertions(+), 32 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index 8ba72971..87f6ef1e 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -64,19 +64,3 @@ lemma RenyiDivergence_mono_sum (x y : ℝ) (α : ℝ) (h : 1 < α) : (Real.exp ( apply le_of_mul_le_mul_left · exact exp_le_exp.mp H · linarith - -lemma RenyiDivergence_exp (p q : T → ENNReal) {α : ℝ} (h : 1 < α) : - Real.exp ((α - 1) * RenyiDivergence p q α) = (∑' x : T, (p x)^α * (q x)^(1 - α)).toReal := by - simp only [RenyiDivergence] - rw [<- mul_assoc] - have test : (α - 1) * (α - 1)⁻¹ = 1 := by - refine mul_inv_cancel ?h - linarith - rw [test] - clear test - simp - rw [Real.exp_log] - apply ENNReal.toReal_pos_iff.mpr - apply And.intro - · sorry - · sorry diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean index b7a9c9cf..f69bc091 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean @@ -40,14 +40,15 @@ def funlike_inst : FunLike (ℝ → ℝ) ℝ ℝ := by case coe_injective' => exact fun ⦃a₁ a₂⦄ a => a + -- set_option pp.all true lemma iSup_exp (f : U -> ℝ) : ⨆ u, Real.exp (f u) = Real.exp (⨆ u, f u) := by - symm - apply (@map_iSup (ℝ -> ℝ) ℝ ℝ U funlike_inst _ _ ?SHC rexp f) - · -- sSupHomClass - -- refine { map_sSup := ?SHC.map_sSup : @sSupHomClass (ℝ -> ℝ) ℝ ℝ _ _ _ funlike_inst} - -- May be circular - sorry + sorry + -- apply (@map_iSup (ℝ -> ℝ) ℝ ℝ U funlike_inst _ _ ?SHC rexp f) + -- · -- sSupHomClass + -- -- refine { map_sSup := ?SHC.map_sSup : @sSupHomClass (ℝ -> ℝ) ℝ ℝ _ _ _ funlike_inst} + -- -- May be circular + -- sorry lemma exp_non_top : ∀ (z : ENNReal) (β : ℝ), z ≠ 0 -> z ≠ ⊤ -> z ^ β ≠ ⊤ := by intro z β Hz0 HzT @@ -60,17 +61,36 @@ lemma exp_non_top : ∀ (z : ENNReal) (β : ℝ), z ≠ 0 -> z ≠ ⊤ -> z ^ β · aesop +lemma RenyiDivergence_exp (p q : SLang T) {α : ℝ} (h : 1 < α) (H1 : 0 < ∑' (x : T), p x ^ α * q x ^ (1 - α)) (H2 : ∑' (x : T), p x ^ α * q x ^ (1 - α) < ⊤): + Real.exp ((α - 1) * RenyiDivergence p q α) = (∑' x : T, (p x)^α * (q x)^(1 - α)).toReal := by + simp only [RenyiDivergence] + rw [<- mul_assoc] + have test : (α - 1) * (α - 1)⁻¹ = 1 := by + refine mul_inv_cancel ?h + linarith + rw [test] + clear test + simp + rw [Real.exp_log] + apply ENNReal.toReal_pos_iff.mpr + apply And.intro H1 H2 + + /-- Bound on Renyi divergence on adaptively composed queries -/ -lemma primComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (α : ℝ) (Hα : 1 < α) (HNT1 : NonTopNQ nq1) (HNTRDNQ2 : ∀ u, NonTopRDNQ (nq2 u)) (HN : Neighbour l₁ l₂) (HNZ1 : NonZeroNQ nq1) (HNZ2 : ∀ u, NonZeroNQ (nq2 u)): +lemma privComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (α : ℝ) (Hα : 1 < α) (HNT1 : NonTopNQ nq1) (HNTRDNQ2 : ∀ u, NonTopRDNQ (nq2 u)) (HN : Neighbour l₁ l₂) (HNZ1 : NonZeroNQ nq1) (HNZ2 : ∀ u, NonZeroNQ (nq2 u)) : RenyiDivergence (privComposeAdaptive nq1 nq2 l₁) (privComposeAdaptive nq1 nq2 l₂) α ≤ RenyiDivergence (nq1 l₁) (nq1 l₂) α + ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α := by apply (RenyiDivergence_mono_sum _ _ α Hα) - rw [RenyiDivergence_exp (privComposeAdaptive nq1 nq2 l₁) (privComposeAdaptive nq1 nq2 l₂) Hα] + rw [RenyiDivergence_exp (privComposeAdaptive nq1 nq2 l₁) (privComposeAdaptive nq1 nq2 l₂) Hα ?H1 ?H2] + case H1 => sorry + case H2 => sorry rw [left_distrib] rw [Real.exp_add] - rw [RenyiDivergence_exp (nq1 l₁) (nq1 l₂) Hα] + rw [RenyiDivergence_exp (nq1 l₁) (nq1 l₂) Hα ?H1 ?H2] + case H1 => sorry + case H2 => sorry have hmono_1 : rexp ((α - 1) * ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) = ⨆ u, rexp ((α - 1) * RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) := by rw [iSup_smul_l, iSup_exp] @@ -107,11 +127,9 @@ lemma primComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> Lis · apply exp_non_top · apply HNZ1 · apply HNT1 - linarith · apply exp_non_top · apply HNZ1 · apply HNT1 - linarith · apply ENNReal.ofReal_le_ofReal rw [iSup_exp] rw [<- iSup_smul_l] @@ -119,6 +137,10 @@ lemma primComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> Lis sorry · linarith + have GH1 : ∀ i, 0 < ∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α) := sorry + have GH2 : ∀ i, ∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α) < ⊤ := sorry + + -- After this point the argument is tight apply Eq.le conv => @@ -126,6 +148,9 @@ lemma primComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> Lis congr intro i rw [RenyiDivergence_exp (nq2 i l₁) (nq2 i l₂) Hα] + rfl + · apply GH1 + · apply GH2 conv => lhs @@ -177,10 +202,10 @@ lemma primComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> Lis simp rw [ENNReal.mul_rpow_of_nonneg _ _ ?sc1] case sc1 => linarith - rw [ENNReal.mul_rpow_of_nonneg _ _ ?sc2] - case sc2 => - sorry - rw [mul_mul_mul_comm] + rw [mul_rpow_of_ne_zero] + · exact mul_mul_mul_comm (nq1 l₁ u ^ α) (nq2 u l₁ v ^ α) (nq1 l₂ u ^ (1 - α)) (nq2 u l₂ v ^ (1 - α)) + · apply HNZ1 + · apply HNZ2 /-- Adaptively Composed queries satisfy zCDP Renyi divergence bound. @@ -217,7 +242,7 @@ theorem privComposeAdaptive_zCDPBound {nq1 : List T → SLang U} {nq2 : U -> Lis ciSup_le fun x => h2 x α Hα l₁ l₂ Hneighbours apply (@LE.le.trans _ _ _ (RenyiDivergence (nq1 l₁) (nq1 l₂) α + ⨆ (u : U), RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) _ _ ?case_alg) case case_alg => linarith - apply (primComposeAdaptive_renyi_bound _ Hα _ _) <;> aesop + apply (privComposeAdaptive_renyi_bound _ Hα _ _) <;> aesop /-- Adaptive composed query distribution is nowhere zero From 546c5654a2ae1064201d0048c17aa928db81be85 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Fri, 7 Jun 2024 10:28:16 -0400 Subject: [PATCH 024/122] more side conditions --- SampCert/DifferentialPrivacy/Abstract.lean | 7 + .../ZeroConcentrated/AdaptiveComposition.lean | 177 +++++++++++------- 2 files changed, 121 insertions(+), 63 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Abstract.lean b/SampCert/DifferentialPrivacy/Abstract.lean index 8808a184..39a5bea2 100644 --- a/SampCert/DifferentialPrivacy/Abstract.lean +++ b/SampCert/DifferentialPrivacy/Abstract.lean @@ -155,6 +155,10 @@ lemma bind_bind_indep (p : Mechanism T U) (q : Mechanism T V) (h : U → V → S subst h4 contradiction + + + + lemma compose_sum_rw (nq1 : List T → SLang U) (nq2 : List T → SLang V) (b : U) (c : V) (l : List T) : (∑' (a : U), nq1 l a * ∑' (a_1 : V), if b = a ∧ c = a_1 then nq2 l a_1 else 0) = nq1 l b * nq2 l c := by have A : ∀ a : U, ∀ b : U, (∑' (a_1 : V), if b = a ∧ c = a_1 then nq2 l a_1 else 0) = if b = a then (∑' (a_1 : V), if c = a_1 then nq2 l a_1 else 0) else 0 := by @@ -211,6 +215,9 @@ lemma compose_sum_rw (nq1 : List T → SLang U) (nq2 : List T → SLang V) (b : rw [C] simp +lemma compose_sum_rw_adaptive (nq1 : List T → SLang U) (nq2 : U -> List T → SLang V) (b : U) (c : V) (l : List T) : + (∑' (a : U), nq1 l a * ∑' (a_1 : V), if x = a ∧ y = a_1 then nq2 a l a_1 else 0) = nq1 l x * nq2 x l y := sorry + /-- Composed queries are normalizable. -/ diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean index f69bc091..868f54df 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean @@ -76,6 +76,16 @@ lemma RenyiDivergence_exp (p q : SLang T) {α : ℝ} (h : 1 < α) (H1 : 0 < ∑' apply And.intro H1 H2 +/-- +Adaptive composed query distribution is nowhere zero +-/ +theorem privComposeAdaptive_NonZeroNQ {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (nt1 : NonZeroNQ nq1) (nt2 : ∀ u, NonZeroNQ (nq2 u)) : + NonZeroNQ (privComposeAdaptive nq1 nq2) := by + simp [NonZeroNQ] at * + simp [privComposeAdaptive] + aesop + + /-- Bound on Renyi divergence on adaptively composed queries -/ @@ -83,7 +93,16 @@ lemma privComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> Lis RenyiDivergence (privComposeAdaptive nq1 nq2 l₁) (privComposeAdaptive nq1 nq2 l₂) α ≤ RenyiDivergence (nq1 l₁) (nq1 l₂) α + ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α := by apply (RenyiDivergence_mono_sum _ _ α Hα) rw [RenyiDivergence_exp (privComposeAdaptive nq1 nq2 l₁) (privComposeAdaptive nq1 nq2 l₂) Hα ?H1 ?H2] - case H1 => sorry + case H1 => + rcases HV with ⟨ v0 ⟩ + rcases HU with ⟨ u0 ⟩ + have Hle : (privComposeAdaptive nq1 nq2 l₁ (u0, v0) ^ α * privComposeAdaptive nq1 nq2 l₂ (u0, v0) ^ (1 - α)) ≤ (∑' (x : U × V), privComposeAdaptive nq1 nq2 l₁ x ^ α * privComposeAdaptive nq1 nq2 l₂ x ^ (1 - α)) := by + exact ENNReal.le_tsum (u0, v0) + apply (LE.le.trans_lt' Hle) + clear Hle + apply ENNReal.mul_pos + · sorry + · sorry case H2 => sorry rw [left_distrib] rw [Real.exp_add] @@ -136,10 +155,21 @@ lemma privComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> Lis · -- Should be easy sorry · linarith - - have GH1 : ∀ i, 0 < ∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α) := sorry - have GH2 : ∀ i, ∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α) < ⊤ := sorry - + have GH1 : ∀ i, 0 < ∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α) := by + intro i + rcases HV with ⟨ v0 ⟩ + have Hle : nq2 i l₁ v0 ^ α * nq2 i l₂ v0 ^ (1 - α) <= ∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α) := ENNReal.le_tsum v0 + apply (LE.le.trans_lt' Hle) + clear Hle + apply ENNReal.mul_pos + · have Hlt : (0 < nq2 i l₁ v0 ^ α) := by + apply ENNReal.rpow_pos + · sorry + · sorry + sorry + · sorry + have GH2 : ∀ i, ∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α) < ⊤ := by + sorry -- After this point the argument is tight apply Eq.le @@ -244,14 +274,6 @@ theorem privComposeAdaptive_zCDPBound {nq1 : List T → SLang U} {nq2 : U -> Lis case case_alg => linarith apply (privComposeAdaptive_renyi_bound _ Hα _ _) <;> aesop -/-- -Adaptive composed query distribution is nowhere zero --/ -theorem privComposeAdaptive_NonZeroNQ {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (nt1 : NonZeroNQ nq1) (nt2 : ∀ u, NonZeroNQ (nq2 u)) : - NonZeroNQ (privComposeAdaptive nq1 nq2) := by - simp [NonZeroNQ] at * - simp [privComposeAdaptive] - aesop /-- All outputs of a adaptive composed query have finite probability. @@ -270,9 +292,37 @@ Adaptive composed query is a proper distribution -/ theorem privComposeAdaptive_NonTopSum {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (nt1 : NonTopSum nq1) (nt2 : ∀ u, NonTopSum (nq2 u)) : NonTopSum (privComposeAdaptive nq1 nq2) := by - rw [NonTopSum] at * - -- Chain rule won't help here + simp [NonTopSum] at * + intro l + simp [privComposeAdaptive] + rw [ENNReal.tsum_prod'] + conv => + right + left + right + intro a + right + intro b + simp + rw [(compose_sum_rw_adaptive _ _ a b)] + conv => + right + left + right + intro a + rw [ENNReal.tsum_mul_left] admit + -- 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 + /-- @@ -281,56 +331,56 @@ Renyi divergence beteeen adaptive composed queries on neighbours are finite. theorem privComposeAdaptive_NonTopRDNQ {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (nt1 : NonTopRDNQ nq1) (nt2 : ∀ u, NonTopRDNQ (nq2 u)) (nn1 : NonTopNQ nq1) (nn2 : ∀ u, NonTopNQ (nq2 u)) : NonTopRDNQ (privComposeAdaptive nq1 nq2) := by rw [NonTopRDNQ] at * - admit - -- simp [NonTopRDNQ] at * - -- intro α h1 l₁ l₂ h2 + intro α h1 l₁ l₂ h2 -- replace nt1 := nt1 α h1 l₁ l₂ h2 -- replace nt2 := nt2 α h1 l₁ l₂ h2 - -- simp [privCompose] - -- 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_nonneg _ _ (le_of_lt (lt_trans zero_lt_one h1))] - -- rw [ENNReal.mul_rpow_of_ne_top (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] + simp [privComposeAdaptive] + rw [ENNReal.tsum_prod'] + simp + + conv => + right + left + right + intro x + right + intro y + congr + . left + rw [(compose_sum_rw_adaptive _ _ x y)] + . left + rw [(compose_sum_rw_adaptive _ _ x y)] + conv => + right + left + right + intro x + right + intro y + rw [ENNReal.mul_rpow_of_nonneg _ _ (le_of_lt (lt_trans zero_lt_one h1))] + rw [ENNReal.mul_rpow_of_ne_top (nn1 l₂ x) (nn2 x 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] + -- Might not be true, terms in the second sum are pointwise bounded but not uniformly bounded + -- intro H -- rw [mul_eq_top] at H -- cases H @@ -342,6 +392,7 @@ theorem privComposeAdaptive_NonTopRDNQ {nq1 : List T → SLang U} {nq2 : U -> Li -- cases h3 -- rename_i h4 h5 -- contradiction + admit /-- ``privComposeAdaptive`` satisfies zCDP From 2bb868f1963df2bd662e78f3c1f97aebb3638b94 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Fri, 7 Jun 2024 12:00:43 -0400 Subject: [PATCH 025/122] try to move away from explicit iSup --- .../ZeroConcentrated/AdaptiveComposition.lean | 370 ++++++++++-------- 1 file changed, 208 insertions(+), 162 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean index 868f54df..ff480e48 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean @@ -22,24 +22,41 @@ variable { T U V : Type } variable [HU : Inhabited U] variable [HV : Inhabited V] + +-- def NQBounded (nq2 : U -> List T -> SLang V) (b : ENNReal) : Prop := +-- ∃ b, ∀ u, ∀ l, ∀ v, nq2 u l v <= b < ⊤ + + +-- Morally, b = ⨆ (u : U), RenyiDivergence .... However, iSup itself does not remember that the supremum +-- exists, setting the value to zero if not. +def RDBounded (nq2 : U -> List T -> SLang V) (α : ℝ) (Hα : 1 < α) (l₁ l₂ : List T) (HN : Neighbour l₁ l₂) (b : ℝ): Prop := + ∀ u, (0 < RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) ∧ (RenyiDivergence (nq2 u l₁) (nq2 u l₂) α ≤ b) + +lemma RDBounded_ofZCDPBound {nq2 : U -> List T → SLang V} {ε₃ ε₄ : ℕ+} (α : ℝ) (Hα : 1 < α) (l₁ l₂ : List T) (HN : Neighbour l₁ l₂) + (h2 : ∀ u, zCDPBound (nq2 u) ((ε₃ : ℝ) / ε₄)) : RDBounded nq2 α Hα l₁ l₂ HN (⨆ (u : U), RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) := by + rw [RDBounded] + intro u + apply And.intro + · sorry + · sorry + + + + -- set_option pp.all true -- Maybe would be better as ENNReal? lemma iSup_smul_l (a : ℝ) (Ha : 0 <= a) (f : U -> ℝ) : a * ⨆ u, f u = ⨆ u, a * f u := by refine (mul_iSup_of_nonneg ?ha fun i => f i) apply Ha --- #check @sSupHomClass (ℝ -> ℝ) ℝ ℝ _ _ funlike_inst --- #check map_sSup - - -def funlike_inst : FunLike (ℝ → ℝ) ℝ ℝ := by - constructor - case coe => - intro f - apply f - case coe_injective' => - exact fun ⦃a₁ a₂⦄ a => a +-- def funlike_inst : FunLike (ℝ → ℝ) ℝ ℝ := by +-- constructor +-- case coe => +-- intro f +-- apply f +-- case coe_injective' => +-- exact fun ⦃a₁ a₂⦄ a => a -- set_option pp.all true lemma iSup_exp (f : U -> ℝ) : ⨆ u, Real.exp (f u) = Real.exp (⨆ u, f u) := by @@ -61,7 +78,8 @@ lemma exp_non_top : ∀ (z : ENNReal) (β : ℝ), z ≠ 0 -> z ≠ ⊤ -> z ^ β · aesop -lemma RenyiDivergence_exp (p q : SLang T) {α : ℝ} (h : 1 < α) (H1 : 0 < ∑' (x : T), p x ^ α * q x ^ (1 - α)) (H2 : ∑' (x : T), p x ^ α * q x ^ (1 - α) < ⊤): +lemma RenyiDivergence_exp (p q : SLang T) {α : ℝ} (h : 1 < α) + (H1 : 0 < ∑' (x : T), p x ^ α * q x ^ (1 - α)) (H2 : ∑' (x : T), p x ^ α * q x ^ (1 - α) < ⊤): Real.exp ((α - 1) * RenyiDivergence p q α) = (∑' x : T, (p x)^α * (q x)^(1 - α)).toReal := by simp only [RenyiDivergence] rw [<- mul_assoc] @@ -85,157 +103,175 @@ theorem privComposeAdaptive_NonZeroNQ {nq1 : List T → SLang U} {nq2 : U -> Lis simp [privComposeAdaptive] aesop - /-- Bound on Renyi divergence on adaptively composed queries -/ -lemma privComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (α : ℝ) (Hα : 1 < α) (HNT1 : NonTopNQ nq1) (HNTRDNQ2 : ∀ u, NonTopRDNQ (nq2 u)) (HN : Neighbour l₁ l₂) (HNZ1 : NonZeroNQ nq1) (HNZ2 : ∀ u, NonZeroNQ (nq2 u)) : - RenyiDivergence (privComposeAdaptive nq1 nq2 l₁) (privComposeAdaptive nq1 nq2 l₂) α ≤ RenyiDivergence (nq1 l₁) (nq1 l₂) α + ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α := by - apply (RenyiDivergence_mono_sum _ _ α Hα) - rw [RenyiDivergence_exp (privComposeAdaptive nq1 nq2 l₁) (privComposeAdaptive nq1 nq2 l₂) Hα ?H1 ?H2] - case H1 => - rcases HV with ⟨ v0 ⟩ - rcases HU with ⟨ u0 ⟩ - have Hle : (privComposeAdaptive nq1 nq2 l₁ (u0, v0) ^ α * privComposeAdaptive nq1 nq2 l₂ (u0, v0) ^ (1 - α)) ≤ (∑' (x : U × V), privComposeAdaptive nq1 nq2 l₁ x ^ α * privComposeAdaptive nq1 nq2 l₂ x ^ (1 - α)) := by - exact ENNReal.le_tsum (u0, v0) - apply (LE.le.trans_lt' Hle) - clear Hle - apply ENNReal.mul_pos - · sorry - · sorry - case H2 => sorry - rw [left_distrib] - rw [Real.exp_add] - - rw [RenyiDivergence_exp (nq1 l₁) (nq1 l₂) Hα ?H1 ?H2] - case H1 => sorry - case H2 => sorry - - have hmono_1 : rexp ((α - 1) * ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) = ⨆ u, rexp ((α - 1) * RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) := by - rw [iSup_smul_l, iSup_exp] - linarith - rw [hmono_1] - clear hmono_1 - - rw [mul_comm] - rw [<- (ENNReal.toReal_ofReal_mul _ _ ?h)] - case h => - refine Real.iSup_nonneg ?hf - intro i - exact exp_nonneg ((α - 1) * RenyiDivergence (nq2 i l₁) (nq2 i l₂) α) - rw [mul_comm] - rw [← ENNReal.tsum_mul_right] - - apply (toReal_mono' _ ?goal2) - case goal2 => - intro H - -- ?? - sorry - - - apply (@LE.le.trans _ _ _ ((∑' (i : U), nq1 l₁ i ^ α * nq1 l₂ i ^ (1 - α) * ENNReal.ofReal (rexp ((α - 1) * RenyiDivergence (nq2 i l₁) (nq2 i l₂) α)))) _ _ ?goal2) - case goal2 => - apply (tsum_le_tsum _ ENNReal.summable ENNReal.summable) - intro i - refine (ENNReal.mul_le_mul_left ?h.h.h0 ?h.h.hinf).mpr ?h.h.a - · apply mul_ne_zero_iff.mpr - apply And.intro - · sorry - · sorry - · apply ENNReal.mul_ne_top - · apply exp_non_top - · apply HNZ1 - · apply HNT1 - · apply exp_non_top - · apply HNZ1 - · apply HNT1 - · apply ENNReal.ofReal_le_ofReal - rw [iSup_exp] - rw [<- iSup_smul_l] - · -- Should be easy - sorry - · linarith - have GH1 : ∀ i, 0 < ∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α) := by - intro i - rcases HV with ⟨ v0 ⟩ - have Hle : nq2 i l₁ v0 ^ α * nq2 i l₂ v0 ^ (1 - α) <= ∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α) := ENNReal.le_tsum v0 - apply (LE.le.trans_lt' Hle) - clear Hle - apply ENNReal.mul_pos - · have Hlt : (0 < nq2 i l₁ v0 ^ α) := by - apply ENNReal.rpow_pos - · sorry - · sorry - sorry - · sorry - have GH2 : ∀ i, ∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α) < ⊤ := by - sorry - - -- After this point the argument is tight - apply Eq.le - conv => - rhs - congr - intro i - rw [RenyiDivergence_exp (nq2 i l₁) (nq2 i l₂) Hα] - rfl - · apply GH1 - · apply GH2 - - conv => - lhs - congr - intro - rw [privComposeChainRule] - rw [privComposeChainRule] - - conv => - rhs - congr - intro x - rw [<- (@ENNReal.ofReal_toReal (nq1 l₁ x ^ α * nq1 l₂ x ^ (1 - α)) ?goal2)] - · rw [<- ENNReal.ofReal_mul] - · rw [<- ENNReal.toReal_mul] - rw [(@ENNReal.ofReal_toReal (nq1 l₁ x ^ α * nq1 l₂ x ^ (1 - α) * ∑' (x_1 : V), nq2 x l₁ x_1 ^ α * nq2 x l₂ x_1 ^ (1 - α)) ?goal4)] - rfl - apply ENNReal.mul_ne_top - · apply ENNReal.mul_ne_top - · apply exp_non_top - · apply HNZ1 - · apply HNT1 - · apply exp_non_top - · apply HNZ1 - · apply HNT1 - · apply HNTRDNQ2 - apply Hα - apply HN - · apply ENNReal.toReal_nonneg - · apply ENNReal.mul_ne_top - · apply exp_non_top - · apply HNZ1 - · apply HNT1 - · apply exp_non_top - · apply HNZ1 - · apply HNT1 +lemma privComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (α : ℝ) (Hα : 1 < α) + (HNT1 : NonTopNQ nq1) (HNTRDNQ1 : NonTopRDNQ nq1) (HNTRDNQ2 : ∀ u, NonTopRDNQ (nq2 u)) + (HN : Neighbour l₁ l₂) (HNZ1 : NonZeroNQ nq1) (HNZ2 : ∀ u, NonZeroNQ (nq2 u)) + (b : ℝ) + (Hubound : RDBounded nq2 α Hα l₁ l₂ HN b) : + RenyiDivergence (privComposeAdaptive nq1 nq2 l₁) (privComposeAdaptive nq1 nq2 l₂) α ≤ + RenyiDivergence (nq1 l₁) (nq1 l₂) α + b := by + sorry - conv => - rhs - arg 1 - intro x - rw [<- ENNReal.tsum_mul_left] + -- apply (RenyiDivergence_mono_sum _ _ α Hα) + -- rw [RenyiDivergence_exp (privComposeAdaptive nq1 nq2 l₁) (privComposeAdaptive nq1 nq2 l₂) Hα ?H1 ?H2] + -- case H1 => + -- rcases HV with ⟨ v0 ⟩ + -- rcases HU with ⟨ u0 ⟩ + -- have Hle : (privComposeAdaptive nq1 nq2 l₁ (u0, v0) ^ α * privComposeAdaptive nq1 nq2 l₂ (u0, v0) ^ (1 - α)) ≤ (∑' (x : U × V), privComposeAdaptive nq1 nq2 l₁ x ^ α * privComposeAdaptive nq1 nq2 l₂ x ^ (1 - α)) := by + -- exact ENNReal.le_tsum (u0, v0) + -- apply (LE.le.trans_lt' Hle) + -- clear Hle + -- apply ENNReal.mul_pos + -- · sorry + -- · sorry + -- case H2 => sorry + -- rw [left_distrib] + -- rw [Real.exp_add] + + -- rw [RenyiDivergence_exp (nq1 l₁) (nq1 l₂) Hα ?H1 ?H2] + -- case H1 => sorry + -- case H2 => sorry + + -- have hmono_1 : rexp ((α - 1) * ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) = ⨆ u, rexp ((α - 1) * RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) := by + -- rw [iSup_smul_l] + -- rw [iSup_exp] + -- linarith + -- rw [hmono_1] + -- clear hmono_1 + + -- rw [mul_comm] + -- rw [<- (ENNReal.toReal_ofReal_mul _ _ ?h)] + -- case h => + -- refine Real.iSup_nonneg ?hf + -- intro i + -- exact exp_nonneg ((α - 1) * RenyiDivergence (nq2 i l₁) (nq2 i l₂) α) + -- rw [mul_comm] + -- rw [← ENNReal.tsum_mul_right] + + -- apply (toReal_mono' _ ?goal2) + -- case goal2 => + -- intro H + -- exfalso + -- rw [ENNReal.tsum_mul_right] at H + -- rw [mul_eq_top] at H + -- cases H + -- · aesop + -- · aesop + + -- apply (@LE.le.trans _ _ _ ((∑' (i : U), nq1 l₁ i ^ α * nq1 l₂ i ^ (1 - α) * ENNReal.ofReal (rexp ((α - 1) * RenyiDivergence (nq2 i l₁) (nq2 i l₂) α)))) _ _ ?goal2) + -- case goal2 => + -- apply (tsum_le_tsum _ ENNReal.summable ENNReal.summable) + -- intro i + -- refine (ENNReal.mul_le_mul_left ?h.h.h0 ?h.h.hinf).mpr ?h.h.a + -- · apply mul_ne_zero_iff.mpr + -- apply And.intro + -- · sorry + -- · sorry + -- · apply ENNReal.mul_ne_top + -- · apply exp_non_top + -- · apply HNZ1 + -- · apply HNT1 + -- · apply exp_non_top + -- · apply HNZ1 + -- · apply HNT1 + -- · apply ENNReal.ofReal_le_ofReal + -- rw [iSup_exp] + -- rw [<- iSup_smul_l] + -- · -- Should be easy + -- -- Definitely easy w/ bounded assumption + -- sorry + -- · linarith + -- have GH1 : ∀ i, 0 < ∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α) := by + -- intro i + -- rcases HV with ⟨ v0 ⟩ + -- have Hle : nq2 i l₁ v0 ^ α * nq2 i l₂ v0 ^ (1 - α) <= ∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α) := ENNReal.le_tsum v0 + -- apply (LE.le.trans_lt' Hle) + -- clear Hle + -- apply ENNReal.mul_pos + -- · have Hlt : (0 < nq2 i l₁ v0 ^ α) := by + -- apply ENNReal.rpow_pos + -- · sorry + -- · sorry + -- intro Hk + -- aesop + -- · have Hlt : (0 < nq2 i l₂ v0 ^ (1 - α)) := by + -- apply ENNReal.rpow_pos + -- · sorry + -- · sorry + -- intro Hk + -- aesop + + -- have GH2 : ∀ i, ∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α) < ⊤ := by + -- sorry - rw [<- ENNReal.tsum_prod] - congr - apply funext - intro p - rcases p with ⟨ u , v ⟩ - simp - rw [ENNReal.mul_rpow_of_nonneg _ _ ?sc1] - case sc1 => linarith - rw [mul_rpow_of_ne_zero] - · exact mul_mul_mul_comm (nq1 l₁ u ^ α) (nq2 u l₁ v ^ α) (nq1 l₂ u ^ (1 - α)) (nq2 u l₂ v ^ (1 - α)) - · apply HNZ1 - · apply HNZ2 + -- -- After this point the argument is tight + -- apply Eq.le + -- conv => + -- rhs + -- congr + -- intro i + -- rw [RenyiDivergence_exp (nq2 i l₁) (nq2 i l₂) Hα] + -- rfl + -- · apply GH1 + -- · apply GH2 + + -- conv => + -- lhs + -- congr + -- intro + -- rw [privComposeChainRule] + -- rw [privComposeChainRule] + + -- conv => + -- rhs + -- congr + -- intro x + -- rw [<- (@ENNReal.ofReal_toReal (nq1 l₁ x ^ α * nq1 l₂ x ^ (1 - α)) ?goal2)] + -- · rw [<- ENNReal.ofReal_mul] + -- · rw [<- ENNReal.toReal_mul] + -- rw [(@ENNReal.ofReal_toReal (nq1 l₁ x ^ α * nq1 l₂ x ^ (1 - α) * ∑' (x_1 : V), nq2 x l₁ x_1 ^ α * nq2 x l₂ x_1 ^ (1 - α)) ?goal4)] + -- rfl + -- apply ENNReal.mul_ne_top + -- · apply ENNReal.mul_ne_top + -- · apply exp_non_top + -- · apply HNZ1 + -- · apply HNT1 + -- · apply exp_non_top + -- · apply HNZ1 + -- · apply HNT1 + -- · apply HNTRDNQ2 + -- apply Hα + -- apply HN + -- · apply ENNReal.toReal_nonneg + -- · apply ENNReal.mul_ne_top + -- · apply exp_non_top + -- · apply HNZ1 + -- · apply HNT1 + -- · apply exp_non_top + -- · apply HNZ1 + -- · apply HNT1 + + -- conv => + -- rhs + -- arg 1 + -- intro x + -- rw [<- ENNReal.tsum_mul_left] + + -- rw [<- ENNReal.tsum_prod] + -- congr + -- apply funext + -- intro p + -- rcases p with ⟨ u , v ⟩ + -- simp + -- rw [ENNReal.mul_rpow_of_nonneg _ _ ?sc1] + -- case sc1 => linarith + -- rw [mul_rpow_of_ne_zero] + -- · exact mul_mul_mul_comm (nq1 l₁ u ^ α) (nq2 u l₁ v ^ α) (nq1 l₂ u ^ (1 - α)) (nq2 u l₂ v ^ (1 - α)) + -- · apply HNZ1 + -- · apply HNZ2 /-- Adaptively Composed queries satisfy zCDP Renyi divergence bound. @@ -270,9 +306,16 @@ theorem privComposeAdaptive_zCDPBound {nq1 : List T → SLang U} {nq2 : U -> Lis have marginal_ub := h1 α Hα l₁ l₂ Hneighbours have conditional_ub : (⨆ (u : U), RenyiDivergence (nq2 u l₁) (nq2 u l₂) α ≤ 1 / 2 * (↑↑ε₃ / ↑↑ε₄) ^ 2 * α) := ciSup_le fun x => h2 x α Hα l₁ l₂ Hneighbours - apply (@LE.le.trans _ _ _ (RenyiDivergence (nq1 l₁) (nq1 l₂) α + ⨆ (u : U), RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) _ _ ?case_alg) + apply (@LE.le.trans _ _ _ (RenyiDivergence (nq1 l₁) (nq1 l₂) α + ⨆ (u : U), RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) _ _ ?case_alg) case case_alg => linarith - apply (privComposeAdaptive_renyi_bound _ Hα _ _) <;> aesop + apply (privComposeAdaptive_renyi_bound _ Hα _ _) + · aesop + · aesop + · aesop + · apply RDBounded_ofZCDPBound <;> aesop + · aesop + · aesop + · aesop /-- @@ -311,7 +354,10 @@ theorem privComposeAdaptive_NonTopSum {nq1 : List T → SLang U} {nq2 : U -> Lis right intro a rw [ENNReal.tsum_mul_left] - admit + + -- Might need the second query to be bounded above + + sorry -- rw [ENNReal.tsum_mul_right] -- rw [mul_eq_top] -- intro H @@ -392,7 +438,7 @@ theorem privComposeAdaptive_NonTopRDNQ {nq1 : List T → SLang U} {nq2 : U -> Li -- cases h3 -- rename_i h4 h5 -- contradiction - admit + sorry /-- ``privComposeAdaptive`` satisfies zCDP From 23c422f482248ce0ad392ecbd14f796b4497fab0 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Fri, 7 Jun 2024 14:28:03 -0400 Subject: [PATCH 026/122] renyi bound proof --- .../ZeroConcentrated/AdaptiveComposition.lean | 496 ++++++++++-------- 1 file changed, 275 insertions(+), 221 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean index ff480e48..1a709978 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean @@ -29,18 +29,25 @@ variable [HV : Inhabited V] -- Morally, b = ⨆ (u : U), RenyiDivergence .... However, iSup itself does not remember that the supremum -- exists, setting the value to zero if not. -def RDBounded (nq2 : U -> List T -> SLang V) (α : ℝ) (Hα : 1 < α) (l₁ l₂ : List T) (HN : Neighbour l₁ l₂) (b : ℝ): Prop := +def RDBound (nq2 : U -> List T -> SLang V) (α : ℝ) (Hα : 1 < α) (l₁ l₂ : List T) (HN : Neighbour l₁ l₂) (b : ℝ) : Prop := ∀ u, (0 < RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) ∧ (RenyiDivergence (nq2 u l₁) (nq2 u l₂) α ≤ b) -lemma RDBounded_ofZCDPBound {nq2 : U -> List T → SLang V} {ε₃ ε₄ : ℕ+} (α : ℝ) (Hα : 1 < α) (l₁ l₂ : List T) (HN : Neighbour l₁ l₂) - (h2 : ∀ u, zCDPBound (nq2 u) ((ε₃ : ℝ) / ε₄)) : RDBounded nq2 α Hα l₁ l₂ HN (⨆ (u : U), RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) := by - rw [RDBounded] +lemma RDBound_ofZCDPBound {nq2 : U -> List T → SLang V} {ε₃ ε₄ : ℕ+} (α : ℝ) (Hα : 1 < α) (l₁ l₂ : List T) (HN : Neighbour l₁ l₂) + (h2 : ∀ u, zCDPBound (nq2 u) ((ε₃ : ℝ) / ε₄)) : RDBound nq2 α Hα l₁ l₂ HN (⨆ (u : U), RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) := by + rw [RDBound] intro u apply And.intro · sorry · sorry +def RDBounded (nq2 : U -> List T -> SLang V) : Prop := + ∀ (α : ℝ) (Hα : 1 < α) (l₁ l₂ : List T) (HN : Neighbour l₁ l₂), ∃ (b : ℝ), RDBound nq2 α Hα l₁ l₂ HN b +lemma RDBounded_ofZCDPBound {nq2 : U -> List T → SLang V} {ε₃ ε₄ : ℕ+} (h2 : ∀ u, zCDPBound (nq2 u) ((ε₃ : ℝ) / ε₄)) : + RDBounded nq2 := by + intro α Hα l₁ l₂ HN + exists (⨆ (u : U), RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) + apply RDBound_ofZCDPBound <;> aesop -- set_option pp.all true @@ -103,225 +110,16 @@ theorem privComposeAdaptive_NonZeroNQ {nq1 : List T → SLang U} {nq2 : U -> Lis simp [privComposeAdaptive] aesop -/-- -Bound on Renyi divergence on adaptively composed queries --/ -lemma privComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (α : ℝ) (Hα : 1 < α) - (HNT1 : NonTopNQ nq1) (HNTRDNQ1 : NonTopRDNQ nq1) (HNTRDNQ2 : ∀ u, NonTopRDNQ (nq2 u)) - (HN : Neighbour l₁ l₂) (HNZ1 : NonZeroNQ nq1) (HNZ2 : ∀ u, NonZeroNQ (nq2 u)) - (b : ℝ) - (Hubound : RDBounded nq2 α Hα l₁ l₂ HN b) : - RenyiDivergence (privComposeAdaptive nq1 nq2 l₁) (privComposeAdaptive nq1 nq2 l₂) α ≤ - RenyiDivergence (nq1 l₁) (nq1 l₂) α + b := by +lemma rpow_ne_zero_iff (x : ENNReal) (y : ℝ ): (¬x = 0 ∨ ¬ 0 < y) ∧ (¬ x = ⊤ ∨ ¬ y < 0) -> x ^ y ≠ 0 := by sorry - -- apply (RenyiDivergence_mono_sum _ _ α Hα) - -- rw [RenyiDivergence_exp (privComposeAdaptive nq1 nq2 l₁) (privComposeAdaptive nq1 nq2 l₂) Hα ?H1 ?H2] - -- case H1 => - -- rcases HV with ⟨ v0 ⟩ - -- rcases HU with ⟨ u0 ⟩ - -- have Hle : (privComposeAdaptive nq1 nq2 l₁ (u0, v0) ^ α * privComposeAdaptive nq1 nq2 l₂ (u0, v0) ^ (1 - α)) ≤ (∑' (x : U × V), privComposeAdaptive nq1 nq2 l₁ x ^ α * privComposeAdaptive nq1 nq2 l₂ x ^ (1 - α)) := by - -- exact ENNReal.le_tsum (u0, v0) - -- apply (LE.le.trans_lt' Hle) - -- clear Hle - -- apply ENNReal.mul_pos - -- · sorry - -- · sorry - -- case H2 => sorry - -- rw [left_distrib] - -- rw [Real.exp_add] - - -- rw [RenyiDivergence_exp (nq1 l₁) (nq1 l₂) Hα ?H1 ?H2] - -- case H1 => sorry - -- case H2 => sorry - - -- have hmono_1 : rexp ((α - 1) * ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) = ⨆ u, rexp ((α - 1) * RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) := by - -- rw [iSup_smul_l] - -- rw [iSup_exp] - -- linarith - -- rw [hmono_1] - -- clear hmono_1 - - -- rw [mul_comm] - -- rw [<- (ENNReal.toReal_ofReal_mul _ _ ?h)] - -- case h => - -- refine Real.iSup_nonneg ?hf - -- intro i - -- exact exp_nonneg ((α - 1) * RenyiDivergence (nq2 i l₁) (nq2 i l₂) α) - -- rw [mul_comm] - -- rw [← ENNReal.tsum_mul_right] - - -- apply (toReal_mono' _ ?goal2) - -- case goal2 => - -- intro H - -- exfalso - -- rw [ENNReal.tsum_mul_right] at H - -- rw [mul_eq_top] at H - -- cases H - -- · aesop - -- · aesop - - -- apply (@LE.le.trans _ _ _ ((∑' (i : U), nq1 l₁ i ^ α * nq1 l₂ i ^ (1 - α) * ENNReal.ofReal (rexp ((α - 1) * RenyiDivergence (nq2 i l₁) (nq2 i l₂) α)))) _ _ ?goal2) - -- case goal2 => - -- apply (tsum_le_tsum _ ENNReal.summable ENNReal.summable) - -- intro i - -- refine (ENNReal.mul_le_mul_left ?h.h.h0 ?h.h.hinf).mpr ?h.h.a - -- · apply mul_ne_zero_iff.mpr - -- apply And.intro - -- · sorry - -- · sorry - -- · apply ENNReal.mul_ne_top - -- · apply exp_non_top - -- · apply HNZ1 - -- · apply HNT1 - -- · apply exp_non_top - -- · apply HNZ1 - -- · apply HNT1 - -- · apply ENNReal.ofReal_le_ofReal - -- rw [iSup_exp] - -- rw [<- iSup_smul_l] - -- · -- Should be easy - -- -- Definitely easy w/ bounded assumption - -- sorry - -- · linarith - -- have GH1 : ∀ i, 0 < ∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α) := by - -- intro i - -- rcases HV with ⟨ v0 ⟩ - -- have Hle : nq2 i l₁ v0 ^ α * nq2 i l₂ v0 ^ (1 - α) <= ∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α) := ENNReal.le_tsum v0 - -- apply (LE.le.trans_lt' Hle) - -- clear Hle - -- apply ENNReal.mul_pos - -- · have Hlt : (0 < nq2 i l₁ v0 ^ α) := by - -- apply ENNReal.rpow_pos - -- · sorry - -- · sorry - -- intro Hk - -- aesop - -- · have Hlt : (0 < nq2 i l₂ v0 ^ (1 - α)) := by - -- apply ENNReal.rpow_pos - -- · sorry - -- · sorry - -- intro Hk - -- aesop - - -- have GH2 : ∀ i, ∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α) < ⊤ := by - -- sorry - - -- -- After this point the argument is tight - -- apply Eq.le - -- conv => - -- rhs - -- congr - -- intro i - -- rw [RenyiDivergence_exp (nq2 i l₁) (nq2 i l₂) Hα] - -- rfl - -- · apply GH1 - -- · apply GH2 - - -- conv => - -- lhs - -- congr - -- intro - -- rw [privComposeChainRule] - -- rw [privComposeChainRule] - - -- conv => - -- rhs - -- congr - -- intro x - -- rw [<- (@ENNReal.ofReal_toReal (nq1 l₁ x ^ α * nq1 l₂ x ^ (1 - α)) ?goal2)] - -- · rw [<- ENNReal.ofReal_mul] - -- · rw [<- ENNReal.toReal_mul] - -- rw [(@ENNReal.ofReal_toReal (nq1 l₁ x ^ α * nq1 l₂ x ^ (1 - α) * ∑' (x_1 : V), nq2 x l₁ x_1 ^ α * nq2 x l₂ x_1 ^ (1 - α)) ?goal4)] - -- rfl - -- apply ENNReal.mul_ne_top - -- · apply ENNReal.mul_ne_top - -- · apply exp_non_top - -- · apply HNZ1 - -- · apply HNT1 - -- · apply exp_non_top - -- · apply HNZ1 - -- · apply HNT1 - -- · apply HNTRDNQ2 - -- apply Hα - -- apply HN - -- · apply ENNReal.toReal_nonneg - -- · apply ENNReal.mul_ne_top - -- · apply exp_non_top - -- · apply HNZ1 - -- · apply HNT1 - -- · apply exp_non_top - -- · apply HNZ1 - -- · apply HNT1 - - -- conv => - -- rhs - -- arg 1 - -- intro x - -- rw [<- ENNReal.tsum_mul_left] - - -- rw [<- ENNReal.tsum_prod] - -- congr - -- apply funext - -- intro p - -- rcases p with ⟨ u , v ⟩ - -- simp - -- rw [ENNReal.mul_rpow_of_nonneg _ _ ?sc1] - -- case sc1 => linarith - -- rw [mul_rpow_of_ne_zero] - -- · exact mul_mul_mul_comm (nq1 l₁ u ^ α) (nq2 u l₁ v ^ α) (nq1 l₂ u ^ (1 - α)) (nq2 u l₂ v ^ (1 - α)) - -- · apply HNZ1 - -- · apply HNZ2 - -/-- -Adaptively Composed queries satisfy zCDP Renyi divergence bound. --/ -theorem privComposeAdaptive_zCDPBound {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} {ε₁ ε₂ ε₃ ε₄ : ℕ+} (h1 : zCDPBound nq1 ((ε₁ : ℝ) / ε₂)) (h2 : ∀ u, zCDPBound (nq2 u) ((ε₃ : ℝ) / ε₄)) (nn1 : NonZeroNQ nq1) (nn2 : ∀ u, NonZeroNQ (nq2 u)) (nt1 : NonTopRDNQ nq1) (nt2 : ∀ u, NonTopRDNQ (nq2 u)) (nts1 : NonTopNQ nq1) (nts2 : ∀ u, NonTopNQ (nq2 u)) : - zCDPBound (privComposeAdaptive nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by - rw [zCDPBound] - intro α Hα l₁ l₂ Hneighbours - -- Loose step - apply (@LE.le.trans _ _ _ (1/2 * (↑↑ε₁ / ↑↑ε₂)^2 * α + 1/2 * (↑↑ε₃ / ↑↑ε₄)^2 * α) _ _ ?case_sq) - case case_sq => - -- Binomial bound - rw [add_sq] - rw [<- right_distrib] - apply (mul_le_mul_of_nonneg_right _ ?goal1) - case goal1 => linarith - rw [<- left_distrib] - apply (mul_le_mul_of_nonneg_left _ ?goal1) - case goal1 => linarith - apply add_le_add_right - have hrw : (↑↑ε₁ / ↑↑ε₂ : ℝ) ^ 2 = (↑↑ε₁ / ↑↑ε₂) ^ 2 + 0 := by linarith - conv => - lhs - rw [hrw] - clear hrw - apply add_le_add_left - have h : 0 <= (↑↑ε₁ / ↑↑ε₂) * (↑↑ε₃ / ↑↑ε₄ : ℝ) := by - apply mul_nonneg <;> apply div_nonneg <;> linarith - linarith - -- Rewrite the upper bounds in terms of Renyi divergences of nq1/nq2 - rw [zCDPBound] at h1 - have marginal_ub := h1 α Hα l₁ l₂ Hneighbours - have conditional_ub : (⨆ (u : U), RenyiDivergence (nq2 u l₁) (nq2 u l₂) α ≤ 1 / 2 * (↑↑ε₃ / ↑↑ε₄) ^ 2 * α) := - ciSup_le fun x => h2 x α Hα l₁ l₂ Hneighbours - apply (@LE.le.trans _ _ _ (RenyiDivergence (nq1 l₁) (nq1 l₂) α + ⨆ (u : U), RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) _ _ ?case_alg) - case case_alg => linarith - apply (privComposeAdaptive_renyi_bound _ Hα _ _) - · aesop - · aesop - · aesop - · apply RDBounded_ofZCDPBound <;> aesop - · aesop - · aesop - · aesop - +lemma ne_top_lt_top (x : ENNReal) : (x ≠ ⊤) -> (x < ⊤) := sorry /-- All outputs of a adaptive composed query have finite probability. -/ -theorem privComposeAdaptive_NonTopNQ {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (nt1 : NonTopNQ nq1) (nt2 : ∀ u, NonTopNQ (nq2 u)) : +theorem privComposeAdaptive_NonTopNQ {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} + (nt1 : NonTopNQ nq1) (nt2 : ∀ u, NonTopNQ (nq2 u)) : NonTopNQ (privComposeAdaptive nq1 nq2) := by simp [NonTopNQ] at * intros l u v @@ -333,7 +131,8 @@ theorem privComposeAdaptive_NonTopNQ {nq1 : List T → SLang U} {nq2 : U -> List /-- Adaptive composed query is a proper distribution -/ -theorem privComposeAdaptive_NonTopSum {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (nt1 : NonTopSum nq1) (nt2 : ∀ u, NonTopSum (nq2 u)) : +theorem privComposeAdaptive_NonTopSum {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} + (nt1 : NonTopSum nq1) (nt2 : ∀ u, NonTopSum (nq2 u)) : NonTopSum (privComposeAdaptive nq1 nq2) := by simp [NonTopSum] at * intro l @@ -374,12 +173,11 @@ theorem privComposeAdaptive_NonTopSum {nq1 : List T → SLang U} {nq2 : U -> Lis /-- Renyi divergence beteeen adaptive composed queries on neighbours are finite. -/ -theorem privComposeAdaptive_NonTopRDNQ {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (nt1 : NonTopRDNQ nq1) (nt2 : ∀ u, NonTopRDNQ (nq2 u)) (nn1 : NonTopNQ nq1) (nn2 : ∀ u, NonTopNQ (nq2 u)) : +theorem privComposeAdaptive_NonTopRDNQ {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} + (nt1 : NonTopRDNQ nq1) (nt2 : ∀ u, NonTopRDNQ (nq2 u)) (nn1 : NonTopNQ nq1) (nn2 : ∀ u, NonTopNQ (nq2 u)) : NonTopRDNQ (privComposeAdaptive nq1 nq2) := by rw [NonTopRDNQ] at * intro α h1 l₁ l₂ h2 - -- replace nt1 := nt1 α h1 l₁ l₂ h2 - -- replace nt2 := nt2 α h1 l₁ l₂ h2 simp [privComposeAdaptive] rw [ENNReal.tsum_prod'] simp @@ -440,6 +238,262 @@ theorem privComposeAdaptive_NonTopRDNQ {nq1 : List T → SLang U} {nq2 : U -> Li -- contradiction sorry + +/-- +Bound on Renyi divergence on adaptively composed queries +-/ +lemma privComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (α : ℝ) (Hα : 1 < α) + (HNT1 : NonTopNQ nq1) (HNTRDNQ1 : NonTopRDNQ nq1) (HNTRDNQ2 : ∀ u, NonTopRDNQ (nq2 u)) + (HN : Neighbour l₁ l₂) (HNZ1 : NonZeroNQ nq1) (HNZ2 : ∀ u, NonZeroNQ (nq2 u)) (HNT2 : ∀ u, NonTopNQ (nq2 u)) + (b : ℝ) + (Hubound : RDBound nq2 α Hα l₁ l₂ HN b) : + RenyiDivergence (privComposeAdaptive nq1 nq2 l₁) (privComposeAdaptive nq1 nq2 l₂) α ≤ + RenyiDivergence (nq1 l₁) (nq1 l₂) α + b := by + apply (RenyiDivergence_mono_sum _ _ α Hα) + rw [RenyiDivergence_exp (privComposeAdaptive nq1 nq2 l₁) (privComposeAdaptive nq1 nq2 l₂) Hα ?H1 ?H2] + case H1 => + rcases HV with ⟨ v0 ⟩ + rcases HU with ⟨ u0 ⟩ + have Hle : (privComposeAdaptive nq1 nq2 l₁ (u0, v0) ^ α * privComposeAdaptive nq1 nq2 l₂ (u0, v0) ^ (1 - α)) ≤ (∑' (x : U × V), privComposeAdaptive nq1 nq2 l₁ x ^ α * privComposeAdaptive nq1 nq2 l₂ x ^ (1 - α)) := by + exact ENNReal.le_tsum (u0, v0) + apply (LE.le.trans_lt' Hle) + clear Hle + apply ENNReal.mul_pos + · apply rpow_ne_zero_iff + apply And.intro + · left + apply privComposeAdaptive_NonZeroNQ <;> aesop + · left + apply privComposeAdaptive_NonTopNQ <;> aesop + · apply rpow_ne_zero_iff + apply And.intro + · left + apply privComposeAdaptive_NonZeroNQ <;> aesop + · left + apply privComposeAdaptive_NonTopNQ <;> aesop + case H2 => + apply ne_top_lt_top + apply privComposeAdaptive_NonTopRDNQ <;> aesop + + rw [left_distrib] + rw [Real.exp_add] + + rw [RenyiDivergence_exp (nq1 l₁) (nq1 l₂) Hα ?H1 ?H2] + case H1 => + rcases HU with ⟨ u0 ⟩ + have Hle : nq1 l₁ u0 ^ α * nq1 l₂ u0 ^ (1 - α) <= ∑' (x : U), nq1 l₁ x ^ α * nq1 l₂ x ^ (1 - α) := ENNReal.le_tsum u0 + apply (LE.le.trans_lt' Hle) + clear Hle + apply ENNReal.mul_pos + · apply rpow_ne_zero_iff + apply And.intro + · left + apply HNZ1 + · left + apply HNT1 + · apply rpow_ne_zero_iff + apply And.intro + · left + apply HNZ1 + · left + apply HNT1 + case H2 => + apply ne_top_lt_top + apply HNTRDNQ1 <;> aesop + + have hexp_b : ∀ u, (rexp ((α - 1) * RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) <= rexp ((α - 1) * b)) := by + rw [RDBound] at Hubound + intro u + rcases (Hubound u) with ⟨ hb1 , hb2 ⟩ + apply Real.exp_le_exp_of_le + aesop + + rw [mul_comm] + rw [<- (ENNReal.toReal_ofReal_mul _ _ ?h)] + case h => + exact exp_nonneg ((α - 1) * b) + rw [mul_comm] + rw [← ENNReal.tsum_mul_right] + + apply (toReal_mono' _ ?goal2) + case goal2 => + intro H + exfalso + rw [ENNReal.tsum_mul_right] at H + rw [mul_eq_top] at H + cases H + · rename_i h + rcases h with ⟨ _ , h1 ⟩ + apply (@ENNReal.top_ne_ofReal (rexp ((α - 1) * b))) + aesop + · rename_i h + rcases h with ⟨ h0 , _ ⟩ + apply (HNTRDNQ1 α Hα l₁ l₂ HN) + apply h0 + + apply (@LE.le.trans _ _ _ ((∑' (i : U), nq1 l₁ i ^ α * nq1 l₂ i ^ (1 - α) * ENNReal.ofReal (rexp ((α - 1) * RenyiDivergence (nq2 i l₁) (nq2 i l₂) α)))) _ _ ?goal2) + case goal2 => + apply (tsum_le_tsum _ ENNReal.summable ENNReal.summable) + intro i + refine (ENNReal.mul_le_mul_left ?h.h.h0 ?h.h.hinf).mpr ?h.h.a + · apply mul_ne_zero_iff.mpr + apply And.intro + · apply rpow_ne_zero_iff + apply And.intro + · left + apply HNZ1 + · left + apply HNT1 + · apply rpow_ne_zero_iff + apply And.intro + · left + apply HNZ1 + · left + apply HNT1 + · apply ENNReal.mul_ne_top + · apply exp_non_top + · apply HNZ1 + · apply HNT1 + · apply exp_non_top + · apply HNZ1 + · apply HNT1 + · apply ENNReal.ofReal_le_ofReal + apply hexp_b + + have GH1 : ∀ i, 0 < ∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α) := by + intro i + rcases HV with ⟨ v0 ⟩ + have Hle : nq2 i l₁ v0 ^ α * nq2 i l₂ v0 ^ (1 - α) <= ∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α) := ENNReal.le_tsum v0 + apply (LE.le.trans_lt' Hle) + clear Hle + apply ENNReal.mul_pos + · have Hlt : (0 < nq2 i l₁ v0 ^ α) := by + apply ENNReal.rpow_pos + · exact pos_iff_ne_zero.mpr (HNZ2 i l₁ v0) + · apply HNT2 + intro Hk + aesop + · have Hlt : (0 < nq2 i l₂ v0 ^ (1 - α)) := by + apply ENNReal.rpow_pos + · exact pos_iff_ne_zero.mpr (HNZ2 i l₂ v0) + · exact HNT2 i l₂ v0 + intro Hk + aesop + + have GH2 : ∀ i, ∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α) < ⊤ := by + exact fun i => ne_top_lt_top (∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α)) (HNTRDNQ2 i α Hα l₁ l₂ HN) + + -- After this point the argument is tight + apply Eq.le + conv => + rhs + congr + intro i + rw [RenyiDivergence_exp (nq2 i l₁) (nq2 i l₂) Hα] + rfl + · apply GH1 + · apply GH2 + + conv => + lhs + congr + intro + rw [privComposeChainRule] + rw [privComposeChainRule] + + conv => + rhs + congr + intro x + rw [<- (@ENNReal.ofReal_toReal (nq1 l₁ x ^ α * nq1 l₂ x ^ (1 - α)) ?goal2)] + · rw [<- ENNReal.ofReal_mul] + · rw [<- ENNReal.toReal_mul] + rw [(@ENNReal.ofReal_toReal (nq1 l₁ x ^ α * nq1 l₂ x ^ (1 - α) * ∑' (x_1 : V), nq2 x l₁ x_1 ^ α * nq2 x l₂ x_1 ^ (1 - α)) ?goal4)] + rfl + apply ENNReal.mul_ne_top + · apply ENNReal.mul_ne_top + · apply exp_non_top + · apply HNZ1 + · apply HNT1 + · apply exp_non_top + · apply HNZ1 + · apply HNT1 + · apply HNTRDNQ2 + apply Hα + apply HN + · apply ENNReal.toReal_nonneg + · apply ENNReal.mul_ne_top + · apply exp_non_top + · apply HNZ1 + · apply HNT1 + · apply exp_non_top + · apply HNZ1 + · apply HNT1 + + conv => + rhs + arg 1 + intro x + rw [<- ENNReal.tsum_mul_left] + + rw [<- ENNReal.tsum_prod] + congr + apply funext + intro p + rcases p with ⟨ u , v ⟩ + simp + rw [ENNReal.mul_rpow_of_nonneg _ _ ?sc1] + case sc1 => linarith + rw [mul_rpow_of_ne_zero] + · exact mul_mul_mul_comm (nq1 l₁ u ^ α) (nq2 u l₁ v ^ α) (nq1 l₂ u ^ (1 - α)) (nq2 u l₂ v ^ (1 - α)) + · apply HNZ1 + · apply HNZ2 + +/-- +Adaptively Composed queries satisfy zCDP Renyi divergence bound. +-/ +theorem privComposeAdaptive_zCDPBound {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} {ε₁ ε₂ ε₃ ε₄ : ℕ+} (h1 : zCDPBound nq1 ((ε₁ : ℝ) / ε₂)) (h2 : ∀ u, zCDPBound (nq2 u) ((ε₃ : ℝ) / ε₄)) (nn1 : NonZeroNQ nq1) (nn2 : ∀ u, NonZeroNQ (nq2 u)) (nt1 : NonTopRDNQ nq1) (nt2 : ∀ u, NonTopRDNQ (nq2 u)) (nts1 : NonTopNQ nq1) (nts2 : ∀ u, NonTopNQ (nq2 u)) : + zCDPBound (privComposeAdaptive nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by + rw [zCDPBound] + intro α Hα l₁ l₂ Hneighbours + -- Loose step + apply (@LE.le.trans _ _ _ (1/2 * (↑↑ε₁ / ↑↑ε₂)^2 * α + 1/2 * (↑↑ε₃ / ↑↑ε₄)^2 * α) _ _ ?case_sq) + case case_sq => + -- Binomial bound + rw [add_sq] + rw [<- right_distrib] + apply (mul_le_mul_of_nonneg_right _ ?goal1) + case goal1 => linarith + rw [<- left_distrib] + apply (mul_le_mul_of_nonneg_left _ ?goal1) + case goal1 => linarith + apply add_le_add_right + have hrw : (↑↑ε₁ / ↑↑ε₂ : ℝ) ^ 2 = (↑↑ε₁ / ↑↑ε₂) ^ 2 + 0 := by linarith + conv => + lhs + rw [hrw] + clear hrw + apply add_le_add_left + have h : 0 <= (↑↑ε₁ / ↑↑ε₂) * (↑↑ε₃ / ↑↑ε₄ : ℝ) := by + apply mul_nonneg <;> apply div_nonneg <;> linarith + linarith + -- Rewrite the upper bounds in terms of Renyi divergences of nq1/nq2 + rw [zCDPBound] at h1 + have marginal_ub := h1 α Hα l₁ l₂ Hneighbours + have conditional_ub : (⨆ (u : U), RenyiDivergence (nq2 u l₁) (nq2 u l₂) α ≤ 1 / 2 * (↑↑ε₃ / ↑↑ε₄) ^ 2 * α) := + ciSup_le fun x => h2 x α Hα l₁ l₂ Hneighbours + apply (@LE.le.trans _ _ _ (RenyiDivergence (nq1 l₁) (nq1 l₂) α + ⨆ (u : U), RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) _ _ ?case_alg) + case case_alg => linarith + apply (privComposeAdaptive_renyi_bound _ Hα _ _) + · aesop + · aesop + · aesop + · aesop + · apply RDBound_ofZCDPBound <;> aesop + · aesop + · aesop + · aesop + /-- ``privComposeAdaptive`` satisfies zCDP -/ From 0c06bb2e957fe9075f740eb57e646afea7f5fe23 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Fri, 7 Jun 2024 14:33:23 -0400 Subject: [PATCH 027/122] some cleanup --- .../ZeroConcentrated/AdaptiveComposition.lean | 32 +++---------------- 1 file changed, 5 insertions(+), 27 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean index 1a709978..e85e5ae8 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean @@ -50,30 +50,6 @@ lemma RDBounded_ofZCDPBound {nq2 : U -> List T → SLang V} {ε₃ ε₄ : ℕ+} apply RDBound_ofZCDPBound <;> aesop --- set_option pp.all true --- Maybe would be better as ENNReal? -lemma iSup_smul_l (a : ℝ) (Ha : 0 <= a) (f : U -> ℝ) : a * ⨆ u, f u = ⨆ u, a * f u := by - refine (mul_iSup_of_nonneg ?ha fun i => f i) - apply Ha - - --- def funlike_inst : FunLike (ℝ → ℝ) ℝ ℝ := by --- constructor --- case coe => --- intro f --- apply f --- case coe_injective' => --- exact fun ⦃a₁ a₂⦄ a => a - --- set_option pp.all true -lemma iSup_exp (f : U -> ℝ) : ⨆ u, Real.exp (f u) = Real.exp (⨆ u, f u) := by - sorry - -- apply (@map_iSup (ℝ -> ℝ) ℝ ℝ U funlike_inst _ _ ?SHC rexp f) - -- · -- sSupHomClass - -- -- refine { map_sSup := ?SHC.map_sSup : @sSupHomClass (ℝ -> ℝ) ℝ ℝ _ _ _ funlike_inst} - -- -- May be circular - -- sorry - lemma exp_non_top : ∀ (z : ENNReal) (β : ℝ), z ≠ 0 -> z ≠ ⊤ -> z ^ β ≠ ⊤ := by intro z β Hz0 HzT intro W @@ -110,10 +86,12 @@ theorem privComposeAdaptive_NonZeroNQ {nq1 : List T → SLang U} {nq2 : U -> Lis simp [privComposeAdaptive] aesop -lemma rpow_ne_zero_iff (x : ENNReal) (y : ℝ ): (¬x = 0 ∨ ¬ 0 < y) ∧ (¬ x = ⊤ ∨ ¬ y < 0) -> x ^ y ≠ 0 := by - sorry +lemma rpow_ne_zero_iff (x : ENNReal) (y : ℝ): (¬x = 0 ∨ ¬ 0 < y) ∧ (¬ x = ⊤ ∨ ¬ y < 0) -> x ^ y ≠ 0 := by + have H1 := (@ENNReal.rpow_eq_zero_iff x y) + aesop -lemma ne_top_lt_top (x : ENNReal) : (x ≠ ⊤) -> (x < ⊤) := sorry +lemma ne_top_lt_top (x : ENNReal) : (x ≠ ⊤) -> (x < ⊤) := by + exact fun a => Ne.lt_top' (id (Ne.symm a)) /-- All outputs of a adaptive composed query have finite probability. From a9e8d0d83986f9617b8f948e515171ca19eb95ac Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Fri, 7 Jun 2024 15:45:40 -0400 Subject: [PATCH 028/122] close some side conditions --- .../ZeroConcentrated/AdaptiveComposition.lean | 110 ++++++++++++------ 1 file changed, 75 insertions(+), 35 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean index e85e5ae8..05f9fbb4 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean @@ -22,10 +22,7 @@ variable { T U V : Type } variable [HU : Inhabited U] variable [HV : Inhabited V] - --- def NQBounded (nq2 : U -> List T -> SLang V) (b : ENNReal) : Prop := --- ∃ b, ∀ u, ∀ l, ∀ v, nq2 u l v <= b < ⊤ - +set_option linter.unusedVariables false -- Morally, b = ⨆ (u : U), RenyiDivergence .... However, iSup itself does not remember that the supremum -- exists, setting the value to zero if not. @@ -93,6 +90,9 @@ lemma rpow_ne_zero_iff (x : ENNReal) (y : ℝ): (¬x = 0 ∨ ¬ 0 < y) ∧ (¬ x lemma ne_top_lt_top (x : ENNReal) : (x ≠ ⊤) -> (x < ⊤) := by exact fun a => Ne.lt_top' (id (Ne.symm a)) +lemma lt_top_ne_top (x : ENNReal) : (x < ⊤) -> ¬ (x = ⊤) := by + exact fun a => LT.lt.ne_top a + /-- All outputs of a adaptive composed query have finite probability. -/ @@ -110,7 +110,8 @@ theorem privComposeAdaptive_NonTopNQ {nq1 : List T → SLang U} {nq2 : U -> List Adaptive composed query is a proper distribution -/ theorem privComposeAdaptive_NonTopSum {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} - (nt1 : NonTopSum nq1) (nt2 : ∀ u, NonTopSum (nq2 u)) : + (nt1 : NonTopSum nq1) (nt2 : ∀ u, NonTopSum (nq2 u)) + (Hubound : RDBounded nq2) : NonTopSum (privComposeAdaptive nq1 nq2) := by simp [NonTopSum] at * intro l @@ -132,27 +133,36 @@ theorem privComposeAdaptive_NonTopSum {nq1 : List T → SLang U} {nq2 : U -> Lis intro a rw [ENNReal.tsum_mul_left] - -- Might need the second query to be bounded above - sorry - -- 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 + -- rcases (Hubound α Hα l₁ l₂ HN) with ⟨ b , Hubound ⟩ + -- apply lt_top_ne_top + -- apply (@LE.le.trans_lt _ _ _ (∑' (a : U), nq1 l a * ENNReal.ofReal b) _ ?goal1) + -- case goal1 => + -- apply ENNReal.tsum_le_tsum + -- intro a + -- -- b is not right here... I think I need nq2 to be uniformly bounded? + -- sorry + + -- rw [ENNReal.tsum_mul_right] + -- apply ne_top_lt_top + -- intro Hcont + -- rw [mul_eq_top] at Hcont + -- cases Hcont + -- · rename_i h + -- rcases h with ⟨ h0 , h1 ⟩ + -- aesop + -- · rename_i h + -- rcases h with ⟨ h0 , h1 ⟩ + -- aesop /-- Renyi divergence beteeen adaptive composed queries on neighbours are finite. -/ theorem privComposeAdaptive_NonTopRDNQ {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} - (nt1 : NonTopRDNQ nq1) (nt2 : ∀ u, NonTopRDNQ (nq2 u)) (nn1 : NonTopNQ nq1) (nn2 : ∀ u, NonTopNQ (nq2 u)) : + (nt1 : NonTopRDNQ nq1) (nt2 : ∀ u, NonTopRDNQ (nq2 u)) (nn1 : NonTopNQ nq1) (nn2 : ∀ u, NonTopNQ (nq2 u)) + (Hubound : RDBounded nq2) : NonTopRDNQ (privComposeAdaptive nq1 nq2) := by rw [NonTopRDNQ] at * intro α h1 l₁ l₂ h2 @@ -203,18 +213,30 @@ theorem privComposeAdaptive_NonTopRDNQ {nq1 : List T → SLang U} {nq2 : U -> Li rw [ENNReal.tsum_mul_left] -- Might not be true, terms in the second sum are pointwise bounded but not uniformly bounded - -- 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 - sorry + rcases (Hubound α h1 l₁ l₂ h2) with ⟨ b , Hubound ⟩ + + apply lt_top_ne_top + apply (@LE.le.trans_lt _ _ _ (∑' (x : U), nq1 l₁ x ^ α * nq1 l₂ x ^ (1 - α) * ENNReal.ofReal (rexp ((α - 1) * b))) _ ?goal1) + + case goal1 => + apply ENNReal.tsum_le_tsum + intro a + refine (ENNReal.mul_le_mul_left ?h.h0 ?h.hinf).mpr ?h.a + · sorry + · sorry + · sorry + + rw [ENNReal.tsum_mul_right] + apply ne_top_lt_top + intro H + rw [mul_eq_top] at H + cases H + . rename_i h3 + rcases h3 with ⟨ h30, h31 ⟩ + aesop + . rename_i h3 + rcases h3 with ⟨ h30, h31 ⟩ + aesop /-- @@ -224,7 +246,8 @@ lemma privComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> Lis (HNT1 : NonTopNQ nq1) (HNTRDNQ1 : NonTopRDNQ nq1) (HNTRDNQ2 : ∀ u, NonTopRDNQ (nq2 u)) (HN : Neighbour l₁ l₂) (HNZ1 : NonZeroNQ nq1) (HNZ2 : ∀ u, NonZeroNQ (nq2 u)) (HNT2 : ∀ u, NonTopNQ (nq2 u)) (b : ℝ) - (Hubound : RDBound nq2 α Hα l₁ l₂ HN b) : + (Hubound : RDBound nq2 α Hα l₁ l₂ HN b) + (Hubound2 : RDBounded nq2) : RenyiDivergence (privComposeAdaptive nq1 nq2 l₁) (privComposeAdaptive nq1 nq2 l₂) α ≤ RenyiDivergence (nq1 l₁) (nq1 l₂) α + b := by apply (RenyiDivergence_mono_sum _ _ α Hα) @@ -251,7 +274,14 @@ lemma privComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> Lis apply privComposeAdaptive_NonTopNQ <;> aesop case H2 => apply ne_top_lt_top - apply privComposeAdaptive_NonTopRDNQ <;> aesop + apply privComposeAdaptive_NonTopRDNQ + · apply HNTRDNQ1 + · apply HNTRDNQ2 + · apply HNT1 + · apply HNT2 + · apply Hubound2 + · apply Hα + · apply HN rw [left_distrib] rw [Real.exp_add] @@ -468,6 +498,7 @@ theorem privComposeAdaptive_zCDPBound {nq1 : List T → SLang U} {nq2 : U -> Lis · aesop · aesop · apply RDBound_ofZCDPBound <;> aesop + · apply RDBounded_ofZCDPBound <;> aesop · aesop · aesop · aesop @@ -481,8 +512,17 @@ theorem privComposeAdaptive_zCDP (nq1 : List T → SLang U) (nq2 : U -> List T repeat any_goals constructor · apply privComposeAdaptive_zCDPBound <;> aesop · apply privComposeAdaptive_NonZeroNQ <;> aesop - · apply privComposeAdaptive_NonTopSum <;> aesop + · apply privComposeAdaptive_NonTopSum + · aesop + · aesop + · apply (@RDBounded_ofZCDPBound _ _ _ nq2 ε₃ ε₄) + · aesop · apply privComposeAdaptive_NonTopNQ <;> aesop - · apply privComposeAdaptive_NonTopRDNQ <;> aesop - + · apply privComposeAdaptive_NonTopRDNQ + · aesop + · aesop + · aesop + · aesop + · apply (@RDBounded_ofZCDPBound _ _ _ nq2 ε₃ ε₄) + · aesop end SLang From df5a7be6ad307b0eca029c384c82778341b11ab3 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Fri, 7 Jun 2024 16:05:14 -0400 Subject: [PATCH 029/122] close more side conditions --- .../ZeroConcentrated/AdaptiveComposition.lean | 42 +++++++++++++++++-- 1 file changed, 38 insertions(+), 4 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean index 05f9fbb4..c6b3e53d 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean @@ -161,7 +161,7 @@ theorem privComposeAdaptive_NonTopSum {nq1 : List T → SLang U} {nq2 : U -> Lis Renyi divergence beteeen adaptive composed queries on neighbours are finite. -/ theorem privComposeAdaptive_NonTopRDNQ {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} - (nt1 : NonTopRDNQ nq1) (nt2 : ∀ u, NonTopRDNQ (nq2 u)) (nn1 : NonTopNQ nq1) (nn2 : ∀ u, NonTopNQ (nq2 u)) + (nt1 : NonTopRDNQ nq1) (nt2 : ∀ u, NonTopRDNQ (nq2 u)) (nn1 : NonTopNQ nq1) (nn2 : ∀ u, NonTopNQ (nq2 u)) (nz1 : NonZeroNQ nq1) (nz2 : ∀ u, NonZeroNQ (nq2 u)) (Hubound : RDBounded nq2) : NonTopRDNQ (privComposeAdaptive nq1 nq2) := by rw [NonTopRDNQ] at * @@ -222,9 +222,39 @@ theorem privComposeAdaptive_NonTopRDNQ {nq1 : List T → SLang U} {nq2 : U -> Li apply ENNReal.tsum_le_tsum intro a refine (ENNReal.mul_le_mul_left ?h.h0 ?h.hinf).mpr ?h.a - · sorry - · sorry - · sorry + · apply mul_ne_zero_iff.mpr + apply And.intro + · refine rpow_ne_zero_iff (nq1 l₁ a) α ?h.h0.left.a + apply And.intro + · left + aesop + · left + aesop + · refine rpow_ne_zero_iff (nq1 l₂ a) (1 - α) ?h.h0.right.a + apply And.intro + · left + aesop + · left + aesop + · intro H + rw [mul_eq_top] at H + cases H + · rename_i Hk + rcases Hk with ⟨ hk1, hk2 ⟩ + have Hcont : nq1 l₂ a ^ (1-α) ≠ ⊤ := by + apply exp_non_top + · aesop + · aesop + aesop + · rename_i Hk + rcases Hk with ⟨ hk1, hk2 ⟩ + have Hcont : nq1 l₁ a ^ (α) ≠ ⊤ := by + apply exp_non_top + · aesop + · apply nn1 + aesop + · -- rw [RenyiDivergence_exp (nq1 l₁) (nq1 l₂) Hα ?H1 ?H2] + sorry rw [ENNReal.tsum_mul_right] apply ne_top_lt_top @@ -279,6 +309,8 @@ lemma privComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> Lis · apply HNTRDNQ2 · apply HNT1 · apply HNT2 + · apply HNZ1 + · apply HNZ2 · apply Hubound2 · apply Hα · apply HN @@ -523,6 +555,8 @@ theorem privComposeAdaptive_zCDP (nq1 : List T → SLang U) (nq2 : U -> List T · aesop · aesop · aesop + · aesop + · aesop · apply (@RDBounded_ofZCDPBound _ _ _ nq2 ε₃ ε₄) · aesop end SLang From 4f043138879dd85c8decc2a8400c542a9163c587 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Fri, 7 Jun 2024 16:16:22 -0400 Subject: [PATCH 030/122] strengthen hypothesis: uniform bound on nq2 --- .../ZeroConcentrated/AdaptiveComposition.lean | 52 +++++++++---------- 1 file changed, 26 insertions(+), 26 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean index c6b3e53d..35f3674c 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean @@ -111,7 +111,7 @@ Adaptive composed query is a proper distribution -/ theorem privComposeAdaptive_NonTopSum {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (nt1 : NonTopSum nq1) (nt2 : ∀ u, NonTopSum (nq2 u)) - (Hubound : RDBounded nq2) : + (Hubound : ∃ (z : ℝ), ∀ (u : U), ∀ l, ∑'(v : V), nq2 u l v ≤ ENNReal.ofReal z) : NonTopSum (privComposeAdaptive nq1 nq2) := by simp [NonTopSum] at * intro l @@ -133,28 +133,26 @@ theorem privComposeAdaptive_NonTopSum {nq1 : List T → SLang U} {nq2 : U -> Lis intro a rw [ENNReal.tsum_mul_left] - sorry - - -- rcases (Hubound α Hα l₁ l₂ HN) with ⟨ b , Hubound ⟩ - -- apply lt_top_ne_top - -- apply (@LE.le.trans_lt _ _ _ (∑' (a : U), nq1 l a * ENNReal.ofReal b) _ ?goal1) - -- case goal1 => - -- apply ENNReal.tsum_le_tsum - -- intro a - -- -- b is not right here... I think I need nq2 to be uniformly bounded? - -- sorry - - -- rw [ENNReal.tsum_mul_right] - -- apply ne_top_lt_top - -- intro Hcont - -- rw [mul_eq_top] at Hcont - -- cases Hcont - -- · rename_i h - -- rcases h with ⟨ h0 , h1 ⟩ - -- aesop - -- · rename_i h - -- rcases h with ⟨ h0 , h1 ⟩ - -- aesop + rcases Hubound with ⟨ z , Hubound ⟩ + apply lt_top_ne_top + apply (@LE.le.trans_lt _ _ _ (∑' (a : U), nq1 l a * ENNReal.ofReal z) _ ?goal1) + case goal1 => + apply ENNReal.tsum_le_tsum + intro a + apply mul_le_mul_of_nonneg_left + · aesop + · aesop + rw [ENNReal.tsum_mul_right] + apply ne_top_lt_top + intro Hcont + rw [mul_eq_top] at Hcont + cases Hcont + · rename_i h + rcases h with ⟨ h0 , h1 ⟩ + aesop + · rename_i h + rcases h with ⟨ h0 , h1 ⟩ + aesop /-- @@ -538,7 +536,8 @@ theorem privComposeAdaptive_zCDPBound {nq1 : List T → SLang U} {nq2 : U -> Lis /-- ``privComposeAdaptive`` satisfies zCDP -/ -theorem privComposeAdaptive_zCDP (nq1 : List T → SLang U) (nq2 : U -> List T → SLang V) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h : zCDP nq1 ((ε₁ : ℝ) / ε₂)) (h' : ∀ u, zCDP (nq2 u) ((ε₃ : ℝ) / ε₄)) : +theorem privComposeAdaptive_zCDP (nq1 : List T → SLang U) (nq2 : U -> List T → SLang V) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h : zCDP nq1 ((ε₁ : ℝ) / ε₂)) (h' : ∀ u, zCDP (nq2 u) ((ε₃ : ℝ) / ε₄)) + (Hubound_nq2 : ∃ (z : ℝ), ∀ (u : U), ∀ l, ∑'(v : V), nq2 u l v ≤ ENNReal.ofReal z) : zCDP (privComposeAdaptive nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by simp [zCDP] at * repeat any_goals constructor @@ -547,8 +546,9 @@ theorem privComposeAdaptive_zCDP (nq1 : List T → SLang U) (nq2 : U -> List T · apply privComposeAdaptive_NonTopSum · aesop · aesop - · apply (@RDBounded_ofZCDPBound _ _ _ nq2 ε₃ ε₄) - · aesop + · apply Hubound_nq2 -- Is there any way out of this? Only used in one place + -- · apply (@RDBounded_ofZCDPBound _ _ _ nq2 ε₃ ε₄) + -- · aesop · apply privComposeAdaptive_NonTopNQ <;> aesop · apply privComposeAdaptive_NonTopRDNQ · aesop From dd2dbadd5e7c113e7fef996a15a06191075701ae Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Fri, 7 Jun 2024 16:27:01 -0400 Subject: [PATCH 031/122] checkpoint --- .../ZeroConcentrated/AdaptiveComposition.lean | 9 +++++++-- 1 file changed, 7 insertions(+), 2 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean index 35f3674c..e849680a 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean @@ -34,8 +34,13 @@ lemma RDBound_ofZCDPBound {nq2 : U -> List T → SLang V} {ε₃ ε₄ : ℕ+} ( rw [RDBound] intro u apply And.intro - · sorry - · sorry + · rw [RenyiDivergence] + apply Real.mul_pos + · apply inv_pos.mpr + linarith + · sorry + · refine le_ciSup_of_le ?right.H ?right.c ?right.h + sorry def RDBounded (nq2 : U -> List T -> SLang V) : Prop := ∀ (α : ℝ) (Hα : 1 < α) (l₁ l₂ : List T) (HN : Neighbour l₁ l₂), ∃ (b : ℝ), RDBound nq2 α Hα l₁ l₂ HN b From 107ccd5f6878e60c34a52d32595033138ee4cdbe Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Fri, 7 Jun 2024 16:38:37 -0400 Subject: [PATCH 032/122] modify boundedness hypothesis --- .../ZeroConcentrated/AdaptiveComposition.lean | 25 ++++++++----------- 1 file changed, 11 insertions(+), 14 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean index e849680a..1cb7b3b0 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean @@ -27,25 +27,21 @@ set_option linter.unusedVariables false -- Morally, b = ⨆ (u : U), RenyiDivergence .... However, iSup itself does not remember that the supremum -- exists, setting the value to zero if not. def RDBound (nq2 : U -> List T -> SLang V) (α : ℝ) (Hα : 1 < α) (l₁ l₂ : List T) (HN : Neighbour l₁ l₂) (b : ℝ) : Prop := - ∀ u, (0 < RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) ∧ (RenyiDivergence (nq2 u l₁) (nq2 u l₂) α ≤ b) + ∀ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α ≤ b lemma RDBound_ofZCDPBound {nq2 : U -> List T → SLang V} {ε₃ ε₄ : ℕ+} (α : ℝ) (Hα : 1 < α) (l₁ l₂ : List T) (HN : Neighbour l₁ l₂) - (h2 : ∀ u, zCDPBound (nq2 u) ((ε₃ : ℝ) / ε₄)) : RDBound nq2 α Hα l₁ l₂ HN (⨆ (u : U), RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) := by + (HNTRDNQ2 : ∀ u, NonTopRDNQ (nq2 u)) (h2 : ∀ u, zCDPBound (nq2 u) ((ε₃ : ℝ) / ε₄)) : + RDBound nq2 α Hα l₁ l₂ HN (⨆ (u : U), RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) := by rw [RDBound] intro u - apply And.intro - · rw [RenyiDivergence] - apply Real.mul_pos - · apply inv_pos.mpr - linarith - · sorry - · refine le_ciSup_of_le ?right.H ?right.c ?right.h - sorry + apply le_ciSup_of_le ?right.H u ?right.h + case right.h => linarith + sorry def RDBounded (nq2 : U -> List T -> SLang V) : Prop := ∀ (α : ℝ) (Hα : 1 < α) (l₁ l₂ : List T) (HN : Neighbour l₁ l₂), ∃ (b : ℝ), RDBound nq2 α Hα l₁ l₂ HN b -lemma RDBounded_ofZCDPBound {nq2 : U -> List T → SLang V} {ε₃ ε₄ : ℕ+} (h2 : ∀ u, zCDPBound (nq2 u) ((ε₃ : ℝ) / ε₄)) : +lemma RDBounded_ofZCDPBound {nq2 : U -> List T → SLang V} {ε₃ ε₄ : ℕ+} (h2 : ∀ u, zCDPBound (nq2 u) ((ε₃ : ℝ) / ε₄)) (HNTRDNQ2 : ∀ u, NonTopRDNQ (nq2 u)) : RDBounded nq2 := by intro α Hα l₁ l₂ HN exists (⨆ (u : U), RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) @@ -347,13 +343,13 @@ lemma privComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> Lis have hexp_b : ∀ u, (rexp ((α - 1) * RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) <= rexp ((α - 1) * b)) := by rw [RDBound] at Hubound intro u - rcases (Hubound u) with ⟨ hb1 , hb2 ⟩ + let hb2 := (Hubound u) apply Real.exp_le_exp_of_le aesop rw [mul_comm] - rw [<- (ENNReal.toReal_ofReal_mul _ _ ?h)] - case h => + rw [<- (ENNReal.toReal_ofReal_mul _ _ ?h1)] + case h1 => exact exp_nonneg ((α - 1) * b) rw [mul_comm] rw [← ENNReal.tsum_mul_right] @@ -564,4 +560,5 @@ theorem privComposeAdaptive_zCDP (nq1 : List T → SLang U) (nq2 : U -> List T · aesop · apply (@RDBounded_ofZCDPBound _ _ _ nq2 ε₃ ε₄) · aesop + · aesop end SLang From bf48995516bcff946b3864d86cef994553d62ced Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Fri, 7 Jun 2024 17:32:47 -0400 Subject: [PATCH 033/122] checkpoint --- .../ZeroConcentrated/AdaptiveComposition.lean | 42 ++++++++++++++++++- 1 file changed, 40 insertions(+), 2 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean index 1cb7b3b0..01abcae1 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean @@ -36,7 +36,9 @@ lemma RDBound_ofZCDPBound {nq2 : U -> List T → SLang V} {ε₃ ε₄ : ℕ+} ( intro u apply le_ciSup_of_le ?right.H u ?right.h case right.h => linarith + apply bddAbove_def.mpr sorry + -- No... NonTopRDNQ does not suffice, you need an actual uniform bound on the sum of nq2 def RDBounded (nq2 : U -> List T -> SLang V) : Prop := ∀ (α : ℝ) (Hα : 1 < α) (l₁ l₂ : List T) (HN : Neighbour l₁ l₂), ∃ (b : ℝ), RDBound nq2 α Hα l₁ l₂ HN b @@ -252,8 +254,44 @@ theorem privComposeAdaptive_NonTopRDNQ {nq1 : List T → SLang U} {nq2 : U -> Li · aesop · apply nn1 aesop - · -- rw [RenyiDivergence_exp (nq1 l₁) (nq1 l₂) Hα ?H1 ?H2] - sorry + · refine ((ENNReal.toReal_le_toReal ?g1 ?g2).mp ?g3) + case g1 => exact nt2 a α h1 l₁ l₂ h2 + case g2 => exact Ne.symm top_ne_ofReal + case g3 => + + have GH1 : ∀ i, 0 < ∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α) := by + intro i + rcases HV with ⟨ v0 ⟩ + have Hle : nq2 i l₁ v0 ^ α * nq2 i l₂ v0 ^ (1 - α) <= ∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α) := ENNReal.le_tsum v0 + apply (LE.le.trans_lt' Hle) + clear Hle + sorry + -- apply ENNReal.mul_pos + -- · have Hlt : (0 < nq2 i l₁ v0 ^ α) := by + -- apply ENNReal.rpow_pos + -- · have H1 : 0 <= nq2 i l₁ v0 := by sorry + -- have H2 : 0 ≠ nq2 i l₁ v0 := by sorry + -- sorry + -- · exact nn2 i l₁ v0 + -- intro Hk + -- rw [Hk] in Hlt + -- aesop + -- · have Hlt : (0 < nq2 i l₂ v0 ^ (1 - α)) := by + -- apply ENNReal.rpow_pos + -- · sorry + -- · exact nn2 i l₂ v0 + -- intro Hk + -- sorry + + have GH2 : ∀ i, ∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α) < ⊤ := by + exact fun i => ne_top_lt_top (∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α)) (nt2 i α h1 l₁ l₂ h2) + + rw [<- RenyiDivergence_exp (nq2 a l₁) (nq2 a l₂) h1 ?H1 ?H2] + case H1 => apply GH1 + case H2 => apply GH2 + rw [ENNReal.toReal_ofReal] + · sorry + · apply (exp_nonneg ((α - 1) * b)) rw [ENNReal.tsum_mul_right] apply ne_top_lt_top From ae0265b2ccb601869c7a417477e284f7c767a7cc Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Mon, 10 Jun 2024 10:18:27 -0400 Subject: [PATCH 034/122] proofs closed --- .../ZeroConcentrated/AdaptiveComposition.lean | 100 ++++++++++-------- 1 file changed, 54 insertions(+), 46 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean index 01abcae1..b03fed6a 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean @@ -29,25 +29,25 @@ set_option linter.unusedVariables false def RDBound (nq2 : U -> List T -> SLang V) (α : ℝ) (Hα : 1 < α) (l₁ l₂ : List T) (HN : Neighbour l₁ l₂) (b : ℝ) : Prop := ∀ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α ≤ b -lemma RDBound_ofZCDPBound {nq2 : U -> List T → SLang V} {ε₃ ε₄ : ℕ+} (α : ℝ) (Hα : 1 < α) (l₁ l₂ : List T) (HN : Neighbour l₁ l₂) - (HNTRDNQ2 : ∀ u, NonTopRDNQ (nq2 u)) (h2 : ∀ u, zCDPBound (nq2 u) ((ε₃ : ℝ) / ε₄)) : - RDBound nq2 α Hα l₁ l₂ HN (⨆ (u : U), RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) := by - rw [RDBound] - intro u - apply le_ciSup_of_le ?right.H u ?right.h - case right.h => linarith - apply bddAbove_def.mpr - sorry - -- No... NonTopRDNQ does not suffice, you need an actual uniform bound on the sum of nq2 +-- lemma RDBound_ofZCDPBound {nq2 : U -> List T → SLang V} {ε₃ ε₄ : ℕ+} (α : ℝ) (Hα : 1 < α) (l₁ l₂ : List T) (HN : Neighbour l₁ l₂) +-- (HNTRDNQ2 : ∀ u, NonTopRDNQ (nq2 u)) (h2 : ∀ u, zCDPBound (nq2 u) ((ε₃ : ℝ) / ε₄)) : +-- RDBound nq2 α Hα l₁ l₂ HN (⨆ (u : U), RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) := by +-- rw [RDBound] +-- intro u +-- apply le_ciSup_of_le ?right.H u ?right.h +-- case right.h => linarith +-- apply bddAbove_def.mpr +-- sorry +-- -- No... NonTopRDNQ does not suffice, you need an actual uniform bound on the sum of nq2 def RDBounded (nq2 : U -> List T -> SLang V) : Prop := ∀ (α : ℝ) (Hα : 1 < α) (l₁ l₂ : List T) (HN : Neighbour l₁ l₂), ∃ (b : ℝ), RDBound nq2 α Hα l₁ l₂ HN b -lemma RDBounded_ofZCDPBound {nq2 : U -> List T → SLang V} {ε₃ ε₄ : ℕ+} (h2 : ∀ u, zCDPBound (nq2 u) ((ε₃ : ℝ) / ε₄)) (HNTRDNQ2 : ∀ u, NonTopRDNQ (nq2 u)) : - RDBounded nq2 := by - intro α Hα l₁ l₂ HN - exists (⨆ (u : U), RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) - apply RDBound_ofZCDPBound <;> aesop +-- lemma RDBounded_ofZCDPBound {nq2 : U -> List T → SLang V} {ε₃ ε₄ : ℕ+} (h2 : ∀ u, zCDPBound (nq2 u) ((ε₃ : ℝ) / ε₄)) (HNTRDNQ2 : ∀ u, NonTopRDNQ (nq2 u)) : +-- RDBounded nq2 := by +-- intro α Hα l₁ l₂ HN +-- exists (⨆ (u : U), RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) +-- apply RDBound_ofZCDPBound <;> aesop lemma exp_non_top : ∀ (z : ENNReal) (β : ℝ), z ≠ 0 -> z ≠ ⊤ -> z ^ β ≠ ⊤ := by @@ -157,7 +157,6 @@ theorem privComposeAdaptive_NonTopSum {nq1 : List T → SLang U} {nq2 : U -> Lis rcases h with ⟨ h0 , h1 ⟩ aesop - /-- Renyi divergence beteeen adaptive composed queries on neighbours are finite. -/ @@ -170,7 +169,6 @@ theorem privComposeAdaptive_NonTopRDNQ {nq1 : List T → SLang U} {nq2 : U -> Li simp [privComposeAdaptive] rw [ENNReal.tsum_prod'] simp - conv => right left @@ -262,26 +260,35 @@ theorem privComposeAdaptive_NonTopRDNQ {nq1 : List T → SLang U} {nq2 : U -> Li have GH1 : ∀ i, 0 < ∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α) := by intro i rcases HV with ⟨ v0 ⟩ - have Hle : nq2 i l₁ v0 ^ α * nq2 i l₂ v0 ^ (1 - α) <= ∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α) := ENNReal.le_tsum v0 + have Hle : nq2 i l₁ v0 ^ α * nq2 i l₂ v0 ^ (1 - α) <= ∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α) := by apply ENNReal.le_tsum v0 apply (LE.le.trans_lt' Hle) clear Hle - sorry - -- apply ENNReal.mul_pos - -- · have Hlt : (0 < nq2 i l₁ v0 ^ α) := by - -- apply ENNReal.rpow_pos - -- · have H1 : 0 <= nq2 i l₁ v0 := by sorry - -- have H2 : 0 ≠ nq2 i l₁ v0 := by sorry - -- sorry - -- · exact nn2 i l₁ v0 - -- intro Hk - -- rw [Hk] in Hlt - -- aesop - -- · have Hlt : (0 < nq2 i l₂ v0 ^ (1 - α)) := by - -- apply ENNReal.rpow_pos - -- · sorry - -- · exact nn2 i l₂ v0 - -- intro Hk - -- sorry + apply ENNReal.mul_pos + · have Hlt : (0 < nq2 i l₁ v0 ^ α) := by + apply ENNReal.rpow_pos + · have H1 : 0 <= nq2 i l₁ v0 := by simp + have H2 : 0 ≠ nq2 i l₁ v0 := by + exact fun a => nz2 i l₁ v0 (id (Eq.symm a)) + apply LE.le.lt_or_eq at H1 + cases H1 + · assumption + · contradiction + · exact nn2 i l₁ v0 + intro Hk + rw [Hk] at Hlt + apply (lt_self_iff_false 0).mp Hlt + · intro Hk + have Hlt : (0 < nq2 i l₂ v0 ^ (1 - α)) := by + apply ENNReal.rpow_pos + · have H1 : 0 <= nq2 i l₂ v0 := by simp + have H2 : 0 ≠ nq2 i l₂ v0 := by exact fun a => nz2 i l₂ v0 (id (Eq.symm a)) + apply LE.le.lt_or_eq at H1 + cases H1 + · assumption + · contradiction + · exact nn2 i l₂ v0 + rw [Hk] at Hlt + apply (lt_self_iff_false 0).mp Hlt have GH2 : ∀ i, ∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α) < ⊤ := by exact fun i => ne_top_lt_top (∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α)) (nt2 i α h1 l₁ l₂ h2) @@ -290,7 +297,11 @@ theorem privComposeAdaptive_NonTopRDNQ {nq1 : List T → SLang U} {nq2 : U -> Li case H1 => apply GH1 case H2 => apply GH2 rw [ENNReal.toReal_ofReal] - · sorry + · apply Real.exp_le_exp_of_le + rw [RDBound] at Hubound + apply mul_le_mul_of_nonneg_left + · apply Hubound + · linarith · apply (exp_nonneg ((α - 1) * b)) rw [ENNReal.tsum_mul_right] @@ -529,7 +540,8 @@ lemma privComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> Lis /-- Adaptively Composed queries satisfy zCDP Renyi divergence bound. -/ -theorem privComposeAdaptive_zCDPBound {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} {ε₁ ε₂ ε₃ ε₄ : ℕ+} (h1 : zCDPBound nq1 ((ε₁ : ℝ) / ε₂)) (h2 : ∀ u, zCDPBound (nq2 u) ((ε₃ : ℝ) / ε₄)) (nn1 : NonZeroNQ nq1) (nn2 : ∀ u, NonZeroNQ (nq2 u)) (nt1 : NonTopRDNQ nq1) (nt2 : ∀ u, NonTopRDNQ (nq2 u)) (nts1 : NonTopNQ nq1) (nts2 : ∀ u, NonTopNQ (nq2 u)) : +theorem privComposeAdaptive_zCDPBound {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} {ε₁ ε₂ ε₃ ε₄ : ℕ+} (h1 : zCDPBound nq1 ((ε₁ : ℝ) / ε₂)) (h2 : ∀ u, zCDPBound (nq2 u) ((ε₃ : ℝ) / ε₄)) (nn1 : NonZeroNQ nq1) (nn2 : ∀ u, NonZeroNQ (nq2 u)) (nt1 : NonTopRDNQ nq1) (nt2 : ∀ u, NonTopRDNQ (nq2 u)) (nts1 : NonTopNQ nq1) (nts2 : ∀ u, NonTopNQ (nq2 u)) (HB : RDBounded nq2) (HB' : ∀ α Hα l₁ l₂ HN, RDBound nq2 α Hα l₁ l₂ HN (⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α)) : + zCDPBound (privComposeAdaptive nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by rw [zCDPBound] intro α Hα l₁ l₂ Hneighbours @@ -566,8 +578,8 @@ theorem privComposeAdaptive_zCDPBound {nq1 : List T → SLang U} {nq2 : U -> Lis · aesop · aesop · aesop - · apply RDBound_ofZCDPBound <;> aesop - · apply RDBounded_ofZCDPBound <;> aesop + · apply HB' + · apply HB · aesop · aesop · aesop @@ -576,7 +588,7 @@ theorem privComposeAdaptive_zCDPBound {nq1 : List T → SLang U} {nq2 : U -> Lis ``privComposeAdaptive`` satisfies zCDP -/ theorem privComposeAdaptive_zCDP (nq1 : List T → SLang U) (nq2 : U -> List T → SLang V) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h : zCDP nq1 ((ε₁ : ℝ) / ε₂)) (h' : ∀ u, zCDP (nq2 u) ((ε₃ : ℝ) / ε₄)) - (Hubound_nq2 : ∃ (z : ℝ), ∀ (u : U), ∀ l, ∑'(v : V), nq2 u l v ≤ ENNReal.ofReal z) : + (Hubound_nq2 : ∃ (z : ℝ), ∀ (u : U), ∀ l, ∑'(v : V), nq2 u l v ≤ ENNReal.ofReal z) (HB : RDBounded nq2) (HB' : ∀ α Hα l₁ l₂ HN, RDBound nq2 α Hα l₁ l₂ HN (⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α)): zCDP (privComposeAdaptive nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by simp [zCDP] at * repeat any_goals constructor @@ -585,9 +597,7 @@ theorem privComposeAdaptive_zCDP (nq1 : List T → SLang U) (nq2 : U -> List T · apply privComposeAdaptive_NonTopSum · aesop · aesop - · apply Hubound_nq2 -- Is there any way out of this? Only used in one place - -- · apply (@RDBounded_ofZCDPBound _ _ _ nq2 ε₃ ε₄) - -- · aesop + · apply Hubound_nq2 · apply privComposeAdaptive_NonTopNQ <;> aesop · apply privComposeAdaptive_NonTopRDNQ · aesop @@ -596,7 +606,5 @@ theorem privComposeAdaptive_zCDP (nq1 : List T → SLang U) (nq2 : U -> List T · aesop · aesop · aesop - · apply (@RDBounded_ofZCDPBound _ _ _ nq2 ε₃ ε₄) - · aesop - · aesop + · aesop end SLang From ee42d5f9b258306650122d1eca5001e266be1fe7 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Mon, 10 Jun 2024 11:22:46 -0400 Subject: [PATCH 035/122] tidy some bound hypotheses --- .../ZeroConcentrated/AdaptiveComposition.lean | 150 +++++++++--------- 1 file changed, 75 insertions(+), 75 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean index b03fed6a..0f3cbb39 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean @@ -22,33 +22,6 @@ variable { T U V : Type } variable [HU : Inhabited U] variable [HV : Inhabited V] -set_option linter.unusedVariables false - --- Morally, b = ⨆ (u : U), RenyiDivergence .... However, iSup itself does not remember that the supremum --- exists, setting the value to zero if not. -def RDBound (nq2 : U -> List T -> SLang V) (α : ℝ) (Hα : 1 < α) (l₁ l₂ : List T) (HN : Neighbour l₁ l₂) (b : ℝ) : Prop := - ∀ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α ≤ b - --- lemma RDBound_ofZCDPBound {nq2 : U -> List T → SLang V} {ε₃ ε₄ : ℕ+} (α : ℝ) (Hα : 1 < α) (l₁ l₂ : List T) (HN : Neighbour l₁ l₂) --- (HNTRDNQ2 : ∀ u, NonTopRDNQ (nq2 u)) (h2 : ∀ u, zCDPBound (nq2 u) ((ε₃ : ℝ) / ε₄)) : --- RDBound nq2 α Hα l₁ l₂ HN (⨆ (u : U), RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) := by --- rw [RDBound] --- intro u --- apply le_ciSup_of_le ?right.H u ?right.h --- case right.h => linarith --- apply bddAbove_def.mpr --- sorry --- -- No... NonTopRDNQ does not suffice, you need an actual uniform bound on the sum of nq2 - -def RDBounded (nq2 : U -> List T -> SLang V) : Prop := - ∀ (α : ℝ) (Hα : 1 < α) (l₁ l₂ : List T) (HN : Neighbour l₁ l₂), ∃ (b : ℝ), RDBound nq2 α Hα l₁ l₂ HN b - --- lemma RDBounded_ofZCDPBound {nq2 : U -> List T → SLang V} {ε₃ ε₄ : ℕ+} (h2 : ∀ u, zCDPBound (nq2 u) ((ε₃ : ℝ) / ε₄)) (HNTRDNQ2 : ∀ u, NonTopRDNQ (nq2 u)) : --- RDBounded nq2 := by --- intro α Hα l₁ l₂ HN --- exists (⨆ (u : U), RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) --- apply RDBound_ofZCDPBound <;> aesop - lemma exp_non_top : ∀ (z : ENNReal) (β : ℝ), z ≠ 0 -> z ≠ ⊤ -> z ^ β ≠ ⊤ := by intro z β Hz0 HzT @@ -60,9 +33,53 @@ lemma exp_non_top : ∀ (z : ENNReal) (β : ℝ), z ≠ 0 -> z ≠ ⊤ -> z ^ β · aesop · aesop +lemma rpow_ne_zero_iff (x : ENNReal) (y : ℝ): (¬x = 0 ∨ ¬ 0 < y) ∧ (¬ x = ⊤ ∨ ¬ y < 0) -> x ^ y ≠ 0 := by + have _ := (@ENNReal.rpow_eq_zero_iff x y) + aesop + +lemma ne_top_lt_top (x : ENNReal) : (x ≠ ⊤) -> (x < ⊤) := by + exact fun a => Ne.lt_top' (id (Ne.symm a)) + +lemma lt_top_ne_top (x : ENNReal) : (x < ⊤) -> ¬ (x = ⊤) := by + exact fun a => LT.lt.ne_top a + + +/-- +``b`` is an upper bound on the Renyi divergence for particular neighbours, uniform across all values of the first query. +-/ +def RDBound (nq2 : U -> List T -> SLang V) (α : ℝ) (l₁ l₂ : List T) (b : ℝ) : Prop := + ∀ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α ≤ b + +/-- +There exsts an upper bound on the Renyi divergence, uniform across all values of the first query. +-/ +def RDBounded (nq2 : U -> List T -> SLang V) : Prop := + ∀ α l₁ l₂, (1 < α) -> (Neighbour l₁ l₂) -> ∃ (b : ℝ), RDBound nq2 α l₁ l₂ b + +/-- +``b`` is an upper bound on the Renyi divergence, uniform across all values of the first query. +-/ +def URDBound (nq2 : U -> List T -> SLang V) (b : ℝ -> List T -> List T -> ℝ) : Prop := + ∀ α l₁ l₂, (1 < α) -> (Neighbour l₁ l₂) -> RDBound nq2 α l₁ l₂ (b α l₁ l₂) + +lemma RDBounded_of_URDBound (nq2 : U -> List T -> SLang V) (b : ℝ -> List T -> List T -> ℝ) : URDBound nq2 b -> RDBounded nq2 := by + rw [URDBound, RDBounded] + aesop + + +lemma URDBound_of_sup (nq2 : U -> List T -> SLang V) : URDBound nq2 (fun α l₁ l₂ => ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) := by + -- Is this provable? + -- I think only if the supremum exists, since we're working in ℝ instead of ENNReal + -- An in ℝ, a supremumj which does not exist is 0. + sorry -lemma RenyiDivergence_exp (p q : SLang T) {α : ℝ} (h : 1 < α) - (H1 : 0 < ∑' (x : T), p x ^ α * q x ^ (1 - α)) (H2 : ∑' (x : T), p x ^ α * q x ^ (1 - α) < ⊤): +/-- +Equation for the Renyi divergence series in terms of the Renyi Divergence +-/ +lemma RenyiDivergence_exp (p q : SLang T) {α : ℝ} + (h : 1 < α) + (H1 : 0 < ∑' (x : T), p x ^ α * q x ^ (1 - α)) + (H2 : ∑' (x : T), p x ^ α * q x ^ (1 - α) < ⊤) : Real.exp ((α - 1) * RenyiDivergence p q α) = (∑' x : T, (p x)^α * (q x)^(1 - α)).toReal := by simp only [RenyiDivergence] rw [<- mul_assoc] @@ -80,21 +97,13 @@ lemma RenyiDivergence_exp (p q : SLang T) {α : ℝ} (h : 1 < α) /-- Adaptive composed query distribution is nowhere zero -/ -theorem privComposeAdaptive_NonZeroNQ {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (nt1 : NonZeroNQ nq1) (nt2 : ∀ u, NonZeroNQ (nq2 u)) : +theorem privComposeAdaptive_NonZeroNQ {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} + (nt1 : NonZeroNQ nq1) (nt2 : ∀ u, NonZeroNQ (nq2 u)) : NonZeroNQ (privComposeAdaptive nq1 nq2) := by simp [NonZeroNQ] at * simp [privComposeAdaptive] aesop -lemma rpow_ne_zero_iff (x : ENNReal) (y : ℝ): (¬x = 0 ∨ ¬ 0 < y) ∧ (¬ x = ⊤ ∨ ¬ y < 0) -> x ^ y ≠ 0 := by - have H1 := (@ENNReal.rpow_eq_zero_iff x y) - aesop - -lemma ne_top_lt_top (x : ENNReal) : (x ≠ ⊤) -> (x < ⊤) := by - exact fun a => Ne.lt_top' (id (Ne.symm a)) - -lemma lt_top_ne_top (x : ENNReal) : (x < ⊤) -> ¬ (x = ⊤) := by - exact fun a => LT.lt.ne_top a /-- All outputs of a adaptive composed query have finite probability. @@ -113,7 +122,7 @@ theorem privComposeAdaptive_NonTopNQ {nq1 : List T → SLang U} {nq2 : U -> List Adaptive composed query is a proper distribution -/ theorem privComposeAdaptive_NonTopSum {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} - (nt1 : NonTopSum nq1) (nt2 : ∀ u, NonTopSum (nq2 u)) + (nt1 : NonTopSum nq1) (Hubound : ∃ (z : ℝ), ∀ (u : U), ∀ l, ∑'(v : V), nq2 u l v ≤ ENNReal.ofReal z) : NonTopSum (privComposeAdaptive nq1 nq2) := by simp [NonTopSum] at * @@ -151,10 +160,10 @@ theorem privComposeAdaptive_NonTopSum {nq1 : List T → SLang U} {nq2 : U -> Lis rw [mul_eq_top] at Hcont cases Hcont · rename_i h - rcases h with ⟨ h0 , h1 ⟩ + rcases h with ⟨ _ , h1 ⟩ aesop · rename_i h - rcases h with ⟨ h0 , h1 ⟩ + rcases h with ⟨ h0 , _ ⟩ aesop /-- @@ -210,9 +219,8 @@ theorem privComposeAdaptive_NonTopRDNQ {nq1 : List T → SLang U} {nq2 : U -> Li right intro x rw [ENNReal.tsum_mul_left] - -- Might not be true, terms in the second sum are pointwise bounded but not uniformly bounded - rcases (Hubound α h1 l₁ l₂ h2) with ⟨ b , Hubound ⟩ + rcases (Hubound α l₁ l₂ h1 h2) with ⟨ b , Hubound ⟩ apply lt_top_ne_top apply (@LE.le.trans_lt _ _ _ (∑' (x : U), nq1 l₁ x ^ α * nq1 l₂ x ^ (1 - α) * ENNReal.ofReal (rexp ((α - 1) * b))) _ ?goal1) @@ -310,21 +318,21 @@ theorem privComposeAdaptive_NonTopRDNQ {nq1 : List T → SLang U} {nq2 : U -> Li rw [mul_eq_top] at H cases H . rename_i h3 - rcases h3 with ⟨ h30, h31 ⟩ + rcases h3 with ⟨ _ , h31 ⟩ aesop . rename_i h3 - rcases h3 with ⟨ h30, h31 ⟩ + rcases h3 with ⟨ h30, _ ⟩ aesop /-- Bound on Renyi divergence on adaptively composed queries -/ -lemma privComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (α : ℝ) (Hα : 1 < α) +lemma privComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} (α : ℝ) (Hα : 1 < α) (l₁ l₂ : List T) (HNT1 : NonTopNQ nq1) (HNTRDNQ1 : NonTopRDNQ nq1) (HNTRDNQ2 : ∀ u, NonTopRDNQ (nq2 u)) (HN : Neighbour l₁ l₂) (HNZ1 : NonZeroNQ nq1) (HNZ2 : ∀ u, NonZeroNQ (nq2 u)) (HNT2 : ∀ u, NonTopNQ (nq2 u)) (b : ℝ) - (Hubound : RDBound nq2 α Hα l₁ l₂ HN b) + (Hubound : RDBound nq2 α l₁ l₂ b) (Hubound2 : RDBounded nq2) : RenyiDivergence (privComposeAdaptive nq1 nq2 l₁) (privComposeAdaptive nq1 nq2 l₂) α ≤ RenyiDivergence (nq1 l₁) (nq1 l₂) α + b := by @@ -392,7 +400,7 @@ lemma privComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> Lis have hexp_b : ∀ u, (rexp ((α - 1) * RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) <= rexp ((α - 1) * b)) := by rw [RDBound] at Hubound intro u - let hb2 := (Hubound u) + let _ := (Hubound u) apply Real.exp_le_exp_of_le aesop @@ -540,8 +548,15 @@ lemma privComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> Lis /-- Adaptively Composed queries satisfy zCDP Renyi divergence bound. -/ -theorem privComposeAdaptive_zCDPBound {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} {ε₁ ε₂ ε₃ ε₄ : ℕ+} (h1 : zCDPBound nq1 ((ε₁ : ℝ) / ε₂)) (h2 : ∀ u, zCDPBound (nq2 u) ((ε₃ : ℝ) / ε₄)) (nn1 : NonZeroNQ nq1) (nn2 : ∀ u, NonZeroNQ (nq2 u)) (nt1 : NonTopRDNQ nq1) (nt2 : ∀ u, NonTopRDNQ (nq2 u)) (nts1 : NonTopNQ nq1) (nts2 : ∀ u, NonTopNQ (nq2 u)) (HB : RDBounded nq2) (HB' : ∀ α Hα l₁ l₂ HN, RDBound nq2 α Hα l₁ l₂ HN (⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α)) : - +theorem privComposeAdaptive_zCDPBound {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} {ε₁ ε₂ ε₃ ε₄ : ℕ+} + (h1 : zCDPBound nq1 ((ε₁ : ℝ) / ε₂)) + (h2 : ∀ u, zCDPBound (nq2 u) ((ε₃ : ℝ) / ε₄)) + (nn1 : NonZeroNQ nq1) + (nn2 : ∀ u, NonZeroNQ (nq2 u)) + (nt1 : NonTopRDNQ nq1) + (nt2 : ∀ u, NonTopRDNQ (nq2 u)) (nts1 : NonTopNQ nq1) (nts2 : ∀ u, NonTopNQ (nq2 u)) + (HB : RDBounded nq2) + (HB' : URDBound nq2 (fun α l₁ l₂ => ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α)) : zCDPBound (privComposeAdaptive nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by rw [zCDPBound] intro α Hα l₁ l₂ Hneighbours @@ -573,38 +588,23 @@ theorem privComposeAdaptive_zCDPBound {nq1 : List T → SLang U} {nq2 : U -> Lis ciSup_le fun x => h2 x α Hα l₁ l₂ Hneighbours apply (@LE.le.trans _ _ _ (RenyiDivergence (nq1 l₁) (nq1 l₂) α + ⨆ (u : U), RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) _ _ ?case_alg) case case_alg => linarith - apply (privComposeAdaptive_renyi_bound _ Hα _ _) - · aesop - · aesop - · aesop - · aesop - · apply HB' - · apply HB - · aesop - · aesop - · aesop + apply (privComposeAdaptive_renyi_bound _ _ _ _) <;> aesop /-- ``privComposeAdaptive`` satisfies zCDP -/ -theorem privComposeAdaptive_zCDP (nq1 : List T → SLang U) (nq2 : U -> List T → SLang V) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h : zCDP nq1 ((ε₁ : ℝ) / ε₂)) (h' : ∀ u, zCDP (nq2 u) ((ε₃ : ℝ) / ε₄)) - (Hubound_nq2 : ∃ (z : ℝ), ∀ (u : U), ∀ l, ∑'(v : V), nq2 u l v ≤ ENNReal.ofReal z) (HB : RDBounded nq2) (HB' : ∀ α Hα l₁ l₂ HN, RDBound nq2 α Hα l₁ l₂ HN (⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α)): +theorem privComposeAdaptive_zCDP {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} {ε₁ ε₂ ε₃ ε₄ : ℕ+} + (h : zCDP nq1 ((ε₁ : ℝ) / ε₂)) + (h' : ∀ u, zCDP (nq2 u) ((ε₃ : ℝ) / ε₄)) + (Hubound_nq2 : ∃ (z : ℝ), ∀ (u : U), ∀ l, ∑'(v : V), nq2 u l v ≤ ENNReal.ofReal z) + (HB' : URDBound nq2 (fun α l₁ l₂ => ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α)) : -- Can I get rid of URDBound? zCDP (privComposeAdaptive nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by simp [zCDP] at * + have _ : RDBounded nq2 := RDBounded_of_URDBound nq2 (fun α l₁ l₂ => ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) HB' repeat any_goals constructor · apply privComposeAdaptive_zCDPBound <;> aesop · apply privComposeAdaptive_NonZeroNQ <;> aesop - · apply privComposeAdaptive_NonTopSum - · aesop - · aesop - · apply Hubound_nq2 + · apply privComposeAdaptive_NonTopSum <;> aesop · apply privComposeAdaptive_NonTopNQ <;> aesop - · apply privComposeAdaptive_NonTopRDNQ - · aesop - · aesop - · aesop - · aesop - · aesop - · aesop - · aesop + · apply privComposeAdaptive_NonTopRDNQ <;> aesop end SLang From ee98f4dc5f5d025a674a56a04e67da5054a1fbe7 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Mon, 10 Jun 2024 12:06:35 -0400 Subject: [PATCH 036/122] some cleanup --- .../ZeroConcentrated/AdaptiveComposition.lean | 53 +++++++------------ 1 file changed, 20 insertions(+), 33 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean index 0f3cbb39..b19c7272 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean @@ -30,8 +30,8 @@ lemma exp_non_top : ∀ (z : ENNReal) (β : ℝ), z ≠ 0 -> z ≠ ⊤ -> z ^ β apply rpow_eq_top_iff.mp apply W cases h - · aesop - · aesop + · simp_all only [ne_eq, not_true_eq_false] + · simp_all only [ne_eq, top_ne_zero, not_false_eq_true, not_true_eq_false] lemma rpow_ne_zero_iff (x : ENNReal) (y : ℝ): (¬x = 0 ∨ ¬ 0 < y) ∧ (¬ x = ⊤ ∨ ¬ y < 0) -> x ^ y ≠ 0 := by have _ := (@ENNReal.rpow_eq_zero_iff x y) @@ -67,11 +67,10 @@ lemma RDBounded_of_URDBound (nq2 : U -> List T -> SLang V) (b : ℝ -> List T -> aesop -lemma URDBound_of_sup (nq2 : U -> List T -> SLang V) : URDBound nq2 (fun α l₁ l₂ => ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) := by - -- Is this provable? - -- I think only if the supremum exists, since we're working in ℝ instead of ENNReal - -- An in ℝ, a supremumj which does not exist is 0. - sorry +-- lemma URDBound_of_sup (nq2 : U -> List T -> SLang V) : URDBound nq2 (fun α l₁ l₂ => ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) := by +-- -- Is this provable? +-- -- I think only if the supremum exists, since we're working in ℝ instead of ENNReal +-- sorry /-- Equation for the Renyi divergence series in terms of the Renyi Divergence @@ -91,7 +90,7 @@ lemma RenyiDivergence_exp (p q : SLang T) {α : ℝ} simp rw [Real.exp_log] apply ENNReal.toReal_pos_iff.mpr - apply And.intro H1 H2 + tauto /-- @@ -102,7 +101,7 @@ theorem privComposeAdaptive_NonZeroNQ {nq1 : List T → SLang U} {nq2 : U -> Lis NonZeroNQ (privComposeAdaptive nq1 nq2) := by simp [NonZeroNQ] at * simp [privComposeAdaptive] - aesop + tauto /-- @@ -114,9 +113,8 @@ theorem privComposeAdaptive_NonTopNQ {nq1 : List T → SLang U} {nq2 : U -> List simp [NonTopNQ] at * intros l u v rw [privComposeChainRule] - apply ENNReal.mul_ne_top - · aesop - · aesop + apply ENNReal.mul_ne_top <;> + simp_all only [ne_eq, not_false_eq_true] /-- Adaptive composed query is a proper distribution @@ -152,8 +150,8 @@ theorem privComposeAdaptive_NonTopSum {nq1 : List T → SLang U} {nq2 : U -> Lis apply ENNReal.tsum_le_tsum intro a apply mul_le_mul_of_nonneg_left - · aesop - · aesop + · simp_all + · simp_all only [_root_.zero_le] rw [ENNReal.tsum_mul_right] apply ne_top_lt_top intro Hcont @@ -161,10 +159,10 @@ theorem privComposeAdaptive_NonTopSum {nq1 : List T → SLang U} {nq2 : U -> Lis cases Hcont · rename_i h rcases h with ⟨ _ , h1 ⟩ - aesop + simp_all only [le_top, implies_true, ne_eq, ENNReal.tsum_eq_zero, not_forall, ofReal_ne_top] · rename_i h rcases h with ⟨ h0 , _ ⟩ - aesop + simp_all only /-- Renyi divergence beteeen adaptive composed queries on neighbours are finite. @@ -360,16 +358,7 @@ lemma privComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> Lis apply privComposeAdaptive_NonTopNQ <;> aesop case H2 => apply ne_top_lt_top - apply privComposeAdaptive_NonTopRDNQ - · apply HNTRDNQ1 - · apply HNTRDNQ2 - · apply HNT1 - · apply HNT2 - · apply HNZ1 - · apply HNZ2 - · apply Hubound2 - · apply Hα - · apply HN + apply privComposeAdaptive_NonTopRDNQ <;> aesop rw [left_distrib] rw [Real.exp_add] @@ -403,14 +392,12 @@ lemma privComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> Lis let _ := (Hubound u) apply Real.exp_le_exp_of_le aesop - rw [mul_comm] rw [<- (ENNReal.toReal_ofReal_mul _ _ ?h1)] case h1 => exact exp_nonneg ((α - 1) * b) rw [mul_comm] rw [← ENNReal.tsum_mul_right] - apply (toReal_mono' _ ?goal2) case goal2 => intro H @@ -468,13 +455,13 @@ lemma privComposeAdaptive_renyi_bound {nq1 : List T → SLang U} {nq2 : U -> Lis · exact pos_iff_ne_zero.mpr (HNZ2 i l₁ v0) · apply HNT2 intro Hk - aesop + simp_all only [exp_le_exp, gt_iff_lt, sub_pos, _root_.mul_le_mul_left, lt_self_iff_false] · have Hlt : (0 < nq2 i l₂ v0 ^ (1 - α)) := by apply ENNReal.rpow_pos · exact pos_iff_ne_zero.mpr (HNZ2 i l₂ v0) · exact HNT2 i l₂ v0 intro Hk - aesop + simp_all only [exp_le_exp, gt_iff_lt, sub_pos, _root_.mul_le_mul_left, lt_self_iff_false] have GH2 : ∀ i, ∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α) < ⊤ := by exact fun i => ne_top_lt_top (∑' (x : V), nq2 i l₁ x ^ α * nq2 i l₂ x ^ (1 - α)) (HNTRDNQ2 i α Hα l₁ l₂ HN) @@ -596,11 +583,11 @@ theorem privComposeAdaptive_zCDPBound {nq1 : List T → SLang U} {nq2 : U -> Lis theorem privComposeAdaptive_zCDP {nq1 : List T → SLang U} {nq2 : U -> List T → SLang V} {ε₁ ε₂ ε₃ ε₄ : ℕ+} (h : zCDP nq1 ((ε₁ : ℝ) / ε₂)) (h' : ∀ u, zCDP (nq2 u) ((ε₃ : ℝ) / ε₄)) - (Hubound_nq2 : ∃ (z : ℝ), ∀ (u : U), ∀ l, ∑'(v : V), nq2 u l v ≤ ENNReal.ofReal z) - (HB' : URDBound nq2 (fun α l₁ l₂ => ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α)) : -- Can I get rid of URDBound? + (NonTopSum_Uniform : ∃ (z : ℝ), ∀ (u : U), ∀ l, ∑'(v : V), nq2 u l v ≤ ENNReal.ofReal z) + (RenyiBound_Uniform : URDBound nq2 (fun α l₁ l₂ => ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α)) : zCDP (privComposeAdaptive nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by simp [zCDP] at * - have _ : RDBounded nq2 := RDBounded_of_URDBound nq2 (fun α l₁ l₂ => ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) HB' + have _ : RDBounded nq2 := RDBounded_of_URDBound nq2 (fun α l₁ l₂ => ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) RenyiBound_Uniform repeat any_goals constructor · apply privComposeAdaptive_zCDPBound <;> aesop · apply privComposeAdaptive_NonZeroNQ <;> aesop From 2312e142668654eca4bbbd37ab8ce08194e114ac Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Mon, 10 Jun 2024 13:20:53 -0400 Subject: [PATCH 037/122] extend system --- SampCert/DifferentialPrivacy/Abstract.lean | 12 +++++++++++- SampCert/DifferentialPrivacy/Pure/System.lean | 6 ++++++ .../DifferentialPrivacy/ZeroConcentrated/System.lean | 6 ++++++ 3 files changed, 23 insertions(+), 1 deletion(-) diff --git a/SampCert/DifferentialPrivacy/Abstract.lean b/SampCert/DifferentialPrivacy/Abstract.lean index 39a5bea2..169dc82c 100644 --- a/SampCert/DifferentialPrivacy/Abstract.lean +++ b/SampCert/DifferentialPrivacy/Abstract.lean @@ -101,7 +101,7 @@ class DPSystem (T : Type) where prop : Mechanism T Z → ℝ → Prop /-- Noise mechanism (eg. Laplace, Discrete Gaussian, etc) - Paramaterized by a query, sensitivity, and the numerator/denominator ofa (rational) security paramater. + Paramaterized by a query, sensitivity, and a rational security paramater. -/ noise : Query T ℤ → ℕ+ → ℕ+ → ℕ+ → Mechanism T ℤ /-- @@ -114,6 +114,15 @@ class DPSystem (T : Type) where compose_prop : {U V : Type} → [MeasurableSpace U] → [Countable U] → [DiscreteMeasurableSpace U] → [Inhabited U] → [MeasurableSpace V] → [Countable V] → [DiscreteMeasurableSpace V] → [Inhabited V] → ∀ m₁ : Mechanism T U, ∀ m₂ : Mechanism T V, ∀ ε₁ ε₂ ε₃ ε₄ : ℕ+, prop m₁ (ε₁ / ε₂) → prop m₂ (ε₃ / ε₄) → prop (privCompose m₁ m₂) ((ε₁ / ε₂) + (ε₃ / ε₄)) /-- + Uniformity conditions for adaptive composition + -/ + adaptive_unif_prop {U : Type} : (U -> Mechanism T Z) -> Prop + /-- + Notion of privacy composes by addition + -/ + adaptive_compose_prop : {U V : Type} → [MeasurableSpace U] → [Countable U] → [DiscreteMeasurableSpace U] → [Inhabited U] → [MeasurableSpace V] → [Countable V] → [DiscreteMeasurableSpace V] → [Inhabited V] → ∀ m₁ : Mechanism T U, ∀ m₂ : U -> Mechanism T V, ∀ ε₁ ε₂ ε₃ ε₄ : ℕ+, + prop m₁ (ε₁ / ε₂) → (∀ u, prop (m₂ u) (ε₃ / ε₄)) -> adaptive_unif_prop m₂ → prop (privComposeAdaptive m₁ m₂) ((ε₁ / ε₂) + (ε₃ / ε₄)) + /-- Notion of privacy is invariant under post-processing. -/ postprocess_prop : {U : Type} → [MeasurableSpace U] → [Countable U] → [DiscreteMeasurableSpace U] → [Inhabited U] → { pp : U → V } → Function.Surjective pp → ∀ m : Mechanism T U, ∀ ε₁ ε₂ : ℕ+, @@ -215,6 +224,7 @@ lemma compose_sum_rw (nq1 : List T → SLang U) (nq2 : List T → SLang V) (b : rw [C] simp +-- lemma compose_sum_rw_adaptive (nq1 : List T → SLang U) (nq2 : U -> List T → SLang V) (b : U) (c : V) (l : List T) : (∑' (a : U), nq1 l a * ∑' (a_1 : V), if x = a ∧ y = a_1 then nq2 a l a_1 else 0) = nq1 l x * nq2 x l y := sorry diff --git a/SampCert/DifferentialPrivacy/Pure/System.lean b/SampCert/DifferentialPrivacy/Pure/System.lean index fc3a415b..999956ec 100644 --- a/SampCert/DifferentialPrivacy/Pure/System.lean +++ b/SampCert/DifferentialPrivacy/Pure/System.lean @@ -19,6 +19,12 @@ noncomputable instance PureDPSystem : DPSystem T where noise := privNoisedQueryPure noise_prop := NoisedQuery_PureDP compose_prop := PureDP_Compose + adaptive_unif_prop := fun _ => True + adaptive_compose_prop := fun {U V} + [MeasurableSpace U] [Countable U] [DiscreteMeasurableSpace U] [Inhabited U] + [MeasurableSpace V] [Countable V] [DiscreteMeasurableSpace V] [Inhabited V] + m₁ m₂ ε₁ ε₂ ε₃ ε₄ a a_1 _ => + PureDP_ComposeAdaptive' m₁ m₂ ε₁ ε₂ ε₃ ε₄ a a_1 postprocess_prop := PureDP_PostProcess end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean index 35977cbb..2435b257 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean @@ -28,6 +28,12 @@ noncomputable instance gaussian_zCDPSystem : DPSystem T where noise := privNoisedQuery noise_prop := privNoisedQuery_zCDP compose_prop := privCompose_zCDP + adaptive_unif_prop := fun {Z U} nq2 => + ((∃ (z : ℝ), ∀ (u : U), ∀ l, ∑'(v : Z), nq2 u l v ≤ ENNReal.ofReal z) ∧ + (URDBound nq2 (fun α l₁ l₂ => ⨆ u, RenyiDivergence (nq2 u l₁) (nq2 u l₂) α))) + adaptive_compose_prop := by + intros + apply privComposeAdaptive_zCDP <;> aesop postprocess_prop := privPostProcess_zCDP end SLang From e62af89ebab7fe34f04772429ef019e219a1e85e Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Mon, 10 Jun 2024 13:35:28 -0400 Subject: [PATCH 038/122] fix sorry --- SampCert/DifferentialPrivacy/Abstract.lean | 29 +++++++++++----------- 1 file changed, 14 insertions(+), 15 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Abstract.lean b/SampCert/DifferentialPrivacy/Abstract.lean index 169dc82c..6a77ad07 100644 --- a/SampCert/DifferentialPrivacy/Abstract.lean +++ b/SampCert/DifferentialPrivacy/Abstract.lean @@ -37,18 +37,8 @@ def privComposeAdaptive (nq1 : Mechanism T U) (nq2 : U -> Mechanism T V) (l : Li return (A, B) --- set_option pp.notation false - -/-- -Chain rule relating the adaptive composition definitions - -The joint distribution decomposes into the conditional and marginal (ie, nq1 l) distributions --/ -lemma privComposeChainRule (nq1 : Mechanism T U) (nq2 : U -> Mechanism T V) (l : List T) : - ∀ (u : U), ∀ (v : V), privComposeAdaptive nq1 nq2 l (u, v) = nq1 l u * nq2 u l v := by - intros u v - simp [privComposeAdaptive] - -- How to simplify cases without explicit ite? this is silly +lemma compose_sum_rw_adaptive (nq1 : List T → SLang U) (nq2 : U -> List T → SLang V) (u : U) (v : V) (l : List T) : + (∑' (a : U), nq1 l a * ∑' (a_1 : V), if u = a ∧ v = a_1 then nq2 a l a_1 else 0) = nq1 l u * nq2 u l v := by have hrw1 : ∀ (a : U), nq1 l a * (∑' (a_1 : V), if u = a ∧ v = a_1 then nq2 a l a_1 else 0) = if (u = a) then (nq1 l a * ∑' (a_1 : V), if u = a ∧ v = a_1 then nq2 a l a_1 else 0) else 0 := by intro a split @@ -75,6 +65,18 @@ lemma privComposeChainRule (nq1 : Mechanism T U) (nq2 : U -> Mechanism T V) (l : rw [tsum_ite_eq] exact MulOneClass.mul_one (nq2 u l v) + +/-- +Chain rule relating the adaptive composition definitions + +The joint distribution decomposes into the conditional and marginal (ie, nq1 l) distributions +-/ +lemma privComposeChainRule (nq1 : Mechanism T U) (nq2 : U -> Mechanism T V) (l : List T) : + ∀ (u : U), ∀ (v : V), privComposeAdaptive nq1 nq2 l (u, v) = nq1 l u * nq2 u l v := by + intros u v + rw [<- compose_sum_rw_adaptive] + simp [privComposeAdaptive] + /-- Composition of independent mechanisms -/ @@ -224,9 +226,6 @@ lemma compose_sum_rw (nq1 : List T → SLang U) (nq2 : List T → SLang V) (b : rw [C] simp --- -lemma compose_sum_rw_adaptive (nq1 : List T → SLang U) (nq2 : U -> List T → SLang V) (b : U) (c : V) (l : List T) : - (∑' (a : U), nq1 l a * ∑' (a_1 : V), if x = a ∧ y = a_1 then nq2 a l a_1 else 0) = nq1 l x * nq2 x l y := sorry /-- Composed queries are normalizable. From 166e321c8c32a52d5ab96798e229f640c8691d14 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Thu, 13 Jun 2024 16:43:34 -0400 Subject: [PATCH 039/122] elog and eexp --- .../DifferentialPrivacy/RenyiDivergence.lean | 3 +- SampCert/Util/Log.lean | 156 ++++++++++++++++++ 2 files changed, 158 insertions(+), 1 deletion(-) create mode 100644 SampCert/Util/Log.lean diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index be404036..77ea6cc1 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -6,6 +6,7 @@ Authors: Jean-Baptiste Tristan import Mathlib.Topology.Algebra.InfiniteSum.Ring import Mathlib.NumberTheory.ModularForms.JacobiTheta.TwoVariable import Mathlib.Analysis.SpecialFunctions.Pow.NNReal +import SampCert.Util.Log /-! Renyi Divergence @@ -27,7 +28,7 @@ noncomputable def RenyiDivergence (p q : T → ENNReal) (α : ℝ) : ℝ := Closed form of the series in the definition of hte Renyi divergence. -/ 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 + (∑' x : T, (p x)^α * (q x)^(1 - α)) = ∑' x: T, (p x / q x)^α * (q x) := by congr 4 ext x rw [ENNReal.rpow_sub] diff --git a/SampCert/Util/Log.lean b/SampCert/Util/Log.lean new file mode 100644 index 00000000..29d5f1f1 --- /dev/null +++ b/SampCert/Util/Log.lean @@ -0,0 +1,156 @@ +/- +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, Markus de Medeiros +-/ + +import Mathlib.Data.ENNReal.Basic +import Mathlib.Data.Real.EReal +import Mathlib.Analysis.SpecialFunctions.Log.Basic +import Mathlib.Analysis.SpecialFunctions.Pow.NNReal + +/-! +# Logarithm on ENNReal + +In this file we extend the logarithm to ``ENNReal``. +-/ + +noncomputable section + +open scoped Classical +open ENNReal EReal Real + +namespace ENNReal + +/-- +The extended logarithm +-/ +def elog (x : ENNReal) : EReal := + match x with + | none => ⊤ + | some r => if r = 0 then ⊥ else Real.log r + +/-- +The extended exponential + +Mathlib's has an extended ``rpow`` function of type ``ℝ≥0∞ → ℝ → ℝ≥0∞``, however we +want the exponent to be of type ``EReal``. +-/ +def eexp (y : EReal) : ENNReal := + match y with + | none => 0 + | (some none) => ⊤ + | (some (some r)) => ENNReal.ofReal (Real.exp r) + +variable {x y : ENNReal} +variable {w z : EReal} +variable {r : Real} + +@[simp] +lemma elog_of_pos_real (H : 0 < r) : elog (ENNReal.ofReal r) = Real.log r := by + rw [elog] + split + · sorry + · split + · rename_i r heq h + simp + sorry + · sorry + +@[simp] +lemma elog_zero : elog (ENNReal.ofReal 0) = ⊥ := by sorry + +@[simp] +lemma elog_top : elog ⊤ = ⊤ := by sorry + +@[simp] +lemma eexp_bot : eexp ⊥ = 0 := by sorry + +@[simp] +lemma eexp_top : eexp ⊤ = ⊤ := by sorry + +@[simp] +lemma eexp_pos_real (H : 0 < r) : eexp r = ENNReal.ofReal (Real.log r) := by sorry + +@[simp] +lemma elog_eexp : eexp (elog x) = x := by + rw [elog] + split + · simp + · rename_i _ r' + split + · simp + rename_i _ h + rw [h] + · rename_i _ H + simp + rw [eexp] + rw [NNReal.toReal, Real.toEReal] + simp + rw [Real.exp_log] + rw [ofReal_coe_nnreal] + rcases r' with ⟨ v , Hv ⟩ + apply lt_of_le_of_ne + · simpa + · simp + intro Hk + apply H + apply NNReal.coe_eq_zero.mp + simp + rw [Hk] + +@[simp] +lemma eexp_elog : (elog (eexp w)) = w := by + sorry + +@[simp] +lemma elog_mul : elog x * elog y = elog (x + y) := by sorry -- checked truth table + +@[simp] +lemma eexp_add : eexp w * eexp z = eexp (w + z) := by sorry -- checked truth table + +-- Log of power, log and exp inverses + +lemma eexp_injective : eexp w = eexp z -> w = z := by + rw [eexp, eexp] + intro H + cases w <;> cases z <;> try tauto + · rename_i v + cases v <;> simp at * + sorry + · rename_i v + cases v <;> simp at * + sorry + · rename_i v₁ v₂ + cases v₁ <;> cases v₂ <;> simp at * + congr + sorry + +lemma elog_injective : elog x = elog y -> x = y := by + rw [elog, elog] + cases x <;> cases y <;> try tauto + · rename_i v + cases v; simp at * + sorry + · rename_i v + cases v; simp at * + sorry + · rename_i v₁ v₂ + cases v₁; cases v₂; simp at * + sorry + +lemma eexp_mono_lt : (w < z) <-> eexp w < eexp z := by + + sorry + +lemma eexp_mono_le : (w <= z) <-> eexp w <= eexp z := by sorry + +lemma elog_mono_lt : (x < y) <-> elog x < elog y := by sorry + +lemma elog_mono_le : (x <= y) <-> elog x <= elog y := by sorry + +-- iff for log positive + +-- Special values + +end ENNReal From 0e711205919388be89f49c89f1450c6ee94bc28f Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Fri, 14 Jun 2024 09:24:04 -0400 Subject: [PATCH 040/122] Change Renyi divergence type and stub proofs --- SampCert/DifferentialPrivacy/RenyiDivergence.lean | 4 ++-- .../DifferentialPrivacy/ZeroConcentrated/Composition.lean | 3 +++ .../ZeroConcentrated/ConcentratedBound.lean | 7 +++++-- SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean | 2 +- .../ZeroConcentrated/Mechanism/Properties.lean | 3 +++ .../ZeroConcentrated/Postprocessing.lean | 3 +++ 6 files changed, 17 insertions(+), 5 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index 77ea6cc1..98da2f2a 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -21,8 +21,8 @@ variable {T : Type} /-- The Renyi divergence. -/ -noncomputable def RenyiDivergence (p q : T → ENNReal) (α : ℝ) : ℝ := - (α - 1)⁻¹ * Real.log (∑' x : T, (p x)^α * (q x)^(1 - α)).toReal +noncomputable def RenyiDivergence (p q : T → ENNReal) (α : ℝ) : ENNReal := sorry + -- (α - 1)⁻¹ * elog (∑' x : T, (p x)^α * (q x)^(1 - α)) /-- Closed form of the series in the definition of hte Renyi divergence. diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean index 92fec776..b2ab5a37 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean @@ -77,6 +77,7 @@ Composed queries satisfy zCDP Renyi divergence bound. -/ theorem privCompose_zCDPBound {nq1 : List T → SLang U} {nq2 : List T → SLang V} {ε₁ ε₂ ε₃ ε₄ : ℕ+} (h1 : zCDPBound nq1 ((ε₁ : ℝ) / ε₂)) (h2 : zCDPBound nq2 ((ε₃ : ℝ) / ε₄)) (nn1 : NonZeroNQ nq1) (nn2 : NonZeroNQ nq2) (nt1 : NonTopRDNQ nq1) (nt2 : NonTopRDNQ nq2) (nts1 : NonTopNQ nq1) (nts2 : NonTopNQ nq2) : zCDPBound (privCompose nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by + /- simp [privCompose, RenyiDivergence, zCDPBound] intro α h3 l₁ l₂ h4 have X := h1 @@ -140,6 +141,8 @@ theorem privCompose_zCDPBound {nq1 : List T → SLang U} {nq2 : List T → SLang . apply lt_trans zero_lt_one h3 . apply Renyi_noised_query_NZ h3 h4 X nn1 nt1 nts1 . apply Renyi_noised_query_NZ h3 h4 Y nn2 nt2 nts2 + -/ + sorry /-- All outputs of a composed query have finite probability. diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean index 04cbc211..6f4f4c44 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean @@ -532,7 +532,8 @@ Upper bound on Renyi divergence between discrete Gaussians. theorem Renyi_Gauss_divergence_bound' {σ α : ℝ} (h : σ ≠ 0) (h' : 1 < α) (μ ν : ℤ) : RenyiDivergence (fun (x : ℤ) => ENNReal.ofReal (discrete_gaussian σ μ x)) (fun (x : ℤ) => ENNReal.ofReal (discrete_gaussian σ ν x)) - α ≤ α * (((μ - ν) : ℤ)^2 / (2 * σ^2)) := by + α ≤ (ENNReal.ofReal α) * (ENNReal.ofReal ((((μ - ν) : ℤ)^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) @@ -563,6 +564,8 @@ theorem Renyi_Gauss_divergence_bound' {σ α : ℝ} (h : σ ≠ 0) (h' : 1 < α) . apply Renyi_Gauss_summable h rw [A] apply Renyi_divergence_bound_pre h h' + -/ + sorry namespace SLang @@ -571,7 +574,7 @@ Upper bound on Renyi divergence between outputs of the ``SLang`` discrete Gaussi -/ theorem discrete_GaussianGenSample_ZeroConcentrated {α : ℝ} (h : 1 < α) (num : PNat) (den : PNat) (μ ν : ℤ) : RenyiDivergence ((DiscreteGaussianGenSample num den μ)) (DiscreteGaussianGenSample num den ν) α ≤ - α * (((μ - ν) : ℤ)^2 / (2 * ((num : ℝ) / (den : ℝ))^2)) := by + (ENNReal.ofReal α) * (ENNReal.ofReal ((((μ - ν) : ℤ)^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 => diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean index 1d8dda8d..07c18561 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean @@ -33,7 +33,7 @@ satisfying this bound are ``ε``-DP). -/ def zCDPBound (q : List T → SLang U) (ε : ℝ) : Prop := ∀ α : ℝ, 1 < α → ∀ l₁ l₂ : List T, Neighbour l₁ l₂ → - RenyiDivergence (q l₁) (q l₂) α ≤ (1/2) * ε ^ 2 * α + RenyiDivergence (q l₁) (q l₂) α ≤ ENNReal.ofReal ((1/2) * ε ^ 2 * α) def NonTopNQ (nq : List T → SLang U) := ∀ l : List T, ∀ n : U, nq l n ≠ ⊤ diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean index 393dc881..5fc90bdf 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean @@ -23,6 +23,7 @@ The zCDP mechanism with bounded sensitivity satisfies the bound for ``(Δε₂/ -/ theorem privNoisedQuery_zCDPBound (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_sensitivity : sensitivity query Δ) : zCDPBound (privNoisedQuery query Δ ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by + /- simp [zCDPBound, privNoisedQuery] intros α h1 l₁ l₂ h2 have A := @discrete_GaussianGenSample_ZeroConcentrated α h1 (Δ * ε₂) ε₁ (query l₁) (query l₂) @@ -72,6 +73,8 @@ theorem privNoisedQuery_zCDPBound (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ rw [_root_.mul_pos_iff] left simp + -/ + sorry /-- All outputs of the zCDP mechanism have nonzero probability. diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean index c89fff25..c4e58001 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean @@ -551,6 +551,7 @@ theorem tsum_ne_zero_of_ne_zero {T : Type} [Inhabited T] (f : T → ENNReal) (h theorem privPostProcess_zCDPBound {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : zCDPBound nq ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (conv : NonTopSum nq) (f : U → V) : zCDPBound (privPostProcess nq f) ((ε₁ : ℝ) / ε₂) := by + /- simp [privPostProcess, zCDPBound, RenyiDivergence] intro α h1 l₁ l₂ h2 have h' := h @@ -636,6 +637,8 @@ theorem privPostProcess_zCDPBound {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} apply lt_of_le_of_lt X Y rw [lt_top_iff_ne_top] at Z contradiction + -/ + sorry theorem privPostProcess_NonTopRDNQ {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (f : U → V) (dp :zCDPBound nq ((ε₁ : ℝ) / ε₂)) (nt : NonTopRDNQ nq) (nz : NonZeroNQ nq) (nts : NonTopNQ nq) (ntsum: NonTopSum nq) : NonTopRDNQ (privPostProcess nq f) := by From 7c64a91b515f6f4e9ec2b4a9dafc272e25667c68 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Fri, 14 Jun 2024 10:57:41 -0400 Subject: [PATCH 041/122] Define ENReal -> ENNReal coercion --- .../DifferentialPrivacy/RenyiDivergence.lean | 25 ++++++++++++++++--- .../ZeroConcentrated/ConcentratedBound.lean | 2 +- SampCert/Util/Log.lean | 24 +++++++++++++++++- 3 files changed, 45 insertions(+), 6 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index 98da2f2a..f4923515 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -7,6 +7,7 @@ import Mathlib.Topology.Algebra.InfiniteSum.Ring import Mathlib.NumberTheory.ModularForms.JacobiTheta.TwoVariable import Mathlib.Analysis.SpecialFunctions.Pow.NNReal import SampCert.Util.Log +import Mathlib.Probability.ProbabilityMassFunction.Basic /-! Renyi Divergence @@ -14,18 +15,34 @@ This file defines the Renyi divergence and equations for evaluating its expectat -/ -open Real ENNReal +open Real ENNReal PMF variable {T : Type} +noncomputable def RenyiDivergence_def (p q : PMF T) (α : ℝ) : EReal := + (α - 1)⁻¹ * elog (∑' x : T, (p x)^α * (q x)^(1 - α)) + +theorem RenyiDivergence_def_nonneg (p q : PMF T) (α : ℝ) : (0 ≤ RenyiDivergence_def p q α) := by + -- See paper + sorry + +theorem RenyiDivergence_def_zero (p q : PMF T) (α : ℝ) : p = q <-> (0 = RenyiDivergence_def p q α) := by + -- See paper + sorry + /-- The Renyi divergence. -/ -noncomputable def RenyiDivergence (p q : T → ENNReal) (α : ℝ) : ENNReal := sorry - -- (α - 1)⁻¹ * elog (∑' x : T, (p x)^α * (q x)^(1 - α)) +noncomputable def RenyiDivergence (p q : PMF T) (α : ℝ) : ENNReal := + ENNReal.ofEReal (RenyiDivergence_def p q α) + +-- MARKUSDE: We evaluate through the ENNReal.ofEReal using ``RenyiDivergence_def_nonneg``, these are some special cases +theorem RenyiDivergence_aux_zero (p q : PMF T) (α : ℝ) : p = q <-> RenyiDivergence p q α = 0 + := sorry + /-- -Closed form of the series in the definition of hte Renyi divergence. +Closed form of the series in the definition of the Renyi divergence. -/ 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 diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean index 6f4f4c44..41bd6570 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean @@ -574,7 +574,7 @@ Upper bound on Renyi divergence between outputs of the ``SLang`` discrete Gaussi -/ theorem discrete_GaussianGenSample_ZeroConcentrated {α : ℝ} (h : 1 < α) (num : PNat) (den : PNat) (μ ν : ℤ) : RenyiDivergence ((DiscreteGaussianGenSample num den μ)) (DiscreteGaussianGenSample num den ν) α ≤ - (ENNReal.ofReal α) * (ENNReal.ofReal ((((μ - ν) : ℤ)^2 : ℝ) / (2 * ((num : ℝ) / (den : ℝ))^2))) := by + (ENNReal.ofReal α) * (ENNReal.ofReal ((((μ - ν) : ℤ)^2 : ENNReal) / (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 => diff --git a/SampCert/Util/Log.lean b/SampCert/Util/Log.lean index 29d5f1f1..403363f8 100644 --- a/SampCert/Util/Log.lean +++ b/SampCert/Util/Log.lean @@ -22,6 +22,17 @@ open ENNReal EReal Real namespace ENNReal +/-- +Coerce a EReal to an ENNReal by truncation +-/ +noncomputable def ofEReal (e : EReal) : ENNReal := + match e with + | none => some 0 + | some none => none + | some (some r) => ENNReal.ofReal r + +instance : Coe EReal ENNReal := ⟨ofEReal⟩ + /-- The extended logarithm -/ @@ -42,6 +53,7 @@ def eexp (y : EReal) : ENNReal := | (some none) => ⊤ | (some (some r)) => ENNReal.ofReal (Real.exp r) + variable {x y : ENNReal} variable {w z : EReal} variable {r : Real} @@ -140,7 +152,6 @@ lemma elog_injective : elog x = elog y -> x = y := by sorry lemma eexp_mono_lt : (w < z) <-> eexp w < eexp z := by - sorry lemma eexp_mono_le : (w <= z) <-> eexp w <= eexp z := by sorry @@ -153,4 +164,15 @@ lemma elog_mono_le : (x <= y) <-> elog x <= elog y := by sorry -- Special values +-- Need to write compat lemmas for ofEReal +-- Not sure which ones we'll need but + + +-- Specialized lemmas for ofEReal when its argument is nonnegative (so no truncation happens) +lemma ofEReal_nonneg_eq_iff (Hw : 0 <= w) (Hz : 0 <= z) : w = z <-> (ofEReal w = ofEReal z) := + sorry + +lemma ofEReal_le_mono : w ≤ z <-> (ofEReal w ≤ ofEReal z) := + sorry + end ENNReal From ee2d88d01f6a1db713848841b9cb8fdeda9ba1a1 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Fri, 14 Jun 2024 11:15:58 -0400 Subject: [PATCH 042/122] copy RenyiDivergence_def_exp proof from other branch --- .../DifferentialPrivacy/RenyiDivergence.lean | 61 +++++++++++++++---- SampCert/Util/Log.lean | 8 ++- 2 files changed, 55 insertions(+), 14 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index f4923515..d20a8e09 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -40,7 +40,7 @@ noncomputable def RenyiDivergence (p q : PMF T) (α : ℝ) : ENNReal := theorem RenyiDivergence_aux_zero (p q : PMF T) (α : ℝ) : p = q <-> RenyiDivergence p q α = 0 := sorry - +-- MARKUSDE: fixme /-- Closed form of the series in the definition of the Renyi divergence. -/ @@ -56,17 +56,52 @@ theorem RenyiDivergenceExpectation (p q : T → ENNReal) {α : ℝ} (h : 1 < α) . apply h1 x . apply h2 x + +-- Unused +/- +Closed form for the Renyi Divergence. +-/ +-- 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 + + /-- -Closed form for the Renti Divergence. +The Renyi divergence is monotonic in the value of its sum. -/ -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 +lemma RenyiDivergence_mono_sum (x y : ℝ) (α : ℝ) (h : 1 < α) : (Real.exp ((α - 1) * x)) ≤ (Real.exp ((α - 1) * y)) -> (x ≤ y) := by + intro H + apply le_of_mul_le_mul_left + · exact exp_le_exp.mp H + · linarith + +/-- +Equation for the Renyi divergence series in terms of the Renyi Divergence +-/ +lemma RenyiDivergence_def_exp (p q : PMF T) {α : ℝ} + (h : 1 < α) + (H1 : 0 < ∑' (x : T), p x ^ α * q x ^ (1 - α)) + (H2 : ∑' (x : T), p x ^ α * q x ^ (1 - α) < ⊤) : + eexp ((α - 1) * RenyiDivergence_def p q α) = (∑' x : T, (p x)^α * (q x)^(1 - α)) := by + sorry + /- + simp only [RenyiDivergence] + rw [<- mul_assoc] + have test : (α - 1) * (α - 1)⁻¹ = 1 := by + refine mul_inv_cancel ?h + linarith + rw [test] + clear test + simp + rw [Real.exp_log] + apply ENNReal.toReal_pos_iff.mpr + tauto + -/ diff --git a/SampCert/Util/Log.lean b/SampCert/Util/Log.lean index 403363f8..97ffdb78 100644 --- a/SampCert/Util/Log.lean +++ b/SampCert/Util/Log.lean @@ -31,7 +31,7 @@ noncomputable def ofEReal (e : EReal) : ENNReal := | some none => none | some (some r) => ENNReal.ofReal r -instance : Coe EReal ENNReal := ⟨ofEReal⟩ +-- instance : Coe EReal ENNReal := ⟨ofEReal⟩ /-- The extended logarithm @@ -175,4 +175,10 @@ lemma ofEReal_nonneg_eq_iff (Hw : 0 <= w) (Hz : 0 <= z) : w = z <-> (ofEReal w = lemma ofEReal_le_mono : w ≤ z <-> (ofEReal w ≤ ofEReal z) := sorry +@[simp] +lemma toEReal_ofENNReal_nonneg (H : 0 ≤ w) : ENNReal.toEReal (ofEReal w) = w := sorry + +@[simp] +lemma ofEReal_toENNReal : ofEReal (ENNReal.toEReal x) = x := by sorry + end ENNReal From 8841b84f75de54780046b9c74afb6ba4609ea617 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Fri, 14 Jun 2024 11:57:37 -0400 Subject: [PATCH 043/122] Implement RenyiDivergence_def_exp --- .../DifferentialPrivacy/RenyiDivergence.lean | 81 +++++++++++-------- 1 file changed, 47 insertions(+), 34 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index d20a8e09..5e00f459 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -20,7 +20,8 @@ open Real ENNReal PMF variable {T : Type} noncomputable def RenyiDivergence_def (p q : PMF T) (α : ℝ) : EReal := - (α - 1)⁻¹ * elog (∑' x : T, (p x)^α * (q x)^(1 - α)) + (α - 1)⁻¹ * (ENNReal.ofEReal (elog (∑' x : T, (p x)^α * (q x)^(1 - α)))) + theorem RenyiDivergence_def_nonneg (p q : PMF T) (α : ℝ) : (0 ≤ RenyiDivergence_def p q α) := by -- See paper @@ -30,6 +31,51 @@ theorem RenyiDivergence_def_zero (p q : PMF T) (α : ℝ) : p = q <-> (0 = Renyi -- See paper sorry +lemma RenyiDivergence_def_log_sum_nonneg (p q : PMF T) (α : ℝ) : (0 ≤ (elog (∑' x : T, (p x)^α * (q x)^(1 - α)))) := by + -- Follows from RenyiDivergence_def_nonneg + sorry + +/-- +The Renyi divergence is monotonic in the value of its sum. +-/ +lemma RenyiDivergence_mono_sum (x y : ℝ) (α : ℝ) (h : 1 < α) : (Real.exp ((α - 1) * x)) ≤ (Real.exp ((α - 1) * y)) -> (x ≤ y) := by + intro H + apply le_of_mul_le_mul_left + · exact exp_le_exp.mp H + · linarith + +/-- +Equation for the Renyi divergence series in terms of the Renyi Divergence +-/ +lemma RenyiDivergence_def_exp (p q : PMF T) {α : ℝ} + (h : 1 < α) + (H1 : 0 < ∑' (x : T), p x ^ α * q x ^ (1 - α)) + (H2 : ∑' (x : T), p x ^ α * q x ^ (1 - α) < ⊤) : + eexp (((α - 1)) * RenyiDivergence_def p q α) = (∑' x : T, (p x)^α * (q x)^(1 - α)) := by + simp only [RenyiDivergence_def] + rw [<- mul_assoc] + have Hinvert : ((↑α - 1) * ↑(α - 1)⁻¹ : EReal) = 1 := by + clear H1 + clear H2 + have Hsub : ((↑α - 1) : EReal) = (α - 1 : ℝ) := by simp + rw [Hsub] + have Hundo_coe (r1 r2 : ℝ) : (r1 : EReal) * (r2 : EReal) = ((r1 * r2 : ℝ) : EReal) := by + rw [EReal.coe_mul] + rw [Hundo_coe] + have Hinv : (α - 1) * (α - 1)⁻¹ = 1 := by + apply mul_inv_cancel + linarith + rw [Hinv] + simp + rw [Hinvert] + clear Hinvert + simp + rw [toEReal_ofENNReal_nonneg ?Helog_nn] + case Helog_nn => + apply RenyiDivergence_def_log_sum_nonneg + apply elog_eexp + + /-- The Renyi divergence. -/ @@ -72,36 +118,3 @@ Closed form for the Renyi Divergence. -- . apply le_of_lt (lt_trans Real.zero_lt_one h ) -- . apply h1 x -- . apply h2 x - - -/-- -The Renyi divergence is monotonic in the value of its sum. --/ -lemma RenyiDivergence_mono_sum (x y : ℝ) (α : ℝ) (h : 1 < α) : (Real.exp ((α - 1) * x)) ≤ (Real.exp ((α - 1) * y)) -> (x ≤ y) := by - intro H - apply le_of_mul_le_mul_left - · exact exp_le_exp.mp H - · linarith - -/-- -Equation for the Renyi divergence series in terms of the Renyi Divergence --/ -lemma RenyiDivergence_def_exp (p q : PMF T) {α : ℝ} - (h : 1 < α) - (H1 : 0 < ∑' (x : T), p x ^ α * q x ^ (1 - α)) - (H2 : ∑' (x : T), p x ^ α * q x ^ (1 - α) < ⊤) : - eexp ((α - 1) * RenyiDivergence_def p q α) = (∑' x : T, (p x)^α * (q x)^(1 - α)) := by - sorry - /- - simp only [RenyiDivergence] - rw [<- mul_assoc] - have test : (α - 1) * (α - 1)⁻¹ = 1 := by - refine mul_inv_cancel ?h - linarith - rw [test] - clear test - simp - rw [Real.exp_log] - apply ENNReal.toReal_pos_iff.mpr - tauto - -/ From 202a6b116e00184cae80e97689b41312f6c9a4d8 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Fri, 14 Jun 2024 14:55:49 -0400 Subject: [PATCH 044/122] prove closed form for expectation formula (absolute continuity assumption) --- .../DifferentialPrivacy/RenyiDivergence.lean | 136 ++++++++++++------ 1 file changed, 96 insertions(+), 40 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index 5e00f459..df651cc0 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -19,30 +19,13 @@ open Real ENNReal PMF variable {T : Type} -noncomputable def RenyiDivergence_def (p q : PMF T) (α : ℝ) : EReal := - (α - 1)⁻¹ * (ENNReal.ofEReal (elog (∑' x : T, (p x)^α * (q x)^(1 - α)))) - - -theorem RenyiDivergence_def_nonneg (p q : PMF T) (α : ℝ) : (0 ≤ RenyiDivergence_def p q α) := by - -- See paper - sorry - -theorem RenyiDivergence_def_zero (p q : PMF T) (α : ℝ) : p = q <-> (0 = RenyiDivergence_def p q α) := by - -- See paper - sorry +/-- +Definition of the Renyi divergence as an EReal. -lemma RenyiDivergence_def_log_sum_nonneg (p q : PMF T) (α : ℝ) : (0 ≤ (elog (∑' x : T, (p x)^α * (q x)^(1 - α)))) := by - -- Follows from RenyiDivergence_def_nonneg - sorry -/-- -The Renyi divergence is monotonic in the value of its sum. -/ -lemma RenyiDivergence_mono_sum (x y : ℝ) (α : ℝ) (h : 1 < α) : (Real.exp ((α - 1) * x)) ≤ (Real.exp ((α - 1) * y)) -> (x ≤ y) := by - intro H - apply le_of_mul_le_mul_left - · exact exp_le_exp.mp H - · linarith +noncomputable def RenyiDivergence_def (p q : PMF T) (α : ℝ) : EReal := + (α - 1)⁻¹ * (elog (∑' x : T, (p x)^α * (q x)^(1 - α))) /-- Equation for the Renyi divergence series in terms of the Renyi Divergence @@ -70,10 +53,98 @@ lemma RenyiDivergence_def_exp (p q : PMF T) {α : ℝ} rw [Hinvert] clear Hinvert simp - rw [toEReal_ofENNReal_nonneg ?Helog_nn] - case Helog_nn => - apply RenyiDivergence_def_log_sum_nonneg - apply elog_eexp + +-- FIXME: where is this in mathlib? +def AbsolutelyContinuous (p q : T -> ENNReal) : Prop := ∀ x : T, q x = 0 -> p x = 0 + +/-- +Closed form of the series in the definition of the Renyi divergence. +-/ +theorem RenyiDivergenceExpectation (p q : T → ENNReal) {α : ℝ} (h : 1 < α) (H : AbsolutelyContinuous p q) : + (∑' x : T, (p x)^α * (q x)^(1 - α)) = ∑' x: T, (p x / q x)^α * (q x) := by + congr 4 + ext x + rw [AbsolutelyContinuous] at H + -- have HK : (q x ≠ 0 ∨ p x = 0) := by + -- apply Decidable.not_or_of_imp + -- apply H + generalize Hvq : q x = vq + cases vq <;> simp + · -- q x = ⊤ + rw [ENNReal.zero_rpow_of_pos ?H] + case H => linarith + simp + right + apply h + · rename_i vq' + cases (Classical.em (vq' = 0)) + · -- q x = 0 + -- by abs. cty. p x = 0 + rename_i Hvq' + have Hp : p x = 0 := by + apply H + simp [Hvq, Hvq'] + simp [Hp, Hvq', Hvq] + left + linarith + · -- q x ∈ ℝ+ + rename_i Hvq' + generalize Hvp : p x = vp + cases vp <;> simp + · -- q x ∈ ℝ+, p x = ⊤ + rw [ENNReal.top_rpow_of_pos ?H] + case H => linarith + rw [top_mul'] + split + · exfalso + rename_i Hcont + apply Hvq' + have Hcont' : (vq' : ENNReal) = 0 ∧ 0 < (1-α) ∨ (vq' : ENNReal) = ⊤ ∧ (1-α)< 0 := by + exact rpow_eq_zero_iff.mp Hcont + cases Hcont' + · simp_all only [some_eq_coe, none_eq_top, zero_ne_top] + · simp_all only [some_eq_coe, none_eq_top, top_rpow_of_neg, coe_ne_top, sub_neg, and_true] + · rw [top_mul'] + split + · simp_all only [some_eq_coe, none_eq_top, zero_ne_top] + · rfl + · rename_i vp + cases (Classical.em (vq' = 0)) + · -- q x ∈ ℝ+, p x = 0 + rename_i Hvp' + simp_all + · -- q x ∈ ℝ+, p x ∈ ℝ+ + rename_i Hvp' + 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 ) + · simp_all only [some_eq_coe, not_false_eq_true, ne_eq, coe_eq_zero] + · simp_all only [some_eq_coe, not_false_eq_true, ne_eq, coe_ne_top] + +-- FIXME +/-- +The Renyi divergence is monotonic in the value of its sum. +-/ +lemma RenyiDivergence_mono_sum (x y : ℝ) (α : ℝ) (h : 1 < α) : (Real.exp ((α - 1) * x)) ≤ (Real.exp ((α - 1) * y)) -> (x ≤ y) := by + intro H + apply le_of_mul_le_mul_left + · exact exp_le_exp.mp H + · linarith + +theorem RenyiDivergence_def_nonneg (p q : PMF T) (α : ℝ) : (0 ≤ RenyiDivergence_def p q α) := by + -- See paper + sorry + +theorem RenyiDivergence_def_zero (p q : PMF T) (α : ℝ) : p = q <-> (0 = RenyiDivergence_def p q α) := by + -- See paper + sorry + +lemma RenyiDivergence_def_log_sum_nonneg (p q : PMF T) (α : ℝ) : (0 ≤ (elog (∑' x : T, (p x)^α * (q x)^(1 - α)))) := by + -- Follows from RenyiDivergence_def_nonneg + sorry + /-- @@ -86,21 +157,6 @@ noncomputable def RenyiDivergence (p q : PMF T) (α : ℝ) : ENNReal := theorem RenyiDivergence_aux_zero (p q : PMF T) (α : ℝ) : p = q <-> RenyiDivergence p q α = 0 := sorry --- MARKUSDE: fixme -/-- -Closed form of the series in the definition of the Renyi divergence. --/ -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 -- Unused From a0c3a4079aa59a3b3ded103c48252486addba22c Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Fri, 14 Jun 2024 15:11:25 -0400 Subject: [PATCH 045/122] Move Jensen's inequality to Renyi --- .../DifferentialPrivacy/RenyiDivergence.lean | 144 ++++++++++++- .../ZeroConcentrated/Postprocessing.lean | 196 +++++++++--------- 2 files changed, 232 insertions(+), 108 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index df651cc0..997c1794 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -8,6 +8,8 @@ import Mathlib.NumberTheory.ModularForms.JacobiTheta.TwoVariable import Mathlib.Analysis.SpecialFunctions.Pow.NNReal import SampCert.Util.Log import Mathlib.Probability.ProbabilityMassFunction.Basic +import Mathlib.Probability.ProbabilityMassFunction.Integrals +import Mathlib.Analysis.Convex.Integral /-! Renyi Divergence @@ -15,14 +17,13 @@ This file defines the Renyi divergence and equations for evaluating its expectat -/ -open Real ENNReal PMF +open Real ENNReal PMF Nat Int MeasureTheory Measure PMF +open Classical variable {T : Type} /-- Definition of the Renyi divergence as an EReal. - - -/ noncomputable def RenyiDivergence_def (p q : PMF T) (α : ℝ) : EReal := (α - 1)⁻¹ * (elog (∑' x : T, (p x)^α * (q x)^(1 - α))) @@ -55,19 +56,19 @@ lemma RenyiDivergence_def_exp (p q : PMF T) {α : ℝ} simp -- FIXME: where is this in mathlib? -def AbsolutelyContinuous (p q : T -> ENNReal) : Prop := ∀ x : T, q x = 0 -> p x = 0 +-- Need to say that p and q are the PMF's of a measure space (how do I use their typeclasses to do that?) +-- PMF.toMeasure +-- Then get AbsCts from AbsolutelyContinuous +def AbsCts (p q : T -> ENNReal) : Prop := ∀ x : T, q x = 0 -> p x = 0 /-- Closed form of the series in the definition of the Renyi divergence. -/ -theorem RenyiDivergenceExpectation (p q : T → ENNReal) {α : ℝ} (h : 1 < α) (H : AbsolutelyContinuous p q) : +theorem RenyiDivergenceExpectation (p q : T → ENNReal) {α : ℝ} (h : 1 < α) (H : AbsCts p q) : (∑' x : T, (p x)^α * (q x)^(1 - α)) = ∑' x: T, (p x / q x)^α * (q x) := by congr 4 ext x - rw [AbsolutelyContinuous] at H - -- have HK : (q x ≠ 0 ∨ p x = 0) := by - -- apply Decidable.not_or_of_imp - -- apply H + rw [AbsCts] at H generalize Hvq : q x = vq cases vq <;> simp · -- q x = ⊤ @@ -123,13 +124,136 @@ theorem RenyiDivergenceExpectation (p q : T → ENNReal) {α : ℝ} (h : 1 < α) · simp_all only [some_eq_coe, not_false_eq_true, ne_eq, coe_eq_zero] · simp_all only [some_eq_coe, not_false_eq_true, ne_eq, coe_ne_top] + + +/-! +## Jensen's inequality +-/ +section Jensen + +variable {T : Type} +variable [t1 : MeasurableSpace T] +variable [t2 : MeasurableSingletonClass T] + +variable {U V : Type} +variable [m2 : MeasurableSpace U] +variable [count : Countable U] +variable [disc : DiscreteMeasurableSpace U] +variable [Inhabited U] + +lemma 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 ℝ t1 μ _ 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] + +-- FIXME: get rid of toReal's +/-- +Jensen's ineuquality for the exponential applied to Renyi's sum +-/ +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 => + 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 ℝ 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 + . 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 t1 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 t1 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 + + + +end Jensen + + + + -- FIXME /-- The Renyi divergence is monotonic in the value of its sum. -/ lemma RenyiDivergence_mono_sum (x y : ℝ) (α : ℝ) (h : 1 < α) : (Real.exp ((α - 1) * x)) ≤ (Real.exp ((α - 1) * y)) -> (x ≤ y) := by intro H - apply le_of_mul_le_mul_left + apply _root_.le_of_mul_le_mul_left · exact exp_le_exp.mp H · linarith diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean index c4e58001..8818374b 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean @@ -28,104 +28,104 @@ variable [count : Countable U] variable [disc : DiscreteMeasurableSpace U] variable [Inhabited U] -lemma 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 ℝ t1 μ _ 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] - -/-- -Jensen's ineuquality for the exponential applied to Renyi's sum --/ -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 => - 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 ℝ 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 - . 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 t1 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 t1 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 +-- lemma 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 ℝ t1 μ _ 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] +-- +-- /-- +-- Jensen's ineuquality for the exponential applied to Renyi's sum +-- -/ +-- 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 => +-- 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 ℝ 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 +-- . 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 t1 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 t1 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 def δ (nq : SLang U) (f : U → V) (a : V) : {n : U | a = f n} → ENNReal := fun x : {n : U | a = f n} => nq x * (∑' (x : {n | a = f n}), nq x)⁻¹ From 538baa331a38980191d6c82062f3e1eadde3df89 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Fri, 14 Jun 2024 16:25:43 -0400 Subject: [PATCH 046/122] Sketch Jensen's inequality for ENNReal --- .../DifferentialPrivacy/RenyiDivergence.lean | 77 ++++++++++++++++++- 1 file changed, 73 insertions(+), 4 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index 997c1794..773acf79 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -165,11 +165,10 @@ lemma Integrable_rpow (f : T → ℝ) (nn : ∀ x : T, 0 ≤ f x) (μ : Measure . rw [← hasFiniteIntegral_norm_iff] simp [X] --- FIXME: get rid of toReal's /-- Jensen's ineuquality for the exponential applied to Renyi's sum -/ -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)) : +theorem Renyi_Jensen_real (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 @@ -240,10 +239,80 @@ theorem Renyi_Jensen (f : T → ℝ) (q : PMF T) (α : ℝ) (h : 1 < α) (h2 : rw [one_le_ofReal] apply le_of_lt h - - end Jensen +-- MARKUSDE: move +noncomputable def Renyi_Jensen_f (p : T -> ENNReal) (q : PMF T) : T -> ℝ := (fun z => ((p z / q z)).toReal) + +-- Except for one case, we can rewrite the ENNReal-valued inequality into the form Jenen's inequality expects. +lemma Renyi_Jensen_rw (p : T → ENNReal) (q : PMF T) {α : ℝ} (h : 1 < α) (H : AbsCts p q) (Hspecial : ∀ x : T, ¬(p x = ⊤ ∧ q x ≠ 0 ∧ q x ≠ ⊤)) (x : T) : + (p x / q x)^α * (q x) = ENNReal.ofReal (((Renyi_Jensen_f p q) x)^α * (q x).toReal) := sorry + +lemma Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] (p : T → ENNReal) (q : PMF T) {α : ℝ} (h : 1 < α) (H : AbsCts p q) : + (∑' x : T, (p x / q x) * q x) ^ α ≤ (∑' x : T, (p x / q x) ^ α * q x) := by + cases (Classical.em (∀ x : T, ¬(p x = ⊤ ∧ q x ≠ 0 ∧ q x ≠ ⊤))) + · -- Typical case + rename_i Hspecial + conv => + rhs + arg 1 + intro x + rw [Renyi_Jensen_rw p q h H Hspecial] + + rw [<- ENNReal.ofReal_tsum_of_nonneg ?Hnonneg ?Hsummable] + case Hnonneg => + -- Comes from PMF + sorry + case Hsummable => + -- Summability + sorry + + apply (le_trans ?G1 ?G2) + case G2 => + apply (ofReal_le_ofReal ?Hle) + case Hle => + apply Renyi_Jensen_real + · apply h + · -- Comes from PMF + sorry + · -- Also summability + sorry + case G1 => + -- We need the latter fn to be summable + sorry + · -- Special case: There exists some element x0 with p x0 = ⊤ but q x0 ∈ ℝ+ + rename_i Hspecial + simp at * + rcases Hspecial with ⟨ x0, ⟨ H1, H2 , H3 ⟩⟩ + have HT1 : (∑' (x : T), p x / q x * q x) ^ α = ⊤ := by + apply rpow_eq_top_iff.mpr + right + apply And.intro + · apply ENNReal.tsum_eq_top_of_eq_top + exists x0 + apply mul_eq_top.mpr + right + apply And.intro + · apply div_eq_top.mpr + simp_all + · simp_all + · linarith + have HT2 : ∑' (x : T), (p x / q x) ^ α * q x = ⊤ := by + apply ENNReal.tsum_eq_top_of_eq_top + exists x0 + apply mul_eq_top.mpr + right + apply And.intro + · apply rpow_eq_top_iff.mpr + right + apply And.intro + · simp_all + exact top_div_of_ne_top H3 + · simp_all + linarith + · simp_all + rw [HT1, HT2] + From e68ed824accb0e6fee3771a5e01f8d977a7cd889 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Fri, 14 Jun 2024 16:56:22 -0400 Subject: [PATCH 047/122] checkpoint log lemmas --- SampCert/Util/Log.lean | 41 +++++++++++++++++++++++++++++------------ 1 file changed, 29 insertions(+), 12 deletions(-) diff --git a/SampCert/Util/Log.lean b/SampCert/Util/Log.lean index 97ffdb78..7dc88ffa 100644 --- a/SampCert/Util/Log.lean +++ b/SampCert/Util/Log.lean @@ -27,8 +27,8 @@ Coerce a EReal to an ENNReal by truncation -/ noncomputable def ofEReal (e : EReal) : ENNReal := match e with - | none => some 0 - | some none => none + | ⊥ => some 0 + | ⊤ => ⊤ | some (some r) => ENNReal.ofReal r -- instance : Coe EReal ENNReal := ⟨ofEReal⟩ @@ -38,7 +38,7 @@ The extended logarithm -/ def elog (x : ENNReal) : EReal := match x with - | none => ⊤ + | ⊤ => ⊤ | some r => if r = 0 then ⊥ else Real.log r /-- @@ -49,9 +49,9 @@ want the exponent to be of type ``EReal``. -/ def eexp (y : EReal) : ENNReal := match y with - | none => 0 - | (some none) => ⊤ - | (some (some r)) => ENNReal.ofReal (Real.exp r) + | ⊥ => 0 + | ⊤ => ⊤ + | some (some r) => ENNReal.ofReal (Real.exp r) variable {x y : ENNReal} @@ -70,19 +70,21 @@ lemma elog_of_pos_real (H : 0 < r) : elog (ENNReal.ofReal r) = Real.log r := by · sorry @[simp] -lemma elog_zero : elog (ENNReal.ofReal 0) = ⊥ := by sorry +lemma elog_zero : elog (ENNReal.ofReal 0) = ⊥ := by simp [elog] @[simp] -lemma elog_top : elog ⊤ = ⊤ := by sorry +lemma elog_top : elog ⊤ = ⊤ := by simp [elog] @[simp] -lemma eexp_bot : eexp ⊥ = 0 := by sorry +lemma eexp_bot : eexp ⊥ = 0 := by simp [eexp] @[simp] -lemma eexp_top : eexp ⊤ = ⊤ := by sorry +lemma eexp_top : eexp ⊤ = ⊤ := by simp [eexp] @[simp] -lemma eexp_pos_real (H : 0 < r) : eexp r = ENNReal.ofReal (Real.log r) := by sorry +lemma eexp_pos_real (H : 0 < r) : eexp r = ENNReal.ofReal (Real.exp r) := by + simp [ENNReal.ofReal, eexp, elog] + rfl @[simp] lemma elog_eexp : eexp (elog x) = x := by @@ -113,7 +115,22 @@ lemma elog_eexp : eexp (elog x) = x := by @[simp] lemma eexp_elog : (elog (eexp w)) = w := by - sorry + cases w + · simp [eexp, elog] + rfl + · simp [eexp, elog] + rename_i v' + cases v' + · simp + rfl + · simp + rename_i v'' + simp [ENNReal.ofReal] + split + · -- exp is nonnegative + sorry + · sorry + @[simp] lemma elog_mul : elog x * elog y = elog (x + y) := by sorry -- checked truth table From 634fa7c2154ee2e124f9e8d36f3ead212e240506 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Mon, 17 Jun 2024 10:31:13 -0400 Subject: [PATCH 048/122] checkpoint --- .../DifferentialPrivacy/RenyiDivergence.lean | 16 +++++++++------- .../ZeroConcentrated/ConcentratedBound.lean | 11 +++++++---- .../DifferentialPrivacy/ZeroConcentrated/DP.lean | 12 ++++++------ SampCert/Samplers/GaussianGen/Properties.lean | 6 ++++++ SampCert/Util/Gaussian/DiscreteGaussian.lean | 15 +++++++++++++++ SampCert/Util/Log.lean | 9 +++++---- 6 files changed, 48 insertions(+), 21 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index 773acf79..f93010af 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -242,13 +242,13 @@ theorem Renyi_Jensen_real (f : T → ℝ) (q : PMF T) (α : ℝ) (h : 1 < α) (h end Jensen -- MARKUSDE: move -noncomputable def Renyi_Jensen_f (p : T -> ENNReal) (q : PMF T) : T -> ℝ := (fun z => ((p z / q z)).toReal) +noncomputable def Renyi_Jensen_f (p q : PMF T) : T -> ℝ := (fun z => ((p z / q z)).toReal) -- Except for one case, we can rewrite the ENNReal-valued inequality into the form Jenen's inequality expects. -lemma Renyi_Jensen_rw (p : T → ENNReal) (q : PMF T) {α : ℝ} (h : 1 < α) (H : AbsCts p q) (Hspecial : ∀ x : T, ¬(p x = ⊤ ∧ q x ≠ 0 ∧ q x ≠ ⊤)) (x : T) : +lemma Renyi_Jensen_rw (p q : PMF T) {α : ℝ} (h : 1 < α) (H : AbsCts p q) (Hspecial : ∀ x : T, ¬(p x = ⊤ ∧ q x ≠ 0 ∧ q x ≠ ⊤)) (x : T) : (p x / q x)^α * (q x) = ENNReal.ofReal (((Renyi_Jensen_f p q) x)^α * (q x).toReal) := sorry -lemma Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] (p : T → ENNReal) (q : PMF T) {α : ℝ} (h : 1 < α) (H : AbsCts p q) : +theorem Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] (p q : PMF T) {α : ℝ} (h : 1 < α) (H : AbsCts p q) : (∑' x : T, (p x / q x) * q x) ^ α ≤ (∑' x : T, (p x / q x) ^ α * q x) := by cases (Classical.em (∀ x : T, ¬(p x = ⊤ ∧ q x ≠ 0 ∧ q x ≠ ⊤))) · -- Typical case @@ -262,7 +262,11 @@ lemma Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] (p : rw [<- ENNReal.ofReal_tsum_of_nonneg ?Hnonneg ?Hsummable] case Hnonneg => -- Comes from PMF - sorry + intro t + apply mul_nonneg + · refine rpow_nonneg ?ha.hx α + simp [Renyi_Jensen_f] + · exact toReal_nonneg case Hsummable => -- Summability sorry @@ -274,7 +278,7 @@ lemma Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] (p : apply Renyi_Jensen_real · apply h · -- Comes from PMF - sorry + simp [Renyi_Jensen_f] · -- Also summability sorry case G1 => @@ -314,8 +318,6 @@ lemma Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] (p : rw [HT1, HT2] - - -- FIXME /-- The Renyi divergence is monotonic in the value of its sum. diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean index 41bd6570..33962dba 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean @@ -530,8 +530,8 @@ theorem Renyi_Gauss_summable {σ : ℝ} (h : σ ≠ 0) (μ ν : ℤ) (α : ℝ) Upper bound on Renyi divergence between discrete Gaussians. -/ theorem Renyi_Gauss_divergence_bound' {σ α : ℝ} (h : σ ≠ 0) (h' : 1 < α) (μ ν : ℤ) : - RenyiDivergence (fun (x : ℤ) => ENNReal.ofReal (discrete_gaussian σ μ x)) - (fun (x : ℤ) => ENNReal.ofReal (discrete_gaussian σ ν x)) + RenyiDivergence (discrete_gaussian_pmf h μ) + (discrete_gaussian_pmf h ν) α ≤ (ENNReal.ofReal α) * (ENNReal.ofReal ((((μ - ν) : ℤ)^2 : ℝ) / (2 * σ^2))) := by /- have A : RenyiDivergence (fun (x : ℤ) => ENNReal.ofReal (discrete_gaussian σ μ x)) @@ -573,8 +573,9 @@ namespace SLang Upper bound on Renyi divergence between outputs of the ``SLang`` discrete Gaussian sampler. -/ theorem discrete_GaussianGenSample_ZeroConcentrated {α : ℝ} (h : 1 < α) (num : PNat) (den : PNat) (μ ν : ℤ) : - RenyiDivergence ((DiscreteGaussianGenSample num den μ)) (DiscreteGaussianGenSample num den ν) α ≤ - (ENNReal.ofReal α) * (ENNReal.ofReal ((((μ - ν) : ℤ)^2 : ENNReal) / (2 * ((num : ℝ) / (den : ℝ))^2))) := by + RenyiDivergence ((DiscreteGaussianGenPMF num den μ)) (DiscreteGaussianGenPMF num den ν) α ≤ + (ENNReal.ofReal α) * (ENNReal.ofReal (((μ - ν) : ℤ)^2 : ℝ) / (((2 : ENNReal) * ((num : ENNReal) / (den : ENNReal))^2 : ENNReal))) := 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 => @@ -586,5 +587,7 @@ theorem discrete_GaussianGenSample_ZeroConcentrated {α : ℝ} (h : 1 < α) (num rw [DiscreteGaussianGenSample_apply] . skip apply Renyi_Gauss_divergence_bound' A h + -/ + sorry end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean index 07c18561..9de8d294 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean @@ -31,7 +31,7 @@ Inequality defining ``(ε^2)/2``-zCDP. All ``ε``-DP mechanisms satisfy this bound (though not all mechanisms satisfying this bound are ``ε``-DP). -/ -def zCDPBound (q : List T → SLang U) (ε : ℝ) : Prop := +def zCDPBound (q : List T → PMF U) (ε : ℝ) : Prop := ∀ α : ℝ, 1 < α → ∀ l₁ l₂ : List T, Neighbour l₁ l₂ → RenyiDivergence (q l₁) (q l₂) α ≤ ENNReal.ofReal ((1/2) * ε ^ 2 * α) @@ -48,9 +48,9 @@ def NonTopRDNQ (nq : List T → SLang U) : Prop := /-- The mechanism ``q`` is ``(ε^2)/2``-zCDP -/ -def zCDP (q : List T → SLang U) (ε : ℝ) : Prop := +def zCDP (q : List T → PMF U) (ε : ℝ) : Prop := zCDPBound q ε - ∧ NonZeroNQ q - ∧ NonTopSum q - ∧ NonTopNQ q - ∧ NonTopRDNQ q + -- ∧ NonZeroNQ q + -- ∧ NonTopSum q + -- ∧ NonTopNQ q + -- ∧ NonTopRDNQ q diff --git a/SampCert/Samplers/GaussianGen/Properties.lean b/SampCert/Samplers/GaussianGen/Properties.lean index 230c88d9..50e7e180 100644 --- a/SampCert/Samplers/GaussianGen/Properties.lean +++ b/SampCert/Samplers/GaussianGen/Properties.lean @@ -53,4 +53,10 @@ theorem DiscreteGaussianGenSample_apply (num : PNat) (den : PNat) (μ x : ℤ) : . simp [gauss_term_ℝ] . rw [shifted_gauss_sum_0 A] + +/-- +``SLang`` program is a proper distribution. +-/ +def DiscreteGaussianGenPMF (num : PNat) (den : PNat) (μ : ℤ) : PMF ℤ := ⟨ DiscreteGaussianGenSample num den μ , sorry ⟩ + end SLang diff --git a/SampCert/Util/Gaussian/DiscreteGaussian.lean b/SampCert/Util/Gaussian/DiscreteGaussian.lean index f522f308..9b4a82ac 100644 --- a/SampCert/Util/Gaussian/DiscreteGaussian.lean +++ b/SampCert/Util/Gaussian/DiscreteGaussian.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jean-Baptiste Tristan -/ import Mathlib.NumberTheory.ModularForms.JacobiTheta.TwoVariable +import Mathlib.Probability.ProbabilityMassFunction.Basic /-! # Discrete Guassian Definitions @@ -334,3 +335,17 @@ theorem discrete_gaussian_normalizes {σ : ℝ} (h : σ ≠ 0) (μ : ℝ) : rw [tsum_div_const] rw [div_eq_one_iff_eq] apply sum_gauss_term_ne_zero h + + + +/-- +The discrete Gaussian as a PMF +-/ +def discrete_gaussian_pmf {σ : ℝ} (h : σ ≠ 0) (μ : ℝ) : PMF ℤ := + ⟨ (fun n => ENNReal.ofReal (discrete_gaussian σ μ n)), + by + let Hsummable := Summable.hasSum (discrete_gaussian_summable' h μ) + rw [discrete_gaussian_normalizes h μ] at Hsummable + -- Might be a lemma to do this already + sorry + ⟩ diff --git a/SampCert/Util/Log.lean b/SampCert/Util/Log.lean index 7dc88ffa..66389c35 100644 --- a/SampCert/Util/Log.lean +++ b/SampCert/Util/Log.lean @@ -82,7 +82,7 @@ lemma eexp_bot : eexp ⊥ = 0 := by simp [eexp] lemma eexp_top : eexp ⊤ = ⊤ := by simp [eexp] @[simp] -lemma eexp_pos_real (H : 0 < r) : eexp r = ENNReal.ofReal (Real.exp r) := by +lemma eexp_ofReal : eexp r = ENNReal.ofReal (Real.exp r) := by simp [ENNReal.ofReal, eexp, elog] rfl @@ -98,8 +98,7 @@ lemma elog_eexp : eexp (elog x) = x := by rw [h] · rename_i _ H simp - rw [eexp] - rw [NNReal.toReal, Real.toEReal] + rw [NNReal.toReal] simp rw [Real.exp_log] rw [ofReal_coe_nnreal] @@ -133,7 +132,9 @@ lemma eexp_elog : (elog (eexp w)) = w := by @[simp] -lemma elog_mul : elog x * elog y = elog (x + y) := by sorry -- checked truth table +lemma elog_mul : elog x * elog y = elog (x + y) := by + + sorry -- checked truth table @[simp] lemma eexp_add : eexp w * eexp z = eexp (w + z) := by sorry -- checked truth table From 73f66b8c09d0094d54a5bc93ee4682545d090a1b Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Mon, 17 Jun 2024 11:42:26 -0400 Subject: [PATCH 049/122] checkpoint --- .../DifferentialPrivacy/RenyiDivergence.lean | 36 ++++++++++++++++--- 1 file changed, 31 insertions(+), 5 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index f93010af..c0af372a 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -55,12 +55,32 @@ lemma RenyiDivergence_def_exp (p q : PMF T) {α : ℝ} clear Hinvert simp + -- FIXME: where is this in mathlib? -- Need to say that p and q are the PMF's of a measure space (how do I use their typeclasses to do that?) -- PMF.toMeasure -- Then get AbsCts from AbsolutelyContinuous def AbsCts (p q : T -> ENNReal) : Prop := ∀ x : T, q x = 0 -> p x = 0 +/-- +Specialize the definitation of AbsolutelyContinuous when singletons are measurable +-/ +lemma PMF_AbsCts [MeasurableSpace T] [MeasurableSingletonClass T] (p q : PMF T) (H : AbsolutelyContinuous (PMF.toMeasure p) (PMF.toMeasure q)) : AbsCts p q := by + rw [AbsolutelyContinuous] at H + rw [AbsCts] + intro x Hx + have Hxm : q.toMeasure { x } = 0 := by + rw [toMeasure] + simp + apply (toOuterMeasure_apply_eq_zero_iff q {x}).mpr + exact Set.disjoint_singleton_right.mpr fun a => a Hx + have H := H Hxm + rw [toMeasure] at H + simp at * + have Hp : Disjoint p.support {x} := (toOuterMeasure_apply_eq_zero_iff p {x}).mp H + simp at Hp + assumption + /-- Closed form of the series in the definition of the Renyi divergence. -/ @@ -125,7 +145,6 @@ theorem RenyiDivergenceExpectation (p q : T → ENNReal) {α : ℝ} (h : 1 < α) · simp_all only [some_eq_coe, not_false_eq_true, ne_eq, coe_ne_top] - /-! ## Jensen's inequality -/ @@ -261,7 +280,6 @@ theorem Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] (p rw [<- ENNReal.ofReal_tsum_of_nonneg ?Hnonneg ?Hsummable] case Hnonneg => - -- Comes from PMF intro t apply mul_nonneg · refine rpow_nonneg ?ha.hx α @@ -269,6 +287,13 @@ theorem Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] (p · exact toReal_nonneg case Hsummable => -- Summability + -- Derivable from PMF for p? + conv => + congr + intro x + rw [Renyi_Jensen_f] + -- If this is not summable, then the RHS sum is infinity, and it's trivial by a classical cases? + -- Is that true? sorry apply (le_trans ?G1 ?G2) @@ -277,9 +302,10 @@ theorem Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] (p case Hle => apply Renyi_Jensen_real · apply h - · -- Comes from PMF - simp [Renyi_Jensen_f] - · -- Also summability + · simp [Renyi_Jensen_f] + · rw [Memℒp] + -- The seminorm condition related to summability + sorry case G1 => -- We need the latter fn to be summable From 7112e0deafcd50263793a2f8466c0e5c73ec317f Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Mon, 17 Jun 2024 14:30:58 -0400 Subject: [PATCH 050/122] main MemLP proof for Renyi --- .../DifferentialPrivacy/RenyiDivergence.lean | 192 +++++++++++------- 1 file changed, 124 insertions(+), 68 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index c0af372a..1087f35f 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -267,82 +267,138 @@ noncomputable def Renyi_Jensen_f (p q : PMF T) : T -> ℝ := (fun z => ((p z / q lemma Renyi_Jensen_rw (p q : PMF T) {α : ℝ} (h : 1 < α) (H : AbsCts p q) (Hspecial : ∀ x : T, ¬(p x = ⊤ ∧ q x ≠ 0 ∧ q x ≠ ⊤)) (x : T) : (p x / q x)^α * (q x) = ENNReal.ofReal (((Renyi_Jensen_f p q) x)^α * (q x).toReal) := sorry -theorem Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] (p q : PMF T) {α : ℝ} (h : 1 < α) (H : AbsCts p q) : +-- set_option pp.all true + +#check Or.elim (Classical.em ?P) ?G1 ?G2 + +theorem Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] [Countable T] (p q : PMF T) {α : ℝ} (h : 1 < α) (H : AbsCts p q) : (∑' x : T, (p x / q x) * q x) ^ α ≤ (∑' x : T, (p x / q x) ^ α * q x) := by - cases (Classical.em (∀ x : T, ¬(p x = ⊤ ∧ q x ≠ 0 ∧ q x ≠ ⊤))) - · -- Typical case - rename_i Hspecial - conv => - rhs - arg 1 - intro x - rw [Renyi_Jensen_rw p q h H Hspecial] - - rw [<- ENNReal.ofReal_tsum_of_nonneg ?Hnonneg ?Hsummable] - case Hnonneg => - intro t - apply mul_nonneg - · refine rpow_nonneg ?ha.hx α - simp [Renyi_Jensen_f] - · exact toReal_nonneg - case Hsummable => - -- Summability - -- Derivable from PMF for p? + have Hdiscr : DiscreteMeasurableSpace T := MeasurableSingletonClass.toDiscreteMeasurableSpace + cases (Classical.em (∑' (a : T), (p a / q a) ^ α * q a ≠ ⊤)) + · rename_i Hnts + cases (Classical.em (∀ x : T, ¬(p x = ⊤ ∧ q x ≠ 0 ∧ q x ≠ ⊤))) + · -- Typical case + rename_i Hspecial conv => - congr + rhs + arg 1 intro x - rw [Renyi_Jensen_f] - -- If this is not summable, then the RHS sum is infinity, and it's trivial by a classical cases? - -- Is that true? - sorry - - apply (le_trans ?G1 ?G2) - case G2 => - apply (ofReal_le_ofReal ?Hle) - case Hle => - apply Renyi_Jensen_real - · apply h - · simp [Renyi_Jensen_f] - · rw [Memℒp] - -- The seminorm condition related to summability - - sorry - case G1 => - -- We need the latter fn to be summable - sorry - · -- Special case: There exists some element x0 with p x0 = ⊤ but q x0 ∈ ℝ+ - rename_i Hspecial - simp at * - rcases Hspecial with ⟨ x0, ⟨ H1, H2 , H3 ⟩⟩ - have HT1 : (∑' (x : T), p x / q x * q x) ^ α = ⊤ := by - apply rpow_eq_top_iff.mpr - right - apply And.intro - · apply ENNReal.tsum_eq_top_of_eq_top - exists x0 - apply mul_eq_top.mpr + rw [Renyi_Jensen_rw p q h H Hspecial] + + rw [<- ENNReal.ofReal_tsum_of_nonneg ?Hnonneg ?Hsummable] + case Hnonneg => + intro t + apply mul_nonneg + · refine rpow_nonneg ?ha.hx α + simp [Renyi_Jensen_f] + · exact toReal_nonneg + case Hsummable => + -- Summability + -- Derivable from PMF for p? + conv => + congr + intro x + rw [Renyi_Jensen_f] + -- If this is not summable, then the RHS sum is infinity, and it's trivial by a classical cases? + -- Is that true? + sorry + + have HRJf_nonneg (a : T) : 0 <= Renyi_Jensen_f p q a := + sorry + have HRJf_nt (a : T) : p a / q a ≠ ⊤ := by + sorry -- This one might need some more cases? + have Hsum_indicator (a : T) : ∑' (i : T), q i * Set.indicator {a} (fun x => 1) i = q a := + sorry + + apply (le_trans ?G1 ?G2) + case G2 => + apply (ofReal_le_ofReal ?Hle) + case Hle => + apply Renyi_Jensen_real + · apply h + · simp [Renyi_Jensen_f] + · 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 + . apply measurable_discrete + . apply Measurable.inv + apply measurable_discrete + · simp [snorm] + split + · simp + · rename_i Hα + simp [snorm'] + rw [MeasureTheory.lintegral_countable'] + rw [toReal_ofReal (le_of_lt (lt_trans zero_lt_one h))] + apply rpow_lt_top_of_nonneg + · simp + apply le_of_not_ge Hα + · conv => + lhs + arg 1 + intro a + arg 1 + arg 1 + rw [<- Real.toNNReal_eq_nnnorm_of_nonneg (HRJf_nonneg a)] + rw [Renyi_Jensen_f] + rw [<- ENNReal.ofReal.eq_1] + rw [ENNReal.ofReal_toReal (HRJf_nt a)] + rfl + conv => + lhs + arg 1 + intro a + rhs + simp [toMeasure] + simp [PMF.toOuterMeasure] + rw [Hsum_indicator] + apply Hnts + case G1 => + -- We need the latter fn to be summable + sorry + · -- Special case: There exists some element x0 with p x0 = ⊤ but q x0 ∈ ℝ+ + rename_i Hspecial + simp at * + rcases Hspecial with ⟨ x0, ⟨ H1, H2 , H3 ⟩⟩ + have HT1 : (∑' (x : T), p x / q x * q x) ^ α = ⊤ := by + apply rpow_eq_top_iff.mpr right apply And.intro - · apply div_eq_top.mpr - simp_all - · simp_all - · linarith - have HT2 : ∑' (x : T), (p x / q x) ^ α * q x = ⊤ := by - apply ENNReal.tsum_eq_top_of_eq_top - exists x0 - apply mul_eq_top.mpr - right - apply And.intro - · apply rpow_eq_top_iff.mpr + · apply ENNReal.tsum_eq_top_of_eq_top + exists x0 + apply mul_eq_top.mpr + right + apply And.intro + · apply div_eq_top.mpr + simp_all + · simp_all + · linarith + have HT2 : ∑' (x : T), (p x / q x) ^ α * q x = ⊤ := by + apply ENNReal.tsum_eq_top_of_eq_top + exists x0 + apply mul_eq_top.mpr right apply And.intro + · apply rpow_eq_top_iff.mpr + right + apply And.intro + · simp_all + exact top_div_of_ne_top H3 + · simp_all + linarith · simp_all - exact top_div_of_ne_top H3 - · simp_all - linarith - · simp_all - rw [HT1, HT2] - + rw [HT1, HT2] + · rename_i HStop + simp at * + rw [HStop] + exact OrderTop.le_top ((∑' (x : T), p x / q x * q x) ^ α) -- FIXME /-- From c2f78710c0a12f36204d9f24f8c93086a3b77736 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Mon, 17 Jun 2024 15:26:26 -0400 Subject: [PATCH 051/122] close side conditition --- .../DifferentialPrivacy/RenyiDivergence.lean | 76 ++++++++++++++----- 1 file changed, 57 insertions(+), 19 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index 1087f35f..3e649377 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -267,10 +267,6 @@ noncomputable def Renyi_Jensen_f (p q : PMF T) : T -> ℝ := (fun z => ((p z / q lemma Renyi_Jensen_rw (p q : PMF T) {α : ℝ} (h : 1 < α) (H : AbsCts p q) (Hspecial : ∀ x : T, ¬(p x = ⊤ ∧ q x ≠ 0 ∧ q x ≠ ⊤)) (x : T) : (p x / q x)^α * (q x) = ENNReal.ofReal (((Renyi_Jensen_f p q) x)^α * (q x).toReal) := sorry --- set_option pp.all true - -#check Or.elim (Classical.em ?P) ?G1 ?G2 - theorem Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] [Countable T] (p q : PMF T) {α : ℝ} (h : 1 < α) (H : AbsCts p q) : (∑' x : T, (p x / q x) * q x) ^ α ≤ (∑' x : T, (p x / q x) ^ α * q x) := by have Hdiscr : DiscreteMeasurableSpace T := MeasurableSingletonClass.toDiscreteMeasurableSpace @@ -284,7 +280,6 @@ theorem Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] [C arg 1 intro x rw [Renyi_Jensen_rw p q h H Hspecial] - rw [<- ENNReal.ofReal_tsum_of_nonneg ?Hnonneg ?Hsummable] case Hnonneg => intro t @@ -293,23 +288,44 @@ theorem Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] [C simp [Renyi_Jensen_f] · exact toReal_nonneg case Hsummable => - -- Summability - -- Derivable from PMF for p? conv => congr intro x rw [Renyi_Jensen_f] - -- If this is not summable, then the RHS sum is infinity, and it's trivial by a classical cases? - -- Is that true? - sorry - - have HRJf_nonneg (a : T) : 0 <= Renyi_Jensen_f p q a := - sorry + conv => + arg 1 + intro x + lhs + rw [ENNReal.toReal_rpow] + conv => + arg 1 + intro x + rw [<- ENNReal.toReal_mul] + apply ENNReal.summable_toReal + assumption + have HRJf_nonneg (a : T) : 0 <= Renyi_Jensen_f p q a := by apply toReal_nonneg have HRJf_nt (a : T) : p a / q a ≠ ⊤ := by - sorry -- This one might need some more cases? - have Hsum_indicator (a : T) : ∑' (i : T), q i * Set.indicator {a} (fun x => 1) i = q a := - sorry - + intro HK + have HK' : (p a ≠ 0 ∧ q a = 0 ∨ p a = ⊤ ∧ q a ≠ ⊤) := by exact div_eq_top.mp HK + cases HK' + · rename_i HK' + rcases HK' with ⟨ HK1 , HK2 ⟩ + rw [AbsCts] at H + simp_all only [ne_eq, not_and, Decidable.not_not, ENNReal.zero_div, zero_ne_top] + · rename_i HK' + rcases HK' with ⟨ HK1 , HK2 ⟩ + apply (Hspecial a) + simp_all + -- We need to eliminate the possibility that (q a = 0) + -- All summads with q a = 0 will have value 0, so they should be removable. + sorry + have Hsum_indicator (a : T) : ∑' (i : T), q i * Set.indicator {a} (fun x => 1) i = q a := by + have Hfun : (fun (i : T) => q i * Set.indicator {a} (fun x => 1) i) = (fun (i : T) => if i = a then q a else 0) := by + funext i + rw [Set.indicator] + split <;> simp <;> split <;> simp_all + rw [Hfun] + exact tsum_ite_eq a (q a) apply (le_trans ?G1 ?G2) case G2 => apply (ofReal_le_ofReal ?Hle) @@ -361,8 +377,30 @@ theorem Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] [C rw [Hsum_indicator] apply Hnts case G1 => - -- We need the latter fn to be summable - sorry + -- We need the latter fn to be summable or else it becomes zero and the inequality does not hold + rw [<- ENNReal.ofReal_rpow_of_nonneg ?Harg ?Hα] + case Harg => + apply tsum_nonneg + intro i + apply mul_nonneg + · apply HRJf_nonneg + · exact toReal_nonneg + case Hα => linarith + apply (ENNReal.rpow_le_rpow _ ?Hα') + case Hα' => linarith + conv => + rhs + arg 1 + arg 1 + intro a + rw [Renyi_Jensen_f] + rw [<- ENNReal.toReal_mul] + rw [<- ENNReal.tsum_toReal_eq] + · rw [ENNReal.ofReal_toReal] + -- Could do another case at the top if not derivable + sorry + · -- Derivable from Hnts + sorry · -- Special case: There exists some element x0 with p x0 = ⊤ but q x0 ∈ ℝ+ rename_i Hspecial simp at * From 3fbe660ffc5962d8b7e94d62096c32d16e0f2289 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Mon, 17 Jun 2024 15:59:54 -0400 Subject: [PATCH 052/122] remove hypotheses from RenyiDivergence_def_exp --- .../DifferentialPrivacy/RenyiDivergence.lean | 64 ++++++++++++------- 1 file changed, 41 insertions(+), 23 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index 3e649377..000a1a2d 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -31,29 +31,45 @@ noncomputable def RenyiDivergence_def (p q : PMF T) (α : ℝ) : EReal := /-- Equation for the Renyi divergence series in terms of the Renyi Divergence -/ -lemma RenyiDivergence_def_exp (p q : PMF T) {α : ℝ} - (h : 1 < α) - (H1 : 0 < ∑' (x : T), p x ^ α * q x ^ (1 - α)) - (H2 : ∑' (x : T), p x ^ α * q x ^ (1 - α) < ⊤) : +lemma RenyiDivergence_def_exp (p q : PMF T) {α : ℝ} (h : 1 < α) : eexp (((α - 1)) * RenyiDivergence_def p q α) = (∑' x : T, (p x)^α * (q x)^(1 - α)) := by - simp only [RenyiDivergence_def] + rw [RenyiDivergence_def] rw [<- mul_assoc] - have Hinvert : ((↑α - 1) * ↑(α - 1)⁻¹ : EReal) = 1 := by - clear H1 - clear H2 - have Hsub : ((↑α - 1) : EReal) = (α - 1 : ℝ) := by simp - rw [Hsub] - have Hundo_coe (r1 r2 : ℝ) : (r1 : EReal) * (r2 : EReal) = ((r1 * r2 : ℝ) : EReal) := by - rw [EReal.coe_mul] - rw [Hundo_coe] - have Hinv : (α - 1) * (α - 1)⁻¹ = 1 := by - apply mul_inv_cancel - linarith - rw [Hinv] - simp - rw [Hinvert] - clear Hinvert - simp + have H1 : (α.toEReal - OfNat.ofNat 1) = (α - OfNat.ofNat 1).toEReal := by + rw [EReal.coe_sub] + congr + have H2 : ((α.toEReal - OfNat.ofNat 1) * (α - OfNat.ofNat 1)⁻¹.toEReal = 1) := by + rw [H1] + rw [← EReal.coe_mul] + rw [mul_inv_cancel] + · simp + · linarith + simp [H2] + + +-- lemma RenyiDivergence_def_exp (p q : PMF T) {α : ℝ} +-- (h : 1 < α) +-- (H1 : 0 < ∑' (x : T), p x ^ α * q x ^ (1 - α)) +-- (H2 : ∑' (x : T), p x ^ α * q x ^ (1 - α) < ⊤) : +-- eexp (((α - 1)) * RenyiDivergence_def p q α) = (∑' x : T, (p x)^α * (q x)^(1 - α)) := by +-- simp only [RenyiDivergence_def] +-- rw [<- mul_assoc] +-- have Hinvert : ((↑α - 1) * ↑(α - 1)⁻¹ : EReal) = 1 := by +-- clear H1 +-- clear H2 +-- have Hsub : ((↑α - 1) : EReal) = (α - 1 : ℝ) := by simp +-- rw [Hsub] +-- have Hundo_coe (r1 r2 : ℝ) : (r1 : EReal) * (r2 : EReal) = ((r1 * r2 : ℝ) : EReal) := by +-- rw [EReal.coe_mul] +-- rw [Hundo_coe] +-- have Hinv : (α - 1) * (α - 1)⁻¹ = 1 := by +-- apply mul_inv_cancel +-- linarith +-- rw [Hinv] +-- simp +-- rw [Hinvert] +-- clear Hinvert +-- simp -- FIXME: where is this in mathlib? @@ -448,8 +464,10 @@ lemma RenyiDivergence_mono_sum (x y : ℝ) (α : ℝ) (h : 1 < α) : (Real.exp ( · exact exp_le_exp.mp H · linarith -theorem RenyiDivergence_def_nonneg (p q : PMF T) (α : ℝ) : (0 ≤ RenyiDivergence_def p q α) := by - -- See paper +theorem RenyiDivergence_def_nonneg (p q : PMF T) {α : ℝ} (Hα : 1 < α): (0 ≤ RenyiDivergence_def p q α) := by + have H1 : eexp (((α - 1)) * 0) ≤ eexp ((α - 1) * RenyiDivergence_def p q α) := by + rw [RenyiDivergence_def_exp p q Hα] + sorry sorry theorem RenyiDivergence_def_zero (p q : PMF T) (α : ℝ) : p = q <-> (0 = RenyiDivergence_def p q α) := by From 59b5fc95b7c4ecf5a1dfe40e10ccb04af669c22a Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Mon, 17 Jun 2024 16:52:07 -0400 Subject: [PATCH 053/122] checkpoint proof of nonnegativity from paper --- .../DifferentialPrivacy/RenyiDivergence.lean | 58 +++++++++---------- SampCert/Util/Log.lean | 11 +++- 2 files changed, 36 insertions(+), 33 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index 000a1a2d..63894ad4 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -46,36 +46,9 @@ lemma RenyiDivergence_def_exp (p q : PMF T) {α : ℝ} (h : 1 < α) : · linarith simp [H2] - --- lemma RenyiDivergence_def_exp (p q : PMF T) {α : ℝ} --- (h : 1 < α) --- (H1 : 0 < ∑' (x : T), p x ^ α * q x ^ (1 - α)) --- (H2 : ∑' (x : T), p x ^ α * q x ^ (1 - α) < ⊤) : --- eexp (((α - 1)) * RenyiDivergence_def p q α) = (∑' x : T, (p x)^α * (q x)^(1 - α)) := by --- simp only [RenyiDivergence_def] --- rw [<- mul_assoc] --- have Hinvert : ((↑α - 1) * ↑(α - 1)⁻¹ : EReal) = 1 := by --- clear H1 --- clear H2 --- have Hsub : ((↑α - 1) : EReal) = (α - 1 : ℝ) := by simp --- rw [Hsub] --- have Hundo_coe (r1 r2 : ℝ) : (r1 : EReal) * (r2 : EReal) = ((r1 * r2 : ℝ) : EReal) := by --- rw [EReal.coe_mul] --- rw [Hundo_coe] --- have Hinv : (α - 1) * (α - 1)⁻¹ = 1 := by --- apply mul_inv_cancel --- linarith --- rw [Hinv] --- simp --- rw [Hinvert] --- clear Hinvert --- simp - - --- FIXME: where is this in mathlib? --- Need to say that p and q are the PMF's of a measure space (how do I use their typeclasses to do that?) --- PMF.toMeasure --- Then get AbsCts from AbsolutelyContinuous +/-- +Simplified consequence of absolute continuity (remove me?) +-/ def AbsCts (p q : T -> ENNReal) : Prop := ∀ x : T, q x = 0 -> p x = 0 /-- @@ -464,10 +437,31 @@ lemma RenyiDivergence_mono_sum (x y : ℝ) (α : ℝ) (h : 1 < α) : (Real.exp ( · exact exp_le_exp.mp H · linarith -theorem RenyiDivergence_def_nonneg (p q : PMF T) {α : ℝ} (Hα : 1 < α): (0 ≤ RenyiDivergence_def p q α) := by +#check True + +theorem RenyiDivergence_def_nonneg [MeasurableSpace T] [MeasurableSingletonClass T] [Countable T] (p q : PMF T) (Hpq : AbsCts p q) {α : ℝ} (Hα : 1 < α) : + (0 ≤ RenyiDivergence_def p q α) := by have H1 : eexp (((α - 1)) * 0) ≤ eexp ((α - 1) * RenyiDivergence_def p q α) := by rw [RenyiDivergence_def_exp p q Hα] - sorry + rw [RenyiDivergenceExpectation p q Hα Hpq] + simp + apply (le_trans ?G1 (Renyi_Jensen_ENNReal p q Hα Hpq)) + have Hone : (∑' (x : T), p x / q x * q x) = 1 := by + -- This argument is also necessary to finish the Jensen's inequality proof (filter out the q x = 0 elements) + sorry + have Hle : (∑' (x : T), p x / q x * q x) ≤ (∑' (x : T), p x / q x * q x) ^ α := by + apply ENNReal.le_rpow_self_of_one_le + · rw [Hone] + · linarith + apply le_trans ?X Hle + rw [Hone] + linarith + apply eexp_mono_le.mpr at H1 + have HX : (0 < (α.toEReal - 1)) := by sorry + have HX1 : (ofEReal ((↑α - 1) * 0) ≤ ofEReal ((↑α - 1) * RenyiDivergence_def p q α)) := by + exact ofEReal_le_mono.mp fun a => H1 + rw [ofEReal_mul] at HX1 + -- Circular side condition? Why is cancelling EReals so hard? Make a lemma with the iff cases in Log sorry theorem RenyiDivergence_def_zero (p q : PMF T) (α : ℝ) : p = q <-> (0 = RenyiDivergence_def p q α) := by diff --git a/SampCert/Util/Log.lean b/SampCert/Util/Log.lean index 66389c35..59da0c49 100644 --- a/SampCert/Util/Log.lean +++ b/SampCert/Util/Log.lean @@ -81,6 +81,10 @@ lemma eexp_bot : eexp ⊥ = 0 := by simp [eexp] @[simp] lemma eexp_top : eexp ⊤ = ⊤ := by simp [eexp] +@[simp] +lemma eexp_zero : eexp 0 = 1 := by simp [eexp] + + @[simp] lemma eexp_ofReal : eexp r = ENNReal.ofReal (Real.exp r) := by simp [ENNReal.ofReal, eexp, elog] @@ -190,12 +194,17 @@ lemma elog_mono_le : (x <= y) <-> elog x <= elog y := by sorry lemma ofEReal_nonneg_eq_iff (Hw : 0 <= w) (Hz : 0 <= z) : w = z <-> (ofEReal w = ofEReal z) := sorry -lemma ofEReal_le_mono : w ≤ z <-> (ofEReal w ≤ ofEReal z) := +lemma ofEReal_le_mono : (0 ≤ w) -> w ≤ z <-> (ofEReal w ≤ ofEReal z) := sorry + +@[simp] +lemma ofEReal_mul (Hw : 0 ≤ w) (Hz : 0 ≤ z) : ofEReal (w * z) = ofEReal w * ofEReal z := sorry + @[simp] lemma toEReal_ofENNReal_nonneg (H : 0 ≤ w) : ENNReal.toEReal (ofEReal w) = w := sorry + @[simp] lemma ofEReal_toENNReal : ofEReal (ENNReal.toEReal x) = x := by sorry From eb3e7de1fb65d102676bf9d27e9b979cc95d9d40 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Tue, 18 Jun 2024 09:26:13 -0400 Subject: [PATCH 054/122] make postprocessing compile --- .../DifferentialPrivacy/RenyiDivergence.lean | 5 +- .../ZeroConcentrated/DP.lean | 8 +- .../ZeroConcentrated/Postprocessing.lean | 268 +++++++++--------- 3 files changed, 146 insertions(+), 135 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index 63894ad4..bba85490 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -455,14 +455,13 @@ theorem RenyiDivergence_def_nonneg [MeasurableSpace T] [MeasurableSingletonClass · linarith apply le_trans ?X Hle rw [Hone] - linarith apply eexp_mono_le.mpr at H1 have HX : (0 < (α.toEReal - 1)) := by sorry have HX1 : (ofEReal ((↑α - 1) * 0) ≤ ofEReal ((↑α - 1) * RenyiDivergence_def p q α)) := by exact ofEReal_le_mono.mp fun a => H1 rw [ofEReal_mul] at HX1 -- Circular side condition? Why is cancelling EReals so hard? Make a lemma with the iff cases in Log - sorry + all_goals sorry theorem RenyiDivergence_def_zero (p q : PMF T) (α : ℝ) : p = q <-> (0 = RenyiDivergence_def p q α) := by -- See paper @@ -472,8 +471,6 @@ lemma RenyiDivergence_def_log_sum_nonneg (p q : PMF T) (α : ℝ) : (0 ≤ (elog -- Follows from RenyiDivergence_def_nonneg sorry - - /-- The Renyi divergence. -/ diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean index 9de8d294..09718ec4 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean @@ -31,9 +31,9 @@ Inequality defining ``(ε^2)/2``-zCDP. All ``ε``-DP mechanisms satisfy this bound (though not all mechanisms satisfying this bound are ``ε``-DP). -/ -def zCDPBound (q : List T → PMF U) (ε : ℝ) : Prop := +def zCDPBound (q : List T → SLang U) (HNorm : ∀ l, HasSum (q l) 1) (ε : ℝ) : Prop := ∀ α : ℝ, 1 < α → ∀ l₁ l₂ : List T, Neighbour l₁ l₂ → - RenyiDivergence (q l₁) (q l₂) α ≤ ENNReal.ofReal ((1/2) * ε ^ 2 * α) + RenyiDivergence (SLang.toPMF (q l₁) (HNorm l₁)) (SLang.toPMF (q l₂) (HNorm l₂)) α ≤ ENNReal.ofReal ((1/2) * ε ^ 2 * α) def NonTopNQ (nq : List T → SLang U) := ∀ l : List T, ∀ n : U, nq l n ≠ ⊤ @@ -48,8 +48,8 @@ def NonTopRDNQ (nq : List T → SLang U) : Prop := /-- The mechanism ``q`` is ``(ε^2)/2``-zCDP -/ -def zCDP (q : List T → PMF U) (ε : ℝ) : Prop := - zCDPBound q ε +def zCDP (q : List T → SLang U) (HNorm : ∀ l, HasSum (q l) 1) (ε : ℝ) : Prop := + zCDPBound q HNorm ε -- ∧ NonZeroNQ q -- ∧ NonTopSum q -- ∧ NonTopNQ q diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean index 8818374b..4c3369c6 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean @@ -28,104 +28,104 @@ variable [count : Countable U] variable [disc : DiscreteMeasurableSpace U] variable [Inhabited U] --- lemma 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 ℝ t1 μ _ 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] --- --- /-- --- Jensen's ineuquality for the exponential applied to Renyi's sum --- -/ --- 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 => --- 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 ℝ 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 --- . 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 t1 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 t1 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 +lemma 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 ℝ t1 μ _ 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] + +/-- +Jensen's ineuquality for the exponential applied to Renyi's sum +-/ +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 => + 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 ℝ 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 + . 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 t1 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 t1 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 def δ (nq : SLang U) (f : U → V) (a : V) : {n : U | a = f n} → ENNReal := fun x : {n : U | a = f n} => nq x * (∑' (x : {n | a = f n}), nq x)⁻¹ @@ -167,10 +167,10 @@ theorem norm_simplify (x : ENNReal) (h : x ≠ ⊤) : 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 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 convergent_subset {p : T → ENNReal} (f : T → V) (conv : ∑' (x : T), p x ≠ ⊤) : ∑' (x : { y : T| x = f y }), p x ≠ ⊤ := by @@ -217,7 +217,11 @@ theorem tsum_pos_int {f : ℤ → ENNReal} (h1 : ∑' x : ℤ, f x ≠ ⊤) (h2 apply toReal_strict_mono h1 apply ENNReal.tsum_pos_int h1 h2 -theorem DPostPocess_pre {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : zCDPBound nq ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (conv : NonTopSum nq) (f : U → V) {α : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (h2 : Neighbour l₁ l₂) : + + +theorem DPostPocess_pre {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) 1} {ε₁ ε₂ : ℕ+} + (h : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) + (conv : NonTopSum nq) (f : U → V) {α : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (Habs : AbsCts (nq l₁) (nq l₂)) (h2 : Neighbour l₁ l₂) : (∑' (x : V), (∑' (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 - α)) ≤ @@ -226,7 +230,7 @@ theorem DPostPocess_pre {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : zCDP simp [zCDPBound, RenyiDivergence] at h -- Rewrite as cascading expectations - rw [@RenyiDivergenceExpectation _ (nq l₁) (nq l₂) _ h1 (nn l₂) (nts l₂)] + rw [@RenyiDivergenceExpectation _ (nq l₁) (nq l₂) _ h1 Habs] -- Shuffle the sum rw [fiberwisation (fun x => (nq l₁ x / nq l₂ x) ^ α * nq l₂ x) f] @@ -318,9 +322,10 @@ theorem DPostPocess_pre {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : zCDP . replace nt := nt α h1 l₁ l₂ h2 apply convergent_subset _ nt . intro x - apply nn - . intro x - apply nts + sorry + -- 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 @@ -549,8 +554,13 @@ theorem tsum_ne_zero_of_ne_zero {T : Type} [Inhabited T] (f : T → ENNReal) (h have B := CONTRA default contradiction -theorem privPostProcess_zCDPBound {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (h : zCDPBound nq ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (conv : NonTopSum nq) (f : U → V) : - zCDPBound (privPostProcess nq f) ((ε₁ : ℝ) / ε₂) := by +lemma pushforward_pmf_norm (nq : List T → SLang U) (HNorm : ∀ l, HasSum (nq l) 1) (f : U -> V) : (∀ l, HasSum (privPostProcess nq f l) 1) := + sorry + +theorem privPostProcess_zCDPBound {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) 1} {ε₁ ε₂ : ℕ+} + (h : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (conv : NonTopSum nq) + (f : U → V) : + zCDPBound (privPostProcess nq f) (pushforward_pmf_norm nq HNorm f) ((ε₁ : ℝ) / ε₂) := by /- simp [privPostProcess, zCDPBound, RenyiDivergence] intro α h1 l₁ l₂ h2 @@ -640,33 +650,37 @@ theorem privPostProcess_zCDPBound {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} -/ sorry -theorem privPostProcess_NonTopRDNQ {nq : List T → SLang U} {ε₁ ε₂ : ℕ+} (f : U → V) (dp :zCDPBound nq ((ε₁ : ℝ) / ε₂)) (nt : NonTopRDNQ nq) (nz : NonZeroNQ nq) (nts : NonTopNQ nq) (ntsum: NonTopSum nq) : +theorem privPostProcess_NonTopRDNQ {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) 1} {ε₁ ε₂ : ℕ+} (f : U → V) + (dp : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) (nt : NonTopRDNQ nq) (nz : NonZeroNQ nq) (nts : NonTopNQ nq) (ntsum: NonTopSum nq) : NonTopRDNQ (privPostProcess nq f) := by simp [NonTopRDNQ, NonTopSum, privPostProcess] at * intros α h1 l₁ l₂ h2 have ntrdnq := nt replace nt := nt α h1 l₁ l₂ h2 - have A := @DPostPocess_pre T U V _ _ _ nq ε₁ ε₂ dp nz ntrdnq nts ntsum f α h1 l₁ l₂ h2 - apply @LT.lt.ne_top _ _ _ _ ⊤ - rw [Eq.comm] at nt - have B := Ne.lt_top' nt - exact lt_of_le_of_lt A B + sorry + -- have A := @DPostPocess_pre T U V _ _ _ nq ε₁ ε₂ dp nz ntrdnq nts ntsum f α h1 l₁ l₂ h2 + -- apply @LT.lt.ne_top _ _ _ _ ⊤ + -- rw [Eq.comm] at nt + -- have B := Ne.lt_top' nt + -- exact lt_of_le_of_lt A B /-- Postprocessing preserves zCDP -/ -theorem privPostProcess_zCDP {f : U → V} (sur : Function.Surjective f) (nq : List T → SLang U) (ε₁ ε₂ : ℕ+) (h : zCDP nq ((ε₁ : ℝ) / ε₂)) : - zCDP (privPostProcess nq f) (((ε₁ : ℝ) / ε₂)) := by +theorem privPostProcess_zCDP {f : U → V} (sur : Function.Surjective f) + (nq : List T → SLang U) (HNorm : ∀ l, HasSum (nq l) 1) (ε₁ ε₂ : ℕ+) (h : zCDP nq HNorm ((ε₁ : ℝ) / ε₂)) : + zCDP (privPostProcess nq f) (pushforward_pmf_norm nq HNorm f) (((ε₁ : ℝ) / ε₂)) := by simp [zCDP] at * - cases h ; rename_i h1 h2 ; cases h2 ; rename_i h2 h3 ; cases h3 ; rename_i h3 h4 ; cases h4 ; rename_i h4 h5 - repeat any_goals constructor - . apply privPostProcess_zCDPBound h1 h2 h5 h4 h3 - . apply privPostProcess_NonZeroNQ h2 sur - . apply privPostProcess_NonTopSum f h3 - . simp [NonTopNQ] - intro l - apply ENNReal.ne_top_of_tsum_ne_top - apply privPostProcess_NonTopSum f h3 - . apply privPostProcess_NonTopRDNQ f h1 h5 h2 h4 h3 + sorry + -- cases h ; rename_i h1 h2 ; cases h2 ; rename_i h2 h3 ; cases h3 ; rename_i h3 h4 ; cases h4 ; rename_i h4 h5 + -- repeat any_goals constructor + -- . apply privPostProcess_zCDPBound h1 h2 h5 h4 h3 + -- . apply privPostProcess_NonZeroNQ h2 sur + -- . apply privPostProcess_NonTopSum f h3 + -- . simp [NonTopNQ] + -- intro l + -- apply ENNReal.ne_top_of_tsum_ne_top + -- apply privPostProcess_NonTopSum f h3 + -- . apply privPostProcess_NonTopRDNQ f h1 h5 h2 h4 h3 end SLang From cc447e5f844c1aa30d8a825e3a8b1644924e3a53 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Tue, 18 Jun 2024 09:32:30 -0400 Subject: [PATCH 055/122] composition compiles --- .../ZeroConcentrated/Composition.lean | 33 +++++++++++-------- 1 file changed, 20 insertions(+), 13 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean index b2ab5a37..26396498 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean @@ -37,7 +37,9 @@ lemma simp_α_1 {α : ℝ} (h : 1 < α) : 0 < α := by /-- The Renyi Divergence between neighbouring inputs of noised queries is nonzero. -/ -theorem Renyi_noised_query_NZ {nq : List T → SLang U} {α ε : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (h2 : Neighbour l₁ l₂) (h3 : zCDPBound nq ε) (h4 : NonZeroNQ nq) (h5 : NonTopRDNQ nq) (nts : NonTopNQ nq) : +theorem Renyi_noised_query_NZ {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) 1} {α ε : ℝ} + (h1 : 1 < α) {l₁ l₂ : List T} (h2 : Neighbour l₁ l₂) (h3 : zCDPBound nq HNorm ε) + (h4 : NonZeroNQ nq) (h5 : NonTopRDNQ nq) (nts : NonTopNQ nq) : (∑' (i : U), nq l₁ i ^ α * nq l₂ i ^ (1 - α)).toReal ≠ 0 := by simp [zCDPBound] at h3 replace h3 := h3 α h1 l₁ l₂ h2 @@ -72,11 +74,15 @@ theorem Renyi_noised_query_NZ {nq : List T → SLang U} {α ε : ℝ} (h1 : 1 < contradiction . exact h5 +lemma compose_norm (nq1 : List T → SLang U) (nq2 : List T → SLang V) (HNorm1 : ∀ l, HasSum (nq1 l) 1) (HNorm2 : ∀ l, HasSum (nq2 l) 1) : ∀ l, HasSum (privCompose nq1 nq2 l) 1 := sorry + + /-- Composed queries satisfy zCDP Renyi divergence bound. -/ -theorem privCompose_zCDPBound {nq1 : List T → SLang U} {nq2 : List T → SLang V} {ε₁ ε₂ ε₃ ε₄ : ℕ+} (h1 : zCDPBound nq1 ((ε₁ : ℝ) / ε₂)) (h2 : zCDPBound nq2 ((ε₃ : ℝ) / ε₄)) (nn1 : NonZeroNQ nq1) (nn2 : NonZeroNQ nq2) (nt1 : NonTopRDNQ nq1) (nt2 : NonTopRDNQ nq2) (nts1 : NonTopNQ nq1) (nts2 : NonTopNQ nq2) : - zCDPBound (privCompose nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by +theorem privCompose_zCDPBound {nq1 : List T → SLang U} {nq2 : List T → SLang V} {HNorm1 : ∀ l, HasSum (nq1 l) 1} {HNorm2 : ∀ l, HasSum (nq2 l) 1} {ε₁ ε₂ ε₃ ε₄ : ℕ+} + (h1 : zCDPBound nq1 HNorm1 ((ε₁ : ℝ) / ε₂)) (h2 : zCDPBound nq2 HNorm2 ((ε₃ : ℝ) / ε₄)) (nn1 : NonZeroNQ nq1) (nn2 : NonZeroNQ nq2) (nt1 : NonTopRDNQ nq1) (nt2 : NonTopRDNQ nq2) (nts1 : NonTopNQ nq1) (nts2 : NonTopNQ nq2) : + zCDPBound (privCompose nq1 nq2) (compose_norm nq1 nq2 HNorm1 HNorm2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by /- simp [privCompose, RenyiDivergence, zCDPBound] intro α h3 l₁ l₂ h4 @@ -234,16 +240,17 @@ theorem privCompose_NonTopRDNQ {nq1 : List T → SLang U} {nq2 : List T → SLan /-- ``privCompose`` satisfies zCDP -/ -theorem privCompose_zCDP (nq1 : List T → SLang U) (nq2 : List T → SLang V) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h : zCDP nq1 ((ε₁ : ℝ) / ε₂)) (h' : zCDP nq2 ((ε₃ : ℝ) / ε₄)) : - zCDP (privCompose nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by +theorem privCompose_zCDP (nq1 : List T → SLang U) (nq2 : List T → SLang V) (HNorm1 : ∀ l, HasSum (nq1 l) 1) (HNorm2 : ∀ l, HasSum (nq2 l) 1) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h : zCDP nq1 HNorm1 ((ε₁ : ℝ) / ε₂)) (h' : zCDP nq2 HNorm2 ((ε₃ : ℝ) / ε₄)) : + zCDP (privCompose nq1 nq2) (compose_norm nq1 nq2 HNorm1 HNorm2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by simp [zCDP] at * - cases h ; rename_i h1 h2 ; cases h2 ; rename_i h2 h3 ; cases h3 ; rename_i h3 h4 ; cases h4 ; rename_i h4 h5 - cases h' ; rename_i h'1 h'2 ; cases h'2 ; rename_i h'2 h'3 ; cases h'3 ; rename_i h'3 h'4 ; cases h'4 ; rename_i h'4 h'5 - repeat any_goals constructor - . apply privCompose_zCDPBound h1 h'1 h2 h'2 h5 h'5 h4 h'4 - . apply privCompose_NonZeroNQ h2 h'2 - . apply privCompose_NonTopSum h3 h'3 - . apply privCompose_NonTopNQ h4 h'4 - . apply privCompose_NonTopRDNQ h5 h'5 h4 h'4 + sorry + -- cases h ; rename_i h1 h2 ; cases h2 ; rename_i h2 h3 ; cases h3 ; rename_i h3 h4 ; cases h4 ; rename_i h4 h5 + -- cases h' ; rename_i h'1 h'2 ; cases h'2 ; rename_i h'2 h'3 ; cases h'3 ; rename_i h'3 h'4 ; cases h'4 ; rename_i h'4 h'5 + -- repeat any_goals constructor + -- . apply privCompose_zCDPBound h1 h'1 h2 h'2 h5 h'5 h4 h'4 + -- . apply privCompose_NonZeroNQ h2 h'2 + -- . apply privCompose_NonTopSum h3 h'3 + -- . apply privCompose_NonTopNQ h4 h'4 + -- . apply privCompose_NonTopRDNQ h5 h'5 h4 h'4 end SLang From ec6c9eb0ce12a7c59fbdc48596cc0d4e5c2e3ff8 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Tue, 18 Jun 2024 09:35:49 -0400 Subject: [PATCH 056/122] Perperties compiles --- .../ZeroConcentrated/Mechanism/Properties.lean | 16 ++++++++++------ 1 file changed, 10 insertions(+), 6 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean index 5fc90bdf..df8d26f9 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean @@ -18,11 +18,15 @@ open Classical Nat Int Real ENNReal MeasureTheory Measure namespace SLang +lemma privNoisedQuery_norm (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_sensitivity : sensitivity query Δ) : + ∀ l, HasSum (privNoisedQuery query Δ ε₁ ε₂ l) 1 := sorry + + /-- The zCDP mechanism with bounded sensitivity satisfies the bound for ``(Δε₂/ε₁)^2``-zCDP. -/ theorem privNoisedQuery_zCDPBound (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_sensitivity : sensitivity query Δ) : - zCDPBound (privNoisedQuery query Δ ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by + zCDPBound (privNoisedQuery query Δ ε₁ ε₂) (privNoisedQuery_norm query Δ ε₁ ε₂ bounded_sensitivity) ((ε₁ : ℝ) / ε₂) := by /- simp [zCDPBound, privNoisedQuery] intros α h1 l₁ l₂ h2 @@ -270,15 +274,15 @@ theorem privNoisedQuery_NonTopRDNQ (query : List T → ℤ) (Δ ε₁ ε₂ : The zCDP mechanism is ``(Δε₂/ε₁)^2``-zCDP. -/ theorem privNoisedQuery_zCDP (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_sensitivity : sensitivity query Δ) : - zCDP (privNoisedQuery query Δ ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by + zCDP (privNoisedQuery query Δ ε₁ ε₂) (privNoisedQuery_norm query Δ ε₁ ε₂ bounded_sensitivity) ((ε₁ : ℝ) / ε₂) := by simp [zCDP] repeat any_goals constructor . apply privNoisedQuery_zCDPBound exact bounded_sensitivity - . apply privNoisedQuery_NonZeroNQ - . apply privNoisedQuery_NonTopSum - . apply privNoisedQuery_NonTopNQ - . apply privNoisedQuery_NonTopRDNQ + -- . apply privNoisedQuery_NonZeroNQ + -- . apply privNoisedQuery_NonTopSum + -- . apply privNoisedQuery_NonTopNQ + -- . apply privNoisedQuery_NonTopRDNQ end SLang From 01c420967405884a7ce4a7274001874c40f81b93 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Tue, 18 Jun 2024 09:54:42 -0400 Subject: [PATCH 057/122] project compiles again --- SampCert/DifferentialPrivacy/Abstract.lean | 15 +++++++++++++++ .../ZeroConcentrated/Composition.lean | 8 +++----- .../DifferentialPrivacy/ZeroConcentrated/DP.lean | 7 +++---- .../ZeroConcentrated/Mechanism/Properties.lean | 5 ++--- .../ZeroConcentrated/Postprocessing.lean | 8 +++----- 5 files changed, 26 insertions(+), 17 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Abstract.lean b/SampCert/DifferentialPrivacy/Abstract.lean index cdf2720d..d69d753a 100644 --- a/SampCert/DifferentialPrivacy/Abstract.lean +++ b/SampCert/DifferentialPrivacy/Abstract.lean @@ -23,11 +23,18 @@ def NonZeroNQ (nq : List T → SLang U) := def NonTopSum (nq : List T → SLang U) := ∀ l : List T, ∑' n : U, nq l n ≠ ⊤ +def NormalMechanism (q : List T -> SLang U) : Prop := + ∀ l, HasSum (q l) 1 + namespace SLang abbrev Query (T U : Type) := List T → U abbrev Mechanism (T U : Type) := List T → SLang U + +-- FIXME: Move composition of normal mechanism here + + /-- Product of mechanisms. @@ -305,4 +312,12 @@ theorem privPostProcess_NonTopSum {nq : List T → SLang U} (f : U → V) (nt : rw [← A] trivial + +-- FIXME: Rename too + +lemma compose_norm (nq1 : List T → SLang U) (nq2 : List T → SLang V) (HNorm1 : NormalMechanism nq1) (HNorm2 : NormalMechanism nq2) : NormalMechanism (privCompose nq1 nq2) := sorry + +lemma pushforward_pmf_norm (nq : List T → SLang U) (HNorm : NormalMechanism nq1) (f : U -> V) : NormalMechanism (privPostProcess nq f) := + sorry + end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean index 26396498..f20e22d0 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean @@ -37,7 +37,7 @@ lemma simp_α_1 {α : ℝ} (h : 1 < α) : 0 < α := by /-- The Renyi Divergence between neighbouring inputs of noised queries is nonzero. -/ -theorem Renyi_noised_query_NZ {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) 1} {α ε : ℝ} +theorem Renyi_noised_query_NZ {nq : List T → SLang U} {HNorm : NormalMechanism nq} {α ε : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (h2 : Neighbour l₁ l₂) (h3 : zCDPBound nq HNorm ε) (h4 : NonZeroNQ nq) (h5 : NonTopRDNQ nq) (nts : NonTopNQ nq) : (∑' (i : U), nq l₁ i ^ α * nq l₂ i ^ (1 - α)).toReal ≠ 0 := by @@ -74,8 +74,6 @@ theorem Renyi_noised_query_NZ {nq : List T → SLang U} {HNorm : ∀ l, HasSum ( contradiction . exact h5 -lemma compose_norm (nq1 : List T → SLang U) (nq2 : List T → SLang V) (HNorm1 : ∀ l, HasSum (nq1 l) 1) (HNorm2 : ∀ l, HasSum (nq2 l) 1) : ∀ l, HasSum (privCompose nq1 nq2 l) 1 := sorry - /-- Composed queries satisfy zCDP Renyi divergence bound. @@ -240,8 +238,8 @@ theorem privCompose_NonTopRDNQ {nq1 : List T → SLang U} {nq2 : List T → SLan /-- ``privCompose`` satisfies zCDP -/ -theorem privCompose_zCDP (nq1 : List T → SLang U) (nq2 : List T → SLang V) (HNorm1 : ∀ l, HasSum (nq1 l) 1) (HNorm2 : ∀ l, HasSum (nq2 l) 1) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h : zCDP nq1 HNorm1 ((ε₁ : ℝ) / ε₂)) (h' : zCDP nq2 HNorm2 ((ε₃ : ℝ) / ε₄)) : - zCDP (privCompose nq1 nq2) (compose_norm nq1 nq2 HNorm1 HNorm2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by +theorem privCompose_zCDP (nq1 : List T → SLang U) (nq2 : List T → SLang V) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h : zCDP nq1 ((ε₁ : ℝ) / ε₂)) (h' : zCDP nq2 ((ε₃ : ℝ) / ε₄)) : + zCDP (privCompose nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by simp [zCDP] at * sorry -- cases h ; rename_i h1 h2 ; cases h2 ; rename_i h2 h3 ; cases h3 ; rename_i h3 h4 ; cases h4 ; rename_i h4 h5 diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean index 09718ec4..4758f1de 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean @@ -31,9 +31,9 @@ Inequality defining ``(ε^2)/2``-zCDP. All ``ε``-DP mechanisms satisfy this bound (though not all mechanisms satisfying this bound are ``ε``-DP). -/ -def zCDPBound (q : List T → SLang U) (HNorm : ∀ l, HasSum (q l) 1) (ε : ℝ) : Prop := +def zCDPBound (q : List T → SLang U) (HN : NormalMechanism q) (ε : ℝ) : Prop := ∀ α : ℝ, 1 < α → ∀ l₁ l₂ : List T, Neighbour l₁ l₂ → - RenyiDivergence (SLang.toPMF (q l₁) (HNorm l₁)) (SLang.toPMF (q l₂) (HNorm l₂)) α ≤ ENNReal.ofReal ((1/2) * ε ^ 2 * α) + RenyiDivergence (SLang.toPMF (q l₁) (HN l₁)) (SLang.toPMF (q l₂) (HN l₂)) α ≤ ENNReal.ofReal ((1/2) * ε ^ 2 * α) def NonTopNQ (nq : List T → SLang U) := ∀ l : List T, ∀ n : U, nq l n ≠ ⊤ @@ -48,8 +48,7 @@ def NonTopRDNQ (nq : List T → SLang U) : Prop := /-- The mechanism ``q`` is ``(ε^2)/2``-zCDP -/ -def zCDP (q : List T → SLang U) (HNorm : ∀ l, HasSum (q l) 1) (ε : ℝ) : Prop := - zCDPBound q HNorm ε +def zCDP (q : List T → SLang U) (ε : ℝ) : Prop := ∃ (HN : NormalMechanism q), zCDPBound q HN ε -- ∧ NonZeroNQ q -- ∧ NonTopSum q -- ∧ NonTopNQ q diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean index df8d26f9..de8055a2 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean @@ -19,8 +19,7 @@ open Classical Nat Int Real ENNReal MeasureTheory Measure namespace SLang lemma privNoisedQuery_norm (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_sensitivity : sensitivity query Δ) : - ∀ l, HasSum (privNoisedQuery query Δ ε₁ ε₂ l) 1 := sorry - + NormalMechanism (privNoisedQuery query Δ ε₁ ε₂) := sorry /-- The zCDP mechanism with bounded sensitivity satisfies the bound for ``(Δε₂/ε₁)^2``-zCDP. @@ -274,7 +273,7 @@ theorem privNoisedQuery_NonTopRDNQ (query : List T → ℤ) (Δ ε₁ ε₂ : The zCDP mechanism is ``(Δε₂/ε₁)^2``-zCDP. -/ theorem privNoisedQuery_zCDP (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_sensitivity : sensitivity query Δ) : - zCDP (privNoisedQuery query Δ ε₁ ε₂) (privNoisedQuery_norm query Δ ε₁ ε₂ bounded_sensitivity) ((ε₁ : ℝ) / ε₂) := by + zCDP (privNoisedQuery query Δ ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by simp [zCDP] repeat any_goals constructor . apply privNoisedQuery_zCDPBound diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean index 4c3369c6..9a0f6f19 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean @@ -554,10 +554,8 @@ theorem tsum_ne_zero_of_ne_zero {T : Type} [Inhabited T] (f : T → ENNReal) (h have B := CONTRA default contradiction -lemma pushforward_pmf_norm (nq : List T → SLang U) (HNorm : ∀ l, HasSum (nq l) 1) (f : U -> V) : (∀ l, HasSum (privPostProcess nq f l) 1) := - sorry -theorem privPostProcess_zCDPBound {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) 1} {ε₁ ε₂ : ℕ+} +theorem privPostProcess_zCDPBound {nq : List T → SLang U} {HNorm : NormalMechanism nq} {ε₁ ε₂ : ℕ+} (h : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (conv : NonTopSum nq) (f : U → V) : zCDPBound (privPostProcess nq f) (pushforward_pmf_norm nq HNorm f) ((ε₁ : ℝ) / ε₂) := by @@ -668,8 +666,8 @@ theorem privPostProcess_NonTopRDNQ {nq : List T → SLang U} {HNorm : ∀ l, Has Postprocessing preserves zCDP -/ theorem privPostProcess_zCDP {f : U → V} (sur : Function.Surjective f) - (nq : List T → SLang U) (HNorm : ∀ l, HasSum (nq l) 1) (ε₁ ε₂ : ℕ+) (h : zCDP nq HNorm ((ε₁ : ℝ) / ε₂)) : - zCDP (privPostProcess nq f) (pushforward_pmf_norm nq HNorm f) (((ε₁ : ℝ) / ε₂)) := by + (nq : List T → SLang U) (ε₁ ε₂ : ℕ+) (h : zCDP nq ((ε₁ : ℝ) / ε₂)) : + zCDP (privPostProcess nq f) (((ε₁ : ℝ) / ε₂)) := by simp [zCDP] at * sorry -- cases h ; rename_i h1 h2 ; cases h2 ; rename_i h2 h3 ; cases h3 ; rename_i h3 h4 ; cases h4 ; rename_i h4 h5 From 42999e7e5220dcac0405989d4f6514060b85e388 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Tue, 18 Jun 2024 11:12:35 -0400 Subject: [PATCH 058/122] checkpoint --- SampCert/DifferentialPrivacy/Abstract.lean | 24 ++++++++++++++----- .../DifferentialPrivacy/RenyiDivergence.lean | 1 + .../ZeroConcentrated/Composition.lean | 2 +- .../ZeroConcentrated/Postprocessing.lean | 2 +- SampCert/Util/Gaussian/DiscreteGaussian.lean | 16 ++++++------- 5 files changed, 29 insertions(+), 16 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Abstract.lean b/SampCert/DifferentialPrivacy/Abstract.lean index d69d753a..f62f63d0 100644 --- a/SampCert/DifferentialPrivacy/Abstract.lean +++ b/SampCert/DifferentialPrivacy/Abstract.lean @@ -312,12 +312,24 @@ theorem privPostProcess_NonTopSum {nq : List T → SLang U} (f : U → V) (nt : rw [← A] trivial +-- Provable using stuff in the adaptive composition PR, wait merge that stuff later +lemma privCompose_norm (nq1 : List T → SLang U) (nq2 : List T → SLang V) (HNorm1 : NormalMechanism nq1) (HNorm2 : NormalMechanism nq2) : NormalMechanism (privCompose nq1 nq2) := by + rw [NormalMechanism] at * + intro l + have HNorm1' := HasSum.tsum_eq (HNorm1 l) + have HNorm2' := HasSum.tsum_eq (HNorm2 l) + have HR : (∑' (x : U × V), privCompose nq1 nq2 l x = 1) := sorry + rw [<- HR] + apply Summable.hasSum + exact ENNReal.summable --- FIXME: Rename too - -lemma compose_norm (nq1 : List T → SLang U) (nq2 : List T → SLang V) (HNorm1 : NormalMechanism nq1) (HNorm2 : NormalMechanism nq2) : NormalMechanism (privCompose nq1 nq2) := sorry - -lemma pushforward_pmf_norm (nq : List T → SLang U) (HNorm : NormalMechanism nq1) (f : U -> V) : NormalMechanism (privPostProcess nq f) := - sorry +lemma privPostProcess_norm (nq : List T → SLang U) (HNorm : NormalMechanism nq) (f : U -> V) : NormalMechanism (privPostProcess nq f) := by + rw [NormalMechanism] at * + intro l + have HNorm' := HasSum.tsum_eq (HNorm l) + have HR : (∑' (b : V), privPostProcess nq f l b = 1) := sorry + rw [<- HR] + apply Summable.hasSum + exact ENNReal.summable end SLang diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index bba85490..95fdd4f9 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -387,6 +387,7 @@ theorem Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] [C rw [<- ENNReal.tsum_toReal_eq] · rw [ENNReal.ofReal_toReal] -- Could do another case at the top if not derivable + -- Want to rewrite to ∑(a : T), p a sorry · -- Derivable from Hnts sorry diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean index f20e22d0..c9192827 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean @@ -80,7 +80,7 @@ Composed queries satisfy zCDP Renyi divergence bound. -/ theorem privCompose_zCDPBound {nq1 : List T → SLang U} {nq2 : List T → SLang V} {HNorm1 : ∀ l, HasSum (nq1 l) 1} {HNorm2 : ∀ l, HasSum (nq2 l) 1} {ε₁ ε₂ ε₃ ε₄ : ℕ+} (h1 : zCDPBound nq1 HNorm1 ((ε₁ : ℝ) / ε₂)) (h2 : zCDPBound nq2 HNorm2 ((ε₃ : ℝ) / ε₄)) (nn1 : NonZeroNQ nq1) (nn2 : NonZeroNQ nq2) (nt1 : NonTopRDNQ nq1) (nt2 : NonTopRDNQ nq2) (nts1 : NonTopNQ nq1) (nts2 : NonTopNQ nq2) : - zCDPBound (privCompose nq1 nq2) (compose_norm nq1 nq2 HNorm1 HNorm2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by + zCDPBound (privCompose nq1 nq2) (privCompose_norm nq1 nq2 HNorm1 HNorm2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by /- simp [privCompose, RenyiDivergence, zCDPBound] intro α h3 l₁ l₂ h4 diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean index 9a0f6f19..4422930d 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean @@ -558,7 +558,7 @@ theorem tsum_ne_zero_of_ne_zero {T : Type} [Inhabited T] (f : T → ENNReal) (h theorem privPostProcess_zCDPBound {nq : List T → SLang U} {HNorm : NormalMechanism nq} {ε₁ ε₂ : ℕ+} (h : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (conv : NonTopSum nq) (f : U → V) : - zCDPBound (privPostProcess nq f) (pushforward_pmf_norm nq HNorm f) ((ε₁ : ℝ) / ε₂) := by + zCDPBound (privPostProcess nq f) (privPostProcess_norm nq HNorm f) ((ε₁ : ℝ) / ε₂) := by /- simp [privPostProcess, zCDPBound, RenyiDivergence] intro α h1 l₁ l₂ h2 diff --git a/SampCert/Util/Gaussian/DiscreteGaussian.lean b/SampCert/Util/Gaussian/DiscreteGaussian.lean index 9b4a82ac..77cc1e7e 100644 --- a/SampCert/Util/Gaussian/DiscreteGaussian.lean +++ b/SampCert/Util/Gaussian/DiscreteGaussian.lean @@ -337,15 +337,15 @@ theorem discrete_gaussian_normalizes {σ : ℝ} (h : σ ≠ 0) (μ : ℝ) : apply sum_gauss_term_ne_zero h +def discrete_gaussian_normal {σ : ℝ} (h : σ ≠ 0) (μ : ℝ) : HasSum (fun z : ℤ => ENNReal.ofReal (discrete_gaussian σ μ z)) 1 := by + let HR := discrete_gaussian_normalizes h μ + have HS := Summable.hasSum (discrete_gaussian_summable' h μ) + simp at HR + rw [HR] at HS + sorry /-- -The discrete Gaussian as a PMF +The discrete Gaussian as a PMF when evaluated over ℤ -/ def discrete_gaussian_pmf {σ : ℝ} (h : σ ≠ 0) (μ : ℝ) : PMF ℤ := - ⟨ (fun n => ENNReal.ofReal (discrete_gaussian σ μ n)), - by - let Hsummable := Summable.hasSum (discrete_gaussian_summable' h μ) - rw [discrete_gaussian_normalizes h μ] at Hsummable - -- Might be a lemma to do this already - sorry - ⟩ + ⟨ (fun z : ℤ => ENNReal.ofReal (discrete_gaussian σ μ z)) , discrete_gaussian_normal h μ ⟩ From b7c99a2293c94748b6c1ab17669ef69f91cda1da Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Tue, 18 Jun 2024 11:12:45 -0400 Subject: [PATCH 059/122] checkpoint --- SampCert/DifferentialPrivacy/RenyiDivergence.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index 95fdd4f9..abc59e74 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -387,9 +387,9 @@ theorem Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] [C rw [<- ENNReal.tsum_toReal_eq] · rw [ENNReal.ofReal_toReal] -- Could do another case at the top if not derivable - -- Want to rewrite to ∑(a : T), p a + -- Want to bound above my ∑' sorry - · -- Derivable from Hnts + · -- Bound above by p a sorry · -- Special case: There exists some element x0 with p x0 = ⊤ but q x0 ∈ ℝ+ rename_i Hspecial From 8529cb723039da8408d5d281b2517d58f5b90d85 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Tue, 18 Jun 2024 11:57:54 -0400 Subject: [PATCH 060/122] PMF cancellation lemmas --- .../DifferentialPrivacy/RenyiDivergence.lean | 59 +++++++++++-------- SampCert/Util/Log.lean | 41 +++++++++++++ 2 files changed, 75 insertions(+), 25 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index abc59e74..58ba70c9 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -20,6 +20,40 @@ This file defines the Renyi divergence and equations for evaluating its expectat open Real ENNReal PMF Nat Int MeasureTheory Measure PMF open Classical +/-- +Simplified consequence of absolute continuity (remove me?) +-/ +def AbsCts (p q : T -> ENNReal) : Prop := ∀ x : T, q x = 0 -> p x = 0 + +/-- +Specialize the definitation of AbsolutelyContinuous when singletons are measurable +-/ +lemma PMF_AbsCts [MeasurableSpace T] [MeasurableSingletonClass T] (p q : PMF T) (H : AbsolutelyContinuous (PMF.toMeasure p) (PMF.toMeasure q)) : AbsCts p q := by + rw [AbsolutelyContinuous] at H + rw [AbsCts] + intro x Hx + have Hxm : q.toMeasure { x } = 0 := by + rw [toMeasure] + simp + apply (toOuterMeasure_apply_eq_zero_iff q {x}).mpr + exact Set.disjoint_singleton_right.mpr fun a => a Hx + have H := H Hxm + rw [toMeasure] at H + simp at * + have Hp : Disjoint p.support {x} := (toOuterMeasure_apply_eq_zero_iff p {x}).mp H + simp at Hp + assumption + +lemma PMF_mul_mul_inv_eq_mul_cancel (p q : PMF T) (HA : AbsCts p q) (a : T) : (p a * q a) / q a = p a := by + apply mul_mul_inv_eq_mul_cancel + · rw [AbsCts] at HA + intro + simp_all + · simp + have HK : (q a ≠ ⊤) := apply_ne_top q a + simp_all only [ne_eq, not_false_eq_true] + simp + variable {T : Type} /-- @@ -46,29 +80,6 @@ lemma RenyiDivergence_def_exp (p q : PMF T) {α : ℝ} (h : 1 < α) : · linarith simp [H2] -/-- -Simplified consequence of absolute continuity (remove me?) --/ -def AbsCts (p q : T -> ENNReal) : Prop := ∀ x : T, q x = 0 -> p x = 0 - -/-- -Specialize the definitation of AbsolutelyContinuous when singletons are measurable --/ -lemma PMF_AbsCts [MeasurableSpace T] [MeasurableSingletonClass T] (p q : PMF T) (H : AbsolutelyContinuous (PMF.toMeasure p) (PMF.toMeasure q)) : AbsCts p q := by - rw [AbsolutelyContinuous] at H - rw [AbsCts] - intro x Hx - have Hxm : q.toMeasure { x } = 0 := by - rw [toMeasure] - simp - apply (toOuterMeasure_apply_eq_zero_iff q {x}).mpr - exact Set.disjoint_singleton_right.mpr fun a => a Hx - have H := H Hxm - rw [toMeasure] at H - simp at * - have Hp : Disjoint p.support {x} := (toOuterMeasure_apply_eq_zero_iff p {x}).mp H - simp at Hp - assumption /-- Closed form of the series in the definition of the Renyi divergence. @@ -438,8 +449,6 @@ lemma RenyiDivergence_mono_sum (x y : ℝ) (α : ℝ) (h : 1 < α) : (Real.exp ( · exact exp_le_exp.mp H · linarith -#check True - theorem RenyiDivergence_def_nonneg [MeasurableSpace T] [MeasurableSingletonClass T] [Countable T] (p q : PMF T) (Hpq : AbsCts p q) {α : ℝ} (Hα : 1 < α) : (0 ≤ RenyiDivergence_def p q α) := by have H1 : eexp (((α - 1)) * 0) ≤ eexp ((α - 1) * RenyiDivergence_def p q α) := by diff --git a/SampCert/Util/Log.lean b/SampCert/Util/Log.lean index 59da0c49..23a9e057 100644 --- a/SampCert/Util/Log.lean +++ b/SampCert/Util/Log.lean @@ -208,4 +208,45 @@ lemma toEReal_ofENNReal_nonneg (H : 0 ≤ w) : ENNReal.toEReal (ofEReal w) = w : @[simp] lemma ofEReal_toENNReal : ofEReal (ENNReal.toEReal x) = x := by sorry + +-- ENNReal inequalities +-- These are needed to prove the extensded version of Jensen's inequality +lemma mul_mul_inv_le_mul_cancel : (x * y) * y⁻¹ ≤ x := by + cases x + · simp_all + rename_i x' + cases (Classical.em (x' = 0)) + · simp_all + rename_i Hx' + cases y + · simp_all + rename_i y' + cases (Classical.em (y' = 0)) + · simp_all + rename_i Hy' + simp + rw [← coe_inv Hy'] + rw [← coe_mul] + rw [← coe_mul] + rw [mul_inv_cancel_right₀ Hy' x'] + +lemma mul_mul_inv_eq_mul_cancel (H : y = 0 -> x = 0) (H2 : ¬(x ≠ 0 ∧ y = ⊤)) : (x * y) * y⁻¹ = x := by + cases x + · simp_all + rename_i x' + cases (Classical.em (x' = 0)) + · simp_all + rename_i Hx' + cases y + · simp_all + rename_i y' + cases (Classical.em (y' = 0)) + · simp_all + rename_i Hy' + simp + rw [← coe_inv Hy'] + rw [← coe_mul] + rw [← coe_mul] + rw [mul_inv_cancel_right₀ Hy' x'] + end ENNReal From 936bc8b24930bbaebaf89fc1194a0e927ca84fa1 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Tue, 18 Jun 2024 12:11:53 -0400 Subject: [PATCH 061/122] close reduct proof --- .../DifferentialPrivacy/RenyiDivergence.lean | 37 +++++++++++++++---- SampCert/Util/Log.lean | 6 ++- 2 files changed, 33 insertions(+), 10 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index 58ba70c9..4fb86f89 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -44,7 +44,7 @@ lemma PMF_AbsCts [MeasurableSpace T] [MeasurableSingletonClass T] (p q : PMF T) simp at Hp assumption -lemma PMF_mul_mul_inv_eq_mul_cancel (p q : PMF T) (HA : AbsCts p q) (a : T) : (p a * q a) / q a = p a := by +lemma PMF_mul_mul_inv_eq_mul_cancel (p q : PMF T) (HA : AbsCts p q) (a : T) : (p a / q a) * q a = p a := by apply mul_mul_inv_eq_mul_cancel · rw [AbsCts] at HA intro @@ -267,7 +267,13 @@ noncomputable def Renyi_Jensen_f (p q : PMF T) : T -> ℝ := (fun z => ((p z / q lemma Renyi_Jensen_rw (p q : PMF T) {α : ℝ} (h : 1 < α) (H : AbsCts p q) (Hspecial : ∀ x : T, ¬(p x = ⊤ ∧ q x ≠ 0 ∧ q x ≠ ⊤)) (x : T) : (p x / q x)^α * (q x) = ENNReal.ofReal (((Renyi_Jensen_f p q) x)^α * (q x).toReal) := sorry -theorem Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] [Countable T] (p q : PMF T) {α : ℝ} (h : 1 < α) (H : AbsCts p q) : + +-- FIXME: might be able to simplify this argument with the new rewrite lemmas +/-- +Jensen's inquality applied to ENNReals, in the case that q is nonzero +-/ +lemma Renyi_Jensen_ENNReal_reduct [MeasurableSpace T] [MeasurableSingletonClass T] [Countable T] + (p q : PMF T) {α : ℝ} (h : 1 < α) (H : AbsCts p q) (Hq : ∀ t, q t ≠ 0): (∑' x : T, (p x / q x) * q x) ^ α ≤ (∑' x : T, (p x / q x) ^ α * q x) := by have Hdiscr : DiscreteMeasurableSpace T := MeasurableSingletonClass.toDiscreteMeasurableSpace cases (Classical.em (∑' (a : T), (p a / q a) ^ α * q a ≠ ⊤)) @@ -313,12 +319,9 @@ theorem Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] [C rw [AbsCts] at H simp_all only [ne_eq, not_and, Decidable.not_not, ENNReal.zero_div, zero_ne_top] · rename_i HK' - rcases HK' with ⟨ HK1 , HK2 ⟩ + rcases HK' with ⟨ HK1 , _ ⟩ apply (Hspecial a) simp_all - -- We need to eliminate the possibility that (q a = 0) - -- All summads with q a = 0 will have value 0, so they should be removable. - sorry have Hsum_indicator (a : T) : ∑' (i : T), q i * Set.indicator {a} (fun x => 1) i = q a := by have Hfun : (fun (i : T) => q i * Set.indicator {a} (fun x => 1) i) = (fun (i : T) => if i = a then q a else 0) := by funext i @@ -399,9 +402,18 @@ theorem Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] [C · rw [ENNReal.ofReal_toReal] -- Could do another case at the top if not derivable -- Want to bound above my ∑' - sorry + conv => + arg 1 + arg 1 + intro a + rw [PMF_mul_mul_inv_eq_mul_cancel p q H] + exact tsum_coe_ne_top p · -- Bound above by p a - sorry + intro a + conv => + arg 1 + rw [PMF_mul_mul_inv_eq_mul_cancel p q H] + exact apply_ne_top p a · -- Special case: There exists some element x0 with p x0 = ⊤ but q x0 ∈ ℝ+ rename_i Hspecial simp at * @@ -439,6 +451,15 @@ theorem Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] [C rw [HStop] exact OrderTop.le_top ((∑' (x : T), p x / q x * q x) ^ α) +/-- +Jensen's inquality applied to ENNReals +-/ +theorem Renyi_Jensen_ENNReal[MeasurableSpace T] [MeasurableSingletonClass T] [Countable T] (p q : PMF T) {α : ℝ} (h : 1 < α) (H : AbsCts p q) : + (∑' x : T, (p x / q x) * q x) ^ α ≤ (∑' x : T, (p x / q x) ^ α * q x) := by + sorry + + + -- FIXME /-- The Renyi divergence is monotonic in the value of its sum. diff --git a/SampCert/Util/Log.lean b/SampCert/Util/Log.lean index 23a9e057..b882466e 100644 --- a/SampCert/Util/Log.lean +++ b/SampCert/Util/Log.lean @@ -211,7 +211,7 @@ lemma ofEReal_toENNReal : ofEReal (ENNReal.toEReal x) = x := by sorry -- ENNReal inequalities -- These are needed to prove the extensded version of Jensen's inequality -lemma mul_mul_inv_le_mul_cancel : (x * y) * y⁻¹ ≤ x := by +lemma mul_mul_inv_le_mul_cancel : (x * y⁻¹) * y ≤ x := by cases x · simp_all rename_i x' @@ -228,9 +228,10 @@ lemma mul_mul_inv_le_mul_cancel : (x * y) * y⁻¹ ≤ x := by rw [← coe_inv Hy'] rw [← coe_mul] rw [← coe_mul] + rw [mul_right_comm] rw [mul_inv_cancel_right₀ Hy' x'] -lemma mul_mul_inv_eq_mul_cancel (H : y = 0 -> x = 0) (H2 : ¬(x ≠ 0 ∧ y = ⊤)) : (x * y) * y⁻¹ = x := by +lemma mul_mul_inv_eq_mul_cancel (H : y = 0 -> x = 0) (H2 : ¬(x ≠ 0 ∧ y = ⊤)) : (x * y⁻¹) * y = x := by cases x · simp_all rename_i x' @@ -247,6 +248,7 @@ lemma mul_mul_inv_eq_mul_cancel (H : y = 0 -> x = 0) (H2 : ¬(x ≠ 0 ∧ y = rw [← coe_inv Hy'] rw [← coe_mul] rw [← coe_mul] + rw [mul_right_comm] rw [mul_inv_cancel_right₀ Hy' x'] end ENNReal From 91db8e07e17d948534ebfa9ff86fc2dcf6df27ef Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Tue, 18 Jun 2024 12:26:15 -0400 Subject: [PATCH 062/122] close some side conditions in Renyi --- .../DifferentialPrivacy/RenyiDivergence.lean | 36 ++++++++++++++++--- 1 file changed, 32 insertions(+), 4 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index 4fb86f89..8b823d44 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -265,7 +265,31 @@ noncomputable def Renyi_Jensen_f (p q : PMF T) : T -> ℝ := (fun z => ((p z / q -- Except for one case, we can rewrite the ENNReal-valued inequality into the form Jenen's inequality expects. lemma Renyi_Jensen_rw (p q : PMF T) {α : ℝ} (h : 1 < α) (H : AbsCts p q) (Hspecial : ∀ x : T, ¬(p x = ⊤ ∧ q x ≠ 0 ∧ q x ≠ ⊤)) (x : T) : - (p x / q x)^α * (q x) = ENNReal.ofReal (((Renyi_Jensen_f p q) x)^α * (q x).toReal) := sorry + (p x / q x)^α * (q x) = ENNReal.ofReal (((Renyi_Jensen_f p q) x)^α * (q x).toReal) := by + simp [Renyi_Jensen_f] + rw [ENNReal.toReal_rpow] + rw [<- ENNReal.toReal_mul] + rw [ENNReal.ofReal_toReal] + apply mul_ne_top + · apply rpow_ne_top_of_nonneg + · linarith + · intro HK + apply ENNReal.div_eq_top.mp at HK + simp at HK + rw [AbsCts] at H + cases HK + · rename_i HK + rcases HK with ⟨ HK1, HK2 ⟩ + simp_all + · rename_i HK + rcases HK with ⟨ HK1, HK2 ⟩ + simp_all + apply HK2 + apply (Hspecial) + · apply HK1 + · intro Hcont1 + simp_all + · exact apply_ne_top q x -- FIXME: might be able to simplify this argument with the new rewrite lemmas @@ -454,7 +478,7 @@ lemma Renyi_Jensen_ENNReal_reduct [MeasurableSpace T] [MeasurableSingletonClass /-- Jensen's inquality applied to ENNReals -/ -theorem Renyi_Jensen_ENNReal[MeasurableSpace T] [MeasurableSingletonClass T] [Countable T] (p q : PMF T) {α : ℝ} (h : 1 < α) (H : AbsCts p q) : +theorem Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] [Countable T] (p q : PMF T) {α : ℝ} (h : 1 < α) (H : AbsCts p q) : (∑' x : T, (p x / q x) * q x) ^ α ≤ (∑' x : T, (p x / q x) ^ α * q x) := by sorry @@ -478,8 +502,12 @@ theorem RenyiDivergence_def_nonneg [MeasurableSpace T] [MeasurableSingletonClass simp apply (le_trans ?G1 (Renyi_Jensen_ENNReal p q Hα Hpq)) have Hone : (∑' (x : T), p x / q x * q x) = 1 := by - -- This argument is also necessary to finish the Jensen's inequality proof (filter out the q x = 0 elements) - sorry + conv => + arg 1 + arg 1 + intro x + rw [PMF_mul_mul_inv_eq_mul_cancel p q Hpq] + exact tsum_coe p have Hle : (∑' (x : T), p x / q x * q x) ≤ (∑' (x : T), p x / q x * q x) ^ α := by apply ENNReal.le_rpow_self_of_one_le · rw [Hone] From 1e76fc88d1f6784bebe281baee5ab2d3ec35a125 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Tue, 18 Jun 2024 16:02:44 -0400 Subject: [PATCH 063/122] Reduction for the full Renyi case --- .../DifferentialPrivacy/RenyiDivergence.lean | 74 ++++++++++++++++++- 1 file changed, 71 insertions(+), 3 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index 8b823d44..5e17a5af 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -297,7 +297,7 @@ lemma Renyi_Jensen_rw (p q : PMF T) {α : ℝ} (h : 1 < α) (H : AbsCts p q) (Hs Jensen's inquality applied to ENNReals, in the case that q is nonzero -/ lemma Renyi_Jensen_ENNReal_reduct [MeasurableSpace T] [MeasurableSingletonClass T] [Countable T] - (p q : PMF T) {α : ℝ} (h : 1 < α) (H : AbsCts p q) (Hq : ∀ t, q t ≠ 0): + (p q : PMF T) {α : ℝ} (h : 1 < α) (H : AbsCts p q) (Hq : ∀ t, q t ≠ 0) : (∑' x : T, (p x / q x) * q x) ^ α ≤ (∑' x : T, (p x / q x) ^ α * q x) := by have Hdiscr : DiscreteMeasurableSpace T := MeasurableSingletonClass.toDiscreteMeasurableSpace cases (Classical.em (∑' (a : T), (p a / q a) ^ α * q a ≠ ⊤)) @@ -332,6 +332,7 @@ lemma Renyi_Jensen_ENNReal_reduct [MeasurableSpace T] [MeasurableSingletonClass intro x rw [<- ENNReal.toReal_mul] apply ENNReal.summable_toReal + assumption have HRJf_nonneg (a : T) : 0 <= Renyi_Jensen_f p q a := by apply toReal_nonneg have HRJf_nt (a : T) : p a / q a ≠ ⊤ := by @@ -475,13 +476,80 @@ lemma Renyi_Jensen_ENNReal_reduct [MeasurableSpace T] [MeasurableSingletonClass rw [HStop] exact OrderTop.le_top ((∑' (x : T), p x / q x * q x) ^ α) + +/-- +Restriction of the PMF f to the support of q +-/ +def reducedPMF_def (f q : PMF T) (x : { t : T // ¬q t = 0 }) : ENNReal := f x.val + +-- Should be provable +lemma acts_refl (q : PMF T) : AbsCts q q := by + rw [AbsCts] + simp + + +lemma reducedPMF_norm_acts (p q : PMF T) (H : AbsCts p q) : HasSum (reducedPMF_def p q) 1 := by + have H1 : Summable (reducedPMF_def p q) := by exact ENNReal.summable + have H2 := Summable.hasSum H1 + have H3 : (∑' (b : { t // q t ≠ 0 }), reducedPMF_def p q b) = 1 := by + -- Provable? + sorry + rw [<- H3] + apply H2 + +noncomputable def reducedPMF {p q : PMF T} (H : AbsCts p q): PMF { t : T // ¬q t = 0 } := + ⟨ reducedPMF_def p q, reducedPMF_norm_acts p q H ⟩ + +lemma reducedPMF_pos {q : PMF T} (H : AbsCts q q) (a : T) (Ha : ¬q a = 0): (reducedPMF H) ⟨a, Ha⟩ ≠ 0 := by + simp + rw [reducedPMF] + unfold reducedPMF_def + rw [DFunLike.coe] + rw [PMF.instFunLike] + simp + apply Ha + /-- Jensen's inquality applied to ENNReals -/ -theorem Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] [Countable T] (p q : PMF T) {α : ℝ} (h : 1 < α) (H : AbsCts p q) : +theorem Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] [Countable T] (p q : PMF T) {α : ℝ} (h : 1 < α) (Hac : AbsCts p q) : (∑' x : T, (p x / q x) * q x) ^ α ≤ (∑' x : T, (p x / q x) ^ α * q x) := by - sorry + have K1 : Function.support (fun x : T => (p x / q x) * q x) ⊆ { t : T | q t ≠ 0 } := sorry + have K2 : Function.support (fun x : T => (p x / q x)^α * q x) ⊆ { t : T | q t ≠ 0 } := sorry + rw [<- tsum_subtype_eq_of_support_subset K1] + rw [<- tsum_subtype_eq_of_support_subset K2] + simp + + have Hq : AbsCts q q := acts_refl q + + have B1 (x : { x // ¬q x = 0 }) : p ↑x / q ↑x * q ↑x = reducedPMF Hac x / reducedPMF Hq x * reducedPMF Hq x := by congr + have B2 (x : { x // ¬q x = 0 }) : (p ↑x / q ↑x)^α * q ↑x = (reducedPMF Hac x / reducedPMF Hq x)^α * reducedPMF Hq x := by congr + conv => + congr + · arg 1 + arg 1 + intro x + rw [B1 x] + · arg 1 + intro x + rw [B2 x] + + clear B1 + clear B2 + clear K1 + clear K2 + + apply Renyi_Jensen_ENNReal_reduct + · apply h + · rw [AbsCts] + simp + intro a Ha Hcont + exfalso + apply (reducedPMF_pos Hq a Ha Hcont) + · intro t + rcases t with ⟨ a , Ha ⟩ + apply (reducedPMF_pos Hq a Ha) -- FIXME From c120a5cf45af4ea2e0615a8f8f40ce8c1a4bc9d2 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Wed, 19 Jun 2024 09:44:21 -0400 Subject: [PATCH 064/122] close nonnegativity proof for Renyi divergence --- .../DifferentialPrivacy/RenyiDivergence.lean | 45 ++++++++--- SampCert/Util/Log.lean | 77 +++++++++++++++++++ 2 files changed, 111 insertions(+), 11 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index 5e17a5af..d4e2d2ec 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -487,13 +487,25 @@ lemma acts_refl (q : PMF T) : AbsCts q q := by rw [AbsCts] simp - lemma reducedPMF_norm_acts (p q : PMF T) (H : AbsCts p q) : HasSum (reducedPMF_def p q) 1 := by have H1 : Summable (reducedPMF_def p q) := by exact ENNReal.summable have H2 := Summable.hasSum H1 have H3 : (∑' (b : { t // q t ≠ 0 }), reducedPMF_def p q b) = 1 := by - -- Provable? - sorry + have K1 : Function.support (fun x => p x) ⊆ { t : T | q t ≠ 0 } := by + rw [Function.support] + simp + intro a Hp Hcont + rw [AbsCts] at H + apply Hp + apply H + apply Hcont + have S1 : ∑' (x : ↑{t | q t ≠ 0}), p ↑x = ∑' (x : T), p x := by + apply tsum_subtype_eq_of_support_subset K1 + have T1 : ∑' (x : T), p x = 1 := by exact tsum_coe p + rw [<- T1] + rw [<- S1] + simp + rfl rw [<- H3] apply H2 @@ -515,8 +527,8 @@ Jensen's inquality applied to ENNReals theorem Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] [Countable T] (p q : PMF T) {α : ℝ} (h : 1 < α) (Hac : AbsCts p q) : (∑' x : T, (p x / q x) * q x) ^ α ≤ (∑' x : T, (p x / q x) ^ α * q x) := by - have K1 : Function.support (fun x : T => (p x / q x) * q x) ⊆ { t : T | q t ≠ 0 } := sorry - have K2 : Function.support (fun x : T => (p x / q x)^α * q x) ⊆ { t : T | q t ≠ 0 } := sorry + have K1 : Function.support (fun x : T => (p x / q x) * q x) ⊆ { t : T | q t ≠ 0 } := by simp [Function.support] + have K2 : Function.support (fun x : T => (p x / q x)^α * q x) ⊆ { t : T | q t ≠ 0 } := by simp [Function.support] rw [<- tsum_subtype_eq_of_support_subset K1] rw [<- tsum_subtype_eq_of_support_subset K2] simp @@ -552,6 +564,8 @@ theorem Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] [C apply (reducedPMF_pos Hq a Ha) + + -- FIXME /-- The Renyi divergence is monotonic in the value of its sum. @@ -562,6 +576,8 @@ lemma RenyiDivergence_mono_sum (x y : ℝ) (α : ℝ) (h : 1 < α) : (Real.exp ( · exact exp_le_exp.mp H · linarith +set_option pp.coercions false + theorem RenyiDivergence_def_nonneg [MeasurableSpace T] [MeasurableSingletonClass T] [Countable T] (p q : PMF T) (Hpq : AbsCts p q) {α : ℝ} (Hα : 1 < α) : (0 ≤ RenyiDivergence_def p q α) := by have H1 : eexp (((α - 1)) * 0) ≤ eexp ((α - 1) * RenyiDivergence_def p q α) := by @@ -583,12 +599,19 @@ theorem RenyiDivergence_def_nonneg [MeasurableSpace T] [MeasurableSingletonClass apply le_trans ?X Hle rw [Hone] apply eexp_mono_le.mpr at H1 - have HX : (0 < (α.toEReal - 1)) := by sorry - have HX1 : (ofEReal ((↑α - 1) * 0) ≤ ofEReal ((↑α - 1) * RenyiDivergence_def p q α)) := by - exact ofEReal_le_mono.mp fun a => H1 - rw [ofEReal_mul] at HX1 - -- Circular side condition? Why is cancelling EReals so hard? Make a lemma with the iff cases in Log - all_goals sorry + have Hone : (OfNat.ofNat 1 = Real.toEReal (1 : ℝ)) := by simp + have Hzero : (OfNat.ofNat 0 = Real.toEReal (0 : ℝ)) := by simp + + apply ereal_smul_l (α.toEReal - OfNat.ofNat 1) + · rw [Hone] + rw [<- EReal.coe_sub] + rw [Hzero] + apply EReal.coe_lt_coe_iff.mpr + exact sub_pos.mpr Hα + · rw [Hone] + rw [<- EReal.coe_sub] + exact EReal.coe_lt_top (α - OfNat.ofNat 1) + · assumption theorem RenyiDivergence_def_zero (p q : PMF T) (α : ℝ) : p = q <-> (0 = RenyiDivergence_def p q α) := by -- See paper diff --git a/SampCert/Util/Log.lean b/SampCert/Util/Log.lean index b882466e..82851234 100644 --- a/SampCert/Util/Log.lean +++ b/SampCert/Util/Log.lean @@ -209,6 +209,7 @@ lemma toEReal_ofENNReal_nonneg (H : 0 ≤ w) : ENNReal.toEReal (ofEReal w) = w : lemma ofEReal_toENNReal : ofEReal (ENNReal.toEReal x) = x := by sorry + -- ENNReal inequalities -- These are needed to prove the extensded version of Jensen's inequality lemma mul_mul_inv_le_mul_cancel : (x * y⁻¹) * y ≤ x := by @@ -251,4 +252,80 @@ lemma mul_mul_inv_eq_mul_cancel (H : y = 0 -> x = 0) (H2 : ¬(x ≠ 0 ∧ y = rw [mul_right_comm] rw [mul_inv_cancel_right₀ Hy' x'] + +-- Could be shortened +lemma ereal_smul_l (s : EReal) (Hr1 : 0 < s) (Hr2 : s < ⊤) (H : s * w ≤ s * z) : w ≤ z := by + have defTop : some none = (⊤ : EReal) := by simp [Top.top] + have defBot : none = (⊥ : EReal) := by simp [Bot.bot] + + cases s + · exfalso + rw [defBot] at Hr1 + simp_all only [not_lt_bot] + rename_i s_nnr + cases s_nnr + · rw [defTop] at Hr2 + exfalso + simp_all only [EReal.zero_lt_top, lt_self_iff_false] + rename_i s_R + have Hsr : some (some s_R) = Real.toEReal s_R := by simp [Real.toEReal] + rw [Hsr] at H + rw [Hsr] at Hr1 + rw [Hsr] at Hr2 + clear Hsr + + cases w + · apply left_eq_inf.mp + rfl + rename_i w_nnr + cases w_nnr + · simp [defTop] at H + rw [EReal.mul_top_of_pos Hr1] at H + have X1 : z = ⊤ := by + cases z + · exfalso + simp at H + rw [defBot] at H + rw [EReal.mul_bot_of_pos] at H + · cases H + · apply Hr1 + rename_i z_nnr + cases z_nnr + · simp [Top.top] + exfalso + apply top_le_iff.mp at H + rename_i z_R + have Hzr : some (some z_R) = Real.toEReal z_R := by simp [Real.toEReal] + rw [Hzr] at H + rw [<- EReal.coe_mul] at H + cases H + rw [defTop, X1] + rename_i w_R + cases z + · simp [defBot] at H + rw [EReal.mul_bot_of_pos] at H + apply le_bot_iff.mp at H + · rw [defBot] + have Hwr : some (some w_R) = Real.toEReal w_R := by simp [Real.toEReal] + rw [Hwr] at H + rw [<- EReal.coe_mul] at H + cases H + · apply Hr1 + rename_i z_nnr + cases z_nnr + · exact right_eq_inf.mp rfl + rename_i z_R + have Hwr : some (some w_R) = Real.toEReal w_R := by simp [Real.toEReal] + have Hzr : some (some z_R) = Real.toEReal z_R := by simp [Real.toEReal] + rw [Hwr, Hzr] at H + rw [Hwr, Hzr] + clear Hwr + clear Hzr + apply EReal.coe_le_coe_iff.mpr + repeat rw [<- EReal.coe_mul] at H + apply EReal.coe_le_coe_iff.mp at H + apply le_of_mul_le_mul_left H + exact EReal.coe_pos.mp Hr1 + + end ENNReal From db56434d3c114cc1841c6e9a6c0a847c7925a5d7 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Wed, 19 Jun 2024 11:46:11 -0400 Subject: [PATCH 065/122] Prove Da(P,P)=0 --- .../DifferentialPrivacy/RenyiDivergence.lean | 60 ++++++++++++++++--- SampCert/Util/Log.lean | 5 +- 2 files changed, 56 insertions(+), 9 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index d4e2d2ec..a6ad130b 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -483,7 +483,7 @@ Restriction of the PMF f to the support of q def reducedPMF_def (f q : PMF T) (x : { t : T // ¬q t = 0 }) : ENNReal := f x.val -- Should be provable -lemma acts_refl (q : PMF T) : AbsCts q q := by +lemma AbsCts_refl (q : PMF T) : AbsCts q q := by rw [AbsCts] simp @@ -533,7 +533,7 @@ theorem Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] [C rw [<- tsum_subtype_eq_of_support_subset K2] simp - have Hq : AbsCts q q := acts_refl q + have Hq : AbsCts q q := AbsCts_refl q have B1 (x : { x // ¬q x = 0 }) : p ↑x / q ↑x * q ↑x = reducedPMF Hac x / reducedPMF Hq x * reducedPMF Hq x := by congr have B2 (x : { x // ¬q x = 0 }) : (p ↑x / q ↑x)^α * q ↑x = (reducedPMF Hac x / reducedPMF Hq x)^α * reducedPMF Hq x := by congr @@ -576,8 +576,6 @@ lemma RenyiDivergence_mono_sum (x y : ℝ) (α : ℝ) (h : 1 < α) : (Real.exp ( · exact exp_le_exp.mp H · linarith -set_option pp.coercions false - theorem RenyiDivergence_def_nonneg [MeasurableSpace T] [MeasurableSingletonClass T] [Countable T] (p q : PMF T) (Hpq : AbsCts p q) {α : ℝ} (Hα : 1 < α) : (0 ≤ RenyiDivergence_def p q α) := by have H1 : eexp (((α - 1)) * 0) ≤ eexp ((α - 1) * RenyiDivergence_def p q α) := by @@ -602,7 +600,7 @@ theorem RenyiDivergence_def_nonneg [MeasurableSpace T] [MeasurableSingletonClass have Hone : (OfNat.ofNat 1 = Real.toEReal (1 : ℝ)) := by simp have Hzero : (OfNat.ofNat 0 = Real.toEReal (0 : ℝ)) := by simp - apply ereal_smul_l (α.toEReal - OfNat.ofNat 1) + apply ereal_smul_le_left (α.toEReal - OfNat.ofNat 1) · rw [Hone] rw [<- EReal.coe_sub] rw [Hzero] @@ -613,8 +611,56 @@ theorem RenyiDivergence_def_nonneg [MeasurableSpace T] [MeasurableSingletonClass exact EReal.coe_lt_top (α - OfNat.ofNat 1) · assumption -theorem RenyiDivergence_def_zero (p q : PMF T) (α : ℝ) : p = q <-> (0 = RenyiDivergence_def p q α) := by - -- See paper +set_option pp.coercions false + +lemma RenyiDivergence_refl_zero (p : PMF T) {α : ℝ} (Hα : 1 < α) : (0 = RenyiDivergence_def p p α) := by + have H1 : 1 = eexp ((α - 1) * RenyiDivergence_def p p α) := by + rw [RenyiDivergence_def_exp p p Hα] + rw [RenyiDivergenceExpectation p p Hα (AbsCts_refl p)] + have HRW (x : T) : ((p.val x) / (p.val x)) ^α * p.val x = p.val x := by + cases (Classical.em (p x = 0)) + · rename_i Hz + simp [DFunLike.coe] at Hz + simp [Hz] + · rename_i Hnz + rw [((@div_eq_one_iff (p.val x) (p.val x) ?GNZ) ?GNT).mpr rfl] + case GNZ => + simp [DFunLike.coe] at Hnz + simp + apply Hnz + case GNT => + have HltTop := PMF.apply_lt_top p x + apply LT.lt.ne_top HltTop + simp + conv => + rhs + arg 1 + intro x + simp [DFunLike.coe] + rw [HRW] + rcases p with ⟨ p' , Hp' ⟩ + exact Eq.symm (HasSum.tsum_eq Hp') + + have Hone : (OfNat.ofNat 1 = Real.toEReal (1 : ℝ)) := by simp + have Hzero : (OfNat.ofNat 0 = Real.toEReal (0 : ℝ)) := by simp + apply ereal_smul_eq_left (α.toEReal - OfNat.ofNat 1) + · rw [Hone] + rw [<- EReal.coe_sub] + rw [Hzero] + apply EReal.coe_lt_coe_iff.mpr + exact sub_pos.mpr Hα + · rw [Hone] + rw [<- EReal.coe_sub] + exact EReal.coe_lt_top (α - OfNat.ofNat 1) + apply eexp_injective + rw [<- H1] + simp + +lemma RenyiDivergence_zero_eq (p q : PMF T) {α : ℝ} (Hα : 1 < α) (Hpq : AbsCts p q) (HR : 0 = RenyiDivergence_def p q α) : p = q:= by + have HR' : 1 = eexp (((α - 1)) * RenyiDivergence_def p q α) := by simp [<- HR] + rw [RenyiDivergence_def_exp p q Hα] at HR' + rw [RenyiDivergenceExpectation p q Hα Hpq] at HR' + -- Is this true? sorry lemma RenyiDivergence_def_log_sum_nonneg (p q : PMF T) (α : ℝ) : (0 ≤ (elog (∑' x : T, (p x)^α * (q x)^(1 - α)))) := by diff --git a/SampCert/Util/Log.lean b/SampCert/Util/Log.lean index 82851234..8634455c 100644 --- a/SampCert/Util/Log.lean +++ b/SampCert/Util/Log.lean @@ -252,9 +252,8 @@ lemma mul_mul_inv_eq_mul_cancel (H : y = 0 -> x = 0) (H2 : ¬(x ≠ 0 ∧ y = rw [mul_right_comm] rw [mul_inv_cancel_right₀ Hy' x'] - -- Could be shortened -lemma ereal_smul_l (s : EReal) (Hr1 : 0 < s) (Hr2 : s < ⊤) (H : s * w ≤ s * z) : w ≤ z := by +lemma ereal_smul_le_left (s : EReal) (Hr1 : 0 < s) (Hr2 : s < ⊤) (H : s * w ≤ s * z) : w ≤ z := by have defTop : some none = (⊤ : EReal) := by simp [Top.top] have defBot : none = (⊥ : EReal) := by simp [Bot.bot] @@ -327,5 +326,7 @@ lemma ereal_smul_l (s : EReal) (Hr1 : 0 < s) (Hr2 : s < ⊤) (H : s * w ≤ s * apply le_of_mul_le_mul_left H exact EReal.coe_pos.mp Hr1 +lemma ereal_smul_eq_left (s : EReal) (Hr1 : 0 < s) (Hr2 : s < ⊤) (H : s * w = s * z) : w = z := by + sorry end ENNReal From fc17bca715917d00530ca0b4fb5dabbb536f8412 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Wed, 19 Jun 2024 13:34:39 -0400 Subject: [PATCH 066/122] Gaussian gen PMF definition --- .../ZeroConcentrated/ConcentratedBound.lean | 73 ++++++++++--------- SampCert/Samplers/GaussianGen/Properties.lean | 23 +++++- 2 files changed, 59 insertions(+), 37 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean index 33962dba..edd7c5a1 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean @@ -85,7 +85,6 @@ lemma SG_Renyi_simplify {σ : ℝ} (h : σ ≠ 0) (μ ν : ℤ) (α : ℝ) : /-- Alternative definition for the Renyi Divergence. -FIXME: is there any reason to not get rid of this? -/ noncomputable def RenyiDivergence' (p q : T → ℝ) (α : ℝ) : ℝ := (1 / (α - 1)) * Real.log (∑' x : T, (p x)^α * (q x)^(1 - α)) @@ -533,39 +532,41 @@ theorem Renyi_Gauss_divergence_bound' {σ α : ℝ} (h : σ ≠ 0) (h' : 1 < α) RenyiDivergence (discrete_gaussian_pmf h μ) (discrete_gaussian_pmf h ν) α ≤ (ENNReal.ofReal α) * (ENNReal.ofReal ((((μ - ν) : ℤ)^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 + have A : RenyiDivergence (discrete_gaussian_pmf h μ) (discrete_gaussian_pmf h ν) α = + ENNReal.ofReal (RenyiDivergence' (fun (x : ℤ) => discrete_gaussian σ μ x) (fun (x : ℤ) => discrete_gaussian σ ν x) α) := by unfold RenyiDivergence + unfold RenyiDivergence_def 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 Renyi_sum_SG_nonneg h - . apply Renyi_sum_SG_nonneg h - . apply Renyi_Gauss_summable h + unfold discrete_gaussian_pmf + sorry + -- 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 Renyi_sum_SG_nonneg h + -- . apply Renyi_sum_SG_nonneg h + -- . apply Renyi_Gauss_summable h rw [A] - apply Renyi_divergence_bound_pre h h' - -/ - sorry + rw [<- ENNReal.ofReal_mul] + apply ENNReal.ofReal_le_ofReal + apply Renyi_divergence_bound_pre h h' μ ν + linarith namespace SLang @@ -575,19 +576,19 @@ Upper bound on Renyi divergence between outputs of the ``SLang`` discrete Gaussi theorem discrete_GaussianGenSample_ZeroConcentrated {α : ℝ} (h : 1 < α) (num : PNat) (den : PNat) (μ ν : ℤ) : RenyiDivergence ((DiscreteGaussianGenPMF num den μ)) (DiscreteGaussianGenPMF num den ν) α ≤ (ENNReal.ofReal α) * (ENNReal.ofReal (((μ - ν) : ℤ)^2 : ℝ) / (((2 : ENNReal) * ((num : ENNReal) / (den : ENNReal))^2 : ENNReal))) := 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] + . rw [DiscreteGaussianGenPMF] + -- rw [DiscreteGaussianGenSample_apply] + . rw [DiscreteGaussianGenPMF] + -- rw [DiscreteGaussianGenSample_apply] . skip - apply Renyi_Gauss_divergence_bound' A h - -/ + sorry + -- apply Renyi_Gauss_divergence_bound' A h + -- sorry end SLang diff --git a/SampCert/Samplers/GaussianGen/Properties.lean b/SampCert/Samplers/GaussianGen/Properties.lean index 50e7e180..88ea7912 100644 --- a/SampCert/Samplers/GaussianGen/Properties.lean +++ b/SampCert/Samplers/GaussianGen/Properties.lean @@ -53,10 +53,31 @@ theorem DiscreteGaussianGenSample_apply (num : PNat) (den : PNat) (μ x : ℤ) : . simp [gauss_term_ℝ] . rw [shifted_gauss_sum_0 A] +theorem DiscreteGaussianGen_sum (num : PNat) (den : PNat) (μ : ℤ) : HasSum (DiscreteGaussianGenSample num den μ) 1 := by + rw [Summable.hasSum_iff ENNReal.summable] + conv => + lhs + arg 1 + intro b + rw [DiscreteGaussianGenSample_apply] + + have Hnz : (num / den : ℝ) ≠ 0 := by simp + have Hcoe : (OfNat.ofNat 1 = ENNReal.ofReal (OfNat.ofNat 1)) := by simp + rw [Hcoe] + rw [<- (discrete_gaussian_normalizes Hnz μ)] + symm + apply (ENNReal.ofReal_tsum_of_nonneg ?hf_nonneg ?hf) + · exact fun n => discrete_gaussian_nonneg Hnz (↑μ) n + · apply discrete_gaussian_summable' + apply Hnz /-- ``SLang`` program is a proper distribution. -/ -def DiscreteGaussianGenPMF (num : PNat) (den : PNat) (μ : ℤ) : PMF ℤ := ⟨ DiscreteGaussianGenSample num den μ , sorry ⟩ +def DiscreteGaussianGenPMF (num : PNat) (den : PNat) (μ : ℤ) : PMF ℤ := + ⟨ DiscreteGaussianGenSample num den μ , DiscreteGaussianGen_sum num den μ ⟩ + + + end SLang From 307568c895ec5edef9c24e532fbfad5bb89ca9f7 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Wed, 19 Jun 2024 13:40:49 -0400 Subject: [PATCH 067/122] discrete gaussian PMF --- SampCert/Util/Gaussian/DiscreteGaussian.lean | 13 ++++++++----- 1 file changed, 8 insertions(+), 5 deletions(-) diff --git a/SampCert/Util/Gaussian/DiscreteGaussian.lean b/SampCert/Util/Gaussian/DiscreteGaussian.lean index 77cc1e7e..17c1eff1 100644 --- a/SampCert/Util/Gaussian/DiscreteGaussian.lean +++ b/SampCert/Util/Gaussian/DiscreteGaussian.lean @@ -338,11 +338,14 @@ theorem discrete_gaussian_normalizes {σ : ℝ} (h : σ ≠ 0) (μ : ℝ) : def discrete_gaussian_normal {σ : ℝ} (h : σ ≠ 0) (μ : ℝ) : HasSum (fun z : ℤ => ENNReal.ofReal (discrete_gaussian σ μ z)) 1 := by - let HR := discrete_gaussian_normalizes h μ - have HS := Summable.hasSum (discrete_gaussian_summable' h μ) - simp at HR - rw [HR] at HS - sorry + rw [Summable.hasSum_iff ENNReal.summable] + rw [<- ENNReal.ofReal_tsum_of_nonneg] + case hf_nonneg => + exact fun n => discrete_gaussian_nonneg h μ n + case hf => + apply discrete_gaussian_summable' + apply h + simp [discrete_gaussian_normalizes h μ] /-- The discrete Gaussian as a PMF when evaluated over ℤ From 40869a9a018bb5b500544ca78a6ac6854e67c03f Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Wed, 19 Jun 2024 14:43:50 -0400 Subject: [PATCH 068/122] sketch Renyi_Gauss_divergence_bound' (it is fixable) --- .../ZeroConcentrated/ConcentratedBound.lean | 69 +++++++++++++------ SampCert/Util/Log.lean | 15 ++++ 2 files changed, 62 insertions(+), 22 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean index edd7c5a1..6fce6e4c 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean @@ -525,6 +525,9 @@ theorem Renyi_Gauss_summable {σ : ℝ} (h : σ ≠ 0) (μ ν : ℤ) (α : ℝ) apply Summable.mul_right apply summable_gauss_term' h + +set_option pp.coercions false +-- set_option pp.notation false /-- Upper bound on Renyi divergence between discrete Gaussians. -/ @@ -532,6 +535,7 @@ theorem Renyi_Gauss_divergence_bound' {σ α : ℝ} (h : σ ≠ 0) (h' : 1 < α) RenyiDivergence (discrete_gaussian_pmf h μ) (discrete_gaussian_pmf h ν) α ≤ (ENNReal.ofReal α) * (ENNReal.ofReal ((((μ - ν) : ℤ)^2 : ℝ) / (2 * σ^2))) := by + -- Fixable have A : RenyiDivergence (discrete_gaussian_pmf h μ) (discrete_gaussian_pmf h ν) α = ENNReal.ofReal (RenyiDivergence' (fun (x : ℤ) => discrete_gaussian σ μ x) (fun (x : ℤ) => discrete_gaussian σ ν x) α) := by unfold RenyiDivergence @@ -540,28 +544,49 @@ theorem Renyi_Gauss_divergence_bound' {σ α : ℝ} (h : σ ≠ 0) (h' : 1 < α) congr simp unfold discrete_gaussian_pmf - sorry - -- 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 Renyi_sum_SG_nonneg h - -- . apply Renyi_sum_SG_nonneg h - -- . apply Renyi_Gauss_summable h + + have Hdg_pos (x : ℤ) (w : ℝ) : OfNat.ofNat 0 < discrete_gaussian σ w x.cast := by sorry + have Hdg_pow_pos (x : ℤ) w : OfNat.ofNat 0 ≤ discrete_gaussian σ w x.cast ^ α := by sorry + + conv => + lhs + arg 1 + arg 2 + arg 1 + arg 1 + intro x + simp [DFunLike.coe] + rw [ENNReal.ofReal_rpow_of_pos (Hdg_pos x μ.cast)] + rw [ENNReal.ofReal_rpow_of_pos (Hdg_pos x ν.cast)] + rw [<- ENNReal.ofReal_mul (Hdg_pow_pos x μ.cast)] + rw [<- ENNReal.ofEReal_ofReal_toENNReal] + simp + congr + -- Here is where I would do the classical I talk about below + cases (Classical.em (0 = ∑' (x : ℤ), discrete_gaussian σ μ.cast x.cast ^ α * discrete_gaussian σ ν.cast x.cast ^ (1 - α))) + · rename_i Hzero + rw [<- Hzero] + simp + rw [<- ENNReal.ofReal_tsum_of_nonneg] + · rw [<- Hzero] + simp + -- Sadly this is not true (ultimately because they set Real.log 0 = 0) + -- Fixable: This series should only ever be zero when μ = ν, if the paper is right. + -- We also should prove that μ ≠ ν <-> (discretVe_gaussian_pmf h μ) ≠ (discretVe_gaussian_pmf h ν) + -- In the outermost proof, that case is trivial. + sorry + · exact fun n => Renyi_sum_SG_nonneg h μ ν n + · exact Renyi_Gauss_summable h μ ν α + · rename_i Hnz + rw [<- ENNReal.elog_ENNReal_ofReal_of_pos] + · rw [ENNReal.ofReal_tsum_of_nonneg ?Hnn] + · exact Renyi_Gauss_summable h μ ν α + · intro n + exact Renyi_sum_SG_nonneg h μ ν n + apply lt_of_le_of_ne + · sorry + · sorry + -- I could to add a classical case to finish this, assuming the rw [A] rw [<- ENNReal.ofReal_mul] apply ENNReal.ofReal_le_ofReal diff --git a/SampCert/Util/Log.lean b/SampCert/Util/Log.lean index 8634455c..73f2535d 100644 --- a/SampCert/Util/Log.lean +++ b/SampCert/Util/Log.lean @@ -208,6 +208,21 @@ lemma toEReal_ofENNReal_nonneg (H : 0 ≤ w) : ENNReal.toEReal (ofEReal w) = w : @[simp] lemma ofEReal_toENNReal : ofEReal (ENNReal.toEReal x) = x := by sorry +lemma ofEReal_ofReal_toENNReal : ENNReal.ofEReal (Real.toEReal r) = ENNReal.ofReal r := by + simp [ofEReal, Real.toEReal, ENNReal.ofReal] + +set_option pp.coercions false + +lemma elog_ENNReal_ofReal_of_pos (x : ℝ) (H : 0 < x): (ENNReal.ofReal x).elog = x.log.toEReal := by + simp [ENNReal.ofReal, ENNReal.elog, ENNReal.toEReal] + rw [ite_eq_iff'] + apply And.intro + · intro + exfalso + linarith + · intro H + simp at H + rw [max_eq_left_of_lt H] -- ENNReal inequalities From 633953fef75828d491209021eba4305efd423823 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Wed, 19 Jun 2024 15:17:57 -0400 Subject: [PATCH 069/122] discrete_GaussianGenSample_ZeroConcentrated --- .../ZeroConcentrated/ConcentratedBound.lean | 48 +++++++++++++++---- 1 file changed, 38 insertions(+), 10 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean index 6fce6e4c..8ad13078 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean @@ -526,8 +526,6 @@ theorem Renyi_Gauss_summable {σ : ℝ} (h : σ ≠ 0) (μ ν : ℤ) (α : ℝ) apply summable_gauss_term' h -set_option pp.coercions false --- set_option pp.notation false /-- Upper bound on Renyi divergence between discrete Gaussians. -/ @@ -603,17 +601,47 @@ theorem discrete_GaussianGenSample_ZeroConcentrated {α : ℝ} (h : 1 < α) (num (ENNReal.ofReal α) * (ENNReal.ofReal (((μ - ν) : ℤ)^2 : ℝ) / (((2 : ENNReal) * ((num : ENNReal) / (den : ENNReal))^2 : ENNReal))) := 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] + have Hpmf (w : ℤ) : (discrete_gaussian_pmf A w = DiscreteGaussianGenPMF num den w) := by + simp [discrete_gaussian_pmf] + simp [DiscreteGaussianGenPMF] + congr + apply funext + intro z + rw [← DiscreteGaussianGenSample_apply] conv => left congr - . rw [DiscreteGaussianGenPMF] - -- rw [DiscreteGaussianGenSample_apply] - . rw [DiscreteGaussianGenPMF] - -- rw [DiscreteGaussianGenSample_apply] + . rw [<- Hpmf] + . rw [<- Hpmf] . skip - - sorry - -- apply Renyi_Gauss_divergence_bound' A h - -- sorry + apply le_trans + · apply (Renyi_Gauss_divergence_bound') + apply h + · have X : (OfNat.ofNat 2 * (num.val.cast / den.val.cast) ^ OfNat.ofNat 2) = ENNReal.ofReal (OfNat.ofNat 2 * (num.val.cast / den.val.cast) ^ OfNat.ofNat 2) := by + simp + rw [ENNReal.ofReal_mul ?Hnn] + case Hnn => simp + congr + · simp + · rw [ENNReal.ofReal_div_of_pos ?Hpos] + case Hpos => simp + repeat rw [sq] + rw [ENNReal.ofReal_mul ?Hnn1] + case Hnn1 => simp + rw [ENNReal.ofReal_mul ?Hnn2] + case Hnn2 => simp + simp + repeat rw [ENNReal.div_eq_inv_mul] + rw [mul_mul_mul_comm] + congr + rw [ENNReal.mul_inv] + · right + simp + · left + simp + rw [X] + rw [ENNReal.ofReal_div_of_pos] + congr + simp end SLang From 749ffd9657dd78d4f6bdee520d225bb388ef021d Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Wed, 19 Jun 2024 16:20:16 -0400 Subject: [PATCH 070/122] reduction of privNoisedQuery_zCDPBound to the Real-valued case --- .../Mechanism/Properties.lean | 140 ++++++++++++++---- 1 file changed, 113 insertions(+), 27 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean index de8055a2..005bff76 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean @@ -21,24 +21,106 @@ namespace SLang lemma privNoisedQuery_norm (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_sensitivity : sensitivity query Δ) : NormalMechanism (privNoisedQuery query Δ ε₁ ε₂) := sorry +set_option pp.coercions false + /-- The zCDP mechanism with bounded sensitivity satisfies the bound for ``(Δε₂/ε₁)^2``-zCDP. -/ theorem privNoisedQuery_zCDPBound (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_sensitivity : sensitivity query Δ) : zCDPBound (privNoisedQuery query Δ ε₁ ε₂) (privNoisedQuery_norm query Δ ε₁ ε₂ bounded_sensitivity) ((ε₁ : ℝ) / ε₂) := by - /- simp [zCDPBound, privNoisedQuery] intros α h1 l₁ l₂ h2 have A := @discrete_GaussianGenSample_ZeroConcentrated α h1 (Δ * ε₂) ε₁ (query l₁) (query l₂) apply le_trans A clear A + + -- Pull out the ENNReal.ofReal, reducing it to a Real case + -- Simplify this argument, it is a mess + rw [ENNReal.div_eq_inv_mul] + have L : (OfNat.ofNat 2 * ((Δ * ε₂).val.cast / ε₁.val.cast) ^ OfNat.ofNat 2)⁻¹ = ENNReal.ofReal ((OfNat.ofNat 2 * ((Δ * ε₂).val.cast / ε₁.val.cast) ^ OfNat.ofNat 2)⁻¹ ) := by + simp + rw [← div_pow] + conv => + rhs + arg 1 + rw [mul_comm] + rw [ENNReal.mul_inv] + · rw [ENNReal.ofReal_mul ?G1] + case G1 => + simp + congr + · rw [ENNReal.ofReal_inv_of_pos] + congr + simp + exact zero_lt_two + · rw [ENNReal.inv_pow] + rw [ENNReal.ofReal_pow ?G] + case G => + apply div_nonneg + · simp + · simp + congr + repeat rewrite [ENNReal.div_eq_inv_mul] + rw [ENNReal.mul_inv ?G1 ?G2] + case G1 => + left + simp + case G2 => + left + simp + simp + rw [div_eq_mul_inv] + rw [ENNReal.ofReal_mul] + · simp + congr + rw [ENNReal.mul_inv] + · rw [ENNReal.ofReal_mul] + · rw [mul_comm] + congr + · rw [ENNReal.ofReal_inv_of_pos] + · congr + simp + · simp + · rw [ENNReal.ofReal_inv_of_pos] + · congr + simp + · simp + · simp + · left ; simp + · left ; simp + · simp + · left ; simp + · left ; simp + rw [L] + have H1 : OfNat.ofNat 0 ≤ ((OfNat.ofNat (OfNat.ofNat 2) * ((Δ * ε₂).val.cast / ε₁.val.cast) ^ OfNat.ofNat (OfNat.ofNat 2))⁻¹ : ℝ) := by + simp + apply div_nonneg + · simp + · simp + have H2 : OfNat.ofNat 0 ≤ α := by linarith + conv => + lhs + arg 2 + rw [<- ENNReal.ofReal_mul] + · skip + . apply H1 + conv => + lhs + rw [<- ENNReal.ofReal_mul] + · skip + · apply H2 + clear H1 + clear H2 + apply ofReal_le_ofReal_iff'.mpr + + replace bounded_sensitivity := bounded_sensitivity l₁ l₂ h2 ring_nf simp conv => left left - right + left rw [mul_pow] conv => left @@ -46,38 +128,42 @@ theorem privNoisedQuery_zCDPBound (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ right rw [mul_comm] rw [← mul_assoc] - conv => - left - rw [mul_assoc] - right - rw [← mul_assoc] - left - rw [mul_comm] - rw [← mul_assoc] + + -- Almost would be easier just to redo it... from here, the main reduction has already been finished + + -- conv => + -- left + -- rw [mul_assoc] + -- right + -- rw [← mul_assoc] + -- left + -- rw [mul_comm] rw [← mul_assoc] rw [← mul_assoc] - simp only [inv_pow] + -- 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 α (zero_le_one' ℝ) (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 - -/ - sorry + sorry + -- apply mul_le_mul + -- 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 + . sorry + -- rw [pow_two] + -- rw [_root_.mul_pos_iff] + -- left + -- simp /-- All outputs of the zCDP mechanism have nonzero probability. From 99c1745604f096adc445b0d55203b156406de4ea Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Wed, 19 Jun 2024 16:23:10 -0400 Subject: [PATCH 071/122] privNoisedQuery_zCDPBound normalizes --- .../ZeroConcentrated/Mechanism/Properties.lean | 6 +++++- 1 file changed, 5 insertions(+), 1 deletion(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean index 005bff76..463ee6da 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean @@ -19,7 +19,11 @@ open Classical Nat Int Real ENNReal MeasureTheory Measure namespace SLang lemma privNoisedQuery_norm (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_sensitivity : sensitivity query Δ) : - NormalMechanism (privNoisedQuery query Δ ε₁ ε₂) := sorry + NormalMechanism (privNoisedQuery query Δ ε₁ ε₂) := by + rw [NormalMechanism] + intro l + rw [privNoisedQuery] + exact DiscreteGaussianGen_sum (Δ * ε₂) ε₁ (query l) set_option pp.coercions false From 9f8e244c854c454c2d4b6f8e13bbad2324d4fcb9 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Thu, 20 Jun 2024 12:13:29 -0400 Subject: [PATCH 072/122] real-valued Jensen converse --- .../DifferentialPrivacy/RenyiDivergence.lean | 154 ++++++++++++++++++ .../ZeroConcentrated/Composition.lean | 2 +- .../ZeroConcentrated/Postprocessing.lean | 7 +- 3 files changed, 160 insertions(+), 3 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index a6ad130b..6698bfe4 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -153,6 +153,7 @@ section Jensen variable {T : Type} variable [t1 : MeasurableSpace T] variable [t2 : MeasurableSingletonClass T] +variable [tcount : Countable T] -- New variable {U V : Type} variable [m2 : MeasurableSpace U] @@ -258,6 +259,146 @@ theorem Renyi_Jensen_real (f : T → ℝ) (q : PMF T) (α : ℝ) (h : 1 < α) (h rw [one_le_ofReal] apply le_of_lt h +-- set_option pp.notation false +-- set_option pp.coercions false +-- set_option pp.all true +-- set_option pp.universes false + +/-- +Jensen's strict ineuquality for the exponential applied to Renyi's sum +-/ +theorem Renyi_Jensen_strict_real (f : T → ℝ) (q : PMF T) (α : ℝ) (h : 1 < α) (h2 : ∀ x : T, 0 ≤ f x) (mem : Memℒp f (ENNReal.ofReal α) (PMF.toMeasure q)) (HT_nz : ∀ t : T, q t ≠ 0): + ((∑' x : T, (f x) * (q x).toReal)) ^ α < (∑' x : T, (f x) ^ α * (q x).toReal) ∨ (∀ x : T, f x = ∑' (x : T), (q x).toReal * f x) := by + conv => + left + left + left + right + intro x + rw [mul_comm] + rw [← smul_eq_mul] + conv => + left + right + right + intro x + rw [mul_comm] + rw [← smul_eq_mul] + rw [← PMF.integral_eq_tsum] + rw [← PMF.integral_eq_tsum] + + have A := strictConvexOn_rpow 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 := @StrictConvexOn.ae_eq_const_or_map_average_lt T ℝ t1 _ _ _ (PMF.toMeasure q) (Set.Ici 0) f (fun (x : ℝ) => x ^ α) ((PMF.toMeasure.isProbabilityMeasure q).toIsFiniteMeasure) A B C ?G1 ?G2 ?G3 + case G1 => + exact MeasureTheory.ae_of_all (PMF.toMeasure q) h2 + case G2 => + apply MeasureTheory.Memℒp.integrable _ mem + rw [one_le_ofReal] + apply le_of_lt h + case G3 => + 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 t1 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 + simp at D + · cases D + · rename_i HR + right + simp at HR + -- Because T is discrete, almost-everywhere equality should become equality + have HR' := @Filter.EventuallyEq.eventually _ _ (q.toMeasure.ae) f (Function.const T (⨍ (x : T), f x ∂q.toMeasure)) HR + simp [Filter.Eventually] at HR' + -- The measure of the compliment of the set in HR' is zero + simp [MeasureTheory.Measure.ae] at HR' + rw [PMF.toMeasure_apply _ _ ?Hmeas] at HR' + case Hmeas => + apply (@measurableSet_discrete _ _ ?DM) + apply MeasurableSingletonClass.toDiscreteMeasurableSpace + -- Sum is zero iff all elements are zero + apply ENNReal.tsum_eq_zero.mp at HR' + -- Indicator is zero when proposition is not true + intro x + have HR' := HR' x + simp at HR' + cases (Classical.em (f x = ⨍ (x : T), f x ∂q.toMeasure)) + · rename_i Heqx + -- Rewrite the average + rw [MeasureTheory.average] at Heqx + rw [MeasureTheory.integral_countable'] at Heqx + · simp at Heqx + conv at Heqx => + rhs + arg 1 + intro x + rw [PMF.toMeasure_apply_singleton] + · skip + · apply measurableSet_singleton + -- Interesting.... is this sum not just 1? + simp at * + apply Heqx + · simp + apply MeasureTheory.Memℒp.integrable _ mem + sorry -- true + · -- At type T, q x is never zero + rename_i Hnex + exfalso + apply (HT_nz x) + apply HR' + apply Hnex + · rename_i HR + left + rw [<- MeasureTheory.integral_average] + rw [<- MeasureTheory.integral_average] + simp + rw [<- MeasureTheory.integral_average] + rw [<- MeasureTheory.integral_average] + simp + apply HR + . 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 t1 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 + + + + end Jensen -- MARKUSDE: move @@ -656,10 +797,23 @@ lemma RenyiDivergence_refl_zero (p : PMF T) {α : ℝ} (Hα : 1 < α) : (0 = Ren rw [<- H1] simp + +-- lemma Jensens_equality_converse_ENNReal (f g h : ENNReal -> ENNReal) (H : h (∑'(x : T), f x * g x) ≤ ∑'(x : T), h (f x) * g x) : +-- (∃ a b : ENNReal, ∀ x : ENNReal, h x = a * x + b) := by +-- sorry + + + + + + lemma RenyiDivergence_zero_eq (p q : PMF T) {α : ℝ} (Hα : 1 < α) (Hpq : AbsCts p q) (HR : 0 = RenyiDivergence_def p q α) : p = q:= by have HR' : 1 = eexp (((α - 1)) * RenyiDivergence_def p q α) := by simp [<- HR] rw [RenyiDivergence_def_exp p q Hα] at HR' rw [RenyiDivergenceExpectation p q Hα Hpq] at HR' + have Hlhs : (∑' (x : T), (p x / q x) * q x)^α = 1:= sorry + rw [<- Hlhs] at HR' + -- Is this true? sorry diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean index c9192827..9111ae8c 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean @@ -79,7 +79,7 @@ theorem Renyi_noised_query_NZ {nq : List T → SLang U} {HNorm : NormalMechanism Composed queries satisfy zCDP Renyi divergence bound. -/ theorem privCompose_zCDPBound {nq1 : List T → SLang U} {nq2 : List T → SLang V} {HNorm1 : ∀ l, HasSum (nq1 l) 1} {HNorm2 : ∀ l, HasSum (nq2 l) 1} {ε₁ ε₂ ε₃ ε₄ : ℕ+} - (h1 : zCDPBound nq1 HNorm1 ((ε₁ : ℝ) / ε₂)) (h2 : zCDPBound nq2 HNorm2 ((ε₃ : ℝ) / ε₄)) (nn1 : NonZeroNQ nq1) (nn2 : NonZeroNQ nq2) (nt1 : NonTopRDNQ nq1) (nt2 : NonTopRDNQ nq2) (nts1 : NonTopNQ nq1) (nts2 : NonTopNQ nq2) : + (h1 : zCDPBound nq1 HNorm1 ((ε₁ : ℝ) / ε₂)) (h2 : zCDPBound nq2 HNorm2 ((ε₃ : ℝ) / ε₄)) : -- (nn1 : NonZeroNQ nq1) (nn2 : NonZeroNQ nq2) (nt1 : NonTopRDNQ nq1) (nt2 : NonTopRDNQ nq2) (nts1 : NonTopNQ nq1) (nts2 : NonTopNQ nq2) : zCDPBound (privCompose nq1 nq2) (privCompose_norm nq1 nq2 HNorm1 HNorm2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by /- simp [privCompose, RenyiDivergence, zCDPBound] diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean index 4422930d..d5aad4e8 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean @@ -559,7 +559,6 @@ theorem privPostProcess_zCDPBound {nq : List T → SLang U} {HNorm : NormalMecha (h : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (conv : NonTopSum nq) (f : U → V) : zCDPBound (privPostProcess nq f) (privPostProcess_norm nq HNorm f) ((ε₁ : ℝ) / ε₂) := by - /- simp [privPostProcess, zCDPBound, RenyiDivergence] intro α h1 l₁ l₂ h2 have h' := h @@ -571,6 +570,10 @@ theorem privPostProcess_zCDPBound {nq : List T → SLang U} {HNorm : NormalMecha apply le_trans _ h' clear h' + + sorry + + /- -- remove the α scaling have A : 0 ≤ (α - 1)⁻¹ := by simp @@ -645,8 +648,8 @@ theorem privPostProcess_zCDPBound {nq : List T → SLang U} {HNorm : NormalMecha apply lt_of_le_of_lt X Y rw [lt_top_iff_ne_top] at Z contradiction - -/ sorry + -/ theorem privPostProcess_NonTopRDNQ {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) 1} {ε₁ ε₂ : ℕ+} (f : U → V) (dp : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) (nt : NonTopRDNQ nq) (nz : NonZeroNQ nq) (nts : NonTopNQ nq) (ntsum: NonTopSum nq) : From 6618cf27c7d278a36fbaa19853593de78b8ddfdb Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Thu, 20 Jun 2024 14:31:38 -0400 Subject: [PATCH 073/122] Renyi_Jensen_ENNReal_converse_reduct --- .../DifferentialPrivacy/RenyiDivergence.lean | 253 +++++++++++++++++- SampCert/Util/Log.lean | 3 + 2 files changed, 250 insertions(+), 6 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index 6698bfe4..33e3c332 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -366,7 +366,11 @@ theorem Renyi_Jensen_strict_real (f : T → ℝ) (q : PMF T) (α : ℝ) (h : 1 < apply Heqx · simp apply MeasureTheory.Memℒp.integrable _ mem - sorry -- true + have X : (1 : ENNReal) = ENNReal.ofReal (1 : ℝ) := by simp + rw [X] + apply ofReal_le_ofReal_iff'.mpr + left + linarith · -- At type T, q x is never zero rename_i Hnex exfalso @@ -473,7 +477,6 @@ lemma Renyi_Jensen_ENNReal_reduct [MeasurableSpace T] [MeasurableSingletonClass intro x rw [<- ENNReal.toReal_mul] apply ENNReal.summable_toReal - assumption have HRJf_nonneg (a : T) : 0 <= Renyi_Jensen_f p q a := by apply toReal_nonneg have HRJf_nt (a : T) : p a / q a ≠ ⊤ := by @@ -618,6 +621,248 @@ lemma Renyi_Jensen_ENNReal_reduct [MeasurableSpace T] [MeasurableSingletonClass exact OrderTop.le_top ((∑' (x : T), p x / q x * q x) ^ α) +-- set_option pp.all true + +-- Lift the properties of Renyi_Jensen_strict_real to ENNReal +-- q is still nonnegative +lemma Renyi_Jensen_ENNReal_converse_reduct [MeasurableSpace T] [MeasurableSingletonClass T] [Countable T] + (p q : PMF T) {α : ℝ} (h : 1 < α) (H : AbsCts p q) (Hq : ∀ t, q t ≠ 0) + (Hsumeq : (∑' x : T, (p x / q x) * q x) ^ α = (∑' x : T, (p x / q x) ^ α * q x)) : + (p = q) := by + have Hdiscr : DiscreteMeasurableSpace T := MeasurableSingletonClass.toDiscreteMeasurableSpace + cases (Classical.em (∑' (a : T), (p a / q a) ^ α * q a ≠ ⊤)) + · -- Preliminary stuff, basically the same as the forward case + rename_i Hnts + cases (Classical.em (∀ x : T, ¬(p x = ⊤ ∧ q x ≠ 0 ∧ q x ≠ ⊤))) + · rename_i Hspecial + conv at Hsumeq => + rhs + arg 1 + intro x + rw [Renyi_Jensen_rw p q h H Hspecial] + rw [<- ENNReal.ofReal_tsum_of_nonneg ?Hnonneg ?Hsummable] at Hsumeq + case Hnonneg => + intro t + apply mul_nonneg + · refine rpow_nonneg ?ha.hx α + simp [Renyi_Jensen_f] + · exact toReal_nonneg + case Hsummable => + conv => + congr + intro x + rw [Renyi_Jensen_f] + conv => + arg 1 + intro x + lhs + rw [ENNReal.toReal_rpow] + conv => + arg 1 + intro x + rw [<- ENNReal.toReal_mul] + apply ENNReal.summable_toReal + assumption + have HRJf_nonneg (a : T) : 0 <= Renyi_Jensen_f p q a := by apply toReal_nonneg + have HRJf_nt (a : T) : p a / q a ≠ ⊤ := by + intro HK + have HK' : (p a ≠ 0 ∧ q a = 0 ∨ p a = ⊤ ∧ q a ≠ ⊤) := by exact div_eq_top.mp HK + cases HK' + · rename_i HK' + rcases HK' with ⟨ HK1 , HK2 ⟩ + rw [AbsCts] at H + simp_all only [ne_eq, not_and, Decidable.not_not, ENNReal.zero_div, zero_ne_top] + · rename_i HK' + rcases HK' with ⟨ HK1 , _ ⟩ + apply (Hspecial a) + simp_all + have Hsum_indicator (a : T) : ∑' (i : T), q i * Set.indicator {a} (fun x => 1) i = q a := by + have Hfun : (fun (i : T) => q i * Set.indicator {a} (fun x => 1) i) = (fun (i : T) => if i = a then q a else 0) := by + funext i + rw [Set.indicator] + split <;> simp <;> split <;> simp_all + rw [Hfun] + exact tsum_ite_eq a (q a) + + -- Apply the converse lemma + have Hieq := Renyi_Jensen_strict_real (Renyi_Jensen_f p q) q α h HRJf_nonneg ?GLp Hq + case GLp => + -- ℒp bound (same as forward proof) + 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 + . apply measurable_discrete + . apply Measurable.inv + apply measurable_discrete + · simp [snorm] + split + · simp + · rename_i Hα + simp [snorm'] + rw [MeasureTheory.lintegral_countable'] + rw [toReal_ofReal (le_of_lt (lt_trans zero_lt_one h))] + apply rpow_lt_top_of_nonneg + · simp + apply le_of_not_ge Hα + · conv => + lhs + arg 1 + intro a + arg 1 + arg 1 + rw [<- Real.toNNReal_eq_nnnorm_of_nonneg (HRJf_nonneg a)] + rw [Renyi_Jensen_f] + rw [<- ENNReal.ofReal.eq_1] + rw [ENNReal.ofReal_toReal (HRJf_nt a)] + rfl + conv => + lhs + arg 1 + intro a + rhs + simp [toMeasure] + simp [PMF.toOuterMeasure] + rw [Hsum_indicator] + apply Hnts + cases Hieq + · rename_i Hk + exfalso + have CG1 (z : T) : q z = 0 → p z = 0 := by apply H + have CG2 (z : T) : ¬(p z ≠ 0 ∧ q z = ⊤) := by + simp + intro + apply PMF.apply_ne_top + conv at Hk => + lhs + arg 1 + arg 1 + intro z + rw [Renyi_Jensen_f] + rw [<- ENNReal.toReal_mul] + arg 1 + rw [division_def] + rw [mul_mul_inv_eq_mul_cancel (CG1 z) (CG2 z)] + clear CG1 + clear CG2 + + -- Convert the LHS of Hsumeq to the ℝ-valued summand, and then contradict + conv at Hsumeq => + lhs + arg 1 + arg 1 + intro x + rw [division_def] + rw [mul_assoc] + rw [ENNReal.inv_mul_cancel] + · skip + · apply Hq + · apply PMF.apply_ne_top + simp at * + rw [<- ENNReal.tsum_toReal_eq ?G1] at Hk + case G1 => + intro + apply PMF.apply_ne_top + simp at * + have Hone' : (1 : ENNReal).toReal = (1 : ℝ) := by simp + rw [<- Hone'] at Hk + rw [Hsumeq] at Hk + clear Hone' + rw [ENNReal.toReal_ofReal ?G1] at Hk + case G1 => + apply tsum_nonneg + intro i + apply mul_nonneg + · apply rpow_nonneg + apply HRJf_nonneg + · exact toReal_nonneg + linarith + · rename_i Hext + -- RHS of Hext is 1, LHS is p/q + apply PMF.ext + intro x + have Hext' := Hext x + rw [Renyi_Jensen_f] at Hext' + have CG1 (z : T) : q z = 0 → p z = 0 := by apply H + have CG2 (z : T) : ¬(p z ≠ 0 ∧ q z = ⊤) := by + simp + intro + apply PMF.apply_ne_top + conv at Hext' => + rhs + arg 1 + intro z + rw [Renyi_Jensen_f] + rw [<- ENNReal.toReal_mul] + arg 1 + rw [mul_comm] + rw [division_def] + rw [mul_mul_inv_eq_mul_cancel (CG1 z) (CG2 z)] + clear CG1 + clear CG2 + rw [<- ENNReal.tsum_toReal_eq] at Hext' + · rw [PMF.tsum_coe] at Hext' + apply (@ENNReal.mul_eq_mul_right _ _ ((q x)⁻¹) ?G1 ?G2).mp + case G1 => + simp + apply PMF.apply_ne_top + case G2 => + simp + apply Hq + rw [ENNReal.mul_inv_cancel ?G1 ?G2] + case G1 => apply Hq + case G2 => apply PMF.apply_ne_top + apply ENNReal_toReal_partial_inj at Hext' + rw [<- division_def] + apply Hext' + · apply HRJf_nt + · simp + · intro + apply PMF.apply_ne_top + + · -- Special case: There exists some element x0 with p x0 = ⊤ but q x0 ∈ ℝ+ + -- This means the sum in Hnts will actually be ⊤ + rename_i Hspecial + exfalso + apply Hnts + apply ENNReal.tsum_eq_top_of_eq_top + simp at Hspecial + rcases Hspecial with ⟨ x , ⟨ Hx1, ⟨ Hx2 , Hx3 ⟩ ⟩ ⟩ + exists x + apply mul_eq_top.mpr + right + apply And.intro + · apply rpow_eq_top_iff.mpr + right + apply And.intro + · apply div_eq_top.mpr + right + apply And.intro Hx1 Hx3 + · linarith + · apply Hx2 + · -- One of the series is Top, so the other series is too + rename_i Hlhs_top + simp at Hlhs_top + rw [Hlhs_top] at Hsumeq + -- This series should actually be 1 by PMF + conv at Hsumeq => + lhs + arg 1 + arg 1 + intro x + rw [PMF_mul_mul_inv_eq_mul_cancel p q H] + conv at Hsumeq => + lhs + arg 1 + rw [PMF.tsum_coe] + simp at Hsumeq + /-- Restriction of the PMF f to the support of q -/ @@ -798,10 +1043,6 @@ lemma RenyiDivergence_refl_zero (p : PMF T) {α : ℝ} (Hα : 1 < α) : (0 = Ren simp --- lemma Jensens_equality_converse_ENNReal (f g h : ENNReal -> ENNReal) (H : h (∑'(x : T), f x * g x) ≤ ∑'(x : T), h (f x) * g x) : --- (∃ a b : ENNReal, ∀ x : ENNReal, h x = a * x + b) := by --- sorry - diff --git a/SampCert/Util/Log.lean b/SampCert/Util/Log.lean index 73f2535d..1c1d677e 100644 --- a/SampCert/Util/Log.lean +++ b/SampCert/Util/Log.lean @@ -344,4 +344,7 @@ lemma ereal_smul_le_left (s : EReal) (Hr1 : 0 < s) (Hr2 : s < ⊤) (H : s * w lemma ereal_smul_eq_left (s : EReal) (Hr1 : 0 < s) (Hr2 : s < ⊤) (H : s * w = s * z) : w = z := by sorry + +lemma ENNReal_toReal_partial_inj (a b : ENNReal) (H : a.toReal = b.toReal) (H1 : a ≠ ⊤) (H2 : b ≠ ⊤) : a = b := sorry + end ENNReal From fc6804910caa7e7051e4486e8a6fdf9ef4073cbf Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Thu, 20 Jun 2024 15:14:59 -0400 Subject: [PATCH 074/122] RenyiDivergence_def_eq_0_iff --- .../DifferentialPrivacy/RenyiDivergence.lean | 110 +++++++++++++++--- 1 file changed, 97 insertions(+), 13 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index 33e3c332..e23b3cb0 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -259,10 +259,6 @@ theorem Renyi_Jensen_real (f : T → ℝ) (q : PMF T) (α : ℝ) (h : 1 < α) (h rw [one_le_ofReal] apply le_of_lt h --- set_option pp.notation false --- set_option pp.coercions false --- set_option pp.all true --- set_option pp.universes false /-- Jensen's strict ineuquality for the exponential applied to Renyi's sum @@ -586,7 +582,7 @@ lemma Renyi_Jensen_ENNReal_reduct [MeasurableSpace T] [MeasurableSingletonClass · -- Special case: There exists some element x0 with p x0 = ⊤ but q x0 ∈ ℝ+ rename_i Hspecial simp at * - rcases Hspecial with ⟨ x0, ⟨ H1, H2 , H3 ⟩⟩ + rcases Hspecial with ⟨ x0, ⟨ H1, _, H3 ⟩⟩ have HT1 : (∑' (x : T), p x / q x * q x) ^ α = ⊤ := by apply rpow_eq_top_iff.mpr right @@ -621,7 +617,6 @@ lemma Renyi_Jensen_ENNReal_reduct [MeasurableSpace T] [MeasurableSingletonClass exact OrderTop.le_top ((∑' (x : T), p x / q x * q x) ^ α) --- set_option pp.all true -- Lift the properties of Renyi_Jensen_strict_real to ENNReal -- q is still nonnegative @@ -949,10 +944,70 @@ theorem Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] [C rcases t with ⟨ a , Ha ⟩ apply (reducedPMF_pos Hq a Ha) +set_option pp.coercions false + +lemma Renyi_Jensen_ENNReal_converse [MeasurableSpace T] [MeasurableSingletonClass T] [Countable T] + (p q : PMF T) {α : ℝ} (h : 1 < α) (H : AbsCts p q) + (Hsumeq : (∑' x : T, (p x / q x) * q x) ^ α = (∑' x : T, (p x / q x) ^ α * q x)) : + (p = q) := by + + have K1 : Function.support (fun x : T => (p x / q x) * q x) ⊆ { t : T | q t ≠ 0 } := by simp [Function.support] + have K2 : Function.support (fun x : T => (p x / q x)^α * q x) ⊆ { t : T | q t ≠ 0 } := by simp [Function.support] + rw [<- tsum_subtype_eq_of_support_subset K1] at Hsumeq + rw [<- tsum_subtype_eq_of_support_subset K2] at Hsumeq + simp at Hsumeq + have Hq : AbsCts q q := AbsCts_refl q + have B1 (x : { x // ¬q x = 0 }) : p ↑x / q ↑x * q ↑x = reducedPMF H x / reducedPMF Hq x * reducedPMF Hq x := by congr + have B2 (x : { x // ¬q x = 0 }) : (p ↑x / q ↑x)^α * q ↑x = (reducedPMF H x / reducedPMF Hq x)^α * reducedPMF Hq x := by congr --- FIXME + conv at Hsumeq => + congr + · arg 1 + arg 1 + intro x + rw [B1 x] + · arg 1 + intro x + rw [B2 x] + + clear B1 + clear B2 + clear K1 + clear K2 + + have Hreduced : (reducedPMF H = reducedPMF Hq) := by + apply (Renyi_Jensen_ENNReal_converse_reduct (reducedPMF H) (reducedPMF Hq) h) + · intro t Ht + exfalso + rcases t with ⟨ a , Ha ⟩ + apply (reducedPMF_pos Hq a Ha) + apply Ht + · intro t + rcases t with ⟨ a , Ha ⟩ + apply (reducedPMF_pos Hq a Ha) + · apply Hsumeq + + apply PMF.ext + intro x + cases (Classical.em (q x = 0)) + · rename_i Hqz + rw [Hqz] + apply H + apply Hqz + · rename_i Hqnz + have Hreduced' : reducedPMF H ⟨ x , Hqnz ⟩ = reducedPMF Hq ⟨ x , Hqnz ⟩ := by + exact congrFun (congrArg DFunLike.coe Hreduced) ⟨ x , Hqnz ⟩ + repeat rw [DFunLike.coe] at Hreduced' + repeat rw [PMF.instFunLike] at Hreduced' + repeat rw [reducedPMF] at Hreduced' + unfold reducedPMF_def at Hreduced' + simp at Hreduced' + assumption + + +-- FIXME (or not ?) /-- The Renyi divergence is monotonic in the value of its sum. -/ @@ -997,8 +1052,6 @@ theorem RenyiDivergence_def_nonneg [MeasurableSpace T] [MeasurableSingletonClass exact EReal.coe_lt_top (α - OfNat.ofNat 1) · assumption -set_option pp.coercions false - lemma RenyiDivergence_refl_zero (p : PMF T) {α : ℝ} (Hα : 1 < α) : (0 = RenyiDivergence_def p p α) := by have H1 : 1 = eexp ((α - 1) * RenyiDivergence_def p p α) := by rw [RenyiDivergence_def_exp p p Hα] @@ -1042,10 +1095,41 @@ lemma RenyiDivergence_refl_zero (p : PMF T) {α : ℝ} (Hα : 1 < α) : (0 = Ren rw [<- H1] simp - - - - +/-- +Renyi divergence attains equality if and only if the distributions are equal +-/ +theorem RenyiDivergence_def_eq_0_iff [MeasurableSpace T] [MeasurableSingletonClass T] [Countable T] + (p q : PMF T) {α : ℝ} (Hα : 1 < α) (Hcts : AbsCts p q) : + (RenyiDivergence_def p q α = 0) <-> (p = q) := by + apply Iff.intro + · intro Hrdeq + apply Renyi_Jensen_ENNReal_converse + · apply Hα + · apply Hcts + · have H1 : eexp (((α - 1)) * 0) = eexp ((α - 1) * RenyiDivergence_def p q α) := by rw [Hrdeq] + simp at H1 + rw [RenyiDivergence_def_exp p q Hα] at H1 + rw [RenyiDivergenceExpectation p q Hα Hcts] at H1 + rw [<- H1] + clear H1 + have CG1 (x : T) : DFunLike.coe q x = OfNat.ofNat 0 → DFunLike.coe p x = OfNat.ofNat 0 := by apply Hcts + have CG2 (x : T) : ¬(DFunLike.coe p x ≠ OfNat.ofNat 0 ∧ DFunLike.coe q x = ⊤) := by + simp + intro + apply PMF.apply_ne_top + conv => + lhs + arg 1 + arg 1 + intro x + rw [division_def] + rw [mul_mul_inv_eq_mul_cancel (CG1 x) (CG2 x)] + simp + · intro Hpq + rw [Hpq] + symm + apply RenyiDivergence_refl_zero + apply Hα lemma RenyiDivergence_zero_eq (p q : PMF T) {α : ℝ} (Hα : 1 < α) (Hpq : AbsCts p q) (HR : 0 = RenyiDivergence_def p q α) : p = q:= by From 3c99563e7f53047e50d9a198c33e7111414a97d8 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Thu, 20 Jun 2024 15:16:20 -0400 Subject: [PATCH 075/122] minor cleanup --- SampCert/DifferentialPrivacy/RenyiDivergence.lean | 11 ----------- 1 file changed, 11 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index e23b3cb0..6345e538 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -1132,16 +1132,6 @@ theorem RenyiDivergence_def_eq_0_iff [MeasurableSpace T] [MeasurableSingletonCla apply Hα -lemma RenyiDivergence_zero_eq (p q : PMF T) {α : ℝ} (Hα : 1 < α) (Hpq : AbsCts p q) (HR : 0 = RenyiDivergence_def p q α) : p = q:= by - have HR' : 1 = eexp (((α - 1)) * RenyiDivergence_def p q α) := by simp [<- HR] - rw [RenyiDivergence_def_exp p q Hα] at HR' - rw [RenyiDivergenceExpectation p q Hα Hpq] at HR' - have Hlhs : (∑' (x : T), (p x / q x) * q x)^α = 1:= sorry - rw [<- Hlhs] at HR' - - -- Is this true? - sorry - lemma RenyiDivergence_def_log_sum_nonneg (p q : PMF T) (α : ℝ) : (0 ≤ (elog (∑' x : T, (p x)^α * (q x)^(1 - α)))) := by -- Follows from RenyiDivergence_def_nonneg sorry @@ -1157,7 +1147,6 @@ theorem RenyiDivergence_aux_zero (p q : PMF T) (α : ℝ) : p = q <-> RenyiDiver := sorry - -- Unused /- Closed form for the Renyi Divergence. From 42be9fb73202b4ed4ead9af80536169a8732d33a Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Thu, 20 Jun 2024 15:58:42 -0400 Subject: [PATCH 076/122] completed concentrated bound --- .../ZeroConcentrated/ConcentratedBound.lean | 43 +++++++++++++------ 1 file changed, 30 insertions(+), 13 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean index 8ad13078..603e99bd 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean @@ -533,7 +533,6 @@ theorem Renyi_Gauss_divergence_bound' {σ α : ℝ} (h : σ ≠ 0) (h' : 1 < α) RenyiDivergence (discrete_gaussian_pmf h μ) (discrete_gaussian_pmf h ν) α ≤ (ENNReal.ofReal α) * (ENNReal.ofReal ((((μ - ν) : ℤ)^2 : ℝ) / (2 * σ^2))) := by - -- Fixable have A : RenyiDivergence (discrete_gaussian_pmf h μ) (discrete_gaussian_pmf h ν) α = ENNReal.ofReal (RenyiDivergence' (fun (x : ℤ) => discrete_gaussian σ μ x) (fun (x : ℤ) => discrete_gaussian σ ν x) α) := by unfold RenyiDivergence @@ -543,8 +542,11 @@ theorem Renyi_Gauss_divergence_bound' {σ α : ℝ} (h : σ ≠ 0) (h' : 1 < α) simp unfold discrete_gaussian_pmf - have Hdg_pos (x : ℤ) (w : ℝ) : OfNat.ofNat 0 < discrete_gaussian σ w x.cast := by sorry - have Hdg_pow_pos (x : ℤ) w : OfNat.ofNat 0 ≤ discrete_gaussian σ w x.cast ^ α := by sorry + have Hdg_pos (x : ℤ) (w : ℝ) : OfNat.ofNat 0 < discrete_gaussian σ w x.cast := by + exact discrete_gaussian_pos h w x + have Hdg_pow_pos (x : ℤ) w : OfNat.ofNat 0 ≤ discrete_gaussian σ w x.cast ^ α := by + apply rpow_nonneg + exact discrete_gaussian_nonneg h w x conv => lhs @@ -560,19 +562,33 @@ theorem Renyi_Gauss_divergence_bound' {σ α : ℝ} (h : σ ≠ 0) (h' : 1 < α) rw [<- ENNReal.ofEReal_ofReal_toENNReal] simp congr - -- Here is where I would do the classical I talk about below cases (Classical.em (0 = ∑' (x : ℤ), discrete_gaussian σ μ.cast x.cast ^ α * discrete_gaussian σ ν.cast x.cast ^ (1 - α))) · rename_i Hzero rw [<- Hzero] simp rw [<- ENNReal.ofReal_tsum_of_nonneg] · rw [<- Hzero] - simp - -- Sadly this is not true (ultimately because they set Real.log 0 = 0) - -- Fixable: This series should only ever be zero when μ = ν, if the paper is right. - -- We also should prove that μ ≠ ν <-> (discretVe_gaussian_pmf h μ) ≠ (discretVe_gaussian_pmf h ν) - -- In the outermost proof, that case is trivial. - sorry + exfalso + symm at Hzero + have Hzero' : (ENNReal.ofReal (∑' (x : ℤ), discrete_gaussian σ ↑μ ↑x ^ α * discrete_gaussian σ ↑ν ↑x ^ (1 - α)) = 0) := by + simp [Hzero] + rw [ENNReal.ofReal_tsum_of_nonneg ?G1 ?G2] at Hzero' + case G1 => exact fun n => Renyi_sum_SG_nonneg h μ ν n + case G2 => exact Renyi_Gauss_summable h μ ν α + apply ENNReal.tsum_eq_zero.mp at Hzero' + have Hzero'' := Hzero' (0 : ℤ) + simp at Hzero'' + have C : (0 < discrete_gaussian σ (↑μ) 0 ^ α * discrete_gaussian σ (↑ν) 0 ^ (1 - α)) := by + apply Real.mul_pos + · apply Real.rpow_pos_of_pos + have A := discrete_gaussian_pos h μ (0 : ℤ) + simp at A + apply A + · apply Real.rpow_pos_of_pos + have A := discrete_gaussian_pos h ν (0 : ℤ) + simp at A + apply A + linarith · exact fun n => Renyi_sum_SG_nonneg h μ ν n · exact Renyi_Gauss_summable h μ ν α · rename_i Hnz @@ -582,9 +598,10 @@ theorem Renyi_Gauss_divergence_bound' {σ α : ℝ} (h : σ ≠ 0) (h' : 1 < α) · intro n exact Renyi_sum_SG_nonneg h μ ν n apply lt_of_le_of_ne - · sorry - · sorry - -- I could to add a classical case to finish this, assuming the + · apply tsum_nonneg + intro i + exact Renyi_sum_SG_nonneg h μ ν i + · apply Hnz rw [A] rw [<- ENNReal.ofReal_mul] apply ENNReal.ofReal_le_ofReal From 9562ab91ebcf5a298564c8a3fed1a66cc4150d9b Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Thu, 20 Jun 2024 17:47:13 -0400 Subject: [PATCH 077/122] checkpoint --- .../ZeroConcentrated/Composition.lean | 72 +++++++++++++------ SampCert/Util/Log.lean | 3 +- 2 files changed, 54 insertions(+), 21 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean index 9111ae8c..fd09374b 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean @@ -74,14 +74,14 @@ theorem Renyi_noised_query_NZ {nq : List T → SLang U} {HNorm : NormalMechanism contradiction . exact h5 +set_option pp.coercions false /-- Composed queries satisfy zCDP Renyi divergence bound. -/ theorem privCompose_zCDPBound {nq1 : List T → SLang U} {nq2 : List T → SLang V} {HNorm1 : ∀ l, HasSum (nq1 l) 1} {HNorm2 : ∀ l, HasSum (nq2 l) 1} {ε₁ ε₂ ε₃ ε₄ : ℕ+} - (h1 : zCDPBound nq1 HNorm1 ((ε₁ : ℝ) / ε₂)) (h2 : zCDPBound nq2 HNorm2 ((ε₃ : ℝ) / ε₄)) : -- (nn1 : NonZeroNQ nq1) (nn2 : NonZeroNQ nq2) (nt1 : NonTopRDNQ nq1) (nt2 : NonTopRDNQ nq2) (nts1 : NonTopNQ nq1) (nts2 : NonTopNQ nq2) : + (h1 : zCDPBound nq1 HNorm1 ((ε₁ : ℝ) / ε₂)) (h2 : zCDPBound nq2 HNorm2 ((ε₃ : ℝ) / ε₄)) : zCDPBound (privCompose nq1 nq2) (privCompose_norm nq1 nq2 HNorm1 HNorm2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by - /- simp [privCompose, RenyiDivergence, zCDPBound] intro α h3 l₁ l₂ h4 have X := h1 @@ -90,6 +90,31 @@ theorem privCompose_zCDPBound {nq1 : List T → SLang U} {nq2 : List T → SLang replace h1 := h1 α h3 l₁ l₂ h4 replace h2 := h2 α h3 l₁ l₂ h4 simp [RenyiDivergence] at h1 h2 + unfold RenyiDivergence_def + unfold RenyiDivergence_def at h1 h2 + rw [DFunLike.coe] + rw [PMF.instFunLike] + simp + repeat rw [SLang.toPMF] + have CG1 (b : U) : nq1 l₂ b ≠ ⊤ := by + have HS : (∑'(u : U), nq1 l₂ u) = 1 := by exact HasSum.tsum_eq (HNorm1 l₂) + have Hs : nq1 l₂ b ≤ (∑'(u : U), nq1 l₂ u) := by exact ENNReal.le_tsum b + rw [HS] at Hs + rename_i inst inst_1 + simp_all only [ne_eq] + apply Aesop.BuiltinRules.not_intro + intro a + simp_all only [top_le_iff, one_ne_top] + have CG2 (b : V) : nq2 l₂ b ≠ ⊤ := by + have HS : (∑'(u : V), nq2 l₂ u) = 1 := by exact HasSum.tsum_eq (HNorm2 l₂) + have Hs : nq2 l₂ b ≤ (∑'(u : V), nq2 l₂ u) := by exact ENNReal.le_tsum b + rw [HS] at Hs + rename_i inst inst_1 + simp_all only [ne_eq] + apply Aesop.BuiltinRules.not_intro + intro a + simp_all only [top_le_iff, one_ne_top] + simp rw [tsum_prod' ENNReal.summable (fun b : U => ENNReal.summable)] . simp conv => @@ -104,13 +129,15 @@ theorem privCompose_zCDPBound {nq1 : List T → SLang U} {nq2 : List T → SLang rw [compose_sum_rw] rw [compose_sum_rw] rw [ENNReal.mul_rpow_of_nonneg _ _ (le_of_lt (lt_trans zero_lt_one h3))] - rw [ENNReal.mul_rpow_of_ne_top (nts1 l₂ b) (nts2 l₂ c)] + rw [ENNReal.mul_rpow_of_ne_top (CG1 b) (CG2 c)] rw [mul_assoc] right rw [mul_comm] rw [mul_assoc] right rw [mul_comm] + clear CG1 + clear CG2 conv => left right @@ -129,24 +156,29 @@ theorem privCompose_zCDPBound {nq1 : List T → SLang U} {nq2 : List T → SLang 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 Renyi_noised_query_NZ h3 h4 X nn1 nt1 nts1 - . apply Renyi_noised_query_NZ h3 h4 Y nn2 nt2 nts2 - -/ - sorry + rw [← elog_mul] + + + -- Need product/sum rule for elog, is it true? + + sorry + -- 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 Renyi_noised_query_NZ h3 h4 X nn1 nt1 nts1 + -- . apply Renyi_noised_query_NZ h3 h4 Y nn2 nt2 nts2 /-- All outputs of a composed query have finite probability. diff --git a/SampCert/Util/Log.lean b/SampCert/Util/Log.lean index 1c1d677e..6881cc99 100644 --- a/SampCert/Util/Log.lean +++ b/SampCert/Util/Log.lean @@ -136,13 +136,14 @@ lemma eexp_elog : (elog (eexp w)) = w := by @[simp] -lemma elog_mul : elog x * elog y = elog (x + y) := by +lemma elog_mul : elog x + elog y = elog (x * y) := by sorry -- checked truth table @[simp] lemma eexp_add : eexp w * eexp z = eexp (w + z) := by sorry -- checked truth table + -- Log of power, log and exp inverses lemma eexp_injective : eexp w = eexp z -> w = z := by From 96e4d1e6a785436c056f5e1186aa1f52152c2b5e Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Fri, 21 Jun 2024 11:34:10 -0400 Subject: [PATCH 078/122] zCDP composition, up to absolute continuity assumption --- .../DifferentialPrivacy/RenyiDivergence.lean | 13 +- .../ZeroConcentrated/Composition.lean | 139 ++++++++++++++---- SampCert/Util/Log.lean | 79 +++++++++- 3 files changed, 199 insertions(+), 32 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index 6345e538..ca292724 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -944,8 +944,6 @@ theorem Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] [C rcases t with ⟨ a , Ha ⟩ apply (reducedPMF_pos Hq a Ha) -set_option pp.coercions false - lemma Renyi_Jensen_ENNReal_converse [MeasurableSpace T] [MeasurableSingletonClass T] [Countable T] (p q : PMF T) {α : ℝ} (h : 1 < α) (H : AbsCts p q) (Hsumeq : (∑' x : T, (p x / q x) * q x) ^ α = (∑' x : T, (p x / q x) ^ α * q x)) : @@ -1052,6 +1050,7 @@ theorem RenyiDivergence_def_nonneg [MeasurableSpace T] [MeasurableSingletonClass exact EReal.coe_lt_top (α - OfNat.ofNat 1) · assumption + lemma RenyiDivergence_refl_zero (p : PMF T) {α : ℝ} (Hα : 1 < α) : (0 = RenyiDivergence_def p p α) := by have H1 : 1 = eexp ((α - 1) * RenyiDivergence_def p p α) := by rw [RenyiDivergence_def_exp p p Hα] @@ -1132,9 +1131,15 @@ theorem RenyiDivergence_def_eq_0_iff [MeasurableSpace T] [MeasurableSingletonCla apply Hα -lemma RenyiDivergence_def_log_sum_nonneg (p q : PMF T) (α : ℝ) : (0 ≤ (elog (∑' x : T, (p x)^α * (q x)^(1 - α)))) := by +lemma RenyiDivergence_def_log_sum_nonneg [MeasurableSpace T] [MeasurableSingletonClass T] [Countable T] + (p q : PMF T) (Hac : AbsCts p q) {α : ℝ} (Hα : 1 < α ): (0 ≤ (elog (∑' x : T, (p x)^α * (q x)^(1 - α)))) := by -- Follows from RenyiDivergence_def_nonneg - sorry + have Hrd := RenyiDivergence_def_nonneg p q Hac Hα + rw [RenyiDivergence_def] at Hrd + apply ofEReal_nonneg_scal_l at Hrd + · assumption + · apply inv_pos_of_pos + linarith /-- The Renyi divergence. diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean index fd09374b..e5fa0435 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean @@ -22,6 +22,8 @@ namespace SLang variable { T U V : Type } variable [Inhabited U] variable [Inhabited V] +variable [MeasurableSpace U] [MeasurableSingletonClass U] [Countable U] +variable [MeasurableSpace V] [MeasurableSingletonClass V] [Countable V] lemma ENNReal_toReal_NZ (x : ENNReal) (h1 : x ≠ 0) (h2 : x ≠ ⊤) : x.toReal ≠ 0 := by @@ -156,29 +158,117 @@ theorem privCompose_zCDPBound {nq1 : List T → SLang U} {nq2 : List T → SLang right intro b rw [ENNReal.tsum_mul_left] - rw [ENNReal.tsum_mul_right] rw [← elog_mul] + conv at h1 => + lhs + arg 1 + arg 2 + arg 1 + arg 1 + intro x + rw [DFunLike.coe] + rw [PMF.instFunLike] + rw [SLang.toPMF] + rw [SLang.toPMF] + simp + conv at h2 => + lhs + arg 1 + arg 2 + arg 1 + arg 1 + intro x + rw [DFunLike.coe] + rw [PMF.instFunLike] + rw [SLang.toPMF] + rw [SLang.toPMF] + simp + + have log_nonneg_1 : 0 ≤ (∑' (i : U), nq1 l₁ i ^ α * nq1 l₂ i ^ (1 - α)).elog := by + have Hac1 : AbsCts ((nq1 l₁).toPMF (HNorm1 l₁)) ((nq1 l₂).toPMF (HNorm1 l₂)) := by sorry + have Hnn1 := (RenyiDivergence_def_log_sum_nonneg ((nq1 l₁).toPMF (HNorm1 l₁)) ((nq1 l₂).toPMF (HNorm1 l₂)) Hac1 h3) + conv at Hnn1 => + rhs + arg 1 + arg 1 + intro x + rw [DFunLike.coe] + rw [PMF.instFunLike] + rw [SLang.toPMF] + rw [SLang.toPMF] + simp + apply Hnn1 + have log_nonneg_2 : 0 ≤ (∑' (i : V), nq2 l₁ i ^ α * nq2 l₂ i ^ (1 - α)).elog := by + have Hac2 : AbsCts ((nq2 l₁).toPMF (HNorm2 l₁)) ((nq2 l₂).toPMF (HNorm2 l₂)) := by sorry + have Hnn2 := (RenyiDivergence_def_log_sum_nonneg ((nq2 l₁).toPMF (HNorm2 l₁)) ((nq2 l₂).toPMF (HNorm2 l₂)) Hac2 h3) + conv at Hnn2 => + rhs + arg 1 + arg 1 + intro x + rw [DFunLike.coe] + rw [PMF.instFunLike] + rw [SLang.toPMF] + rw [SLang.toPMF] + simp + apply Hnn2 - -- Need product/sum rule for elog, is it true? + -- Split up the series + rw [ofEReal_mul_nonneg] + · simp only [ofEReal_real] + -- In order to distribute ofReal, we need the logarithms to be nonegative + rw [ofEReal_plus_nonneg log_nonneg_1 log_nonneg_2] - sorry - -- 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 Renyi_noised_query_NZ h3 h4 X nn1 nt1 nts1 - -- . apply Renyi_noised_query_NZ h3 h4 Y nn2 nt2 nts2 + -- Distribute + rw [CanonicallyOrderedCommSemiring.left_distrib] + apply (@le_trans _ _ _ (ENNReal.ofReal (2⁻¹ * (↑↑ε₁ ^ 2 / ↑↑ε₂ ^ 2) * α) + ENNReal.ofReal (2⁻¹ * (↑↑ε₃ ^ 2 / ↑↑ε₄ ^ 2) * α))) + · -- apply? + apply _root_.add_le_add + · rw [ofEReal_mul_nonneg] at h1 + · apply h1 + · apply EReal.coe_nonneg.mpr + apply inv_nonneg.mpr + linarith + · apply log_nonneg_1 + · rw [ofEReal_mul_nonneg] at h2 + · apply h2 + · apply EReal.coe_nonneg.mpr + apply inv_nonneg.mpr + linarith + · apply log_nonneg_2 + · clear h1 h2 + rw [<- ENNReal.ofReal_add] + · apply ofReal_le_ofReal_iff'.mpr + left + 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 only [inv_pos, Nat.ofNat_pos] + · linarith + · apply mul_nonneg + · apply mul_nonneg + · simp + · apply div_nonneg + · exact sq_nonneg ε₁.val.cast + · exact sq_nonneg ε₂.val.cast + · linarith + · apply mul_nonneg + · apply mul_nonneg + · simp + · apply div_nonneg + · exact sq_nonneg ε₃.val.cast + · exact sq_nonneg ε₄.val.cast + · linarith + · simp + linarith + · apply Left.add_nonneg + · apply log_nonneg_1 + · apply log_nonneg_2 /-- All outputs of a composed query have finite probability. @@ -273,14 +363,9 @@ theorem privCompose_NonTopRDNQ {nq1 : List T → SLang U} {nq2 : List T → SLan theorem privCompose_zCDP (nq1 : List T → SLang U) (nq2 : List T → SLang V) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h : zCDP nq1 ((ε₁ : ℝ) / ε₂)) (h' : zCDP nq2 ((ε₃ : ℝ) / ε₄)) : zCDP (privCompose nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by simp [zCDP] at * - sorry - -- cases h ; rename_i h1 h2 ; cases h2 ; rename_i h2 h3 ; cases h3 ; rename_i h3 h4 ; cases h4 ; rename_i h4 h5 - -- cases h' ; rename_i h'1 h'2 ; cases h'2 ; rename_i h'2 h'3 ; cases h'3 ; rename_i h'3 h'4 ; cases h'4 ; rename_i h'4 h'5 - -- repeat any_goals constructor - -- . apply privCompose_zCDPBound h1 h'1 h2 h'2 h5 h'5 h4 h'4 - -- . apply privCompose_NonZeroNQ h2 h'2 - -- . apply privCompose_NonTopSum h3 h'3 - -- . apply privCompose_NonTopNQ h4 h'4 - -- . apply privCompose_NonTopRDNQ h5 h'5 h4 h'4 + rcases h with ⟨ Hn1, Hb1 ⟩ + rcases h' with ⟨ Hn2, Hb2 ⟩ + exists (privCompose_norm nq1 nq2 Hn1 Hn2) + exact privCompose_zCDPBound Hb1 Hb2 end SLang diff --git a/SampCert/Util/Log.lean b/SampCert/Util/Log.lean index 6881cc99..62f0858b 100644 --- a/SampCert/Util/Log.lean +++ b/SampCert/Util/Log.lean @@ -31,6 +31,35 @@ noncomputable def ofEReal (e : EReal) : ENNReal := | ⊤ => ⊤ | some (some r) => ENNReal.ofReal r +lemma EReal_cases (w : EReal) : w = ⊥ ∨ w = ⊤ ∨ (∃ v : ℝ, w = v) := by + cases w + · left + rfl + rename_i w' + cases w' + · right + left + rfl + rename_i wR + right + right + exists wR + +@[simp] +lemma ofEReal_bot : ofEReal ⊥ = 0 := by simp [ofEReal] + +@[simp] +lemma ofEReal_top : ofEReal ⊤ = ⊤ := by simp [ofEReal] + +@[simp] +lemma ofEReal_real (r : ℝ) : ofEReal r = ENNReal.ofReal r := by + simp [Real.toEReal] + simp [ofEReal] + +lemma ofEReal_nonpos (w : EReal) (HW : w ≤ 0): ofEReal w = 0 := by + rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> simp_all + + -- instance : Coe EReal ENNReal := ⟨ofEReal⟩ /-- @@ -144,6 +173,7 @@ lemma elog_mul : elog x + elog y = elog (x * y) := by lemma eexp_add : eexp w * eexp z = eexp (w + z) := by sorry -- checked truth table + -- Log of power, log and exp inverses lemma eexp_injective : eexp w = eexp z -> w = z := by @@ -198,9 +228,56 @@ lemma ofEReal_nonneg_eq_iff (Hw : 0 <= w) (Hz : 0 <= z) : w = z <-> (ofEReal w = lemma ofEReal_le_mono : (0 ≤ w) -> w ≤ z <-> (ofEReal w ≤ ofEReal z) := sorry +#check 3 @[simp] -lemma ofEReal_mul (Hw : 0 ≤ w) (Hz : 0 ≤ z) : ofEReal (w * z) = ofEReal w * ofEReal z := sorry +lemma ofEReal_plus_nonneg (Hw : 0 ≤ w) (Hz : 0 ≤ z) : ofEReal (w + z) = ofEReal w + ofEReal z := by + rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> + rcases (EReal_cases z) with Hz' | (Hz' | ⟨ z', Hz' ⟩) + all_goals rw [Hw', Hz'] + all_goals simp_all + sorry + +@[simp] +lemma ofEReal_mul_nonneg (Hw : 0 ≤ w) (Hz : 0 ≤ z) : ofEReal (w * z) = ofEReal w * ofEReal z := by + have HBle: (⊥ : EReal) ≤ 0 := by exact OrderBot.bot_le 0 + rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> + rcases (EReal_cases z) with Hz' | (Hz' | ⟨ z', Hz' ⟩) + all_goals rw [Hw', Hz'] + all_goals simp_all + · rw [top_mul'] + split + · simp_all + apply ofEReal_nonpos + sorry + · cases (LE.le.lt_or_eq Hz) + · sorry + · sorry + · sorry + · sorry + +lemma ofEReal_nonneg_scal_l (H1 : 0 < r) (H2 : 0 ≤ r * w) : 0 ≤ w := by + rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) + · simp_all + sorry + · simp_all + · simp_all + sorry + + + + -- <;> rcases (EReal_cases z) with ⟨ Hz , Hz , ⟨ z' , Hz ⟩⟩ + + + + + -- cases w <;> cases z + -- · exfalso + -- rw [WithBot.none_eq_bot] at Hw + -- simp at * + -- · sorry + -- · sorry + -- · sorry @[simp] lemma toEReal_ofENNReal_nonneg (H : 0 ≤ w) : ENNReal.toEReal (ofEReal w) = w := sorry From 2b8058ff282e0ff5d98c7f9f983c6d77f192789d Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Fri, 21 Jun 2024 12:38:08 -0400 Subject: [PATCH 079/122] mechanism zCDP --- .../ZeroConcentrated/Composition.lean | 258 ++++++----- .../Mechanism/Properties.lean | 432 ++++++++++-------- 2 files changed, 357 insertions(+), 333 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean index e5fa0435..005efa52 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean @@ -25,58 +25,56 @@ variable [Inhabited V] variable [MeasurableSpace U] [MeasurableSingletonClass U] [Countable U] variable [MeasurableSpace V] [MeasurableSingletonClass V] [Countable V] -lemma ENNReal_toReal_NZ (x : ENNReal) (h1 : x ≠ 0) (h2 : x ≠ ⊤) : - x.toReal ≠ 0 := by - unfold ENNReal.toReal - unfold ENNReal.toNNReal - simp - constructor ; any_goals trivial - -lemma simp_α_1 {α : ℝ} (h : 1 < α) : 0 < α := by - apply @lt_trans _ _ _ 1 _ _ h - simp only [zero_lt_one] +-- lemma ENNReal_toReal_NZ (x : ENNReal) (h1 : x ≠ 0) (h2 : x ≠ ⊤) : +-- x.toReal ≠ 0 := by +-- unfold ENNReal.toReal +-- unfold ENNReal.toNNReal +-- simp +-- constructor ; any_goals trivial +-- +-- lemma simp_α_1 {α : ℝ} (h : 1 < α) : 0 < α := by +-- apply @lt_trans _ _ _ 1 _ _ h +-- simp only [zero_lt_one] -/-- +/- The Renyi Divergence between neighbouring inputs of noised queries is nonzero. -/ -theorem Renyi_noised_query_NZ {nq : List T → SLang U} {HNorm : NormalMechanism nq} {α ε : ℝ} - (h1 : 1 < α) {l₁ l₂ : List T} (h2 : Neighbour l₁ l₂) (h3 : zCDPBound nq HNorm ε) - (h4 : NonZeroNQ nq) (h5 : NonTopRDNQ nq) (nts : NonTopNQ nq) : - (∑' (i : U), nq l₁ i ^ α * nq l₂ i ^ (1 - α)).toReal ≠ 0 := by - simp [zCDPBound] 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_toReal_NZ - . by_contra CONTRA - rw [ENNReal.tsum_eq_zero] at CONTRA - replace CONTRA := CONTRA default - replace h6 := h6 default - replace h7 := h7 default - 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₂ default - contradiction - . exact h5 - -set_option pp.coercions false +-- theorem Renyi_noised_query_NZ {nq : List T → SLang U} {HNorm : NormalMechanism nq} {α ε : ℝ} +-- (h1 : 1 < α) {l₁ l₂ : List T} (h2 : Neighbour l₁ l₂) (h3 : zCDPBound nq HNorm ε) +-- (h4 : NonZeroNQ nq) (h5 : NonTopRDNQ nq) (nts : NonTopNQ nq) : +-- (∑' (i : U), nq l₁ i ^ α * nq l₂ i ^ (1 - α)).toReal ≠ 0 := by +-- simp [zCDPBound] 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_toReal_NZ +-- . by_contra CONTRA +-- rw [ENNReal.tsum_eq_zero] at CONTRA +-- replace CONTRA := CONTRA default +-- replace h6 := h6 default +-- replace h7 := h7 default +-- 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₂ default +-- contradiction +-- . exact h5 /-- Composed queries satisfy zCDP Renyi divergence bound. @@ -270,92 +268,92 @@ theorem privCompose_zCDPBound {nq1 : List T → SLang U} {nq2 : List T → SLang · apply log_nonneg_1 · apply log_nonneg_2 -/-- +/- All outputs of a composed query have finite probability. -/ -theorem privCompose_NonTopNQ {nq1 : List T → SLang U} {nq2 : List T → SLang V} (nt1 : NonTopNQ nq1) (nt2 : NonTopNQ nq2) : - NonTopNQ (privCompose nq1 nq2) := by - simp [NonTopNQ] at * - intro l a b - replace nt1 := nt1 l a - replace nt2 := nt2 l b - simp [privCompose] - 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 privCompose_NonTopNQ {nq1 : List T → SLang U} {nq2 : List T → SLang V} (nt1 : NonTopNQ nq1) (nt2 : NonTopNQ nq2) : +-- NonTopNQ (privCompose nq1 nq2) := by +-- simp [NonTopNQ] at * +-- intro l a b +-- replace nt1 := nt1 l a +-- replace nt2 := nt2 l b +-- simp [privCompose] +-- rw [compose_sum_rw] +-- rw [mul_eq_top] +-- intro H +-- cases H +-- . rename_i H +-- cases H +-- contradiction +-- . rename_i H +-- cases H +-- contradiction -/-- +/- Renyi divergence beteeen composed queries on neighbours are finite. -/ -theorem privCompose_NonTopRDNQ {nq1 : List T → SLang U} {nq2 : List T → SLang V} (nt1 : NonTopRDNQ nq1) (nt2 : NonTopRDNQ nq2) (nn1 : NonTopNQ nq1) (nn2 : NonTopNQ nq2) : - NonTopRDNQ (privCompose 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 [privCompose] - 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_nonneg _ _ (le_of_lt (lt_trans zero_lt_one h1))] - rw [ENNReal.mul_rpow_of_ne_top (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 +-- theorem privCompose_NonTopRDNQ {nq1 : List T → SLang U} {nq2 : List T → SLang V} (nt1 : NonTopRDNQ nq1) (nt2 : NonTopRDNQ nq2) (nn1 : NonTopNQ nq1) (nn2 : NonTopNQ nq2) : +-- NonTopRDNQ (privCompose 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 [privCompose] +-- 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_nonneg _ _ (le_of_lt (lt_trans zero_lt_one h1))] +-- rw [ENNReal.mul_rpow_of_ne_top (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 /-- ``privCompose`` satisfies zCDP diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean index 463ee6da..5bbd126c 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean @@ -18,7 +18,7 @@ open Classical Nat Int Real ENNReal MeasureTheory Measure namespace SLang -lemma privNoisedQuery_norm (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_sensitivity : sensitivity query Δ) : +lemma privNoisedQuery_norm (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : -- (bounded_sensitivity : sensitivity query Δ) : NormalMechanism (privNoisedQuery query Δ ε₁ ε₂) := by rw [NormalMechanism] intro l @@ -26,12 +26,17 @@ lemma privNoisedQuery_norm (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bou exact DiscreteGaussianGen_sum (Δ * ε₂) ε₁ (query l) set_option pp.coercions false +-- set_option pp.notation false +-- set_option pp.all true +-- set_option pp.universes false + +-- lemma L (x : ℝ) : x = 1 * x := by exact? /-- The zCDP mechanism with bounded sensitivity satisfies the bound for ``(Δε₂/ε₁)^2``-zCDP. -/ theorem privNoisedQuery_zCDPBound (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_sensitivity : sensitivity query Δ) : - zCDPBound (privNoisedQuery query Δ ε₁ ε₂) (privNoisedQuery_norm query Δ ε₁ ε₂ bounded_sensitivity) ((ε₁ : ℝ) / ε₂) := by + zCDPBound (privNoisedQuery query Δ ε₁ ε₂) (privNoisedQuery_norm query Δ ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by simp [zCDPBound, privNoisedQuery] intros α h1 l₁ l₂ h2 have A := @discrete_GaussianGenSample_ZeroConcentrated α h1 (Δ * ε₂) ε₁ (query l₁) (query l₂) @@ -134,83 +139,104 @@ theorem privNoisedQuery_zCDPBound (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ rw [← mul_assoc] -- Almost would be easier just to redo it... from here, the main reduction has already been finished + left - -- 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 + -- Move α to left + conv => + rhs + rw [mul_assoc] + rw [mul_comm] + -- Move ε₁^2 to the left + conv => + lhs + rw [<- mul_assoc] + arg 1 + rw [mul_assoc] + arg 2 + rw [mul_comm] + conv => + lhs + rw [<- mul_assoc] + rw [<- mul_assoc] + -- Add factor of 1 on the right + conv => + rhs + rw [<- (mul_one (α * ε₁ ^ 2 * ((ε₂ : ℝ)^ 2)⁻¹ : ℝ))] + -- Reassocaite on left + conv => + lhs + rw [mul_assoc] + apply mul_le_mul + · simp only [inv_pow, le_refl] + · simp + rw [inv_mul_eq_div] + rw [div_le_one_iff] + left + apply And.intro + · apply sq_pos_iff.mpr simp - apply @le_trans ℝ _ 0 1 α (zero_le_one' ℝ) (le_of_lt h1) - sorry - -- apply mul_le_mul - -- 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 - . sorry - -- rw [pow_two] - -- rw [_root_.mul_pos_iff] - -- left - -- simp + · apply sq_le_sq.mpr + clear L h2 h1 α + rw [<- Int.dist_eq] + rw [Int.dist_eq'] + have X : ((query l₁ - query l₂).natAbs : ℝ) ≤ (Δ.val : ℝ) := by + apply Nat.cast_le.mpr + apply bounded_sensitivity + rw [abs_eq_natAbs] + simp only [abs_cast] + exact X + · apply mul_nonneg + · apply sq_nonneg + · apply sq_nonneg + · apply mul_nonneg + · apply mul_nonneg + · linarith + · apply sq_nonneg + · apply inv_nonneg.mpr + apply sq_nonneg -/-- +/- All outputs of the zCDP mechanism have nonzero probability. -/ -theorem privNoisedQuery_NonZeroNQ (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : - NonZeroNQ (privNoisedQuery query Δ ε₁ ε₂) := by - simp [NonZeroNQ, privNoisedQuery, 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 privNoisedQuery_NonZeroNQ (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : +-- NonZeroNQ (privNoisedQuery query Δ ε₁ ε₂) := by +-- simp [NonZeroNQ, privNoisedQuery, 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 -/-- +/- All outputs of the zCDP mechanism have finite probability. -/ -theorem privNoisedQuery_NonTopNQ (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : - NonTopNQ (privNoisedQuery query Δ ε₁ ε₂) := by - simp [NonTopNQ, privNoisedQuery, 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) (n.instDecidableEq (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 privNoisedQuery_NonTopNQ (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : +-- NonTopNQ (privNoisedQuery query Δ ε₁ ε₂) := by +-- simp [NonTopNQ, privNoisedQuery, 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) (n.instDecidableEq (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 lemma discrete_gaussian_shift {σ : ℝ} (h : σ ≠ 0) (μ : ℝ) (τ x : ℤ) : discrete_gaussian σ μ (x - τ) = discrete_gaussian σ (μ + τ) (x) := by @@ -221,143 +247,143 @@ lemma discrete_gaussian_shift {σ : ℝ} (h : σ ≠ 0) (μ : ℝ) (τ x : ℤ) ring_nf . rw [shifted_gauss_sum h] -/-- +/- The zCDP mechanism is normalizable. -/ -theorem privNoisedQuery_NonTopSum (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : - NonTopSum (privNoisedQuery query Δ ε₁ ε₂) := by - simp [NonTopSum, privNoisedQuery, 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) (n.instDecidableEq (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 privNoisedQuery_NonTopSum (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : +-- NonTopSum (privNoisedQuery query Δ ε₁ ε₂) := by +-- simp [NonTopSum, privNoisedQuery, 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) (n.instDecidableEq (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) -/-- +/- The Renyi divergence of the zCDP mechanism is finite on neighbouring inputs. -/ -theorem privNoisedQuery_NonTopRDNQ (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : - NonTopRDNQ (privNoisedQuery query Δ ε₁ ε₂) := by - simp [NonTopRDNQ, privNoisedQuery, 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₁) (x.instDecidableEq (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₂) (x.instDecidableEq (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 := @Renyi_sum_SG_nonneg _ α P (query l₁) (query l₂) n - rw [discrete_gaussian_shift P] - rw [discrete_gaussian_shift P] - simp [X] - . have X := @Renyi_Gauss_summable _ P (query l₁) (query l₂) - conv => - right - intro x - rw [discrete_gaussian_shift P] - rw [discrete_gaussian_shift P] - simp [X] +-- theorem privNoisedQuery_NonTopRDNQ (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : +-- NonTopRDNQ (privNoisedQuery query Δ ε₁ ε₂) := by +-- simp [NonTopRDNQ, privNoisedQuery, 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₁) (x.instDecidableEq (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₂) (x.instDecidableEq (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 := @Renyi_sum_SG_nonneg _ α P (query l₁) (query l₂) n +-- rw [discrete_gaussian_shift P] +-- rw [discrete_gaussian_shift P] +-- simp [X] +-- . have X := @Renyi_Gauss_summable _ P (query l₁) (query l₂) +-- conv => +-- right +-- intro x +-- rw [discrete_gaussian_shift P] +-- rw [discrete_gaussian_shift P] +-- simp [X] /-- The zCDP mechanism is ``(Δε₂/ε₁)^2``-zCDP. From f8c66d599ca257f63768be3f5082cddebd0637ca Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Fri, 21 Jun 2024 13:53:20 -0400 Subject: [PATCH 080/122] checkpoint: before changing postprocessing --- .../ZeroConcentrated/Postprocessing.lean | 88 +++++++++++++++---- SampCert/Util/Log.lean | 18 +--- 2 files changed, 73 insertions(+), 33 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean index d5aad4e8..8e7819d4 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean @@ -321,7 +321,8 @@ theorem DPostPocess_pre {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) . rw [← RenyiDivergenceExpectation _ _ h1] . replace nt := nt α h1 l₁ l₂ h2 apply convergent_subset _ nt - . intro x + . -- Need absolute continuity between neighbours + intro x sorry -- apply nn -- . intro x @@ -554,6 +555,7 @@ theorem tsum_ne_zero_of_ne_zero {T : Type} [Inhabited T] (f : T → ENNReal) (h have B := CONTRA default contradiction +set_option pp.coercions false theorem privPostProcess_zCDPBound {nq : List T → SLang U} {HNorm : NormalMechanism nq} {ε₁ ε₂ : ℕ+} (h : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (conv : NonTopSum nq) @@ -566,13 +568,67 @@ theorem privPostProcess_zCDPBound {nq : List T → SLang U} {HNorm : NormalMecha replace h' := h' α h1 l₁ l₂ h2 -- Part 1, removing fluff - apply le_trans _ h' clear h' + unfold RenyiDivergence_def + rw [DFunLike.coe] + rw [PMF.instFunLike] + simp + -- Clean up coercions + conv => + lhs + arg 1 + arg 2 + arg 1 + arg 1 + intro x + repeat rw [SLang.toPMF] + simp + conv => + rhs + arg 1 + arg 2 + arg 1 + arg 1 + intro x + repeat rw [SLang.toPMF] + repeat rw [DFunLike.coe] + repeat rw [PMF.instFunLike] + simp + -- Split multiplication in ofReals + rw [ofEReal_mul_nonneg ?G1 ?G2] + case G1 => + sorry + case G2 => + -- Use normalization for PostProcessing for the arguments (might have to move around + -- the coercion cleanup step) + #check RenyiDivergence_def_log_sum_nonneg + sorry + rw [ofEReal_mul_nonneg ?G1 ?G2] + case G1 => + sorry + case G2 => + sorry + simp + + -- Remove the α scaling + apply mul_le_mul_of_nonneg_left _ ?G1 + case G1 => simp + + apply (ofEReal_le_mono_nonneg ?G1).mp + case G1 => sorry + + #check DPostPocess_pre h + + -- have B := DPostPocess_pre h nn nt nts conv f h1 h2 + -- have B' : ∑' (x : V), (∑' (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 sorry - /- -- remove the α scaling have A : 0 ≤ (α - 1)⁻¹ := by @@ -651,19 +707,19 @@ theorem privPostProcess_zCDPBound {nq : List T → SLang U} {HNorm : NormalMecha sorry -/ -theorem privPostProcess_NonTopRDNQ {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) 1} {ε₁ ε₂ : ℕ+} (f : U → V) - (dp : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) (nt : NonTopRDNQ nq) (nz : NonZeroNQ nq) (nts : NonTopNQ nq) (ntsum: NonTopSum nq) : - NonTopRDNQ (privPostProcess nq f) := by - simp [NonTopRDNQ, NonTopSum, privPostProcess] at * - intros α h1 l₁ l₂ h2 - have ntrdnq := nt - replace nt := nt α h1 l₁ l₂ h2 - sorry - -- have A := @DPostPocess_pre T U V _ _ _ nq ε₁ ε₂ dp nz ntrdnq nts ntsum f α h1 l₁ l₂ h2 - -- apply @LT.lt.ne_top _ _ _ _ ⊤ - -- rw [Eq.comm] at nt - -- have B := Ne.lt_top' nt - -- exact lt_of_le_of_lt A B +-- theorem privPostProcess_NonTopRDNQ {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) 1} {ε₁ ε₂ : ℕ+} (f : U → V) +-- (dp : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) (nt : NonTopRDNQ nq) (nz : NonZeroNQ nq) (nts : NonTopNQ nq) (ntsum: NonTopSum nq) : +-- NonTopRDNQ (privPostProcess nq f) := by +-- simp [NonTopRDNQ, NonTopSum, privPostProcess] at * +-- intros α h1 l₁ l₂ h2 +-- have ntrdnq := nt +-- replace nt := nt α h1 l₁ l₂ h2 +-- sorry +-- -- have A := @DPostPocess_pre T U V _ _ _ nq ε₁ ε₂ dp nz ntrdnq nts ntsum f α h1 l₁ l₂ h2 +-- -- apply @LT.lt.ne_top _ _ _ _ ⊤ +-- -- rw [Eq.comm] at nt +-- -- have B := Ne.lt_top' nt +-- -- exact lt_of_le_of_lt A B /-- Postprocessing preserves zCDP diff --git a/SampCert/Util/Log.lean b/SampCert/Util/Log.lean index 62f0858b..6cfb90dd 100644 --- a/SampCert/Util/Log.lean +++ b/SampCert/Util/Log.lean @@ -225,11 +225,9 @@ lemma elog_mono_le : (x <= y) <-> elog x <= elog y := by sorry lemma ofEReal_nonneg_eq_iff (Hw : 0 <= w) (Hz : 0 <= z) : w = z <-> (ofEReal w = ofEReal z) := sorry -lemma ofEReal_le_mono : (0 ≤ w) -> w ≤ z <-> (ofEReal w ≤ ofEReal z) := +lemma ofEReal_le_mono_nonneg (Hw : 0 ≤ w) : w ≤ z <-> (ofEReal w ≤ ofEReal z) := sorry -#check 3 - @[simp] lemma ofEReal_plus_nonneg (Hw : 0 ≤ w) (Hz : 0 ≤ z) : ofEReal (w + z) = ofEReal w + ofEReal z := by rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> @@ -265,20 +263,6 @@ lemma ofEReal_nonneg_scal_l (H1 : 0 < r) (H2 : 0 ≤ r * w) : 0 ≤ w := by sorry - - -- <;> rcases (EReal_cases z) with ⟨ Hz , Hz , ⟨ z' , Hz ⟩⟩ - - - - - -- cases w <;> cases z - -- · exfalso - -- rw [WithBot.none_eq_bot] at Hw - -- simp at * - -- · sorry - -- · sorry - -- · sorry - @[simp] lemma toEReal_ofENNReal_nonneg (H : 0 ≤ w) : ENNReal.toEReal (ofEReal w) = w := sorry From 2825ee0b9009dc87e390536870cc818cd61c4eee Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Fri, 21 Jun 2024 15:45:21 -0400 Subject: [PATCH 081/122] High-level sketch for postprocessing reduction --- .../ZeroConcentrated/Postprocessing.lean | 101 +++++++----------- SampCert/Util/Log.lean | 14 +++ 2 files changed, 54 insertions(+), 61 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean index 8e7819d4..eac014d7 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean @@ -57,7 +57,6 @@ Jensen's ineuquality for the exponential applied to Renyi's sum -/ 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 => left left @@ -219,9 +218,10 @@ theorem tsum_pos_int {f : ℤ → ENNReal} (h1 : ∑' x : ℤ, f x ≠ ⊤) (h2 -theorem DPostPocess_pre {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) 1} {ε₁ ε₂ : ℕ+} - (h : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) - (conv : NonTopSum nq) (f : U → V) {α : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (Habs : AbsCts (nq l₁) (nq l₂)) (h2 : Neighbour l₁ l₂) : +theorem DPostPocess_pre_reduct {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) 1} {ε₁ ε₂ : ℕ+} + (h : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) + (nn : ∀(l : List T), ∀(u : U), nq l u ≠ 0) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (conv : NonTopSum nq) + (f : U → V) {α : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (Habs : AbsCts (nq l₁) (nq l₂)) (h2 : Neighbour l₁ l₂) : (∑' (x : V), (∑' (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 - α)) ≤ @@ -555,27 +555,46 @@ theorem tsum_ne_zero_of_ne_zero {T : Type} [Inhabited T] (f : T → ENNReal) (h have B := CONTRA default contradiction -set_option pp.coercions false + + + + + + + +theorem DPostPocess_pre {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) 1} {ε₁ ε₂ : ℕ+} + (h : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) + (f : U → V) {α : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (Habs : AbsCts (nq l₁) (nq l₂)) (h2 : Neighbour l₁ l₂) : + (∑' (x : V), + (∑' (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 + + -- First step is to reduce to the case where (nq l) is nonzero + -- Next I need to generalize the fiberwisation argument to eliminate the side conditions and work in + -- the extended reals + -- I'll also have to generalize the δpmf normalization, though that should mstly just be making sure the correct terms are PMFs + -- That should bring be down to the extended Renyi divergence + + -- This lemma should be provable as stated + sorry + + theorem privPostProcess_zCDPBound {nq : List T → SLang U} {HNorm : NormalMechanism nq} {ε₁ ε₂ : ℕ+} - (h : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) (nn : NonZeroNQ nq) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (conv : NonTopSum nq) - (f : U → V) : + (h : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) (f : U → V) : zCDPBound (privPostProcess nq f) (privPostProcess_norm nq HNorm f) ((ε₁ : ℝ) / ε₂) := by simp [privPostProcess, zCDPBound, RenyiDivergence] intro α h1 l₁ l₂ h2 have h' := h simp [zCDPBound, RenyiDivergence] at h' replace h' := h' α h1 l₁ l₂ h2 - - -- Part 1, removing fluff apply le_trans _ h' clear h' unfold RenyiDivergence_def rw [DFunLike.coe] rw [PMF.instFunLike] simp - - -- Clean up coercions conv => lhs arg 1 @@ -596,59 +615,19 @@ theorem privPostProcess_zCDPBound {nq : List T → SLang U} {HNorm : NormalMecha repeat rw [DFunLike.coe] repeat rw [PMF.instFunLike] simp - - -- Split multiplication in ofReals - rw [ofEReal_mul_nonneg ?G1 ?G2] - case G1 => - sorry - case G2 => - -- Use normalization for PostProcessing for the arguments (might have to move around - -- the coercion cleanup step) - #check RenyiDivergence_def_log_sum_nonneg + apply ofEReal_le_mono + apply ereal_smul_left_le + · apply EReal.coe_pos.mpr + apply inv_pos_of_pos + linarith + · exact EReal.coe_lt_top (α - 1)⁻¹ + apply elog_mono_le.mp + apply (DPostPocess_pre h f h1 ?Gabs h2) + case Gabs => + -- absolute continuity sorry - rw [ofEReal_mul_nonneg ?G1 ?G2] - case G1 => - sorry - case G2 => - sorry - simp - - -- Remove the α scaling - apply mul_le_mul_of_nonneg_left _ ?G1 - case G1 => simp - apply (ofEReal_le_mono_nonneg ?G1).mp - case G1 => sorry - - #check DPostPocess_pre h - - -- have B := DPostPocess_pre h nn nt nts conv f h1 h2 - -- have B' : ∑' (x : V), (∑' (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 - sorry /- - -- 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 : V), (∑' (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' diff --git a/SampCert/Util/Log.lean b/SampCert/Util/Log.lean index 6cfb90dd..78a61c77 100644 --- a/SampCert/Util/Log.lean +++ b/SampCert/Util/Log.lean @@ -228,6 +228,16 @@ lemma ofEReal_nonneg_eq_iff (Hw : 0 <= w) (Hz : 0 <= z) : w = z <-> (ofEReal w = lemma ofEReal_le_mono_nonneg (Hw : 0 ≤ w) : w ≤ z <-> (ofEReal w ≤ ofEReal z) := sorry + +lemma ofEReal_le_mono (H : w ≤ z) : ofEReal w ≤ ofEReal z := by + rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> + rcases (EReal_cases z) with Hz' | (Hz' | ⟨ z', Hz' ⟩) + all_goals rw [Hw', Hz'] + all_goals simp_all [ENNReal.ofEReal] + simp [Real.toEReal] + exact ofReal_le_ofReal H + + @[simp] lemma ofEReal_plus_nonneg (Hw : 0 ≤ w) (Hz : 0 ≤ z) : ofEReal (w + z) = ofEReal w + ofEReal z := by rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> @@ -407,6 +417,10 @@ lemma ereal_smul_eq_left (s : EReal) (Hr1 : 0 < s) (Hr2 : s < ⊤) (H : s * w = sorry +lemma ereal_smul_left_le (s : EReal) (Hr1 : 0 < s) (Hr2 : s < ⊤) (H : w ≤ z) : s * w ≤ s * z := by + sorry + + lemma ENNReal_toReal_partial_inj (a b : ENNReal) (H : a.toReal = b.toReal) (H1 : a ≠ ⊤) (H2 : b ≠ ⊤) : a = b := sorry end ENNReal From fbecc40ebbb3d1b24b434442dc5627f22faef01d Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Fri, 21 Jun 2024 17:02:43 -0400 Subject: [PATCH 082/122] more log lemmas --- .../ZeroConcentrated/Postprocessing.lean | 6 --- SampCert/Util/Log.lean | 48 +++++++++++++++++-- 2 files changed, 43 insertions(+), 11 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean index eac014d7..8d9574fb 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean @@ -556,12 +556,6 @@ theorem tsum_ne_zero_of_ne_zero {T : Type} [Inhabited T] (f : T → ENNReal) (h contradiction - - - - - - theorem DPostPocess_pre {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) 1} {ε₁ ε₂ : ℕ+} (h : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) (f : U → V) {α : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (Habs : AbsCts (nq l₁) (nq l₂)) (h2 : Neighbour l₁ l₂) : diff --git a/SampCert/Util/Log.lean b/SampCert/Util/Log.lean index 78a61c77..1316e6eb 100644 --- a/SampCert/Util/Log.lean +++ b/SampCert/Util/Log.lean @@ -414,13 +414,51 @@ lemma ereal_smul_le_left (s : EReal) (Hr1 : 0 < s) (Hr2 : s < ⊤) (H : s * w exact EReal.coe_pos.mp Hr1 lemma ereal_smul_eq_left (s : EReal) (Hr1 : 0 < s) (Hr2 : s < ⊤) (H : s * w = s * z) : w = z := by - sorry - + rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> + rcases (EReal_cases z) with Hz' | (Hz' | ⟨ z', Hz' ⟩) <;> + rcases (EReal_cases s) with Hs' | (Hs' | ⟨ s', Hs' ⟩) + all_goals rw [Hw', Hz'] + all_goals simp_all + · rw [mul_top_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H + rw [mul_bot_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H + cases H + · rw [mul_bot_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H + cases H + · rw [mul_top_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H + rw [mul_bot_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H + cases H + · rw [mul_top_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H + cases H + · rw [mul_bot_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H + repeat rw [<- EReal.coe_mul] at H + cases H + · rw [mul_top_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H + repeat rw [<- EReal.coe_mul] at H + cases H + · repeat rw [<- EReal.coe_mul] at H + apply EReal.coe_eq_coe_iff.mp at H + sorry lemma ereal_smul_left_le (s : EReal) (Hr1 : 0 < s) (Hr2 : s < ⊤) (H : w ≤ z) : s * w ≤ s * z := by - sorry - + rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> + rcases (EReal_cases z) with Hz' | (Hz' | ⟨ z', Hz' ⟩) + all_goals rw [Hw', Hz'] + all_goals simp_all + · rw [mul_top_of_pos Hr1] + exact OrderTop.le_top (s * ⊥) + · rw [mul_bot_of_pos Hr1] + exact OrderBot.bot_le (s * z'.toEReal) + · rw [mul_top_of_pos Hr1] + exact OrderTop.le_top (s * w'.toEReal) + rcases (EReal_cases s) with Hs' | (Hs' | ⟨ s', Hs' ⟩) + all_goals simp_all + repeat rw [← EReal.coe_mul] + apply EReal.coe_le_coe_iff.mpr + exact (mul_le_mul_iff_of_pos_left Hr1).mpr H -lemma ENNReal_toReal_partial_inj (a b : ENNReal) (H : a.toReal = b.toReal) (H1 : a ≠ ⊤) (H2 : b ≠ ⊤) : a = b := sorry +lemma ENNReal_toReal_partial_inj (a b : ENNReal) (H : a.toReal = b.toReal) (H1 : a ≠ ⊤) (H2 : b ≠ ⊤) : a = b := by + rcases a with Ha' | ⟨ a' | Ha' ⟩ <;> + rcases b with Hb' | ⟨ b' | Hb' ⟩ + all_goals simp_all end ENNReal From b1bda804360c42ec0a9cd83e7aa81c671f164ebb Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Mon, 24 Jun 2024 10:02:03 -0400 Subject: [PATCH 083/122] privCompose normalizes --- SampCert/DifferentialPrivacy/Abstract.lean | 22 ++++++++++++++++++++-- 1 file changed, 20 insertions(+), 2 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Abstract.lean b/SampCert/DifferentialPrivacy/Abstract.lean index f62f63d0..e5cebb88 100644 --- a/SampCert/DifferentialPrivacy/Abstract.lean +++ b/SampCert/DifferentialPrivacy/Abstract.lean @@ -312,13 +312,31 @@ theorem privPostProcess_NonTopSum {nq : List T → SLang U} (f : U → V) (nt : rw [← A] trivial --- Provable using stuff in the adaptive composition PR, wait merge that stuff later lemma privCompose_norm (nq1 : List T → SLang U) (nq2 : List T → SLang V) (HNorm1 : NormalMechanism nq1) (HNorm2 : NormalMechanism nq2) : NormalMechanism (privCompose nq1 nq2) := by rw [NormalMechanism] at * intro l have HNorm1' := HasSum.tsum_eq (HNorm1 l) have HNorm2' := HasSum.tsum_eq (HNorm2 l) - have HR : (∑' (x : U × V), privCompose nq1 nq2 l x = 1) := sorry + have HR : (∑' (x : U × V), privCompose nq1 nq2 l x = 1) := by + rw [privCompose] + rw [ENNReal.tsum_prod'] + simp + conv => + lhs + arg 1 + intro a + arg 1 + intro b + rw [compose_sum_rw nq1 (fun l a_1 => nq2 l a_1) a b l] + conv => + lhs + arg 1 + intro a + rw [ENNReal.tsum_mul_left] + rw [ENNReal.tsum_mul_right] + rw [HNorm1'] + rw [HNorm2'] + simp rw [<- HR] apply Summable.hasSum exact ENNReal.summable From b23d6f1ce894cbf3d7dd019087cc4a539d0cf770 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Mon, 24 Jun 2024 10:36:39 -0400 Subject: [PATCH 084/122] privPostProcess normalizes --- SampCert/DifferentialPrivacy/Abstract.lean | 17 ++++++++++++++--- 1 file changed, 14 insertions(+), 3 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Abstract.lean b/SampCert/DifferentialPrivacy/Abstract.lean index e5cebb88..e21e4c82 100644 --- a/SampCert/DifferentialPrivacy/Abstract.lean +++ b/SampCert/DifferentialPrivacy/Abstract.lean @@ -341,11 +341,22 @@ lemma privCompose_norm (nq1 : List T → SLang U) (nq2 : List T → SLang V) (HN apply Summable.hasSum exact ENNReal.summable -lemma privPostProcess_norm (nq : List T → SLang U) (HNorm : NormalMechanism nq) (f : U -> V) : NormalMechanism (privPostProcess nq f) := by +lemma privPostProcess_norm (nq : List T → SLang U) (HNorm : NormalMechanism nq) (f : U -> V) : NormalMechanism (privPostProcess nq f) := by rw [NormalMechanism] at * intro l - have HNorm' := HasSum.tsum_eq (HNorm l) - have HR : (∑' (b : V), privPostProcess nq f l b = 1) := sorry + have HR : (∑' (b : V), privPostProcess nq f l b = 1) := by + rw [privPostProcess] + simp + rw [<- HasSum.tsum_eq (ENNReal.HasSum_fiberwise (HNorm l) f)] + apply tsum_congr + intro v + rw [condition_to_subset] + have Htyeq : ({a | v = f a}) = (f ⁻¹' {v}) := by + rw [Set.ext_iff] + intro x + simp + exact eq_comm + rw [Htyeq] rw [<- HR] apply Summable.hasSum exact ENNReal.summable From abae0a542fc977e80efac3734a165f9b1867ba92 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Mon, 24 Jun 2024 10:51:39 -0400 Subject: [PATCH 085/122] Add absolute continuity side condition --- .../ZeroConcentrated/Composition.lean | 18 +++++++++--------- .../ZeroConcentrated/DP.lean | 12 +++++++++++- .../ZeroConcentrated/Mechanism/Properties.lean | 12 +++++++++--- 3 files changed, 29 insertions(+), 13 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean index 005efa52..f824669c 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean @@ -80,12 +80,10 @@ The Renyi Divergence between neighbouring inputs of noised queries is nonzero. Composed queries satisfy zCDP Renyi divergence bound. -/ theorem privCompose_zCDPBound {nq1 : List T → SLang U} {nq2 : List T → SLang V} {HNorm1 : ∀ l, HasSum (nq1 l) 1} {HNorm2 : ∀ l, HasSum (nq2 l) 1} {ε₁ ε₂ ε₃ ε₄ : ℕ+} - (h1 : zCDPBound nq1 HNorm1 ((ε₁ : ℝ) / ε₂)) (h2 : zCDPBound nq2 HNorm2 ((ε₃ : ℝ) / ε₄)) : + (h1 : zCDPBound nq1 HNorm1 ((ε₁ : ℝ) / ε₂)) (h2 : zCDPBound nq2 HNorm2 ((ε₃ : ℝ) / ε₄)) (Ha1 : ACNeighbour nq1) (Ha2 : ACNeighbour nq2): zCDPBound (privCompose nq1 nq2) (privCompose_norm nq1 nq2 HNorm1 HNorm2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by simp [privCompose, RenyiDivergence, zCDPBound] intro α h3 l₁ l₂ h4 - have X := h1 - have Y := h2 simp [zCDPBound] at h1 h2 replace h1 := h1 α h3 l₁ l₂ h4 replace h2 := h2 α h3 l₁ l₂ h4 @@ -185,7 +183,7 @@ theorem privCompose_zCDPBound {nq1 : List T → SLang U} {nq2 : List T → SLang simp have log_nonneg_1 : 0 ≤ (∑' (i : U), nq1 l₁ i ^ α * nq1 l₂ i ^ (1 - α)).elog := by - have Hac1 : AbsCts ((nq1 l₁).toPMF (HNorm1 l₁)) ((nq1 l₂).toPMF (HNorm1 l₂)) := by sorry + have Hac1 : AbsCts ((nq1 l₁).toPMF (HNorm1 l₁)) ((nq1 l₂).toPMF (HNorm1 l₂)) := by exact Ha1 l₁ l₂ h4 have Hnn1 := (RenyiDivergence_def_log_sum_nonneg ((nq1 l₁).toPMF (HNorm1 l₁)) ((nq1 l₂).toPMF (HNorm1 l₂)) Hac1 h3) conv at Hnn1 => rhs @@ -199,7 +197,7 @@ theorem privCompose_zCDPBound {nq1 : List T → SLang U} {nq2 : List T → SLang simp apply Hnn1 have log_nonneg_2 : 0 ≤ (∑' (i : V), nq2 l₁ i ^ α * nq2 l₂ i ^ (1 - α)).elog := by - have Hac2 : AbsCts ((nq2 l₁).toPMF (HNorm2 l₁)) ((nq2 l₂).toPMF (HNorm2 l₂)) := by sorry + have Hac2 : AbsCts ((nq2 l₁).toPMF (HNorm2 l₁)) ((nq2 l₂).toPMF (HNorm2 l₂)) := by exact Ha2 l₁ l₂ h4 have Hnn2 := (RenyiDivergence_def_log_sum_nonneg ((nq2 l₁).toPMF (HNorm2 l₁)) ((nq2 l₂).toPMF (HNorm2 l₂)) Hac2 h3) conv at Hnn2 => rhs @@ -361,9 +359,11 @@ Renyi divergence beteeen composed queries on neighbours are finite. theorem privCompose_zCDP (nq1 : List T → SLang U) (nq2 : List T → SLang V) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h : zCDP nq1 ((ε₁ : ℝ) / ε₂)) (h' : zCDP nq2 ((ε₃ : ℝ) / ε₄)) : zCDP (privCompose nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by simp [zCDP] at * - rcases h with ⟨ Hn1, Hb1 ⟩ - rcases h' with ⟨ Hn2, Hb2 ⟩ - exists (privCompose_norm nq1 nq2 Hn1 Hn2) - exact privCompose_zCDPBound Hb1 Hb2 + rcases h with ⟨ Hac1, ⟨ Hn1, Hb1 ⟩⟩ + rcases h' with ⟨ Hac2, ⟨ Hn2, Hb2 ⟩⟩ + apply And.intro + · sorry + · exists (privCompose_norm nq1 nq2 Hn1 Hn2) + exact privCompose_zCDPBound Hb1 Hb2 Hac1 Hac2 end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean index 4758f1de..c412f328 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean @@ -45,10 +45,20 @@ 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 - α) ≠ ⊤ +def ACNeighbour (p : List T -> SLang U) : Prop := ∀ l₁ l₂ , Neighbour l₁ l₂ -> AbsCts (p l₁) (p l₂) + + + /-- The mechanism ``q`` is ``(ε^2)/2``-zCDP -/ -def zCDP (q : List T → SLang U) (ε : ℝ) : Prop := ∃ (HN : NormalMechanism q), zCDPBound q HN ε +def zCDP (q : List T → SLang U) (ε : ℝ) : Prop := + ACNeighbour q + -- FIXME: Instead of a side condition, make normalization intrinsic to the Mechanism type. + ∧ ∃ (HN : NormalMechanism q), zCDPBound q HN ε + + + -- ∧ NonZeroNQ q -- ∧ NonTopSum q -- ∧ NonTopNQ q diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean index 5bbd126c..94dd6712 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean @@ -385,15 +385,21 @@ The Renyi divergence of the zCDP mechanism is finite on neighbouring inputs. -- rw [discrete_gaussian_shift P] -- simp [X] +def privNoisedQuery_AC (query : List T -> ℤ) (Δ ε₁ ε₂ : ℕ+) : ACNeighbour (privNoisedQuery query Δ ε₁ ε₂) := by + rw [ACNeighbour] + sorry + /-- The zCDP mechanism is ``(Δε₂/ε₁)^2``-zCDP. -/ theorem privNoisedQuery_zCDP (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_sensitivity : sensitivity query Δ) : zCDP (privNoisedQuery query Δ ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by simp [zCDP] - repeat any_goals constructor - . apply privNoisedQuery_zCDPBound - exact bounded_sensitivity + apply And.intro + · sorry + · repeat any_goals constructor + . apply privNoisedQuery_zCDPBound + exact bounded_sensitivity -- . apply privNoisedQuery_NonZeroNQ -- . apply privNoisedQuery_NonTopSum -- . apply privNoisedQuery_NonTopNQ From 98b15fd08c42b56f2b8cfae5d047a7b9aa798723 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Mon, 24 Jun 2024 10:56:26 -0400 Subject: [PATCH 086/122] Mechanism absolute continuity --- .../Mechanism/Properties.lean | 32 +++++++------------ 1 file changed, 12 insertions(+), 20 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean index 94dd6712..63265cb0 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean @@ -195,20 +195,6 @@ theorem privNoisedQuery_zCDPBound (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ · apply inv_nonneg.mpr apply sq_nonneg -/- -All outputs of the zCDP mechanism have nonzero probability. --/ --- theorem privNoisedQuery_NonZeroNQ (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : --- NonZeroNQ (privNoisedQuery query Δ ε₁ ε₂) := by --- simp [NonZeroNQ, privNoisedQuery, 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 /- All outputs of the zCDP mechanism have finite probability. @@ -387,7 +373,17 @@ The Renyi divergence of the zCDP mechanism is finite on neighbouring inputs. def privNoisedQuery_AC (query : List T -> ℤ) (Δ ε₁ ε₂ : ℕ+) : ACNeighbour (privNoisedQuery query Δ ε₁ ε₂) := by rw [ACNeighbour] - sorry + intro l₁ l₂ _ + rw [AbsCts] + intro n Hk + exfalso + simp [privNoisedQuery, DiscreteGaussianGenSample] at Hk + have Hk := Hk (n - query l₂) + simp at Hk + have A : ((Δ : ℝ) * ε₂ / ε₁) ≠ 0 := by simp + have X := @discrete_gaussian_pos (↑↑Δ * ↑↑ε₂ / ↑↑ε₁) A 0 (n - query l₂) + simp at X + linarith /-- The zCDP mechanism is ``(Δε₂/ε₁)^2``-zCDP. @@ -396,14 +392,10 @@ theorem privNoisedQuery_zCDP (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (b zCDP (privNoisedQuery query Δ ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by simp [zCDP] apply And.intro - · sorry + · exact privNoisedQuery_AC query Δ ε₁ ε₂ · repeat any_goals constructor . apply privNoisedQuery_zCDPBound exact bounded_sensitivity - -- . apply privNoisedQuery_NonZeroNQ - -- . apply privNoisedQuery_NonTopSum - -- . apply privNoisedQuery_NonTopNQ - -- . apply privNoisedQuery_NonTopRDNQ end SLang From 8057170c3e26d9836e598ffd56b310edba3e4a4d Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Mon, 24 Jun 2024 11:08:00 -0400 Subject: [PATCH 087/122] privCompose absolute continuity --- .../ZeroConcentrated/Composition.lean | 27 ++++++++++++++++++- 1 file changed, 26 insertions(+), 1 deletion(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean index f824669c..013a9867 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean @@ -353,6 +353,31 @@ Renyi divergence beteeen composed queries on neighbours are finite. -- rename_i h4 h5 -- contradiction + +def privCompose_AC (nq1 : List T → SLang U) (nq2 : List T → SLang V) (Hac1 : ACNeighbour nq1) (Hac2 : ACNeighbour nq2) : ACNeighbour (privCompose nq1 nq2) := by + rw [ACNeighbour] at * + intro l₁ l₂ Hn + have Hac1 := Hac1 l₁ l₂ Hn + have Hac2 := Hac2 l₁ l₂ Hn + rw [AbsCts] at * + intro x + rcases x with ⟨ u, v ⟩ + intro Hk + simp [privCompose] at * + intro i + cases (Hk i) + · rename_i Hk + left + apply Hac1 + apply Hk + · rename_i Hk + right + intro H + apply Hac2 + apply Hk + apply H + + /-- ``privCompose`` satisfies zCDP -/ @@ -362,7 +387,7 @@ theorem privCompose_zCDP (nq1 : List T → SLang U) (nq2 : List T → SLang V) ( rcases h with ⟨ Hac1, ⟨ Hn1, Hb1 ⟩⟩ rcases h' with ⟨ Hac2, ⟨ Hn2, Hb2 ⟩⟩ apply And.intro - · sorry + · exact privCompose_AC nq1 nq2 Hac1 Hac2 · exists (privCompose_norm nq1 nq2 Hn1 Hn2) exact privCompose_zCDPBound Hb1 Hb2 Hac1 Hac2 From e386b98fa1175b323483fe7b06f5e0772c4bc218 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Mon, 24 Jun 2024 11:24:44 -0400 Subject: [PATCH 088/122] RenyiDivergence_aux_zero --- .../DifferentialPrivacy/RenyiDivergence.lean | 19 ++++++++++++++++--- SampCert/Util/Log.lean | 7 +++++++ 2 files changed, 23 insertions(+), 3 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index ca292724..ebabfe2f 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -1147,10 +1147,23 @@ The Renyi divergence. noncomputable def RenyiDivergence (p q : PMF T) (α : ℝ) : ENNReal := ENNReal.ofEReal (RenyiDivergence_def p q α) --- MARKUSDE: We evaluate through the ENNReal.ofEReal using ``RenyiDivergence_def_nonneg``, these are some special cases -theorem RenyiDivergence_aux_zero (p q : PMF T) (α : ℝ) : p = q <-> RenyiDivergence p q α = 0 - := sorry +-- Lifted property of RenyiDivergence_def +theorem RenyiDivergence_aux_zero [MeasurableSpace T] [MeasurableSingletonClass T] [Countable T] + (p q : PMF T) {α : ℝ} (Hα : 1 < α) (Hac : AbsCts p q) : p = q <-> RenyiDivergence p q α = 0 := by + apply Iff.intro + · intro Heq + rw [Heq] + rw [RenyiDivergence] + rw [<- RenyiDivergence_refl_zero _ Hα] + simp + · intro H + apply (RenyiDivergence_def_eq_0_iff p q Hα Hac).mp + symm + apply (ofEReal_nonneg_zero ?G1).mpr + case G1 => + exact RenyiDivergence_def_nonneg p q Hac Hα + exact id (Eq.symm H) -- Unused /- diff --git a/SampCert/Util/Log.lean b/SampCert/Util/Log.lean index 1316e6eb..5880e934 100644 --- a/SampCert/Util/Log.lean +++ b/SampCert/Util/Log.lean @@ -51,6 +51,10 @@ lemma ofEReal_bot : ofEReal ⊥ = 0 := by simp [ofEReal] @[simp] lemma ofEReal_top : ofEReal ⊤ = ⊤ := by simp [ofEReal] +@[simp] +lemma ofEReal_zero : ofEReal 0 = 0 := by simp [ofEReal] + + @[simp] lemma ofEReal_real (r : ℝ) : ofEReal r = ENNReal.ofReal r := by simp [Real.toEReal] @@ -228,6 +232,9 @@ lemma ofEReal_nonneg_eq_iff (Hw : 0 <= w) (Hz : 0 <= z) : w = z <-> (ofEReal w = lemma ofEReal_le_mono_nonneg (Hw : 0 ≤ w) : w ≤ z <-> (ofEReal w ≤ ofEReal z) := sorry +-- Use above +lemma ofEReal_nonneg_zero (Hz : 0 ≤ z) : (0 = z) <-> 0 = ofEReal z := sorry + lemma ofEReal_le_mono (H : w ≤ z) : ofEReal w ≤ ofEReal z := by rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> From c16857e2b9c2b00b1db2c4e4acd31f74a14a2980 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Mon, 24 Jun 2024 11:56:05 -0400 Subject: [PATCH 089/122] privPostProccess absolute continuity (removes surjectivity requirement?) --- .../ZeroConcentrated/Postprocessing.lean | 32 ++++++++++++------- .../ZeroConcentrated/System.lean | 2 +- 2 files changed, 21 insertions(+), 13 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean index 8d9574fb..4d9167cc 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean @@ -694,23 +694,31 @@ theorem privPostProcess_zCDPBound {nq : List T → SLang U} {HNorm : NormalMecha -- -- have B := Ne.lt_top' nt -- -- exact lt_of_le_of_lt A B + + + +def privPostProcess_AC {f : U -> V} (nq : List T → SLang U) (Hac : ACNeighbour nq) : ACNeighbour (privPostProcess nq f) := by + rw [ACNeighbour] at * + unfold AbsCts at * + intro l₁ l₂ Hn v + have Hac := Hac l₁ l₂ Hn + simp [privPostProcess] + intro Hppz i fi + apply Hac + apply Hppz + apply fi + /-- Postprocessing preserves zCDP -/ -theorem privPostProcess_zCDP {f : U → V} (sur : Function.Surjective f) +theorem privPostProcess_zCDP {f : U → V} (nq : List T → SLang U) (ε₁ ε₂ : ℕ+) (h : zCDP nq ((ε₁ : ℝ) / ε₂)) : zCDP (privPostProcess nq f) (((ε₁ : ℝ) / ε₂)) := by + rcases h with ⟨ Hac1, ⟨ Hn1, Hb1 ⟩⟩ simp [zCDP] at * - sorry - -- cases h ; rename_i h1 h2 ; cases h2 ; rename_i h2 h3 ; cases h3 ; rename_i h3 h4 ; cases h4 ; rename_i h4 h5 - -- repeat any_goals constructor - -- . apply privPostProcess_zCDPBound h1 h2 h5 h4 h3 - -- . apply privPostProcess_NonZeroNQ h2 sur - -- . apply privPostProcess_NonTopSum f h3 - -- . simp [NonTopNQ] - -- intro l - -- apply ENNReal.ne_top_of_tsum_ne_top - -- apply privPostProcess_NonTopSum f h3 - -- . apply privPostProcess_NonTopRDNQ f h1 h5 h2 h4 h3 + apply And.intro + · exact privPostProcess_AC nq Hac1 + · exists (privPostProcess_norm nq Hn1 f) + exact (privPostProcess_zCDPBound Hb1 f) end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean index a7bbbe8b..36b699a5 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean @@ -27,6 +27,6 @@ noncomputable instance gaussian_zCDPSystem : DPSystem T where noise := privNoisedQuery noise_prop := privNoisedQuery_zCDP compose_prop := privCompose_zCDP - postprocess_prop := privPostProcess_zCDP + postprocess_prop := (fun _ => privPostProcess_zCDP) -- FIXME: Remove surjectivity from system definition? end SLang From 149d62dcaf8460a9380e491a955f176496723a08 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Mon, 24 Jun 2024 16:09:29 -0400 Subject: [PATCH 090/122] (broken build) sketch PostProcess_pre_reduct --- .../ZeroConcentrated/Postprocessing.lean | 493 +++++++----------- 1 file changed, 189 insertions(+), 304 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean index 4d9167cc..689d90bb 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean @@ -217,335 +217,199 @@ theorem tsum_pos_int {f : ℤ → ENNReal} (h1 : ∑' x : ℤ, f x ≠ ⊤) (h2 apply ENNReal.tsum_pos_int h1 h2 +-- set_option pp.coercions false theorem DPostPocess_pre_reduct {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) 1} {ε₁ ε₂ : ℕ+} (h : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) - (nn : ∀(l : List T), ∀(u : U), nq l u ≠ 0) (nt : NonTopRDNQ nq) (nts : NonTopNQ nq) (conv : NonTopSum nq) + (nn : ∀(l : List T), ∀(u : U), nq l u ≠ 0) (f : U → V) {α : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (Habs : AbsCts (nq l₁) (nq l₂)) (h2 : Neighbour l₁ l₂) : (∑' (x : V), (∑' (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 [zCDPBound, RenyiDivergence] at h + -- Turn everything into an explicit PMF + let nq_PMF (l : List T) : PMF U := ⟨ nq l, HNorm l ⟩ + have nq_PMF_coe (l : List T) (u : U) : nq l u = nq_PMF l u := by + rw [DFunLike.coe] + simp [PMF.instFunLike] + conv => + congr + · arg 1 + intro x + congr + · arg 1 + arg 1 + intro i + rw [nq_PMF_coe] + · arg 1 + arg 1 + intro a + rw [nq_PMF_coe] + · arg 1 + intro x + rw [nq_PMF_coe] + + -- Derive the assumptions from the old proof + have nq_nts : ∀ l : List T, ∑' n : U, nq_PMF l n ≠ ⊤ := by + intro l + exact PMF.tsum_coe_ne_top (nq_PMF l) + + have nq_nt : ∀ l : List T, ∀ u : U, nq_PMF l u ≠ ⊤ := by + exact fun l u => PMF.apply_ne_top (nq_PMF l) u + + -- TODO (if needed): Renyi divergence is not Top -- Rewrite as cascading expectations - rw [@RenyiDivergenceExpectation _ (nq l₁) (nq l₂) _ h1 Habs] + simp [zCDPBound, RenyiDivergence] at h + rw [RenyiDivergenceExpectation] + case h => apply h1 + case H => sorry + simp + rw [RenyiDivergenceExpectation (fun x => DFunLike.coe (nq_PMF l₁) x) (fun x => nq l₂ x) h1 Habs] + conv => + rhs + arg 1 + intro x + rw [nq_PMF_coe] -- Shuffle the sum - rw [fiberwisation (fun x => (nq l₁ x / nq l₂ x) ^ α * nq l₂ x) f] - + rw [fiberwisation (fun x => (nq_PMF l₁ x / nq_PMF l₂ x) ^ α * nq_PMF l₂ x) f] apply ENNReal.tsum_le_tsum - intro i + simp - -- Get rid of elements with probability 0 in the pushforward + -- Eliminate epements with probability zero split - . rename_i empty - rw [condition_to_subset] - 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 + case h.inl => + rename_i H + repeat rw [condition_to_subset] + rw [H] + simp - 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 - . -- Need absolute continuity between neighbours - intro x - sorry - -- 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 [condition_to_subset] - rw [condition_to_subset] - - -- 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₂ + -- Part 2: Apply Jensen's inequality to the normalized fibers + rename_i NotEmpty + repeat rw [condition_to_subset] - conv => - right - right - intro a - rw [← mul_one ((nq l₁ ↑a / nq l₂ ↑a) ^ α * nq l₂ ↑a)] - right - rw [← P4] - clear P4 - simp only [κ] + have MasterRW : ∀ l : List T, ∑' (a : ↑{a | i = f a}), nq_PMF l ↑a ≠ ⊤ := by + intro l + exact convergent_subset (fun y => f y) (nq_nts l) - conv => - right - right - intro a - right - rw [division_def] - rw [mul_comm] + have MasterZero : ∀ l : List T, ∑' (a : ↑{a | i = f a}), nq_PMF l ↑a ≠ 0 := by + intro l + simp + have T := witness NotEmpty + cases T + rename_i z w + exists z + constructor + . trivial + . apply nn l + simp - conv => - right - right - intro a - rw [← mul_assoc] + have Hac : AbsCts (δpmf (nq_PMF l₁) f i (MasterZero l₁) (MasterRW l₁)) (δpmf (nq_PMF l₂) f i (MasterZero l₂) (MasterRW l₂)) := by sorry + rw [AbsCts] at Hac - rw [ENNReal.tsum_mul_right] + have HJ := Renyi_Jensen_ENNReal (δpmf (nq_PMF l₁) f i (MasterZero l₁) (MasterRW l₁)) (δpmf (nq_PMF l₂) f i (MasterZero l₂) (MasterRW l₂)) h1 Hac - -- Jensen's inequality + have δnt (x : { x // i = f x }) (l : List T) : (δpmf (nq_PMF l) f i (MasterZero l) (MasterRW l) x) ≠ ⊤ := by sorry - 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 := @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 - simp only [toReal_nonneg] - have A' := mul_le_mul_of_nonneg_left Jensen's P6 - clear Jensen's P6 + have Hspecial (x : { x // i = f x }) : ¬((δpmf (nq_PMF l₁) f i (MasterZero l₁) (MasterRW l₁) x) ≠ 0 ∧ (δpmf (nq_PMF l₂) f i (MasterZero l₂) (MasterRW l₂) x) = ⊤) := by + simp + intro _ Hcont + exfalso + apply δnt + apply Hcont - conv => - right - rw [mul_comm] - right - right - intro a - rw [mul_assoc] - rw [δpmf_conv] + -- Cancel the δPMF on the LHS of HJ, its sum is 1 + conv at HJ => + lhs + arg 1 + arg 1 + intro x + rw [division_def] + rw [mul_mul_inv_eq_mul_cancel (Hac x) (Hspecial x)] + conv at HJ => + lhs + arg 1 + rw [PMF.tsum_coe] + simp at HJ + + -- Unfold normalization constants from δpmf + let N (l : List T) := (∑' (x : {n // i = f n}), nq_PMF l x)⁻¹ + have N_def (l : List T) : N l = (∑' (x : {n // i = f n}), nq_PMF l x)⁻¹ := by sorry -- exact rfl + have N_inv (l : List T) : (∑' (x : {n // i = f n}), nq_PMF l x) = (N l)⁻¹ := by sorry -- exact rfl + have δpmf_N (l : List T) (x : { x // i = f x }) : (δpmf (nq_PMF l) f i (MasterZero l) (MasterRW l)) x = (N l) * nq_PMF l x := by + simp [δpmf] + unfold δ + repeat rw [DFunLike.coe] + repeat rw [PMF.instFunLike] + simp + repeat rw [DFunLike.coe] + repeat rw [PMF.instFunLike] + simp + rw [N_def] + rw [mul_comm] + repeat rw [DFunLike.coe] + repeat rw [PMF.instFunLike] + conv at HJ => + rhs + arg 1 + intro x + simp [δpmf] + unfold δ + repeat rw [DFunLike.coe] + repeat rw [PMF.instFunLike] + simp + rw [<- N_def] + rw [<- N_def] - -- Here + -- Simplify constants in goal + rw [N_inv] + rw [N_inv] - 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' + -- Simplify constants in HJ + have C (x : { n // i = f n }) : (((nq_PMF l₁) ↑x * N l₁ / ((nq_PMF l₂) ↑x * N l₂)) : ENNReal) = ((N l₁ / N l₂) * ((nq_PMF l₁) ↑x / ((nq_PMF l₂) ↑x )) : ENNReal) := by sorry + conv at HJ => + rhs + arg 1 + intro x + arg 1 + arg 1 + rw [C] + clear C - 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' + have Hα' : 0 <= α := by + linarith - apply le_trans _ A' - clear A' - apply le_of_eq + conv at HJ => + rhs + arg 1 + intro x + arg 1 + rw [ENNReal.mul_rpow_of_nonneg _ _ Hα'] + conv at HJ => + rhs + arg 1 + intro x + rw [mul_assoc] + arg 2 + rw [<- mul_assoc] + rw [mul_comm] + rw [ENNReal.tsum_mul_left] at HJ + rw [ENNReal.tsum_mul_left] at HJ - -- Part 3: + -- Apply transitivity with HJ + rw [<- mul_assoc] at HJ + rw [mul_comm] at HJ + apply (ENNReal.div_le_iff_le_mul ?G1 ?G2).mpr at HJ + apply (le_trans ?G3 HJ) - 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] + -- These are equal + sorry - 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 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 @@ -556,15 +420,15 @@ theorem tsum_ne_zero_of_ne_zero {T : Type} [Inhabited T] (f : T → ENNReal) (h contradiction -theorem DPostPocess_pre {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) 1} {ε₁ ε₂ : ℕ+} - (h : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) + +theorem DPostPocess_pre_reduct_1 {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) 1} {ε₁ ε₂ : ℕ+} + (h : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) (nn : ∀(l : List T), ∀(u : U), nq l u ≠ 0) (f : U → V) {α : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (Habs : AbsCts (nq l₁) (nq l₂)) (h2 : Neighbour l₁ l₂) : (∑' (x : V), (∑' (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 - -- First step is to reduce to the case where (nq l) is nonzero -- Next I need to generalize the fiberwisation argument to eliminate the side conditions and work in -- the extended reals -- I'll also have to generalize the δpmf normalization, though that should mstly just be making sure the correct terms are PMFs @@ -575,8 +439,29 @@ theorem DPostPocess_pre {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) +theorem DPostPocess_pre {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) 1} {ε₁ ε₂ : ℕ+} + (h : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) + (f : U → V) {α : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (Habs : AbsCts (nq l₁) (nq l₂)) (h2 : Neighbour l₁ l₂) : + (∑' (x : V), + (∑' (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 + -- First step is to reduce to the case where (nq l) is nonzero (reduct_1) + + + -- have K1 : Function.support (fun (a : U) => if x = f a then nq l₁ a else 0) ⊆ { t : T | q t ≠ 0 } := by simp [Function.support] + + -- have K2 : Function.support (fun x : T => (p x / q x)^α * q x) ⊆ { t : T | q t ≠ 0 } := by simp [Function.support] + -- rw [<- tsum_subtype_eq_of_support_subset K1] at Hsumeq + -- rw [<- tsum_subtype_eq_of_support_subset K2] at Hsumeq + -- simp at Hsumeq + + sorry + + + theorem privPostProcess_zCDPBound {nq : List T → SLang U} {HNorm : NormalMechanism nq} {ε₁ ε₂ : ℕ+} - (h : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) (f : U → V) : + (h : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) (f : U → V) (Hac : ACNeighbour nq) : zCDPBound (privPostProcess nq f) (privPostProcess_norm nq HNorm f) ((ε₁ : ℝ) / ε₂) := by simp [privPostProcess, zCDPBound, RenyiDivergence] intro α h1 l₁ l₂ h2 @@ -617,9 +502,9 @@ theorem privPostProcess_zCDPBound {nq : List T → SLang U} {HNorm : NormalMecha · exact EReal.coe_lt_top (α - 1)⁻¹ apply elog_mono_le.mp apply (DPostPocess_pre h f h1 ?Gabs h2) - case Gabs => - -- absolute continuity - sorry + case Gabs => exact Hac l₁ l₂ h2 + + /- -- remove the log @@ -718,7 +603,7 @@ theorem privPostProcess_zCDP {f : U → V} simp [zCDP] at * apply And.intro · exact privPostProcess_AC nq Hac1 - · exists (privPostProcess_norm nq Hn1 f) - exact (privPostProcess_zCDPBound Hb1 f) + · exists (privPostProcess_norm nq Hn1 f) + exact (privPostProcess_zCDPBound Hb1 f Hac1) end SLang From db76175e6833b6f8407be508f1d1f3d7d78cac11 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Mon, 24 Jun 2024 17:10:41 -0400 Subject: [PATCH 091/122] (build fixed) checkpoint progress on PostProcess_pre_reduct --- .../ZeroConcentrated/Postprocessing.lean | 275 +++++++++--------- 1 file changed, 136 insertions(+), 139 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean index 689d90bb..33e9b167 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean @@ -28,105 +28,106 @@ variable [count : Countable U] variable [disc : DiscreteMeasurableSpace U] variable [Inhabited U] -lemma 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 ℝ t1 μ _ 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] - -/-- -Jensen's ineuquality for the exponential applied to Renyi's sum --/ -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 => - 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 ℝ 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 - . 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 t1 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 t1 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 - -def δ (nq : SLang U) (f : U → V) (a : V) : {n : U | a = f n} → ENNReal := fun x : {n : U | a = f n} => nq x * (∑' (x : {n | a = f n}), nq x)⁻¹ +-- lemma 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 ℝ t1 μ _ 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] + +-- /- +-- Jensen's ineuquality for the exponential applied to Renyi's sum +-- -/ +-- 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 => +-- 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 ℝ 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 +-- . 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 t1 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 t1 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 + +def δ (nq : SLang U) (f : U → V) (a : V) : {n : U | a = f n} → ENNReal := + fun x : {n : U | a = f n} => nq x * (∑' (x : {n | a = f n}), nq x)⁻¹ lemma δ_normalizes (nq : SLang U) (f : U → V) (a : V) (h1 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ 0) (h2 : ∑' (i : ↑{n | a = f n}), nq ↑i ≠ ⊤) : HasSum (δ nq f a) 1 := by @@ -216,23 +217,20 @@ theorem tsum_pos_int {f : ℤ → ENNReal} (h1 : ∑' x : ℤ, f x ≠ ⊤) (h2 apply toReal_strict_mono h1 apply ENNReal.tsum_pos_int h1 h2 - --- set_option pp.coercions false - -theorem DPostPocess_pre_reduct {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) 1} {ε₁ ε₂ : ℕ+} +theorem PostPocess_pre_reduct {nq : List T → SLang U} {HNorm : ∀ (l : List T), HasSum (nq l) 1} {ε₁ ε₂ : ℕ+} (h : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) (nn : ∀(l : List T), ∀(u : U), nq l u ≠ 0) (f : U → V) {α : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (Habs : AbsCts (nq l₁) (nq l₂)) (h2 : Neighbour l₁ l₂) : - (∑' (x : V), - (∑' (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 + (∑' (x : V), (∑' (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 + + -- Turn everything into an explicit PMF (is this necessary? I don't think so, since we use δpmf for Renyi lemma) + let nq_PMF : List T -> PMF U := (fun l => SLang.toPMF (nq l) (HNorm l)) - -- Turn everything into an explicit PMF - let nq_PMF (l : List T) : PMF U := ⟨ nq l, HNorm l ⟩ have nq_PMF_coe (l : List T) (u : U) : nq l u = nq_PMF l u := by rw [DFunLike.coe] simp [PMF.instFunLike] + exact rfl + conv => congr · arg 1 @@ -258,13 +256,13 @@ theorem DPostPocess_pre_reduct {nq : List T → SLang U} {HNorm : ∀ l, HasSum have nq_nt : ∀ l : List T, ∀ u : U, nq_PMF l u ≠ ⊤ := by exact fun l u => PMF.apply_ne_top (nq_PMF l) u - -- TODO (if needed): Renyi divergence is not Top - -- Rewrite as cascading expectations simp [zCDPBound, RenyiDivergence] at h rw [RenyiDivergenceExpectation] case h => apply h1 - case H => sorry + case H => + -- Unfold the other top-level continuity proof for postprocessing + sorry simp rw [RenyiDivergenceExpectation (fun x => DFunLike.coe (nq_PMF l₁) x) (fun x => nq l₂ x) h1 Habs] conv => @@ -337,8 +335,9 @@ theorem DPostPocess_pre_reduct {nq : List T → SLang U} {HNorm : ∀ l, HasSum -- Unfold normalization constants from δpmf let N (l : List T) := (∑' (x : {n // i = f n}), nq_PMF l x)⁻¹ - have N_def (l : List T) : N l = (∑' (x : {n // i = f n}), nq_PMF l x)⁻¹ := by sorry -- exact rfl - have N_inv (l : List T) : (∑' (x : {n // i = f n}), nq_PMF l x) = (N l)⁻¹ := by sorry -- exact rfl + have N_def (l : List T) : N l = (∑' (x : {n // i = f n}), nq_PMF l x)⁻¹ := by exact rfl + have N_inv (l : List T) : (∑' (x : {n // i = f n}), nq_PMF l x) = (N l)⁻¹ := by + exact Eq.symm (inv_inv (∑' (x : { n // i = f n }), (nq_PMF l) ↑x)) have δpmf_N (l : List T) (x : { x // i = f x }) : (δpmf (nq_PMF l) f i (MasterZero l) (MasterRW l)) x = (N l) * nq_PMF l x := by simp [δpmf] unfold δ @@ -402,15 +401,34 @@ theorem DPostPocess_pre_reduct {nq : List T → SLang U} {HNorm : ∀ l, HasSum -- Apply transitivity with HJ rw [<- mul_assoc] at HJ rw [mul_comm] at HJ - apply (ENNReal.div_le_iff_le_mul ?G1 ?G2).mpr at HJ + + -- Super haunted bug: When I apply this as normal to HJ (with placeholders) + -- Lean it lights up all of my "have" and "let" statements because it \"doesn't + -- know how to synthesize\" a placeholder. The placeholder it points me to is in + -- Pure/Postprocessing, where the same lemma is also applied with placeholders. + -- + -- So I can't use placeholders here, and I have to write out the terms explicitly + have W := + @ENNReal.div_le_iff_le_mul + 1 + ((N l₁ / N l₂) ^ α * N l₂) + (∑' (i_1 : { x // i = f x }), ((nq_PMF l₁) ↑i_1 / (nq_PMF l₂) ↑i_1) ^ α * (nq_PMF l₂) ↑i_1) + ?G1 + ?G2 + case G1 => sorry -- Probably right? + case G2 => sorry -- left + apply W.mpr at HJ + clear W + apply (le_trans ?G3 HJ) + clear HJ -- These are equal + apply Eq.le + repeat rw [division_def] + simp sorry - - - 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 @@ -420,25 +438,6 @@ theorem tsum_ne_zero_of_ne_zero {T : Type} [Inhabited T] (f : T → ENNReal) (h contradiction - -theorem DPostPocess_pre_reduct_1 {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) 1} {ε₁ ε₂ : ℕ+} - (h : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) (nn : ∀(l : List T), ∀(u : U), nq l u ≠ 0) - (f : U → V) {α : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (Habs : AbsCts (nq l₁) (nq l₂)) (h2 : Neighbour l₁ l₂) : - (∑' (x : V), - (∑' (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 - - -- Next I need to generalize the fiberwisation argument to eliminate the side conditions and work in - -- the extended reals - -- I'll also have to generalize the δpmf normalization, though that should mstly just be making sure the correct terms are PMFs - -- That should bring be down to the extended Renyi divergence - - -- This lemma should be provable as stated - sorry - - - theorem DPostPocess_pre {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) 1} {ε₁ ε₂ : ℕ+} (h : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) (f : U → V) {α : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (Habs : AbsCts (nq l₁) (nq l₂)) (h2 : Neighbour l₁ l₂) : @@ -458,8 +457,6 @@ theorem DPostPocess_pre {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) sorry - - theorem privPostProcess_zCDPBound {nq : List T → SLang U} {HNorm : NormalMechanism nq} {ε₁ ε₂ : ℕ+} (h : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) (f : U → V) (Hac : ACNeighbour nq) : zCDPBound (privPostProcess nq f) (privPostProcess_norm nq HNorm f) ((ε₁ : ℝ) / ε₂) := by From 9967bb758f4e12a7ea9fc08fe496021f81f92b32 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Tue, 25 Jun 2024 11:15:09 -0400 Subject: [PATCH 092/122] strengthen PostProcess_pre_reduct for the reduction and cleanup --- .../ZeroConcentrated/Postprocessing.lean | 341 +++++++++++------- 1 file changed, 214 insertions(+), 127 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean index 33e9b167..d60f9984 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean @@ -126,6 +126,17 @@ variable [Inhabited U] -- rw [one_le_ofReal] -- apply le_of_lt h +def privPostProcess_AC {f : U -> V} (nq : List T → SLang U) (Hac : ACNeighbour nq) : ACNeighbour (privPostProcess nq f) := by + rw [ACNeighbour] at * + unfold AbsCts at * + intro l₁ l₂ Hn v + have Hac := Hac l₁ l₂ Hn + simp [privPostProcess] + intro Hppz i fi + apply Hac + apply Hppz + apply fi + def δ (nq : SLang U) (f : U → V) (a : V) : {n : U | a = f n} → ENNReal := fun x : {n : U | a = f n} => nq x * (∑' (x : {n | a = f n}), nq x)⁻¹ @@ -217,67 +228,58 @@ theorem tsum_pos_int {f : ℤ → ENNReal} (h1 : ∑' x : ℤ, f x ≠ ⊤) (h2 apply toReal_strict_mono h1 apply ENNReal.tsum_pos_int h1 h2 + +lemma rpow_nonzero (x : ENNReal) (y : ℝ) (H : ¬(x = 0 ∧ 0 < y ∨ x = ⊤ ∧ y < 0)) : x ^ y ≠ 0 := by + intro Hk + apply H + apply (ENNReal.rpow_eq_zero_iff).mp + apply Hk + + + + + +-- set_option pp.coercions false + theorem PostPocess_pre_reduct {nq : List T → SLang U} {HNorm : ∀ (l : List T), HasSum (nq l) 1} {ε₁ ε₂ : ℕ+} - (h : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) - (nn : ∀(l : List T), ∀(u : U), nq l u ≠ 0) - (f : U → V) {α : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (Habs : AbsCts (nq l₁) (nq l₂)) (h2 : Neighbour l₁ l₂) : + (f : U → V) {α : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} + (Habs : AbsCts (nq l₁) (nq l₂)) + (Hnq2 : ∀ (u : U), nq l₁ u ≠ 0) + (h2 : Neighbour l₁ l₂) : (∑' (x : V), (∑' (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 - -- Turn everything into an explicit PMF (is this necessary? I don't think so, since we use δpmf for Renyi lemma) - let nq_PMF : List T -> PMF U := (fun l => SLang.toPMF (nq l) (HNorm l)) + -- By absolute continuity, nq1 is nonzero + have Hnql1 : (∀ (u : U), nq l₂ u ≠ 0) := by + rw [AbsCts] at Habs + intro u HK + apply Hnq2 + apply Habs + apply HK - have nq_PMF_coe (l : List T) (u : U) : nq l u = nq_PMF l u := by - rw [DFunLike.coe] - simp [PMF.instFunLike] - exact rfl - - conv => - congr - · arg 1 - intro x - congr - · arg 1 - arg 1 - intro i - rw [nq_PMF_coe] - · arg 1 - arg 1 - intro a - rw [nq_PMF_coe] - · arg 1 - intro x - rw [nq_PMF_coe] - - -- Derive the assumptions from the old proof - have nq_nts : ∀ l : List T, ∑' n : U, nq_PMF l n ≠ ⊤ := by + -- Noised query is not ⊤ + have nq_nts : ∀ l : List T, ∑' n : U, nq l n ≠ ⊤ := by intro l - exact PMF.tsum_coe_ne_top (nq_PMF l) - - have nq_nt : ∀ l : List T, ∀ u : U, nq_PMF l u ≠ ⊤ := by - exact fun l u => PMF.apply_ne_top (nq_PMF l) u + simp [HasSum.tsum_eq (HNorm l)] -- Rewrite as cascading expectations - simp [zCDPBound, RenyiDivergence] at h rw [RenyiDivergenceExpectation] case h => apply h1 case H => - -- Unfold the other top-level continuity proof for postprocessing - sorry + simp [AbsCts] + intro v Hv i vEq + apply Habs + apply Hv + apply vEq simp - rw [RenyiDivergenceExpectation (fun x => DFunLike.coe (nq_PMF l₁) x) (fun x => nq l₂ x) h1 Habs] - conv => - rhs - arg 1 - intro x - rw [nq_PMF_coe] + rw [RenyiDivergenceExpectation (fun x => nq l₁ x) (fun x => nq l₂ x) h1 Habs] -- Shuffle the sum - rw [fiberwisation (fun x => (nq_PMF l₁ x / nq_PMF l₂ x) ^ α * nq_PMF l₂ x) f] + rw [fiberwisation (fun x => (nq l₁ x / nq l₂ x) ^ α * nq l₂ x) f] apply ENNReal.tsum_le_tsum intro i simp - -- Eliminate epements with probability zero + -- Eliminate elements with probability zero split case h.inl => rename_i H @@ -285,16 +287,14 @@ theorem PostPocess_pre_reduct {nq : List T → SLang U} {HNorm : ∀ (l : List T rw [H] simp - -- Part 2: Apply Jensen's inequality to the normalized fibers + -- Normalize each fiber into a PMF rename_i NotEmpty repeat rw [condition_to_subset] - have MasterRW : ∀ l : List T, ∑' (a : ↑{a | i = f a}), nq_PMF l ↑a ≠ ⊤ := by - intro l - exact convergent_subset (fun y => f y) (nq_nts l) + have nq_restriction_nts (l' : List T) : ∑' (a : ↑{a | i = f a}), nq l' ↑a ≠ ⊤ := by + exact convergent_subset (fun y => f y) (nq_nts l') - have MasterZero : ∀ l : List T, ∑' (a : ↑{a | i = f a}), nq_PMF l ↑a ≠ 0 := by - intro l + have nq_restriction_nzs (l' : List T) (Hl : ∀ u : U, nq l' u ≠ 0) : ∑' (a : ↑{a | i = f a}), nq l' ↑a ≠ 0 := by simp have T := witness NotEmpty cases T @@ -302,59 +302,83 @@ theorem PostPocess_pre_reduct {nq : List T → SLang U} {HNorm : ∀ (l : List T exists z constructor . trivial - . apply nn l + . exact Hl z simp - have Hac : AbsCts (δpmf (nq_PMF l₁) f i (MasterZero l₁) (MasterRW l₁)) (δpmf (nq_PMF l₂) f i (MasterZero l₂) (MasterRW l₂)) := by sorry - rw [AbsCts] at Hac - - have HJ := Renyi_Jensen_ENNReal (δpmf (nq_PMF l₁) f i (MasterZero l₁) (MasterRW l₁)) (δpmf (nq_PMF l₂) f i (MasterZero l₂) (MasterRW l₂)) h1 Hac - - have δnt (x : { x // i = f x }) (l : List T) : (δpmf (nq_PMF l) f i (MasterZero l) (MasterRW l) x) ≠ ⊤ := by sorry - - have Hspecial (x : { x // i = f x }) : ¬((δpmf (nq_PMF l₁) f i (MasterZero l₁) (MasterRW l₁) x) ≠ 0 ∧ (δpmf (nq_PMF l₂) f i (MasterZero l₂) (MasterRW l₂) x) = ⊤) := by + let δF₁ := (δpmf (nq l₁) f i (nq_restriction_nzs l₁ Hnq2) (nq_restriction_nts l₁)) + let δF₂ := (δpmf (nq l₂) f i (nq_restriction_nzs l₂ Hnql1) (nq_restriction_nts l₂)) + have δF₁_Eq : δF₁ = (δpmf (nq l₁) f i (nq_restriction_nzs l₁ Hnq2) (nq_restriction_nts l₁)) := by exact rfl + have δF₂_Eq : δF₂ = (δpmf (nq l₂) f i (nq_restriction_nzs l₂ Hnql1) (nq_restriction_nts l₂)) := by exact rfl + + -- Normalized fibers are absolutely continuous + have HAC_Fiber : AbsCts δF₁ δF₂ := by + simp [AbsCts] + rw [δF₁_Eq] + rw [δF₂_Eq] + intro a b + repeat rw [δpmf] + unfold δ + simp + rw [DFunLike.coe] + simp [PMF.instFunLike] + intro H + cases H + · rename_i Hl2z + left + apply Habs + apply Hl2z + · exfalso + apply nq_restriction_nts l₂ + simp + assumption + + have δF₂_NT (x : { x // i = f x }) : δF₂ x ≠ ⊤ := by + rw [δF₂_Eq] + exact PMF.apply_ne_top ((nq l₂).δpmf f i (nq_restriction_nzs l₂ Hnql1) (nq_restriction_nts l₂)) x + + -- Normalized fibers avoid the bad case for rewriting the Jensen term + have Hspecial (x : { x // i = f x }) : ¬(δF₁ x ≠ 0 ∧ δF₂ x = ⊤) := by simp intro _ Hcont exfalso - apply δnt + apply δF₂_NT apply Hcont - -- Cancel the δPMF on the LHS of HJ, its sum is 1 + -- Apply Jensen's inequality to the normalized fibers + have HJ := Renyi_Jensen_ENNReal δF₁ δF₂ h1 HAC_Fiber + + -- Cancel and simplify the LHS of HJ to 1 conv at HJ => lhs arg 1 arg 1 intro x rw [division_def] - rw [mul_mul_inv_eq_mul_cancel (Hac x) (Hspecial x)] + rw [mul_mul_inv_eq_mul_cancel (HAC_Fiber x) (Hspecial x)] conv at HJ => lhs arg 1 rw [PMF.tsum_coe] simp at HJ - -- Unfold normalization constants from δpmf - let N (l : List T) := (∑' (x : {n // i = f n}), nq_PMF l x)⁻¹ - have N_def (l : List T) : N l = (∑' (x : {n // i = f n}), nq_PMF l x)⁻¹ := by exact rfl - have N_inv (l : List T) : (∑' (x : {n // i = f n}), nq_PMF l x) = (N l)⁻¹ := by - exact Eq.symm (inv_inv (∑' (x : { n // i = f n }), (nq_PMF l) ↑x)) - have δpmf_N (l : List T) (x : { x // i = f x }) : (δpmf (nq_PMF l) f i (MasterZero l) (MasterRW l)) x = (N l) * nq_PMF l x := by - simp [δpmf] - unfold δ - repeat rw [DFunLike.coe] - repeat rw [PMF.instFunLike] - simp - repeat rw [DFunLike.coe] - repeat rw [PMF.instFunLike] - simp - rw [N_def] - rw [mul_comm] - repeat rw [DFunLike.coe] - repeat rw [PMF.instFunLike] + -- Name the normalization constants for each fiber + let N (l : List T) := (∑' (x : {n // i = f n}), nq l x)⁻¹ + have N_def (l : List T) : N l = (∑' (x : {n // i = f n}), nq l x)⁻¹ := by exact rfl + have N_inv (l : List T) : (∑' (x : {n // i = f n}), nq l x) = (N l)⁻¹ := by + exact Eq.symm (inv_inv (∑' (x : { n // i = f n }), (nq l) ↑x)) + have N1_nz : N l₁ ≠ 0 := ENNReal.inv_ne_zero.mpr (nq_restriction_nts l₁) + have N2_nz : N l₂ ≠ 0 := ENNReal.inv_ne_zero.mpr (nq_restriction_nts l₂) + + -- Unfold normalization constants in HJ conv at HJ => rhs arg 1 intro x + repeat rw [DFunLike.coe] + repeat rw [PMF.instFunLike] + simp + rw [δF₁_Eq] + rw [δF₂_Eq] simp [δpmf] unfold δ repeat rw [DFunLike.coe] @@ -363,71 +387,148 @@ theorem PostPocess_pre_reduct {nq : List T → SLang U} {HNorm : ∀ (l : List T rw [<- N_def] rw [<- N_def] - -- Simplify constants in goal + -- Fold constants in goal rw [N_inv] rw [N_inv] - -- Simplify constants in HJ - have C (x : { n // i = f n }) : (((nq_PMF l₁) ↑x * N l₁ / ((nq_PMF l₂) ↑x * N l₂)) : ENNReal) = ((N l₁ / N l₂) * ((nq_PMF l₁) ↑x / ((nq_PMF l₂) ↑x )) : ENNReal) := by sorry - conv at HJ => - rhs - arg 1 - intro x - arg 1 - arg 1 - rw [C] - clear C - - have Hα' : 0 <= α := by - linarith - - conv at HJ => - rhs - arg 1 - intro x - arg 1 - rw [ENNReal.mul_rpow_of_nonneg _ _ Hα'] + -- Pull out constants from the sum in HJ + have conv1 (x : { x // i = f x }) : (nq l₁ x.val * N l₁ / (nq l₂ x.val * N l₂)) ^ α = (N l₁ / N l₂) ^ α * (nq l₁ x.val / nq l₂ x.val) ^ α := by + simp [division_def] + rw [ENNReal.mul_rpow_of_ne_top ?G1 ?G2] + case G1 => + apply mul_ne_top + · exact ENNReal.ne_top_of_tsum_ne_top (nq_nts l₁) ↑x + · exact inv_ne_top.mpr (nq_restriction_nzs l₁ Hnq2) + case G2 => + apply inv_ne_top.mpr + exact mul_ne_zero (Hnql1 ↑x) N2_nz + rw [ENNReal.mul_rpow_of_ne_top ?G1 ?G2] + case G1 => + exact ENNReal.ne_top_of_tsum_ne_top (nq_nts l₁) ↑x + case G2 => + exact inv_ne_top.mpr (nq_restriction_nzs l₁ Hnq2) + rw [ENNReal.mul_rpow_of_ne_top ?G1 ?G2] + case G1 => + exact inv_ne_top.mpr (nq_restriction_nzs l₁ Hnq2) + case G2 => + exact inv_ne_top.mpr N2_nz + rw [ENNReal.mul_rpow_of_ne_top ?G1 ?G2] + case G1 => + exact ENNReal.ne_top_of_tsum_ne_top (nq_nts l₁) ↑x + case G2 => + exact inv_ne_top.mpr (Hnql1 ↑x) + conv => + arg 1 + rw [mul_assoc] + rw [mul_comm] + skip + repeat rw [mul_assoc] + congr 1 + conv => + arg 2 + arg 2 + rw [mul_comm] + rw [<- mul_assoc] + congr 1 + rw [ENNReal.mul_inv ?G1 ?G2] + case G1 => + right + exact inv_ne_top.mpr (nq_restriction_nzs l₂ Hnql1) + case G2 => + right + exact N2_nz + rw [ENNReal.mul_rpow_of_ne_top ?G1 ?G2] + case G1 => exact inv_ne_top.mpr (Hnql1 ↑x) + case G2 => exact inv_ne_top.mpr N2_nz + rw [mul_comm] conv at HJ => rhs arg 1 intro x + rw [conv1 x] rw [mul_assoc] arg 2 - rw [<- mul_assoc] rw [mul_comm] + rw [mul_assoc] + rw [mul_comm] + rw [mul_assoc] + clear conv1 rw [ENNReal.tsum_mul_left] at HJ rw [ENNReal.tsum_mul_left] at HJ - - -- Apply transitivity with HJ rw [<- mul_assoc] at HJ - rw [mul_comm] at HJ + -- Move constants to the left-hand side of HJ -- Super haunted bug: When I apply this as normal to HJ (with placeholders) -- Lean it lights up all of my "have" and "let" statements because it \"doesn't -- know how to synthesize\" a placeholder. The placeholder it points me to is in -- Pure/Postprocessing, where the same lemma is also applied with placeholders. - -- - -- So I can't use placeholders here, and I have to write out the terms explicitly have W := @ENNReal.div_le_iff_le_mul 1 ((N l₁ / N l₂) ^ α * N l₂) - (∑' (i_1 : { x // i = f x }), ((nq_PMF l₁) ↑i_1 / (nq_PMF l₂) ↑i_1) ^ α * (nq_PMF l₂) ↑i_1) + (∑' (i_1 : { x // i = f x }), (nq l₁ i_1.val / nq l₂ i_1.val) ^ α * nq l₂ i_1.val) ?G1 ?G2 - case G1 => sorry -- Probably right? - case G2 => sorry -- left + case G1 => + left + apply mul_ne_zero_iff.mpr + apply And.intro + · apply (rpow_nonzero (N l₁ / N l₂) α) + intro Hk + rcases Hk with ⟨ H1 , _ ⟩ | ⟨ _ , H2 ⟩ + · have Hk' : 0 < N l₁ / N l₂ := by + apply ENNReal.div_pos_iff.mpr + apply And.intro + · exact N1_nz + · exact inv_ne_top.mpr (nq_restriction_nzs l₂ Hnql1) + rw [H1] at Hk' + simp at Hk' + · linarith + · apply N2_nz + case G2 => + left + apply mul_ne_top + · apply rpow_ne_top_of_nonneg + · linarith + · intro Hk + have Hk' : (N l₁ / N l₂ < ⊤) := by + apply ENNReal.div_lt_top + · exact inv_ne_top.mpr (nq_restriction_nzs l₁ Hnq2) + · exact N2_nz + rw [Hk] at Hk' + simp at Hk' + · exact inv_ne_top.mpr (nq_restriction_nzs l₂ Hnql1) + rw [mul_comm] at HJ apply W.mpr at HJ clear W + -- Apply transitivity apply (le_trans ?G3 HJ) clear HJ - -- These are equal + -- These terms are equal apply Eq.le repeat rw [division_def] simp - sorry + rw [ENNReal.mul_inv ?G1 ?G2] + case G1 => + right + exact inv_ne_top.mpr (nq_restriction_nzs l₂ Hnql1) + case G2 => + right + exact N2_nz + congr + rw [<- ENNReal.inv_rpow] + congr + rw [ENNReal.mul_inv ?G1 ?G2] + case G1 => + left + exact N1_nz + case G2 => + left + exact inv_ne_top.mpr (nq_restriction_nzs l₁ Hnq2) + simp + 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 @@ -447,8 +548,7 @@ theorem DPostPocess_pre {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) (∑' (x : U), nq l₁ x ^ α * nq l₂ x ^ (1 - α)) := by -- First step is to reduce to the case where (nq l) is nonzero (reduct_1) - - -- have K1 : Function.support (fun (a : U) => if x = f a then nq l₁ a else 0) ⊆ { t : T | q t ≠ 0 } := by simp [Function.support] + -- have K1_pre (x : V) : Function.support (fun (a : U) => if x = f a then nq l₁ a else 0) ⊆ { t : T | nq t ≠ 0 } := by -- have K2 : Function.support (fun x : T => (p x / q x)^α * q x) ⊆ { t : T | q t ≠ 0 } := by simp [Function.support] -- rw [<- tsum_subtype_eq_of_support_subset K1] at Hsumeq @@ -577,19 +677,6 @@ theorem privPostProcess_zCDPBound {nq : List T → SLang U} {HNorm : NormalMecha -- -- exact lt_of_le_of_lt A B - - -def privPostProcess_AC {f : U -> V} (nq : List T → SLang U) (Hac : ACNeighbour nq) : ACNeighbour (privPostProcess nq f) := by - rw [ACNeighbour] at * - unfold AbsCts at * - intro l₁ l₂ Hn v - have Hac := Hac l₁ l₂ Hn - simp [privPostProcess] - intro Hppz i fi - apply Hac - apply Hppz - apply fi - /-- Postprocessing preserves zCDP -/ From 895e0d390974055dfc3192d7e7b4e3191fc9215d Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Tue, 25 Jun 2024 12:01:40 -0400 Subject: [PATCH 093/122] checkpoint reduction: needs neighbours to be symmetric --- SampCert/DifferentialPrivacy/Neighbours.lean | 10 ++++ .../ZeroConcentrated/Postprocessing.lean | 52 ++++++++++++++++--- 2 files changed, 55 insertions(+), 7 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Neighbours.lean b/SampCert/DifferentialPrivacy/Neighbours.lean index fe541bba..7c4c57b0 100644 --- a/SampCert/DifferentialPrivacy/Neighbours.lean +++ b/SampCert/DifferentialPrivacy/Neighbours.lean @@ -20,3 +20,13 @@ 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₂ + + +def Neighbour_symm (l₁ l₂ : List T) (H : Neighbour l₁ l₂) : Neighbour l₂ l₁ := by + cases H + · rename_i _ _ _ Hl1 Hl2 + exact Neighbour.Deletion Hl2 Hl1 + · rename_i _ _ _ Hl1 Hl2 + exact Neighbour.Addition Hl2 Hl1 + · rename_i _ _ _ _ Hl1 Hl2 + exact Neighbour.Update Hl2 Hl1 diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean index d60f9984..d4a968cb 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean @@ -241,7 +241,8 @@ lemma rpow_nonzero (x : ENNReal) (y : ℝ) (H : ¬(x = 0 ∧ 0 < y ∨ x = ⊤ -- set_option pp.coercions false -theorem PostPocess_pre_reduct {nq : List T → SLang U} {HNorm : ∀ (l : List T), HasSum (nq l) 1} {ε₁ ε₂ : ℕ+} +theorem PostPocess_pre_reduct {U : Type} [m2 : MeasurableSpace U] [count : Countable U] [disc : DiscreteMeasurableSpace U] [Inhabited U] + {nq : List T → SLang U} {HNorm : ∀ (l : List T), HasSum (nq l) 1} (f : U → V) {α : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (Habs : AbsCts (nq l₁) (nq l₂)) (Hnq2 : ∀ (u : U), nq l₁ u ≠ 0) @@ -538,24 +539,57 @@ theorem tsum_ne_zero_of_ne_zero {T : Type} [Inhabited T] (f : T → ENNReal) (h have B := CONTRA default contradiction - +-- Note: Relies on the symmetry of neighbours theorem DPostPocess_pre {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) 1} {ε₁ ε₂ : ℕ+} (h : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) - (f : U → V) {α : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (Habs : AbsCts (nq l₁) (nq l₂)) (h2 : Neighbour l₁ l₂) : + (f : U → V) {α : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (Habs : AbsCts (nq l₁) (nq l₂)) (Habs' : AbsCts (nq l₂) (nq l₁)) : (∑' (x : V), (∑' (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 - -- First step is to reduce to the case where (nq l) is nonzero (reduct_1) + -- Restruct the sum to a type where nq l₁ is nonzero + have HSup1 (x : V) : Function.support (fun (a : U) => if x = f a then nq l₁ a else 0) ⊆ { u : U | nq l₁ u ≠ 0 } := by simp [Function.support] + have HSup2 (x : V) : Function.support (fun (a : U) => if x = f a then nq l₂ a else 0) ⊆ { u : U | nq l₁ u ≠ 0 } := by + simp [Function.support] + exact fun a a_1 a_2 a_3 => a_2 (Habs' a a_3) + have HSup3 : Function.support (fun (a : U) => nq l₁ a ^ α * nq l₂ a ^ (1 - α)) ⊆ { u : U | nq l₁ u ≠ 0 } := by + simp only [Function.support, Set.setOf_subset_setOf] + intro a Hnz + apply mul_ne_zero_iff.mp at Hnz + rcases Hnz with ⟨ H1 , _ ⟩ + intro Hk + apply H1 + apply (ENNReal.rpow_eq_zero_iff).mpr + left + apply And.intro + · apply Hk + · linarith + rw [<- tsum_subtype_eq_of_support_subset HSup3] + conv => + lhs + arg 1 + intro v + rw [<- tsum_subtype_eq_of_support_subset (HSup1 v)] + rw [<- tsum_subtype_eq_of_support_subset (HSup2 v)] + clear HSup1 + clear HSup2 + clear HSup3 + simp + + -- Now rewrite nq and f to be reduced functions over the type {x // ¬ nq l₁ x = 0} + + -- #check @PostPocess_pre_reduct T {x // ¬ nq l₁ x = 0} V ?TC1 ?TC2 ?TC3 ?TC4 + + sorry + -- First step is to reduce to the case where (nq l₁) is nonzero (reduct_1) - -- have K1_pre (x : V) : Function.support (fun (a : U) => if x = f a then nq l₁ a else 0) ⊆ { t : T | nq t ≠ 0 } := by -- have K2 : Function.support (fun x : T => (p x / q x)^α * q x) ⊆ { t : T | q t ≠ 0 } := by simp [Function.support] -- rw [<- tsum_subtype_eq_of_support_subset K1] at Hsumeq -- rw [<- tsum_subtype_eq_of_support_subset K2] at Hsumeq -- simp at Hsumeq - sorry + -- repeat rw [<- condition_to_subset] theorem privPostProcess_zCDPBound {nq : List T → SLang U} {HNorm : NormalMechanism nq} {ε₁ ε₂ : ℕ+} (h : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) (f : U → V) (Hac : ACNeighbour nq) : @@ -598,8 +632,12 @@ theorem privPostProcess_zCDPBound {nq : List T → SLang U} {HNorm : NormalMecha linarith · exact EReal.coe_lt_top (α - 1)⁻¹ apply elog_mono_le.mp - apply (DPostPocess_pre h f h1 ?Gabs h2) + apply (DPostPocess_pre h f h1 ?Gabs ?Gabs') case Gabs => exact Hac l₁ l₂ h2 + case Gabs' => + apply Hac + apply Neighbour_symm + apply h2 From 5db5fd1ffe0c1de039577bdac66031552d54b025 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Tue, 25 Jun 2024 13:29:15 -0400 Subject: [PATCH 094/122] All major theorems completed --- .../ZeroConcentrated/Postprocessing.lean | 180 ++++++++---------- 1 file changed, 79 insertions(+), 101 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean index d4a968cb..6d3a00a7 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean @@ -242,8 +242,10 @@ lemma rpow_nonzero (x : ENNReal) (y : ℝ) (H : ¬(x = 0 ∧ 0 < y ∨ x = ⊤ -- set_option pp.coercions false theorem PostPocess_pre_reduct {U : Type} [m2 : MeasurableSpace U] [count : Countable U] [disc : DiscreteMeasurableSpace U] [Inhabited U] - {nq : List T → SLang U} {HNorm : ∀ (l : List T), HasSum (nq l) 1} + {nq : List T → SLang U} (f : U → V) {α : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} + (HNorm1 : HasSum (nq l₁) 1) + (HNorm2 : HasSum (nq l₂) 1) (Habs : AbsCts (nq l₁) (nq l₂)) (Hnq2 : ∀ (u : U), nq l₁ u ≠ 0) (h2 : Neighbour l₁ l₂) : @@ -258,9 +260,8 @@ theorem PostPocess_pre_reduct {U : Type} [m2 : MeasurableSpace U] [count : Count apply HK -- Noised query is not ⊤ - have nq_nts : ∀ l : List T, ∑' n : U, nq l n ≠ ⊤ := by - intro l - simp [HasSum.tsum_eq (HNorm l)] + have nq_nts1 : ∑' n : U, nq l₁ n ≠ ⊤ := by simp [HasSum.tsum_eq HNorm1] + have nq_nts2 : ∑' n : U, nq l₂ n ≠ ⊤ := by simp [HasSum.tsum_eq HNorm2] -- Rewrite as cascading expectations rw [RenyiDivergenceExpectation] @@ -292,8 +293,8 @@ theorem PostPocess_pre_reduct {U : Type} [m2 : MeasurableSpace U] [count : Count rename_i NotEmpty repeat rw [condition_to_subset] - have nq_restriction_nts (l' : List T) : ∑' (a : ↑{a | i = f a}), nq l' ↑a ≠ ⊤ := by - exact convergent_subset (fun y => f y) (nq_nts l') + have nq_restriction_nts1 : ∑' (a : ↑{a | i = f a}), nq l₁ ↑a ≠ ⊤ := by exact convergent_subset (fun y => f y) nq_nts1 + have nq_restriction_nts2 : ∑' (a : ↑{a | i = f a}), nq l₂ ↑a ≠ ⊤ := by exact convergent_subset (fun y => f y) nq_nts2 have nq_restriction_nzs (l' : List T) (Hl : ∀ u : U, nq l' u ≠ 0) : ∑' (a : ↑{a | i = f a}), nq l' ↑a ≠ 0 := by simp @@ -306,10 +307,10 @@ theorem PostPocess_pre_reduct {U : Type} [m2 : MeasurableSpace U] [count : Count . exact Hl z simp - let δF₁ := (δpmf (nq l₁) f i (nq_restriction_nzs l₁ Hnq2) (nq_restriction_nts l₁)) - let δF₂ := (δpmf (nq l₂) f i (nq_restriction_nzs l₂ Hnql1) (nq_restriction_nts l₂)) - have δF₁_Eq : δF₁ = (δpmf (nq l₁) f i (nq_restriction_nzs l₁ Hnq2) (nq_restriction_nts l₁)) := by exact rfl - have δF₂_Eq : δF₂ = (δpmf (nq l₂) f i (nq_restriction_nzs l₂ Hnql1) (nq_restriction_nts l₂)) := by exact rfl + let δF₁ := (δpmf (nq l₁) f i (nq_restriction_nzs l₁ Hnq2) nq_restriction_nts1) + let δF₂ := (δpmf (nq l₂) f i (nq_restriction_nzs l₂ Hnql1) nq_restriction_nts2) + have δF₁_Eq : δF₁ = (δpmf (nq l₁) f i (nq_restriction_nzs l₁ Hnq2) nq_restriction_nts1) := by exact rfl + have δF₂_Eq : δF₂ = (δpmf (nq l₂) f i (nq_restriction_nzs l₂ Hnql1) nq_restriction_nts2) := by exact rfl -- Normalized fibers are absolutely continuous have HAC_Fiber : AbsCts δF₁ δF₂ := by @@ -329,13 +330,13 @@ theorem PostPocess_pre_reduct {U : Type} [m2 : MeasurableSpace U] [count : Count apply Habs apply Hl2z · exfalso - apply nq_restriction_nts l₂ + apply nq_restriction_nts2 simp assumption have δF₂_NT (x : { x // i = f x }) : δF₂ x ≠ ⊤ := by rw [δF₂_Eq] - exact PMF.apply_ne_top ((nq l₂).δpmf f i (nq_restriction_nzs l₂ Hnql1) (nq_restriction_nts l₂)) x + exact PMF.apply_ne_top ((nq l₂).δpmf f i (nq_restriction_nzs l₂ Hnql1) (nq_restriction_nts2)) x -- Normalized fibers avoid the bad case for rewriting the Jensen term have Hspecial (x : { x // i = f x }) : ¬(δF₁ x ≠ 0 ∧ δF₂ x = ⊤) := by @@ -367,8 +368,8 @@ theorem PostPocess_pre_reduct {U : Type} [m2 : MeasurableSpace U] [count : Count have N_def (l : List T) : N l = (∑' (x : {n // i = f n}), nq l x)⁻¹ := by exact rfl have N_inv (l : List T) : (∑' (x : {n // i = f n}), nq l x) = (N l)⁻¹ := by exact Eq.symm (inv_inv (∑' (x : { n // i = f n }), (nq l) ↑x)) - have N1_nz : N l₁ ≠ 0 := ENNReal.inv_ne_zero.mpr (nq_restriction_nts l₁) - have N2_nz : N l₂ ≠ 0 := ENNReal.inv_ne_zero.mpr (nq_restriction_nts l₂) + have N1_nz : N l₁ ≠ 0 := ENNReal.inv_ne_zero.mpr (nq_restriction_nts1) + have N2_nz : N l₂ ≠ 0 := ENNReal.inv_ne_zero.mpr (nq_restriction_nts2) -- Unfold normalization constants in HJ conv at HJ => @@ -398,14 +399,14 @@ theorem PostPocess_pre_reduct {U : Type} [m2 : MeasurableSpace U] [count : Count rw [ENNReal.mul_rpow_of_ne_top ?G1 ?G2] case G1 => apply mul_ne_top - · exact ENNReal.ne_top_of_tsum_ne_top (nq_nts l₁) ↑x + · exact ENNReal.ne_top_of_tsum_ne_top (nq_nts1) ↑x · exact inv_ne_top.mpr (nq_restriction_nzs l₁ Hnq2) case G2 => apply inv_ne_top.mpr exact mul_ne_zero (Hnql1 ↑x) N2_nz rw [ENNReal.mul_rpow_of_ne_top ?G1 ?G2] case G1 => - exact ENNReal.ne_top_of_tsum_ne_top (nq_nts l₁) ↑x + exact ENNReal.ne_top_of_tsum_ne_top (nq_nts1) ↑x case G2 => exact inv_ne_top.mpr (nq_restriction_nzs l₁ Hnq2) rw [ENNReal.mul_rpow_of_ne_top ?G1 ?G2] @@ -415,7 +416,7 @@ theorem PostPocess_pre_reduct {U : Type} [m2 : MeasurableSpace U] [count : Count exact inv_ne_top.mpr N2_nz rw [ENNReal.mul_rpow_of_ne_top ?G1 ?G2] case G1 => - exact ENNReal.ne_top_of_tsum_ne_top (nq_nts l₁) ↑x + exact ENNReal.ne_top_of_tsum_ne_top (nq_nts1) ↑x case G2 => exact inv_ne_top.mpr (Hnql1 ↑x) conv => @@ -540,9 +541,9 @@ theorem tsum_ne_zero_of_ne_zero {T : Type} [Inhabited T] (f : T → ENNReal) (h contradiction -- Note: Relies on the symmetry of neighbours -theorem DPostPocess_pre {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) 1} {ε₁ ε₂ : ℕ+} - (h : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) - (f : U → V) {α : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (Habs : AbsCts (nq l₁) (nq l₂)) (Habs' : AbsCts (nq l₂) (nq l₁)) : +theorem DPostPocess_pre {nq : List T → SLang U} (HNorm : ∀ l, HasSum (nq l) 1) + (f : U → V) {α : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (HN : Neighbour l₁ l₂) + (Habs : AbsCts (nq l₁) (nq l₂)) (Habs' : AbsCts (nq l₂) (nq l₁)) : (∑' (x : V), (∑' (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 - α)) ≤ @@ -551,7 +552,7 @@ theorem DPostPocess_pre {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) have HSup1 (x : V) : Function.support (fun (a : U) => if x = f a then nq l₁ a else 0) ⊆ { u : U | nq l₁ u ≠ 0 } := by simp [Function.support] have HSup2 (x : V) : Function.support (fun (a : U) => if x = f a then nq l₂ a else 0) ⊆ { u : U | nq l₁ u ≠ 0 } := by simp [Function.support] - exact fun a a_1 a_2 a_3 => a_2 (Habs' a a_3) + exact fun a _ a_2 a_3 => a_2 (Habs' a a_3) have HSup3 : Function.support (fun (a : U) => nq l₁ a ^ α * nq l₂ a ^ (1 - α)) ⊆ { u : U | nq l₁ u ≠ 0 } := by simp only [Function.support, Set.setOf_subset_setOf] intro a Hnz @@ -576,20 +577,59 @@ theorem DPostPocess_pre {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) clear HSup3 simp - -- Now rewrite nq and f to be reduced functions over the type {x // ¬ nq l₁ x = 0} - - -- #check @PostPocess_pre_reduct T {x // ¬ nq l₁ x = 0} V ?TC1 ?TC2 ?TC3 ?TC4 - - sorry - -- First step is to reduce to the case where (nq l₁) is nonzero (reduct_1) - - - -- have K2 : Function.support (fun x : T => (p x / q x)^α * q x) ⊆ { t : T | q t ≠ 0 } := by simp [Function.support] - -- rw [<- tsum_subtype_eq_of_support_subset K1] at Hsumeq - -- rw [<- tsum_subtype_eq_of_support_subset K2] at Hsumeq - -- simp at Hsumeq - - -- repeat rw [<- condition_to_subset] + cases (Classical.em (∃ x : U, ¬ nq l₁ x = 0)) + · rename_i x_witness + have HR := @PostPocess_pre_reduct T V {x // ¬ nq l₁ x = 0} + _ _ _ ?TC + (fun t u => nq t u) (fun x => f x) α h1 + l₁ l₂ ?GNorm1 ?GNorm2 ?Gac ?Gnz HN + case TC => + apply inhabited_of_nonempty + exact Set.Nonempty.to_subtype x_witness + case GNorm1 => + simp + rw [<- HasSum.tsum_eq (HNorm l₁)] + have HSup4 : Function.support (fun u => nq l₁ u) ⊆ { u | nq l₁ u ≠ 0 } := by simp [Function.support] + rw [<- tsum_subtype_eq_of_support_subset HSup4] + simp + apply Summable.hasSum + exact ENNReal.summable + case GNorm2 => + simp + rw [<- HasSum.tsum_eq (HNorm l₂)] + have HSup4 : Function.support (fun u => nq l₂ u) ⊆ { u | nq l₁ u ≠ 0 } := by + simp [Function.support] + intro a Hk1 Hk2 + apply Hk1 + apply Habs' + apply Hk2 + rw [<- tsum_subtype_eq_of_support_subset HSup4] + simp + apply Summable.hasSum + exact ENNReal.summable + case Gac => + rw [AbsCts] + simp + intro a _ H2 + exact Habs a (Habs' a (Habs a H2)) + case Gnz => simp + simp at HR + apply HR + · rename_i x_empty + simp at * + have Hempty : IsEmpty {x // ¬nq l₁ x = 0} := by + exact Subtype.isEmpty_of_false fun a a_1 => a_1 (x_empty a) + rw [@tsum_empty _ _ _ _ _ Hempty] + conv => + lhs + arg 1 + intro + rw [@tsum_empty _ _ _ _ _ Hempty] + rw [@tsum_empty _ _ _ _ _ Hempty] + simp + intro + left + linarith theorem privPostProcess_zCDPBound {nq : List T → SLang U} {HNorm : NormalMechanism nq} {ε₁ ε₂ : ℕ+} (h : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) (f : U → V) (Hac : ACNeighbour nq) : @@ -632,73 +672,11 @@ theorem privPostProcess_zCDPBound {nq : List T → SLang U} {HNorm : NormalMecha linarith · exact EReal.coe_lt_top (α - 1)⁻¹ apply elog_mono_le.mp - apply (DPostPocess_pre h f h1 ?Gabs ?Gabs') - case Gabs => exact Hac l₁ l₂ h2 - case Gabs' => - apply Hac - apply Neighbour_symm - apply h2 - - - - /- - -- 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) + apply (DPostPocess_pre HNorm (fun a => f a) h1 h2 (Hac l₁ l₂ h2)) + apply Hac + apply Neighbour_symm + apply h2 - 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 - sorry - -/ -- theorem privPostProcess_NonTopRDNQ {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) 1} {ε₁ ε₂ : ℕ+} (f : U → V) -- (dp : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) (nt : NonTopRDNQ nq) (nz : NonZeroNQ nq) (nts : NonTopNQ nq) (ntsum: NonTopSum nq) : From 782b41d43adc498c4e5735863d671371a65f419a Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Tue, 25 Jun 2024 16:55:21 -0400 Subject: [PATCH 095/122] checkpoint conversion lemmas --- SampCert/Util/Log.lean | 207 ++++++++++++++++++++++++++++++++++++----- 1 file changed, 185 insertions(+), 22 deletions(-) diff --git a/SampCert/Util/Log.lean b/SampCert/Util/Log.lean index 5880e934..77d8d51d 100644 --- a/SampCert/Util/Log.lean +++ b/SampCert/Util/Log.lean @@ -45,6 +45,7 @@ lemma EReal_cases (w : EReal) : w = ⊥ ∨ w = ⊤ ∨ (∃ v : ℝ, w = v) := right exists wR + @[simp] lemma ofEReal_bot : ofEReal ⊥ = 0 := by simp [ofEReal] @@ -91,16 +92,25 @@ variable {x y : ENNReal} variable {w z : EReal} variable {r : Real} +set_option pp.coercions false + @[simp] lemma elog_of_pos_real (H : 0 < r) : elog (ENNReal.ofReal r) = Real.log r := by rw [elog] split - · sorry + · simp at * · split - · rename_i r heq h - simp - sorry - · sorry + · rename_i r' heq h + exfalso + rw [h] at heq + simp at heq + linarith + · rename_i h' r' heq h + simp_all + congr + simp [ENNReal.ofReal] at heq + rw [<- heq] + exact (Real.coe_toNNReal r (le_of_lt H)) @[simp] lemma elog_zero : elog (ENNReal.ofReal 0) = ⊥ := by simp [elog] @@ -117,7 +127,6 @@ lemma eexp_top : eexp ⊤ = ⊤ := by simp [eexp] @[simp] lemma eexp_zero : eexp 0 = 1 := by simp [eexp] - @[simp] lemma eexp_ofReal : eexp r = ENNReal.ofReal (Real.exp r) := by simp [ENNReal.ofReal, eexp, elog] @@ -163,15 +172,93 @@ lemma eexp_elog : (elog (eexp w)) = w := by rename_i v'' simp [ENNReal.ofReal] split - · -- exp is nonnegative - sorry - · sorry + · rename_i Hcont + have Hcont' : 0 < rexp v'' := by exact exp_pos v'' + linarith + · rename_i H + have RW : (max (rexp v'') 0) = (rexp v'') := by + apply max_eq_left_iff.mpr + linarith + simp [RW] + clear RW + simp [Real.toEReal] @[simp] lemma elog_mul : elog x + elog y = elog (x * y) := by + rcases x with _ | ⟨ rx , Hrx ⟩ <;> + rcases y with _ | ⟨ ry , Hry ⟩ <;> + simp_all + · rcases (LE.le.lt_or_eq Hry) with Hry' | Hry' + · simp [elog] + split + · exfalso + rename_i h + cases h + linarith + · simp_all + · simp [top_mul'] + split + · simp_all [elog] + · exfalso + rename_i H + apply H + congr + rw [Hry'] + · rcases (LE.le.lt_or_eq Hrx) with Hrx' | Hrx' + · simp [elog] + split + · exfalso + rename_i h + cases h + linarith + · simp_all + · simp [mul_top'] + split + · simp_all [elog] + · exfalso + rename_i H + apply H + congr + rw [Hrx'] + · rcases (LE.le.lt_or_eq Hry) with Hry' | Hry' <;> + rcases (LE.le.lt_or_eq Hrx) with Hrx' | Hrx' + · rw [ENNReal.ofNNReal] + simp_all [elog] + sorry + · sorry + · simp [elog] + split + · rename_i H + rw [H] + simp + · rename_i H + split + · rename_i H' + rw [H'] + simp + · split + · exfalso + simp_all [or_self] + · simp_all + sorry + · simp [elog] + split + · split + · rename_i H1 H2 + rw [H1] + simp + · exfalso + rename_i H + apply H + congr + rw [Hry'] + · exfalso + rename_i H + apply H + congr + rw [Hrx'] - sorry -- checked truth table @[simp] lemma eexp_add : eexp w * eexp z = eexp (w + z) := by sorry -- checked truth table @@ -186,20 +273,35 @@ lemma eexp_injective : eexp w = eexp z -> w = z := by cases w <;> cases z <;> try tauto · rename_i v cases v <;> simp at * - sorry + rename_i v' + have Hv' := exp_pos v' + linarith · rename_i v cases v <;> simp at * - sorry + rename_i v' + have Hv' := exp_pos v' + linarith · rename_i v₁ v₂ cases v₁ <;> cases v₂ <;> simp at * congr - sorry + rename_i v₁' v₂' + simp [ENNReal.ofReal] at H + apply NNReal.coe_inj.mpr at H + simp at H + have RW (r : ℝ) : (max (rexp r) 0) = (rexp r) := by + apply max_eq_left_iff.mpr + exact exp_nonneg r + rw [RW v₁'] at H + rw [RW v₂'] at H + exact exp_eq_exp.mp H + lemma elog_injective : elog x = elog y -> x = y := by rw [elog, elog] cases x <;> cases y <;> try tauto · rename_i v cases v; simp at * + sorry · rename_i v cases v; simp at * @@ -226,11 +328,38 @@ lemma elog_mono_le : (x <= y) <-> elog x <= elog y := by sorry -- Specialized lemmas for ofEReal when its argument is nonnegative (so no truncation happens) -lemma ofEReal_nonneg_eq_iff (Hw : 0 <= w) (Hz : 0 <= z) : w = z <-> (ofEReal w = ofEReal z) := - sorry - -lemma ofEReal_le_mono_nonneg (Hw : 0 ≤ w) : w ≤ z <-> (ofEReal w ≤ ofEReal z) := - sorry +lemma ofEReal_nonneg_eq_iff (Hw : 0 <= w) (Hz : 0 <= z) : w = z <-> (ofEReal w = ofEReal z) := by + apply Iff.intro + · intro H1 + rw [H1] + · intro H1 + rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> + rcases (EReal_cases z) with Hz' | (Hz' | ⟨ z', Hz' ⟩) <;> + simp_all + +lemma ofEReal_le_mono_nonneg (Hw : 0 ≤ w) : w ≤ z <-> (ofEReal w ≤ ofEReal z) := by + apply Iff.intro + · intro H1 + rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> + rcases (EReal_cases z) with Hz' | (Hz' | ⟨ z', Hz' ⟩) <;> + simp_all + exact ofReal_le_ofReal H1 + · intro H1 + rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> + rcases (EReal_cases z) with Hz' | (Hz' | ⟨ z', Hz' ⟩) <;> + simp_all + · sorry + · simp [ENNReal.ofReal] at H1 + apply toNNReal_le_toNNReal_iff'.mp at H1 + cases H1 + · assumption + · rename_i H + have H' : w' = 0 := by apply le_antisymm <;> assumption + simp_all + apply EReal.coe_nonneg.mp + simp_all + -- ??? + sorry -- Use above lemma ofEReal_nonneg_zero (Hz : 0 ≤ z) : (0 = z) <-> 0 = ofEReal z := sorry @@ -274,18 +403,48 @@ lemma ofEReal_mul_nonneg (Hw : 0 ≤ w) (Hz : 0 ≤ z) : ofEReal (w * z) = ofERe lemma ofEReal_nonneg_scal_l (H1 : 0 < r) (H2 : 0 ≤ r * w) : 0 ≤ w := by rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) · simp_all - sorry + rw [EReal.mul_bot_of_pos] at H2 + all_goals simp_all · simp_all · simp_all sorry @[simp] -lemma toEReal_ofENNReal_nonneg (H : 0 ≤ w) : ENNReal.toEReal (ofEReal w) = w := sorry +lemma toEReal_ofENNReal_nonneg (H : 0 ≤ w) : ENNReal.toEReal (ofEReal w) = w := by + rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> simp_all @[simp] -lemma ofEReal_toENNReal : ofEReal (ENNReal.toEReal x) = x := by sorry +lemma ofEReal_toENNReal : ofEReal (ENNReal.toEReal x) = x := by + rw [ENNReal.ofEReal] + split + · rename_i x' Hx' + simp [toEReal] at Hx' + cases x <;> simp_all + rename_i x'' + exfalso + rw [NNReal.toReal] at Hx' + rw [Real.toEReal] at Hx' + simp_all + · rename_i x' Hx' + simp [toEReal] at Hx' + cases x <;> simp_all + rw [NNReal.toReal] at Hx' + rw [Real.toEReal] at Hx' + simp_all + cases Hx' + · rename_i e x' Hx' + simp [toEReal] at Hx' + cases x <;> simp_all + · cases Hx' + · rw [ENNReal.ofReal] + simp_all + rw [NNReal.toReal] at Hx' + rw [Real.toEReal] at Hx' + simp_all + cases Hx' + simp_all lemma ofEReal_ofReal_toENNReal : ENNReal.ofEReal (Real.toEReal r) = ENNReal.ofReal r := by simp [ofEReal, Real.toEReal, ENNReal.ofReal] @@ -444,7 +603,11 @@ lemma ereal_smul_eq_left (s : EReal) (Hr1 : 0 < s) (Hr2 : s < ⊤) (H : s * w = cases H · repeat rw [<- EReal.coe_mul] at H apply EReal.coe_eq_coe_iff.mp at H - sorry + simp_all only [mul_eq_mul_left_iff] + cases H + · assumption + · exfalso + linarith lemma ereal_smul_left_le (s : EReal) (Hr1 : 0 < s) (Hr2 : s < ⊤) (H : w ≤ z) : s * w ≤ s * z := by rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> From 0517c461bc5a5235bfe679c22f9da87eec97a548 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Wed, 26 Jun 2024 13:51:33 -0400 Subject: [PATCH 096/122] case analysis tactics for remaining log lemmas --- .../DifferentialPrivacy/RenyiDivergence.lean | 9 +- SampCert/Util/Log.lean | 1107 +++++++++-------- 2 files changed, 620 insertions(+), 496 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index ebabfe2f..e79c413d 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -1160,10 +1160,11 @@ theorem RenyiDivergence_aux_zero [MeasurableSpace T] [MeasurableSingletonClass T · intro H apply (RenyiDivergence_def_eq_0_iff p q Hα Hac).mp symm - apply (ofEReal_nonneg_zero ?G1).mpr - case G1 => - exact RenyiDivergence_def_nonneg p q Hac Hα - exact id (Eq.symm H) + rw [RenyiDivergence] at H + have H' := RenyiDivergence_def_nonneg p q Hac Hα + refine (ofEReal_nonneg_eq_iff ?mpr.Hw H').mpr ?mpr.a + · simp + simp [H] -- Unused /- diff --git a/SampCert/Util/Log.lean b/SampCert/Util/Log.lean index 77d8d51d..854d4504 100644 --- a/SampCert/Util/Log.lean +++ b/SampCert/Util/Log.lean @@ -8,6 +8,8 @@ import Mathlib.Data.ENNReal.Basic import Mathlib.Data.Real.EReal import Mathlib.Analysis.SpecialFunctions.Log.Basic import Mathlib.Analysis.SpecialFunctions.Pow.NNReal +import Lean.Elab.Tactic + /-! # Logarithm on ENNReal @@ -20,18 +22,87 @@ noncomputable section open scoped Classical open ENNReal EReal Real -namespace ENNReal + + +section EReal_conv_cases +/-! +### Case analysis lemmas + +Most conversion proofs follow by splitting into cases, and then simplifying. However +explicitly performing case analysis can be unwieldy and lead to lots of duplicate work, +depending on which simplification rules are used. These tactics allow us to fine-tune the +case splits at the start of a conversion proof in order to reduce the number of cases we must +prove by hand. +-/ + /-- -Coerce a EReal to an ENNReal by truncation +A real number is either zero or nonzero -/ -noncomputable def ofEReal (e : EReal) : ENNReal := - match e with - | ⊥ => some 0 - | ⊤ => ⊤ - | some (some r) => ENNReal.ofReal r +lemma Real_cases_zero (r : ℝ) : r = 0 ∨ r ≠ 0 := by + exact eq_or_ne r (OfNat.ofNat 0) + +syntax "case_Real_zero" term : tactic +macro_rules +| `(tactic| case_Real_zero $r:term ) => + `(tactic| rcases (eq_or_ne $r (OfNat.ofNat 0)) with _ | _ <;> try simp_all) + + + +/-- +A real number is either negative, zero, or postive +-/ +lemma Real_cases_sign (r : ℝ) : r < 0 ∨ r = 0 ∨ 0 < r := by exact lt_trichotomy r (OfNat.ofNat 0) + +syntax "case_Real_sign" term : tactic +macro_rules +| `(tactic| case_Real_sign $r:term ) => + `(tactic| rcases (Real_cases_sign $r) with _ | _ | _ <;> try simp_all) + +/-- +A nonnegative number is either zero or positive +-/ +syntax "case_nonneg_zero" term : tactic +macro_rules +| `(tactic| case_nonneg_zero $H:term ) => + `(tactic| rcases (LE.le.eq_or_gt $H) with _ | _ <;> try simp_all) + + + +/-- +A real number is either negative, or nonzero +-/ +lemma Real_cases_nonnegative (r : ℝ) : r < 0 ∨ 0 ≤ r := by exact lt_or_le r (OfNat.ofNat 0) + +syntax "case_Real_nonnegative " term : tactic +macro_rules +| `(tactic| case_Real_nonnegative $r:term ) => + `(tactic| rcases (Real_cases_nonnegative $r) with _ | _ <;> try simp_all) + + +lemma ENNReal_isReal_cases (x : ENNReal) : x = ⊤ ∨ (∃ v : ℝ, x = ENNReal.ofReal v ∧ 0 ≤ v) := by + cases x + · left + simp + . right + rename_i v + rcases v with ⟨ r, Hr ⟩ + exists r + apply And.intro + · simp [ENNReal.ofReal, Real.toNNReal] + congr + rw [max_eq_left Hr] + · assumption -lemma EReal_cases (w : EReal) : w = ⊥ ∨ w = ⊤ ∨ (∃ v : ℝ, w = v) := by +syntax "case_ENNReal_isReal" term : tactic +macro_rules +| `(tactic| case_ENNReal_isReal $w:term ) => + `(tactic| rcases (ENNReal_isReal_cases $w) with _ | ⟨ _, _, _⟩ <;> try simp_all) + +/-- +An EReal is either ⊤, ⊥, or the lift of some real number. +-/ +lemma EReal_isReal_cases (w : EReal) : w = ⊥ ∨ w = ⊤ ∨ (∃ v : ℝ, w = Real.toEReal v) := by cases w · left rfl @@ -45,6 +116,51 @@ lemma EReal_cases (w : EReal) : w = ⊥ ∨ w = ⊤ ∨ (∃ v : ℝ, w = v) := right exists wR +syntax "case_EReal_isReal" term : tactic +macro_rules +| `(tactic| case_EReal_isReal $w:term ) => + `(tactic| rcases (EReal_isReal_cases $w) with _ | _ | ⟨ _, _ ⟩ <;> try simp_all) + +/-- +An EReal is either negative, or the lift of an ENNReal +-/ +lemma EReal_isENNReal_cases (w : EReal) : (w < 0) ∨ (∃ v : ENNReal, w = ENNReal.toEReal v) := by + case_EReal_isReal w + · exists ⊤ + · rename_i w' Hw' + case_Real_nonnegative w' + rename_i Hw'' + right + exists (ENNReal.ofReal w') + simp only [coe_ennreal_ofReal, EReal.coe_eq_coe_iff] + rw [max_eq_left Hw''] + +syntax "case_EReal_isENNReal" term : tactic +macro_rules +| `(tactic| case_EReal_isENNReal $w:term ) => + `(tactic| rcases (EReal_isENNReal_cases $w) with _ | ⟨ _, _ ⟩ <;> try simp_all) + + +end EReal_conv_cases + + + +namespace ENNReal + +section ofEReal +/-! +### Coercion from EReals to ENNreals +-/ + + +/-- +Truncate an `EReal` to an `ENNReal` +-/ +noncomputable def ofEReal (e : EReal) : ENNReal := + match e with + | ⊥ => some 0 + | ⊤ => ⊤ + | some (some r) => ENNReal.ofReal r @[simp] lemma ofEReal_bot : ofEReal ⊥ = 0 := by simp [ofEReal] @@ -55,17 +171,125 @@ lemma ofEReal_top : ofEReal ⊤ = ⊤ := by simp [ofEReal] @[simp] lemma ofEReal_zero : ofEReal 0 = 0 := by simp [ofEReal] +@[simp] +lemma ofEReal_real (r : ℝ) : ofEReal r = ENNReal.ofReal r := by simp [Real.toEReal, ofEReal] + +-- Not sure if I want this to be an actual coercion or not, but I'm leaning towards no because +-- of the truncation. +-- instance : Coe EReal ENNReal := ⟨ofEReal⟩ + +-- MARKUSDE: formerly: ofEReal_nonpos +lemma ofEReal_eq_zero_iff (w : EReal) : w ≤ 0 <-> ofEReal w = 0 := by + apply Iff.intro + · intro _ + case_EReal_isReal w + · intro _ + case_EReal_isReal w + + + +-- MARKUSDE: Rename me to ofEReal_nonneg_inj +lemma ofEReal_nonneg_eq_iff {w z : EReal} (Hw : 0 <= w) (Hz : 0 <= z) : + w = z <-> (ofEReal w = ofEReal z) := by + apply Iff.intro + · intro _ + simp_all + · intro Heq + all_goals case_EReal_isReal w + all_goals case_EReal_isReal z + +set_option pp.coercions false +-- MARKUSDE: Rename me @[simp] -lemma ofEReal_real (r : ℝ) : ofEReal r = ENNReal.ofReal r := by - simp [Real.toEReal] - simp [ofEReal] +lemma toEReal_ofENNReal_nonneg {w : EReal} (H : 0 ≤ w) : ENNReal.toEReal (ofEReal w) = w := by case_EReal_isReal w -lemma ofEReal_nonpos (w : EReal) (HW : w ≤ 0): ofEReal w = 0 := by - rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> simp_all +-- MARKUSDE: Rename me +@[simp] +lemma ofEReal_toENNReal {x : ENNReal} : ofEReal (ENNReal.toEReal x) = x := by case_ENNReal_isReal x +-- MARKUSDE: Rename me +@[simp] +lemma ofEReal_ofReal_toENNReal : ENNReal.ofEReal (Real.toEReal r) = ENNReal.ofReal r := by + simp [ofEReal, Real.toEReal, ENNReal.ofReal] --- instance : Coe EReal ENNReal := ⟨ofEReal⟩ + +lemma ofEReal_le_mono {w z : EReal} (H : w ≤ z) : ofEReal w ≤ ofEReal z := by + all_goals case_EReal_isReal w + all_goals case_EReal_isReal z + apply ofReal_le_ofReal + assumption + + +lemma ofEReal_le_mono_conv_nonneg {w z : EReal} (Hw : 0 ≤ w) (Hle : ofEReal w ≤ ofEReal z) : w ≤ z := by + -- Reduce to comparison between ENNReals? + + -- case_EReal_isENNReal w + -- · -- w < 0 => ofEReal w = 0 + -- rw [(ofEReal_eq_zero_iff w).mp ?G0] at Hle + -- case G0 => exact le_of_lt Hzlz + -- skip + -- sorry + -- rename_i w' Hw' + -- case_EReal_isENNReal z + -- · rename_i Hzlz + -- -- z < 0 => ofEReal z = 0 + -- rw [(ofEReal_eq_zero_iff z).mp ?G1] at Hle + -- case G1 => exact le_of_lt Hzlz + -- skip + -- sorry + -- rename_i z' Hz' + sorry + + + @[simp] + lemma ofEReal_plus_nonneg (Hw : 0 ≤ w) (Hz : 0 ≤ z) : ofEReal (w + z) = ofEReal w + ofEReal z := by + all_goals case_EReal_isReal w + all_goals case_EReal_isReal z + rename_i w' z' _ _ + rw [← EReal.coe_add] + rw [ofEReal_ofReal_toENNReal] + rw [ofReal_add Hw Hz] + + +@[simp] +lemma ofEReal_mul_nonneg (Hw : 0 ≤ w) (Hz : 0 ≤ z) : ofEReal (w * z) = ofEReal w * ofEReal z := by + all_goals case_EReal_isReal w + all_goals case_EReal_isReal z + · rename_i r Hr Hz' + case_nonneg_zero Hz + rename_i Hr_nz + simp [top_mul_coe_of_pos Hr_nz] + · rename_i r Hr Hz' + case_nonneg_zero Hw + rename_i Hr_nz + -- simp [top_mul_coe_of_pos Hr_nz] + sorry + · rename_i r₁ r₂ Hr₁ Hr₂ + rw [← EReal.coe_mul] + rw [ofEReal_ofReal_toENNReal] + rw [ofReal_mul' Hz] + +-- ??? +lemma ofEReal_nonneg_scal_l {r : ℝ} {w : EReal} (H1 : 0 < r) (H2 : 0 ≤ r * w) : 0 ≤ w := by + sorry +-- lemma ofEReal_nonneg_scal_l (H1 : 0 < r) (H2 : 0 ≤ r * w) : 0 ≤ w := by +-- rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) +-- · simp_all +-- rw [EReal.mul_bot_of_pos] at H2 +-- all_goals simp_all +-- · simp_all +-- · simp_all +-- sorry +-- +-- + + + +end ofEReal + + +section elog_eexp /-- The extended logarithm @@ -88,29 +312,25 @@ def eexp (y : EReal) : ENNReal := | some (some r) => ENNReal.ofReal (Real.exp r) -variable {x y : ENNReal} -variable {w z : EReal} -variable {r : Real} - -set_option pp.coercions false @[simp] -lemma elog_of_pos_real (H : 0 < r) : elog (ENNReal.ofReal r) = Real.log r := by - rw [elog] - split - · simp at * - · split - · rename_i r' heq h - exfalso - rw [h] at heq - simp at heq - linarith - · rename_i h' r' heq h - simp_all - congr - simp [ENNReal.ofReal] at heq - rw [<- heq] - exact (Real.coe_toNNReal r (le_of_lt H)) +lemma elog_of_pos_real {r : ℝ} (H : 0 < r) : elog (ENNReal.ofReal r) = Real.log r := by + sorry + -- rw [elog] + -- split + -- · simp at * + -- · split + -- · rename_i r' heq h + -- exfalso + -- rw [h] at heq + -- simp at heq + -- linarith + -- · rename_i h' r' heq h + -- simp_all + -- congr + -- simp [ENNReal.ofReal] at heq + -- rw [<- heq] + -- exact (Real.coe_toNNReal r (le_of_lt H)) @[simp] lemma elog_zero : elog (ENNReal.ofReal 0) = ⊥ := by simp [elog] @@ -128,196 +348,202 @@ lemma eexp_top : eexp ⊤ = ⊤ := by simp [eexp] lemma eexp_zero : eexp 0 = 1 := by simp [eexp] @[simp] -lemma eexp_ofReal : eexp r = ENNReal.ofReal (Real.exp r) := by - simp [ENNReal.ofReal, eexp, elog] - rfl +lemma eexp_ofReal {r : ℝ} : eexp r = ENNReal.ofReal (Real.exp r) := by + sorry + -- simp [ENNReal.ofReal, eexp, elog] + -- rfl @[simp] -lemma elog_eexp : eexp (elog x) = x := by - rw [elog] - split - · simp - · rename_i _ r' - split - · simp - rename_i _ h - rw [h] - · rename_i _ H - simp - rw [NNReal.toReal] - simp - rw [Real.exp_log] - rw [ofReal_coe_nnreal] - rcases r' with ⟨ v , Hv ⟩ - apply lt_of_le_of_ne - · simpa - · simp - intro Hk - apply H - apply NNReal.coe_eq_zero.mp - simp - rw [Hk] +lemma elog_eexp {x : ENNReal} : eexp (elog x) = x := by + sorry + -- rw [elog] + -- split + -- · simp + -- · rename_i _ r' + -- split + -- · simp + -- rename_i _ h + -- rw [h] + -- · rename_i _ H + -- simp + -- rw [NNReal.toReal] + -- simp + -- rw [Real.exp_log] + -- rw [ofReal_coe_nnreal] + -- rcases r' with ⟨ v , Hv ⟩ + -- apply lt_of_le_of_ne + -- · simpa + -- · simp + -- intro Hk + -- apply H + -- apply NNReal.coe_eq_zero.mp + -- simp + -- rw [Hk] @[simp] -lemma eexp_elog : (elog (eexp w)) = w := by - cases w - · simp [eexp, elog] - rfl - · simp [eexp, elog] - rename_i v' - cases v' - · simp - rfl - · simp - rename_i v'' - simp [ENNReal.ofReal] - split - · rename_i Hcont - have Hcont' : 0 < rexp v'' := by exact exp_pos v'' - linarith - · rename_i H - have RW : (max (rexp v'') 0) = (rexp v'') := by - apply max_eq_left_iff.mpr - linarith - simp [RW] - clear RW - simp [Real.toEReal] +lemma eexp_elog {w : EReal} : (elog (eexp w)) = w := by + sorry + -- cases w + -- · simp [eexp, elog] + -- rfl + -- · simp [eexp, elog] + -- rename_i v' + -- cases v' + -- · simp + -- rfl + -- · simp + -- rename_i v'' + -- simp [ENNReal.ofReal] + -- split + -- · rename_i Hcont + -- have Hcont' : 0 < rexp v'' := by exact exp_pos v'' + -- linarith + -- · rename_i H + -- have RW : (max (rexp v'') 0) = (rexp v'') := by + -- apply max_eq_left_iff.mpr + -- linarith + -- simp [RW] + -- clear RW + -- simp [Real.toEReal] @[simp] -lemma elog_mul : elog x + elog y = elog (x * y) := by - rcases x with _ | ⟨ rx , Hrx ⟩ <;> - rcases y with _ | ⟨ ry , Hry ⟩ <;> - simp_all - · rcases (LE.le.lt_or_eq Hry) with Hry' | Hry' - · simp [elog] - split - · exfalso - rename_i h - cases h - linarith - · simp_all - · simp [top_mul'] - split - · simp_all [elog] - · exfalso - rename_i H - apply H - congr - rw [Hry'] - · rcases (LE.le.lt_or_eq Hrx) with Hrx' | Hrx' - · simp [elog] - split - · exfalso - rename_i h - cases h - linarith - · simp_all - · simp [mul_top'] - split - · simp_all [elog] - · exfalso - rename_i H - apply H - congr - rw [Hrx'] - · rcases (LE.le.lt_or_eq Hry) with Hry' | Hry' <;> - rcases (LE.le.lt_or_eq Hrx) with Hrx' | Hrx' - · rw [ENNReal.ofNNReal] - simp_all [elog] - sorry - · sorry - · simp [elog] - split - · rename_i H - rw [H] - simp - · rename_i H - split - · rename_i H' - rw [H'] - simp - · split - · exfalso - simp_all [or_self] - · simp_all - sorry - · simp [elog] - split - · split - · rename_i H1 H2 - rw [H1] - simp - · exfalso - rename_i H - apply H - congr - rw [Hry'] - · exfalso - rename_i H - apply H - congr - rw [Hrx'] +lemma elog_mul {x y : ENNReal} : elog x + elog y = elog (x * y) := by + sorry + -- rcases x with _ | ⟨ rx , Hrx ⟩ <;> + -- rcases y with _ | ⟨ ry , Hry ⟩ <;> + -- simp_all + -- · rcases (LE.le.lt_or_eq Hry) with Hry' | Hry' + -- · simp [elog] + -- split + -- · exfalso + -- rename_i h + -- cases h + -- linarith + -- · simp_all + -- · simp [top_mul'] + -- split + -- · simp_all [elog] + -- · exfalso + -- rename_i H + -- apply H + -- congr + -- rw [Hry'] + -- · rcases (LE.le.lt_or_eq Hrx) with Hrx' | Hrx' + -- · simp [elog] + -- split + -- · exfalso + -- rename_i h + -- cases h + -- linarith + -- · simp_all + -- · simp [mul_top'] + -- split + -- · simp_all [elog] + -- · exfalso + -- rename_i H + -- apply H + -- congr + -- rw [Hrx'] + -- · rcases (LE.le.lt_or_eq Hry) with Hry' | Hry' <;> + -- rcases (LE.le.lt_or_eq Hrx) with Hrx' | Hrx' + -- · rw [ENNReal.ofNNReal] + -- simp_all [elog] + -- sorry + -- · sorry + -- · simp [elog] + -- split + -- · rename_i H + -- rw [H] + -- simp + -- · rename_i H + -- split + -- · rename_i H' + -- rw [H'] + -- simp + -- · split + -- · exfalso + -- simp_all [or_self] + -- · simp_all + -- sorry + -- · simp [elog] + -- split + -- · split + -- · rename_i H1 H2 + -- rw [H1] + -- simp + -- · exfalso + -- rename_i H + -- apply H + -- congr + -- rw [Hry'] + -- · exfalso + -- rename_i H + -- apply H + -- congr + -- rw [Hrx'] @[simp] -lemma eexp_add : eexp w * eexp z = eexp (w + z) := by sorry -- checked truth table +lemma eexp_add {w z : EReal} : eexp w * eexp z = eexp (w + z) := by sorry -- checked truth table -- Log of power, log and exp inverses -lemma eexp_injective : eexp w = eexp z -> w = z := by - rw [eexp, eexp] - intro H - cases w <;> cases z <;> try tauto - · rename_i v - cases v <;> simp at * - rename_i v' - have Hv' := exp_pos v' - linarith - · rename_i v - cases v <;> simp at * - rename_i v' - have Hv' := exp_pos v' - linarith - · rename_i v₁ v₂ - cases v₁ <;> cases v₂ <;> simp at * - congr - rename_i v₁' v₂' - simp [ENNReal.ofReal] at H - apply NNReal.coe_inj.mpr at H - simp at H - have RW (r : ℝ) : (max (rexp r) 0) = (rexp r) := by - apply max_eq_left_iff.mpr - exact exp_nonneg r - rw [RW v₁'] at H - rw [RW v₂'] at H - exact exp_eq_exp.mp H - - -lemma elog_injective : elog x = elog y -> x = y := by - rw [elog, elog] - cases x <;> cases y <;> try tauto - · rename_i v - cases v; simp at * - - sorry - · rename_i v - cases v; simp at * - sorry - · rename_i v₁ v₂ - cases v₁; cases v₂; simp at * - sorry - -lemma eexp_mono_lt : (w < z) <-> eexp w < eexp z := by +lemma eexp_injective {w z : EReal} : eexp w = eexp z -> w = z := by + sorry + -- rw [eexp, eexp] + -- intro H + -- cases w <;> cases z <;> try tauto + -- · rename_i v + -- cases v <;> simp at * + -- rename_i v' + -- have Hv' := exp_pos v' + -- linarith + -- · rename_i v + -- cases v <;> simp at * + -- rename_i v' + -- have Hv' := exp_pos v' + -- linarith + -- · rename_i v₁ v₂ + -- cases v₁ <;> cases v₂ <;> simp at * + -- congr + -- rename_i v₁' v₂' + -- simp [ENNReal.ofReal] at H + -- apply NNReal.coe_inj.mpr at H + -- simp at H + -- have RW (r : ℝ) : (max (rexp r) 0) = (rexp r) := by + -- apply max_eq_left_iff.mpr + -- exact exp_nonneg r + -- rw [RW v₁'] at H + -- rw [RW v₂'] at H + -- exact exp_eq_exp.mp H + + +lemma elog_injective {x y : ENNReal} : elog x = elog y -> x = y := by + sorry + -- rw [elog, elog] + -- cases x <;> cases y <;> try tauto + -- · rename_i v + -- cases v; simp at * + + -- sorry + -- · rename_i v + -- cases v; simp at * + -- sorry + -- · rename_i v₁ v₂ + -- cases v₁; cases v₂; simp at * + -- sorry + +lemma eexp_mono_lt {w z : EReal} : (w < z) <-> eexp w < eexp z := by sorry -lemma eexp_mono_le : (w <= z) <-> eexp w <= eexp z := by sorry +lemma eexp_mono_le {w z : EReal} : (w <= z) <-> eexp w <= eexp z := by sorry -lemma elog_mono_lt : (x < y) <-> elog x < elog y := by sorry +lemma elog_mono_lt {x y : ENNReal} : (x < y) <-> elog x < elog y := by sorry -lemma elog_mono_le : (x <= y) <-> elog x <= elog y := by sorry +lemma elog_mono_le {x y : ENNReal} : (x <= y) <-> elog x <= elog y := by sorry -- iff for log positive @@ -327,308 +553,205 @@ lemma elog_mono_le : (x <= y) <-> elog x <= elog y := by sorry -- Not sure which ones we'll need but --- Specialized lemmas for ofEReal when its argument is nonnegative (so no truncation happens) -lemma ofEReal_nonneg_eq_iff (Hw : 0 <= w) (Hz : 0 <= z) : w = z <-> (ofEReal w = ofEReal z) := by - apply Iff.intro - · intro H1 - rw [H1] - · intro H1 - rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> - rcases (EReal_cases z) with Hz' | (Hz' | ⟨ z', Hz' ⟩) <;> - simp_all - -lemma ofEReal_le_mono_nonneg (Hw : 0 ≤ w) : w ≤ z <-> (ofEReal w ≤ ofEReal z) := by - apply Iff.intro - · intro H1 - rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> - rcases (EReal_cases z) with Hz' | (Hz' | ⟨ z', Hz' ⟩) <;> - simp_all - exact ofReal_le_ofReal H1 - · intro H1 - rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> - rcases (EReal_cases z) with Hz' | (Hz' | ⟨ z', Hz' ⟩) <;> - simp_all - · sorry - · simp [ENNReal.ofReal] at H1 - apply toNNReal_le_toNNReal_iff'.mp at H1 - cases H1 - · assumption - · rename_i H - have H' : w' = 0 := by apply le_antisymm <;> assumption - simp_all - apply EReal.coe_nonneg.mp - simp_all - -- ??? - sorry - --- Use above -lemma ofEReal_nonneg_zero (Hz : 0 ≤ z) : (0 = z) <-> 0 = ofEReal z := sorry - - -lemma ofEReal_le_mono (H : w ≤ z) : ofEReal w ≤ ofEReal z := by - rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> - rcases (EReal_cases z) with Hz' | (Hz' | ⟨ z', Hz' ⟩) - all_goals rw [Hw', Hz'] - all_goals simp_all [ENNReal.ofEReal] - simp [Real.toEReal] - exact ofReal_le_ofReal H +set_option pp.coercions false -@[simp] -lemma ofEReal_plus_nonneg (Hw : 0 ≤ w) (Hz : 0 ≤ z) : ofEReal (w + z) = ofEReal w + ofEReal z := by - rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> - rcases (EReal_cases z) with Hz' | (Hz' | ⟨ z', Hz' ⟩) - all_goals rw [Hw', Hz'] - all_goals simp_all +lemma elog_ENNReal_ofReal_of_pos {x : ℝ} (H : 0 < x) : (ENNReal.ofReal x).elog = x.log.toEReal := by sorry + -- simp [ENNReal.ofReal, ENNReal.elog, ENNReal.toEReal] + -- rw [ite_eq_iff'] + -- apply And.intro + -- · intro + -- exfalso + -- linarith + -- · intro H + -- simp at H + -- rw [max_eq_left_of_lt H] -@[simp] -lemma ofEReal_mul_nonneg (Hw : 0 ≤ w) (Hz : 0 ≤ z) : ofEReal (w * z) = ofEReal w * ofEReal z := by - have HBle: (⊥ : EReal) ≤ 0 := by exact OrderBot.bot_le 0 - rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> - rcases (EReal_cases z) with Hz' | (Hz' | ⟨ z', Hz' ⟩) - all_goals rw [Hw', Hz'] - all_goals simp_all - · rw [top_mul'] - split - · simp_all - apply ofEReal_nonpos - sorry - · cases (LE.le.lt_or_eq Hz) - · sorry - · sorry - · sorry - · sorry - -lemma ofEReal_nonneg_scal_l (H1 : 0 < r) (H2 : 0 ≤ r * w) : 0 ≤ w := by - rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) - · simp_all - rw [EReal.mul_bot_of_pos] at H2 - all_goals simp_all - · simp_all - · simp_all - sorry +end elog_eexp -@[simp] -lemma toEReal_ofENNReal_nonneg (H : 0 ≤ w) : ENNReal.toEReal (ofEReal w) = w := by - rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> simp_all -@[simp] -lemma ofEReal_toENNReal : ofEReal (ENNReal.toEReal x) = x := by - rw [ENNReal.ofEReal] - split - · rename_i x' Hx' - simp [toEReal] at Hx' - cases x <;> simp_all - rename_i x'' - exfalso - rw [NNReal.toReal] at Hx' - rw [Real.toEReal] at Hx' - simp_all - · rename_i x' Hx' - simp [toEReal] at Hx' - cases x <;> simp_all - rw [NNReal.toReal] at Hx' - rw [Real.toEReal] at Hx' - simp_all - cases Hx' - · rename_i e x' Hx' - simp [toEReal] at Hx' - cases x <;> simp_all - · cases Hx' - · rw [ENNReal.ofReal] - simp_all - rw [NNReal.toReal] at Hx' - rw [Real.toEReal] at Hx' - simp_all - cases Hx' - simp_all -lemma ofEReal_ofReal_toENNReal : ENNReal.ofEReal (Real.toEReal r) = ENNReal.ofReal r := by - simp [ofEReal, Real.toEReal, ENNReal.ofReal] -set_option pp.coercions false - -lemma elog_ENNReal_ofReal_of_pos (x : ℝ) (H : 0 < x): (ENNReal.ofReal x).elog = x.log.toEReal := by - simp [ENNReal.ofReal, ENNReal.elog, ENNReal.toEReal] - rw [ite_eq_iff'] - apply And.intro - · intro - exfalso - linarith - · intro H - simp at H - rw [max_eq_left_of_lt H] +section misc -- ENNReal inequalities -- These are needed to prove the extensded version of Jensen's inequality -lemma mul_mul_inv_le_mul_cancel : (x * y⁻¹) * y ≤ x := by - cases x - · simp_all - rename_i x' - cases (Classical.em (x' = 0)) - · simp_all - rename_i Hx' - cases y - · simp_all - rename_i y' - cases (Classical.em (y' = 0)) - · simp_all - rename_i Hy' - simp - rw [← coe_inv Hy'] - rw [← coe_mul] - rw [← coe_mul] - rw [mul_right_comm] - rw [mul_inv_cancel_right₀ Hy' x'] - -lemma mul_mul_inv_eq_mul_cancel (H : y = 0 -> x = 0) (H2 : ¬(x ≠ 0 ∧ y = ⊤)) : (x * y⁻¹) * y = x := by - cases x - · simp_all - rename_i x' - cases (Classical.em (x' = 0)) - · simp_all - rename_i Hx' - cases y - · simp_all - rename_i y' - cases (Classical.em (y' = 0)) - · simp_all - rename_i Hy' - simp - rw [← coe_inv Hy'] - rw [← coe_mul] - rw [← coe_mul] - rw [mul_right_comm] - rw [mul_inv_cancel_right₀ Hy' x'] +lemma mul_mul_inv_le_mul_cancel {x y : ENNReal} : (x * y⁻¹) * y ≤ x := by + sorry + -- cases x + -- · simp_all + -- rename_i x' + -- cases (Classical.em (x' = 0)) + -- · simp_all + -- rename_i Hx' + -- cases y + -- · simp_all + -- rename_i y' + -- cases (Classical.em (y' = 0)) + -- · simp_all + -- rename_i Hy' + -- simp + -- rw [← coe_inv Hy'] + -- rw [← coe_mul] + -- rw [← coe_mul] + -- rw [mul_right_comm] + -- rw [mul_inv_cancel_right₀ Hy' x'] + +lemma mul_mul_inv_eq_mul_cancel {x y : ENNReal} (H : y = 0 -> x = 0) (H2 : ¬(x ≠ 0 ∧ y = ⊤)) : (x * y⁻¹) * y = x := by + sorry + -- cases x + -- · simp_all + -- rename_i x' + -- cases (Classical.em (x' = 0)) + -- · simp_all + -- rename_i Hx' + -- cases y + -- · simp_all + -- rename_i y' + -- cases (Classical.em (y' = 0)) + -- · simp_all + -- rename_i Hy' + -- simp + -- rw [← coe_inv Hy'] + -- rw [← coe_mul] + -- rw [← coe_mul] + -- rw [mul_right_comm] + -- rw [mul_inv_cancel_right₀ Hy' x'] -- Could be shortened -lemma ereal_smul_le_left (s : EReal) (Hr1 : 0 < s) (Hr2 : s < ⊤) (H : s * w ≤ s * z) : w ≤ z := by - have defTop : some none = (⊤ : EReal) := by simp [Top.top] - have defBot : none = (⊥ : EReal) := by simp [Bot.bot] - - cases s - · exfalso - rw [defBot] at Hr1 - simp_all only [not_lt_bot] - rename_i s_nnr - cases s_nnr - · rw [defTop] at Hr2 - exfalso - simp_all only [EReal.zero_lt_top, lt_self_iff_false] - rename_i s_R - have Hsr : some (some s_R) = Real.toEReal s_R := by simp [Real.toEReal] - rw [Hsr] at H - rw [Hsr] at Hr1 - rw [Hsr] at Hr2 - clear Hsr +lemma ereal_smul_le_left {w z : EReal} (s : EReal) (Hr1 : 0 < s) (Hr2 : s < ⊤) (H : s * w ≤ s * z) : w ≤ z := by + sorry + -- have defTop : some none = (⊤ : EReal) := by simp [Top.top] + -- have defBot : none = (⊥ : EReal) := by simp [Bot.bot] + + -- cases s + -- · exfalso + -- rw [defBot] at Hr1 + -- simp_all only [not_lt_bot] + -- rename_i s_nnr + -- cases s_nnr + -- · rw [defTop] at Hr2 + -- exfalso + -- simp_all only [EReal.zero_lt_top, lt_self_iff_false] + -- rename_i s_R + -- have Hsr : some (some s_R) = Real.toEReal s_R := by simp [Real.toEReal] + -- rw [Hsr] at H + -- rw [Hsr] at Hr1 + -- rw [Hsr] at Hr2 + -- clear Hsr + + -- cases w + -- · apply left_eq_inf.mp + -- rfl + -- rename_i w_nnr + -- cases w_nnr + -- · simp [defTop] at H + -- rw [EReal.mul_top_of_pos Hr1] at H + -- have X1 : z = ⊤ := by + -- cases z + -- · exfalso + -- simp at H + -- rw [defBot] at H + -- rw [EReal.mul_bot_of_pos] at H + -- · cases H + -- · apply Hr1 + -- rename_i z_nnr + -- cases z_nnr + -- · simp [Top.top] + -- exfalso + -- apply top_le_iff.mp at H + -- rename_i z_R + -- have Hzr : some (some z_R) = Real.toEReal z_R := by simp [Real.toEReal] + -- rw [Hzr] at H + -- rw [<- EReal.coe_mul] at H + -- cases H + -- rw [defTop, X1] + -- rename_i w_R + -- cases z + -- · simp [defBot] at H + -- rw [EReal.mul_bot_of_pos] at H + -- apply le_bot_iff.mp at H + -- · rw [defBot] + -- have Hwr : some (some w_R) = Real.toEReal w_R := by simp [Real.toEReal] + -- rw [Hwr] at H + -- rw [<- EReal.coe_mul] at H + -- cases H + -- · apply Hr1 + -- rename_i z_nnr + -- cases z_nnr + -- · exact right_eq_inf.mp rfl + -- rename_i z_R + -- have Hwr : some (some w_R) = Real.toEReal w_R := by simp [Real.toEReal] + -- have Hzr : some (some z_R) = Real.toEReal z_R := by simp [Real.toEReal] + -- rw [Hwr, Hzr] at H + -- rw [Hwr, Hzr] + -- clear Hwr + -- clear Hzr + -- apply EReal.coe_le_coe_iff.mpr + -- repeat rw [<- EReal.coe_mul] at H + -- apply EReal.coe_le_coe_iff.mp at H + -- apply le_of_mul_le_mul_left H + -- exact EReal.coe_pos.mp Hr1 + +lemma ereal_smul_eq_left {w z : EReal} (s : EReal) (Hr1 : 0 < s) (Hr2 : s < ⊤) (H : s * w = s * z) : w = z := by + sorry + -- rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> + -- rcases (EReal_cases z) with Hz' | (Hz' | ⟨ z', Hz' ⟩) <;> + -- rcases (EReal_cases s) with Hs' | (Hs' | ⟨ s', Hs' ⟩) + -- all_goals rw [Hw', Hz'] + -- all_goals simp_all + -- · rw [mul_top_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H + -- rw [mul_bot_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H + -- cases H + -- · rw [mul_bot_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H + -- cases H + -- · rw [mul_top_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H + -- rw [mul_bot_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H + -- cases H + -- · rw [mul_top_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H + -- cases H + -- · rw [mul_bot_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H + -- repeat rw [<- EReal.coe_mul] at H + -- cases H + -- · rw [mul_top_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H + -- repeat rw [<- EReal.coe_mul] at H + -- cases H + -- · repeat rw [<- EReal.coe_mul] at H + -- apply EReal.coe_eq_coe_iff.mp at H + -- simp_all only [mul_eq_mul_left_iff] + -- cases H + -- · assumption + -- · exfalso + -- linarith + +lemma ereal_smul_left_le {s : EReal} {w z : EReal} (Hr1 : 0 < s) (Hr2 : s < ⊤) (H : w ≤ z) : s * w ≤ s * z := by + sorry + -- rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> + -- rcases (EReal_cases z) with Hz' | (Hz' | ⟨ z', Hz' ⟩) + -- all_goals rw [Hw', Hz'] + -- all_goals simp_all + -- · rw [mul_top_of_pos Hr1] + -- exact OrderTop.le_top (s * ⊥) + -- · rw [mul_bot_of_pos Hr1] + -- exact OrderBot.bot_le (s * z'.toEReal) + -- · rw [mul_top_of_pos Hr1] + -- exact OrderTop.le_top (s * w'.toEReal) + -- rcases (EReal_cases s) with Hs' | (Hs' | ⟨ s', Hs' ⟩) + -- all_goals simp_all + -- repeat rw [← EReal.coe_mul] + -- apply EReal.coe_le_coe_iff.mpr + -- exact (mul_le_mul_iff_of_pos_left Hr1).mpr H + +lemma ENNReal_toReal_partial_inj {a b : ENNReal} (H : a.toReal = b.toReal) (H1 : a ≠ ⊤) (H2 : b ≠ ⊤) : a = b := by + sorry + -- rcases a with Ha' | ⟨ a' | Ha' ⟩ <;> + -- rcases b with Hb' | ⟨ b' | Hb' ⟩ + -- all_goals simp_all - cases w - · apply left_eq_inf.mp - rfl - rename_i w_nnr - cases w_nnr - · simp [defTop] at H - rw [EReal.mul_top_of_pos Hr1] at H - have X1 : z = ⊤ := by - cases z - · exfalso - simp at H - rw [defBot] at H - rw [EReal.mul_bot_of_pos] at H - · cases H - · apply Hr1 - rename_i z_nnr - cases z_nnr - · simp [Top.top] - exfalso - apply top_le_iff.mp at H - rename_i z_R - have Hzr : some (some z_R) = Real.toEReal z_R := by simp [Real.toEReal] - rw [Hzr] at H - rw [<- EReal.coe_mul] at H - cases H - rw [defTop, X1] - rename_i w_R - cases z - · simp [defBot] at H - rw [EReal.mul_bot_of_pos] at H - apply le_bot_iff.mp at H - · rw [defBot] - have Hwr : some (some w_R) = Real.toEReal w_R := by simp [Real.toEReal] - rw [Hwr] at H - rw [<- EReal.coe_mul] at H - cases H - · apply Hr1 - rename_i z_nnr - cases z_nnr - · exact right_eq_inf.mp rfl - rename_i z_R - have Hwr : some (some w_R) = Real.toEReal w_R := by simp [Real.toEReal] - have Hzr : some (some z_R) = Real.toEReal z_R := by simp [Real.toEReal] - rw [Hwr, Hzr] at H - rw [Hwr, Hzr] - clear Hwr - clear Hzr - apply EReal.coe_le_coe_iff.mpr - repeat rw [<- EReal.coe_mul] at H - apply EReal.coe_le_coe_iff.mp at H - apply le_of_mul_le_mul_left H - exact EReal.coe_pos.mp Hr1 - -lemma ereal_smul_eq_left (s : EReal) (Hr1 : 0 < s) (Hr2 : s < ⊤) (H : s * w = s * z) : w = z := by - rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> - rcases (EReal_cases z) with Hz' | (Hz' | ⟨ z', Hz' ⟩) <;> - rcases (EReal_cases s) with Hs' | (Hs' | ⟨ s', Hs' ⟩) - all_goals rw [Hw', Hz'] - all_goals simp_all - · rw [mul_top_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H - rw [mul_bot_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H - cases H - · rw [mul_bot_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H - cases H - · rw [mul_top_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H - rw [mul_bot_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H - cases H - · rw [mul_top_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H - cases H - · rw [mul_bot_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H - repeat rw [<- EReal.coe_mul] at H - cases H - · rw [mul_top_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H - repeat rw [<- EReal.coe_mul] at H - cases H - · repeat rw [<- EReal.coe_mul] at H - apply EReal.coe_eq_coe_iff.mp at H - simp_all only [mul_eq_mul_left_iff] - cases H - · assumption - · exfalso - linarith - -lemma ereal_smul_left_le (s : EReal) (Hr1 : 0 < s) (Hr2 : s < ⊤) (H : w ≤ z) : s * w ≤ s * z := by - rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> - rcases (EReal_cases z) with Hz' | (Hz' | ⟨ z', Hz' ⟩) - all_goals rw [Hw', Hz'] - all_goals simp_all - · rw [mul_top_of_pos Hr1] - exact OrderTop.le_top (s * ⊥) - · rw [mul_bot_of_pos Hr1] - exact OrderBot.bot_le (s * z'.toEReal) - · rw [mul_top_of_pos Hr1] - exact OrderTop.le_top (s * w'.toEReal) - rcases (EReal_cases s) with Hs' | (Hs' | ⟨ s', Hs' ⟩) - all_goals simp_all - repeat rw [← EReal.coe_mul] - apply EReal.coe_le_coe_iff.mpr - exact (mul_le_mul_iff_of_pos_left Hr1).mpr H - -lemma ENNReal_toReal_partial_inj (a b : ENNReal) (H : a.toReal = b.toReal) (H1 : a ≠ ⊤) (H2 : b ≠ ⊤) : a = b := by - rcases a with Ha' | ⟨ a' | Ha' ⟩ <;> - rcases b with Hb' | ⟨ b' | Hb' ⟩ - all_goals simp_all +end misc end ENNReal From 76f9bb20e0aa5cafd57b3a921ab66bd2a408257b Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Wed, 26 Jun 2024 14:45:44 -0400 Subject: [PATCH 097/122] All but one ofEReal lemma --- SampCert/Util/Log.lean | 68 ++++++++++++++++++++---------------------- 1 file changed, 33 insertions(+), 35 deletions(-) diff --git a/SampCert/Util/Log.lean b/SampCert/Util/Log.lean index 854d4504..8b996bb8 100644 --- a/SampCert/Util/Log.lean +++ b/SampCert/Util/Log.lean @@ -214,6 +214,12 @@ lemma ofEReal_ofReal_toENNReal : ENNReal.ofEReal (Real.toEReal r) = ENNReal.ofRe simp [ofEReal, Real.toEReal, ENNReal.ofReal] +lemma ofReal_injective_nonneg {r s : ℝ} (HR : 0 ≤ r) (HS : 0 ≤ s) (H : ENNReal.ofReal r = ENNReal.ofReal s) : r = s := by + apply (Real.toNNReal_eq_toNNReal_iff HR HS).mp + simp [ENNReal.ofReal] at H + assumption + + lemma ofEReal_le_mono {w z : EReal} (H : w ≤ z) : ofEReal w ≤ ofEReal z := by all_goals case_EReal_isReal w all_goals case_EReal_isReal z @@ -222,24 +228,21 @@ lemma ofEReal_le_mono {w z : EReal} (H : w ≤ z) : ofEReal w ≤ ofEReal z := b lemma ofEReal_le_mono_conv_nonneg {w z : EReal} (Hw : 0 ≤ w) (Hle : ofEReal w ≤ ofEReal z) : w ≤ z := by - -- Reduce to comparison between ENNReals? - - -- case_EReal_isENNReal w - -- · -- w < 0 => ofEReal w = 0 - -- rw [(ofEReal_eq_zero_iff w).mp ?G0] at Hle - -- case G0 => exact le_of_lt Hzlz - -- skip - -- sorry - -- rename_i w' Hw' - -- case_EReal_isENNReal z - -- · rename_i Hzlz - -- -- z < 0 => ofEReal z = 0 - -- rw [(ofEReal_eq_zero_iff z).mp ?G1] at Hle - -- case G1 => exact le_of_lt Hzlz - -- skip - -- sorry - -- rename_i z' Hz' + all_goals case_EReal_isENNReal w + all_goals case_EReal_isENNReal z + · exfalso + rename_i Hw' _ + -- 0 <= w < 0 + sorry + · exfalso + rename_i r Hw' Hr + -- 0 <= w < 0 sorry + · exfalso + rename_i r Hr Hz' + -- 0 <= w <= z < 0 + sorry + @[simp] @@ -252,39 +255,34 @@ lemma ofEReal_le_mono_conv_nonneg {w z : EReal} (Hw : 0 ≤ w) (Hle : ofEReal w rw [ofReal_add Hw Hz] + @[simp] lemma ofEReal_mul_nonneg (Hw : 0 ≤ w) (Hz : 0 ≤ z) : ofEReal (w * z) = ofEReal w * ofEReal z := by all_goals case_EReal_isReal w all_goals case_EReal_isReal z - · rename_i r Hr Hz' + · rename_i r _ _ case_nonneg_zero Hz rename_i Hr_nz simp [top_mul_coe_of_pos Hr_nz] · rename_i r Hr Hz' case_nonneg_zero Hw rename_i Hr_nz - -- simp [top_mul_coe_of_pos Hr_nz] - sorry - · rename_i r₁ r₂ Hr₁ Hr₂ - rw [← EReal.coe_mul] + simp [coe_mul_top_of_pos Hr_nz] + · rw [← EReal.coe_mul] rw [ofEReal_ofReal_toENNReal] rw [ofReal_mul' Hz] --- ??? -lemma ofEReal_nonneg_scal_l {r : ℝ} {w : EReal} (H1 : 0 < r) (H2 : 0 ≤ r * w) : 0 ≤ w := by - sorry --- lemma ofEReal_nonneg_scal_l (H1 : 0 < r) (H2 : 0 ≤ r * w) : 0 ≤ w := by --- rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) --- · simp_all --- rw [EReal.mul_bot_of_pos] at H2 --- all_goals simp_all --- · simp_all --- · simp_all --- sorry --- --- +lemma ofEReal_nonneg_scal_l {r : ℝ} {w : EReal} (H1 : 0 < r) (H2 : 0 ≤ r * w) : 0 ≤ w := by + all_goals case_EReal_isReal w + · exfalso + rw [EReal.mul_bot_of_pos (EReal.coe_pos.mpr H1)] at H2 + simp at H2 + · rename_i w' Hw' + rw [← EReal.coe_mul] at H2 + apply EReal.coe_nonneg.mp at H2 + exact nonneg_of_mul_nonneg_right H2 H1 end ofEReal From 5dbbcc90166b2284052bf62a8f00ba552bbb92a2 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Wed, 26 Jun 2024 15:31:35 -0400 Subject: [PATCH 098/122] checkpoint --- SampCert/Util/Log.lean | 585 +++++++++++++++++++---------------------- 1 file changed, 274 insertions(+), 311 deletions(-) diff --git a/SampCert/Util/Log.lean b/SampCert/Util/Log.lean index 8b996bb8..a2387001 100644 --- a/SampCert/Util/Log.lean +++ b/SampCert/Util/Log.lean @@ -99,6 +99,21 @@ macro_rules | `(tactic| case_ENNReal_isReal $w:term ) => `(tactic| rcases (ENNReal_isReal_cases $w) with _ | ⟨ _, _, _⟩ <;> try simp_all) +/-- +An ENNReal is either top, zero, or the lift if a positive real +-/ +lemma ENNReal_isReal_zero_cases (x : ENNReal) : x = ⊤ ∨ x = 0 ∨ (∃ v : ℝ, x = ENNReal.ofReal v ∧ 0 < v) := by + case_ENNReal_isReal x + rename_i r Hr1 Hr2 + case_nonneg_zero Hr2 + right + exists r + +syntax "case_ENNReal_isReal_zero" term : tactic +macro_rules +| `(tactic| case_ENNReal_isReal_zero $w:term ) => + `(tactic| rcases (ENNReal_isReal_zero_cases $w) with _ | _ | ⟨ _, _, _⟩ <;> try simp_all) + /-- An EReal is either ⊤, ⊥, or the lift of some real number. -/ @@ -309,29 +324,27 @@ def eexp (y : EReal) : ENNReal := | ⊤ => ⊤ | some (some r) => ENNReal.ofReal (Real.exp r) - - +-- MARKUSDE: cleanup? @[simp] lemma elog_of_pos_real {r : ℝ} (H : 0 < r) : elog (ENNReal.ofReal r) = Real.log r := by - sorry - -- rw [elog] - -- split - -- · simp at * - -- · split - -- · rename_i r' heq h - -- exfalso - -- rw [h] at heq - -- simp at heq - -- linarith - -- · rename_i h' r' heq h - -- simp_all - -- congr - -- simp [ENNReal.ofReal] at heq - -- rw [<- heq] - -- exact (Real.coe_toNNReal r (le_of_lt H)) + rw [elog] + split + · simp at * + · split + · rename_i r' heq h + exfalso + rw [h] at heq + simp at heq + linarith + · rename_i h' r' heq h + simp_all + congr + simp [ENNReal.ofReal] at heq + rw [<- heq] + exact (Real.coe_toNNReal r (le_of_lt H)) @[simp] -lemma elog_zero : elog (ENNReal.ofReal 0) = ⊥ := by simp [elog] +lemma elog_zero : elog 0 = ⊥ := by simp [elog] @[simp] lemma elog_top : elog ⊤ = ⊤ := by simp [elog] @@ -347,192 +360,152 @@ lemma eexp_zero : eexp 0 = 1 := by simp [eexp] @[simp] lemma eexp_ofReal {r : ℝ} : eexp r = ENNReal.ofReal (Real.exp r) := by - sorry - -- simp [ENNReal.ofReal, eexp, elog] - -- rfl + simp [ENNReal.ofReal, eexp, elog] + rfl +-- MARKUSDE: Cleanup @[simp] lemma elog_eexp {x : ENNReal} : eexp (elog x) = x := by - sorry - -- rw [elog] - -- split - -- · simp - -- · rename_i _ r' - -- split - -- · simp - -- rename_i _ h - -- rw [h] - -- · rename_i _ H - -- simp - -- rw [NNReal.toReal] - -- simp - -- rw [Real.exp_log] - -- rw [ofReal_coe_nnreal] - -- rcases r' with ⟨ v , Hv ⟩ - -- apply lt_of_le_of_ne - -- · simpa - -- · simp - -- intro Hk - -- apply H - -- apply NNReal.coe_eq_zero.mp - -- simp - -- rw [Hk] - + rw [elog] + split + · simp + · rename_i _ r' + split + · simp + rename_i _ h + rw [h] + · rename_i _ H + simp + rw [NNReal.toReal] + simp + rw [Real.exp_log] + rw [ofReal_coe_nnreal] + rcases r' with ⟨ v , Hv ⟩ + apply lt_of_le_of_ne + · simpa + · simp + intro Hk + apply H + apply NNReal.coe_eq_zero.mp + simp + rw [Hk] + +-- MARKUSDE: Cleanup @[simp] lemma eexp_elog {w : EReal} : (elog (eexp w)) = w := by - sorry - -- cases w - -- · simp [eexp, elog] - -- rfl - -- · simp [eexp, elog] - -- rename_i v' - -- cases v' - -- · simp - -- rfl - -- · simp - -- rename_i v'' - -- simp [ENNReal.ofReal] - -- split - -- · rename_i Hcont - -- have Hcont' : 0 < rexp v'' := by exact exp_pos v'' - -- linarith - -- · rename_i H - -- have RW : (max (rexp v'') 0) = (rexp v'') := by - -- apply max_eq_left_iff.mpr - -- linarith - -- simp [RW] - -- clear RW - -- simp [Real.toEReal] + cases w + · simp [eexp, elog] + rfl + · simp [eexp, elog] + rename_i v' + cases v' + · simp + rfl + · simp + rename_i v'' + simp [ENNReal.ofReal] + split + · rename_i Hcont + have Hcont' : 0 < rexp v'' := by exact exp_pos v'' + linarith + · rename_i H + have RW : (max (rexp v'') 0) = (rexp v'') := by + apply max_eq_left_iff.mpr + linarith + simp [RW] + clear RW + simp [Real.toEReal] + + +-- MARKUSDE: cleanup +lemma elog_ENNReal_ofReal_of_pos {x : ℝ} (H : 0 < x) : (ENNReal.ofReal x).elog = x.log.toEReal := by + simp [ENNReal.ofReal, ENNReal.elog, ENNReal.toEReal] + rw [ite_eq_iff'] + apply And.intro + · intro + exfalso + linarith + · intro H + simp at H + rw [max_eq_left_of_lt H] +set_option pp.coercions false @[simp] lemma elog_mul {x y : ENNReal} : elog x + elog y = elog (x * y) := by - sorry - -- rcases x with _ | ⟨ rx , Hrx ⟩ <;> - -- rcases y with _ | ⟨ ry , Hry ⟩ <;> - -- simp_all - -- · rcases (LE.le.lt_or_eq Hry) with Hry' | Hry' - -- · simp [elog] - -- split - -- · exfalso - -- rename_i h - -- cases h - -- linarith - -- · simp_all - -- · simp [top_mul'] - -- split - -- · simp_all [elog] - -- · exfalso - -- rename_i H - -- apply H - -- congr - -- rw [Hry'] - -- · rcases (LE.le.lt_or_eq Hrx) with Hrx' | Hrx' - -- · simp [elog] - -- split - -- · exfalso - -- rename_i h - -- cases h - -- linarith - -- · simp_all - -- · simp [mul_top'] - -- split - -- · simp_all [elog] - -- · exfalso - -- rename_i H - -- apply H - -- congr - -- rw [Hrx'] - -- · rcases (LE.le.lt_or_eq Hry) with Hry' | Hry' <;> - -- rcases (LE.le.lt_or_eq Hrx) with Hrx' | Hrx' - -- · rw [ENNReal.ofNNReal] - -- simp_all [elog] - -- sorry - -- · sorry - -- · simp [elog] - -- split - -- · rename_i H - -- rw [H] - -- simp - -- · rename_i H - -- split - -- · rename_i H' - -- rw [H'] - -- simp - -- · split - -- · exfalso - -- simp_all [or_self] - -- · simp_all - -- sorry - -- · simp [elog] - -- split - -- · split - -- · rename_i H1 H2 - -- rw [H1] - -- simp - -- · exfalso - -- rename_i H - -- apply H - -- congr - -- rw [Hry'] - -- · exfalso - -- rename_i H - -- apply H - -- congr - -- rw [Hrx'] + all_goals case_ENNReal_isReal_zero x + all_goals case_ENNReal_isReal_zero y + rename_i r₁ Hr₁ HPr₁ r₂ Hr₂ HPr₂ + rw [← EReal.coe_add] + rw [<- Real.log_mul ?G1 ?G2] + case G1 => linarith + case G2 => linarith + rw [<- elog_ENNReal_ofReal_of_pos ?G1] + case G1 => exact Real.mul_pos HPr₁ HPr₂ + rw [ENNReal.ofReal_mul] + linarith @[simp] -lemma eexp_add {w z : EReal} : eexp w * eexp z = eexp (w + z) := by sorry -- checked truth table - +lemma eexp_add {w z : EReal} : eexp w * eexp z = eexp (w + z) := by + all_goals case_EReal_isReal w + all_goals case_EReal_isReal z + · apply top_mul + simp + apply exp_pos + · apply mul_top + simp + apply exp_pos + rw [← EReal.coe_add] + rw [<- ENNReal.ofReal_mul ?G1] + case G1 => apply exp_nonneg + rw [← exp_add] + rw [eexp_ofReal] --- Log of power, log and exp inverses +-- TODO (maybe): Log of power, log and exp inverses +-- MARKUSDE: cleanup lemma eexp_injective {w z : EReal} : eexp w = eexp z -> w = z := by - sorry - -- rw [eexp, eexp] - -- intro H - -- cases w <;> cases z <;> try tauto - -- · rename_i v - -- cases v <;> simp at * - -- rename_i v' - -- have Hv' := exp_pos v' - -- linarith - -- · rename_i v - -- cases v <;> simp at * - -- rename_i v' - -- have Hv' := exp_pos v' - -- linarith - -- · rename_i v₁ v₂ - -- cases v₁ <;> cases v₂ <;> simp at * - -- congr - -- rename_i v₁' v₂' - -- simp [ENNReal.ofReal] at H - -- apply NNReal.coe_inj.mpr at H - -- simp at H - -- have RW (r : ℝ) : (max (rexp r) 0) = (rexp r) := by - -- apply max_eq_left_iff.mpr - -- exact exp_nonneg r - -- rw [RW v₁'] at H - -- rw [RW v₂'] at H - -- exact exp_eq_exp.mp H + rw [eexp, eexp] + intro H + cases w <;> cases z <;> try tauto + · rename_i v + cases v <;> simp at * + rename_i v' + have Hv' := exp_pos v' + linarith + · rename_i v + cases v <;> simp at * + rename_i v' + have Hv' := exp_pos v' + linarith + · rename_i v₁ v₂ + cases v₁ <;> cases v₂ <;> simp at * + congr + rename_i v₁' v₂' + simp [ENNReal.ofReal] at H + apply NNReal.coe_inj.mpr at H + simp at H + have RW (r : ℝ) : (max (rexp r) 0) = (rexp r) := by + apply max_eq_left_iff.mpr + exact exp_nonneg r + rw [RW v₁'] at H + rw [RW v₂'] at H + exact exp_eq_exp.mp H + lemma elog_injective {x y : ENNReal} : elog x = elog y -> x = y := by - sorry - -- rw [elog, elog] - -- cases x <;> cases y <;> try tauto - -- · rename_i v - -- cases v; simp at * - - -- sorry - -- · rename_i v - -- cases v; simp at * - -- sorry - -- · rename_i v₁ v₂ - -- cases v₁; cases v₂; simp at * - -- sorry + all_goals case_ENNReal_isReal_zero x + all_goals case_ENNReal_isReal_zero y + rename_i r₁ Hr₁ HPr₁ r₂ Hr₂ HPr₂ + intro Hlog_eq + suffices r₁ = r₁ by simp + apply Real.log_injOn_pos + all_goals simp_all + + lemma eexp_mono_lt {w z : EReal} : (w < z) <-> eexp w < eexp z := by sorry @@ -554,17 +527,6 @@ lemma elog_mono_le {x y : ENNReal} : (x <= y) <-> elog x <= elog y := by sorry set_option pp.coercions false -lemma elog_ENNReal_ofReal_of_pos {x : ℝ} (H : 0 < x) : (ENNReal.ofReal x).elog = x.log.toEReal := by - sorry - -- simp [ENNReal.ofReal, ENNReal.elog, ENNReal.toEReal] - -- rw [ite_eq_iff'] - -- apply And.intro - -- · intro - -- exfalso - -- linarith - -- · intro H - -- simp at H - -- rw [max_eq_left_of_lt H] end elog_eexp @@ -576,127 +538,126 @@ end elog_eexp section misc +-- MARKUSDE: cleanup -- ENNReal inequalities -- These are needed to prove the extensded version of Jensen's inequality lemma mul_mul_inv_le_mul_cancel {x y : ENNReal} : (x * y⁻¹) * y ≤ x := by - sorry - -- cases x - -- · simp_all - -- rename_i x' - -- cases (Classical.em (x' = 0)) - -- · simp_all - -- rename_i Hx' - -- cases y - -- · simp_all - -- rename_i y' - -- cases (Classical.em (y' = 0)) - -- · simp_all - -- rename_i Hy' - -- simp - -- rw [← coe_inv Hy'] - -- rw [← coe_mul] - -- rw [← coe_mul] - -- rw [mul_right_comm] - -- rw [mul_inv_cancel_right₀ Hy' x'] - + cases x + · simp_all + rename_i x' + cases (Classical.em (x' = 0)) + · simp_all + rename_i Hx' + cases y + · simp_all + rename_i y' + cases (Classical.em (y' = 0)) + · simp_all + rename_i Hy' + simp + rw [← coe_inv Hy'] + rw [← coe_mul] + rw [← coe_mul] + rw [mul_right_comm] + rw [mul_inv_cancel_right₀ Hy' x'] + +-- MARKUSDE: cleanup lemma mul_mul_inv_eq_mul_cancel {x y : ENNReal} (H : y = 0 -> x = 0) (H2 : ¬(x ≠ 0 ∧ y = ⊤)) : (x * y⁻¹) * y = x := by - sorry - -- cases x - -- · simp_all - -- rename_i x' - -- cases (Classical.em (x' = 0)) - -- · simp_all - -- rename_i Hx' - -- cases y - -- · simp_all - -- rename_i y' - -- cases (Classical.em (y' = 0)) - -- · simp_all - -- rename_i Hy' - -- simp - -- rw [← coe_inv Hy'] - -- rw [← coe_mul] - -- rw [← coe_mul] - -- rw [mul_right_comm] - -- rw [mul_inv_cancel_right₀ Hy' x'] - --- Could be shortened + cases x + · simp_all + rename_i x' + cases (Classical.em (x' = 0)) + · simp_all + rename_i Hx' + cases y + · simp_all + rename_i y' + cases (Classical.em (y' = 0)) + · simp_all + rename_i Hy' + simp + rw [← coe_inv Hy'] + rw [← coe_mul] + rw [← coe_mul] + rw [mul_right_comm] + rw [mul_inv_cancel_right₀ Hy' x'] + +-- MARKUSDE: Cleanup lemma ereal_smul_le_left {w z : EReal} (s : EReal) (Hr1 : 0 < s) (Hr2 : s < ⊤) (H : s * w ≤ s * z) : w ≤ z := by - sorry - -- have defTop : some none = (⊤ : EReal) := by simp [Top.top] - -- have defBot : none = (⊥ : EReal) := by simp [Bot.bot] - - -- cases s - -- · exfalso - -- rw [defBot] at Hr1 - -- simp_all only [not_lt_bot] - -- rename_i s_nnr - -- cases s_nnr - -- · rw [defTop] at Hr2 - -- exfalso - -- simp_all only [EReal.zero_lt_top, lt_self_iff_false] - -- rename_i s_R - -- have Hsr : some (some s_R) = Real.toEReal s_R := by simp [Real.toEReal] - -- rw [Hsr] at H - -- rw [Hsr] at Hr1 - -- rw [Hsr] at Hr2 - -- clear Hsr - - -- cases w - -- · apply left_eq_inf.mp - -- rfl - -- rename_i w_nnr - -- cases w_nnr - -- · simp [defTop] at H - -- rw [EReal.mul_top_of_pos Hr1] at H - -- have X1 : z = ⊤ := by - -- cases z - -- · exfalso - -- simp at H - -- rw [defBot] at H - -- rw [EReal.mul_bot_of_pos] at H - -- · cases H - -- · apply Hr1 - -- rename_i z_nnr - -- cases z_nnr - -- · simp [Top.top] - -- exfalso - -- apply top_le_iff.mp at H - -- rename_i z_R - -- have Hzr : some (some z_R) = Real.toEReal z_R := by simp [Real.toEReal] - -- rw [Hzr] at H - -- rw [<- EReal.coe_mul] at H - -- cases H - -- rw [defTop, X1] - -- rename_i w_R - -- cases z - -- · simp [defBot] at H - -- rw [EReal.mul_bot_of_pos] at H - -- apply le_bot_iff.mp at H - -- · rw [defBot] - -- have Hwr : some (some w_R) = Real.toEReal w_R := by simp [Real.toEReal] - -- rw [Hwr] at H - -- rw [<- EReal.coe_mul] at H - -- cases H - -- · apply Hr1 - -- rename_i z_nnr - -- cases z_nnr - -- · exact right_eq_inf.mp rfl - -- rename_i z_R - -- have Hwr : some (some w_R) = Real.toEReal w_R := by simp [Real.toEReal] - -- have Hzr : some (some z_R) = Real.toEReal z_R := by simp [Real.toEReal] - -- rw [Hwr, Hzr] at H - -- rw [Hwr, Hzr] - -- clear Hwr - -- clear Hzr - -- apply EReal.coe_le_coe_iff.mpr - -- repeat rw [<- EReal.coe_mul] at H - -- apply EReal.coe_le_coe_iff.mp at H - -- apply le_of_mul_le_mul_left H - -- exact EReal.coe_pos.mp Hr1 + have defTop : some none = (⊤ : EReal) := by simp [Top.top] + have defBot : none = (⊥ : EReal) := by simp [Bot.bot] + cases s + · exfalso + rw [defBot] at Hr1 + simp_all only [not_lt_bot] + rename_i s_nnr + cases s_nnr + · rw [defTop] at Hr2 + exfalso + simp_all only [EReal.zero_lt_top, lt_self_iff_false] + rename_i s_R + have Hsr : some (some s_R) = Real.toEReal s_R := by simp [Real.toEReal] + rw [Hsr] at H + rw [Hsr] at Hr1 + rw [Hsr] at Hr2 + clear Hsr + + cases w + · apply left_eq_inf.mp + rfl + rename_i w_nnr + cases w_nnr + · simp [defTop] at H + rw [EReal.mul_top_of_pos Hr1] at H + have X1 : z = ⊤ := by + cases z + · exfalso + simp at H + rw [defBot] at H + rw [EReal.mul_bot_of_pos] at H + · cases H + · apply Hr1 + rename_i z_nnr + cases z_nnr + · simp [Top.top] + exfalso + apply top_le_iff.mp at H + rename_i z_R + have Hzr : some (some z_R) = Real.toEReal z_R := by simp [Real.toEReal] + rw [Hzr] at H + rw [<- EReal.coe_mul] at H + cases H + rw [defTop, X1] + rename_i w_R + cases z + · simp [defBot] at H + rw [EReal.mul_bot_of_pos] at H + apply le_bot_iff.mp at H + · rw [defBot] + have Hwr : some (some w_R) = Real.toEReal w_R := by simp [Real.toEReal] + rw [Hwr] at H + rw [<- EReal.coe_mul] at H + cases H + · apply Hr1 + rename_i z_nnr + cases z_nnr + · exact right_eq_inf.mp rfl + rename_i z_R + have Hwr : some (some w_R) = Real.toEReal w_R := by simp [Real.toEReal] + have Hzr : some (some z_R) = Real.toEReal z_R := by simp [Real.toEReal] + rw [Hwr, Hzr] at H + rw [Hwr, Hzr] + clear Hwr + clear Hzr + apply EReal.coe_le_coe_iff.mpr + repeat rw [<- EReal.coe_mul] at H + apply EReal.coe_le_coe_iff.mp at H + apply le_of_mul_le_mul_left H + exact EReal.coe_pos.mp Hr1 + +-- MARKUSDE: Cleanup lemma ereal_smul_eq_left {w z : EReal} (s : EReal) (Hr1 : 0 < s) (Hr2 : s < ⊤) (H : s * w = s * z) : w = z := by - sorry -- rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> -- rcases (EReal_cases z) with Hz' | (Hz' | ⟨ z', Hz' ⟩) <;> -- rcases (EReal_cases s) with Hs' | (Hs' | ⟨ s', Hs' ⟩) @@ -725,7 +686,9 @@ lemma ereal_smul_eq_left {w z : EReal} (s : EReal) (Hr1 : 0 < s) (Hr2 : s < ⊤) -- · assumption -- · exfalso -- linarith + sorry + -- Hr2 should not be needed? lemma ereal_smul_left_le {s : EReal} {w z : EReal} (Hr1 : 0 < s) (Hr2 : s < ⊤) (H : w ≤ z) : s * w ≤ s * z := by sorry -- rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> From 211c0de30677e0384b78f56082a87cf619ec6102 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Wed, 26 Jun 2024 15:36:58 -0400 Subject: [PATCH 099/122] Strengthen ereal_smul_left_le --- .../DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean | 5 ++--- SampCert/Util/Log.lean | 4 ++-- 2 files changed, 4 insertions(+), 5 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean index 6d3a00a7..f2bce1ba 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean @@ -667,10 +667,9 @@ theorem privPostProcess_zCDPBound {nq : List T → SLang U} {HNorm : NormalMecha simp apply ofEReal_le_mono apply ereal_smul_left_le - · apply EReal.coe_pos.mpr - apply inv_pos_of_pos + · apply EReal.coe_nonneg.mpr + apply inv_nonneg_of_nonneg linarith - · exact EReal.coe_lt_top (α - 1)⁻¹ apply elog_mono_le.mp apply (DPostPocess_pre HNorm (fun a => f a) h1 h2 (Hac l₁ l₂ h2)) apply Hac diff --git a/SampCert/Util/Log.lean b/SampCert/Util/Log.lean index a2387001..c6aa1671 100644 --- a/SampCert/Util/Log.lean +++ b/SampCert/Util/Log.lean @@ -501,7 +501,7 @@ lemma elog_injective {x y : ENNReal} : elog x = elog y -> x = y := by all_goals case_ENNReal_isReal_zero y rename_i r₁ Hr₁ HPr₁ r₂ Hr₂ HPr₂ intro Hlog_eq - suffices r₁ = r₁ by simp + suffices r₁ = r₂ by simp [this] apply Real.log_injOn_pos all_goals simp_all @@ -689,7 +689,7 @@ lemma ereal_smul_eq_left {w z : EReal} (s : EReal) (Hr1 : 0 < s) (Hr2 : s < ⊤) sorry -- Hr2 should not be needed? -lemma ereal_smul_left_le {s : EReal} {w z : EReal} (Hr1 : 0 < s) (Hr2 : s < ⊤) (H : w ≤ z) : s * w ≤ s * z := by +lemma ereal_smul_left_le {s : EReal} {w z : EReal} (Hr1 : 0 ≤ s) /- (Hr2 : s < ⊤) -/ (H : w ≤ z) : s * w ≤ s * z := by sorry -- rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> -- rcases (EReal_cases z) with Hz' | (Hz' | ⟨ z', Hz' ⟩) From cb95d8f5d67d0bea557d123dd357333f8f6f7519 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Wed, 26 Jun 2024 16:23:24 -0400 Subject: [PATCH 100/122] Inline single use misc log lemmas --- .../DifferentialPrivacy/RenyiDivergence.lean | 8 +-- .../ZeroConcentrated/Postprocessing.lean | 5 +- SampCert/Util/Log.lean | 58 +------------------ 3 files changed, 9 insertions(+), 62 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index e79c413d..4afcb1d1 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -813,11 +813,9 @@ lemma Renyi_Jensen_ENNReal_converse_reduct [MeasurableSpace T] [MeasurableSingl rw [ENNReal.mul_inv_cancel ?G1 ?G2] case G1 => apply Hq case G2 => apply PMF.apply_ne_top - apply ENNReal_toReal_partial_inj at Hext' - rw [<- division_def] + apply (toReal_eq_toReal_iff' (HRJf_nt x) ?G3).mp + case G3 => simp apply Hext' - · apply HRJf_nt - · simp · intro apply PMF.apply_ne_top @@ -1164,7 +1162,7 @@ theorem RenyiDivergence_aux_zero [MeasurableSpace T] [MeasurableSingletonClass T have H' := RenyiDivergence_def_nonneg p q Hac Hα refine (ofEReal_nonneg_eq_iff ?mpr.Hw H').mpr ?mpr.a · simp - simp [H] + simp [H] -- Unused /- diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean index f2bce1ba..5a02428e 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean @@ -666,8 +666,9 @@ theorem privPostProcess_zCDPBound {nq : List T → SLang U} {HNorm : NormalMecha repeat rw [PMF.instFunLike] simp apply ofEReal_le_mono - apply ereal_smul_left_le - · apply EReal.coe_nonneg.mpr + apply mul_le_mul_of_nonneg_left ?G1 ?G2 + case G2 => + apply EReal.coe_nonneg.mpr apply inv_nonneg_of_nonneg linarith apply elog_mono_le.mp diff --git a/SampCert/Util/Log.lean b/SampCert/Util/Log.lean index c6aa1671..f3df3d45 100644 --- a/SampCert/Util/Log.lean +++ b/SampCert/Util/Log.lean @@ -656,62 +656,10 @@ lemma ereal_smul_le_left {w z : EReal} (s : EReal) (Hr1 : 0 < s) (Hr2 : s < ⊤) apply le_of_mul_le_mul_left H exact EReal.coe_pos.mp Hr1 --- MARKUSDE: Cleanup lemma ereal_smul_eq_left {w z : EReal} (s : EReal) (Hr1 : 0 < s) (Hr2 : s < ⊤) (H : s * w = s * z) : w = z := by - -- rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> - -- rcases (EReal_cases z) with Hz' | (Hz' | ⟨ z', Hz' ⟩) <;> - -- rcases (EReal_cases s) with Hs' | (Hs' | ⟨ s', Hs' ⟩) - -- all_goals rw [Hw', Hz'] - -- all_goals simp_all - -- · rw [mul_top_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H - -- rw [mul_bot_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H - -- cases H - -- · rw [mul_bot_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H - -- cases H - -- · rw [mul_top_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H - -- rw [mul_bot_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H - -- cases H - -- · rw [mul_top_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H - -- cases H - -- · rw [mul_bot_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H - -- repeat rw [<- EReal.coe_mul] at H - -- cases H - -- · rw [mul_top_of_pos (by exact EReal.coe_pos.mpr Hr1)] at H - -- repeat rw [<- EReal.coe_mul] at H - -- cases H - -- · repeat rw [<- EReal.coe_mul] at H - -- apply EReal.coe_eq_coe_iff.mp at H - -- simp_all only [mul_eq_mul_left_iff] - -- cases H - -- · assumption - -- · exfalso - -- linarith - sorry - - -- Hr2 should not be needed? -lemma ereal_smul_left_le {s : EReal} {w z : EReal} (Hr1 : 0 ≤ s) /- (Hr2 : s < ⊤) -/ (H : w ≤ z) : s * w ≤ s * z := by - sorry - -- rcases (EReal_cases w) with Hw' | (Hw' | ⟨ w', Hw' ⟩) <;> - -- rcases (EReal_cases z) with Hz' | (Hz' | ⟨ z', Hz' ⟩) - -- all_goals rw [Hw', Hz'] - -- all_goals simp_all - -- · rw [mul_top_of_pos Hr1] - -- exact OrderTop.le_top (s * ⊥) - -- · rw [mul_bot_of_pos Hr1] - -- exact OrderBot.bot_le (s * z'.toEReal) - -- · rw [mul_top_of_pos Hr1] - -- exact OrderTop.le_top (s * w'.toEReal) - -- rcases (EReal_cases s) with Hs' | (Hs' | ⟨ s', Hs' ⟩) - -- all_goals simp_all - -- repeat rw [← EReal.coe_mul] - -- apply EReal.coe_le_coe_iff.mpr - -- exact (mul_le_mul_iff_of_pos_left Hr1).mpr H - -lemma ENNReal_toReal_partial_inj {a b : ENNReal} (H : a.toReal = b.toReal) (H1 : a ≠ ⊤) (H2 : b ≠ ⊤) : a = b := by - sorry - -- rcases a with Ha' | ⟨ a' | Ha' ⟩ <;> - -- rcases b with Hb' | ⟨ b' | Hb' ⟩ - -- all_goals simp_all + apply LE.le.antisymm + · apply ereal_smul_le_left s Hr1 Hr2 (le_of_eq H) + · apply ereal_smul_le_left s Hr1 Hr2 (le_of_eq (id (Eq.symm H))) end misc From 0996fe78dd98f975c2f1d9987df52aa84225056c Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Wed, 26 Jun 2024 16:59:44 -0400 Subject: [PATCH 101/122] all proofs closed --- SampCert/Util/Log.lean | 140 +++++++++++++++++++++++++++++++++-------- 1 file changed, 113 insertions(+), 27 deletions(-) diff --git a/SampCert/Util/Log.lean b/SampCert/Util/Log.lean index f3df3d45..cf523050 100644 --- a/SampCert/Util/Log.lean +++ b/SampCert/Util/Log.lean @@ -241,22 +241,22 @@ lemma ofEReal_le_mono {w z : EReal} (H : w ≤ z) : ofEReal w ≤ ofEReal z := b apply ofReal_le_ofReal assumption - -lemma ofEReal_le_mono_conv_nonneg {w z : EReal} (Hw : 0 ≤ w) (Hle : ofEReal w ≤ ofEReal z) : w ≤ z := by - all_goals case_EReal_isENNReal w - all_goals case_EReal_isENNReal z - · exfalso - rename_i Hw' _ - -- 0 <= w < 0 - sorry - · exfalso - rename_i r Hw' Hr - -- 0 <= w < 0 - sorry - · exfalso - rename_i r Hr Hz' - -- 0 <= w <= z < 0 - sorry +-- True, but unused +-- lemma ofEReal_le_mono_conv_nonneg {w z : EReal} (Hw : 0 ≤ w) (Hle : ofEReal w ≤ ofEReal z) : w ≤ z := by +-- all_goals case_EReal_isENNReal w +-- all_goals case_EReal_isENNReal z +-- · exfalso +-- rename_i Hw' _ +-- -- 0 <= w < 0 +-- sorry +-- · exfalso +-- rename_i r Hw' Hr +-- -- 0 <= w < 0 +-- sorry +-- · exfalso +-- rename_i r Hr Hz' +-- -- 0 <= w <= z < 0 +-- sorry @@ -506,27 +506,115 @@ lemma elog_injective {x y : ENNReal} : elog x = elog y -> x = y := by all_goals simp_all +lemma eexp_zero_iff {w : EReal} : eexp w = 0 <-> w = ⊥ := by + apply Iff.intro + · intro H + all_goals case_EReal_isReal w + rename_i r _ + have Hcont := exp_pos r + linarith + · simp_all + + +lemma elog_bot_iff {x : ENNReal} : elog x = ⊥ <-> x = 0 := by + apply Iff.intro + · intro H + all_goals case_ENNReal_isReal_zero x + · simp_all + + lemma eexp_mono_lt {w z : EReal} : (w < z) <-> eexp w < eexp z := by - sorry + apply Iff.intro + · intro H + all_goals case_EReal_isReal w + · apply bot_lt_iff_ne_bot.mpr + apply bot_lt_iff_ne_bot.mp at H + intro HK + apply H + apply eexp_zero_iff.mp + simp_all + all_goals case_EReal_isReal z + apply (ENNReal.ofReal_lt_ofReal_iff_of_nonneg ?G1).mpr + case G1 => apply exp_nonneg + exact exp_lt_exp.mpr H + · intro H + all_goals case_EReal_isReal w + · apply bot_lt_iff_ne_bot.mpr + intro HK + have H'' : 0 ≠ eexp z := by exact ne_of_lt H + apply H'' + symm + apply eexp_zero_iff.mpr + assumption + all_goals case_EReal_isReal z + rename_i a _ _ _ + have C1 : OfNat.ofNat 0 ≤ rexp a := by exact exp_nonneg a + apply (ENNReal.ofReal_lt_ofReal_iff_of_nonneg C1).mp at H + exact exp_lt_exp.mp H -lemma eexp_mono_le {w z : EReal} : (w <= z) <-> eexp w <= eexp z := by sorry -lemma elog_mono_lt {x y : ENNReal} : (x < y) <-> elog x < elog y := by sorry -lemma elog_mono_le {x y : ENNReal} : (x <= y) <-> elog x <= elog y := by sorry +lemma eexp_mono_le {w z : EReal} : (w <= z) <-> eexp w <= eexp z := by + apply Iff.intro + · intro H + cases (LE.le.lt_or_eq H) + · apply LT.lt.le + apply eexp_mono_lt.mp + assumption + · simp_all + · intro H + cases (LE.le.lt_or_eq H) + · apply LT.lt.le + apply eexp_mono_lt.mpr + assumption + · apply Eq.le + apply eexp_injective + assumption --- iff for log positive --- Special values --- Need to write compat lemmas for ofEReal --- Not sure which ones we'll need but +lemma elog_mono_lt {x y : ENNReal} : (x < y) <-> elog x < elog y := by + apply Iff.intro + · intro H + all_goals case_ENNReal_isReal_zero x + · apply Ne.bot_lt' + intro HK + symm at HK + apply elog_bot_iff.mp at HK + simp [HK] at H + all_goals case_ENNReal_isReal_zero y + apply log_lt_log + · assumption + · assumption + · intro H + all_goals case_ENNReal_isReal_zero x + · apply Ne.bot_lt' + intro HK + symm at HK + simp [HK] at H + all_goals case_ENNReal_isReal_zero y + apply (Real.log_lt_log_iff ?G1 ?G2).mp <;> assumption -set_option pp.coercions false +lemma elog_mono_le {x y : ENNReal} : (x <= y) <-> elog x <= elog y := by + apply Iff.intro + · intro H + cases (LE.le.lt_or_eq H) + · apply LT.lt.le + apply elog_mono_lt.mp + assumption + · simp_all + · intro H + cases (LE.le.lt_or_eq H) + · apply LT.lt.le + apply elog_mono_lt.mpr + assumption + · apply Eq.le + apply elog_injective + assumption end elog_eexp @@ -539,8 +627,6 @@ section misc -- MARKUSDE: cleanup --- ENNReal inequalities --- These are needed to prove the extensded version of Jensen's inequality lemma mul_mul_inv_le_mul_cancel {x y : ENNReal} : (x * y⁻¹) * y ≤ x := by cases x · simp_all From 0d4cb66ae5417ede37d5980d057734fe76006dc4 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Thu, 27 Jun 2024 11:05:31 -0400 Subject: [PATCH 102/122] Change return type of mechanism to PMF --- SampCert/DifferentialPrivacy/Abstract.lean | 300 +++++++++--------- .../DifferentialPrivacy/Pure/Composition.lean | 4 +- .../Pure/Mechanism/Code.lean | 6 +- .../Pure/Mechanism/Properties.lean | 142 +++++---- .../Pure/Postprocessing.lean | 2 +- .../Queries/BoundedMean/Code.lean | 2 +- .../Queries/BoundedMean/Properties.lean | 11 +- .../Queries/BoundedSum/Code.lean | 2 +- .../Queries/Count/Code.lean | 2 +- .../ZeroConcentrated/Composition.lean | 118 +------ .../ZeroConcentrated/DP.lean | 23 +- .../ZeroConcentrated/Mechanism/Code.lean | 6 +- .../Mechanism/Properties.lean | 20 +- .../ZeroConcentrated/Postprocessing.lean | 24 +- SampCert/Samplers/LaplaceGen/Properties.lean | 11 + 15 files changed, 292 insertions(+), 381 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Abstract.lean b/SampCert/DifferentialPrivacy/Abstract.lean index e21e4c82..11ef47b3 100644 --- a/SampCert/DifferentialPrivacy/Abstract.lean +++ b/SampCert/DifferentialPrivacy/Abstract.lean @@ -17,31 +17,28 @@ noncomputable section open Classical Nat Int Real ENNReal -def NonZeroNQ (nq : List T → SLang U) := +def NonZeroNQ (nq : List T → PMF U) := ∀ l : List T, ∀ n : U, nq l n ≠ 0 -def NonTopSum (nq : List T → SLang U) := +def NonTopSum (nq : List T → PMF U) := ∀ l : List T, ∑' n : U, nq l n ≠ ⊤ -def NormalMechanism (q : List T -> SLang U) : Prop := - ∀ l, HasSum (q l) 1 +-- -- FIXME: Remove +-- def NormalMechanism (q : List T -> PMF U) : Prop := +-- ∀ l, HasSum (q l) 1 namespace SLang abbrev Query (T U : Type) := List T → U -abbrev Mechanism (T U : Type) := List T → SLang U - - --- FIXME: Move composition of normal mechanism here +abbrev Mechanism (T U : Type) := List T → PMF U +-- instance : Coe (Mechanism T U) (List T -> SLang U) where +-- coe m l := m l /-- Product of mechanisms. - -Note that the second mechanism does not depend on the output of the first; this is in currently -in contrast to the notions of composition found in the DP literature. -/ -def privCompose (nq1 : Mechanism T U) (nq2 : Mechanism T V) (l : List T) : SLang (U × V) := do +def privCompose (nq1 : Mechanism T U) (nq2 : Mechanism T V) (l : List T) : PMF (U × V) := do let A ← nq1 l let B ← nq2 l return (A,B) @@ -49,7 +46,7 @@ def privCompose (nq1 : Mechanism T U) (nq2 : Mechanism T V) (l : List T) : SLang /-- Mechanism obtained by applying a post-processing function to a mechanism. -/ -def privPostProcess (nq : Mechanism T U) (pp : U → V) (l : List T) : SLang V := do +def privPostProcess (nq : Mechanism T U) (pp : U → V) (l : List T) : PMF V := do let A ← nq l return pp A @@ -83,40 +80,40 @@ class DPSystem (T : Type) where prop m (ε₁ / ε₂) → prop (privPostProcess m pp) (ε₁ / ε₂) @[simp] -lemma bind_bind_indep (p : Mechanism T U) (q : Mechanism T V) (h : U → V → SLang A) : - (fun l => (p l).probBind (fun a : U => (q l).probBind fun b : V => h a b)) +lemma bind_bind_indep (p : Mechanism T U) (q : Mechanism T V) (h : U → V → PMF A) : + (fun l => (p l) >>= (fun a : U => (q l) >>= fun b : V => h a b)) = - fun l => (privCompose p q l).probBind (fun z => h z.1 z.2) := by + fun l => (privCompose p q l) >>= (fun z => h z.1 z.2) := by ext l x simp [privCompose, tsum_prod'] - apply tsum_congr - intro b - rw [← ENNReal.tsum_mul_left] - apply tsum_congr - intro c - rw [← mul_assoc] - congr 1 - rw [tsum_eq_single b] - . congr 1 - rw [tsum_eq_single c] - . simp - . intro b' h1 - simp - intro h2 - subst h2 - contradiction - . intro b' h1 - rw [tsum_eq_single c] - . simp - intro h2 - subst h2 - contradiction - . intro b'' h2 - simp - intro h3 h4 - subst h3 - subst h4 - contradiction + -- apply tsum_congr + -- intro b + -- rw [← ENNReal.tsum_mul_left] + -- apply tsum_congr + -- intro c + -- rw [← mul_assoc] + -- congr 1 + -- rw [tsum_eq_single b] + -- . congr 1 + -- rw [tsum_eq_single c] + -- . simp + -- . intro b' h1 + -- simp + -- intro h2 + -- subst h2 + -- contradiction + -- . intro b' h1 + -- rw [tsum_eq_single c] + -- . simp + -- intro h2 + -- subst h2 + -- contradiction + -- . intro b'' h2 + -- simp + -- intro h3 h4 + -- subst h3 + -- subst h4 + -- contradiction lemma compose_sum_rw (nq1 : List T → SLang U) (nq2 : List T → SLang V) (b : U) (c : V) (l : List T) : (∑' (a : U), nq1 l a * ∑' (a_1 : V), if b = a ∧ c = a_1 then nq2 l a_1 else 0) = nq1 l b * nq2 l c := by @@ -174,47 +171,48 @@ lemma compose_sum_rw (nq1 : List T → SLang U) (nq2 : List T → SLang V) (b : rw [C] simp -/-- +-- Unused +/- Composed queries are normalizable. -/ -theorem privCompose_NonTopSum {nq1 : List T → SLang U} {nq2 : List T → SLang V} (nt1 : NonTopSum nq1) (nt2 : NonTopSum nq2) : - NonTopSum (privCompose nq1 nq2) := by - simp [NonTopSum] at * - intro l - replace nt1 := nt1 l - replace nt2 := nt2 l - simp [privCompose] - 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 privCompose_NonTopSum {nq1 : Mechanism T U} {nq2 : Mechanism T V} (nt1 : NonTopSum nq1) (nt2 : NonTopSum nq2) : +-- NonTopSum (privCompose nq1 nq2) := by +-- simp [NonTopSum] at * +-- intro l +-- replace nt1 := nt1 l +-- replace nt2 := nt2 l +-- simp [privCompose] +-- 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 /-- All outputs of a composed query have nonzero probability. -/ -theorem privCompose_NonZeroNQ {nq1 : List T → SLang U} {nq2 : List T → SLang V} (nn1 : NonZeroNQ nq1) (nn2 : NonZeroNQ nq2) : +theorem privCompose_NonZeroNQ {nq1 : Mechanism T U} {nq2 : Mechanism T V} (nn1 : NonZeroNQ nq1) (nn2 : NonZeroNQ nq2) : NonZeroNQ (privCompose nq1 nq2) := by simp [NonZeroNQ] at * intro l a b @@ -278,7 +276,7 @@ lemma condition_to_subset (f : U → V) (g : U → ENNReal) (x : V) : simp rw [B] -theorem privPostProcess_NonZeroNQ {nq : List T → SLang U} {f : U → V} (nn : NonZeroNQ nq) (sur : Function.Surjective f) : +theorem privPostProcess_NonZeroNQ {nq : Mechanism T U} {f : U → V} (nn : NonZeroNQ nq) (sur : Function.Surjective f) : NonZeroNQ (privPostProcess nq f) := by simp [NonZeroNQ, Function.Surjective, privPostProcess] at * intros l n @@ -290,75 +288,75 @@ theorem privPostProcess_NonZeroNQ {nq : List T → SLang U} {f : U → V} (nn : . rw [h] . apply nn -theorem privPostProcess_NonTopSum {nq : List T → SLang U} (f : U → V) (nt : NonTopSum nq) : - NonTopSum (privPostProcess nq f) := by - simp [NonTopSum, privPostProcess] at * - intros l - have nt := nt l - rw [← ENNReal.tsum_fiberwise _ f] at nt - conv => - right - left - right - intro n - rw [condition_to_subset] - have A : ∀ x : V, f ⁻¹' {x} = {y | x = f y} := by - aesop - conv => - right - left - right - intro x - rw [← A] - trivial - -lemma privCompose_norm (nq1 : List T → SLang U) (nq2 : List T → SLang V) (HNorm1 : NormalMechanism nq1) (HNorm2 : NormalMechanism nq2) : NormalMechanism (privCompose nq1 nq2) := by - rw [NormalMechanism] at * - intro l - have HNorm1' := HasSum.tsum_eq (HNorm1 l) - have HNorm2' := HasSum.tsum_eq (HNorm2 l) - have HR : (∑' (x : U × V), privCompose nq1 nq2 l x = 1) := by - rw [privCompose] - rw [ENNReal.tsum_prod'] - simp - conv => - lhs - arg 1 - intro a - arg 1 - intro b - rw [compose_sum_rw nq1 (fun l a_1 => nq2 l a_1) a b l] - conv => - lhs - arg 1 - intro a - rw [ENNReal.tsum_mul_left] - rw [ENNReal.tsum_mul_right] - rw [HNorm1'] - rw [HNorm2'] - simp - rw [<- HR] - apply Summable.hasSum - exact ENNReal.summable - -lemma privPostProcess_norm (nq : List T → SLang U) (HNorm : NormalMechanism nq) (f : U -> V) : NormalMechanism (privPostProcess nq f) := by - rw [NormalMechanism] at * - intro l - have HR : (∑' (b : V), privPostProcess nq f l b = 1) := by - rw [privPostProcess] - simp - rw [<- HasSum.tsum_eq (ENNReal.HasSum_fiberwise (HNorm l) f)] - apply tsum_congr - intro v - rw [condition_to_subset] - have Htyeq : ({a | v = f a}) = (f ⁻¹' {v}) := by - rw [Set.ext_iff] - intro x - simp - exact eq_comm - rw [Htyeq] - rw [<- HR] - apply Summable.hasSum - exact ENNReal.summable +-- theorem privPostProcess_NonTopSum {nq : Mechanism T U} (f : U → V) (nt : NonTopSum nq) : +-- NonTopSum (privPostProcess nq f) := by +-- simp [NonTopSum, privPostProcess] at * +-- intros l +-- have nt := nt l +-- rw [← ENNReal.tsum_fiberwise _ f] at nt +-- conv => +-- right +-- left +-- right +-- intro n +-- rw [condition_to_subset] +-- have A : ∀ x : V, f ⁻¹' {x} = {y | x = f y} := by +-- aesop +-- conv => +-- right +-- left +-- right +-- intro x +-- rw [← A] +-- trivial -end SLang +-- lemma privCompose_norm (nq1 : List T → PMF U) (nq2 : List T → PMF V) (HNorm1 : NormalMechanism nq1) (HNorm2 : NormalMechanism nq2) : NormalMechanism (privCompose nq1 nq2) := by +-- rw [NormalMechanism] at * +-- intro l +-- have HNorm1' := HasSum.tsum_eq (HNorm1 l) +-- have HNorm2' := HasSum.tsum_eq (HNorm2 l) +-- have HR : (∑' (x : U × V), privCompose nq1 nq2 l x = 1) := by +-- rw [privCompose] +-- rw [ENNReal.tsum_prod'] +-- simp +-- conv => +-- lhs +-- arg 1 +-- intro a +-- arg 1 +-- intro b +-- rw [compose_sum_rw nq1 (fun l a_1 => nq2 l a_1) a b l] +-- conv => +-- lhs +-- arg 1 +-- intro a +-- rw [ENNReal.tsum_mul_left] +-- rw [ENNReal.tsum_mul_right] +-- rw [HNorm1'] +-- rw [HNorm2'] +-- simp +-- rw [<- HR] +-- apply Summable.hasSum +-- exact ENNReal.summable +-- +-- lemma privPostProcess_norm (nq : List T → SLang U) (HNorm : NormalMechanism nq) (f : U -> V) : NormalMechanism (privPostProcess nq f) := by +-- rw [NormalMechanism] at * +-- intro l +-- have HR : (∑' (b : V), privPostProcess nq f l b = 1) := by +-- rw [privPostProcess] +-- simp +-- rw [<- HasSum.tsum_eq (ENNReal.HasSum_fiberwise (HNorm l) f)] +-- apply tsum_congr +-- intro v +-- rw [condition_to_subset] +-- have Htyeq : ({a | v = f a}) = (f ⁻¹' {v}) := by +-- rw [Set.ext_iff] +-- intro x +-- simp +-- exact eq_comm +-- rw [Htyeq] +-- rw [<- HR] +-- apply Summable.hasSum +-- exact ENNReal.summable +-- +-- end SLang diff --git a/SampCert/DifferentialPrivacy/Pure/Composition.lean b/SampCert/DifferentialPrivacy/Pure/Composition.lean index 18504502..8bcfd820 100644 --- a/SampCert/DifferentialPrivacy/Pure/Composition.lean +++ b/SampCert/DifferentialPrivacy/Pure/Composition.lean @@ -14,7 +14,7 @@ open Classical Set namespace SLang -theorem PureDP_Compose' {nq1 : Mechanism T U} {nq2 : List T → SLang V} {ε₁ ε₂ ε₃ ε₄ : ℕ+} (h1 : PureDP nq1 ((ε₁ : ℝ) / ε₂)) (h2 : PureDP nq2 ((ε₃ : ℝ) / ε₄)) : +theorem PureDP_Compose' {nq1 : Mechanism T U} {nq2 : Mechanism T V} {ε₁ ε₂ ε₃ ε₄ : ℕ+} (h1 : PureDP nq1 ((ε₁ : ℝ) / ε₂)) (h2 : PureDP nq2 ((ε₃ : ℝ) / ε₄)) : DP (privCompose nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by simp [PureDP] at * rcases h1 with ⟨h1a, _⟩ @@ -57,7 +57,7 @@ theorem PureDP_Compose' {nq1 : Mechanism T U} {nq2 : List T → SLang V} {ε₁ . aesop . aesop -theorem PureDP_Compose (nq1 : List T → SLang U) (nq2 : List T → SLang V) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h : PureDP nq1 ((ε₁ : ℝ) / ε₂)) (h' : PureDP nq2 ((ε₃ : ℝ) / ε₄)) : +theorem PureDP_Compose (nq1 : Mechanism T U) (nq2 : Mechanism T V) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h : PureDP nq1 ((ε₁ : ℝ) / ε₂)) (h' : PureDP nq2 ((ε₃ : ℝ) / ε₄)) : PureDP (privCompose nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by simp [PureDP] at * have hc := h diff --git a/SampCert/DifferentialPrivacy/Pure/Mechanism/Code.lean b/SampCert/DifferentialPrivacy/Pure/Mechanism/Code.lean index b8ff5a4e..dbe43e3a 100644 --- a/SampCert/DifferentialPrivacy/Pure/Mechanism/Code.lean +++ b/SampCert/DifferentialPrivacy/Pure/Mechanism/Code.lean @@ -3,13 +3,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 SampCert.Samplers.LaplaceGen.Code +import SampCert.Samplers.LaplaceGen.Properties noncomputable section namespace SLang -def privNoisedQueryPure (query : List T → ℤ) (Δ : ℕ+) (ε₁ ε₂ : ℕ+) (l : List T) : SLang ℤ := do - DiscreteLaplaceGenSample (Δ * ε₂) ε₁ (query l) +def privNoisedQueryPure (query : List T → ℤ) (Δ : ℕ+) (ε₁ ε₂ : ℕ+) (l : List T) : PMF ℤ := do + DiscreteLaplaceGenSamplePMF (Δ * ε₂) ε₁ (query l) end SLang diff --git a/SampCert/DifferentialPrivacy/Pure/Mechanism/Properties.lean b/SampCert/DifferentialPrivacy/Pure/Mechanism/Properties.lean index ff3d3605..8b3ecefc 100644 --- a/SampCert/DifferentialPrivacy/Pure/Mechanism/Properties.lean +++ b/SampCert/DifferentialPrivacy/Pure/Mechanism/Properties.lean @@ -15,17 +15,18 @@ namespace SLang theorem NoisedQuery_NonZeroNQPureDP (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : NonZeroNQ (privNoisedQueryPure query Δ ε₁ ε₂) := by - simp [NonZeroNQ, privNoisedQueryPure] - intro l n - apply Real.mul_pos - . rw [_root_.div_pos_iff] - left - constructor - . aesop - . have A : 0 < rexp (↑↑ε₁ / (↑↑Δ * ↑↑ε₂)) := by - apply exp_pos - apply add_pos A Real.zero_lt_one - . apply exp_pos + sorry + -- simp [NonZeroNQ, privNoisedQueryPure] + -- intro l n + -- apply Real.mul_pos + -- . rw [_root_.div_pos_iff] + -- left + -- constructor + -- . aesop + -- . have A : 0 < rexp (↑↑ε₁ / (↑↑Δ * ↑↑ε₂)) := by + -- apply exp_pos + -- apply add_pos A Real.zero_lt_one + -- . apply exp_pos theorem natAbs_to_abs (a b : ℤ) : (a - b).natAbs = |(a : ℝ) - (b : ℝ)| := by @@ -53,69 +54,70 @@ theorem NoisedQuery_PureDP' (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bo simp [DP_singleton] at * intros l₁ l₂ neighbours x simp [privNoisedQueryPure] - rw [← ENNReal.ofReal_div_of_pos] - . apply ofReal_le_ofReal - rw [division_def] - rw [mul_inv] - rw [← mul_assoc] - conv => - left - left - rw [mul_assoc] - right - rw [mul_comm] - conv => - left - left - rw [← mul_assoc] - left - rw [mul_inv_cancel (normalizing_constant_nonzero ε₁ ε₂ Δ)] - simp only [one_mul] - rw [← division_def] - rw [← exp_sub] - simp only [sub_neg_eq_add, exp_le_exp] - rw [neg_div'] - rw [div_add_div_same] - rw [division_def] - apply (mul_inv_le_iff' _).mpr - . have B : (ε₁ : ℝ) / ε₂ * (Δ * ε₂ / ε₁) = Δ := by - ring_nf - simp - field_simp - rw [B] - clear B + sorry + -- rw [← ENNReal.ofReal_div_of_pos] + -- . apply ofReal_le_ofReal + -- rw [division_def] + -- rw [mul_inv] + -- rw [← mul_assoc] + -- conv => + -- left + -- left + -- rw [mul_assoc] + -- right + -- rw [mul_comm] + -- conv => + -- left + -- left + -- rw [← mul_assoc] + -- left + -- rw [mul_inv_cancel (normalizing_constant_nonzero ε₁ ε₂ Δ)] + -- simp only [one_mul] + -- rw [← division_def] + -- rw [← exp_sub] + -- simp only [sub_neg_eq_add, exp_le_exp] + -- rw [neg_div'] + -- rw [div_add_div_same] + -- rw [division_def] + -- apply (mul_inv_le_iff' _).mpr + -- . have B : (ε₁ : ℝ) / ε₂ * (Δ * ε₂ / ε₁) = Δ := by + -- ring_nf + -- simp + -- field_simp + -- rw [B] + -- clear B - rw [add_comm] - ring_nf - -- Triangle inequality - have C := @abs_sub_abs_le_abs_sub ℝ _ ((x : ℝ) - (query l₂)) ((x : ℝ) - (query l₁)) - apply le_trans C - clear C - simp + -- rw [add_comm] + -- ring_nf + -- -- Triangle inequality + -- have C := @abs_sub_abs_le_abs_sub ℝ _ ((x : ℝ) - (query l₂)) ((x : ℝ) - (query l₁)) + -- apply le_trans C + -- clear C + -- simp - simp [sensitivity] at bounded_sensitivity - replace bounded_sensitivity := bounded_sensitivity l₁ l₂ neighbours + -- simp [sensitivity] at bounded_sensitivity + -- replace bounded_sensitivity := bounded_sensitivity l₁ l₂ neighbours - rw [← natAbs_to_abs] - exact Nat.cast_le.mpr bounded_sensitivity + -- rw [← natAbs_to_abs] + -- exact Nat.cast_le.mpr bounded_sensitivity - . simp - . rw [_root_.mul_pos_iff] - left - constructor - . rw [_root_.div_pos_iff] - left - have A : 1 < rexp ((ε₁ : ℝ) / (Δ * ε₂)) := by - rw [← exp_zero] - apply exp_lt_exp.mpr - simp - constructor - . simp [A] - . apply @lt_trans _ _ _ 2 _ - . simp - . rw [← one_add_one_eq_two] - exact (add_lt_add_iff_right 1).mpr A - . apply exp_pos + -- . simp + -- . rw [_root_.mul_pos_iff] + -- left + -- constructor + -- . rw [_root_.div_pos_iff] + -- left + -- have A : 1 < rexp ((ε₁ : ℝ) / (Δ * ε₂)) := by + -- rw [← exp_zero] + -- apply exp_lt_exp.mpr + -- simp + -- constructor + -- . simp [A] + -- . apply @lt_trans _ _ _ 2 _ + -- . simp + -- . rw [← one_add_one_eq_two] + -- exact (add_lt_add_iff_right 1).mpr A + -- . apply exp_pos theorem NoisedQuery_PureDP (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_sensitivity : sensitivity query Δ) : PureDP (privNoisedQueryPure query Δ ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by diff --git a/SampCert/DifferentialPrivacy/Pure/Postprocessing.lean b/SampCert/DifferentialPrivacy/Pure/Postprocessing.lean index d8c12a4d..10146a97 100644 --- a/SampCert/DifferentialPrivacy/Pure/Postprocessing.lean +++ b/SampCert/DifferentialPrivacy/Pure/Postprocessing.lean @@ -37,7 +37,7 @@ theorem PureDP_PostProcess' {nq : Mechanism T U} {ε₁ ε₂ : ℕ+} (h : PureD exact Real.exp_pos ((ε₁: ℝ) / ε₂) . simp -theorem PureDP_PostProcess {f : U → V} (sur : Function.Surjective f) (nq : List T → SLang U) (ε₁ ε₂ : ℕ+) (h : PureDP nq ((ε₁ : ℝ) / ε₂)) : +theorem PureDP_PostProcess {f : U → V} (sur : Function.Surjective f) (nq : Mechanism T U) (ε₁ ε₂ : ℕ+) (h : PureDP nq ((ε₁ : ℝ) / ε₂)) : PureDP (privPostProcess nq f) (((ε₁ : ℝ) / ε₂)) := by simp [PureDP] at * have hc := h diff --git a/SampCert/DifferentialPrivacy/Queries/BoundedMean/Code.lean b/SampCert/DifferentialPrivacy/Queries/BoundedMean/Code.lean index 66fab2ec..82699dc0 100644 --- a/SampCert/DifferentialPrivacy/Queries/BoundedMean/Code.lean +++ b/SampCert/DifferentialPrivacy/Queries/BoundedMean/Code.lean @@ -22,7 +22,7 @@ variable [dps : DPSystem ℕ] /-- Compute a noised mean using a noised sum and noised count -/ -def privNoisedBoundedMean (U : ℕ+) (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℚ := do +def privNoisedBoundedMean (U : ℕ+) (ε₁ ε₂ : ℕ+) (l : List ℕ) : PMF ℚ := do let S ← privNoisedBoundedSum U ε₁ (2 * ε₂) l let C ← privNoisedCount ε₁ (2 * ε₂) l return S / C diff --git a/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean b/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean index 43befc61..f71eca56 100644 --- a/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean @@ -46,10 +46,11 @@ theorem privNoisedBoundedMean_DP (U : ℕ+) (ε₁ ε₂ : ℕ+) : dps.prop (privNoisedBoundedMean U ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by unfold privNoisedBoundedMean simp - apply dps.postprocess_prop div_surjective - rw [budget_split] - apply dps.compose_prop - . apply privNoisedBoundedSum_DP - . apply privNoisedCount_DP + sorry + -- apply dps.postprocess_prop div_surjective + -- rw [budget_split] + -- apply dps.compose_prop + -- . apply privNoisedBoundedSum_DP + -- . apply privNoisedCount_DP end SLang diff --git a/SampCert/DifferentialPrivacy/Queries/BoundedSum/Code.lean b/SampCert/DifferentialPrivacy/Queries/BoundedSum/Code.lean index 96c10995..c6875dcc 100644 --- a/SampCert/DifferentialPrivacy/Queries/BoundedSum/Code.lean +++ b/SampCert/DifferentialPrivacy/Queries/BoundedSum/Code.lean @@ -29,7 +29,7 @@ def exactBoundedSum (U : ℕ+) (l : List ℕ) : ℤ := /-- Noised bounded sum query obtained by applying the DP noise mechanism to the bounded sum. -/ -def privNoisedBoundedSum (U : ℕ+) (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℤ := do +def privNoisedBoundedSum (U : ℕ+) (ε₁ ε₂ : ℕ+) (l : List ℕ) : PMF ℤ := do dps.noise (exactBoundedSum U) U ε₁ ε₂ l end SLang diff --git a/SampCert/DifferentialPrivacy/Queries/Count/Code.lean b/SampCert/DifferentialPrivacy/Queries/Count/Code.lean index 7e841c6e..6e44aa23 100644 --- a/SampCert/DifferentialPrivacy/Queries/Count/Code.lean +++ b/SampCert/DifferentialPrivacy/Queries/Count/Code.lean @@ -26,7 +26,7 @@ def exactCount (l : List T) : ℤ := List.length l /-- Noised counting mechanism from the DP system -/ -def privNoisedCount (ε₁ ε₂ : ℕ+) (l : List T) : SLang ℤ := do +def privNoisedCount (ε₁ ε₂ : ℕ+) (l : List T) : PMF ℤ := do dps.noise exactCount 1 ε₁ ε₂ l end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean index 013a9867..163b82cc 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean @@ -25,63 +25,13 @@ variable [Inhabited V] variable [MeasurableSpace U] [MeasurableSingletonClass U] [Countable U] variable [MeasurableSpace V] [MeasurableSingletonClass V] [Countable V] --- lemma ENNReal_toReal_NZ (x : ENNReal) (h1 : x ≠ 0) (h2 : x ≠ ⊤) : --- x.toReal ≠ 0 := by --- unfold ENNReal.toReal --- unfold ENNReal.toNNReal --- simp --- constructor ; any_goals trivial --- --- lemma simp_α_1 {α : ℝ} (h : 1 < α) : 0 < α := by --- apply @lt_trans _ _ _ 1 _ _ h --- simp only [zero_lt_one] - -/- -The Renyi Divergence between neighbouring inputs of noised queries is nonzero. --/ --- theorem Renyi_noised_query_NZ {nq : List T → SLang U} {HNorm : NormalMechanism nq} {α ε : ℝ} --- (h1 : 1 < α) {l₁ l₂ : List T} (h2 : Neighbour l₁ l₂) (h3 : zCDPBound nq HNorm ε) --- (h4 : NonZeroNQ nq) (h5 : NonTopRDNQ nq) (nts : NonTopNQ nq) : --- (∑' (i : U), nq l₁ i ^ α * nq l₂ i ^ (1 - α)).toReal ≠ 0 := by --- simp [zCDPBound] 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_toReal_NZ --- . by_contra CONTRA --- rw [ENNReal.tsum_eq_zero] at CONTRA --- replace CONTRA := CONTRA default --- replace h6 := h6 default --- replace h7 := h7 default --- 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₂ default --- contradiction --- . exact h5 - /-- Composed queries satisfy zCDP Renyi divergence bound. -/ -theorem privCompose_zCDPBound {nq1 : List T → SLang U} {nq2 : List T → SLang V} {HNorm1 : ∀ l, HasSum (nq1 l) 1} {HNorm2 : ∀ l, HasSum (nq2 l) 1} {ε₁ ε₂ ε₃ ε₄ : ℕ+} - (h1 : zCDPBound nq1 HNorm1 ((ε₁ : ℝ) / ε₂)) (h2 : zCDPBound nq2 HNorm2 ((ε₃ : ℝ) / ε₄)) (Ha1 : ACNeighbour nq1) (Ha2 : ACNeighbour nq2): - zCDPBound (privCompose nq1 nq2) (privCompose_norm nq1 nq2 HNorm1 HNorm2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by +theorem privCompose_zCDPBound {nq1 : Mechanism T U} {nq2 : Mechanism T V} {ε₁ ε₂ ε₃ ε₄ : ℕ+} + (h1 : zCDPBound nq1 ((ε₁ : ℝ) / ε₂)) (h2 : zCDPBound nq2 ((ε₃ : ℝ) / ε₄)) + (Ha1 : ACNeighbour nq1) (Ha2 : ACNeighbour nq2) : + zCDPBound (privCompose nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by simp [privCompose, RenyiDivergence, zCDPBound] intro α h3 l₁ l₂ h4 simp [zCDPBound] at h1 h2 @@ -94,27 +44,13 @@ theorem privCompose_zCDPBound {nq1 : List T → SLang U} {nq2 : List T → SLang rw [PMF.instFunLike] simp repeat rw [SLang.toPMF] - have CG1 (b : U) : nq1 l₂ b ≠ ⊤ := by - have HS : (∑'(u : U), nq1 l₂ u) = 1 := by exact HasSum.tsum_eq (HNorm1 l₂) - have Hs : nq1 l₂ b ≤ (∑'(u : U), nq1 l₂ u) := by exact ENNReal.le_tsum b - rw [HS] at Hs - rename_i inst inst_1 - simp_all only [ne_eq] - apply Aesop.BuiltinRules.not_intro - intro a - simp_all only [top_le_iff, one_ne_top] - have CG2 (b : V) : nq2 l₂ b ≠ ⊤ := by - have HS : (∑'(u : V), nq2 l₂ u) = 1 := by exact HasSum.tsum_eq (HNorm2 l₂) - have Hs : nq2 l₂ b ≤ (∑'(u : V), nq2 l₂ u) := by exact ENNReal.le_tsum b - rw [HS] at Hs - rename_i inst inst_1 - simp_all only [ne_eq] - apply Aesop.BuiltinRules.not_intro - intro a - simp_all only [top_le_iff, one_ne_top] + have CG1 (b : U) : nq1 l₂ b ≠ ⊤ := by exact PMF.apply_ne_top (nq1 l₂) b + have CG2 (b : V) : nq2 l₂ b ≠ ⊤ := by exact PMF.apply_ne_top (nq2 l₂) b simp rw [tsum_prod' ENNReal.summable (fun b : U => ENNReal.summable)] - . simp + sorry + /- + simp conv => left right @@ -265,27 +201,7 @@ theorem privCompose_zCDPBound {nq1 : List T → SLang U} {nq2 : List T → SLang · apply Left.add_nonneg · apply log_nonneg_1 · apply log_nonneg_2 - -/- -All outputs of a composed query have finite probability. --/ --- theorem privCompose_NonTopNQ {nq1 : List T → SLang U} {nq2 : List T → SLang V} (nt1 : NonTopNQ nq1) (nt2 : NonTopNQ nq2) : --- NonTopNQ (privCompose nq1 nq2) := by --- simp [NonTopNQ] at * --- intro l a b --- replace nt1 := nt1 l a --- replace nt2 := nt2 l b --- simp [privCompose] --- rw [compose_sum_rw] --- rw [mul_eq_top] --- intro H --- cases H --- . rename_i H --- cases H --- contradiction --- . rename_i H --- cases H --- contradiction + -/ /- Renyi divergence beteeen composed queries on neighbours are finite. @@ -354,7 +270,8 @@ Renyi divergence beteeen composed queries on neighbours are finite. -- contradiction -def privCompose_AC (nq1 : List T → SLang U) (nq2 : List T → SLang V) (Hac1 : ACNeighbour nq1) (Hac2 : ACNeighbour nq2) : ACNeighbour (privCompose nq1 nq2) := by +def privCompose_AC (nq1 : Mechanism T U) (nq2 : Mechanism T V) (Hac1 : ACNeighbour nq1) (Hac2 : ACNeighbour nq2) : + ACNeighbour (privCompose nq1 nq2) := by rw [ACNeighbour] at * intro l₁ l₂ Hn have Hac1 := Hac1 l₁ l₂ Hn @@ -381,14 +298,13 @@ def privCompose_AC (nq1 : List T → SLang U) (nq2 : List T → SLang V) (Hac1 : /-- ``privCompose`` satisfies zCDP -/ -theorem privCompose_zCDP (nq1 : List T → SLang U) (nq2 : List T → SLang V) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h : zCDP nq1 ((ε₁ : ℝ) / ε₂)) (h' : zCDP nq2 ((ε₃ : ℝ) / ε₄)) : - zCDP (privCompose nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by +theorem privCompose_zCDP (nq1 : Mechanism T U) (nq2 : Mechanism T V) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h : zCDP nq1 ((ε₁ : ℝ) / ε₂)) (h' : zCDP nq2 ((ε₃ : ℝ) / ε₄)) : + zCDP (privCompose nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by simp [zCDP] at * - rcases h with ⟨ Hac1, ⟨ Hn1, Hb1 ⟩⟩ - rcases h' with ⟨ Hac2, ⟨ Hn2, Hb2 ⟩⟩ + rcases h with ⟨ Hac1, Hb1 ⟩ + rcases h' with ⟨ Hac2, Hb2 ⟩ apply And.intro · exact privCompose_AC nq1 nq2 Hac1 Hac2 - · exists (privCompose_norm nq1 nq2 Hn1 Hn2) - exact privCompose_zCDPBound Hb1 Hb2 Hac1 Hac2 + · exact privCompose_zCDPBound Hb1 Hb2 Hac1 Hac2 end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean index c412f328..70f3ae86 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean @@ -31,12 +31,9 @@ Inequality defining ``(ε^2)/2``-zCDP. All ``ε``-DP mechanisms satisfy this bound (though not all mechanisms satisfying this bound are ``ε``-DP). -/ -def zCDPBound (q : List T → SLang U) (HN : NormalMechanism q) (ε : ℝ) : Prop := +def zCDPBound (q : List T → PMF U) (ε : ℝ) : Prop := ∀ α : ℝ, 1 < α → ∀ l₁ l₂ : List T, Neighbour l₁ l₂ → - RenyiDivergence (SLang.toPMF (q l₁) (HN l₁)) (SLang.toPMF (q l₂) (HN l₂)) α ≤ ENNReal.ofReal ((1/2) * ε ^ 2 * α) - -def NonTopNQ (nq : List T → SLang U) := - ∀ l : List T, ∀ n : U, nq l n ≠ ⊤ + RenyiDivergence (q l₁) (q l₂) α ≤ ENNReal.ofReal ((1/2) * ε ^ 2 * α) /-- The Renyi divergence between neighbouring elements of the output of ``nq`` is finite. @@ -45,21 +42,9 @@ 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 - α) ≠ ⊤ -def ACNeighbour (p : List T -> SLang U) : Prop := ∀ l₁ l₂ , Neighbour l₁ l₂ -> AbsCts (p l₁) (p l₂) - - +def ACNeighbour (p : List T -> PMF U) : Prop := ∀ l₁ l₂, Neighbour l₁ l₂ -> AbsCts (p l₁) (p l₂) /-- The mechanism ``q`` is ``(ε^2)/2``-zCDP -/ -def zCDP (q : List T → SLang U) (ε : ℝ) : Prop := - ACNeighbour q - -- FIXME: Instead of a side condition, make normalization intrinsic to the Mechanism type. - ∧ ∃ (HN : NormalMechanism q), zCDPBound q HN ε - - - - -- ∧ NonZeroNQ q - -- ∧ NonTopSum q - -- ∧ NonTopNQ q - -- ∧ NonTopRDNQ q +def zCDP (q : List T → PMF U) (ε : ℝ) : Prop := ACNeighbour q ∧ zCDPBound q ε diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Code.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Code.lean index 7af53520..232c510e 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Code.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Code.lean @@ -3,7 +3,7 @@ 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.Samplers.GaussianGen.Code +import SampCert.Samplers.GaussianGen.Properties /-! # ``privNoisedQuery`` Implementation @@ -18,7 +18,7 @@ namespace SLang /-- Mechanism obtained by postcomposing query with by gaussian noise (of variance ``(Δ ε₁ / ε₂)^2``). -/ -def privNoisedQuery (query : List T → ℤ) (Δ : ℕ+) (ε₁ ε₂ : ℕ+) (l : List T) : SLang ℤ := do - DiscreteGaussianGenSample (Δ * ε₂) ε₁ (query l) +def privNoisedQuery (query : List T → ℤ) (Δ : ℕ+) (ε₁ ε₂ : ℕ+) (l : List T) : PMF ℤ := do + DiscreteGaussianGenPMF (Δ * ε₂) ε₁ (query l) end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean index 63265cb0..edb8d992 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean @@ -18,25 +18,23 @@ open Classical Nat Int Real ENNReal MeasureTheory Measure namespace SLang -lemma privNoisedQuery_norm (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : -- (bounded_sensitivity : sensitivity query Δ) : - NormalMechanism (privNoisedQuery query Δ ε₁ ε₂) := by - rw [NormalMechanism] - intro l - rw [privNoisedQuery] - exact DiscreteGaussianGen_sum (Δ * ε₂) ε₁ (query l) +-- lemma privNoisedQuery_norm (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : -- (bounded_sensitivity : sensitivity query Δ) : +-- NormalMechanism (privNoisedQuery query Δ ε₁ ε₂) := by +-- rw [NormalMechanism] +-- intro l +-- rw [privNoisedQuery] +-- exact DiscreteGaussianGen_sum (Δ * ε₂) ε₁ (query l) -set_option pp.coercions false +-- set_option pp.coercions false -- set_option pp.notation false -- set_option pp.all true -- set_option pp.universes false --- lemma L (x : ℝ) : x = 1 * x := by exact? - /-- The zCDP mechanism with bounded sensitivity satisfies the bound for ``(Δε₂/ε₁)^2``-zCDP. -/ theorem privNoisedQuery_zCDPBound (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_sensitivity : sensitivity query Δ) : - zCDPBound (privNoisedQuery query Δ ε₁ ε₂) (privNoisedQuery_norm query Δ ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by + zCDPBound (privNoisedQuery query Δ ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by simp [zCDPBound, privNoisedQuery] intros α h1 l₁ l₂ h2 have A := @discrete_GaussianGenSample_ZeroConcentrated α h1 (Δ * ε₂) ε₁ (query l₁) (query l₂) @@ -377,7 +375,7 @@ def privNoisedQuery_AC (query : List T -> ℤ) (Δ ε₁ ε₂ : ℕ+) : ACNeigh rw [AbsCts] intro n Hk exfalso - simp [privNoisedQuery, DiscreteGaussianGenSample] at Hk + simp [privNoisedQuery, DiscreteGaussianGenPMF, DiscreteGaussianGenSample, DFunLike.coe] at Hk have Hk := Hk (n - query l₂) simp at Hk have A : ((Δ : ℝ) * ε₂ / ε₁) ≠ 0 := by simp diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean index 5a02428e..b9444f06 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean @@ -126,7 +126,7 @@ variable [Inhabited U] -- rw [one_le_ofReal] -- apply le_of_lt h -def privPostProcess_AC {f : U -> V} (nq : List T → SLang U) (Hac : ACNeighbour nq) : ACNeighbour (privPostProcess nq f) := by +def privPostProcess_AC {f : U -> V} (nq : Mechanism T U) (Hac : ACNeighbour nq) : ACNeighbour (privPostProcess nq f) := by rw [ACNeighbour] at * unfold AbsCts at * intro l₁ l₂ Hn v @@ -631,9 +631,9 @@ theorem DPostPocess_pre {nq : List T → SLang U} (HNorm : ∀ l, HasSum (nq l) left linarith -theorem privPostProcess_zCDPBound {nq : List T → SLang U} {HNorm : NormalMechanism nq} {ε₁ ε₂ : ℕ+} - (h : zCDPBound nq HNorm ((ε₁ : ℝ) / ε₂)) (f : U → V) (Hac : ACNeighbour nq) : - zCDPBound (privPostProcess nq f) (privPostProcess_norm nq HNorm f) ((ε₁ : ℝ) / ε₂) := by +theorem privPostProcess_zCDPBound {nq : Mechanism T U} {ε₁ ε₂ : ℕ+} + (h : zCDPBound nq ((ε₁ : ℝ) / ε₂)) (f : U → V) (Hac : ACNeighbour nq) : + zCDPBound (privPostProcess nq f) ((ε₁ : ℝ) / ε₂) := by simp [privPostProcess, zCDPBound, RenyiDivergence] intro α h1 l₁ l₂ h2 have h' := h @@ -672,10 +672,11 @@ theorem privPostProcess_zCDPBound {nq : List T → SLang U} {HNorm : NormalMecha apply inv_nonneg_of_nonneg linarith apply elog_mono_le.mp - apply (DPostPocess_pre HNorm (fun a => f a) h1 h2 (Hac l₁ l₂ h2)) - apply Hac - apply Neighbour_symm - apply h2 + sorry + -- apply (DPostPocess_pre (fun a => f a) h1 h2 (Hac l₁ l₂ h2)) + -- apply Hac + -- apply Neighbour_symm + -- apply h2 -- theorem privPostProcess_NonTopRDNQ {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) 1} {ε₁ ε₂ : ℕ+} (f : U → V) @@ -697,13 +698,12 @@ theorem privPostProcess_zCDPBound {nq : List T → SLang U} {HNorm : NormalMecha Postprocessing preserves zCDP -/ theorem privPostProcess_zCDP {f : U → V} - (nq : List T → SLang U) (ε₁ ε₂ : ℕ+) (h : zCDP nq ((ε₁ : ℝ) / ε₂)) : + (nq : Mechanism T U) (ε₁ ε₂ : ℕ+) (h : zCDP nq ((ε₁ : ℝ) / ε₂)) : zCDP (privPostProcess nq f) (((ε₁ : ℝ) / ε₂)) := by - rcases h with ⟨ Hac1, ⟨ Hn1, Hb1 ⟩⟩ + rcases h with ⟨ Hac1, Hb1 ⟩ simp [zCDP] at * apply And.intro · exact privPostProcess_AC nq Hac1 - · exists (privPostProcess_norm nq Hn1 f) - exact (privPostProcess_zCDPBound Hb1 f Hac1) + · exact (privPostProcess_zCDPBound Hb1 f Hac1) end SLang diff --git a/SampCert/Samplers/LaplaceGen/Properties.lean b/SampCert/Samplers/LaplaceGen/Properties.lean index 3870123b..81f6c11c 100644 --- a/SampCert/Samplers/LaplaceGen/Properties.lean +++ b/SampCert/Samplers/LaplaceGen/Properties.lean @@ -126,6 +126,17 @@ theorem DiscreteLaplaceGenSample_periodic (num : PNat) (den : PNat) (μ x : ℤ) -- -- sorrrryyyy +theorem DiscreteLaplaceGenSample_sum (num : PNat) (den : PNat) (μ : ℤ) : HasSum (DiscreteLaplaceGenSample num den μ) 1 := by + rw [Summable.hasSum_iff ENNReal.summable] + sorry + -- conv => + -- lhs + -- arg 1 + -- intro b + -- rw [DiscreteGaussianGenSample_apply] + +def DiscreteLaplaceGenSamplePMF (num : PNat) (den : PNat) (μ : ℤ) : PMF ℤ := + ⟨ DiscreteLaplaceGenSample num den μ , DiscreteLaplaceGenSample_sum num den μ ⟩ end SLang From 72e3736dfc1979252d7b31ef958f29dbcff44577 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Thu, 27 Jun 2024 11:37:50 -0400 Subject: [PATCH 103/122] close composition --- SampCert/DifferentialPrivacy/Abstract.lean | 10 +- .../ZeroConcentrated/Composition.lean | 277 ++++++++---------- 2 files changed, 132 insertions(+), 155 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Abstract.lean b/SampCert/DifferentialPrivacy/Abstract.lean index 11ef47b3..9beff83e 100644 --- a/SampCert/DifferentialPrivacy/Abstract.lean +++ b/SampCert/DifferentialPrivacy/Abstract.lean @@ -115,9 +115,9 @@ lemma bind_bind_indep (p : Mechanism T U) (q : Mechanism T V) (h : U → V → P -- subst h4 -- contradiction -lemma compose_sum_rw (nq1 : List T → SLang U) (nq2 : List T → SLang V) (b : U) (c : V) (l : List T) : - (∑' (a : U), nq1 l a * ∑' (a_1 : V), if b = a ∧ c = a_1 then nq2 l a_1 else 0) = nq1 l b * nq2 l c := by - have A : ∀ a : U, ∀ b : U, (∑' (a_1 : V), if b = a ∧ c = a_1 then nq2 l a_1 else 0) = if b = a then (∑' (a_1 : V), if c = a_1 then nq2 l a_1 else 0) else 0 := by +lemma compose_sum_rw (nq1 : U -> ENNReal) (nq2 : V -> ENNReal) (b : U) (c : V) : + (∑' (a : U), nq1 a * ∑' (a_1 : V), if b = a ∧ c = a_1 then nq2 a_1 else 0) = nq1 b * nq2 c := by + have A : ∀ a : U, ∀ b : U, (∑' (a_1 : V), if b = a ∧ c = a_1 then nq2 a_1 else 0) = if b = a then (∑' (a_1 : V), if c = a_1 then nq2 a_1 else 0) else 0 := by intro x y split . rename_i h @@ -135,7 +135,7 @@ lemma compose_sum_rw (nq1 : List T → SLang U) (nq2 : List T → SLang V) (b : rw [A] rw [ENNReal.tsum_eq_add_tsum_ite b] simp - have B : ∀ x : U, (if x = b then 0 else if b = x then nq1 l x * ∑' (a_1 : V), if c = a_1 then nq2 l a_1 else 0 else 0) = 0 := by + have B : ∀ x : U, (if x = b then 0 else if b = x then nq1 x * ∑' (a_1 : V), if c = a_1 then nq2 a_1 else 0 else 0) = 0 := by intro x split . simp @@ -154,7 +154,7 @@ lemma compose_sum_rw (nq1 : List T → SLang U) (nq2 : List T → SLang V) (b : congr 1 rw [ENNReal.tsum_eq_add_tsum_ite c] simp - have C :∀ x : V, (if x = c then 0 else if c = x then nq2 l x else 0) = 0 := by + have C :∀ x : V, (if x = c then 0 else if c = x then nq2 x else 0) = 0 := by intro x split . simp diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean index 163b82cc..1246fa95 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean @@ -25,6 +25,7 @@ variable [Inhabited V] variable [MeasurableSpace U] [MeasurableSingletonClass U] [Countable U] variable [MeasurableSpace V] [MeasurableSingletonClass V] [Countable V] +set_option pp.coercions false /-- Composed queries satisfy zCDP Renyi divergence bound. -/ @@ -44,164 +45,140 @@ theorem privCompose_zCDPBound {nq1 : Mechanism T U} {nq2 : Mechanism T V} {ε₁ rw [PMF.instFunLike] simp repeat rw [SLang.toPMF] - have CG1 (b : U) : nq1 l₂ b ≠ ⊤ := by exact PMF.apply_ne_top (nq1 l₂) b - have CG2 (b : V) : nq2 l₂ b ≠ ⊤ := by exact PMF.apply_ne_top (nq2 l₂) b + have CG1 (b : U) : (nq1 l₂).val b ≠ ⊤ := by + have H := PMF.apply_ne_top (nq1 l₂) b + simp [DFunLike.coe, PMF.instFunLike] at H + apply H + have CG2 (b : V) : (nq2 l₂).val b ≠ ⊤ := by + have H := PMF.apply_ne_top (nq2 l₂) b + simp [DFunLike.coe, PMF.instFunLike] at H + apply H simp rw [tsum_prod' ENNReal.summable (fun b : U => ENNReal.summable)] - sorry - /- - 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_nonneg _ _ (le_of_lt (lt_trans zero_lt_one h3))] - rw [ENNReal.mul_rpow_of_ne_top (CG1 b) (CG2 c)] - rw [mul_assoc] - right - rw [mul_comm] - rw [mul_assoc] - right - rw [mul_comm] - clear CG1 - clear CG2 - 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 [← elog_mul] + simp [PMF.bind, PMF.pure, DFunLike.coe, PMF.instFunLike] + conv => + enter [1, 1, 2, 1, 1, b, 1, c] + rw [compose_sum_rw] + rw [compose_sum_rw] + rw [ENNReal.mul_rpow_of_nonneg _ _ (le_of_lt (lt_trans zero_lt_one h3))] + rw [ENNReal.mul_rpow_of_ne_top (CG1 b) (CG2 c)] + rw [mul_assoc] + right + rw [mul_comm] + rw [mul_assoc] + right + rw [mul_comm] + clear CG1 + clear CG2 + 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 [← elog_mul] - conv at h1 => - lhs - arg 1 - arg 2 - arg 1 - arg 1 - intro x - rw [DFunLike.coe] - rw [PMF.instFunLike] - rw [SLang.toPMF] - rw [SLang.toPMF] - simp - conv at h2 => - lhs - arg 1 - arg 2 - arg 1 - arg 1 - intro x - rw [DFunLike.coe] - rw [PMF.instFunLike] - rw [SLang.toPMF] - rw [SLang.toPMF] - simp + -- conv at h1 => + -- lhs + -- arg 1 + -- arg 2 + -- arg 1 + -- arg 1 + -- intro x + -- rw [DFunLike.coe] + -- rw [PMF.instFunLike] + -- rw [SLang.toPMF] + -- rw [SLang.toPMF] + -- simp + -- conv at h2 => + -- lhs + -- arg 1 + -- arg 2 + -- arg 1 + -- arg 1 + -- intro x + -- rw [DFunLike.coe] + -- rw [PMF.instFunLike] + -- rw [SLang.toPMF] + -- rw [SLang.toPMF] + -- simp - have log_nonneg_1 : 0 ≤ (∑' (i : U), nq1 l₁ i ^ α * nq1 l₂ i ^ (1 - α)).elog := by - have Hac1 : AbsCts ((nq1 l₁).toPMF (HNorm1 l₁)) ((nq1 l₂).toPMF (HNorm1 l₂)) := by exact Ha1 l₁ l₂ h4 - have Hnn1 := (RenyiDivergence_def_log_sum_nonneg ((nq1 l₁).toPMF (HNorm1 l₁)) ((nq1 l₂).toPMF (HNorm1 l₂)) Hac1 h3) - conv at Hnn1 => - rhs - arg 1 - arg 1 - intro x - rw [DFunLike.coe] - rw [PMF.instFunLike] - rw [SLang.toPMF] - rw [SLang.toPMF] - simp - apply Hnn1 - have log_nonneg_2 : 0 ≤ (∑' (i : V), nq2 l₁ i ^ α * nq2 l₂ i ^ (1 - α)).elog := by - have Hac2 : AbsCts ((nq2 l₁).toPMF (HNorm2 l₁)) ((nq2 l₂).toPMF (HNorm2 l₂)) := by exact Ha2 l₁ l₂ h4 - have Hnn2 := (RenyiDivergence_def_log_sum_nonneg ((nq2 l₁).toPMF (HNorm2 l₁)) ((nq2 l₂).toPMF (HNorm2 l₂)) Hac2 h3) - conv at Hnn2 => - rhs - arg 1 - arg 1 - intro x - rw [DFunLike.coe] - rw [PMF.instFunLike] - rw [SLang.toPMF] - rw [SLang.toPMF] - simp - apply Hnn2 + have log_nonneg_1 : 0 ≤ (∑' (i : U), (nq1 l₁).val i ^ α * (nq1 l₂).val i ^ (1 - α)).elog := by + have Hac1 : AbsCts (nq1 l₁) (nq1 l₂) := by exact Ha1 l₁ l₂ h4 + apply (RenyiDivergence_def_log_sum_nonneg (nq1 l₁) (nq1 l₂) Hac1 h3) + have log_nonneg_2 : 0 ≤ (∑' (i : V), (nq2 l₁).val i ^ α * (nq2 l₂).val i ^ (1 - α)).elog := by + have Hac2 : AbsCts (nq2 l₁) (nq2 l₂) := by exact Ha2 l₁ l₂ h4 + have Hnn2 := (RenyiDivergence_def_log_sum_nonneg (nq2 l₁) (nq2 l₂) Hac2 h3) + apply Hnn2 - -- Split up the series - rw [ofEReal_mul_nonneg] - · simp only [ofEReal_real] - -- In order to distribute ofReal, we need the logarithms to be nonegative - rw [ofEReal_plus_nonneg log_nonneg_1 log_nonneg_2] + -- Split up the series + rw [ofEReal_mul_nonneg] + · simp only [ofEReal_real] + -- In order to distribute ofReal, we need the logarithms to be nonegative + rw [ofEReal_plus_nonneg log_nonneg_1 log_nonneg_2] - -- Distribute - rw [CanonicallyOrderedCommSemiring.left_distrib] - apply (@le_trans _ _ _ (ENNReal.ofReal (2⁻¹ * (↑↑ε₁ ^ 2 / ↑↑ε₂ ^ 2) * α) + ENNReal.ofReal (2⁻¹ * (↑↑ε₃ ^ 2 / ↑↑ε₄ ^ 2) * α))) - · -- apply? - apply _root_.add_le_add - · rw [ofEReal_mul_nonneg] at h1 - · apply h1 - · apply EReal.coe_nonneg.mpr - apply inv_nonneg.mpr - linarith - · apply log_nonneg_1 - · rw [ofEReal_mul_nonneg] at h2 - · apply h2 - · apply EReal.coe_nonneg.mpr - apply inv_nonneg.mpr - linarith - · apply log_nonneg_2 - · clear h1 h2 - rw [<- ENNReal.ofReal_add] - · apply ofReal_le_ofReal_iff'.mpr - left - 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 only [inv_pos, Nat.ofNat_pos] - · linarith + -- Distribute + rw [CanonicallyOrderedCommSemiring.left_distrib] + apply (@le_trans _ _ _ (ENNReal.ofReal (2⁻¹ * (↑↑ε₁ ^ 2 / ↑↑ε₂ ^ 2) * α) + ENNReal.ofReal (2⁻¹ * (↑↑ε₃ ^ 2 / ↑↑ε₄ ^ 2) * α))) + · -- apply? + apply _root_.add_le_add + · rw [ofEReal_mul_nonneg] at h1 + · apply h1 + · apply EReal.coe_nonneg.mpr + apply inv_nonneg.mpr + linarith + · apply log_nonneg_1 + · rw [ofEReal_mul_nonneg] at h2 + · apply h2 + · apply EReal.coe_nonneg.mpr + apply inv_nonneg.mpr + linarith + · apply log_nonneg_2 + · clear h1 h2 + rw [<- ENNReal.ofReal_add] + · apply ofReal_le_ofReal_iff'.mpr + left + 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 only [inv_pos, Nat.ofNat_pos] + · linarith + · apply mul_nonneg · apply mul_nonneg - · apply mul_nonneg - · simp - · apply div_nonneg - · exact sq_nonneg ε₁.val.cast - · exact sq_nonneg ε₂.val.cast - · linarith + · simp + · apply div_nonneg + · exact sq_nonneg ε₁.val.cast + · exact sq_nonneg ε₂.val.cast + · linarith + · apply mul_nonneg · apply mul_nonneg - · apply mul_nonneg - · simp - · apply div_nonneg - · exact sq_nonneg ε₃.val.cast - · exact sq_nonneg ε₄.val.cast - · linarith - · simp - linarith - · apply Left.add_nonneg - · apply log_nonneg_1 - · apply log_nonneg_2 - -/ + · simp + · apply div_nonneg + · exact sq_nonneg ε₃.val.cast + · exact sq_nonneg ε₄.val.cast + · linarith + · simp + linarith + · apply Left.add_nonneg + · apply log_nonneg_1 + · apply log_nonneg_2 + /- Renyi divergence beteeen composed queries on neighbours are finite. From 49166b6c31b8d3b6d6b79cdb72e039ad9009960f Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Thu, 27 Jun 2024 11:43:40 -0400 Subject: [PATCH 104/122] Close pure mechanism properties --- .../Pure/Mechanism/Properties.lean | 121 +++++++++--------- 1 file changed, 61 insertions(+), 60 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Pure/Mechanism/Properties.lean b/SampCert/DifferentialPrivacy/Pure/Mechanism/Properties.lean index 8b3ecefc..0cbd83fb 100644 --- a/SampCert/DifferentialPrivacy/Pure/Mechanism/Properties.lean +++ b/SampCert/DifferentialPrivacy/Pure/Mechanism/Properties.lean @@ -54,70 +54,71 @@ theorem NoisedQuery_PureDP' (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bo simp [DP_singleton] at * intros l₁ l₂ neighbours x simp [privNoisedQueryPure] - sorry - -- rw [← ENNReal.ofReal_div_of_pos] - -- . apply ofReal_le_ofReal - -- rw [division_def] - -- rw [mul_inv] - -- rw [← mul_assoc] - -- conv => - -- left - -- left - -- rw [mul_assoc] - -- right - -- rw [mul_comm] - -- conv => - -- left - -- left - -- rw [← mul_assoc] - -- left - -- rw [mul_inv_cancel (normalizing_constant_nonzero ε₁ ε₂ Δ)] - -- simp only [one_mul] - -- rw [← division_def] - -- rw [← exp_sub] - -- simp only [sub_neg_eq_add, exp_le_exp] - -- rw [neg_div'] - -- rw [div_add_div_same] - -- rw [division_def] - -- apply (mul_inv_le_iff' _).mpr - -- . have B : (ε₁ : ℝ) / ε₂ * (Δ * ε₂ / ε₁) = Δ := by - -- ring_nf - -- simp - -- field_simp - -- rw [B] - -- clear B + simp [DiscreteLaplaceGenSamplePMF] + simp [DFunLike.coe, PMF.instFunLike] + rw [← ENNReal.ofReal_div_of_pos] + . apply ofReal_le_ofReal + rw [division_def] + rw [mul_inv] + rw [← mul_assoc] + conv => + left + left + rw [mul_assoc] + right + rw [mul_comm] + conv => + left + left + rw [← mul_assoc] + left + rw [mul_inv_cancel (normalizing_constant_nonzero ε₁ ε₂ Δ)] + simp only [one_mul] + rw [← division_def] + rw [← exp_sub] + simp only [sub_neg_eq_add, exp_le_exp] + rw [neg_div'] + rw [div_add_div_same] + rw [division_def] + apply (mul_inv_le_iff' _).mpr + . have B : (ε₁ : ℝ) / ε₂ * (Δ * ε₂ / ε₁) = Δ := by + ring_nf + simp + field_simp + rw [B] + clear B - -- rw [add_comm] - -- ring_nf - -- -- Triangle inequality - -- have C := @abs_sub_abs_le_abs_sub ℝ _ ((x : ℝ) - (query l₂)) ((x : ℝ) - (query l₁)) - -- apply le_trans C - -- clear C - -- simp + rw [add_comm] + ring_nf + -- Triangle inequality + have C := @abs_sub_abs_le_abs_sub ℝ _ ((x : ℝ) - (query l₂)) ((x : ℝ) - (query l₁)) + apply le_trans C + clear C + simp - -- simp [sensitivity] at bounded_sensitivity - -- replace bounded_sensitivity := bounded_sensitivity l₁ l₂ neighbours + simp [sensitivity] at bounded_sensitivity + replace bounded_sensitivity := bounded_sensitivity l₁ l₂ neighbours - -- rw [← natAbs_to_abs] - -- exact Nat.cast_le.mpr bounded_sensitivity + rw [← natAbs_to_abs] + exact Nat.cast_le.mpr bounded_sensitivity - -- . simp - -- . rw [_root_.mul_pos_iff] - -- left - -- constructor - -- . rw [_root_.div_pos_iff] - -- left - -- have A : 1 < rexp ((ε₁ : ℝ) / (Δ * ε₂)) := by - -- rw [← exp_zero] - -- apply exp_lt_exp.mpr - -- simp - -- constructor - -- . simp [A] - -- . apply @lt_trans _ _ _ 2 _ - -- . simp - -- . rw [← one_add_one_eq_two] - -- exact (add_lt_add_iff_right 1).mpr A - -- . apply exp_pos + . simp + . rw [_root_.mul_pos_iff] + left + constructor + . rw [_root_.div_pos_iff] + left + have A : 1 < rexp ((ε₁ : ℝ) / (Δ * ε₂)) := by + rw [← exp_zero] + apply exp_lt_exp.mpr + simp + constructor + . simp [A] + . apply @lt_trans _ _ _ 2 _ + . simp + . rw [← one_add_one_eq_two] + exact (add_lt_add_iff_right 1).mpr A + . apply exp_pos theorem NoisedQuery_PureDP (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_sensitivity : sensitivity query Δ) : PureDP (privNoisedQueryPure query Δ ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by From 474dec50fa392706d70a88b4c50b7c5094568517 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Thu, 27 Jun 2024 11:44:28 -0400 Subject: [PATCH 105/122] nit --- .../Pure/Mechanism/Properties.lean | 25 ++++++++++--------- 1 file changed, 13 insertions(+), 12 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Pure/Mechanism/Properties.lean b/SampCert/DifferentialPrivacy/Pure/Mechanism/Properties.lean index 0cbd83fb..6924fc7d 100644 --- a/SampCert/DifferentialPrivacy/Pure/Mechanism/Properties.lean +++ b/SampCert/DifferentialPrivacy/Pure/Mechanism/Properties.lean @@ -15,18 +15,19 @@ namespace SLang theorem NoisedQuery_NonZeroNQPureDP (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : NonZeroNQ (privNoisedQueryPure query Δ ε₁ ε₂) := by - sorry - -- simp [NonZeroNQ, privNoisedQueryPure] - -- intro l n - -- apply Real.mul_pos - -- . rw [_root_.div_pos_iff] - -- left - -- constructor - -- . aesop - -- . have A : 0 < rexp (↑↑ε₁ / (↑↑Δ * ↑↑ε₂)) := by - -- apply exp_pos - -- apply add_pos A Real.zero_lt_one - -- . apply exp_pos + simp [NonZeroNQ, privNoisedQueryPure] + simp [DiscreteLaplaceGenSamplePMF] + simp [DFunLike.coe, PMF.instFunLike] + intro l n + apply Real.mul_pos + . rw [_root_.div_pos_iff] + left + constructor + . aesop + . have A : 0 < rexp (↑↑ε₁ / (↑↑Δ * ↑↑ε₂)) := by + apply exp_pos + apply add_pos A Real.zero_lt_one + . apply exp_pos theorem natAbs_to_abs (a b : ℤ) : (a - b).natAbs = |(a : ℝ) - (b : ℝ)| := by From aad9994ff0f3b289d3a2354e4da0ddaaca12ebc5 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Thu, 27 Jun 2024 11:52:19 -0400 Subject: [PATCH 106/122] close postprocessing --- .../ZeroConcentrated/Postprocessing.lean | 18 ++++++++++++------ 1 file changed, 12 insertions(+), 6 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean index b9444f06..1e9789a3 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean @@ -541,7 +541,7 @@ theorem tsum_ne_zero_of_ne_zero {T : Type} [Inhabited T] (f : T → ENNReal) (h contradiction -- Note: Relies on the symmetry of neighbours -theorem DPostPocess_pre {nq : List T → SLang U} (HNorm : ∀ l, HasSum (nq l) 1) +theorem DPostPocess_pre {nq : List T → PMF U} (HNorm : ∀ l, HasSum (nq l) 1) (f : U → V) {α : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (HN : Neighbour l₁ l₂) (Habs : AbsCts (nq l₁) (nq l₂)) (Habs' : AbsCts (nq l₂) (nq l₁)) : (∑' (x : V), @@ -631,6 +631,8 @@ theorem DPostPocess_pre {nq : List T → SLang U} (HNorm : ∀ l, HasSum (nq l) left linarith + + theorem privPostProcess_zCDPBound {nq : Mechanism T U} {ε₁ ε₂ : ℕ+} (h : zCDPBound nq ((ε₁ : ℝ) / ε₂)) (f : U → V) (Hac : ACNeighbour nq) : zCDPBound (privPostProcess nq f) ((ε₁ : ℝ) / ε₂) := by @@ -672,11 +674,15 @@ theorem privPostProcess_zCDPBound {nq : Mechanism T U} {ε₁ ε₂ : ℕ+} apply inv_nonneg_of_nonneg linarith apply elog_mono_le.mp - sorry - -- apply (DPostPocess_pre (fun a => f a) h1 h2 (Hac l₁ l₂ h2)) - -- apply Hac - -- apply Neighbour_symm - -- apply h2 + simp [PMF.bind, PMF.pure] + simp [PMF.instFunLike] + apply DPostPocess_pre + · exact fun l => PMF.hasSum_coe_one (nq l) + · exact h1 + · exact h2 + · exact Hac l₁ l₂ h2 + · apply Hac l₂ l₁ + exact Neighbour_symm l₁ l₂ h2 -- theorem privPostProcess_NonTopRDNQ {nq : List T → SLang U} {HNorm : ∀ l, HasSum (nq l) 1} {ε₁ ε₂ : ℕ+} (f : U → V) From 562ce39a68d2d6d17e7fe4feff94bcc55decd774 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Thu, 27 Jun 2024 14:06:41 -0400 Subject: [PATCH 107/122] close bounded mean proof --- .../Queries/BoundedMean/Properties.lean | 23 ++++++++++++------- 1 file changed, 15 insertions(+), 8 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean b/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean index f71eca56..88de069e 100644 --- a/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean @@ -39,18 +39,25 @@ lemma budget_split (ε₁ ε₂ : ℕ+) : field_simp ring_nf + /-- DP bound for noised mean. -/ theorem privNoisedBoundedMean_DP (U : ℕ+) (ε₁ ε₂ : ℕ+) : dps.prop (privNoisedBoundedMean U ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by - unfold privNoisedBoundedMean - simp - sorry - -- apply dps.postprocess_prop div_surjective - -- rw [budget_split] - -- apply dps.compose_prop - -- . apply privNoisedBoundedSum_DP - -- . apply privNoisedCount_DP + have Hbm_alt : + (privNoisedBoundedMean U ε₁ ε₂) = + (privPostProcess + (privCompose (privNoisedBoundedSum U ε₁ (2 * ε₂)) (privNoisedCount ε₁ (2 * ε₂))) + (fun a : ℤ × ℤ => (a.1 : ℚ) / (a.2 : ℚ))) := by + apply funext + simp [privNoisedBoundedMean, privPostProcess, privCompose] + rw [Hbm_alt] + + apply dps.postprocess_prop div_surjective + rw [budget_split] + apply dps.compose_prop + . apply privNoisedBoundedSum_DP + . apply privNoisedCount_DP end SLang From a7afdc065c4cc72a1b82c8a2b1774b297ea04f71 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Thu, 27 Jun 2024 14:06:58 -0400 Subject: [PATCH 108/122] laplace gen normalizes --- SampCert/Samplers/LaplaceGen/Properties.lean | 22 ++++++++++++++------ 1 file changed, 16 insertions(+), 6 deletions(-) diff --git a/SampCert/Samplers/LaplaceGen/Properties.lean b/SampCert/Samplers/LaplaceGen/Properties.lean index 81f6c11c..5e3a50f4 100644 --- a/SampCert/Samplers/LaplaceGen/Properties.lean +++ b/SampCert/Samplers/LaplaceGen/Properties.lean @@ -126,14 +126,24 @@ theorem DiscreteLaplaceGenSample_periodic (num : PNat) (den : PNat) (μ x : ℤ) -- -- sorrrryyyy + theorem DiscreteLaplaceGenSample_sum (num : PNat) (den : PNat) (μ : ℤ) : HasSum (DiscreteLaplaceGenSample num den μ) 1 := by rw [Summable.hasSum_iff ENNReal.summable] - sorry - -- conv => - -- lhs - -- arg 1 - -- intro b - -- rw [DiscreteGaussianGenSample_apply] + conv => + enter [1, 1, b] + rw [DiscreteLaplaceGenSample_periodic] + rw [<- DiscreteLaplaceSample_normalizes num den] + apply tsum_eq_tsum_of_ne_zero_bij ?Bij + case Bij => exact (fun w => w + μ) + · intro _ _ H + simp at H + exact SetCoe.ext H + · simp + intro z HZ + exists (z - μ) + simp + apply HZ + · simp def DiscreteLaplaceGenSamplePMF (num : PNat) (den : PNat) (μ : ℤ) : PMF ℤ := ⟨ DiscreteLaplaceGenSample num den μ , DiscreteLaplaceGenSample_sum num den μ ⟩ From 41d8b58698e11402b5679a5c735145919a4f78a3 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Thu, 27 Jun 2024 14:17:46 -0400 Subject: [PATCH 109/122] simplify bounded mean proof --- .../Queries/BoundedMean/Properties.lean | 11 ++--------- 1 file changed, 2 insertions(+), 9 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean b/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean index 88de069e..2dd0fb65 100644 --- a/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean @@ -45,15 +45,8 @@ DP bound for noised mean. -/ theorem privNoisedBoundedMean_DP (U : ℕ+) (ε₁ ε₂ : ℕ+) : dps.prop (privNoisedBoundedMean U ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by - have Hbm_alt : - (privNoisedBoundedMean U ε₁ ε₂) = - (privPostProcess - (privCompose (privNoisedBoundedSum U ε₁ (2 * ε₂)) (privNoisedCount ε₁ (2 * ε₂))) - (fun a : ℤ × ℤ => (a.1 : ℚ) / (a.2 : ℚ))) := by - apply funext - simp [privNoisedBoundedMean, privPostProcess, privCompose] - rw [Hbm_alt] - + unfold privNoisedBoundedMean + rw [bind_bind_indep] apply dps.postprocess_prop div_surjective rw [budget_split] apply dps.compose_prop From 75d8d34c3c821b22e204b206e6e3afa6a52376b1 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Thu, 27 Jun 2024 14:26:59 -0400 Subject: [PATCH 110/122] Parameterize postprocessing function requirement --- SampCert/DifferentialPrivacy/Abstract.lean | 6 +++++- SampCert/DifferentialPrivacy/Pure/System.lean | 1 + .../DifferentialPrivacy/Queries/BoundedMean/Properties.lean | 5 +++-- SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean | 3 ++- 4 files changed, 11 insertions(+), 4 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Abstract.lean b/SampCert/DifferentialPrivacy/Abstract.lean index 9beff83e..44045149 100644 --- a/SampCert/DifferentialPrivacy/Abstract.lean +++ b/SampCert/DifferentialPrivacy/Abstract.lean @@ -74,9 +74,13 @@ class DPSystem (T : Type) where compose_prop : {U V : Type} → [MeasurableSpace U] → [Countable U] → [DiscreteMeasurableSpace U] → [Inhabited U] → [MeasurableSpace V] → [Countable V] → [DiscreteMeasurableSpace V] → [Inhabited V] → ∀ m₁ : Mechanism T U, ∀ m₂ : Mechanism T V, ∀ ε₁ ε₂ ε₃ ε₄ : ℕ+, prop m₁ (ε₁ / ε₂) → prop m₂ (ε₃ / ε₄) → prop (privCompose m₁ m₂) ((ε₁ / ε₂) + (ε₃ / ε₄)) /-- + Predicate which must hold on a function for postcompostion + -/ + postprocess_prop_f : {U V : Type} -> (f : U -> V) -> Prop + /-- Notion of privacy is invariant under post-processing. -/ - postprocess_prop : {U : Type} → [MeasurableSpace U] → [Countable U] → [DiscreteMeasurableSpace U] → [Inhabited U] → { pp : U → V } → Function.Surjective pp → ∀ m : Mechanism T U, ∀ ε₁ ε₂ : ℕ+, + postprocess_prop : {U : Type} → [MeasurableSpace U] → [Countable U] → [DiscreteMeasurableSpace U] → [Inhabited U] → { pp : U → V } → postprocess_prop_f pp → ∀ m : Mechanism T U, ∀ ε₁ ε₂ : ℕ+, prop m (ε₁ / ε₂) → prop (privPostProcess m pp) (ε₁ / ε₂) @[simp] diff --git a/SampCert/DifferentialPrivacy/Pure/System.lean b/SampCert/DifferentialPrivacy/Pure/System.lean index 0106fc7e..80dd0f25 100644 --- a/SampCert/DifferentialPrivacy/Pure/System.lean +++ b/SampCert/DifferentialPrivacy/Pure/System.lean @@ -18,6 +18,7 @@ noncomputable instance PureDPSystem : DPSystem T where noise := privNoisedQueryPure noise_prop := NoisedQuery_PureDP compose_prop := PureDP_Compose + postprocess_prop_f := Function.Surjective postprocess_prop := PureDP_PostProcess end SLang diff --git a/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean b/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean index 2dd0fb65..7fa55eaa 100644 --- a/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean @@ -43,11 +43,12 @@ lemma budget_split (ε₁ ε₂ : ℕ+) : /-- DP bound for noised mean. -/ -theorem privNoisedBoundedMean_DP (U : ℕ+) (ε₁ ε₂ : ℕ+) : +theorem privNoisedBoundedMean_DP (U : ℕ+) (ε₁ ε₂ : ℕ+) + (H : dps.postprocess_prop_f (fun a : ℤ × ℤ => (a.1 : ℚ) / (a.2 : ℚ) )) : dps.prop (privNoisedBoundedMean U ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by unfold privNoisedBoundedMean rw [bind_bind_indep] - apply dps.postprocess_prop div_surjective + apply dps.postprocess_prop H rw [budget_split] apply dps.compose_prop . apply privNoisedBoundedSum_DP diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean index 36b699a5..0bb6a14c 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean @@ -27,6 +27,7 @@ noncomputable instance gaussian_zCDPSystem : DPSystem T where noise := privNoisedQuery noise_prop := privNoisedQuery_zCDP compose_prop := privCompose_zCDP - postprocess_prop := (fun _ => privPostProcess_zCDP) -- FIXME: Remove surjectivity from system definition? + postprocess_prop_f _ := True + postprocess_prop _ := privPostProcess_zCDP end SLang From 048e20f60c46cd9e8af58fa5723ebb358a508153 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Thu, 27 Jun 2024 14:52:33 -0400 Subject: [PATCH 111/122] cleanup DifferentialPrivacy/Abstract --- SampCert/DifferentialPrivacy/Abstract.lean | 192 +++--------------- .../Queries/BoundedMean/Properties.lean | 2 +- 2 files changed, 31 insertions(+), 163 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Abstract.lean b/SampCert/DifferentialPrivacy/Abstract.lean index 44045149..d5e36a99 100644 --- a/SampCert/DifferentialPrivacy/Abstract.lean +++ b/SampCert/DifferentialPrivacy/Abstract.lean @@ -10,30 +10,24 @@ import SampCert.Foundations.Basic /-! # Differential Privacy -This file defines a notion of a differential private system. +This file defines an abstract system of differentially private operators. -/ noncomputable section open Classical Nat Int Real ENNReal +/-- +A mechanism is nonzero on its entire domain +-/ def NonZeroNQ (nq : List T → PMF U) := ∀ l : List T, ∀ n : U, nq l n ≠ 0 -def NonTopSum (nq : List T → PMF U) := - ∀ l : List T, ∑' n : U, nq l n ≠ ⊤ - --- -- FIXME: Remove --- def NormalMechanism (q : List T -> PMF U) : Prop := --- ∀ l, HasSum (q l) 1 - namespace SLang abbrev Query (T U : Type) := List T → U -abbrev Mechanism (T U : Type) := List T → PMF U --- instance : Coe (Mechanism T U) (List T -> SLang U) where --- coe m l := m l +abbrev Mechanism (T U : Type) := List T → PMF U /-- Product of mechanisms. @@ -50,74 +44,44 @@ def privPostProcess (nq : Mechanism T U) (pp : U → V) (l : List T) : PMF V := let A ← nq l return pp A - /-- Abstract definition of a differentially private systemm. -/ class DPSystem (T : Type) where /-- - Notion of differential privacy with a paramater (ε-DP, ε-zCDP, etc) + Differential privacy proposition, with one real paramater (ε-DP, ε-zCDP, etc) -/ prop : Mechanism T Z → ℝ → Prop /-- - Noise mechanism (eg. Laplace, Discrete Gaussian, etc) - Paramaterized by a query, sensitivity, and the numerator/denominator ofa (rational) security paramater. + A noise mechanism (eg. Laplace, Discrete Gaussian, etc) + Paramaterized by a query, sensitivity, and a (rational) security paramater. -/ - noise : Query T ℤ → ℕ+ → ℕ+ → ℕ+ → Mechanism T ℤ + noise : Query T ℤ → (sensitivity : ℕ+) → (num : ℕ+) → (den : ℕ+) → Mechanism T ℤ /-- - Adding noise to a query makes it differentially private. + Adding noise to a query makes it private. -/ noise_prop : ∀ q : List T → ℤ, ∀ Δ εn εd : ℕ+, sensitivity q Δ → prop (noise q Δ εn εd) (εn / εd) /-- - Notion of privacy composes by addition. + Privacy composes by addition. -/ compose_prop : {U V : Type} → [MeasurableSpace U] → [Countable U] → [DiscreteMeasurableSpace U] → [Inhabited U] → [MeasurableSpace V] → [Countable V] → [DiscreteMeasurableSpace V] → [Inhabited V] → ∀ m₁ : Mechanism T U, ∀ m₂ : Mechanism T V, ∀ ε₁ ε₂ ε₃ ε₄ : ℕ+, prop m₁ (ε₁ / ε₂) → prop m₂ (ε₃ / ε₄) → prop (privCompose m₁ m₂) ((ε₁ / ε₂) + (ε₃ / ε₄)) /-- - Predicate which must hold on a function for postcompostion + Requirement for postcomposition to hold. -/ postprocess_prop_f : {U V : Type} -> (f : U -> V) -> Prop /-- - Notion of privacy is invariant under post-processing. + Privacy is invariant under post-processing. -/ postprocess_prop : {U : Type} → [MeasurableSpace U] → [Countable U] → [DiscreteMeasurableSpace U] → [Inhabited U] → { pp : U → V } → postprocess_prop_f pp → ∀ m : Mechanism T U, ∀ ε₁ ε₂ : ℕ+, prop m (ε₁ / ε₂) → prop (privPostProcess m pp) (ε₁ / ε₂) @[simp] -lemma bind_bind_indep (p : Mechanism T U) (q : Mechanism T V) (h : U → V → PMF A) : - (fun l => (p l) >>= (fun a : U => (q l) >>= fun b : V => h a b)) - = - fun l => (privCompose p q l) >>= (fun z => h z.1 z.2) := by +lemma bind_bind_indep (p : Mechanism T U) (q : Mechanism T V) (h : U → V → PMF A) : + (fun l => (p l) >>= (fun a : U => (q l) >>= fun b : V => h a b)) = + fun l => (privCompose p q l) >>= (fun z => h z.1 z.2) := by ext l x simp [privCompose, tsum_prod'] - -- apply tsum_congr - -- intro b - -- rw [← ENNReal.tsum_mul_left] - -- apply tsum_congr - -- intro c - -- rw [← mul_assoc] - -- congr 1 - -- rw [tsum_eq_single b] - -- . congr 1 - -- rw [tsum_eq_single c] - -- . simp - -- . intro b' h1 - -- simp - -- intro h2 - -- subst h2 - -- contradiction - -- . intro b' h1 - -- rw [tsum_eq_single c] - -- . simp - -- intro h2 - -- subst h2 - -- contradiction - -- . intro b'' h2 - -- simp - -- intro h3 h4 - -- subst h3 - -- subst h4 - -- contradiction lemma compose_sum_rw (nq1 : U -> ENNReal) (nq2 : V -> ENNReal) (b : U) (c : V) : (∑' (a : U), nq1 a * ∑' (a_1 : V), if b = a ∧ c = a_1 then nq2 a_1 else 0) = nq1 b * nq2 c := by @@ -175,44 +139,6 @@ lemma compose_sum_rw (nq1 : U -> ENNReal) (nq2 : V -> ENNReal) (b : U) (c : V) : rw [C] simp --- Unused -/- -Composed queries are normalizable. --/ --- theorem privCompose_NonTopSum {nq1 : Mechanism T U} {nq2 : Mechanism T V} (nt1 : NonTopSum nq1) (nt2 : NonTopSum nq2) : --- NonTopSum (privCompose nq1 nq2) := by --- simp [NonTopSum] at * --- intro l --- replace nt1 := nt1 l --- replace nt2 := nt2 l --- simp [privCompose] --- 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 - /-- All outputs of a composed query have nonzero probability. -/ @@ -225,6 +151,9 @@ theorem privCompose_NonZeroNQ {nq1 : Mechanism T U} {nq2 : Mechanism T V} (nn1 : simp [privCompose] exists a +/-- +Partition series into fibers. `g` maps an element to its fiber. +-/ theorem ENNReal.HasSum_fiberwise {f : T → ENNReal} {a : ENNReal} (hf : HasSum f a) (g : T → V) : HasSum (fun c : V ↦ ∑' b : g ⁻¹' {c}, f b) a := by let A := (Equiv.sigmaFiberEquiv g) @@ -238,6 +167,9 @@ theorem ENNReal.HasSum_fiberwise {f : T → ENNReal} {a : ENNReal} (hf : HasSum . rfl . apply ENNReal.summable +/-- +Partition series into fibers. `g` maps an element to its fiber. +-/ theorem ENNReal.tsum_fiberwise (p : T → ENNReal) (f : T → V) : ∑' (x : V), ∑' (b : (f ⁻¹' {x})), p b = ∑' i : T, p i := by @@ -246,6 +178,9 @@ theorem ENNReal.tsum_fiberwise (p : T → ENNReal) (f : T → V) : apply Summable.hasSum exact ENNReal.summable +/-- +Rewrite a series into a sum over fibers. `f` maps an element into its fiber. +-/ theorem fiberwisation (p : T → ENNReal) (f : T → V) : (∑' i : T, p i) = ∑' (x : V), if {a : T | x = f a} = {} then 0 else ∑'(i : {a : T | x = f a}), p i := by @@ -280,6 +215,10 @@ lemma condition_to_subset (f : U → V) (g : U → ENNReal) (x : V) : simp rw [B] + +/-- +Postprocessing by a surjection preserves `NonZeroNQ`. +-/ theorem privPostProcess_NonZeroNQ {nq : Mechanism T U} {f : U → V} (nn : NonZeroNQ nq) (sur : Function.Surjective f) : NonZeroNQ (privPostProcess nq f) := by simp [NonZeroNQ, Function.Surjective, privPostProcess] at * @@ -292,75 +231,4 @@ theorem privPostProcess_NonZeroNQ {nq : Mechanism T U} {f : U → V} (nn : NonZe . rw [h] . apply nn --- theorem privPostProcess_NonTopSum {nq : Mechanism T U} (f : U → V) (nt : NonTopSum nq) : --- NonTopSum (privPostProcess nq f) := by --- simp [NonTopSum, privPostProcess] at * --- intros l --- have nt := nt l --- rw [← ENNReal.tsum_fiberwise _ f] at nt --- conv => --- right --- left --- right --- intro n --- rw [condition_to_subset] --- have A : ∀ x : V, f ⁻¹' {x} = {y | x = f y} := by --- aesop --- conv => --- right --- left --- right --- intro x --- rw [← A] --- trivial - --- lemma privCompose_norm (nq1 : List T → PMF U) (nq2 : List T → PMF V) (HNorm1 : NormalMechanism nq1) (HNorm2 : NormalMechanism nq2) : NormalMechanism (privCompose nq1 nq2) := by --- rw [NormalMechanism] at * --- intro l --- have HNorm1' := HasSum.tsum_eq (HNorm1 l) --- have HNorm2' := HasSum.tsum_eq (HNorm2 l) --- have HR : (∑' (x : U × V), privCompose nq1 nq2 l x = 1) := by --- rw [privCompose] --- rw [ENNReal.tsum_prod'] --- simp --- conv => --- lhs --- arg 1 --- intro a --- arg 1 --- intro b --- rw [compose_sum_rw nq1 (fun l a_1 => nq2 l a_1) a b l] --- conv => --- lhs --- arg 1 --- intro a --- rw [ENNReal.tsum_mul_left] --- rw [ENNReal.tsum_mul_right] --- rw [HNorm1'] --- rw [HNorm2'] --- simp --- rw [<- HR] --- apply Summable.hasSum --- exact ENNReal.summable --- --- lemma privPostProcess_norm (nq : List T → SLang U) (HNorm : NormalMechanism nq) (f : U -> V) : NormalMechanism (privPostProcess nq f) := by --- rw [NormalMechanism] at * --- intro l --- have HR : (∑' (b : V), privPostProcess nq f l b = 1) := by --- rw [privPostProcess] --- simp --- rw [<- HasSum.tsum_eq (ENNReal.HasSum_fiberwise (HNorm l) f)] --- apply tsum_congr --- intro v --- rw [condition_to_subset] --- have Htyeq : ({a | v = f a}) = (f ⁻¹' {v}) := by --- rw [Set.ext_iff] --- intro x --- simp --- exact eq_comm --- rw [Htyeq] --- rw [<- HR] --- apply Summable.hasSum --- exact ENNReal.summable --- --- end SLang +end SLang diff --git a/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean b/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean index 7fa55eaa..5cd28b1b 100644 --- a/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean @@ -44,7 +44,7 @@ lemma budget_split (ε₁ ε₂ : ℕ+) : DP bound for noised mean. -/ theorem privNoisedBoundedMean_DP (U : ℕ+) (ε₁ ε₂ : ℕ+) - (H : dps.postprocess_prop_f (fun a : ℤ × ℤ => (a.1 : ℚ) / (a.2 : ℚ) )) : + (H : dps.postprocess_prop_f (fun a : ℤ × ℤ => (a.1 : ℚ) / (a.2 : ℚ))) : dps.prop (privNoisedBoundedMean U ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by unfold privNoisedBoundedMean rw [bind_bind_indep] From 8ab5ad077cf26018d05577178b3ef973ae194fa6 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Thu, 27 Jun 2024 14:58:56 -0400 Subject: [PATCH 112/122] cleanup DifferentialPrivacy/Neighbours --- SampCert/DifferentialPrivacy/Neighbours.lean | 9 +++++++-- 1 file changed, 7 insertions(+), 2 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Neighbours.lean b/SampCert/DifferentialPrivacy/Neighbours.lean index 7c4c57b0..e9bb31fc 100644 --- a/SampCert/DifferentialPrivacy/Neighbours.lean +++ b/SampCert/DifferentialPrivacy/Neighbours.lean @@ -7,14 +7,16 @@ Authors: Jean-Baptiste Tristan /-! # Neighbours -This file defines the notion of neighbouring lists, in order to define a notion of sensitivity. +This file defines neighbouring lists. -/ variable {T : Type} /-- -The notion of "difference" between lists for differential privacy. +Lists which differ by the inclusion or modification of an element. + +This is SampCert's private property. -/ inductive Neighbour (l₁ l₂ : List T) : Prop where | Addition : l₁ = a ++ b → l₂ = a ++ [n] ++ b → Neighbour l₁ l₂ @@ -22,6 +24,9 @@ inductive Neighbour (l₁ l₂ : List T) : Prop where | Update : l₁ = a ++ [n] ++ b → l₂ = a ++ [m] ++ b -> Neighbour l₁ l₂ +/-- +Neighbour relation is symmetric. +-/ def Neighbour_symm (l₁ l₂ : List T) (H : Neighbour l₁ l₂) : Neighbour l₂ l₁ := by cases H · rename_i _ _ _ Hl1 Hl2 From 19360fd78386a1d807634a56b4bc5cc198c7182d Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Thu, 27 Jun 2024 15:03:31 -0400 Subject: [PATCH 113/122] cleanup DifferentialPrivacy/Pure/Composition --- .../DifferentialPrivacy/Pure/Composition.lean | 19 ++++++++++++++++--- SampCert/DifferentialPrivacy/Pure/System.lean | 2 +- 2 files changed, 17 insertions(+), 4 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Pure/Composition.lean b/SampCert/DifferentialPrivacy/Pure/Composition.lean index 8bcfd820..3a2cc4bd 100644 --- a/SampCert/DifferentialPrivacy/Pure/Composition.lean +++ b/SampCert/DifferentialPrivacy/Pure/Composition.lean @@ -8,13 +8,22 @@ import SampCert.DifferentialPrivacy.Pure.DP import Mathlib.Data.Set.Defs import Mathlib.Data.Set.Prod +/- +# Pure Composition in Pure Differential Privacy + +This file proves a pure DP privacy bound on composed independent queries. +-/ + noncomputable section open Classical Set namespace SLang -theorem PureDP_Compose' {nq1 : Mechanism T U} {nq2 : Mechanism T V} {ε₁ ε₂ ε₃ ε₄ : ℕ+} (h1 : PureDP nq1 ((ε₁ : ℝ) / ε₂)) (h2 : PureDP nq2 ((ε₃ : ℝ) / ε₄)) : +/-- +Pure DP privacy bound for ``privCompose``. +-/ +theorem privCompose_DP_bound {nq1 : Mechanism T U} {nq2 : Mechanism T V} {ε₁ ε₂ ε₃ ε₄ : ℕ+} (h1 : PureDP nq1 ((ε₁ : ℝ) / ε₂)) (h2 : PureDP nq2 ((ε₃ : ℝ) / ε₄)) : DP (privCompose nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by simp [PureDP] at * rcases h1 with ⟨h1a, _⟩ @@ -57,7 +66,11 @@ theorem PureDP_Compose' {nq1 : Mechanism T U} {nq2 : Mechanism T V} {ε₁ ε₂ . aesop . aesop -theorem PureDP_Compose (nq1 : Mechanism T U) (nq2 : Mechanism T V) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h : PureDP nq1 ((ε₁ : ℝ) / ε₂)) (h' : PureDP nq2 ((ε₃ : ℝ) / ε₄)) : + +/-- +Pure DP satisfies pure differential privacy. +-/ +theorem privCompose_DP (nq1 : Mechanism T U) (nq2 : Mechanism T V) (ε₁ ε₂ ε₃ ε₄ : ℕ+) (h : PureDP nq1 ((ε₁ : ℝ) / ε₂)) (h' : PureDP nq2 ((ε₃ : ℝ) / ε₄)) : PureDP (privCompose nq1 nq2) (((ε₁ : ℝ) / ε₂) + ((ε₃ : ℝ) / ε₄)) := by simp [PureDP] at * have hc := h @@ -65,7 +78,7 @@ theorem PureDP_Compose (nq1 : Mechanism T U) (nq2 : Mechanism T V) (ε₁ ε₂ rcases h with ⟨ _ , h2 ⟩ rcases h' with ⟨ _ , h'2 ⟩ constructor - . apply PureDP_Compose' hc h'c + . apply privCompose_DP_bound hc h'c . apply privCompose_NonZeroNQ h2 h'2 end SLang diff --git a/SampCert/DifferentialPrivacy/Pure/System.lean b/SampCert/DifferentialPrivacy/Pure/System.lean index 80dd0f25..da08ae75 100644 --- a/SampCert/DifferentialPrivacy/Pure/System.lean +++ b/SampCert/DifferentialPrivacy/Pure/System.lean @@ -17,7 +17,7 @@ noncomputable instance PureDPSystem : DPSystem T where prop := PureDP noise := privNoisedQueryPure noise_prop := NoisedQuery_PureDP - compose_prop := PureDP_Compose + compose_prop := privCompose_DP postprocess_prop_f := Function.Surjective postprocess_prop := PureDP_PostProcess From 17072f9e5b819c683a4d8f681f3af80c0c7345ce Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Thu, 27 Jun 2024 15:06:21 -0400 Subject: [PATCH 114/122] cleanup DifferentialPrivacy/Pure/Mechanism/Code --- SampCert/DifferentialPrivacy/Pure/Composition.lean | 2 +- SampCert/DifferentialPrivacy/Pure/Mechanism/Code.lean | 7 +++++++ 2 files changed, 8 insertions(+), 1 deletion(-) diff --git a/SampCert/DifferentialPrivacy/Pure/Composition.lean b/SampCert/DifferentialPrivacy/Pure/Composition.lean index 3a2cc4bd..f9010122 100644 --- a/SampCert/DifferentialPrivacy/Pure/Composition.lean +++ b/SampCert/DifferentialPrivacy/Pure/Composition.lean @@ -8,7 +8,7 @@ import SampCert.DifferentialPrivacy.Pure.DP import Mathlib.Data.Set.Defs import Mathlib.Data.Set.Prod -/- +/-! # Pure Composition in Pure Differential Privacy This file proves a pure DP privacy bound on composed independent queries. diff --git a/SampCert/DifferentialPrivacy/Pure/Mechanism/Code.lean b/SampCert/DifferentialPrivacy/Pure/Mechanism/Code.lean index dbe43e3a..9656934f 100644 --- a/SampCert/DifferentialPrivacy/Pure/Mechanism/Code.lean +++ b/SampCert/DifferentialPrivacy/Pure/Mechanism/Code.lean @@ -5,10 +5,17 @@ Authors: Jean-Baptiste Tristan -/ import SampCert.Samplers.LaplaceGen.Properties +/-! +# Noise Mechanism for Pure Differential Privacy +-/ + noncomputable section namespace SLang +/-- +Add noise to a a query from the discrete Laplace distribution, to obtain (Δε₁/ε₂)-DP. +-/ def privNoisedQueryPure (query : List T → ℤ) (Δ : ℕ+) (ε₁ ε₂ : ℕ+) (l : List T) : PMF ℤ := do DiscreteLaplaceGenSamplePMF (Δ * ε₂) ε₁ (query l) From 94509f0e1f96b95ea8ce044d0341b8827aef9f06 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Thu, 27 Jun 2024 15:12:44 -0400 Subject: [PATCH 115/122] cleanup DifferentialPrivacy/Pure/Mechanism --- .../Pure/Mechanism/Code.lean | 4 +-- .../Pure/Mechanism/Properties.lean | 27 +++++++++++++++---- SampCert/DifferentialPrivacy/Pure/System.lean | 2 +- 3 files changed, 25 insertions(+), 8 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Pure/Mechanism/Code.lean b/SampCert/DifferentialPrivacy/Pure/Mechanism/Code.lean index 9656934f..6d82adf7 100644 --- a/SampCert/DifferentialPrivacy/Pure/Mechanism/Code.lean +++ b/SampCert/DifferentialPrivacy/Pure/Mechanism/Code.lean @@ -6,7 +6,7 @@ Authors: Jean-Baptiste Tristan import SampCert.Samplers.LaplaceGen.Properties /-! -# Noise Mechanism for Pure Differential Privacy +# Implementation of ``privNoisedQueryPure`` -/ noncomputable section @@ -14,7 +14,7 @@ noncomputable section namespace SLang /-- -Add noise to a a query from the discrete Laplace distribution, to obtain (Δε₁/ε₂)-DP. +Add noise to a a query from the discrete Laplace distribution in order to obtain a (ε₁/ε₂)-DP mechanism from a Δ-sensitive query. -/ def privNoisedQueryPure (query : List T → ℤ) (Δ : ℕ+) (ε₁ ε₂ : ℕ+) (l : List T) : PMF ℤ := do DiscreteLaplaceGenSamplePMF (Δ * ε₂) ε₁ (query l) diff --git a/SampCert/DifferentialPrivacy/Pure/Mechanism/Properties.lean b/SampCert/DifferentialPrivacy/Pure/Mechanism/Properties.lean index 6924fc7d..7d1e3214 100644 --- a/SampCert/DifferentialPrivacy/Pure/Mechanism/Properties.lean +++ b/SampCert/DifferentialPrivacy/Pure/Mechanism/Properties.lean @@ -7,12 +7,21 @@ import SampCert.DifferentialPrivacy.Pure.DP import SampCert.DifferentialPrivacy.Pure.Mechanism.Code import SampCert.Samplers.LaplaceGen.Basic +/-! +# Properties of ``privNoisedQueryPure`` + +This file proves pure differential privacy for ``privNoisedQueryPure``. +-/ + noncomputable section open Classical Nat Int Real ENNReal MeasureTheory Measure namespace SLang +/-- +Noised query has nonzero probability over entire domain. +-/ theorem NoisedQuery_NonZeroNQPureDP (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : NonZeroNQ (privNoisedQueryPure query Δ ε₁ ε₂) := by simp [NonZeroNQ, privNoisedQueryPure] @@ -29,12 +38,12 @@ theorem NoisedQuery_NonZeroNQPureDP (query : List T → ℤ) (Δ ε₁ ε₂ : apply add_pos A Real.zero_lt_one . apply exp_pos -theorem natAbs_to_abs (a b : ℤ) : +lemma natAbs_to_abs (a b : ℤ) : (a - b).natAbs = |(a : ℝ) - (b : ℝ)| := by rw [Int.cast_natAbs] simp only [cast_abs, Int.cast_sub] -theorem normalizing_constant_nonzero (ε₁ ε₂ Δ : ℕ+) : +lemma normalizing_constant_nonzero (ε₁ ε₂ Δ : ℕ+) : (rexp (ε₁ / (Δ * ε₂)) - 1) / (rexp (ε₁ / (Δ * ε₂)) + 1) ≠ 0 := by field_simp intro h @@ -49,7 +58,11 @@ theorem normalizing_constant_nonzero (ε₁ ε₂ Δ : ℕ+) : rw [h] at C contradiction -theorem NoisedQuery_PureDP' (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_sensitivity : sensitivity query Δ) : + +/-- +Differential privacy bound for a ``privNoisedQueryPure`` +-/ +theorem privNoisedQueryPure_DP_bound (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_sensitivity : sensitivity query Δ) : DP (privNoisedQueryPure query Δ ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by rw [event_eq_singleton] at * simp [DP_singleton] at * @@ -121,11 +134,15 @@ theorem NoisedQuery_PureDP' (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bo exact (add_lt_add_iff_right 1).mpr A . apply exp_pos -theorem NoisedQuery_PureDP (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_sensitivity : sensitivity query Δ) : + +/-- +Laplace noising mechanism ``privNoisedQueryPure`` produces a pure ``ε₁/ε₂``-DP mechanism from a Δ-sensitive query. +-/ +theorem privNoisedQueryPure_DP (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_sensitivity : sensitivity query Δ) : PureDP (privNoisedQueryPure query Δ ε₁ ε₂) ((ε₁ : ℝ) / ε₂) := by simp [PureDP] constructor - . apply NoisedQuery_PureDP' + . apply privNoisedQueryPure_DP_bound apply bounded_sensitivity . apply NoisedQuery_NonZeroNQPureDP diff --git a/SampCert/DifferentialPrivacy/Pure/System.lean b/SampCert/DifferentialPrivacy/Pure/System.lean index da08ae75..bf8e1d97 100644 --- a/SampCert/DifferentialPrivacy/Pure/System.lean +++ b/SampCert/DifferentialPrivacy/Pure/System.lean @@ -16,7 +16,7 @@ variable { T : Type } noncomputable instance PureDPSystem : DPSystem T where prop := PureDP noise := privNoisedQueryPure - noise_prop := NoisedQuery_PureDP + noise_prop := privNoisedQueryPure_DP compose_prop := privCompose_DP postprocess_prop_f := Function.Surjective postprocess_prop := PureDP_PostProcess From 65f9f00fb4e24dcb10ace48cefb8b3340dc202ae Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Thu, 27 Jun 2024 15:18:59 -0400 Subject: [PATCH 116/122] cleanup DifferentialPrivacy/Pure --- .../DifferentialPrivacy/Pure/Postprocessing.lean | 13 +++++++++++-- SampCert/DifferentialPrivacy/Pure/System.lean | 7 +++++++ 2 files changed, 18 insertions(+), 2 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Pure/Postprocessing.lean b/SampCert/DifferentialPrivacy/Pure/Postprocessing.lean index 10146a97..c71230a6 100644 --- a/SampCert/DifferentialPrivacy/Pure/Postprocessing.lean +++ b/SampCert/DifferentialPrivacy/Pure/Postprocessing.lean @@ -8,13 +8,19 @@ import SampCert.DifferentialPrivacy.Pure.DP import Mathlib.Data.Set.Defs import Mathlib.Data.Set.Prod +/-! +# ``privPostProcess`` pure DP + +This file establishes a pure DP bound for ``privPostProcess`` +-/ + noncomputable section open Classical Set namespace SLang -theorem PureDP_PostProcess' {nq : Mechanism T U} {ε₁ ε₂ : ℕ+} (h : PureDP nq ((ε₁ : ℝ) / ε₂)) (f : U → V) : +lemma privPostProcess_DP_bound {nq : Mechanism T U} {ε₁ ε₂ : ℕ+} (h : PureDP nq ((ε₁ : ℝ) / ε₂)) (f : U → V) : DP (privPostProcess nq f) (((ε₁ : ℝ) / ε₂)) := by simp [PureDP] at * rcases h with ⟨ha, _⟩ @@ -37,13 +43,16 @@ theorem PureDP_PostProcess' {nq : Mechanism T U} {ε₁ ε₂ : ℕ+} (h : PureD exact Real.exp_pos ((ε₁: ℝ) / ε₂) . simp +/-- +``privPostProcess`` satisfies pure DP, for any surjective postprocessing function. +-/ theorem PureDP_PostProcess {f : U → V} (sur : Function.Surjective f) (nq : Mechanism T U) (ε₁ ε₂ : ℕ+) (h : PureDP nq ((ε₁ : ℝ) / ε₂)) : PureDP (privPostProcess nq f) (((ε₁ : ℝ) / ε₂)) := by simp [PureDP] at * have hc := h rcases h with ⟨ _ , h2 ⟩ constructor - . apply PureDP_PostProcess' hc + . apply privPostProcess_DP_bound hc . apply privPostProcess_NonZeroNQ h2 sur end SLang diff --git a/SampCert/DifferentialPrivacy/Pure/System.lean b/SampCert/DifferentialPrivacy/Pure/System.lean index bf8e1d97..8daa99ce 100644 --- a/SampCert/DifferentialPrivacy/Pure/System.lean +++ b/SampCert/DifferentialPrivacy/Pure/System.lean @@ -9,10 +9,17 @@ import SampCert.DifferentialPrivacy.Pure.Mechanism.Basic import SampCert.DifferentialPrivacy.Pure.Composition import SampCert.DifferentialPrivacy.Pure.Postprocessing +/-! +# Pure DP system +-/ + namespace SLang variable { T : Type } +/-- +Pure ε-DP with noise drawn from the discrete Laplace distribution. +-/ noncomputable instance PureDPSystem : DPSystem T where prop := PureDP noise := privNoisedQueryPure From ed1857470fb9aea8debd5e56e31cdf8a294bd69c Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Thu, 27 Jun 2024 15:23:48 -0400 Subject: [PATCH 117/122] cleanup DifferentialPrivacy/Queries --- SampCert/DifferentialPrivacy/Queries/BoundedMean/Code.lean | 4 ++-- SampCert/DifferentialPrivacy/Queries/BoundedSum/Code.lean | 2 +- SampCert/DifferentialPrivacy/Queries/Count/Code.lean | 4 ++-- 3 files changed, 5 insertions(+), 5 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/BoundedMean/Code.lean b/SampCert/DifferentialPrivacy/Queries/BoundedMean/Code.lean index 82699dc0..2cf226d8 100644 --- a/SampCert/DifferentialPrivacy/Queries/BoundedMean/Code.lean +++ b/SampCert/DifferentialPrivacy/Queries/BoundedMean/Code.lean @@ -10,7 +10,7 @@ import SampCert.DifferentialPrivacy.Queries.BoundedSum.Code /-! # ``privNoisedBoundedMean`` Implementation -This file defines a differentially private noising of a bounded mean query. +This file defines a differentially private bounded mean query. -/ noncomputable section @@ -20,7 +20,7 @@ namespace SLang variable [dps : DPSystem ℕ] /-- -Compute a noised mean using a noised sum and noised count +Compute a noised mean using a noised sum and noised count. -/ def privNoisedBoundedMean (U : ℕ+) (ε₁ ε₂ : ℕ+) (l : List ℕ) : PMF ℚ := do let S ← privNoisedBoundedSum U ε₁ (2 * ε₂) l diff --git a/SampCert/DifferentialPrivacy/Queries/BoundedSum/Code.lean b/SampCert/DifferentialPrivacy/Queries/BoundedSum/Code.lean index c6875dcc..12a37d7a 100644 --- a/SampCert/DifferentialPrivacy/Queries/BoundedSum/Code.lean +++ b/SampCert/DifferentialPrivacy/Queries/BoundedSum/Code.lean @@ -11,7 +11,7 @@ import Init.Data.Int.Order /-! # ``privNoisedBoundedSum`` Implementation -This file defines a differentially private noising of a bounded sum query. +This file defines a differentially private bounded sum query. -/ noncomputable section diff --git a/SampCert/DifferentialPrivacy/Queries/Count/Code.lean b/SampCert/DifferentialPrivacy/Queries/Count/Code.lean index 6e44aa23..a94db31e 100644 --- a/SampCert/DifferentialPrivacy/Queries/Count/Code.lean +++ b/SampCert/DifferentialPrivacy/Queries/Count/Code.lean @@ -19,12 +19,12 @@ variable {T : Type} variable [dps : DPSystem T] /-- -Query to count the size of the input +Query to count the size of the input list. -/ def exactCount (l : List T) : ℤ := List.length l /-- -Noised counting mechanism from the DP system +A noised counting mechanism. -/ def privNoisedCount (ε₁ ε₂ : ℕ+) (l : List T) : PMF ℤ := do dps.noise exactCount 1 ε₁ ε₂ l From edc4422612c3d85bd841e01f0464c40b3653103b Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Thu, 27 Jun 2024 16:26:08 -0400 Subject: [PATCH 118/122] Some cleanup in DifferentialPrivacy/RenyiDivergence --- .../DifferentialPrivacy/RenyiDivergence.lean | 194 +++++++----------- 1 file changed, 79 insertions(+), 115 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index 4afcb1d1..332d1a97 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -13,20 +13,29 @@ import Mathlib.Analysis.Convex.Integral /-! Renyi Divergence -This file defines the Renyi divergence and equations for evaluating its expectation. +This file defines the Renyi divergence, and equations for evaluating it. -/ open Real ENNReal PMF Nat Int MeasureTheory Measure PMF open Classical + /-- -Simplified consequence of absolute continuity (remove me?) +Simplified consequence of absolute continuity between PMF's. -/ def AbsCts (p q : T -> ENNReal) : Prop := ∀ x : T, q x = 0 -> p x = 0 + /-- -Specialize the definitation of AbsolutelyContinuous when singletons are measurable +All PMF's are absolutely continuous with respect to themselves. +-/ +lemma AbsCts_refl (q : PMF T) : AbsCts q q := by + rw [AbsCts] + simp + +/-- +Obtain simplified absolute continuity from the measure-theoretic version of absolute continuity in a discrete space. -/ lemma PMF_AbsCts [MeasurableSpace T] [MeasurableSingletonClass T] (p q : PMF T) (H : AbsolutelyContinuous (PMF.toMeasure p) (PMF.toMeasure q)) : AbsCts p q := by rw [AbsolutelyContinuous] at H @@ -63,7 +72,7 @@ noncomputable def RenyiDivergence_def (p q : PMF T) (α : ℝ) : EReal := (α - 1)⁻¹ * (elog (∑' x : T, (p x)^α * (q x)^(1 - α))) /-- -Equation for the Renyi divergence series in terms of the Renyi Divergence +Rearrange the definition of ``RenyiDivergence_def`` to obtain an equation for the inner series. -/ lemma RenyiDivergence_def_exp (p q : PMF T) {α : ℝ} (h : 1 < α) : eexp (((α - 1)) * RenyiDivergence_def p q α) = (∑' x : T, (p x)^α * (q x)^(1 - α)) := by @@ -82,7 +91,7 @@ lemma RenyiDivergence_def_exp (p q : PMF T) {α : ℝ} (h : 1 < α) : /-- -Closed form of the series in the definition of the Renyi divergence. +Renyi Divergence series written as a conditional expectation. -/ theorem RenyiDivergenceExpectation (p q : T → ENNReal) {α : ℝ} (h : 1 < α) (H : AbsCts p q) : (∑' x : T, (p x)^α * (q x)^(1 - α)) = ∑' x: T, (p x / q x)^α * (q x) := by @@ -146,7 +155,7 @@ theorem RenyiDivergenceExpectation (p q : T → ENNReal) {α : ℝ} (h : 1 < α) /-! -## Jensen's inequality +## Jensen's inequalitya -/ section Jensen @@ -185,22 +194,19 @@ lemma Integrable_rpow (f : T → ℝ) (nn : ∀ x : T, 0 ≤ f x) (μ : Measure . rw [← hasFiniteIntegral_norm_iff] simp [X] +-- MARKUSDE: This lemma is derivable from ``Renyi_Jensen_strict_real``, however it requires a reduction +-- to first eliminate all elements (t : T) where q t = 0 from the series. /-- -Jensen's ineuquality for the exponential applied to Renyi's sum +Jensen's inequality for the exponential applied to the real-valued function ``(⬝)^α``. -/ theorem Renyi_Jensen_real (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 + enter [1, 1, 1, x] rw [mul_comm] rw [← smul_eq_mul] conv => - right - right - intro x + enter [2, 1, x] rw [mul_comm] rw [← smul_eq_mul] rw [← PMF.integral_eq_tsum] @@ -261,23 +267,16 @@ theorem Renyi_Jensen_real (f : T → ℝ) (q : PMF T) (α : ℝ) (h : 1 < α) (h /-- -Jensen's strict ineuquality for the exponential applied to Renyi's sum +Strict version of Jensen't inequality applied to the function ``(⬝)^α``. -/ theorem Renyi_Jensen_strict_real (f : T → ℝ) (q : PMF T) (α : ℝ) (h : 1 < α) (h2 : ∀ x : T, 0 ≤ f x) (mem : Memℒp f (ENNReal.ofReal α) (PMF.toMeasure q)) (HT_nz : ∀ t : T, q t ≠ 0): ((∑' x : T, (f x) * (q x).toReal)) ^ α < (∑' x : T, (f x) ^ α * (q x).toReal) ∨ (∀ x : T, f x = ∑' (x : T), (q x).toReal * f x) := by conv => - left - left - left - right - intro x + enter [1, 1, 1, 1, x] rw [mul_comm] rw [← smul_eq_mul] conv => - left - right - right - intro x + enter [1, 2, 1, x] rw [mul_comm] rw [← smul_eq_mul] rw [← PMF.integral_eq_tsum] @@ -396,15 +395,17 @@ theorem Renyi_Jensen_strict_real (f : T → ℝ) (q : PMF T) (α : ℝ) (h : 1 < rw [one_le_ofReal] apply le_of_lt h - - - end Jensen --- MARKUSDE: move -noncomputable def Renyi_Jensen_f (p q : PMF T) : T -> ℝ := (fun z => ((p z / q z)).toReal) --- Except for one case, we can rewrite the ENNReal-valued inequality into the form Jenen's inequality expects. +/-- +Quotient from the Real-valued Jenen's inequality applied to the series in the Renyi divergence. +-/ +noncomputable def Renyi_Jensen_f (p q : PMF T) : T -> ℝ := (fun z => (p z / q z).toReal) + +/-- +Summand from the Renyi divergence equals a real-valued summand, except in a special case. +-/ lemma Renyi_Jensen_rw (p q : PMF T) {α : ℝ} (h : 1 < α) (H : AbsCts p q) (Hspecial : ∀ x : T, ¬(p x = ⊤ ∧ q x ≠ 0 ∧ q x ≠ ⊤)) (x : T) : (p x / q x)^α * (q x) = ENNReal.ofReal (((Renyi_Jensen_f p q) x)^α * (q x).toReal) := by simp [Renyi_Jensen_f] @@ -433,9 +434,10 @@ lemma Renyi_Jensen_rw (p q : PMF T) {α : ℝ} (h : 1 < α) (H : AbsCts p q) (Hs · exact apply_ne_top q x --- FIXME: might be able to simplify this argument with the new rewrite lemmas +-- MARKUSDE: I think it might be possible to use `Renyi_Jensen_strict_real` in this proof instead, +-- this would eliminate the need for `Renyi_Jensen_real`. /-- -Jensen's inquality applied to ENNReals, in the case that q is nonzero +Jensen's inquality applied to ENNReals, in the case that q is nonzero. -/ lemma Renyi_Jensen_ENNReal_reduct [MeasurableSpace T] [MeasurableSingletonClass T] [Countable T] (p q : PMF T) {α : ℝ} (h : 1 < α) (H : AbsCts p q) (Hq : ∀ t, q t ≠ 0) : @@ -525,21 +527,14 @@ lemma Renyi_Jensen_ENNReal_reduct [MeasurableSpace T] [MeasurableSingletonClass · simp apply le_of_not_ge Hα · conv => - lhs - arg 1 - intro a - arg 1 - arg 1 + enter [1, 1, a, 1, 1] rw [<- Real.toNNReal_eq_nnnorm_of_nonneg (HRJf_nonneg a)] rw [Renyi_Jensen_f] rw [<- ENNReal.ofReal.eq_1] rw [ENNReal.ofReal_toReal (HRJf_nt a)] rfl conv => - lhs - arg 1 - intro a - rhs + enter [1, 1, a, 2] simp [toMeasure] simp [PMF.toOuterMeasure] rw [Hsum_indicator] @@ -565,16 +560,11 @@ lemma Renyi_Jensen_ENNReal_reduct [MeasurableSpace T] [MeasurableSingletonClass rw [<- ENNReal.toReal_mul] rw [<- ENNReal.tsum_toReal_eq] · rw [ENNReal.ofReal_toReal] - -- Could do another case at the top if not derivable - -- Want to bound above my ∑' conv => - arg 1 - arg 1 - intro a + enter [1, 1, a] rw [PMF_mul_mul_inv_eq_mul_cancel p q H] exact tsum_coe_ne_top p - · -- Bound above by p a - intro a + · intro a conv => arg 1 rw [PMF_mul_mul_inv_eq_mul_cancel p q H] @@ -617,10 +607,10 @@ lemma Renyi_Jensen_ENNReal_reduct [MeasurableSpace T] [MeasurableSingletonClass exact OrderTop.le_top ((∑' (x : T), p x / q x * q x) ^ α) - --- Lift the properties of Renyi_Jensen_strict_real to ENNReal --- q is still nonnegative -lemma Renyi_Jensen_ENNReal_converse_reduct [MeasurableSpace T] [MeasurableSingletonClass T] [Countable T] +/-- +On types where q is nonzero, Jensen's inequality for the Renyi divergence sum is tight only for equal distributions. +-/ +lemma Renyi_Jensen_ENNReal_converse_reduct [MeasurableSpace T] [MeasurableSingletonClass T] [Countable T] (p q : PMF T) {α : ℝ} (h : 1 < α) (H : AbsCts p q) (Hq : ∀ t, q t ≠ 0) (Hsumeq : (∑' x : T, (p x / q x) * q x) ^ α = (∑' x : T, (p x / q x) ^ α * q x)) : (p = q) := by @@ -648,13 +638,10 @@ lemma Renyi_Jensen_ENNReal_converse_reduct [MeasurableSpace T] [MeasurableSingl intro x rw [Renyi_Jensen_f] conv => - arg 1 - intro x - lhs + enter [1, x, 1] rw [ENNReal.toReal_rpow] conv => - arg 1 - intro x + enter [1, x] rw [<- ENNReal.toReal_mul] apply ENNReal.summable_toReal assumption @@ -707,21 +694,14 @@ lemma Renyi_Jensen_ENNReal_converse_reduct [MeasurableSpace T] [MeasurableSingl · simp apply le_of_not_ge Hα · conv => - lhs - arg 1 - intro a - arg 1 - arg 1 + enter [1, 1, a, 1, 1] rw [<- Real.toNNReal_eq_nnnorm_of_nonneg (HRJf_nonneg a)] rw [Renyi_Jensen_f] rw [<- ENNReal.ofReal.eq_1] rw [ENNReal.ofReal_toReal (HRJf_nt a)] rfl conv => - lhs - arg 1 - intro a - rhs + enter [1, 1, a, 2] simp [toMeasure] simp [PMF.toOuterMeasure] rw [Hsum_indicator] @@ -735,10 +715,7 @@ lemma Renyi_Jensen_ENNReal_converse_reduct [MeasurableSpace T] [MeasurableSingl intro apply PMF.apply_ne_top conv at Hk => - lhs - arg 1 - arg 1 - intro z + enter [1, 1, 1, z] rw [Renyi_Jensen_f] rw [<- ENNReal.toReal_mul] arg 1 @@ -749,10 +726,7 @@ lemma Renyi_Jensen_ENNReal_converse_reduct [MeasurableSpace T] [MeasurableSingl -- Convert the LHS of Hsumeq to the ℝ-valued summand, and then contradict conv at Hsumeq => - lhs - arg 1 - arg 1 - intro x + enter [1, 1, 1, x] rw [division_def] rw [mul_assoc] rw [ENNReal.inv_mul_cancel] @@ -857,15 +831,13 @@ lemma Renyi_Jensen_ENNReal_converse_reduct [MeasurableSpace T] [MeasurableSingl simp at Hsumeq /-- -Restriction of the PMF f to the support of q +Restriction of the PMF f to the support of q. -/ def reducedPMF_def (f q : PMF T) (x : { t : T // ¬q t = 0 }) : ENNReal := f x.val --- Should be provable -lemma AbsCts_refl (q : PMF T) : AbsCts q q := by - rw [AbsCts] - simp - +/-- +Restricted PMF has sum 1 +-/ lemma reducedPMF_norm_acts (p q : PMF T) (H : AbsCts p q) : HasSum (reducedPMF_def p q) 1 := by have H1 : Summable (reducedPMF_def p q) := by exact ENNReal.summable have H2 := Summable.hasSum H1 @@ -888,9 +860,15 @@ lemma reducedPMF_norm_acts (p q : PMF T) (H : AbsCts p q) : HasSum (reducedPMF_d rw [<- H3] apply H2 +/-- +Restriction of the PMF f to the support of q +-/ noncomputable def reducedPMF {p q : PMF T} (H : AbsCts p q): PMF { t : T // ¬q t = 0 } := ⟨ reducedPMF_def p q, reducedPMF_norm_acts p q H ⟩ +/-- +`reducedPMF` is nonzero everywhere +-/ lemma reducedPMF_pos {q : PMF T} (H : AbsCts q q) (a : T) (Ha : ¬q a = 0): (reducedPMF H) ⟨a, Ha⟩ ≠ 0 := by simp rw [reducedPMF] @@ -901,7 +879,7 @@ lemma reducedPMF_pos {q : PMF T} (H : AbsCts q q) (a : T) (Ha : ¬q a = 0): (red apply Ha /-- -Jensen's inquality applied to ENNReals +Jensen's inquality for the Renyi divergence sum between absolutely continuous PMFs -/ theorem Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] [Countable T] (p q : PMF T) {α : ℝ} (h : 1 < α) (Hac : AbsCts p q) : (∑' x : T, (p x / q x) * q x) ^ α ≤ (∑' x : T, (p x / q x) ^ α * q x) := by @@ -918,9 +896,7 @@ theorem Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] [C have B2 (x : { x // ¬q x = 0 }) : (p ↑x / q ↑x)^α * q ↑x = (reducedPMF Hac x / reducedPMF Hq x)^α * reducedPMF Hq x := by congr conv => congr - · arg 1 - arg 1 - intro x + · enter [1, 1, x] rw [B1 x] · arg 1 intro x @@ -942,6 +918,9 @@ theorem Renyi_Jensen_ENNReal [MeasurableSpace T] [MeasurableSingletonClass T] [C rcases t with ⟨ a , Ha ⟩ apply (reducedPMF_pos Hq a Ha) +/-- +Converse of Jensen's inquality for the Renyi divergence sum between absolutely continuous PMFs +-/ lemma Renyi_Jensen_ENNReal_converse [MeasurableSpace T] [MeasurableSingletonClass T] [Countable T] (p q : PMF T) {α : ℝ} (h : 1 < α) (H : AbsCts p q) (Hsumeq : (∑' x : T, (p x / q x) * q x) ^ α = (∑' x : T, (p x / q x) ^ α * q x)) : @@ -1002,17 +981,12 @@ lemma Renyi_Jensen_ENNReal_converse [MeasurableSpace T] [MeasurableSingletonCla simp at Hreduced' assumption - --- FIXME (or not ?) /-- -The Renyi divergence is monotonic in the value of its sum. --/ -lemma RenyiDivergence_mono_sum (x y : ℝ) (α : ℝ) (h : 1 < α) : (Real.exp ((α - 1) * x)) ≤ (Real.exp ((α - 1) * y)) -> (x ≤ y) := by - intro H - apply _root_.le_of_mul_le_mul_left - · exact exp_le_exp.mp H - · linarith +The ``EReal``-valued Renyi divergence is nonnegative. +Use this lemma to perform rewrites through the ``ENNReal.ofEReal`` in the +``ENNReal``-valued ``RenyiDivergence`` +-/ theorem RenyiDivergence_def_nonneg [MeasurableSpace T] [MeasurableSingletonClass T] [Countable T] (p q : PMF T) (Hpq : AbsCts p q) {α : ℝ} (Hα : 1 < α) : (0 ≤ RenyiDivergence_def p q α) := by have H1 : eexp (((α - 1)) * 0) ≤ eexp ((α - 1) * RenyiDivergence_def p q α) := by @@ -1048,7 +1022,9 @@ theorem RenyiDivergence_def_nonneg [MeasurableSpace T] [MeasurableSingletonClass exact EReal.coe_lt_top (α - OfNat.ofNat 1) · assumption - +/-- +Renyi divergence between identical distributions is zero +-/ lemma RenyiDivergence_refl_zero (p : PMF T) {α : ℝ} (Hα : 1 < α) : (0 = RenyiDivergence_def p p α) := by have H1 : 1 = eexp ((α - 1) * RenyiDivergence_def p p α) := by rw [RenyiDivergence_def_exp p p Hα] @@ -1093,7 +1069,7 @@ lemma RenyiDivergence_refl_zero (p : PMF T) {α : ℝ} (Hα : 1 < α) : (0 = Ren simp /-- -Renyi divergence attains equality if and only if the distributions are equal +Renyi divergence is zero if and only if the distributions are equal -/ theorem RenyiDivergence_def_eq_0_iff [MeasurableSpace T] [MeasurableSingletonClass T] [Countable T] (p q : PMF T) {α : ℝ} (Hα : 1 < α) (Hcts : AbsCts p q) : @@ -1128,10 +1104,11 @@ theorem RenyiDivergence_def_eq_0_iff [MeasurableSpace T] [MeasurableSingletonCla apply RenyiDivergence_refl_zero apply Hα - +/-- +The logarithm in the definition of the Renyi divergence is nonnegative. +-/ lemma RenyiDivergence_def_log_sum_nonneg [MeasurableSpace T] [MeasurableSingletonClass T] [Countable T] (p q : PMF T) (Hac : AbsCts p q) {α : ℝ} (Hα : 1 < α ): (0 ≤ (elog (∑' x : T, (p x)^α * (q x)^(1 - α)))) := by - -- Follows from RenyiDivergence_def_nonneg have Hrd := RenyiDivergence_def_nonneg p q Hac Hα rw [RenyiDivergence_def] at Hrd apply ofEReal_nonneg_scal_l at Hrd @@ -1140,13 +1117,16 @@ lemma RenyiDivergence_def_log_sum_nonneg [MeasurableSpace T] [MeasurableSingleto linarith /-- -The Renyi divergence. +The ``ENNReal``-valued Renyi divergence between PMF's. -/ noncomputable def RenyiDivergence (p q : PMF T) (α : ℝ) : ENNReal := ENNReal.ofEReal (RenyiDivergence_def p q α) --- Lifted property of RenyiDivergence_def +/-- +The Renyi divergence between absolutely continuous distributions is zero if and only if the +distributions are equal. +-/ theorem RenyiDivergence_aux_zero [MeasurableSpace T] [MeasurableSingletonClass T] [Countable T] (p q : PMF T) {α : ℝ} (Hα : 1 < α) (Hac : AbsCts p q) : p = q <-> RenyiDivergence p q α = 0 := by apply Iff.intro @@ -1163,19 +1143,3 @@ theorem RenyiDivergence_aux_zero [MeasurableSpace T] [MeasurableSingletonClass T refine (ofEReal_nonneg_eq_iff ?mpr.Hw H').mpr ?mpr.a · simp simp [H] - --- Unused -/- -Closed form for the Renyi Divergence. --/ --- 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 8f2fc72221d9d60c9462f170aee17f0bbb6ba802 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Fri, 28 Jun 2024 08:54:23 -0400 Subject: [PATCH 119/122] cleanup DifferentialPrivacy/ZeroConcentrated/Composition --- .../ZeroConcentrated/Composition.lean | 52 +++---------------- .../ZeroConcentrated/DP.lean | 3 ++ 2 files changed, 11 insertions(+), 44 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean index 1246fa95..c74d6647 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Composition.lean @@ -25,7 +25,6 @@ variable [Inhabited V] variable [MeasurableSpace U] [MeasurableSingletonClass U] [Countable U] variable [MeasurableSpace V] [MeasurableSingletonClass V] [Countable V] -set_option pp.coercions false /-- Composed queries satisfy zCDP Renyi divergence bound. -/ @@ -71,51 +70,14 @@ theorem privCompose_zCDPBound {nq1 : Mechanism T U} {nq2 : Mechanism T V} {ε₁ clear CG1 clear CG2 conv => - left - right - right - right - right - intro b - right - intro c + enter [1, 1, 2, 1, 1, b, 1, c] rw [← mul_assoc] conv => - left - right - right - right - right - intro b + enter [1, 1, 2, 1, 1, b] rw [ENNReal.tsum_mul_left] rw [ENNReal.tsum_mul_right] rw [← elog_mul] - -- conv at h1 => - -- lhs - -- arg 1 - -- arg 2 - -- arg 1 - -- arg 1 - -- intro x - -- rw [DFunLike.coe] - -- rw [PMF.instFunLike] - -- rw [SLang.toPMF] - -- rw [SLang.toPMF] - -- simp - -- conv at h2 => - -- lhs - -- arg 1 - -- arg 2 - -- arg 1 - -- arg 1 - -- intro x - -- rw [DFunLike.coe] - -- rw [PMF.instFunLike] - -- rw [SLang.toPMF] - -- rw [SLang.toPMF] - -- simp - have log_nonneg_1 : 0 ≤ (∑' (i : U), (nq1 l₁).val i ^ α * (nq1 l₂).val i ^ (1 - α)).elog := by have Hac1 : AbsCts (nq1 l₁) (nq1 l₂) := by exact Ha1 l₁ l₂ h4 apply (RenyiDivergence_def_log_sum_nonneg (nq1 l₁) (nq1 l₂) Hac1 h3) @@ -133,8 +95,7 @@ theorem privCompose_zCDPBound {nq1 : Mechanism T U} {nq2 : Mechanism T V} {ε₁ -- Distribute rw [CanonicallyOrderedCommSemiring.left_distrib] apply (@le_trans _ _ _ (ENNReal.ofReal (2⁻¹ * (↑↑ε₁ ^ 2 / ↑↑ε₂ ^ 2) * α) + ENNReal.ofReal (2⁻¹ * (↑↑ε₃ ^ 2 / ↑↑ε₄ ^ 2) * α))) - · -- apply? - apply _root_.add_le_add + · apply _root_.add_le_add · rw [ofEReal_mul_nonneg] at h1 · apply h1 · apply EReal.coe_nonneg.mpr @@ -156,7 +117,8 @@ theorem privCompose_zCDPBound {nq1 : Mechanism T U} {nq2 : Mechanism T V} {ε₁ · rw [← mul_add] rw [mul_le_mul_iff_of_pos_left] · ring_nf - simp + simp only [inv_pow, add_le_add_iff_right, le_add_iff_nonneg_left, gt_iff_lt, inv_pos, Nat.cast_pos, PNat.pos, + mul_pos_iff_of_pos_right, mul_pos_iff_of_pos_left, mul_nonneg_iff_of_pos_left, Nat.ofNat_nonneg] · simp only [inv_pos, Nat.ofNat_pos] · linarith · apply mul_nonneg @@ -246,7 +208,9 @@ Renyi divergence beteeen composed queries on neighbours are finite. -- rename_i h4 h5 -- contradiction - +/-- +``privCompose`` preserves absolute continuity between neighbours +-/ def privCompose_AC (nq1 : Mechanism T U) (nq2 : Mechanism T V) (Hac1 : ACNeighbour nq1) (Hac2 : ACNeighbour nq2) : ACNeighbour (privCompose nq1 nq2) := by rw [ACNeighbour] at * diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean index 70f3ae86..d9f6a36f 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean @@ -42,6 +42,9 @@ 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 - α) ≠ ⊤ +/-- +All neighbouring queries are absolutely continuous +-/ def ACNeighbour (p : List T -> PMF U) : Prop := ∀ l₁ l₂, Neighbour l₁ l₂ -> AbsCts (p l₁) (p l₂) /-- From ea00129823e56113ef3adb2aaf30287c46c6a0b7 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Fri, 28 Jun 2024 09:09:06 -0400 Subject: [PATCH 120/122] cleanup DifferentialPrivacy/ZeroConcentrated --- .../ZeroConcentrated/ConcentratedBound.lean | 6 +- .../Mechanism/Properties.lean | 93 +---------- .../ZeroConcentrated/Postprocessing.lean | 153 ++++-------------- 3 files changed, 34 insertions(+), 218 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean index 603e99bd..9d360b3c 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/ConcentratedBound.lean @@ -14,8 +14,8 @@ import SampCert.Samplers.GaussianGen.Basic /-! # Concentrated Bound -This file derives a cDP bound for the discrete Gaussian. In particular, it bounds -the Renyi divergence between discrete Gaussian evaluations with integer means. +This file derives a cDP bound on the discrete Gaussian, and for the discrete Gaussian +sampler. -/ open Real Nat @@ -84,7 +84,7 @@ lemma SG_Renyi_simplify {σ : ℝ} (h : σ ≠ 0) (μ ν : ℤ) (α : ℝ) : rfl /-- -Alternative definition for the Renyi Divergence. +Real-valued Renyi Divergence. -/ noncomputable def RenyiDivergence' (p q : T → ℝ) (α : ℝ) : ℝ := (1 / (α - 1)) * Real.log (∑' x : T, (p x)^α * (q x)^(1 - α)) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean index edb8d992..381ec5f1 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean @@ -18,18 +18,6 @@ open Classical Nat Int Real ENNReal MeasureTheory Measure namespace SLang --- lemma privNoisedQuery_norm (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : -- (bounded_sensitivity : sensitivity query Δ) : --- NormalMechanism (privNoisedQuery query Δ ε₁ ε₂) := by --- rw [NormalMechanism] --- intro l --- rw [privNoisedQuery] --- exact DiscreteGaussianGen_sum (Δ * ε₂) ε₁ (query l) - --- set_option pp.coercions false --- set_option pp.notation false --- set_option pp.all true --- set_option pp.universes false - /-- The zCDP mechanism with bounded sensitivity satisfies the bound for ``(Δε₂/ε₁)^2``-zCDP. -/ @@ -42,7 +30,6 @@ theorem privNoisedQuery_zCDPBound (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ clear A -- Pull out the ENNReal.ofReal, reducing it to a Real case - -- Simplify this argument, it is a mess rw [ENNReal.div_eq_inv_mul] have L : (OfNat.ofNat 2 * ((Δ * ε₂).val.cast / ε₁.val.cast) ^ OfNat.ofNat 2)⁻¹ = ENNReal.ofReal ((OfNat.ofNat 2 * ((Δ * ε₂).val.cast / ε₁.val.cast) ^ OfNat.ofNat 2)⁻¹ ) := by simp @@ -136,9 +123,7 @@ theorem privNoisedQuery_zCDPBound (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ rw [mul_comm] rw [← mul_assoc] - -- Almost would be easier just to redo it... from here, the main reduction has already been finished left - -- Move α to left conv => rhs @@ -193,35 +178,6 @@ theorem privNoisedQuery_zCDPBound (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ · apply inv_nonneg.mpr apply sq_nonneg - -/- -All outputs of the zCDP mechanism have finite probability. --/ --- theorem privNoisedQuery_NonTopNQ (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : --- NonTopNQ (privNoisedQuery query Δ ε₁ ε₂) := by --- simp [NonTopNQ, privNoisedQuery, 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) (n.instDecidableEq (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 - lemma discrete_gaussian_shift {σ : ℝ} (h : σ ≠ 0) (μ : ℝ) (τ x : ℤ) : discrete_gaussian σ μ (x - τ) = discrete_gaussian σ (μ + τ) (x) := by simp [discrete_gaussian] @@ -231,51 +187,6 @@ lemma discrete_gaussian_shift {σ : ℝ} (h : σ ≠ 0) (μ : ℝ) (τ x : ℤ) ring_nf . rw [shifted_gauss_sum h] -/- -The zCDP mechanism is normalizable. --/ --- theorem privNoisedQuery_NonTopSum (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) : --- NonTopSum (privNoisedQuery query Δ ε₁ ε₂) := by --- simp [NonTopSum, privNoisedQuery, 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) (n.instDecidableEq (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) - /- The Renyi divergence of the zCDP mechanism is finite on neighbouring inputs. -/ @@ -369,6 +280,10 @@ The Renyi divergence of the zCDP mechanism is finite on neighbouring inputs. -- rw [discrete_gaussian_shift P] -- simp [X] + +/-- +privNoisedQuery preserves absolute continuity between neighbours +-/ def privNoisedQuery_AC (query : List T -> ℤ) (Δ ε₁ ε₂ : ℕ+) : ACNeighbour (privNoisedQuery query Δ ε₁ ε₂) := by rw [ACNeighbour] intro l₁ l₂ _ diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean index 1e9789a3..4ed59926 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean @@ -28,104 +28,9 @@ variable [count : Countable U] variable [disc : DiscreteMeasurableSpace U] variable [Inhabited U] --- lemma 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 ℝ t1 μ _ 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] - --- /- --- Jensen's ineuquality for the exponential applied to Renyi's sum --- -/ --- 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 => --- 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 ℝ 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 --- . 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 t1 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 t1 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 - +/-- +privPostProcess preserves absolute continuity between neighbours +-/ def privPostProcess_AC {f : U -> V} (nq : Mechanism T U) (Hac : ACNeighbour nq) : ACNeighbour (privPostProcess nq f) := by rw [ACNeighbour] at * unfold AbsCts at * @@ -137,6 +42,9 @@ def privPostProcess_AC {f : U -> V} (nq : Mechanism T U) (Hac : ACNeighbour nq) apply Hppz apply fi +/-- +Normalized fiber +-/ def δ (nq : SLang U) (f : U → V) (a : V) : {n : U | a = f n} → ENNReal := fun x : {n : U | a = f n} => nq x * (∑' (x : {n | a = f n}), nq x)⁻¹ @@ -147,6 +55,9 @@ lemma δ_normalizes (nq : SLang U) (f : U → V) (a : V) (h1 : ∑' (i : ↑{n | rw [ENNReal.tsum_mul_right] rw [ENNReal.mul_inv_cancel h1 h2] +/-- +Normalized fiber distribution +-/ def δpmf (nq : SLang U) (f : U → V) (a : V) (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 ⟩ @@ -178,10 +89,6 @@ theorem norm_simplify (x : ENNReal) (h : x ≠ ⊤) : 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 convergent_subset {p : T → ENNReal} (f : T → V) (conv : ∑' (x : T), p x ≠ ⊤) : ∑' (x : { y : T| x = f y }), p x ≠ ⊤ := by @@ -236,12 +143,10 @@ lemma rpow_nonzero (x : ENNReal) (y : ℝ) (H : ¬(x = 0 ∧ 0 < y ∨ x = ⊤ apply Hk - - - --- set_option pp.coercions false - -theorem PostPocess_pre_reduct {U : Type} [m2 : MeasurableSpace U] [count : Countable U] [disc : DiscreteMeasurableSpace U] [Inhabited U] +/-- +Jensen's inequality for privPostProcess, restructed to types where ``nq l₁`` is nonzero +-/ +theorem privPostPocess_DP_pre_reduct {U : Type} [m2 : MeasurableSpace U] [count : Countable U] [disc : DiscreteMeasurableSpace U] [Inhabited U] {nq : List T → SLang U} (f : U → V) {α : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (HNorm1 : HasSum (nq l₁) 1) @@ -540,8 +445,13 @@ theorem tsum_ne_zero_of_ne_zero {T : Type} [Inhabited T] (f : T → ENNReal) (h have B := CONTRA default contradiction --- Note: Relies on the symmetry of neighbours -theorem DPostPocess_pre {nq : List T → PMF U} (HNorm : ∀ l, HasSum (nq l) 1) +/-- +Jensen's inequality for privPostProcess. + +Implementation note: This reduction only works because neighbours are symmetric. A different reduction may +be able to relax this requirement. +-/ +theorem privPostPocess_DP_pre {nq : List T → PMF U} (HNorm : ∀ l, HasSum (nq l) 1) (f : U → V) {α : ℝ} (h1 : 1 < α) {l₁ l₂ : List T} (HN : Neighbour l₁ l₂) (Habs : AbsCts (nq l₁) (nq l₂)) (Habs' : AbsCts (nq l₂) (nq l₁)) : (∑' (x : V), @@ -579,7 +489,7 @@ theorem DPostPocess_pre {nq : List T → PMF U} (HNorm : ∀ l, HasSum (nq l) 1) cases (Classical.em (∃ x : U, ¬ nq l₁ x = 0)) · rename_i x_witness - have HR := @PostPocess_pre_reduct T V {x // ¬ nq l₁ x = 0} + have HR := @privPostPocess_DP_pre_reduct T V {x // ¬ nq l₁ x = 0} _ _ _ ?TC (fun t u => nq t u) (fun x => f x) α h1 l₁ l₂ ?GNorm1 ?GNorm2 ?Gac ?Gnz HN @@ -631,8 +541,9 @@ theorem DPostPocess_pre {nq : List T → PMF U} (HNorm : ∀ l, HasSum (nq l) 1) left linarith - - +/-- +privPostProcess satisfies the zCDP bound +-/ theorem privPostProcess_zCDPBound {nq : Mechanism T U} {ε₁ ε₂ : ℕ+} (h : zCDPBound nq ((ε₁ : ℝ) / ε₂)) (f : U → V) (Hac : ACNeighbour nq) : zCDPBound (privPostProcess nq f) ((ε₁ : ℝ) / ε₂) := by @@ -648,21 +559,11 @@ theorem privPostProcess_zCDPBound {nq : Mechanism T U} {ε₁ ε₂ : ℕ+} rw [PMF.instFunLike] simp conv => - lhs - arg 1 - arg 2 - arg 1 - arg 1 - intro x + enter [1, 1, 2, 1, 1, x] repeat rw [SLang.toPMF] simp conv => - rhs - arg 1 - arg 2 - arg 1 - arg 1 - intro x + enter [2, 1, 2, 1, 1, x] repeat rw [SLang.toPMF] repeat rw [DFunLike.coe] repeat rw [PMF.instFunLike] @@ -676,7 +577,7 @@ theorem privPostProcess_zCDPBound {nq : Mechanism T U} {ε₁ ε₂ : ℕ+} apply elog_mono_le.mp simp [PMF.bind, PMF.pure] simp [PMF.instFunLike] - apply DPostPocess_pre + apply privPostPocess_DP_pre · exact fun l => PMF.hasSum_coe_one (nq l) · exact h1 · exact h2 From 5adaaa492327735431d6641e54005e6c2f08da40 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Fri, 28 Jun 2024 09:11:47 -0400 Subject: [PATCH 121/122] cleanup Samplers --- SampCert/Samplers/GaussianGen/Properties.lean | 5 +- SampCert/Samplers/LaplaceGen/Properties.lean | 104 +----------------- 2 files changed, 10 insertions(+), 99 deletions(-) diff --git a/SampCert/Samplers/GaussianGen/Properties.lean b/SampCert/Samplers/GaussianGen/Properties.lean index 88ea7912..9db662f5 100644 --- a/SampCert/Samplers/GaussianGen/Properties.lean +++ b/SampCert/Samplers/GaussianGen/Properties.lean @@ -53,6 +53,9 @@ theorem DiscreteGaussianGenSample_apply (num : PNat) (den : PNat) (μ x : ℤ) : . simp [gauss_term_ℝ] . rw [shifted_gauss_sum_0 A] +/-- +DiscreteGaussianGen has sum 1 +-/ theorem DiscreteGaussianGen_sum (num : PNat) (den : PNat) (μ : ℤ) : HasSum (DiscreteGaussianGenSample num den μ) 1 := by rw [Summable.hasSum_iff ENNReal.summable] conv => @@ -72,7 +75,7 @@ theorem DiscreteGaussianGen_sum (num : PNat) (den : PNat) (μ : ℤ) : HasSum (D apply Hnz /-- -``SLang`` program is a proper distribution. +DiscreteGaussianGen as a PMF -/ def DiscreteGaussianGenPMF (num : PNat) (den : PNat) (μ : ℤ) : PMF ℤ := ⟨ DiscreteGaussianGenSample num den μ , DiscreteGaussianGen_sum num den μ ⟩ diff --git a/SampCert/Samplers/LaplaceGen/Properties.lean b/SampCert/Samplers/LaplaceGen/Properties.lean index 5e3a50f4..b63dda5d 100644 --- a/SampCert/Samplers/LaplaceGen/Properties.lean +++ b/SampCert/Samplers/LaplaceGen/Properties.lean @@ -29,104 +29,10 @@ theorem DiscreteLaplaceGenSample_periodic (num : PNat) (den : PNat) (μ x : ℤ) rw [tsum_eq_single (x - μ) (by aesop)] simp --- theorem DiscreteLaplaceSampleGen_normalizes (num den : PNat) (μ : ℤ) : --- ∑' x : ℤ, (DiscreteLaplaceGenSample num den μ) x = 1 := by --- conv => --- left --- right --- intro x --- rw [DiscreteLaplaceGenSample_periodic] --- rw [Int.sub_eq_add_neg] --- have X : ∀ (μ : ℤ), Summable fun x => (DiscreteLaplaceSample num den (x + μ)).toReal := by --- intro μ --- simp --- conv => --- right --- intro x --- rw [ENNReal.ofReal_mul sorrrrry] --- apply Summable.mul_left --- sorrrryyyy - --- have A : (∑' (x : ℤ), DiscreteLaplaceSample num den (x + -μ)).toReal = (1 : ENNReal).toReal := by --- rw [ENNReal.tsum_toReal_eq] --- . - --- have B := @tsum_shift (fun a : ℤ => (DiscreteLaplaceSample num den a).toReal) (- μ) X --- simp only at B --- rw [B] --- rw [← ENNReal.tsum_toReal_eq] --- . rw [DiscreteLaplaceSample_normalizes] --- . sorrrryyyy -- OK --- . sorrrryyyy -- OK - --- exact (ENNReal.toReal_eq_one_iff (∑' (x : ℤ), DiscreteLaplaceSample num den (x + -μ))).mp A - - - --- -- simp only [DiscreteLaplaceGenSample, Bind.bind, Pure.pure, bind_apply, pure_apply, mul_ite, --- -- mul_one, mul_zero] --- -- conv => --- -- left --- -- right --- -- intro x --- -- rw [tsum_eq_single (x - μ) (by aesop)] --- -- simp only [sub_add_cancel, ↓reduceIte] - --- -- simp only [DiscreteLaplaceSample, Bind.bind, Pure.pure, SLang.bind_apply] --- -- have A := DiscreteLaplaceSampleLoop_normalizes num den --- -- conv => --- -- left --- -- right --- -- intro x --- -- right --- -- intro a --- -- rw [prob_until_apply_norm _ _ _ A] --- -- simp only [ENNReal.tsum_prod'] - --- -- rw [ENNReal.tsum_comm] --- -- conv => --- -- left --- -- right --- -- intro b --- -- rw [ENNReal.tsum_comm] --- -- right --- -- intro c --- -- rw [ENNReal.tsum_mul_left] --- -- right --- -- simp - --- -- conv => --- -- left --- -- right --- -- intro b --- -- right --- -- intro x --- -- left --- -- rw [mul_comm] --- -- conv => --- -- left --- -- right --- -- intro b --- -- right --- -- intro x --- -- rw [mul_assoc] --- -- conv => --- -- left --- -- right --- -- intro b --- -- rw [ENNReal.tsum_mul_left] --- -- conv => --- -- left --- -- rw [ENNReal.tsum_mul_left] - --- -- simp only [not_and, decide_implies, decide_not, dite_eq_ite, Bool.if_true_right, --- -- Bool.decide_eq_true, Bool.or_eq_true, Bool.not_eq_true', decide_eq_false_iff_not, tsum_bool, --- -- true_or, ↓reduceIte, Bool.true_eq_false, false_or, ite_not, ite_mul, zero_mul, --- -- Bool.false_eq_true] - --- -- sorrrryyyy - +/-- +Discrete Laplace has sum 1 +-/ theorem DiscreteLaplaceGenSample_sum (num : PNat) (den : PNat) (μ : ℤ) : HasSum (DiscreteLaplaceGenSample num den μ) 1 := by rw [Summable.hasSum_iff ENNReal.summable] conv => @@ -145,8 +51,10 @@ theorem DiscreteLaplaceGenSample_sum (num : PNat) (den : PNat) (μ : ℤ) : HasS apply HZ · simp +/-- +Discrete Laplace as a PMF +-/ def DiscreteLaplaceGenSamplePMF (num : PNat) (den : PNat) (μ : ℤ) : PMF ℤ := ⟨ DiscreteLaplaceGenSample num den μ , DiscreteLaplaceGenSample_sum num den μ ⟩ - end SLang From 96560eeaab5f6b43e4ccd9e7e430502818571423 Mon Sep 17 00:00:00 2001 From: de Medeiros Date: Fri, 28 Jun 2024 09:30:46 -0400 Subject: [PATCH 122/122] cleanup Util --- .../DifferentialPrivacy/RenyiDivergence.lean | 2 +- SampCert/Util/Gaussian/DiscreteGaussian.lean | 5 +- SampCert/Util/Log.lean | 73 +++++++++++-------- 3 files changed, 46 insertions(+), 34 deletions(-) diff --git a/SampCert/DifferentialPrivacy/RenyiDivergence.lean b/SampCert/DifferentialPrivacy/RenyiDivergence.lean index 332d1a97..ee9c18dd 100644 --- a/SampCert/DifferentialPrivacy/RenyiDivergence.lean +++ b/SampCert/DifferentialPrivacy/RenyiDivergence.lean @@ -1140,6 +1140,6 @@ theorem RenyiDivergence_aux_zero [MeasurableSpace T] [MeasurableSingletonClass T symm rw [RenyiDivergence] at H have H' := RenyiDivergence_def_nonneg p q Hac Hα - refine (ofEReal_nonneg_eq_iff ?mpr.Hw H').mpr ?mpr.a + refine (ofEReal_nonneg_inj ?mpr.Hw H').mpr ?mpr.a · simp simp [H] diff --git a/SampCert/Util/Gaussian/DiscreteGaussian.lean b/SampCert/Util/Gaussian/DiscreteGaussian.lean index 17c1eff1..33623a93 100644 --- a/SampCert/Util/Gaussian/DiscreteGaussian.lean +++ b/SampCert/Util/Gaussian/DiscreteGaussian.lean @@ -337,6 +337,9 @@ theorem discrete_gaussian_normalizes {σ : ℝ} (h : σ ≠ 0) (μ : ℝ) : apply sum_gauss_term_ne_zero h +/-- +Discrete Gaussian cast to ENNReal has sum 1 +-/ def discrete_gaussian_normal {σ : ℝ} (h : σ ≠ 0) (μ : ℝ) : HasSum (fun z : ℤ => ENNReal.ofReal (discrete_gaussian σ μ z)) 1 := by rw [Summable.hasSum_iff ENNReal.summable] rw [<- ENNReal.ofReal_tsum_of_nonneg] @@ -348,7 +351,7 @@ def discrete_gaussian_normal {σ : ℝ} (h : σ ≠ 0) (μ : ℝ) : HasSum (fun simp [discrete_gaussian_normalizes h μ] /-- -The discrete Gaussian as a PMF when evaluated over ℤ +The discrete Gaussian as a PMF -/ def discrete_gaussian_pmf {σ : ℝ} (h : σ ≠ 0) (μ : ℝ) : PMF ℤ := ⟨ (fun z : ℤ => ENNReal.ofReal (discrete_gaussian σ μ z)) , discrete_gaussian_normal h μ ⟩ diff --git a/SampCert/Util/Log.lean b/SampCert/Util/Log.lean index cf523050..064db04b 100644 --- a/SampCert/Util/Log.lean +++ b/SampCert/Util/Log.lean @@ -15,6 +15,11 @@ import Lean.Elab.Tactic # Logarithm on ENNReal In this file we extend the logarithm to ``ENNReal``. + +The main definitions in this file are +- ``ofEReal : EReal -> ENNReal`` : Casting ``EReal`` to ``ENNReal`` by truncation +- ``eexp : EReal -> ENNReal`` : Exponential extended to the ``EReal``s +- ``elog : ENNReal -> EReal`` : Logarithm extended to the ``ENNReal``s -/ noncomputable section @@ -33,6 +38,23 @@ explicitly performing case analysis can be unwieldy and lead to lots of duplicat depending on which simplification rules are used. These tactics allow us to fine-tune the case splits at the start of a conversion proof in order to reduce the number of cases we must prove by hand. + + +Tactic overview: + +Real numbers: +- ``case_Real_zero``: A real number is zero or nonzero +- ``case_Real_sign``: A real number is negative, zero, or positive +- ``case_nonneg_zero``: Given ``0 ≤ r``, `r` is zero or positive +- ``case_Real_nonnegative`` : A real number is negative or nonnegative + +Extended nonnegative real numbers: +- ``case_ENNReal_isReal``: An ``ENNReal`` is ⊤, or the cast of a real number +- ``case_ENNReal_isReal_zero``: An ``ENNReal`` is ⊤, zero, or the cast of a real number + +Extended reals: +- ``case_EReal_isReal``: An ``EReal`` is ⊤, ⊥, or the cast of a real number +- ``case_EReal_isENNReal``: An `EReal`` is negative, or the cast of an ``ENNReal`` -/ @@ -189,11 +211,7 @@ lemma ofEReal_zero : ofEReal 0 = 0 := by simp [ofEReal] @[simp] lemma ofEReal_real (r : ℝ) : ofEReal r = ENNReal.ofReal r := by simp [Real.toEReal, ofEReal] --- Not sure if I want this to be an actual coercion or not, but I'm leaning towards no because --- of the truncation. --- instance : Coe EReal ENNReal := ⟨ofEReal⟩ --- MARKUSDE: formerly: ofEReal_nonpos lemma ofEReal_eq_zero_iff (w : EReal) : w ≤ 0 <-> ofEReal w = 0 := by apply Iff.intro · intro _ @@ -201,10 +219,10 @@ lemma ofEReal_eq_zero_iff (w : EReal) : w ≤ 0 <-> ofEReal w = 0 := by · intro _ case_EReal_isReal w - - --- MARKUSDE: Rename me to ofEReal_nonneg_inj -lemma ofEReal_nonneg_eq_iff {w z : EReal} (Hw : 0 <= w) (Hz : 0 <= z) : +/-- +``ofEReal`` is injective for for positive EReals +-/ +lemma ofEReal_nonneg_inj {w z : EReal} (Hw : 0 <= w) (Hz : 0 <= z) : w = z <-> (ofEReal w = ofEReal z) := by apply Iff.intro · intro _ @@ -213,35 +231,27 @@ lemma ofEReal_nonneg_eq_iff {w z : EReal} (Hw : 0 <= w) (Hz : 0 <= z) : all_goals case_EReal_isReal w all_goals case_EReal_isReal z -set_option pp.coercions false - --- MARKUSDE: Rename me @[simp] lemma toEReal_ofENNReal_nonneg {w : EReal} (H : 0 ≤ w) : ENNReal.toEReal (ofEReal w) = w := by case_EReal_isReal w --- MARKUSDE: Rename me @[simp] lemma ofEReal_toENNReal {x : ENNReal} : ofEReal (ENNReal.toEReal x) = x := by case_ENNReal_isReal x --- MARKUSDE: Rename me +/- +`ENNReal.ofReal` is the composition of cases from Real to EReal to ENNReal +-/ @[simp] lemma ofEReal_ofReal_toENNReal : ENNReal.ofEReal (Real.toEReal r) = ENNReal.ofReal r := by simp [ofEReal, Real.toEReal, ENNReal.ofReal] -lemma ofReal_injective_nonneg {r s : ℝ} (HR : 0 ≤ r) (HS : 0 ≤ s) (H : ENNReal.ofReal r = ENNReal.ofReal s) : r = s := by - apply (Real.toNNReal_eq_toNNReal_iff HR HS).mp - simp [ENNReal.ofReal] at H - assumption - - lemma ofEReal_le_mono {w z : EReal} (H : w ≤ z) : ofEReal w ≤ ofEReal z := by all_goals case_EReal_isReal w all_goals case_EReal_isReal z apply ofReal_le_ofReal assumption --- True, but unused +-- True, and provable, but unused -- lemma ofEReal_le_mono_conv_nonneg {w z : EReal} (Hw : 0 ≤ w) (Hle : ofEReal w ≤ ofEReal z) : w ≤ z := by -- all_goals case_EReal_isENNReal w -- all_goals case_EReal_isENNReal z @@ -302,6 +312,17 @@ lemma ofEReal_nonneg_scal_l {r : ℝ} {w : EReal} (H1 : 0 < r) (H2 : 0 ≤ r * w end ofEReal +/-- +``ENNReal.ofReal`` is injective for for positive EReals +-/ +lemma ofReal_injective_nonneg {r s : ℝ} (HR : 0 ≤ r) (HS : 0 ≤ s) (H : ENNReal.ofReal r = ENNReal.ofReal s) : r = s := by + apply (Real.toNNReal_eq_toNNReal_iff HR HS).mp + simp [ENNReal.ofReal] at H + assumption + + + + section elog_eexp /-- @@ -363,7 +384,6 @@ lemma eexp_ofReal {r : ℝ} : eexp r = ENNReal.ofReal (Real.exp r) := by simp [ENNReal.ofReal, eexp, elog] rfl --- MARKUSDE: Cleanup @[simp] lemma elog_eexp {x : ENNReal} : eexp (elog x) = x := by rw [elog] @@ -390,7 +410,6 @@ lemma elog_eexp {x : ENNReal} : eexp (elog x) = x := by simp rw [Hk] --- MARKUSDE: Cleanup @[simp] lemma eexp_elog {w : EReal} : (elog (eexp w)) = w := by cases w @@ -417,7 +436,6 @@ lemma eexp_elog {w : EReal} : (elog (eexp w)) = w := by simp [Real.toEReal] --- MARKUSDE: cleanup lemma elog_ENNReal_ofReal_of_pos {x : ℝ} (H : 0 < x) : (ENNReal.ofReal x).elog = x.log.toEReal := by simp [ENNReal.ofReal, ENNReal.elog, ENNReal.toEReal] rw [ite_eq_iff'] @@ -429,8 +447,6 @@ lemma elog_ENNReal_ofReal_of_pos {x : ℝ} (H : 0 < x) : (ENNReal.ofReal x).elog simp at H rw [max_eq_left_of_lt H] -set_option pp.coercions false - @[simp] lemma elog_mul {x y : ENNReal} : elog x + elog y = elog (x * y) := by all_goals case_ENNReal_isReal_zero x @@ -463,9 +479,6 @@ lemma eexp_add {w z : EReal} : eexp w * eexp z = eexp (w + z) := by rw [eexp_ofReal] --- TODO (maybe): Log of power, log and exp inverses - --- MARKUSDE: cleanup lemma eexp_injective {w z : EReal} : eexp w = eexp z -> w = z := by rw [eexp, eexp] intro H @@ -495,7 +508,6 @@ lemma eexp_injective {w z : EReal} : eexp w = eexp z -> w = z := by exact exp_eq_exp.mp H - lemma elog_injective {x y : ENNReal} : elog x = elog y -> x = y := by all_goals case_ENNReal_isReal_zero x all_goals case_ENNReal_isReal_zero y @@ -626,7 +638,6 @@ end elog_eexp section misc --- MARKUSDE: cleanup lemma mul_mul_inv_le_mul_cancel {x y : ENNReal} : (x * y⁻¹) * y ≤ x := by cases x · simp_all @@ -647,7 +658,6 @@ lemma mul_mul_inv_le_mul_cancel {x y : ENNReal} : (x * y⁻¹) * y ≤ x := by rw [mul_right_comm] rw [mul_inv_cancel_right₀ Hy' x'] --- MARKUSDE: cleanup lemma mul_mul_inv_eq_mul_cancel {x y : ENNReal} (H : y = 0 -> x = 0) (H2 : ¬(x ≠ 0 ∧ y = ⊤)) : (x * y⁻¹) * y = x := by cases x · simp_all @@ -668,7 +678,6 @@ lemma mul_mul_inv_eq_mul_cancel {x y : ENNReal} (H : y = 0 -> x = 0) (H2 : ¬(x rw [mul_right_comm] rw [mul_inv_cancel_right₀ Hy' x'] --- MARKUSDE: Cleanup lemma ereal_smul_le_left {w z : EReal} (s : EReal) (Hr1 : 0 < s) (Hr2 : s < ⊤) (H : s * w ≤ s * z) : w ≤ z := by have defTop : some none = (⊤ : EReal) := by simp [Top.top] have defBot : none = (⊥ : EReal) := by simp [Bot.bot]